Theory BicategoryOfSpans

(*  Title:       BicategoryOfSpans
    Author:      Eugene W. Stark <stark@cs.stonybrook.edu>, 2019
    Maintainer:  Eugene W. Stark <stark@cs.stonybrook.edu>
*)

section "Bicategories of Spans"

theory BicategoryOfSpans
imports Category3.ConcreteCategory IsomorphismClass CanonicalIsos EquivalenceOfBicategories
        SpanBicategory Tabulation
begin

text ‹
  In this section, we prove CKS Theorem 4, which characterizes up to equivalence the
  bicategories of spans in a category with pullbacks.
  The characterization consists of three conditions:
  BS1: ``Every 1-cell is isomorphic to a composition g ⋆ f*, where f and g are maps'';
  BS2: ``For every span of maps (f, g)› there is a 2-cell ρ› such that (f, ρ, g)›
  is a tabulation''; and
  BS3: ``Any two 2-cells between the same pair of maps are equal and invertible.''
  One direction of the proof, which is the easier direction once it is established that
  BS1 and BS3 are respected by equivalence of bicategories, shows that if a bicategory B›
  is biequivalent to the bicategory of spans in some category C› with pullbacks,
  then it satisfies BS1 -- BS3.
  The other direction, which is harder, shows that a bicategory B› satisfying BS1 -- BS3 is
  biequivalent to the bicategory of spans in a certain category with pullbacks that
  is constructed from the sub-bicategory of maps of B›.
›

  subsection "Definition"

  text ‹
    We define a \emph{bicategory of spans} to be a bicategory that satisfies the conditions
    BS1› -- BS3› stated informally above.
  ›

  locale bicategory_of_spans =
    bicategory + chosen_right_adjoints +
  assumes BS1: "ide r  f g. is_left_adjoint f  is_left_adjoint g  isomorphic r (g  f*)"
  and BS2: " is_left_adjoint f; is_left_adjoint g; src f = src g 
                       r ρ. tabulation V H 𝖺 𝗂 src trg r ρ f g"
  and BS3: " is_left_adjoint f; is_left_adjoint f'; «μ : f  f'»; «μ' : f  f'» 
                              iso μ  iso μ'  μ = μ'"

  text ‹
    Using the already-established fact equivalence_pseudofunctor.reflects_tabulation›
    that tabulations are reflected by equivalence pseudofunctors, it is not difficult to prove
    that the notion `bicategory of spans' respects equivalence of bicategories.
  ›
  
  lemma bicategory_of_spans_respects_equivalence:
  assumes "equivalent_bicategories VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD"
  and "bicategory_of_spans VC HC 𝖺C 𝗂C srcC trgC"
  shows "bicategory_of_spans VD HD 𝖺D 𝗂D srcD trgD"
  proof -
    interpret C: bicategory_of_spans VC HC 𝖺C 𝗂C srcC trgC
      using assms by simp
    interpret C: chosen_right_adjoints VC HC 𝖺C 𝗂C srcC trgC ..
    interpret D: bicategory VD HD 𝖺D 𝗂D srcD trgD
      using assms equivalent_bicategories_def equivalence_pseudofunctor.axioms(1)
            pseudofunctor.axioms(2)
      by fast
    interpret D: chosen_right_adjoints VD HD 𝖺D 𝗂D srcD trgD ..
    obtain F Φ where F: "equivalence_pseudofunctor
                           VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD F Φ"
      using assms equivalent_bicategories_def by blast
    interpret F: equivalence_pseudofunctor
                   VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD F Φ
      using F by simp
    interpret E: equivalence_of_bicategories
                   VD HD 𝖺D 𝗂D srcD trgD VC HC 𝖺C 𝗂C srcC trgC  (* 17 sec *)
                F Φ F.right_map F.right_cmp F.unit0 F.unit1 F.counit0 F.counit1
      using F.extends_to_equivalence_of_bicategories by simp
    interpret E': converse_equivalence_of_bicategories
                    VD HD 𝖺D 𝗂D srcD trgD VC HC 𝖺C 𝗂C srcC trgC
                    F Φ F.right_map F.right_cmp F.unit0 F.unit1 F.counit0 F.counit1
      ..
    interpret G: equivalence_pseudofunctor
                   VD HD 𝖺D 𝗂D srcD trgD VC HC 𝖺C 𝗂C srcC trgC
                   F.right_map F.right_cmp
      using E'.equivalence_pseudofunctor_left by simp

    write VC          (infixr "C" 55)
    write VD          (infixr "D" 55)
    write HC          (infixr "C" 53)
    write HD          (infixr "D" 53)
    write 𝖺C          ("𝖺C[_, _, _]")
    write 𝖺D          ("𝖺D[_, _, _]")
    write C.in_hhom   ("«_ : _ C _»")
    write C.in_hom    ("«_ : _ C _»")
    write D.in_hhom   ("«_ : _ D _»")
    write D.in_hom    ("«_ : _ D _»")
    write C.isomorphic (infix "C" 50)
    write D.isomorphic (infix "D" 50)
    write C.some_right_adjoint ("_*C" [1000] 1000)
    write D.some_right_adjoint ("_*D" [1000] 1000)

    show "bicategory_of_spans VD HD 𝖺D 𝗂D srcD trgD"
    proof
      show "r'. D.ide r' 
                 f' g'. D.is_left_adjoint f'  D.is_left_adjoint g'  r' D g' D (f')*D"
      proof -
        fix r'
        assume r': "D.ide r'"
        obtain f g
          where fg: "C.is_left_adjoint f  C.is_left_adjoint g  F.right_map r' C g C f*C"
          using r' C.BS1 [of "F.right_map r'"] by auto
        have trg_g: "trgC g = E.G.map0 (trgD r')"
          using fg r' C.isomorphic_implies_ide C.isomorphic_implies_hpar
          by (metis C.ideD(1) C.trg_hcomp D.ideD(1) E.G.preserves_trg)
        have trg_f: "trgC f = E.G.map0 (srcD r')"
          using fg r' C.isomorphic_implies_ide C.isomorphic_implies_hpar
          by (metis C.ideD(1) C.right_adjoint_simps(2) C.src_hcomp D.ideD(1) E.G.preserves_src)

        define d_src where "d_src  F.counit0 (srcD r')"
        define e_src where "e_src  (F.counit0 (srcD r'))~D"
        have d_src: "«d_src : F.map0 (E.G.map0 (srcD r')) D srcD r'» 
                     D.equivalence_map d_src"
          using d_src_def r' E.ε.map0_in_hhom E.ε.components_are_equivalences by simp
        have e_src: "«e_src : srcD r' D F.map0 (E.G.map0 (srcD r'))» 
                     D.equivalence_map e_src"
          using e_src_def r' E.ε.map0_in_hhom E.ε.components_are_equivalences by simp
        obtain η_src ε_src
        where eq_src: "equivalence_in_bicategory VD HD 𝖺D 𝗂D srcD trgD d_src e_src η_src ε_src"
          using d_src_def e_src_def d_src e_src D.quasi_inverses_some_quasi_inverse
                D.quasi_inverses_def
          by blast
        interpret eq_src: equivalence_in_bicategory
                            VD HD 𝖺D 𝗂D srcD trgD d_src e_src η_src ε_src
          using eq_src by simp

        define d_trg where "d_trg  F.counit0 (trgD r')"
        define e_trg where "e_trg  (F.counit0 (trgD r'))~D"
        have d_trg: "«d_trg : F.map0 (E.G.map0 (trgD r')) D trgD r'» 
                     D.equivalence_map d_trg"
          using d_trg_def r' E.ε.map0_in_hhom E.ε.components_are_equivalences by simp
        have e_trg: "«e_trg : trgD r' D F.map0 (E.G.map0 (trgD r'))» 
                     D.equivalence_map e_trg"
          using e_trg_def r' E.ε.map0_in_hhom E.ε.components_are_equivalences by simp
        obtain η_trg ε_trg
        where eq_trg: "equivalence_in_bicategory VD HD 𝖺D 𝗂D srcD trgD d_trg e_trg η_trg ε_trg"
          using d_trg_def e_trg_def d_trg e_trg D.quasi_inverses_some_quasi_inverse
                D.quasi_inverses_def
          by blast
        interpret eq_trg: equivalence_in_bicategory VD HD 𝖺D 𝗂D srcD trgD d_trg e_trg η_trg ε_trg
          using eq_trg by simp

        interpret eqs: two_equivalences_in_bicategory VD HD 𝖺D 𝗂D srcD trgD
                            d_src e_src η_src ε_src d_trg e_trg η_trg ε_trg
          ..
         
        interpret hom: subcategory VD λμ. «μ : trgD d_src D trgD d_trg»
          using D.hhom_is_subcategory by simp
        interpret hom': subcategory VD λμ. «μ : srcD d_src D srcD d_trg»
          using D.hhom_is_subcategory by simp
        interpret e: equivalence_of_categories hom.comp hom'.comp eqs.F eqs.G eqs.φ eqs.ψ
          using eqs.induces_equivalence_of_hom_categories by simp

        have r'_in_hhom: "D.in_hhom r' (srcD e_src) (srcD e_trg)"
          using r' e_src e_trg by (simp add: D.in_hhom_def)

        define g'
        where "g' = d_trg D F g"
        have g': "D.is_left_adjoint g'"
          unfolding g'_def
          using fg r' d_trg trg_g C.left_adjoint_is_ide D.equivalence_is_adjoint
                D.left_adjoints_compose F.preserves_left_adjoint C.ideD(1) D.in_hhom_def
                F.preserves_trg
          by metis
        have 1: "D.is_right_adjoint (F f*C D e_src)"
        proof -
          have "D.is_right_adjoint e_src"
            using r' e_src D.equivalence_is_adjoint by simp
          moreover have "D.is_right_adjoint (F f*C)"
            using fg C.left_adjoint_extends_to_adjoint_pair F.preserves_adjoint_pair by blast
          moreover have "srcD (F f*C) = trgD e_src"
            using fg r' trg_f C.right_adjoint_is_ide e_src by auto
          ultimately show ?thesis
            using fg r' D.right_adjoints_compose F.preserves_right_adjoint by blast
        qed
        obtain f' where f': "D.adjoint_pair f' (F f*C D e_src)"
          using 1 by auto
        have f': "D.is_left_adjoint f'  F f*C D e_src D (f')*D"
          using f' D.left_adjoint_determines_right_up_to_iso D.left_adjoint_extends_to_adjoint_pair
          by blast

        have "r' D d_trg D (e_trg D r' D d_src) D e_src"
          using r' r'_in_hhom D.isomorphic_def eqs.ψ_in_hom eqs.ψ_components_are_iso
                D.isomorphic_symmetric D.ide_char eq_src.antipar(2) eq_trg.antipar(2)
          by metis
        also have 1: "... D d_trg D F (F.right_map r') D e_src"
        proof -
          have "e_trg D r' D d_src D F (F.right_map r')"
          proof -
            have "D.in_hom (F.counit1 r')
                    (r' D d_src) (F.counit0 (trgD r') D F (F.right_map r'))"
              unfolding d_src_def
              using r' E.ε.map1_in_hom(2) [of r'] by simp
            hence "r' D d_src D F.counit0 (trgD r') D F (F.right_map r')"
              using r' D.isomorphic_def E.ε.iso_map1_ide by auto
            thus ?thesis
              using r' e_trg_def E.ε.components_are_equivalences D.isomorphic_symmetric
                    D.quasi_inverse_transpose(2)
              by (metis D.isomorphic_implies_hpar(1) F.preserves_isomorphic d_trg d_trg_def
                  eq_trg.ide_left fg)
          qed
          thus ?thesis
            using D.hcomp_ide_isomorphic D.hcomp_isomorphic_ide D.in_hhom_def
                  D.isomorphic_implies_hpar(4) d_trg e_src eq_src.antipar(1-2)
                  eq_trg.antipar(2) r'
            by force
        qed
        also have 2: "... D d_trg D (F g D F f*C) D e_src"
        proof -
          have "F (F.right_map r') D F g D F f*C"
            by (meson C.hseq_char C.ideD(1) C.isomorphic_implies_ide(2) C.left_adjoint_is_ide
                C.right_adjoint_simps(1) D.isomorphic_symmetric D.isomorphic_transitive
                F.preserves_isomorphic F.weakly_preserves_hcomp fg)
          thus ?thesis
            using D.hcomp_ide_isomorphic D.hcomp_isomorphic_ide
            by (metis 1 D.hseqE D.ideD(1) D.isomorphic_implies_hpar(2)
                eq_src.ide_right eq_trg.ide_left)
        qed
        also have 3: "... D (d_trg D F g) D F f*C D e_src"
        proof -
          have "... D d_trg D F g D F f*C D e_src"
            by (metis C.left_adjoint_is_ide C.right_adjoint_simps(1) D.hcomp_assoc_isomorphic
                D.hcomp_ide_isomorphic D.hcomp_simps(1) D.hseq_char D.ideD(1)
                D.isomorphic_implies_hpar(2) F.preserves_ide calculation eq_src.ide_right
                eq_trg.ide_left fg)
          also have "... D (d_trg D F g) D F f*C D e_src"
            by (metis C.left_adjoint_is_ide D.hcomp_assoc_isomorphic D.hcomp_simps(2)
                D.hseq_char D.ideD(1) D.isomorphic_implies_ide(1) D.isomorphic_symmetric
                F.preserves_ide calculation eq_trg.ide_left f' fg)
          finally show ?thesis by blast
        qed
        also have "... D g' D f'*D"
          using g'_def f'
          by (metis 3 D.adjoint_pair_antipar(1) D.hcomp_ide_isomorphic D.hseq_char D.ideD(1)
              D.isomorphic_implies_ide(2) g')
        finally have "D.isomorphic r' (g' D f'*D)"
          by simp
        thus "f' g'. D.is_left_adjoint f'  D.is_left_adjoint g'  r' D g' D f'*D"
          using f' g' by auto
      qed
      show "f f' μ μ'.  D.is_left_adjoint f; D.is_left_adjoint f';
                           «μ : f D f'»; «μ' : f D f'»   D.iso μ  D.iso μ'  μ = μ'"
      proof -
        fix f f' μ μ'
        assume f: "D.is_left_adjoint f"
        and f': "D.is_left_adjoint f'"
        and μ: "«μ : f D f'»"
        and μ': "«μ' : f D f'»"
        have "C.is_left_adjoint (F.right_map f)  C.is_left_adjoint (F.right_map f')"
          using f f' E.G.preserves_left_adjoint by blast
        moreover have "«F.right_map μ : F.right_map f C F.right_map f'» 
                       «F.right_map μ' : F.right_map f C F.right_map f'»"
          using μ μ' E.G.preserves_hom by simp
        ultimately have "C.iso (F.right_map μ)  C.iso (F.right_map μ') 
                         F.right_map μ = F.right_map μ'"
          using C.BS3 by blast
        thus "D.iso μ  D.iso μ'  μ = μ'"
          using μ μ' G.reflects_iso G.is_faithful by blast
      qed
      show "f g.  D.is_left_adjoint f; D.is_left_adjoint g; srcD f = srcD g 
                      r ρ. tabulation VD HD 𝖺D 𝗂D srcD trgD r ρ f g"
      proof -
        fix f g
        assume f: "D.is_left_adjoint f"
        assume g: "D.is_left_adjoint g"
        assume fg: "srcD f = srcD g"
        have "C.is_left_adjoint (F.right_map f)"
          using f E.G.preserves_left_adjoint by blast
        moreover have "C.is_left_adjoint (F.right_map g)"
          using g E.G.preserves_left_adjoint by blast
        moreover have "srcC (F.right_map f) = srcC (F.right_map g)"
          using f g D.left_adjoint_is_ide fg by simp
        ultimately have
            1: "r ρ. tabulation VC HC 𝖺C 𝗂C srcC trgC r ρ (F.right_map f) (F.right_map g)"
          using C.BS2 by simp
        obtain r ρ where
            ρ: "tabulation VC HC 𝖺C 𝗂C srcC trgC r ρ (F.right_map f) (F.right_map g)"
          using 1 by auto
        interpret ρ: tabulation VC HC 𝖺C 𝗂C srcC trgC r ρ F.right_map f F.right_map g
          using ρ by simp
        obtain r' where
          r': "D.ide r'  D.in_hhom r' (trgD f) (trgD g)  C.isomorphic (F.right_map r') r"
          using f g ρ.ide_base ρ.tab_in_hom G.locally_essentially_surjective
          by (metis D.obj_trg E.G.preserves_reflects_arr E.G.preserves_trg ρ.leg0_simps(2-3)
              ρ.leg1_simps(2,4) ρ.base_in_hom(1))
        obtain φ where φ: "«φ : r C F.right_map r'»  C.iso φ"
          using r' C.isomorphic_symmetric by blast
        have σ: "tabulation VC HC 𝖺C 𝗂C srcC trgC
                   (F.right_map r') ((φ C F.right_map f) C ρ) (F.right_map f) (F.right_map g)"
          using φ ρ.is_preserved_by_base_iso by simp
        have 1: "ρ'. «ρ' : g D HD r' f» 
                      F.right_map ρ' = F.right_cmp (r', f) C (φ C F.right_map f) C ρ"
        proof -
          have "D.ide g"
            by (simp add: D.left_adjoint_is_ide g)
          moreover have "D.ide (HD r' f)"
            using f r' D.left_adjoint_is_ide by auto
          moreover have "srcD g = srcD (HD r' f)"
            using fg by (simp add: calculation(2))
          moreover have "trgD g = trgD (HD r' f)"
            using calculation(2) r' by auto
          moreover have "«F.right_cmp (r', f) C (φ C F.right_map f) C ρ :
                            F.right_map g C F.right_map (r' D f)»"
            using f g r' φ D.left_adjoint_is_ide ρ.ide_base
            by (intro C.comp_in_homI, auto)
          ultimately show ?thesis
            using G.locally_full by simp
        qed
        obtain ρ' where ρ': "«ρ' : g D HD r' f» 
                             F.right_map ρ' = F.right_cmp (r', f) C (φ C F.right_map f) C ρ"
          using 1 by auto
        have "tabulation VD HD 𝖺D 𝗂D srcD trgD r' ρ' f g"
        proof -
          have "C.inv (F.right_cmp (r', f)) C F.right_map ρ' = (φ C F.right_map f) C ρ"
            using r' f ρ' C.comp_assoc C.comp_cod_arr C.invert_side_of_triangle(1)
            by (metis D.adjoint_pair_antipar(1) D.arrI D.in_hhomE E.G.cmp_components_are_iso
                E.G.preserves_arr)
          thus ?thesis
            using σ ρ' G.reflects_tabulation
            by (simp add: D.left_adjoint_is_ide f r')
        qed
        thus "r' ρ'. tabulation VD HD 𝖺D 𝗂D srcD trgD r' ρ' f g"
          by auto
      qed
    qed
  qed

  subsection "Span(C) is a Bicategory of Spans"

  text ‹
    We first consider an arbitrary 1-cell r› in Span(C)›, and show that it has a tabulation
    as a span of maps.  This is CKS Proposition 3 (stated more strongly to assert that
    the ``output leg'' can also be taken to be a map, which the proof shows already).
  ›

  locale identity_arrow_in_span_bicategory =  (* 20 sec *)
    span_bicategory C prj0 prj1 +
    r: identity_arrow_of_spans C r
    for C :: "'a comp"   (infixr "" 55)
    and prj0 :: "'a  'a  'a"  ("𝗉0[_, _]")
    and prj1 :: "'a  'a  'a"  ("𝗉1[_, _]")
    and r :: "'a arrow_of_spans_data"
  begin
    text ‹
      CKS say: ``Suppose r = (r0, R, r1): A → B› and put f = (1, R, r0)›, g = (1, R, r1)›.
      Let k0, k1 form a kernel pair for r0 and define ρ› by k0ρ = k1ρ = 1R.''
    ›
    abbreviation ra where "ra  r.dom.apex"
    abbreviation r0 where "r0  r.dom.leg0"
    abbreviation r1 where "r1  r.dom.leg1"
    abbreviation f where "f  mkIde ra r0"
    abbreviation g where "g  mkIde ra r1"
    abbreviation k0 where "k0  𝗉0[r0, r0]"
    abbreviation k1 where "k1  𝗉1[r0, r0]"
    text ‹
      Here ra› is the apex R› of the span (r0, R, r1)›, and the spans f› and g› also have
      that same 0-cell as their apex.  The tabulation 2-cell ρ› has to be an arrow of spans
      from g = (1, R, r1)› to r ⋆ f›, which is the span (k0, r1 ⋅ k1)›.
    ›
    abbreviation ρ where "ρ  Chn = ra r0, r0 ra,
                               Dom = Leg0 = ra, Leg1 = r1,
                               Cod = Leg0 = k0, Leg1 = r1  k1"

    lemma has_tabulation:
    shows "tabulation vcomp hcomp assoc unit src trg r ρ f g"
    and "is_left_adjoint f" and "is_left_adjoint g"
    proof -
      have ide_f: "ide f"
        using ide_mkIde r.dom.leg_in_hom(1) C.arr_dom C.dom_dom r.dom.apex_def r.dom.leg_simps(1)
        by presburger
      interpret f: identity_arrow_of_spans C f
        using ide_f ide_char' by auto
      have ide_g: "ide g"
        using ide_mkIde r.dom.leg_in_hom
        by (metis C.arr_dom C.dom_dom r.dom.leg_simps(3) r.dom.leg_simps(4))
      interpret g: identity_arrow_of_spans C g
        using ide_g ide_char' by auto

      show "is_left_adjoint f"
        using is_left_adjoint_char [of f] ide_f by simp
      show "is_left_adjoint g"
        using is_left_adjoint_char [of g] ide_g by simp
      
      have ide_r: "ide r"
        using ide_char' r.identity_arrow_of_spans_axioms by auto
      have src_r: "src r = mkObj (C.cod r0)"
        by (simp add: ide_r src_def)
      have trg_r: "trg r = mkObj (C.cod r1)"
        by (simp add: ide_r trg_def)
      have src_f: "src f = mkObj ra"
        using ide_f src_def by auto
      have trg_f: "trg f = mkObj (C.cod r0)"
        using ide_f trg_def by auto
      have src_g: "src g = mkObj ra"
        using ide_g src_def by auto
      have trg_g: "trg g = mkObj (C.cod r1)"
        using ide_g trg_def by auto

      have hseq_rf: "hseq r f"
        using ide_r ide_f src_r trg_f by simp
      interpret rf: two_composable_arrows_of_spans C prj0 prj1 r f
        using hseq_rf hseq_char by unfold_locales auto
      interpret rf: two_composable_identity_arrows_of_spans C prj0 prj1 r f ..
      interpret rf: identity_arrow_of_spans C r  f 
        using rf.ide_composite ide_char' by auto
      let ?rf = "r  f"
      (* TODO: Put this expansion into two_composable_identity_arrows_of_spans. *)
      have rf: "?rf = Chn = r0 ↓↓ r0,
                       Dom = Leg0 = k0, Leg1 = r1  k1,
                       Cod = Leg0 = k0, Leg1 = r1  k1"
        unfolding hcomp_def chine_hcomp_def
        using hseq_rf C.comp_cod_arr by auto

      interpret Cod_rf: span_in_category C Leg0 = k0, Leg1 = r1  k1
        using ide_r ide_f rf C.comp_cod_arr
        by unfold_locales auto

      have Dom_g: "Dom g = Leg0 = ra, Leg1 = r1" by simp
      interpret Dom_g: span_in_category C Leg0 = ra, Leg1 = r1
        using Dom_g g.dom.span_in_category_axioms by simp
      interpret Dom_ρ: span_in_category C Dom ρ
        using Dom_g g.dom.span_in_category_axioms by simp
      interpret Cod_ρ: span_in_category C Cod ρ
        using rf Cod_rf.span_in_category_axioms by simp
      interpret ρ: arrow_of_spans C ρ
        using Dom_ρ.apex_def Cod_ρ.apex_def C.comp_assoc C.comp_arr_dom
        by unfold_locales auto
      have ρ: "«ρ : g  r  f»"
        using rf ide_g arr_char dom_char cod_char ρ.arrow_of_spans_axioms ideD(2)
              Cod_rf.apex_def g.dom.leg_simps(4)
        by auto

      show "tabulation vcomp hcomp assoc unit src trg r ρ f g"
      proof -
        interpret T: tabulation_data vcomp hcomp assoc unit src trg r ρ f g
          using ide_f ρ by unfold_locales auto
        show ?thesis
        proof
          show T1: "u ω.
                        ide u; «ω : dom ω  r  u»  
                       w θ ν. ide w  «θ : f  w  u»  «ν : dom ω  g  w»  iso ν 
                               T.composite_cell w θ  ν = ω"
          proof -
            fix u ω
            assume u: "ide u"
            assume ω: "«ω : dom ω  r  u»"
            show "w θ ν. ide w  «θ : f  w  u»  «ν : dom ω  g  w»  iso ν 
                          T.composite_cell w θ  ν = ω"
            proof -
              interpret u: identity_arrow_of_spans C u
                using u ide_char' by auto
              have v: "ide (dom ω)"
                using ω by auto
              interpret v: identity_arrow_of_spans C dom ω
                using v ide_char' by auto
              interpret ω: arrow_of_spans C ω
                using ω arr_char by auto
              have hseq_ru: "hseq r u"
                using u ω ide_cod by fastforce
              interpret ru: two_composable_arrows_of_spans C prj0 prj1 r u
                using hseq_ru hseq_char by unfold_locales auto
              interpret ru: two_composable_identity_arrows_of_spans C prj0 prj1 r u ..
              text ‹
                CKS say:
                ``We must show that (f, ρ, g)› is a wide tabulation of r›.
                Take u = (u0, U, u1): X → A›, v = (v0, V, v1): X → B›,
                ω: v ⇒ ru› as in T1›.  Let P› be the pullback of u1, r0.
                Let w = (v0, V, p1ω): X → R›, θ = p0ω: fw ⇒ u›,
                ν = 1: v ⇒ gw›; so ω = (rθ)(ρw)ν› as required.''
              ›
              let ?R = "r.apex"
              let ?A = "C.cod r0"
              let ?B = "C.cod r1"
              let ?U = "u.apex"
              let ?u0 = "u.leg0"
              let ?u1 = "u.leg1"
              let ?X = "C.cod ?u0"
              let ?V = "v.apex"
              let ?v0 = "v.leg0"
              let ?v1 = "v.leg1"
              let  = "ω.chine"
              let ?P = "r0 ↓↓ ?u1"
              let ?p0 = "𝗉0[r0, ?u1]"
              let ?p1 = "𝗉1[r0, ?u1]"
              let ?w1 = "?p1  "
              define w where "w = mkIde ?v0 ?w1"
              let ?Q = "?R ↓↓ ?w1"
              let ?q1 = "𝗉1[?R, ?w1]"
              let  = "?R r0, r0 ?R"

              have P: "?P = ru.apex"
                using ru.apex_composite by auto

              have Chn_ω: "« : ?V C ?P»"
                using P Chn_in_hom ω by force
              have p0ω: "«?p0   : ?V C ?U»"
                using Chn_ω ru.legs_form_cospan by auto
              have w1: "«?w1 : ?V C ?R»"
                 using Chn_ω ru.legs_form_cospan r.dom.apex_def by blast
              have r1w1: "r1  ?w1 = ?v1"
                by (metis C.comp_assoc T.base_simps(3) ω ω.leg1_commutes
                    arrow_of_spans_data.select_convs(3) cod_char dom_char ideD(1) ideD(2)
                    in_homE ru.composite_in_hom ru.leg1_composite u v)

              have w: "ide w"
                unfolding w_def
                using P ω w1 by (intro ide_mkIde, auto) 
              interpret w: identity_arrow_of_spans C w
                using w_def w ide_char' by auto

              have hseq_fw: "hseq f w"
                using w_def ide_f w src_def trg_def w1 by auto
              interpret fw: two_composable_arrows_of_spans C prj0 prj1 f w
                using w_def hseq_fw hseq_char by unfold_locales auto
              interpret fw: two_composable_identity_arrows_of_spans C prj0 prj1 f w ..

              have hseq_gw: "hseq g w"
                using w_def ide_g w src_def trg_def w1 by auto
              interpret gw: two_composable_arrows_of_spans C prj0 prj1 g w
                using hseq_gw hseq_char by unfold_locales auto
              interpret gw: two_composable_identity_arrows_of_spans C prj0 prj1 g w ..

              interpret rfw: three_composable_arrows_of_spans C prj0 prj1 r f w ..
              interpret rfw: three_composable_identity_arrows_of_spans C prj0 prj1 r f w ..
              have arfw: "«𝖺[r, f, w] : (r  f)  w  r  f  w»"
                using fw.composable ide_f ide_r w rf.composable by auto

              have fw_apex_eq: "fw.apex = ?R ↓↓ ?w1"
                using w_def fw.dom.apex_def hcomp_def hseq_fw hseq_char ω C.arr_dom_iff_arr
                      C.pbdom_def fw.chine_eq_apex fw.chine_simps(1)
                by auto
              have gw_apex_eq: "gw.apex = ?R ↓↓ ?w1"
                using w_def ω w1 gw.dom.apex_def hcomp_def hseq_gw hseq_char by auto
              text ‹
                Well, CKS say take θ = p0ω›, but taking this literally and trying to define
                θ› so that Chn θ = ?p0 ⋅ ?ω›, does not yield the required dom θ = ?R ↓↓ ?w1›.
                We need «Chn θ : ?R ↓↓ ?w1 →C ?U»›, so we have to compose with a
                projection, which leads to:  Chn θ = ?p0 ⋅ ?ω ⋅ 𝗉0[?R, ?w1]›.
                What CKS say is only correct if the projection 𝗉0[?R, ?w1]› is an identity,
                which can't be guaranteed for an arbitrary choice of pullbacks.
              ›
              define θ
                where
                  "θ = Chn = ?p0    𝗉0[?R, ?w1], Dom = Dom (f  w), Cod = Cod u"

              interpret Dom_θ: span_in_category C Dom θ
                using θ_def fw.dom.span_in_category_axioms by simp
              interpret Cod_θ: span_in_category C Cod θ
                using θ_def u.cod.span_in_category_axioms by simp

              have Dom_θ_leg0_eq: "Dom_θ.leg0 = ?v0  𝗉0[?R, ?w1]"
                using w_def θ_def hcomp_def hseq_fw hseq_char by simp
              have Dom_θ_leg1_eq: "Dom_θ.leg1 = r0  ?q1"
                using w_def θ_def hcomp_def hseq_fw hseq_char by simp
              have Cod_θ_leg0_eq: "Cod_θ.leg0 = ?u0"
                using w_def θ_def hcomp_def hseq_fw hseq_char by simp
              have Cod_θ_leg1_eq: "Cod_θ.leg1 = ?u1"
                using w_def θ_def hcomp_def hseq_fw hseq_char by simp
              have Chn_θ_eq: "Chn θ = ?p0    𝗉0[?R, ?w1]"
                using θ_def by simp

              interpret θ: arrow_of_spans C θ
              proof
                show 1: "«Chn θ : Dom_θ.apex C Cod_θ.apex»"
                  using θ_def Chn_ω ru.legs_form_cospan fw_apex_eq
                  by (intro C.in_homI, auto)
                show "Cod_θ.leg0  Chn θ = Dom_θ.leg0"
                proof -
                  have "Cod_θ.leg0  Chn θ = ?u0  ?p0    𝗉0[?R, ?w1]"
                    using Cod_θ_leg0_eq Chn_θ_eq by simp
                  also have "... = ?v0  𝗉0[?R, ?w1]"
                  proof -
                    have "?u0  ?p0    𝗉0[?R, ?w1] = (?u0  ?p0  )  𝗉0[?R, ?w1]"
                      using C.comp_assoc by simp
                    also have "... = ?v0  𝗉0[?R, ?w1]"
                    proof -
                      have "?u0  ?p0   = (?u0  ?p0)  "
                        using C.comp_assoc by simp
                      also have "... = ω.cod.leg0  "
                        by (metis ω arrow_of_spans_data.select_convs(2) cod_char in_homE
                            ru.leg0_composite)
                      also have "... = ω.dom.leg0"
                        using ω.leg0_commutes by simp
                      also have "... = ?v0"
                        using ω dom_char by auto
                      finally show ?thesis by simp
                    qed
                    finally show ?thesis by simp
                  qed
                  also have "... = Dom_θ.leg0"
                    using Dom_θ_leg0_eq by simp
                  finally show ?thesis by blast
                qed
                show "Cod_θ.leg1  Chn θ = Dom_θ.leg1"
                proof -
                  have "Cod_θ.leg1  Chn θ = ?u1  ?p0    𝗉0[?R, ?w1]"
                    using Cod_θ_leg1_eq Chn_θ_eq by simp
                  also have "... = r0  ?q1"
                  proof -
                    have "?u1  ?p0    𝗉0[?R, ?w1] = ((?u1  ?p0)  )  𝗉0[?R, ?w1]"
                      using C.comp_assoc by fastforce
                    also have "... = ((r0  ?p1)  )  𝗉0[?R, ?w1]"
                      using C.pullback_commutes' ru.legs_form_cospan by simp
                    also have "... = r0  ?w1  𝗉0[?R, ?w1]"
                      using C.comp_assoc by fastforce
                    also have "... = r0  ?R  ?q1"
                      using ω C.in_homE C.pullback_commutes' w1 by auto
                    also have "... = r0  ?q1"
                      using ω w1 C.comp_cod_arr by auto
                    finally show ?thesis by simp
                  qed
                  also have "... = Dom_θ.leg1"
                    using Dom_θ_leg1_eq by simp
                  finally show ?thesis by blast
                qed
              qed
              text ‹
                Similarly, CKS say to take ν = 1: v ⇒ gw›, but obviously this can't be
                interpreted literally, either, because v.apex› and gw.apex› are not equal.
                Instead, we have to define ν› so that Chn ν = C.inv 𝗉0[?R, ?w1]›,
                noting that 𝗉0[?R, ?w1]› is the pullback of an identity,
                hence is an isomorphism. Then ?v0 ⋅ 𝗉0[?R, ?w1] ⋅ Chn ν = ?v0› and
                ?v1 ⋅ 𝗉1[?R, ?w1] ⋅ Chn ν = ?v1 ⋅ ?w1›, showing that ν› is an arrow
                of spans.
              ›
              let ?ν' = "𝗉0[?R, ?w1]"
              define ν
                where
                  "ν = Chn = C.inv ?ν', Dom = Dom (dom ω), Cod = Cod (g  w)"
              have iso_ν: "C.inverse_arrows ?ν' (Chn ν)"
                using ν_def ω w1 C.iso_pullback_ide
                by (metis C.inv_is_inverse C.seqE arrow_of_spans_data.select_convs(1)
                    r.chine_eq_apex r.chine_simps(1) r.chine_simps(3) r.cod_simps(1)
                    r.dom.apex_def r.dom.ide_apex r.dom.is_span r1w1 v.dom.leg_simps(3))
              text ‹
$$
\xymatrix{
  && X \\
  && V \ar[u]_{v_0} \ar[d]_{\omega}  \ar@/ul50pt/[ddddll]_{v_1} \ar@/l30pt/[dd]_<>(0.7){w_1}\\
  & R\downarrow\downarrow w_1 \ar[ur]^{\nu'} \ar[dd]_{q_1}
  & r_0\downarrow\downarrow u_1 \ar[d]_{p_1} \ar@/dl10pt/[drr]_<>(0.4){p_0}
  & R\downarrow\downarrow w_1 \ar[ul]_{\nu'} \ar[dd]^<>(0.7){q_1} \ar@ {.>}[dr]_{\theta}\\
  && R && U \ar@/ur20pt/[uuull]_{u_0} \ar[dd]^{u_1} \\
  & R \ar[dl]_{r_1} \ar@ {<->}[ur]_{R} \ar@ {.>}[dr]^{\rho} \ar@/dl5pt/[ddr]_<>(0.4){R}
  && R \ar@ {<->}[ul]^{R} \ar[dr]^{r_0} \ar[ur]_{r_1}\\
  B && r_0\downarrow\downarrow r_0 \ar[uu]_{k_0} \ar[d]^{k_1} \ar[uu] \ar[ur]_{k_0} && A \\
  & & R \ar[ull]^{r_1} \ar[urr]_{r_0} \\
}
$$
              ›
              have w1_eq: "?w1 = ?q1  C.inv ?ν'"
              proof -
                have "?R  ?q1 = ?w1  ?ν'"
                  using iso_ν ω w1 C.pullback_commutes [of ?R ?w1] by blast
                hence "?q1 = ?w1  ?ν'"
                  using ω w1 C.comp_cod_arr by auto
                thus ?thesis
                  using iso_ν C.invert_side_of_triangle(2)
                  by (metis C.isoI C.prj1_simps(1) arrow_of_spans_data.select_convs(3)
                      fw.legs_form_cospan(2) span_data.simps(1-2) w_def)
              qed

              interpret Dom_ν: span_in_category C Dom ν
                using ν_def v.dom.span_in_category_axioms by simp
              interpret Cod_ν: span_in_category C Cod ν
                using ν_def gw.cod.span_in_category_axioms by simp
              interpret ν: arrow_of_spans C ν
              proof
                show 1: "«Chn ν : Dom_ν.apex C Cod_ν.apex»"
                proof
                  show "C.arr (Chn ν)"
                    using iso_ν by auto
                  show "C.dom (Chn ν) = Dom_ν.apex"
                    using ν_def iso_ν w1 gw_apex_eq by fastforce
                  show "C.cod (Chn ν) = Cod_ν.apex"
                    using ν_def iso_ν gw_apex_eq C.comp_inv_arr C.pbdom_def
                    by (metis C.cod_comp arrow_of_spans_data.select_convs(3)
                        gw.apex_composite gw.chine_composite gw.chine_simps(1,3))
                qed
                show "Cod_ν.leg0  Chn ν = Dom_ν.leg0"
                  using w_def ν_def 1 iso_ν hcomp_def hseq_gw C.comp_arr_inv
                        C.comp_assoc v.leg0_commutes
                  by auto
                show "Cod_ν.leg1  Chn ν = Dom_ν.leg1"
                  using w_def ν_def hcomp_def hseq_gw C.comp_assoc w1_eq r1w1
                  by auto
              qed
              text ‹
                Now we can proceed to establishing the conclusions of T1›.
              ›
              have "ide w  «θ : f  w  u»  «ν : dom ω  dom ρ  w»  iso ν 
                    T.composite_cell w θ  ν = ω"
              proof (intro conjI)
                show ide_w: "ide w"
                  using w by blast
                show θ: "«θ : f  w  u»"
                  using θ_def θ.arrow_of_spans_axioms arr_char dom_char cod_char hseq_fw hseq_char
                        hcomp_def fw.chine_eq_apex
                  by auto
                show ν: "«ν : dom ω  dom ρ  w»"
                proof -
                  have "«ν : dom ω  g  w»"
                    using ν_def ω ν.arrow_of_spans_axioms arr_char dom_char cod_char hseq_gw
                          hseq_char hcomp_def gw.chine_eq_apex
                    by auto
                  thus ?thesis
                    using T.tab_in_hom by simp
                qed
                show "iso ν"
                  using iso_ν iso_char arr_char ν.arrow_of_spans_axioms by auto
                show "T.composite_cell w θ  ν = ω"
                proof (intro arr_eqI)
                  have ρw: "«ρ  w : g  w  (r  f)  w»"
                    using w_def ρ ide_w hseq_rf hseq_fw hseq_gw by auto
                  have : "«r  θ : r  f  w  r  u»"
                    using arfw ide_r θ fw.composite_simps(2) rf.composable by auto 
                  have 1: "«T.composite_cell w θ  ν : dom ω  r  u»"
                    using ν ρw arfw  by auto
                  show 3: "par (T.composite_cell w θ  ν) ω"
                    using 1 ω by (elim in_homE, auto)
                  show "Chn (T.composite_cell w θ  ν) = "
                  proof -
                    have 2: "Chn (T.composite_cell w θ  ν) =
                             Chn (r  θ)  Chn 𝖺[r, f, w]  Chn (ρ  w)  Chn ν"
                      using 1 3 Chn_vcomp C.comp_assoc
                      by (metis (full_types) seqE)
                    also have "... = "
                    proof -
                      let ?LHS = "Chn (r  θ)  Chn 𝖺[r, f, w]  Chn (ρ  w)  Chn ν"
                      have Chn_rθ: "Chn (r  θ) = r.chine  𝗉1[r0, r0  ?q1]
                                                     r0, ?u1
                                                   θ.chine  𝗉0[r0, r0  ?q1]"
                        using  hcomp_def θ_def chine_hcomp_def Dom_θ_leg1_eq
                        by (metis arrI arrow_of_spans_data.select_convs(1,3)
                            hseq_char r.cod_simps(2) u.cod_simps(3))
                      have Chn_arfw: "Chn 𝖺[r, f, w] = rfw.chine_assoc"
                        using α_ide ide_f rf.composable fw.composable w by auto
                      have Chn_ρw: "Chn (ρ  w) =   ?q1 k0, ?w1 ?ν'"
                      proof -
                        have "Chn (ρ  w) =
                              chine_hcomp
                                Chn = ,
                                 Dom = Leg0 = ?R, Leg1 = r1,
                                 Cod = Leg0 = k0, Leg1 = r1  k1
                                Chn = v.apex,
                                 Dom = Leg0 = ?v0, Leg1 = ?w1,
                                 Cod = Leg0 = ?v0, Leg1 = ?w1"
                          using ρ ide_w hseq_rf hseq_char hcomp_def src_def trg_def
                          by (metis (no_types, lifting) ρw arrI arrow_of_spans_data.select_convs(1)
                              v.dom.apex_def w_def)
                        also have "... =   ?q1 k0, ?w1 ?V  ?ν'"
                          unfolding chine_hcomp_def by simp
                        also have "... =   ?q1 k0, ?w1 ?ν'"
                        proof -
                          have "?V  ?ν' = ?ν'"
                            using C.comp_ide_arr v.dom.ide_apex ρ w1 by auto
                          thus ?thesis by simp
                        qed
                        finally show ?thesis by blast
                      qed

                      have 3: "C.seq ?p1 "
                        using w1 by blast
                      moreover have 4: "C.seq ?p1 ?LHS"
                      proof
                        show "«?LHS : v.apex C ru.apex»"
                          by (metis (no_types, lifting) 1 2 Chn_in_hom ru.chine_eq_apex
                              v.chine_eq_apex)
                        show "«?p1 : ru.apex C ?R»"
                          using P C.prj1_in_hom ru.legs_form_cospan by fastforce
                      qed
                      moreover have "?p0  ?LHS = ?p0  "
                      proof -
                        have "?p0  ?LHS =
                              (?p0  Chn (r  θ))  Chn 𝖺[r, f, w]  Chn (ρ  w)  Chn ν"
                          using C.comp_assoc by simp
                        also have "... = (θ.chine  𝗉0[r0, r0  ?q1]) 
                                           Chn 𝖺[r, f, w]  Chn (ρ  w)  Chn ν"
                        proof -
                          have "?p0  Chn (r  θ) = θ.chine  𝗉0[r0, r0  ?q1]"
                            by (metis C.prj_tuple(1) Chn_rθ θ_def arrI Dom_θ_leg1_eq
                                arrow_of_spans_data.select_convs(3) chine_hcomp_props(2)
                                hseq_char r.cod_simps(2)  u.cod_simps(3))
                          thus ?thesis by argo
                        qed
                        also have
                          "... = ?p0    (rfw.Prj00  Chn 𝖺[r, f, w])  Chn (ρ  w)  Chn ν"
                          using w_def θ_def C.comp_assoc by simp
                        also have "... = ?p0    (rfw.Prj0   Chn (ρ  w))  Chn ν"
                          using Chn_arfw rfw.prj_chine_assoc C.comp_assoc by simp
                        also have "... = ?p0    ?ν'  Chn ν"
                        proof -
                          have "rfw.Prj0  Chn (ρ  w) = 𝗉0[k0, ?w1]    ?q1 k0, ?w1 ?ν'"
                            using w_def Chn_ρw C.comp_cod_arr by simp
                          also have "... = ?ν'"
                            by (metis (no_types, lifting) C.not_arr_null C.prj_tuple(1) C.seqE
                                C.tuple_is_extensional Chn_ρw 4)
                          finally have "rfw.Prj0  Chn (ρ  w) = ?ν'"
                            by blast
                          thus ?thesis by simp
                        qed
                        also have "... = ?p0  "
                          using iso_ν C.comp_arr_dom
                          by (metis (no_types, lifting) C.comp_arr_inv C.dom_comp ν_def
                              ω.chine_simps(1) 3 arrow_of_spans_data.simps(1) w1_eq)
                        finally show ?thesis by blast
                      qed
                      moreover have "?p1  ?LHS = ?w1"
                      proof -
                        have "?p1  ?LHS =
                              (?p1  Chn (r  θ))  Chn 𝖺[r, f, w]  Chn (ρ  w)  Chn ν"
                          using C.comp_assoc by simp
                        also have "... = (r.chine  𝗉1[r0, r0  ?q1])  Chn 𝖺[r, f, w] 
                                           Chn (ρ  w)  Chn ν"
                          by (metis (no_types, lifting) C.not_arr_null C.prj_tuple(2) C.seqE
                              C.tuple_is_extensional Chn_rθ 4)
                        also have "... = r.chine  (rfw.Prj1  Chn 𝖺[r, f, w])  Chn (ρ  w)  Chn ν"
                          using w_def Dom_θ_leg1_eq C.comp_assoc by simp
                        also have "... = r.chine  (rfw.Prj11  Chn (ρ  w))  Chn ν"
                          using Chn_arfw rfw.prj_chine_assoc(1) C.comp_assoc by simp
                        also have "... = r.chine  ?q1  Chn ν"
                        proof -
                          have "rfw.Prj11  Chn (ρ  w) =
                                 (k1  𝗉1[k0, ?w1])    ?q1 k0, ?w1 ?ν'"
                            using w_def Chn_ρw C.comp_cod_arr by simp
                          also have "... = k1  𝗉1[k0, ?w1]    ?q1 k0, ?w1 ?ν'"
                            using C.comp_assoc by simp
                          also have "... = k1    ?q1"
                            by (metis (no_types, lifting) C.not_arr_null C.prj_tuple(2)
                                C.seqE C.tuple_is_extensional Chn_ρw 4)
                          also have "... = (k1  )  ?q1"
                            using C.comp_assoc by presburger
                          also have "... = ?R  ?q1"
                            by simp
                          also have "... = ?q1"
                            by (metis Dom_θ_leg1_eq C.comp_ide_arr C.prj1_simps(3)
                                C.prj1_simps_arr C.seqE C.seqI Dom_θ.leg_simps(3)
                                r.dom.ide_apex)
                          finally have "rfw.Prj11  Chn (ρ  w) = ?q1"
                            by blast
                          thus ?thesis by simp
                        qed
                        also have "... = (r.chine  ?p1)  "
                          using ν_def w1_eq C.comp_assoc by simp
                        also have "... = ?w1"
                          using C.comp_cod_arr r.chine_eq_apex ru.prj_simps(1) by auto
                        finally show ?thesis by blast
                      qed
                      ultimately show ?thesis
                        using ru.legs_form_cospan C.prj_joint_monic by blast
                    qed
                    finally show ?thesis by argo
                  qed
                qed
              qed
              thus ?thesis
                using w_def by auto
            qed
          qed

          show T2: "u w w' θ θ' β.
                        ide w; ide w';
                         «θ : f  w  u»; «θ' : f  w'  u»; «β : g  w  g  w'»;
                         T.composite_cell w θ = T.composite_cell w' θ'  β  
                       ∃!γ. «γ : w  w'»  β = g  γ  θ = θ'  (f  γ)"
          proof -
            fix u w w' θ θ' β
            assume ide_w: "ide w"
            assume ide_w': "ide w'"
            assume θ: "«θ : f  w  u»"
            assume θ': "«θ' : f  w'  u»"
            assume β: "«β : g  w  g  w'»"
            assume E: "T.composite_cell w θ = T.composite_cell w' θ'  β"
            interpret T: uwθw'θ'β vcomp hcomp assoc unit src trg r ρ f g u w θ w' θ' β
              using ide_w ide_w' θ θ' β E comp_assoc
              by unfold_locales auto

            show "∃!γ. «γ : w  w'»  β = g  γ  θ = θ'  (f  γ)"
            proof
              interpret u: identity_arrow_of_spans C u
                using T.uwθ.u_simps(1) ide_char' by auto
              interpret w: identity_arrow_of_spans C w
                using ide_w ide_char' by auto
              interpret w': identity_arrow_of_spans C w'
                using ide_w' ide_char' by auto
              let ?u0 = u.leg0
              let ?u1 = u.leg1
              let ?w0 = w.leg0
              let ?w1 = w.leg1
              let ?wa = "w.apex"
              let ?w0' = w'.leg0
              let ?w1' = w'.leg1
              let ?wa' = "w'.apex"
              let ?R = ra
              let ?p0 = "𝗉0[?R, ?w1]"
              let ?p0' = "𝗉0[?R, ?w1']"
              let ?p1 = "𝗉1[?R, ?w1]"
              let ?p1' = "𝗉1[?R, ?w1']"

              interpret fw: two_composable_identity_arrows_of_spans C prj0 prj1 f w
                using hseq_char by unfold_locales auto
              interpret fw': two_composable_identity_arrows_of_spans C prj0 prj1 f w'
                using hseq_char by unfold_locales auto

              have hseq_gw: "hseq g w"
                using T.leg1_in_hom by auto
              interpret gw: two_composable_identity_arrows_of_spans C prj0 prj1 g w
                using hseq_gw hseq_char by unfold_locales auto

              have hseq_gw': "hseq g w'"
                using T.leg1_in_hom by auto
              interpret gw': two_composable_identity_arrows_of_spans C prj0 prj1 g w'
                using hseq_gw' hseq_char by unfold_locales auto

              interpret rfw: three_composable_identity_arrows_of_spans C prj0 prj1 r f w ..
              interpret rfw: identity_arrow_of_spans C r  f  w
                using rfw.composites_are_identities ide_char' by auto
              interpret rfw': three_composable_arrows_of_spans C prj0 prj1 r f w' ..
              interpret rfw': three_composable_identity_arrows_of_spans C prj0 prj1 r f w' ..
              interpret rfw': identity_arrow_of_spans C r  f  w'
                using rfw'.composites_are_identities ide_char' by auto

              have ρw: "«ρ  w : g  w  (r  f)  w»"
                using ρ hseq_gw by blast
              interpret ρw: two_composable_arrows_of_spans C prj0 prj1 ρ w
                using ρw by unfold_locales auto
              have ρw': "«ρ  w' : g  w'  (r  f)  w'»"
                using ρ hseq_gw' by blast
              interpret ρw': two_composable_arrows_of_spans C prj0 prj1 ρ w'
                using ρw' by unfold_locales auto

              have arfw: "«𝖺[r, f, w] : (r  f)  w  r  f  w»"
                using fw.composable ide_f ide_r ide_w rf.composable by auto
              have arfw': "«𝖺[r, f, w'] : (r  f)  w'  r  f  w'»"
                using fw'.composable ide_f ide_r ide_w' rf.composable by auto

              have : "«r  θ : r  f  w  r  u»"
                by fastforce
              interpret Dom_θ: span_in_category C Dom θ
                using fw.dom.span_in_category_axioms
                by (metis θ arrow_of_spans_data.select_convs(2) in_homE dom_char)
              interpret Cod_θ: span_in_category C Cod θ
                using θ u.cod.span_in_category_axioms cod_char by auto
              interpret θ: arrow_of_spans C θ
                using arr_char T.uwθ.θ_simps(1) by auto
              interpret: two_composable_arrows_of_spans C prj0 prj1 r θ
                using  by unfold_locales auto

              have rθ': "«r  θ' : r  f  w'  r  u»"
                by fastforce
              interpret Dom_θ': span_in_category C Dom θ'
                using fw'.dom.span_in_category_axioms
                by (metis θ' arrow_of_spans_data.select_convs(2) in_homE dom_char)
              interpret Cod_θ': span_in_category C Cod θ'
                using θ' u.cod.span_in_category_axioms cod_char by auto
              interpret θ': arrow_of_spans C θ'
                using arr_char T.uw'θ'.θ_simps(1) by auto
              interpret rθ': two_composable_arrows_of_spans C prj0 prj1 r θ'
                using rθ' by unfold_locales auto

              have 7: "«T.composite_cell w' θ'  β : g  w  r  u»"
                using β ρw' arfw' rθ' by auto
              have 8: "«T.composite_cell w θ : g  w  r  u»"
                using ρw arfw  by auto

              interpret ru: two_composable_identity_arrows_of_spans C prj0 prj1 r u
                using hseq_char by unfold_locales auto

              interpret Dom_β: span_in_category C Dom β
                using β fw.dom.span_in_category_axioms arr_char
                by (metis comp_arr_dom in_homE gw.cod.span_in_category_axioms seq_char)
              interpret Cod_β: span_in_category C Cod β
                using β fw.cod.span_in_category_axioms arr_char
                by (metis (no_types, lifting) comp_arr_dom ideD(2) in_homI
                    gw'.cod.span_in_category_axioms gw'.chine_is_identity hseq_gw' seqI'
                    seq_char ide_char)
              interpret β: arrow_of_spans C β
                using β arr_char by auto
              text ‹
                CKS say: ``Take u›, w›, w'›, θ›, θ'› as in T2› and note that
                fw = (w0, W, r0 w1)›, gw = (w0, W, r1 w1)›, \emph{etc}.
                So β: W → W'› satisfies w0 = w0' β›, r1 w1 = r1 w1' β›.
                But the equation (rθ)(ρw) = (rθ')(ρw')β› gives w1 = w1'›.
                So γ = β : w ⇒ w'› is unique with β = g γ, θ = θ' (f γ).›''

                Once again, there is substantial punning in the proof sketch given by CKS.
                We can express fw› and gw› almost in the form they indicate, but projections
                are required.
              ›
              have cospan: "C.cospan ?R ?w1"
                using hseq_char [of ρ w] src_def trg_def by auto
              have cospan': "C.cospan ?R ?w1'"
                using hseq_char [of ρ w'] src_def trg_def by auto
              have fw: "f  w = Chn = ?R ↓↓ ?w1,
                                 Dom = Leg0 = ?w0  ?p0, Leg1 = r0  ?p1,
                                 Cod = Leg0 = ?w0  ?p0, Leg1 = r0  ?p1"
                using ide_f hseq_char hcomp_def chine_hcomp_def fw.dom.apex_def cospan
                      fw.chine_eq_apex
                by auto
              have gw: "g  w = Chn = ?R ↓↓ ?w1,
                                 Dom = Leg0 = ?w0  ?p0, Leg1 = r1  ?p1,
                                 Cod = Leg0 = ?w0  ?p0, Leg1 = r1  ?p1"
                using hseq_gw hseq_char hcomp_def chine_hcomp_def gw.dom.apex_def cospan
                      gw.chine_eq_apex
                by auto
              have fw': "f  w' = Chn = ?R ↓↓ ?w1',
                                   Dom = Leg0 = ?w0'  ?p0', Leg1 = r0  ?p1',
                                   Cod = Leg0 = ?w0'  ?p0', Leg1 = r0  ?p1'"
                using ide_f hseq_char hcomp_def chine_hcomp_def fw'.dom.apex_def cospan'
                      fw'.chine_eq_apex
                by auto
              have gw': "g  w' = Chn = ?R ↓↓ ?w1',
                                   Dom = Leg0 = ?w0'  ?p0', Leg1 = r1  ?p1',
                                   Cod = Leg0 = ?w0'  ?p0', Leg1 = r1  ?p1'"
                using hseq_gw' hseq_char hcomp_def chine_hcomp_def gw'.dom.apex_def cospan'
                      gw'.chine_eq_apex
                by auto
              text ‹
                Note that ?p0› and ?p0'› are only isomorphisms, not identities,
                and we have ?p1› (which equals ?w1 ⋅ ?p0›) and ?p1'› (which equals ?w1' ⋅ ?p0'›)
                in place of ?w1› and ?w1'›.
              ›
              text ‹
                The following diagram summarizes the
                various given and defined arrows involved in the proof.
                We have deviated slightly here from the nomenclature used in in CKS.
                We prefer to use W› and W'› to denote the apexes of
                w› and w'›, respectively.
                We already have the expressions ?R ↓↓ ?w1› and ?R ↓↓ ?w1'› for the
                apexes of fw› and fw'› (which are the same as the apexes of
                gw› and gw'›, respectively) and we will not use any abbreviation for them.
              ›
              text ‹
$$
\xymatrix{
  &&& X \\
  && W \ar[ur]^{w_0} \ar[dr]_{w_1} \ar@ {.>}[rr]^{\gamma}
  && W' \ar[ul]_{w_0'} \ar[dl]^{w_1'} && U \ar@/r10pt/[dddl]^{u_1} \ar@/u7pt/[ulll]_{u_0}\\
  & R\downarrow\downarrow w_1 \ar[ur]_{p_0} \ar[dr]^{p_1} \ar@/d15pt/[rrrr]_{\beta}
    \ar@/u100pt/[urrrrr]^{\theta}
  && R && R \downarrow\downarrow w_1' \ar[ul]^{p_0'} \ar[dl]^{p_1'} \ar[ur]_{\theta'} \\
  && R \ar@ {.>}[dr]_{\rho} \ar@/dl7pt/[ddr]_{R} \ar[ur]_{R} \ar[dl]_{r_1} \ar@ {<->}[rr]_{R}
  && R \ar[ul]^{R} \ar[dr]_{r_0} \\
  & B && r_0 \downarrow\downarrow r_0 \ar[d]^{k_1} \ar[ur]_{k_0} && A \\
  &&& R \ar@/dr10pt/[urr]_{r_0} \ar@/dl5pt/[ull]^{r_1}
}
$$
              ›
              have Chn_β: "«β.chine: ?R ↓↓ ?w1 C ?R ↓↓ ?w1'»"
                using gw gw' Chn_in_hom β gw'.chine_eq_apex gw.chine_eq_apex by force
              have β_eq: "β = Chn = β.chine,
                               Dom = Leg0 = ?w0  ?p0, Leg1 = r1  ?p1,
                               Cod = Leg0 = ?w0'  ?p0', Leg1 = r1  ?p1'"
                using β gw gw' dom_char cod_char by auto
              have Dom_β_eq: "Dom β = Leg0 = ?w0  ?p0, Leg1 = r1  ?p1"
                using β gw gw' dom_char cod_char by auto
              have Cod_β_eq: "Cod β = Leg0 = ?w0'  ?p0', Leg1 = r1  ?p1'"
                using β gw gw' dom_char cod_char by auto

              have β0: "?w0  ?p0 = ?w0'  ?p0'  β.chine"
                using Dom_β_eq Cod_β_eq β.leg0_commutes C.comp_assoc by simp
              have β1: "r1  ?p1 = r1  ?p1'  β.chine"
                using Dom_β_eq Cod_β_eq β.leg1_commutes C.comp_assoc by simp

              have Dom_θ_0: "Dom_θ.leg0 = ?w0  ?p0"
                using arrI dom_char fw T.uwθ.θ_simps(4) by auto
              have Cod_θ_0: "Cod_θ.leg0 = ?u0"
                using θ cod_char by auto
              have Dom_θ_1: "Dom_θ.leg1 = r0  ?p1"
                using arrI dom_char fw T.uwθ.θ_simps(4) by auto
              have Cod_θ_1: "Cod_θ.leg1 = ?u1"
                using T.uwθ.θ_simps(5) cod_char by auto
              have Dom_θ'_0: "Dom_θ'.leg0 = ?w0'  ?p0'"
                using dom_char fw' T.uw'θ'.θ_simps(4) by auto
              have Cod_θ'_0: "Cod_θ'.leg0 = ?u0"
                using T.uw'θ'.θ_simps(5) cod_char by auto
              have Dom_θ'_1: "Dom_θ'.leg1 = r0  ?p1'"
                using dom_char fw' T.uw'θ'.θ_simps(4) by auto
              have Cod_θ'_1: "Cod_θ'.leg1 = ?u1"
                using T.uw'θ'.θ_simps(5) cod_char by auto
              have Dom_ρ_0: "Dom_ρ.leg0 = ?R"
                by simp
              have Dom_ρ_1: "Dom_ρ.leg1 = r1"
                by simp
              have Cod_ρ_0: "Cod_ρ.leg0 = k0"
                by simp
              have Cod_ρ_1: "Cod_ρ.leg1 = r1  k1"
                by simp

              have Chn_rθ: "«rθ.chine : rfw.chine C ru.chine»"
                using rθ.chine_composite_in_hom ru.chine_composite rfw.chine_composite
                      Cod_θ_1 Dom_θ_1 fw.leg1_composite
                by auto
              have Chn_rθ_eq: "rθ.chine = 𝗉1[r0, r0  ?p1] r0, ?u1 θ.chine  𝗉0[r0, r0  ?p1]"
                using rθ.chine_composite Cod_θ_1 Dom_θ_1 fw.leg1_composite C.comp_cod_arr
                by (metis arrow_of_spans_data.simps(2) fw r.chine_eq_apex r.cod_simps(2)
                    rfw.prj_simps(10) rfw.prj_simps(16) span_data.simps(2))

              have rθ_cod_apex_eq: "rθ.cod.apex = r0 ↓↓ ?u1"
                using Cod_θ_1 rθ.chine_composite_in_hom by auto
              hence rθ'_cod_apex_eq: "rθ'.cod.apex = r0 ↓↓ ?u1"
                using Cod_θ'_1 rθ'.chine_composite_in_hom by auto
 
              have Chn_rθ': "«rθ'.chine : rfw'.chine C ru.chine»"
                using rθ'.chine_composite_in_hom ru.chine_composite rfw'.chine_composite
                      Cod_θ'_1 Dom_θ'_1 fw'.leg1_composite
                by auto
              have Chn_rθ'_eq: "rθ'.chine =
                                𝗉1[r0, r0  ?p1'] r0, ?u1 θ'.chine  𝗉0[r0, r0  ?p1']"
                using rθ'.chine_composite Cod_θ'_1 Dom_θ'_1 fw'.leg1_composite C.comp_cod_arr
                by (metis arrow_of_spans_data.simps(2) fw' r.chine_eq_apex r.cod_simps(2)
                    rfw'.prj_simps(10) rfw'.prj_simps(16) span_data.simps(2))

              have Chn_ρw: "«ρw.chine : ?R ↓↓ ?w1 C k0 ↓↓ ?w1»"
                using ρw.chine_composite_in_hom by simp
              have Chn_ρw_eq: "ρw.chine = ρ.chine  ?p1 k0, ?w1 ?p0"
                using ρw.chine_composite C.comp_cod_arr ide_w
                by (simp add: chine_hcomp_arr_ide hcomp_def)
              have Chn_ρw': "«ρw'.chine : ?R ↓↓ ?w1' C k0 ↓↓ ?w1'»"
                using ρw'.chine_composite_in_hom by simp
              have Chn_ρw'_eq: "ρw'.chine = ρ.chine  ?p1' k0, ?w1' ?p0'"
                using ρw'.chine_composite C.comp_cod_arr ide_w' Dom_ρ_0 Cod_ρ_0
                by (metis ρw'.composite_is_arrow chine_hcomp_arr_ide chine_hcomp_def hseq_char
                    w'.cod_simps(3))

              text ‹
                The following are some collected commutativity properties that are used
                subsequently.
              ›
              have "C.commutative_square r0 ?u1 ?p1 θ.chine"
                using ru.legs_form_cospan(1) Dom_θ.is_span Dom_θ_1 Cod_θ_1 θ.leg1_commutes
                by (intro C.commutative_squareI) auto
              have "C.commutative_square r0 ?u1 (?p1'  β.chine) (θ'.chine  β.chine)"
                by (metis (mono_tags, lifting) C.commutative_square_comp_arr C.dom_comp
                    C.seqE Cod_θ'_1 Dom_β.leg_simps(3) Dom_β_eq Dom_θ'.leg_simps(3)
                    Dom_θ'_1 β1 θ'.leg1_commutes C.commutative_squareI
                    ru.legs_form_cospan(1) span_data.simps(2))
              have "C.commutative_square r0 ?u1 𝗉1[r0, r0  ?p1] (θ.chine  𝗉0[r0, r0  ?p1])"
                using ru.legs_form_cospan(1) Dom_θ.is_span Dom_θ_1
                      C.comp_assoc C.pullback_commutes' rθ.legs_form_cospan(1)
                apply (intro C.commutative_squareI)
                   apply auto
                by (metis C.comp_assoc Cod_θ_1 θ.leg1_commutes)
              hence "C.commutative_square r0 ?u1 𝗉1[r0, r0  ?p1] (θ.chine  𝗉0[r0, r0  ?p1])"
                using fw.leg1_composite by auto
              have "C.commutative_square r0 ?u1 𝗉1[r0, r0  ?p1'] (θ'.chine  𝗉0[r0, r0  ?p1'])"
                using C.tuple_is_extensional Chn_rθ'_eq rθ'.chine_simps(1) fw' by force
              have "C.commutative_square ra ?w1 rfw.Prj01 rfw.Prj0"
                using C.pullback_commutes' gw.legs_form_cospan(1) rfw.prj_simps(2) C.comp_assoc
                      C.comp_cod_arr
                by (intro C.commutative_squareI) auto
              have "C.commutative_square ?R ?w1' rfw'.Prj01 rfw'.Prj0"
                by (metis (no_types, lifting) C.commutative_square_comp_arr C.comp_assoc
                    C.pullback_commutes select_convs(2) rfw'.cospan_νπ
                    rfw'.prj_chine_assoc(2) rfw'.prj_chine_assoc(3) rfw'.prj_simps(2)
                    span_data.select_convs(1))
              have "C.commutative_square r0 (r0  ?p1) rfw.Prj11 rfw.Prj01 ra, ?w1 rfw.Prj0"
              proof -
                have "C.arr rfw.chine_assoc"
                  by (metis C.seqE rfw.prj_chine_assoc(1) rfw.prj_simps(1))
                thus ?thesis
                  using C.tuple_is_extensional rfw.chine_assoc_def by fastforce
              qed
              have "C.commutative_square r0 (r0  ?p1') rfw'.Prj11 rfw'.Prj01 ra, ?w1' rfw'.Prj0"
                by (metis (no_types, lifting) C.not_arr_null C.seqE C.tuple_is_extensional
                    arrow_of_spans_data.select_convs(2) rfw'.chine_assoc_def
                    rfw'.prj_chine_assoc(1) rfw'.prj_simps(1) span_data.select_convs(1-2))
              have "C.commutative_square k0 ?w1 (ρ.chine  ?p1) ?p0"
                using C.tuple_is_extensional Chn_ρw_eq ρw.chine_simps(1) by fastforce
              have "C.commutative_square k0 ?w1' (ρ.chine  ?p1') (w'.chine  ?p0')"
                using C.tuple_is_extensional ρw'.chine_composite ρw'.chine_simps(1) by force
              have "C.commutative_square k0 ?w1' (ρ.chine  ?p1') ?p0'"
                using C.tuple_is_extensional Chn_ρw'_eq ρw'.chine_simps(1) by force
              text ‹
                Now, derive the consequences of the equation:
                \[
                  (r ⋆ θ) ∙ 𝖺[r, ?f, w] ∙ (?ρ ⋆ w) = (r ⋆ θ') ∙ 𝖺[r, ?f, w'] ∙ (?ρ ⋆ w') ∙ β›
                \]
                The strategy is to expand and simplify the left and right hand side to tuple form,
                then compose with projections and equate corresponding components.

                We first work on the right-hand side.
              ›
              have R: "Chn (T.composite_cell w' θ'  β) =
                       ?p1'  β.chine r0, ?u1 θ'.chine  β.chine"
              proof -
                have "Chn (T.composite_cell w' θ'  β) =
                      rθ'.chine  Chn 𝖺[r, f, w']  ρw'.chine  β.chine"
                proof -
                  have 1: "«T.composite_cell w' θ'  β : g  w  r  u»"
                    using β ρw' arfw' rθ' by auto
                  have "Chn (T.composite_cell w' θ'  β) = Chn (T.composite_cell w' θ')  β.chine"
                    using 1 Chn_vcomp by blast
                  also have "... = (rθ'.chine  Chn (𝖺[r, f, w']  (ρ  w')))  β.chine"
                  proof -
                    have "seq (r  θ') (𝖺[r, f, w']  (ρ  w'))"
                      using 1 by blast
                    thus ?thesis
                      using 1 Chn_vcomp by presburger
                  qed
                  also have "... = (rθ'.chine  Chn 𝖺[r, f, w']  ρw'.chine)  β.chine"
                  proof -
                    have "seq 𝖺[r, f, w'] (ρ  w')"
                      using 1 by blast
                    thus ?thesis
                      using 1 Chn_vcomp by presburger
                  qed
                  finally show ?thesis
                    using C.comp_assoc by auto
                qed
                also have "... = ?p1'  β.chine r0, ?u1 θ'.chine  β.chine"
                proof -
                  let ?LHS = "rθ'.chine  Chn 𝖺[r, f, w']  ρw'.chine  β.chine"
                  let ?RHS = "?p1'  β.chine r0, ?u1 θ'.chine  β.chine"
         
                  have LHS: "«?LHS : ?R ↓↓ ?w1 C rθ'.cod.apex»"
                  proof (intro C.comp_in_homI)
                    show "«β.chine : ?R ↓↓ ?w1 C ?R ↓↓ ?w1'»"
                      using Chn_β by simp
                    show "«ρw'.chine : ?R ↓↓ ?w1' C Cod_ρ.leg0 ↓↓ w'.cod.leg1»"
                      using Chn_ρw' by simp
                    show "«Chn 𝖺[r, f, w'] : Cod_ρ.leg0 ↓↓ w'.cod.leg1 C rfw'.chine»"
                      using arfw'
                      by (metis (no_types, lifting) Chn_in_hom Cod_ρ_0
                          arrow_of_spans_data.simps(2) rf rf.leg0_composite rfw'.chine_composite(1)
                          span_data.select_convs(1) w'.cod_simps(3))
                    show "«rθ'.chine : rfw'.chine C rθ'.cod.apex»"
                      using Chn_rθ' by auto
                  qed
                  have 2: "C.commutative_square r0 ?u1
                             (?p1'  β.chine) (θ'.chine  β.chine)"
                    by fact
                  have RHS: "«?RHS : ?R ↓↓ ?w1 C rθ'.cod.apex»"
                    using 2 Chn_β rθ'_cod_apex_eq
                          C.tuple_in_hom [of r0 ?u1 "?p1'  β.chine" "θ'.chine  β.chine"]
                    by fastforce

                  show ?thesis
                  proof (intro C.prj_joint_monic [of r0 ?u1 ?LHS ?RHS])
                    show "C.cospan r0 ?u1"
                      using ru.legs_form_cospan(1) by blast
                    show "C.seq ru.prj1 ?LHS"
                      using LHS rθ'_cod_apex_eq by auto
                    show "C.seq ru.prj1 ?RHS"
                      using RHS rθ'_cod_apex_eq by auto
                    show "ru.prj0  ?LHS = ru.prj0  ?RHS"
                    proof -
                      have "ru.prj0  ?LHS =
                            (ru.prj0  rθ'.chine)  Chn 𝖺[r, f, w']  ρw'.chine  β.chine"
                        using C.comp_assoc by simp
                      also have "... = ((θ'.chine  𝗉0[r0, r0  ?p1'])  Chn 𝖺[r, f, w']) 
                                         ρw'.chine  β.chine"
                        using Chn_rθ'_eq C.comp_assoc fw'
                              C.commutative_square r0 ?u1 𝗉1[r0, r0  ?p1']
                                  (θ'.chine  𝗉0[r0, r0  ?p1'])
                        by simp
                      also have "... = θ'.chine  (𝗉0[r0, r0  ?p1']  Chn 𝖺[r, f, w']) 
                                         ρw'.chine  β.chine"
                        using C.comp_assoc by simp
                      also have "... = θ'.chine  (rfw'.Prj01 ?R, ?w1' rfw'.Prj0  ρw'.chine) 
                                         β.chine"
                        using ide_f hseq_rf hseq_char α_ide C.comp_assoc
                              rfw'.chine_assoc_def fw'.leg1_composite C.prj_tuple(1)
                              C.commutative_square r0 (r0  ?p1')
                                 rfw'.Prj11 rfw'.Prj01 ?R, ?w1' rfw'.Prj0
                        by simp
                      also have "... = θ'.chine  β.chine"
                      proof -
                        have "rfw'.Prj01 ?R, ?w1' rfw'.Prj0  ρw'.chine = gw'.apex"
                        proof (intro C.prj_joint_monic
                                       [of ?R ?w1' "rfw'.Prj01 ?R, ?w1' rfw'.Prj0  ρw'.chine"
                                           gw'.apex])
                          show "C.cospan ?R ?w1'"
                            using fw'.legs_form_cospan(1) by simp
                          show "C.seq ?p1' (rfw'.Prj01 ?R, ?w1' rfw'.Prj0  ρw'.chine)"
                          proof (intro C.seqI' C.comp_in_homI)
                            show "«ρw'.chine : Dom_ρ.leg0 ↓↓ w'.leg1 C Cod_ρ.leg0 ↓↓ w'.cod.leg1»"
                              using ρw'.chine_composite_in_hom by simp
                            show "«rfw'.Prj01 ?R, w'.leg1 rfw'.Prj0 :
                                      Cod_ρ.leg0 ↓↓ w'.cod.leg1 C ?R ↓↓ w'.leg1»"
                              using C.commutative_square ?R ?w1' rfw'.Prj01 rfw'.Prj0
                                    C.tuple_in_hom [of ?R ?w1' rfw'.Prj01 rfw'.Prj0]
                                    rf rf.leg0_composite
                              by auto
                            show "«?p1' : ?R ↓↓ w'.leg1 C f.apex»"
                              using fw'.prj_in_hom(1) by auto
                          qed
                          show "C.seq ?p1' gw'.apex"
                            using gw'.dom.apex_def gw'.leg0_composite fw'.prj_in_hom by auto
                          show "?p0'  rfw'.Prj01 ?R, ?w1' rfw'.Prj0  ρw'.chine =
                                         ?p0'  gw'.apex"
                          proof -
                            have "?p0'  rfw'.Prj01 ?R, ?w1' rfw'.Prj0  ρw'.chine =
                                  (?p0'  rfw'.Prj01 ?R, ?w1' rfw'.Prj0)  ρw'.chine"
                              using C.comp_assoc by simp
                            also have "... = rfw'.Prj0  ρw'.chine"
                              using C.commutative_square ?R ?w1' rfw'.Prj01 rfw'.Prj0 by auto
                            also have
                              "... = 𝗉0[k0, ?w1']  ρ.chine  ?p1' k0, ?w1' w'.chine  ?p0'"
                              using ρw'.chine_composite Dom_ρ_0 Cod_ρ_0 C.comp_cod_arr by simp
                            also have "... = w'.chine  ?p0'"
                              using C.commutative_square k0 ?w1'
                                      (ρ.chine  ?p1') (w'.chine  ?p0')
                              by simp
                            also have "... = ?p0'  gw'.apex"
                              using cospan C.comp_cod_arr C.comp_arr_dom gw'.chine_is_identity
                                    gw'.chine_eq_apex gw'.chine_composite fw'.prj_in_hom
                              by auto
                            finally show ?thesis by simp
                          qed
                          show "?p1'  rfw'.Prj01 ra, ?w1' rfw'.Prj0  ρw'.chine =
                                         ?p1'  gw'.apex"
                          proof -
                            have "?p1'  rfw'.Prj01 ra, ?w1' rfw'.Prj0  ρw'.chine =
                                  (?p1'  rfw'.Prj01 ra, ?w1' rfw'.Prj0)  ρw'.chine"
                              using C.comp_assoc by simp
                            also have "... = rfw'.Prj01  ρw'.chine"
                              using C.commutative_square ?R ?w1' rfw'.Prj01 rfw'.Prj0 by simp
                            also have
                              "... = k0  𝗉1[k0, ?w1']  ρ.chine  ?p1' k0, ?w1' w'.chine  ?p0'"
                              using ρw'.chine_composite Cod_ρ_0 C.comp_assoc C.comp_cod_arr
                              by simp
                            also have "... = k0  ρ.chine  ?p1'"
                              using C.commutative_square k0 ?w1'
                                      (ρ.chine  ?p1') (w'.chine  ?p0')
                              by simp
                            also have "... = (k0  ρ.chine)  ?p1'"
                              using C.comp_assoc by metis
                            also have "... = ?p1'"
                              using ρ.leg0_commutes C.comp_cod_arr cospan' by simp
                            also have "... = ?p1'  gw'.apex"
                              using C.comp_arr_dom cospan' gw'.chine_eq_apex gw'.chine_composite
                              by simp
                            finally show ?thesis by simp
                          qed
                        qed
                        thus ?thesis
                          using Chn_β C.comp_cod_arr gw'.apex_composite by auto
                      qed
                      also have "... = 𝗉0[r0, ?u1]  ?RHS"
                        using RHS 2 C.prj_tuple [of r0 ?u1] by simp
                      finally show ?thesis by simp
                    qed
                    show "ru.prj1  ?LHS = ru.prj1  ?RHS"
                    proof -
                      have "ru.prj1  ?LHS =
                            (ru.prj1  rθ'.chine)  Chn 𝖺[r, f, w']  ρw'.chine  β.chine"
                        using C.comp_assoc by simp
                      also have "... = 𝗉1[r0, fw'.leg1]  Chn 𝖺[r, f, w']  ρw'.chine  β.chine"
                        using Chn_rθ' Chn_rθ'_eq fw'
                              C.commutative_square r0 ?u1 𝗉1[r0, r0  ?p1']
                                 (θ'.chine  𝗉0[r0, r0  ?p1'])
                        by simp
                      also have "... = (rfw'.Prj1  rfw'.chine_assoc)  ρw'.chine  β.chine"
                        using ide_f ide_w' hseq_rf hseq_char α_ide fw'.leg1_composite C.comp_assoc
                        by auto
                      also have "... = (rfw'.Prj11  ρw'.chine)  β.chine"
                        using rfw'.prj_chine_assoc C.comp_assoc by simp
                      also have "... = ((k1  𝗉1[k0, ?w1'])  ρw'.chine)  β.chine"
                        using C.comp_cod_arr by simp
                      also have "... = (k1  𝗉1[k0, ?w1']  ρw'.chine)  β.chine"
                        using C.comp_assoc by simp
                      also have "... = (k1  ρ.chine  ?p1')  β.chine"
                        using Chn_ρw'_eq Dom_ρ_0 Cod_ρ_0
                              C.commutative_square k0 ?w1' (ρ.chine  ?p1') ?p0'
                        by simp
                      also have "... = (k1  ρ.chine)  ?p1'  β.chine"
                        using C.comp_assoc by metis
                      also have "... = (?R  ?p1')  β.chine"
                        using C.comp_assoc by simp
                      also have "... = ?p1'  β.chine"
                        using C.comp_cod_arr C.prj1_in_hom [of ?R ?w1'] cospan' by simp
                     also have "... = ru.prj1  ?RHS"
                        using RHS 2 by simp
                      finally show ?thesis by simp
                    qed
                  qed
                qed
                finally show ?thesis by simp
              qed

              text ‹
                Now we work on the left-hand side.
              ›
              have L: "Chn (T.composite_cell w θ) = ?p1 r0, ?u1 θ.chine"
              proof -
                have "Chn (T.composite_cell w θ) = rθ.chine  Chn 𝖺[r, f, w]  ρw.chine"
                  using Chn_vcomp arfw C.comp_assoc by auto
                moreover have "... = ?p1 r0, ?u1 θ.chine"
                proof -
                  let ?LHS = "rθ.chine  Chn 𝖺[r, f, w]  ρw.chine"
                  let ?RHS = "?p1 r0, ?u1 θ.chine"
                  have 2: "C.commutative_square r0 ?u1 ?p1 θ.chine" by fact
                  have LHS: "«?LHS : ?R ↓↓ ?w1 C r0 ↓↓ ?u1»"
                    using Chn_rθ Chn_ρw rfw.chine_assoc_in_hom
                    by (metis (no_types, lifting) "8" Chn_in_hom Dom_ρ_0
                        arrow_of_spans_data.simps(2) calculation gw.chine_composite
                        rθ_cod_apex_eq ru.chine_composite)
                  have RHS: "«?RHS : ?R ↓↓ ?w1 C r0 ↓↓ ?u1»"
                    using 2 C.tuple_in_hom [of r0 ?u1 "?p1" θ.chine] cospan rθ_cod_apex_eq
                    by simp
                  show ?thesis
                  proof (intro C.prj_joint_monic [of r0 ?u1 ?LHS ?RHS])
                    show "C.cospan r0 ?u1"
                      using ru.legs_form_cospan(1) by blast
                    show "C.seq ru.prj1 ?LHS"
                      using LHS rθ_cod_apex_eq by auto
                    show "C.seq ru.prj1 ?RHS"
                      using RHS rθ_cod_apex_eq by auto
                    show "ru.prj0  ?LHS = ru.prj0  ?RHS"
                    proof -
                      have "ru.prj0  ?LHS = (ru.prj0  rθ.chine)  Chn 𝖺[r, f, w]  ρw.chine"
                        using C.comp_assoc by simp
                      also have "... = (θ.chine  𝗉0[r0, f.leg1  fw.prj1]) 
                                         Chn 𝖺[r, f, w]  ρw.chine"
                        using Chn_rθ_eq Dom_θ_1 Cod_θ_1 fw.leg1_composite
                              C.commutative_square r0 ?u1 𝗉1[r0, r0  ?p1]
                                 (θ.chine  𝗉0[r0, r0  ?p1])
                        by simp
                      also have "... = θ.chine  (𝗉0[r0, r0  ?p1]  Chn 𝖺[r, f, w])  ρw.chine"
                        using C.comp_assoc by simp
                      also have "... = θ.chine  rfw.Prj01 ?R, ?w1 rfw.Prj0  ρw.chine"
                      proof -
                        have "Chn 𝖺[r, f, w] = rfw.chine_assoc"
                          using ide_f ide_w hseq_rf hseq_char α_ide by auto
                        moreover have "𝗉0[r0, r0  ?p1]  rfw.chine_assoc =
                                       rfw.Prj01 ?R, ?w1 rfw.Prj0"
                          using rfw.chine_assoc_def
                                C.commutative_square r0 (r0  ?p1) rfw.Prj11
                                   rfw.Prj01 ?R, ?w1 rfw.Prj0
                          by simp
                        ultimately show ?thesis by simp
                      qed
                      also have "... = θ.chine  (?R ↓↓ ?w1)"
                      proof -
                        have "rfw.Prj01 ?R, ?w1 rfw.Prj0  ρw.chine = ?R ↓↓ ?w1"
                        proof (intro C.prj_joint_monic
                                       [of ?R ?w1 "rfw.Prj01 ?R, ?w1 rfw.Prj0  ρw.chine"
                                           "?R ↓↓ ?w1"])
                          show "C.cospan ?R ?w1" by fact
                          show "C.seq ?p1 (rfw.Prj01 ?R, ?w1 rfw.Prj0  ρw.chine)"
                          proof -
                            have "C.seq rfw.Prj01 ρw.chine"
                              by (meson C.seqI' Chn_in_hom ρw rfw.prj_in_hom(2)
                                  C.commutative_square ?R ?w1 rfw.Prj01 rfw.Prj0)
                            thus ?thesis
                              using C.commutative_square ?R ?w1 rfw.Prj01 rfw.Prj0
                              by (metis (no_types) C.comp_assoc C.prj_tuple(2))
                          qed
                          show "C.seq ?p1 (?R ↓↓ ?w1)"
                            using gw.dom.apex_def gw.leg0_composite gw.prj_in_hom by auto
                          show "?p0  rfw.Prj01 ?R, ?w1 rfw.Prj0  ρw.chine =
                                ?p0  (?R ↓↓ ?w1)"
                          proof -
                            have "?p0  rfw.Prj01 ?R, ?w1 rfw.Prj0  ρw.chine =
                                  (?p0  rfw.Prj01 ?R, ?w1 rfw.Prj0)  ρw.chine"
                              using C.comp_assoc by simp
                            also have "... = rfw.Prj0  ρw.chine"
                              using C.commutative_square ?R ?w1 rfw.Prj01 rfw.Prj0 by simp
                            also have "... = 𝗉0[k0, ?w1]  ρ.chine  ?p1 k0, ?w1 ?p0"
                              using Chn_ρw_eq C.comp_cod_arr by simp
                            also have "... = ?p0"
                              using C.commutative_square k0 ?w1 (ρ.chine  ?p1) ?p0
                                    C.prj_tuple(1)
                              by blast
                            also have "... = ?p0  (?R ↓↓ ?w1)"
                              using C.comp_arr_dom gw.chine_eq_apex gw.chine_is_identity
                              by (metis C.arr_dom_iff_arr C.pbdom_def Dom_g gw.chine_composite
                                  gw.chine_simps(1) span_data.select_convs(1))
                            finally show ?thesis by simp
                          qed
                          show "?p1  rfw.Prj01 ?R, ?w1 rfw.Prj0  ρw.chine =
                                ?p1  (?R ↓↓ ?w1)"
                          proof -
                            have "?p1  rfw.Prj01 ?R, ?w1 rfw.Prj0  ρw.chine =
                                  (?p1  rfw.Prj01 ?R, ?w1 rfw.Prj0)  ρw.chine"
                              using C.comp_assoc by simp
                            also have "... = rfw.Prj01  ρw.chine"
                              using C.commutative_square ?R ?w1 rfw.Prj01 rfw.Prj0 by simp
                            also have "... = (k0  𝗉1[k0, ?w1])  ρ.chine  ?p1 k0, ?w1 ?p0"
                              using Chn_ρw_eq C.comp_cod_arr by simp
                            also have "... = k0  𝗉1[k0, ?w1]  ρ.chine  ?p1 k0, ?w1 ?p0"
                              using C.comp_assoc by simp
                            also have "... = k0  ρ.chine  ?p1"
                              using C.commutative_square k0 ?w1 (ρ.chine  ?p1) ?p0 by simp
                            also have "... = (k0  ρ.chine)  ?p1"
                              using C.comp_assoc by metis
                            also have "... = ?p1  (?R ↓↓ ?w1)"
                              using C.comp_arr_dom C.comp_cod_arr cospan by simp
                            finally show ?thesis by blast
                          qed
                        qed
                        thus ?thesis by simp
                      qed
                      also have "... = θ.chine"
                        using C.comp_arr_dom θ.chine_in_hom gw.chine_eq_apex gw.chine_is_identity
                              Dom_θ_0 Cod_θ_0 Dom_θ.apex_def Cod_θ.apex_def
                        by (metis Dom_g θ.chine_simps(1) θ.chine_simps(2) gw.chine_composite
                            gw.dom.apex_def gw.leg0_composite span_data.select_convs(1))
                      also have "... = ru.prj0  ?RHS"
                        using 2 by simp
                      finally show ?thesis by blast
                    qed
                    show "ru.prj1  ?LHS = ru.prj1  ?RHS"
                    proof -
                      have "ru.prj1  ?LHS = (ru.prj1  rθ.chine)  Chn 𝖺[r, f, w]  ρw.chine"
                        using C.comp_assoc by simp
                      also have "... = (r.chine  𝗉1[r0, r0  ?p1])  Chn 𝖺[r, f, w]  ρw.chine"
                      proof -
                        have "rθ.chine  C.null 
                                𝗉1[r.cod.leg0, Cod_θ.leg1]  rθ.chine =
                                r.chine  𝗉1[r0, Dom_θ.leg1]"
                          by (metis (lifting) C.prj_tuple(2) C.tuple_is_extensional r.cod_simps(2)
                              rθ.chine_composite)
                        thus ?thesis
                          using Cod_θ_1 Dom_θ_1 rθ.chine_simps(1) fw by fastforce
                      qed
                      also have "... = r.chine  (rfw.Prj1  Chn 𝖺[r, f, w])  ρw.chine"
                        using C.comp_assoc fw.leg1_composite by simp
                      also have "... = r.chine  rfw.Prj11  ρw.chine"
                        using ide_f ide_w hseq_rf hseq_char α_ide
                              rfw.prj_chine_assoc(1)
                        by auto
                      also have "... = r.chine  k1  𝗉1[k0, ?w1]  ρw.chine"
                        using C.comp_cod_arr C.comp_assoc by simp
                      also have "... = r.chine  k1  ρ.chine  𝗉1[Dom_ρ.leg0, ?w1]"
                        using Chn_ρw_eq
                              C.commutative_square k0 ?w1
                                (ρ.chine  𝗉1[ra, w.leg1]) 𝗉0[ra, w.leg1]
                        by auto
                      also have "... = r.chine  (k1  ρ.chine)  ?p1"
                        using C.comp_assoc Dom_ρ_0 by metis
                      also have "... = r.chine  ra  ?p1"
                        by simp
                      also have "... = r.chine  ?p1"
                        using C.comp_cod_arr
                        by (metis C.comp_assoc r.cod_simps(1) r.chine_eq_apex r.chine_simps(1)
                            r.chine_simps(3))
                      also have "... = ?p1"
                        using C.comp_cod_arr r.chine_eq_apex r.chine_is_identity
                        by (metis 2 C.commutative_squareE r.dom.apex_def)
                      also have "... = ru.prj1  ?RHS"
                        using 2 by simp
                      finally show ?thesis by simp
                    qed
                  qed
                qed
                ultimately show ?thesis
                  by simp
              qed
              text ‹
                This is the main point: the equation E boils down to the following:
                \[
                   ?p1' ⋅ β.chine = ?p1 ∧ θ'.chine ⋅ β.chine = θ.chine›
                \]
                The first equation gets us close to what we need, but we still need
                ?p1 ⋅ C.inv ?p0 = ?w1›, which follows from the fact that ?p0 is the
                pullback of ?R.
              ›
              have *: "?p1'  β.chine r0, ?u1 θ'.chine  β.chine = ?p1 r0, ?u1 θ.chine"
                using L R E by simp
              have **: "?p1'  β.chine = ?p1"
                by (metis "*" C.in_homE C.not_arr_null C.prj_tuple(2) C.tuple_in_hom
                    C.tuple_is_extensional
                    C.commutative_square r0 u.leg1
                       (𝗉1[ra, w'.leg1]  β.chine) (θ'.chine  β.chine))
              have ***: "θ'.chine  β.chine = θ.chine"
                by (metis "*" C.prj_tuple(1) C.commutative_square r0 ?u1
                    (?p1'  β.chine) (θ'.chine  β.chine)
                    C.commutative_square r0 ?u1 ?p1 θ.chine)
              text ‹
                CKS say to take γ = β›, but obviously this cannot work as
                literally described, because «β : g ⋆ w ⇒ g ⋆ w'»›, whereas we must have
                «γ : w ⇒ w'»›.  Instead, we have to define γ› by transporting β› along the
                projections from ?R ↓↓ ?w1› to ?W› and ?R ↓↓ ?w1'› to ?W'›.
                These are isomorphisms by virtue of their being pullbacks of identities,
                but they are not themselves necessarily identities.
                Specifically, we take Chn γ = ?p0' ⋅ Chn β ⋅ C.inv ?p0›.
              ›
              let  = "Chn = ?p0'  β.chine  C.inv ?p0, Dom = Dom w, Cod = Cod w'"
              interpret Dom_γ: span_in_category C Dom 
                using w.dom.span_in_category_axioms by simp
              interpret Cod_γ: span_in_category C Cod 
                using w'.cod.span_in_category_axioms by simp
              text ‹
                It has to be shown that γ› is an arrow of spans.
              ›
              interpret γ: arrow_of_spans C 
              proof
                show "«Chn  : Dom_γ.apex C Cod_γ.apex»"
                proof -
                  have "«Chn β: gw.apex C gw'.apex»"
                    using Chn_in_hom β gw'.chine_eq_apex gw.chine_eq_apex by force
                  moreover have "«?p0' : gw'.apex C w'.apex»"
                    using cospan' hseq_gw' hseq_char hcomp_def gw'.dom.apex_def w'.dom.apex_def
                    by auto
                  moreover have "«C.inv ?p0 : w.apex C gw.apex»"
                    using cospan hseq_gw hseq_char hcomp_def gw.dom.apex_def w.dom.apex_def
                          C.iso_pullback_ide
                    by auto
                  ultimately show ?thesis
                    using Dom_γ.apex_def Cod_γ.apex_def by auto
                qed
                text ‹
                  The commutativity property for the ``input leg'' follows directly from that
                  for β›.
                ›
                show "Cod_γ.leg0  Chn  = Dom_γ.leg0"
                  using C.comp_assoc C.comp_arr_dom cospan C.iso_pullback_ide C.comp_arr_inv'
                  by (metis C.invert_side_of_triangle(2) Dom_β.leg_simps(1) Dom_β_eq β0
                      arrow_of_spans_data.select_convs(1,3) arrow_of_spans_data.simps(2)
                      r.dom.ide_apex span_data.select_convs(1) w'.cod_simps(2))
                text ‹
                  The commutativity property for the ``output leg'' is a bit more subtle.
                ›
                show "Cod_γ.leg1  Chn  = Dom_γ.leg1"
                proof -
                  have "Cod_γ.leg1  Chn  = ((?w1'  ?p0')  β.chine)  C.inv ?p0"
                    using C.comp_assoc by simp
                  also have "... = ((?R  ?p1')  Chn β)  C.inv ?p0"
                    using cospan' C.pullback_commutes [of ?R ?w1'] by auto
                  also have "... = (?p1'  β.chine)  C.inv ?p0"
                    using cospan' C.comp_cod_arr by simp
                  also have "... = ?p1  C.inv ?p0"
                    using ** by simp
                  also have "... = ?w1"
                  text ‹
                     Sledgehammer found this at a time when I was still struggling to
                     understand what was going on.
                  ›
                    by (metis C.comp_cod_arr C.invert_side_of_triangle(2) C.iso_pullback_ide
                        C.prj1_simps(1,3) C.pullback_commutes' cospan r.dom.ide_apex
                        r.chine_eq_apex r.chine_simps(2))
                  also have "... = Dom_γ.leg1" by auto
                  finally show ?thesis by simp
                qed
              qed
              text ‹
                What remains to be shown is that γ› is unique with the properties asserted
                by T2›; \emph{i.e.} «γ : w ⇒ w'» ∧ β = g ⋆ γ ∧ θ = θ' ∙ (f ⋆ γ)›.
                CKS' assertion that the equation (rθ)(ρw) = (rθ')(ρw')β› gives w1 = w1'›
                does not really seem to be true.  The reason γ› is unique is because it is
                obtained by transporting β› along isomorphisms.
              ›
              have γ: "« : w  w'»"
                using γ.arrow_of_spans_axioms arr_char dom_char cod_char by auto
              have hseq_fγ: "hseq f "
                using γ src_def trg_def arrI fw.composable rf.are_arrows(2) by auto
              have hseq_gγ: "hseq g "
                using γ src_def trg_def fw.composable gw.are_arrows(1) src_f by auto
              interpret: two_composable_arrows_of_spans C prj0 prj1 f 
                using hseq_fγ hseq_char by (unfold_locales, simp)
              interpret: arrow_of_spans C f  
                using fγ.composite_is_arrow arr_char by simp
              interpret: two_composable_arrows_of_spans C prj0 prj1 g 
                using hseq_gγ hseq_char by (unfold_locales, simp)
              interpret: arrow_of_spans C g  
                using gγ.composite_is_arrow arr_char by simp
              have Chn_gγ: "Chn (g  ) = ?p1 ?R, ?w1' ?p0'  β.chine"
              proof -
                have "Chn (g  ) = ?R  ?p1 ?R, ?w1' (?p0'  β.chine  C.inv ?p0)  ?p0"
                    using gγ.chine_composite by simp
                also have "... = ?p1 ?R, ?w1' (?p0'  β.chine  C.inv ?p0)  ?p0"
                  using C.comp_cod_arr cospan by simp
                also have "... = ?p1 ?R, ?w1' ?p0'  β.chine"
                proof -
                  have "(?p0'  β.chine  C.inv ?p0)  ?p0 = ?p0'  β.chine"
                    using C.comp_assoc C.iso_pullback_ide [of ?R ?w1] C.comp_inv_arr
                          C.comp_arr_dom Chn_β
                    by (metis C.comp_inv_arr' C.in_homE C.pbdom_def cospan r.dom.ide_apex)
                  thus ?thesis by simp
                qed
                ultimately show ?thesis by simp
              qed
              have Chn_β_eq: "β.chine = Chn (g  )"
                by (metis "**" C.span_prj C.tuple_prj Chn_gγ cospan cospan')
              have β_eq_gγ: "β = g  "
              proof (intro arr_eqI)
                show "par β (g  )"
                proof -
                  have "«g   : g  w  g  w'»"
                    using ide_g γ T.leg1_simps(3)
                    by (intro hcomp_in_vhom, auto)
                  thus ?thesis
                    using β by (elim in_homE, auto)
                qed
                show "β.chine = Chn (g  )"
                  using Chn_β_eq by simp
              qed
              moreover have "θ = θ'  (f  )"
              proof (intro arr_eqI)
                have : "«f   : f  w  f  w'»"
                  using γ ide_f by auto
                show par: "par θ (θ'  (f  ))"
                  using θ θ'  by (elim in_homE, auto)
                show "θ.chine = Chn (θ'  (f  ))"
                  using par "***" Chn_vcomp calculation fγ.chine_composite gγ.chine_composite
                  by auto
              qed
              ultimately show 2: "« : w  w'»  β = g    θ = θ'  (f  )"
                using γ by simp
              show "γ'. «γ' : w  w'»  β = g  γ'  θ = θ'  (f  γ')  γ' = "
              proof -
                fix γ'
                assume 1: "«γ' : w  w'»  β = g  γ'  θ = θ'  (f  γ')"
                interpret γ': arrow_of_spans C γ'
                  using 1 arr_char by auto
                have hseq_gγ': hseq g γ'
                  using 1 β by auto
                interpret gγ': two_composable_arrows_of_spans C prj0 prj1 g γ'
                  using hseq_gγ' hseq_char by unfold_locales auto
                interpret gγ': arrow_of_spans C g  γ'
                  using gγ'.composite_is_arrow arr_char by simp
                show "γ' = "
                proof (intro arr_eqI)
                  show par: "par γ' "
                    using 1 γ by fastforce
                  show "γ'.chine = γ.chine"
                  proof -
                    have "C.commutative_square ?R ?w1' (g.chine  ?p1) (γ'.chine  ?p0)"
                    proof
                      show "C.cospan ?R ?w1'" by fact
                      show 3: "C.span (g.chine  ?p1) (γ'.chine  ?p0)"
                      proof (intro conjI)
                        show "C.seq g.chine ?p1"
                          using cospan by auto
                        show "C.seq γ'.chine ?p0"
                          using cospan 2 par arrow_of_spans_data.simps(1)
                                dom_char in_homE w.chine_eq_apex
                          by auto
                        thus "C.dom (g.chine  ?p1) = C.dom (γ'.chine  ?p0)"
                          using g.chine_eq_apex cospan by simp
                      qed
                      show "C.dom ra = C.cod (g.chine  ?p1)"
                        using cospan by auto
                      show "?R  g.chine  ?p1 = ?w1'  γ'.chine  ?p0"
                      proof -
                        have "?w1'  γ'.chine  ?p0 = (?w1'  γ'.chine)  ?p0"
                          using C.comp_assoc by simp
                        moreover have "... = ?w1  ?p0"
                          using 1 γ'.leg1_commutes dom_char cod_char by auto
                        also have "... = ?R  ?p1"
                          using cospan C.pullback_commutes [of ra ?w1] by auto
                        also have "... = ?R  g.chine  ?p1"
                          using 3 C.comp_cod_arr g.chine_is_identity g.chine_eq_apex g.dom.apex_def
                          by auto
                        finally show ?thesis by auto
                      qed
                    qed
                    have "C.commutative_square ?R ?w1' (g.chine  ?p1) (γ.chine  ?p0)"
                    proof
                      show "C.cospan ?R ?w1'" by fact
                      show 3: "C.span (g.chine  ?p1) (γ.chine  ?p0)"
                        using cospan γ.chine_in_hom by auto
                      show "C.dom ?R = C.cod (g.chine  ?p1)"
                        using cospan by auto
                      show "?R  g.chine  ?p1 = ?w1'  γ.chine  ?p0"
                      proof -
                        have "?w1'  γ.chine  ?p0 = (?w1'  γ.chine)  ?p0"
                          using C.comp_assoc by simp
                        moreover have "... = ?w1  ?p0"
                          using 1 γ.leg1_commutes dom_char cod_char by auto
                        also have "... = ?R  ?p1"
                          using cospan C.pullback_commutes [of ra ?w1] by auto
                        also have "... = ?R  g.chine  ?p1"
                          using 3 C.comp_cod_arr g.chine_is_identity g.chine_eq_apex g.dom.apex_def
                          by auto
                        finally show ?thesis by auto
                      qed
                    qed
                    have "γ'.chine  ?p0 = γ.chine  ?p0"
                    proof -
                      have "γ'.chine  ?p0 = ?p0'  gγ'.chine"
                        using 1 dom_char cod_char gγ'.chine_composite
                              C.commutative_square ?R ?w1' (g.chine  ?p1) (γ'.chine  ?p0)
                        by auto
                      also have "... = ?p0'  β.chine"
                        using 1 by simp
                      also have "... = ?p0'  gγ.chine"
                        using Chn_β_eq by simp
                      also have "... = γ.chine  ?p0"
                        using gγ.chine_composite
                              C.commutative_square ?R ?w1' (g.chine  ?p1) (γ.chine  ?p0)
                        by simp
                      finally show ?thesis by simp
                    qed
                    thus ?thesis
                      using C.iso_pullback_ide C.iso_is_retraction C.retraction_is_epi
                            C.epiE [of "?p0" γ'.chine γ.chine] cospan γ.chine_in_hom
                            γ'.chine_in_hom
                      by auto
                  qed
                qed
              qed
            qed
          qed
        qed
      qed
    qed

  end

  context span_bicategory
  begin

    interpretation chosen_right_adjoints vcomp hcomp assoc unit src trg ..
    notation some_right_adjoint  ("_*" [1000] 1000)  (* TODO: Why is this needed? *)
    notation isomorphic  (infix "" 50)

    text Span(C)› is a bicategory of spans.
    ›

    lemma is_bicategory_of_spans:
    shows "bicategory_of_spans vcomp hcomp assoc unit src trg"
    proof
      text ‹
        Every 1-cell r› is isomorphic to the composition of a map and the right adjoint
        of a map.  The proof is to obtain a tabulation of r› as a span of maps (f, g)›
        and then observe that r› is isomorphic to g ⋆ f*.
      ›
      show "r. ide r  f g. is_left_adjoint f  is_left_adjoint g  r  g  f*"
      proof -
        fix r
        assume r: "ide r"
        interpret r: identity_arrow_of_spans C r
          using r ide_char' by auto
        interpret r: identity_arrow_in_span_bicategory C prj0 prj1 r ..
        have ρ: "tabulation (∙) (⋆) assoc unit src trg r r.ρ r.f r.g 
                 is_left_adjoint r.f  is_left_adjoint r.g"
          using r r.has_tabulation by blast
        interpret ρ: tabulation vcomp hcomp assoc unit src trg r r.ρ r.f r.g
          using ρ by fast
        have 1: "r  r.g  r.f*"
          using ρ ρ.yields_isomorphic_representation' ρ.T0.is_map
                left_adjoint_extends_to_adjoint_pair
                isomorphic_def [of "r.g  r.f*" r] isomorphic_symmetric
          by auto
        thus "f g. is_left_adjoint f  is_left_adjoint g  r  g  f*"
          using ρ by blast
      qed
      text ‹
        Every span of maps extends to a tabulation.
      ›
      show "f g.  is_left_adjoint f; is_left_adjoint g; src f = src g  
                   r ρ. tabulation (∙) (⋆) assoc unit src trg r ρ f g"
      proof -
        text ‹
          The proof idea is as follows:  Let maps f = (f1, f0)› and g = (g1, g0)› be given.
          Let f' = (f1 ⋅ C.inv f0, C.cod f0)› and g' = (g1 ⋅ C.inv g0, C.cod g0)›;
          then f'› and g'› are maps isomorphic to f› and g›, respectively.
          By a previous result, f'› and g'› extend to a tabulation (f', τ, g')› of
          r = (f1 ⋅ C.inv f0, g1 ⋅ C.inv g0)›.
          Compose with isomorphisms «φ : f' ⇒ f»› and «ψ : g ⇒ g'»› to obtain
          (f, (r ⋆ φ) ⋅ τ ⋅ ψ, g)› and show it must also be a tabulation.
        ›
        fix f g
        assume f: "is_left_adjoint f"
        assume g: "is_left_adjoint g"
        assume fg: "src f = src g"
        show "r ρ. tabulation (∙) (⋆) assoc unit src trg r ρ f g"
        proof -
          text ‹We have to unpack the hypotheses to get information about f and g.›
          obtain fa ηf εf
            where ffa: "adjunction_in_bicategory vcomp hcomp assoc unit src trg f fa ηf εf"
            using f adjoint_pair_def by auto
          interpret ffa: adjunction_in_bicategory vcomp hcomp assoc unit src trg f fa ηf εf
            using ffa by simp
          interpret f: arrow_of_spans C f
            using ide_char [of f] by simp
          interpret f: identity_arrow_of_spans C f
            using ide_char [of f] by unfold_locales auto
          obtain ga ηg εg
            where G: "adjunction_in_bicategory vcomp hcomp assoc unit src trg g ga ηg εg"
            using g adjoint_pair_def by auto
          interpret gga: adjunction_in_bicategory vcomp hcomp assoc unit src trg g ga ηg εg
            using G by simp
          interpret g: arrow_of_spans C g
            using ide_char [of g] by simp
          interpret g: identity_arrow_of_spans C g
            using ide_char [of g] by unfold_locales auto

          let ?f' = "mkIde (C.cod f.leg0) (f.dom.leg1  C.inv f.leg0)"
          have f': "ide ?f'"
          proof -
            have "C.span (C.cod f.leg0) (f.leg1  C.inv f.leg0)"
              using f is_left_adjoint_char by auto
            thus ?thesis
              using ide_mkIde by blast
          qed
          interpret f': arrow_of_spans C ?f'
            using f' ide_char by blast
          interpret f': identity_arrow_of_spans C ?f'
            using f' ide_char by unfold_locales auto

          let ?g' = "mkIde (C.cod g.leg0) (g.dom.leg1  C.inv g.leg0)"
          have g': "ide ?g'"
          proof -
            have "C.span (C.cod g.leg0) (g.leg1  C.inv g.leg0)"
              using g is_left_adjoint_char by auto
            thus ?thesis
              using ide_mkIde by blast
          qed
          interpret g': arrow_of_spans C ?g'
            using g' ide_char by blast
          interpret g': identity_arrow_of_spans C ?g'
            using g' ide_char by unfold_locales auto
        
          let ?r = "mkIde (f'.leg1) (g'.leg1)"
          have r: "ide ?r"
          proof -
            have "C.span (f'.leg1) (g'.leg1)"
              using f g fg src_def is_left_adjoint_char by simp
            thus ?thesis
              using ide_mkIde by blast
          qed
          interpret r: arrow_of_spans C ?r
            using r ide_char by blast
          interpret r: identity_arrow_of_spans C ?r
            using r ide_char by unfold_locales auto
          interpret r: identity_arrow_in_span_bicategory C prj0 prj1 ?r ..

          have "r.f = ?f'"
            using f r.chine_eq_apex is_left_adjoint_char by auto
          have "r.g = ?g'"
            using f r.chine_eq_apex fg src_def is_left_adjoint_char by simp

          interpret ρ: tabulation (∙) (⋆) assoc unit src trg ?r r.ρ r.f r.g
            using r.has_tabulation by simp
          have ρ_eq: "r.ρ = Chn = C.cod f.leg0 f'.leg1, f'.leg1 C.cod f.leg0,
                             Dom = Leg0 = C.cod f.leg0, Leg1 = g'.leg1,
                             Cod = Leg0 = 𝗉0[f'.leg1, f'.leg1],
                                    Leg1 = g'.leg1  𝗉1[f'.leg1, f'.leg1]"
            using r.f = ?f' by auto

          text ‹Obtain the isomorphism from f'› to f›.›
          let  = "Chn = C.inv f.leg0, Dom = Dom ?f', Cod = Dom f"
          interpret Dom_φ: span_in_category C
                              Dom Chn = C.inv f.leg0,
                                    Dom = Dom (mkIde f.dsrc (f.leg1  C.inv f.leg0)),
                                    Cod = Dom f
            using f'.dom.span_in_category_axioms by simp
          interpret Cod_φ: span_in_category C
                              Cod Chn = C.inv f.leg0,
                                    Dom = Dom (mkIde f.dsrc (f.leg1  C.inv f.leg0)),
                                    Cod = Dom f
            using f.dom.span_in_category_axioms by simp
          interpret φ: arrow_of_spans C 
          proof
            show "«Chn Chn = C.inv f.leg0,
                        Dom = Dom (mkIde f.dsrc (f.leg1  C.inv f.leg0)),
                        Cod = Dom f : Dom_φ.apex C Cod_φ.apex»"
              using f f.dom.apex_def f'.dom.apex_def is_left_adjoint_char by auto
            show "Cod_φ.leg0  Chn Chn = C.inv f.leg0,
                                    Dom = Dom (mkIde f.dsrc (f.leg1  C.inv f.leg0)),
                                    Cod = Dom f =
                  Dom_φ.leg0"
              using f f.dom.apex_def is_left_adjoint_char C.comp_arr_inv C.inv_is_inverse
              by simp
            show "Cod_φ.leg1  Chn Chn = C.inv f.leg0,
                                    Dom = Dom (mkIde f.dsrc (f.leg1  C.inv f.leg0)),
                                    Cod = Dom f =
                  Dom_φ.leg1"
              by simp
          qed
          have φ: "« : ?f'  f»  iso "
            using f is_left_adjoint_char iso_char arr_char dom_char cod_char
                  φ.arrow_of_spans_axioms f'.dom.apex_def f.dom.apex_def
            by auto

          text ‹
            Obtain the isomorphism from g› to g'›.
            Recall: g' = mkIde (C.cod g.leg0) (g.dom.leg1 ⋅ C.inv g.leg0)›.
            The isomorphism is given by g.leg0›.
          ›
          let  = "Chn = g.leg0, Dom = Dom g, Cod = Dom ?g'"
          interpret Dom_ψ: span_in_category C
                              Dom Chn = g.leg0,
                                    Dom = Dom g,
                                    Cod = Dom (mkIde g.dsrc (g.leg1  C.inv g.leg0))
            using g.dom.span_in_category_axioms by simp
          interpret Cod_ψ: span_in_category C
                              Cod Chn = g.leg0,
                                    Dom = Dom g,
                                    Cod = Dom (mkIde g.dsrc (g.leg1  C.inv g.leg0))
            using g'.dom.span_in_category_axioms by simp
          interpret ψ: arrow_of_spans C 
          proof
            show "«Chn Chn = g.leg0,
                        Dom = Dom g,
                        Cod = Dom (mkIde g.dsrc (g.leg1  C.inv g.leg0)) :
                          Dom_ψ.apex C Cod_ψ.apex»"
              using g g.dom.apex_def g'.dom.apex_def is_left_adjoint_char by auto
            show "Cod_ψ.leg0  Chn Chn = g.leg0,
                                    Dom = Dom g,
                                    Cod = Dom (mkIde g.dsrc (g.leg1  C.inv g.leg0)) =
                  Dom_ψ.leg0"
              using C.comp_cod_arr by simp
            show "Cod_ψ.leg1  Chn Chn = g.leg0,
                                    Dom = Dom g,
                                    Cod = Dom (mkIde g.dsrc (g.leg1  C.inv g.leg0)) =
                  Dom_ψ.leg1"
              using g g.dom.apex_def is_left_adjoint_char C.comp_inv_arr C.inv_is_inverse
                    C.comp_assoc C.comp_arr_dom
              by simp
          qed
          have ψ: "« : g  ?g'»  iso "
            using g is_left_adjoint_char iso_char arr_char dom_char cod_char
                  ψ.arrow_of_spans_axioms g.dom.apex_def g'.dom.apex_def
            by auto
          have ρψ: "tabulation (∙) (⋆) assoc unit src trg ?r (r.ρ  ) r.f g"
            using ψ r.g = ?g' r.has_tabulation ρ.preserved_by_output_iso by simp
          interpret τψ: tabulation vcomp hcomp assoc unit src trg ?r r.ρ   r.f g
            using ρψ by auto
          have "tabulation (∙) (⋆) assoc unit src trg ?r ((?r  )  r.ρ  ) f g"
            using φ r.f = ?f' τψ.preserved_by_input_iso [of  f] by argo
          thus ?thesis by auto
        qed
      qed

      text ‹The sub-bicategory of maps is locally essentially discrete.›
      show "f f' μ μ'.  is_left_adjoint f; is_left_adjoint f'; «μ : f  f'»; «μ' : f  f'» 
                             iso μ  iso μ'  μ = μ'"
      proof -
        fix f f' μ μ'
        assume f: "is_left_adjoint f" and f': "is_left_adjoint f'"
        assume μ: "«μ : f  f'»" and μ': "«μ' : f  f'»"
        obtain fa η ε
          where fa: "adjunction_in_bicategory vcomp hcomp assoc unit src trg f fa η ε"
          using f adjoint_pair_def by auto
        obtain f'a η' ε'
          where f'a: "adjunction_in_bicategory vcomp hcomp assoc unit src trg f' f'a η' ε'"
          using f' adjoint_pair_def adjunction_def by auto
        interpret fa: adjunction_in_bicategory vcomp hcomp assoc unit src trg f fa η ε
          using fa by simp
        interpret f'a: adjunction_in_bicategory vcomp hcomp assoc unit src trg f' f'a η' ε'
          using f'a by simp
        interpret f: identity_arrow_of_spans C f
          using ide_char' [of f] by simp
        interpret f': identity_arrow_of_spans C f'
          using ide_char' [of f'] by simp
        interpret μ: arrow_of_spans C μ using μ arr_char by auto
        interpret μ': arrow_of_spans C μ' using μ' arr_char by auto
        have 1: "C.iso f.leg0  C.iso f'.leg0"
          using f f' is_left_adjoint_char by simp
        have 2: "μ.chine = C.inv f'.leg0  f.leg0"
          using μ 1 dom_char cod_char μ.leg0_commutes C.invert_side_of_triangle by auto
        moreover have "μ'.chine = C.inv f'.leg0  f.leg0"
          using μ' 1 dom_char cod_char μ'.leg0_commutes C.invert_side_of_triangle by auto
        ultimately have 3: "μ.chine = μ'.chine" by simp
        have "iso μ"
          using 1 2 μ C.isos_compose dom_char cod_char iso_char arr_char by auto
        hence "iso μ'"
          using 3 iso_char arr_char μ'.arrow_of_spans_axioms by simp
        moreover have "μ = μ'"
          using 3 μ μ' dom_char cod_char by fastforce
        ultimately show "iso μ  iso μ'  μ = μ'"
          by simp
      qed
    qed

    text ‹
      We can now prove the easier half of the main result (CKS Theorem 4):
      If B› is biequivalent to Span(C)›, where C› is a category with pullbacks,
      then B› is a bicategory of spans.
      (Well, it is easier given that we have already done the work to show that the notion
      ``bicategory of spans'' is respected by equivalence of bicategories.)
    ›

    theorem equivalent_implies_bicategory_of_spans:
    assumes "equivalent_bicategories vcomp hcomp assoc unit src trg V1 H1 𝖺1 𝗂1 src1 trg1"
    shows "bicategory_of_spans V1 H1 𝖺1 𝗂1 src1 trg1"
      using assms is_bicategory_of_spans bicategory_of_spans_respects_equivalence by blast

  end

  subsection "Properties of Bicategories of Spans"

  text ‹
    We now develop consequences of the axioms for a bicategory of spans, in preparation for
    proving the other half of the main result.
  ›

  context bicategory_of_spans
  begin

    notation isomorphic  (infix "" 50)

    text ‹
      The following is a convenience version of BS2› that gives us what we generally want:
      given specified f, g› obtain ρ› that makes (f, ρ, g)› a tabulation of g ⋆ f*,
      not a tabulation of some r› isomorphic to g ⋆ f*.
    ›

    lemma BS2':
    assumes "is_left_adjoint f" and "is_left_adjoint g" and "src f = src g"
    and "isomorphic (g  f*) r"
    shows "ρ. tabulation V H 𝖺 𝗂 src trg r ρ f g"
    proof -
      have 1: "is_left_adjoint f  is_left_adjoint g  g  f*  r"
        using assms BS1 by simp
      obtain φ where φ: "«φ : g  f*  r»  iso φ"
        using 1 isomorphic_def by blast
      obtain r' ρ' where ρ': "tabulation V H 𝖺 𝗂 src trg r' ρ' f g"
        using assms 1 BS2 by blast
      interpret ρ': tabulation V H 𝖺 𝗂 src trg r' ρ' f g
        using ρ' by simp
      let  = "ρ'.T0.trnrε r' ρ'"
      have ψ: "« : g  f*  r'»  iso "
        using ρ'.yields_isomorphic_representation by blast
      have "«φ  inv  : r'  r»  iso (φ  inv )"
        using φ ψ isos_compose by blast
      hence 3: "tabulation V H 𝖺 𝗂 src trg r ((φ  inv   f)  ρ') f g"
        using ρ'.is_preserved_by_base_iso by blast
      hence "ρ. tabulation V H 𝖺 𝗂 src trg r ρ f g"
        by blast
      thus ?thesis
        using someI_ex [of "λρ. tabulation V H 𝖺 𝗂 src trg r ρ f g"] by simp
    qed

    text ‹
      The following observation is made by CKS near the beginning of the proof of Theorem 4:
      If w› is an arbitrary 1-cell, and g› and g ⋆ w› are maps, then w› is in fact a map.
      It is applied frequently.
    ›

    lemma BS4:
    assumes "is_left_adjoint g" and "ide w" and "is_left_adjoint (g  w)"
    shows "is_left_adjoint w"
    proof -
      text ‹
        CKS say: ``by (i) there are maps m, n› with w ≅ nm*, so, by (ii), we have two
        tabulations (1, ρ, gw)›, (m, σ, gn)› of gw›; since tabulations are unique
        up to equivalence, m› is invertible and w ≅ nm* is a map.''
      ›
      have ex_ρ: "ρ. tabulation V H 𝖺 𝗂 src trg (g  w) ρ (src w) (g  w)"
      proof -
        have "(g  w)  src w  g  w"
          by (metis assms(3) iso_runit ideD(1) isomorphic_def left_adjoint_is_ide
              runit_in_hom(2) src_hcomp)
        moreover have "(g  w)  (src w)*  g  w"
        proof -
          have "(g  w)  src (g  w)  g  w"
            using calculation isomorphic_implies_ide(2) by auto
          moreover have "src (g  w)  (src w)*"
          proof -
            interpret src_w: map_in_bicategory V H 𝖺 𝗂 src trg src w
              using assms obj_is_self_adjoint by unfold_locales auto
            interpret src_w: adjunction_in_bicategory V H 𝖺 𝗂 src trg
                               src w (src w)* src_w.η src_w.ε
              using src_w.is_map left_adjoint_extends_to_adjunction by simp
            have "adjoint_pair (src w) (src w)"
              using assms obj_is_self_adjoint by simp
            moreover have "adjoint_pair (src w) (src w)*"
              using adjoint_pair_def src_w.adjunction_in_bicategory_axioms by auto
            ultimately have "src w  (src w)*"
              using left_adjoint_determines_right_up_to_iso by simp
            moreover have "src w = src (g  w)"
              using assms isomorphic_def hcomp_simps(1) left_adjoint_is_ide by simp
            ultimately show ?thesis by simp
          qed
          moreover have "src (g  w) = trg (src (g  w))"
            using assms left_adjoint_is_ide by simp
          ultimately show ?thesis
            using assms left_adjoint_is_ide isomorphic_transitive isomorphic_symmetric
                  hcomp_ide_isomorphic
            by blast
        qed
        ultimately show ?thesis
          using assms obj_is_self_adjoint
                left_adjoint_is_ide BS2' [of "src w" "g  w" "g  w"]
          by auto
      qed
      obtain ρ where ρ: "tabulation V H 𝖺 𝗂 src trg (g  w) ρ (src w) (g  w)"
        using ex_ρ by auto
      obtain m n where mn: "is_left_adjoint m  is_left_adjoint n  isomorphic w (n  m*)"
        using assms BS1 [of w] by auto
      have ma: "adjoint_pair m m*  isomorphic w (n  m*)"
        using mn adjoint_pair_def left_adjoint_extends_to_adjoint_pair by blast
      have ex_σ: "σ. tabulation V H 𝖺 𝗂 src trg (g  w) σ m (g  n)"
      proof -
        have "hseq n m*"
          using mn isomorphic_implies_ide by auto
        have "trg (n  m*) = trg w"
          using mn ma isomorphic_def
          by (metis (no_types, lifting) dom_inv in_homE trg_dom trg_inv)
        hence "trg n = trg w"
          using mn by (metis assms(2) ideD(1) trg.preserves_reflects_arr trg_hcomp)
        hence "hseq g n"
          using assms mn left_adjoint_is_ide ideD(1)
          by (metis hseq_char)
        have "hseq g w"
          using assms left_adjoint_is_ide by simp
        have "src m = src n"
          using mn ma hseq n m* adjoint_pair_antipar [of m "m*"] by fastforce

        have "is_left_adjoint (g  n)"
          using assms mn left_adjoints_compose hseq g n by blast
        moreover have "src m = src (g  n)"
          using assms mn hseq g n src m = src n by simp
        moreover have "(g  n)  m*  g  w"
        proof -
          have 1: "src g = trg (n  m*)"
            using assms trg (n  m*) = trg w hseq g w by fastforce
          hence "(g  n)  m*  g  n  m*"
            using assms mn ma assoc_in_hom iso_assoc hseq g n hseq n m*
                  isomorphic_def left_adjoint_is_ide right_adjoint_is_ide
            by (metis hseqE ideD(2) ideD(3))
          also have "...  g  w"
            using assms 1 mn ma isomorphic_symmetric hcomp_ide_isomorphic left_adjoint_is_ide
            by simp
          finally show ?thesis
            using isomorphic_transitive by blast
        qed
        ultimately show ?thesis
          using assms mn ma BS2' by blast
      qed
      obtain σ where σ: "tabulation V H 𝖺 𝗂 src trg (g  w) σ m (g  n)"
        using ex_σ by auto

      interpret ρ: tabulation V H 𝖺 𝗂 src trg g  w ρ src w g  w
        using ρ by auto
      interpret σ: tabulation V H 𝖺 𝗂 src trg g  w σ m g  n
        using σ by auto
      text ‹
        As usual, the sketch given by CKS seems more suggestive than it is a precise recipe.
        We can obtain an equivalence map «e : src w → src m»› and θ› such that
        «θ : m ⋆ e ⇒ src w»›.
        We can also obtain an equivalence map «e' : src m → src w»› and θ'› such that
        «θ' : src w ⋆ e' ⇒ m»›.  If θ'› can be taken to be an isomorphism; then we have
        e' ≅ src w ⋆ e' ≅ m›.  Since e'› is an equivalence, this shows m› is an equivalence,
        hence its right adjoint m* is also an equivalence and therefore a map.
        But w = n ⋆ ma, so this shows that w› is a map.

        Now, we may assume without loss of generality that e› and e'› are part of an
        adjoint equivalence.
        We have «θ : m ⋆ e ⇒ src w»› and «θ' : src w ⋆ e' ⇒ m»›.
        We may take the transpose of θ› to obtain «ζ : m ⇒ src w ⋆ e'»›;
        then «θ' ⋅ ζ : m ⇒ m»› and «ζ ⋅ θ' : src w ⋆ e' ⇒ src w ⋆ e'»›.
        Since m› and src w ⋆ e'› are maps, by BS3› it must be that ζ› and θ'› are inverses.
        ›
        text ‹
          {\bf Note:} CKS don't cite BS3› here.  I am not sure whether this result can be proved
          without BS3›.  For example, I am interested in knowing whether it can still be
          proved under the the assumption that 2-cells between maps are unique, but not
          necessarily invertible, or maybe even in a more general situation.  It looks like
          the invertibility part of BS3› is not used in the proof below.
        ›
      have 2: "e e' φ ψ θ ν θ' ν'.
                  equivalence_in_bicategory (⋅) (⋆) 𝖺 𝗂 src trg e e' φ ψ 
                  «θ' : src w  e'  m»  «ν : g  n  (g  w)  e'»  iso ν 
                  σ = ρ.composite_cell e' θ'  ν 
                  «θ : m  e  src w»  «ν' : g  w  (g  n)  e»  iso ν' 
                  ρ = ((g  w)  θ)  𝖺[g  w, m, e]  (σ  e)  ν'"
        using ρ σ.apex_unique_up_to_equivalence [of ρ "src w" "g  w"] comp_assoc
        by metis
      obtain e e' φ ψ θ ν θ' ν'
        where *: "equivalence_in_bicategory (⋅) (⋆) 𝖺 𝗂 src trg e e' φ ψ 
                  «θ' : src w  e'  m»  «ν : g  n  (g  w)  e'»  iso ν 
                  σ = ρ.composite_cell e' θ'  ν 
                  «θ : m  e  src w»  «ν' : g  w  (g  n)  e»  iso ν' 
                  ρ = σ.composite_cell e θ  ν'"
        using 2 comp_assoc by auto
      interpret ee': equivalence_in_bicategory (⋅) (⋆) 𝖺 𝗂 src trg e e' φ ψ
        using * by simp

      have equiv_e: "equivalence_map e"
        using ee'.equivalence_in_bicategory_axioms equivalence_map_def by auto
      obtain ψ' where ψ': "adjoint_equivalence_in_bicategory (⋅) (⋆) 𝖺 𝗂 src trg e e' φ ψ'"
        using equivalence_refines_to_adjoint_equivalence [of e e' φ]
              ee'.unit_in_hom(2) ee'.unit_is_iso ee'.antipar equiv_e
        by auto
      interpret ee': adjoint_equivalence_in_bicategory (⋅) (⋆) 𝖺 𝗂 src trg e e' φ ψ'
        using ψ' by simp
      interpret e'e: adjoint_equivalence_in_bicategory (⋅) (⋆) 𝖺 𝗂 src trg e' e inv ψ' inv φ
        using * ee'.dual_adjoint_equivalence by simp
      have equiv_e': "equivalence_map e'"
        using e'e.equivalence_in_bicategory_axioms equivalence_map_def by auto

      have "hseq m e"
        using * ide_dom [of θ]
        by (elim conjE in_homE) simp
      have "hseq (src w) e'"
        using * ide_dom [of θ']
        by (elim conjE in_homE) simp

      have "e'e.trnrη m θ  hom m (src w  e')"
      proof -
        have "src m = trg e"
          using hseq m e by auto
        moreover have "src (src w) = trg e'"
          using hseq (src w) e' by auto
        moreover have "ide m"
          using mn left_adjoint_is_ide by simp
        moreover have "ide (src w)"
          using assms by simp
        ultimately show ?thesis
          using * e'e.adjoint_transpose_right(1) by blast
      qed
      hence 3: "«e'e.trnrη m θ : m  src w  e'»"
        by simp
      hence "«θ'  e'e.trnrη m θ : m  m»  «e'e.trnrη m θ  θ' : src w  e'  src w  e'»"
        using * by auto
      moreover have "«m : m  m»  «src w  e' : src w  e'  src w  e'»"
        using mn 3 ide_cod [of "e'e.trnrη m θ"] left_adjoint_is_ide by fastforce
      moreover have 4: "is_left_adjoint (src w  e')"
      proof -
        have "is_left_adjoint (src w)"
          using assms obj_is_self_adjoint by simp
        moreover have "is_left_adjoint e'"
          using e'e.adjunction_in_bicategory_axioms adjoint_pair_def by auto
        ultimately show ?thesis
          using left_adjoints_compose hseq (src w) e' by auto
      qed
      ultimately have "θ'  e'e.trnrη m θ = m  e'e.trnrη m θ  θ' = src w  e'"
        using mn BS3 [of m m "θ'  e'e.trnrη m θ" m]
              BS3 [of "src w  e'" "src w  e'" "e'e.trnrη m θ  θ'" "src w  e'"]
        by auto
      hence "inverse_arrows θ' (e'e.trnrη m θ)"
        using mn 4 left_adjoint_is_ide inverse_arrows_def by simp
      hence 5: "iso θ'"
        by auto
      have "equivalence_map (src w  e')"
        using assms obj_is_equivalence_map equiv_e' hseq (src w) e' equivalence_maps_compose
        by auto
      hence "equivalence_map m"
        using * 5 equivalence_map_preserved_by_iso isomorphic_def by auto
      hence "equivalence_map m*"
        using mn ma right_adjoint_to_equivalence_is_equivalence by simp
      hence "is_left_adjoint m*"
        using equivalence_is_left_adjoint by simp
      moreover have "hseq n m*"
        using mn isomorphic_implies_ide by auto
      ultimately have "is_left_adjoint (n  m*)"
        using mn left_adjoints_compose by blast
      thus ?thesis
        using mn left_adjoint_preserved_by_iso isomorphic_def isomorphic_symmetric
        by metis
    qed

  end

  subsection "Choosing Tabulations"

  context bicategory_of_spans
  begin

    notation isomorphic  (infix "" 50)
    notation iso_class ("_")

    text ‹
      We will ultimately need to have chosen a specific tabulation for each 1-cell.
      This has to be done carefully, to avoid unnecessary choices.
      We start out by using BS1› to choose a specific factorization of the form
      r ≅ tab1 r ⋆ (tab0 r)* for each 1-cell r›.  This has to be done in such a way
      that all elements of an isomorphism class are assigned the same factorization.
    ›

    abbreviation isomorphic_rep
    where "isomorphic_rep r f g  is_left_adjoint f  is_left_adjoint g  g  f*  r"

    definition tab0
    where "tab0 r  SOME f. g. isomorphic_rep (iso_class_rep r) f g"

    definition tab1
    where "tab1 r  SOME g. isomorphic_rep (iso_class_rep r) (tab0 r) g"

    definition rep
    where "rep r  SOME φ. «φ : tab1 r  (tab0 r)*  r»  iso φ"

    lemma rep_props:
    assumes "ide r"
    shows "«rep r : tab1 r  (tab0 r)*  r»" and "iso (rep r)"
    and "r  iso_class_rep r"
    and "isomorphic_rep r (tab0 r) (tab1 r)"
    and "tab1 r  (tab0 r)*  r"
    proof -
      have 1: "isomorphic_rep r (tab0 r) (tab1 r)"
      proof -
        have "f g. isomorphic_rep (iso_class_rep r) f g"
          using assms BS1 isomorphic_symmetric rep_iso_class isomorphic_transitive
          by blast
        hence "isomorphic_rep (iso_class_rep r) (tab0 r) (tab1 r)"
          using assms tab0_def tab1_def 
                someI_ex [of "λf. g. isomorphic_rep (iso_class_rep r) f g"]
                someI_ex [of "λg. isomorphic_rep (iso_class_rep r) (tab0 r) g"]
          by simp
        thus ?thesis
          using assms isomorphic_symmetric isomorphic_transitive rep_iso_class by blast
      qed
      hence "φ. «φ : tab1 r  (tab0 r)*  r»  iso φ"
        using isomorphic_def by blast
      hence 2: "«rep r : tab1 r  (tab0 r)*  r»  iso (rep r)"
        using someI_ex [of "λφ. «φ : tab1 r  (tab0 r)*  r»  iso φ"] rep_def by auto
      show "«rep r : tab1 r  (tab0 r)*  r»"
        using 2 by simp
      show "iso (rep r)"
        using 2 by simp
      show "r  iso_class_rep r"
        using assms rep_iso_class isomorphic_symmetric by simp
      thus "isomorphic_rep r (tab0 r) (tab1 r)"
        using 1 isomorphic_transitive by blast
      thus "tab1 r  (tab0 r)*  r"
        by simp
    qed

    lemma tab0_in_hom [intro]:
    assumes "ide r"
    shows "«tab0 r : src (tab0 r)  src r»"
    and "«tab0 r : tab0 r  tab0 r»"
    proof -
      show "«tab0 r : tab0 r  tab0 r»"
        using assms rep_props left_adjoint_is_ide by auto
      have "trg (tab0 r) = src r"
        using assms rep_props
        by (metis ideD(1) isomorphic_implies_hpar(1) isomorphic_implies_hpar(3)
            right_adjoint_simps(2) src_hcomp)
      thus "«tab0 r : src (tab0 r)  src r»"
        using assms rep_props left_adjoint_is_ide
        by (intro in_hhomI, auto)
    qed

    lemma tab0_simps [simp]:
    assumes "ide r"
    shows "ide (tab0 r)"
    and "is_left_adjoint (tab0 r)"
    and "trg (tab0 r) = src r"
    and "dom (tab0 r) = tab0 r" and "cod (tab0 r) = tab0 r"
      using assms tab0_in_hom rep_props ide_dom left_adjoint_is_ide by auto

    lemma tab1_in_hom [intro]:
    assumes "ide r"
    shows "«tab1 r : src (tab0 r)  trg r»"
    and "«tab1 r : tab1 r  tab1 r»"
    proof -
      show "«tab1 r : tab1 r  tab1 r»"
        using assms rep_props left_adjoint_is_ide by auto
      have "trg (tab1 r) = trg r"
        using assms rep_props
        by (metis ideD(1) isomorphic_implies_hpar(1) isomorphic_implies_hpar(4) trg_hcomp)
      moreover have "src (tab0 r) = src (tab1 r)"
        using assms rep_props by fastforce
      ultimately show "«tab1 r : src (tab0 r)  trg r»"
        using assms rep_props left_adjoint_is_ide
        by (intro in_hhomI, auto)
    qed

    lemma tab1_simps [simp]:
    assumes "ide r"
    shows "ide (tab1 r)"
    and "is_left_adjoint (tab1 r)"
    and "src (tab1 r) = src (tab0 r)" and "trg (tab1 r) = trg r"
    and "dom (tab1 r) = tab1 r" and "cod (tab1 r) = tab1 r"
      using assms tab1_in_hom rep_props ide_dom left_adjoint_is_ide by auto

    lemma rep_in_hom [intro]:
    assumes "ide r"
    shows "«rep r : src r  trg r»"
    and "«rep r : tab1 r  (tab0 r)*  r»"
    proof -
      show "«rep r : tab1 r  (tab0 r)*  r»"
        using assms rep_props by auto
      thus "«rep r : src r  trg r»"
        using arrI vconn_implies_hpar(1-4) by force
    qed

    lemma rep_simps [simp]:
    assumes "ide r"
    shows "arr (rep r)"
    and "src (rep r) = src r" and "trg (rep r) = trg r"
    and "dom (rep r) = tab1 r  (tab0 r)*" and "cod (rep r) = r"
     using assms rep_in_hom by auto

    lemma iso_rep:
    assumes "ide r"
    shows "iso (rep r)"
      using assms rep_props by simp

  end

  text ‹
    Next, we assign a specific tabulation to each 1-cell r.
    We can't just do this any old way if we ultimately expect to obtain a mapping that is
    functorial with respect to vertical composition.  What we have to do is to assign the
    representative tab1 r ⋆ (tab0 r)* its canonical tabulation, obtained as the adjoint
    transpose of the identity, and then translate this to a tabulation of r› via the chosen
    isomorphism «rep r : tab1 r ⋆ (tab0 r)* ⇒ r»›.
  ›

  locale identity_in_bicategory_of_spans =
    bicategory_of_spans +
  fixes r :: 'a
  assumes is_ide: "ide r"
  begin

    interpretation tab0: map_in_bicategory V H 𝖺 𝗂 src trg tab0 r
      using is_ide rep_props by unfold_locales auto
    interpretation tab1: map_in_bicategory V H 𝖺 𝗂 src trg tab1 r
      using is_ide rep_props by unfold_locales auto

    text ‹
      A tabulation (tab0 r, tab, tab1 r)› of r› can be obtained as the adjoint transpose
      of the isomorphism «rep r : (tab1 r) ⋆ (tab0 r)* ⇒ r»›.  It is essential to define
      it this way if we expect the mapping from 2-cells of the underlying bicategory
      to arrows of spans to preserve vertical composition.
    ›

    definition tab
    where "tab  tab0.trnrη (tab1 r) (rep r)"

    text ‹
      In view of BS2'›, the 1-cell (tab1 r) ⋆ (tab0 r)* has the canonical tabulation
      obtained via adjoint transpose of an identity.  In fact, this tabulation generates the
      chosen tabulation of r› in the same isomorphism class by translation along the
      isomorphism «rep r : (tab1 r) ⋆ (tab0 r)* ⇒ r»›.  This fact is used to show that the
      mapping from 2-cells to arrows of spans preserves identities.
    ›

    lemma canonical_tabulation:
    shows "tabulation V H 𝖺 𝗂 src trg
             ((tab1 r)  (tab0 r)*) (tab0.trnrη (tab1 r) ((tab1 r)  (tab0 r)*)) (tab0 r) (tab1 r)"
    proof -
      have "ρ. tabulation V H 𝖺 𝗂 src trg ((tab1 r)  (tab0 r)*) ρ (tab0 r) (tab1 r)"
        by (simp add: bicategory_of_spans.BS2' bicategory_of_spans_axioms is_ide
            isomorphic_reflexive)
      thus ?thesis
        using is_ide tab0.canonical_tabulation by simp
    qed

    lemma tab_def_alt:
    shows "tab = (rep r  tab0 r)  tab0.trnrη (tab1 r) ((tab1 r)  (tab0 r)*)"
    and "(inv (rep r)  tab0 r)  tab = tab0.trnrη (tab1 r) ((tab1 r)  (tab0 r)*)"
    proof -
      have "tab = tab0.trnrη (tab1 r) (rep r  ((tab1 r)  (tab0 r)*))"
        using tab_def is_ide rep_in_hom [of r] comp_arr_dom by auto
      also have "... = (rep r  tab0 r)  tab0.trnrη (tab1 r) ((tab1 r)  (tab0 r)*)"
        using is_ide tab0.trnrη_comp by auto
      finally show 1: "tab = (rep r  tab0 r)  tab0.trnrη (tab1 r) ((tab1 r)  (tab0 r)*)" by simp
      have "(inv (rep r)  tab0 r)  tab =
            ((inv (rep r)  tab0 r)  (rep r  tab0 r))  tab0.trnrη (tab1 r) ((tab1 r)  (tab0 r)*)"
        unfolding 1 using comp_assoc by presburger
      also have "... = tab0.trnrη (tab1 r) ((tab1 r)  (tab0 r)*)"
      proof - 
        have 1: "(inv (rep r)  tab0 r)  (rep r  tab0 r) = ((tab1 r)  (tab0 r)*)  tab0 r"
          using whisker_right [of "tab0 r" "inv (rep r)" "rep r"] iso_rep rep_in_hom
                inv_is_inverse comp_inv_arr
          by (simp add: comp_inv_arr' is_ide)
        show ?thesis
        proof -
          have "«tab0.trnrη (tab1 r) ((tab1 r)  (tab0 r)*) :
                   tab1 r  (tab1 r  (tab0 r)*)  tab0 r»"
            by (meson canonical_tabulation tabulation_data.tab_in_hom(2) tabulation_def)
          hence "((tab1 r  (tab0 r)*)  tab0 r)  tab0.trnrη (tab1 r) ((tab1 r)  (tab0 r)*) =
                 tab0.trnrη (tab1 r) ((tab1 r)  (tab0 r)*)"
            using 1 comp_cod_arr by blast
          thus ?thesis
            using 1 by simp
        qed
      qed
      finally show "(inv (rep r)  tab0 r)  tab = tab0.trnrη (tab1 r) ((tab1 r)  (tab0 r)*)"
        by blast
    qed

    lemma tab_is_tabulation:
    shows "tabulation V H 𝖺 𝗂 src trg r tab (tab0 r) (tab1 r)"
      by (metis bicategory_of_spans.iso_rep bicategory_of_spans.rep_in_hom(2)
          bicategory_of_spans_axioms is_ide canonical_tabulation tab_def_alt(1)
          tabulation.is_preserved_by_base_iso)

    (*
     * TODO: If I pull the interpretation "tab" out of the following, Isabelle warns that
     * the lemma is a redundant introduction rule and is being "ignored" for that purpose.
     * However, the redundancy is only in the present context: if the enclosing locale is
     * interpreted elsewhere, then the rule is not redundant.  In order to make sure that
     * the rule is not "ignored", I have put the interpretation "tab" into the proof to
     * avoid the warning.
     *)

    lemma tab_in_hom [intro]:
    shows "«tab : src (tab0 r)  trg r»"
    and "«tab : tab1 r  r  tab0 r»"
    proof -
      interpret tab: tabulation V H 𝖺 𝗂 src trg r tab tab0 r tab1 r
        using tab_is_tabulation by simp
      show "«tab : src (tab0 r)  trg r»"
        using tab.tab_in_hom by auto
      show "«tab : tab1 r  r  tab0 r»"
        using tab.tab_in_hom by auto
    qed

    lemma tab_simps [simp]:
    shows "arr tab"
    and "src tab = src (tab0 r)" and "trg tab = trg r"
    and "dom tab = tab1 r" and "cod tab = r  tab0 r"
      using tab_in_hom by auto

  end

  text ‹
    The following makes the chosen tabulation conveniently available whenever we are
    considering a particular 1-cell.
  ›

  sublocale identity_in_bicategory_of_spans  tabulation V H 𝖺 𝗂 src trg r tab tab0 r tab1 r
    using is_ide tab_is_tabulation by simp

  context identity_in_bicategory_of_spans
  begin

    interpretation tab0: map_in_bicategory V H 𝖺 𝗂 src trg tab0 r
      using is_ide rep_props by unfold_locales auto
    interpretation tab1: map_in_bicategory V H 𝖺 𝗂 src trg tab1 r
      using is_ide rep_props by unfold_locales auto

    text ‹
      The fact that adjoint transpose is a bijection allows us to invert the definition
      of tab› in terms of rep› to express rep in terms of tab.
    ›

    lemma rep_in_terms_of_tab:
    shows "rep r = T0.trnrε r tab"
      using is_ide T0.adjoint_transpose_right(3) [of r "tab1 r" "rep r"] tab_def
      by fastforce

    lemma isomorphic_implies_same_tab:
    assumes "isomorphic r r'"
    shows "tab0 r = tab0 r'" and "tab1 r = tab1 r'"
      using assms tab0_def tab1_def iso_class_eqI by auto

    text ‹
      ``Every 1-cell has a tabulation as a span of maps.''
      Has a nice simple ring to it, but maybe not so useful for us, since we generally
      really need to know that the tabulation has a specific form.
    ›

    lemma has_tabulation:
    shows "ρ f g. is_left_adjoint f  is_left_adjoint g  tabulation V H 𝖺 𝗂 src trg r ρ f g"
      using is_ide tab_is_tabulation rep_props by blast

  end

  subsection "Tabulations in a Bicategory of Spans"

  context bicategory_of_spans
  begin

    abbreviation tab_of_ide
    where "tab_of_ide r  identity_in_bicategory_of_spans.tab V H 𝖺 𝗂 src trg r"

    abbreviation prj0
    where "prj0 h k  tab0 (k*  h)"

    abbreviation prj1
    where "prj1 h k  tab1 (k*  h)"

    lemma prj_in_hom [intro]:
    assumes "ide h" and "is_left_adjoint k" and "trg h = trg k"
    shows "«prj0 h k : src (prj0 h k)  src h»"
    and "«prj1 h k : src (prj0 h k)  src k»"
    and "«prj0 h k : prj0 h k  prj0 h k»"
    and "«prj1 h k : prj1 h k  prj1 h k»"
      by (intro in_hhomI, auto simp add: assms(1-3))

    lemma prj_simps [simp]:
    assumes "ide h" and "is_left_adjoint k" and "trg h = trg k"
    shows "trg (prj0 h k) = src h"
    and "src (prj1 h k) = src (prj0 h k)" and "trg (prj1 h k) = src k"
    and "dom (prj0 h k) = prj0 h k" and "cod (prj0 h k) = prj0 h k"
    and "dom (prj1 h k) = prj1 h k" and "cod (prj1 h k) = prj1 h k"
    and "is_left_adjoint (prj0 h k)" and "is_left_adjoint (prj1 h k)"
      using assms prj_in_hom by auto

  end

  text ‹
    Many of the commutativity conditions that we would otherwise have to worry about
    when working with tabulations in a bicategory of spans reduce to trivialities.
    The following locales try to exploit this to make our life more manageable.
  ›

  locale span_of_maps =
    bicategory_of_spans +
  fixes leg0 :: 'a
  and leg1 :: 'a
  assumes leg0_is_map: "is_left_adjoint leg0"
  and leg1_is_map : "is_left_adjoint leg1"

  text ‹
    The purpose of the somewhat strange-looking assumptions in this locale is
    to cater to the form of data that we obtain from T1›.  Under the assumption
    that we are in a bicategory of spans and that the legs of r› and s› are maps,
    the hypothesized 2-cells will be uniquely determined isomorphisms, and an
    arrow of spans w› from r› to s› will be a map.  We want to prove this once and
    for all under the weakest assumptions we can manage.
  ›

  locale arrow_of_spans_of_maps =
    bicategory_of_spans V H 𝖺 𝗂 src trg +
    r: span_of_maps V H 𝖺 𝗂 src trg r0 r1 +
    s: span_of_maps V H 𝖺 𝗂 src trg s0 s1
  for V :: "'a comp"                 (infixr "" 55)
  and H :: "'a  'a  'a"          (infixr "" 53)
  and 𝖺 :: "'a  'a  'a  'a"    ("𝖺[_, _, _]")
  and 𝗂 :: "'a  'a"                ("𝗂[_]")
  and src :: "'a  'a"
  and trg :: "'a  'a"
  and r0 :: 'a
  and r1 :: 'a
  and s0 :: 'a
  and s1 :: 'a
  and w :: 'a +
  assumes is_ide: "ide w"
  and leg0_lax: "θ. «θ : s0  w  r0»"
  and leg1_iso: "ν. «ν : r1  s1  w»  iso ν"
  begin

    notation isomorphic  (infix "" 50)

    lemma composite_leg1_is_map:
    shows "is_left_adjoint (s1  w)"
      using r.leg1_is_map leg1_iso left_adjoint_preserved_by_iso' isomorphic_def
            isomorphic_symmetric
      by auto

    lemma is_map:
    shows "is_left_adjoint w"
      using is_ide composite_leg1_is_map s.leg1_is_map BS4 [of s1 w] by auto

    lemma hseq_leg0:
    shows "hseq s0 w"
      by (metis ideD(1) ide_dom in_homE leg0_lax)

    lemma composite_with_leg0_is_map:
    shows "is_left_adjoint (s0  w)"
      using left_adjoints_compose is_map s.leg0_is_map hseq_leg0 by blast

    lemma leg0_uniquely_isomorphic:
    shows "s0  w  r0"
    and "∃!θ. «θ : s0  w  r0»"
    proof -
      show 1: "s0  w  r0"
        using leg0_lax composite_with_leg0_is_map r.leg0_is_map BS3 [of "s0  w" r0]
              isomorphic_def
        by auto
      have "θ. «θ : s0  w  r0»  iso θ"
        using 1 isomorphic_def by simp
      moreover have "θ θ'. «θ : s0  w  r0»  «θ' : s0  w  r0»  θ = θ'"
        using BS3 r.leg0_is_map composite_with_leg0_is_map by blast
      ultimately show "∃!θ. «θ : s0  w  r0»" by blast
    qed

    lemma leg1_uniquely_isomorphic:
    shows "r1  s1  w"
    and "∃!ν. «ν : r1  s1  w»"
    proof -
      show 1: "r1  s1  w"
        using leg1_iso isomorphic_def by auto
      have "ν. «ν : r1  s1  w»  iso ν"
        using leg1_iso isomorphic_def isomorphic_symmetric by simp
      moreover have "ν ν'. «ν : r1  s1  w»  «ν' : r1  s1  w»  ν = ν'"
        using BS3 r.leg1_is_map composite_leg1_is_map by blast
      ultimately show "∃!ν. «ν : r1  s1  w»" by blast
    qed

    definition the_θ
    where "the_θ  THE θ. «θ : s0  w  r0»"

    definition the_ν
    where "the_ν  THE ν. «ν : r1  s1  w»"

    lemma the_θ_props:
    shows "«the_θ : s0  w  r0»" and "iso the_θ"
    proof -
      show "«the_θ : s0  w  r0»"
        unfolding the_θ_def
        using the1I2 [of "λθ. «θ : s0  w  r0»" "λθ. «θ : s0  w  r0»"]
              leg0_uniquely_isomorphic
        by simp
      thus "iso the_θ"
        using BS3 r.leg0_is_map composite_with_leg0_is_map by simp
    qed

    lemma the_θ_in_hom [intro]:
    shows "«the_θ : src r0  trg r0»"
    and "«the_θ : s0  w  r0»"
      using the_θ_props apply auto
      by (metis cod_trg in_hhom_def in_homE isomorphic_implies_hpar(3) leg0_uniquely_isomorphic(1)
          src_dom trg.preserves_cod)

    lemma the_θ_simps [simp]:
    shows "arr the_θ"
    and "src the_θ = src r0" and "trg the_θ = trg r0"
    and "dom the_θ = s0  w" and "cod the_θ = r0"
      using the_θ_in_hom by auto

    lemma the_ν_props:
    shows "«the_ν : r1  s1  w»" and "iso the_ν"
    proof -
      show "«the_ν : r1  s1  w»"
        unfolding the_ν_def
        using the1I2 [of "λν. «ν : r1  s1  w»" "λν. «ν : r1  s1  w»"]
              leg1_uniquely_isomorphic
        by simp
      thus "iso the_ν"
        using BS3 r.leg1_is_map composite_leg1_is_map by simp
    qed

    lemma the_ν_in_hom [intro]:
    shows "«the_ν : src r1  trg r1»"
    and "«the_ν : r1  s1  w»"
      using the_ν_props apply auto
      by (metis in_hhom_def in_homE isomorphic_implies_hpar(3) leg1_uniquely_isomorphic(1)
          src_cod trg_dom)

    lemma the_ν_simps [simp]:
    shows "arr the_ν"
    and "src the_ν = src r1" and "trg the_ν = trg r1"
    and "dom the_ν = r1" and "cod the_ν = s1  w"
      using the_ν_in_hom by auto

  end

  (*
   * TODO: I could probably avoid repeating the declarations of the locale parameters
   * if I was willing to accept them being given in their order of appearance.
   *)

  locale arrow_of_spans_of_maps_to_tabulation_data =
    bicategory_of_spans V H 𝖺 𝗂 src trg +
    arrow_of_spans_of_maps V H 𝖺 𝗂 src trg r0 r1 s0 s1 w +
    σ: tabulation_data V H 𝖺 𝗂 src trg s σ s0 s1
  for V :: "'a comp"                 (infixr "" 55)
  and H :: "'a  'a  'a"          (infixr "" 53)
  and 𝖺 :: "'a  'a  'a  'a"     ("𝖺[_, _, _]")
  and 𝗂 :: "'a  'a"                 ("𝗂[_]")
  and src :: "'a  'a"
  and trg :: "'a  'a"
  and r0 :: 'a
  and r1 :: 'a
  and s :: 'a
  and σ :: 'a
  and s0 :: 'a
  and s1 :: 'a
  and w :: 'a

  text ‹
    The following declaration allows us to inherit the rules and other facts defined in
    locale @{locale uwθ}.  It is tedious to prove very much without these in place.
  ›

  sublocale arrow_of_spans_of_maps_to_tabulation_data  uwθ V H 𝖺 𝗂 src trg s σ s0 s1 r0 w the_θ
    using σ.tab_in_hom is_ide the_θ_props by unfold_locales auto

  locale arrow_of_spans_of_maps_to_tabulation =
    arrow_of_spans_of_maps_to_tabulation_data +
    tabulation V H 𝖺 𝗂 src trg s σ s0 s1

  locale tabulation_in_maps =
    span_of_maps V H 𝖺 𝗂 src trg s0 s1 +
    tabulation V H 𝖺 𝗂 src trg s σ s0 s1
  for V :: "'a comp"                 (infixr "" 55)
  and H :: "'a  'a  'a"          (infixr "" 53)
  and 𝖺 :: "'a  'a  'a  'a"     ("𝖺[_, _, _]")
  and 𝗂 :: "'a  'a"                 ("𝗂[_]")
  and src :: "'a  'a"
  and trg :: "'a  'a"
  and s :: 'a
  and σ :: 'a
  and s0 :: 'a
  and s1 :: 'a

  sublocale tabulation_in_maps  tabulation V H 𝖺 𝗂 src trg s σ s0 s1 ..

  sublocale identity_in_bicategory_of_spans 
              tabulation_in_maps V H 𝖺 𝗂 src trg r tab tab0 r tab1 r
    using is_ide rep_props by unfold_locales auto

  locale cospan_of_maps_in_bicategory_of_spans =
    bicategory_of_spans +
    fixes h :: 'a
    and k :: 'a
    assumes h_is_map: "is_left_adjoint h"
    and k_is_map: "is_left_adjoint k"
    and cospan: "trg h = trg k"
  begin

    text ‹
      The following sublocale declaration is perhaps pushing the limits of sensibility,
      but the purpose is, given a cospan of maps (h, k)›, to obtain ready access to the
      composite k* ⋆ h› and its chosen tabulation.
    ›

    sublocale identity_in_bicategory_of_spans V H 𝖺 𝗂 src trg k*  h
      using h_is_map k_is_map cospan left_adjoint_is_ide
      by unfold_locales auto

    notation isomorphic  (infix "" 50)

    interpretation E: self_evaluation_map V H 𝖺 𝗂 src trg ..
    notation E.eval ("_")

    interpretation h: map_in_bicategory V H 𝖺 𝗂 src trg h
      using h_is_map by unfold_locales auto
    interpretation k: map_in_bicategory V H 𝖺 𝗂 src trg k
      using k_is_map by unfold_locales auto

    text ‹
      Our goal here is to reformulate the biuniversal properties of the chosen tabulation
      of k* ⋆ h› in terms of its transpose, which yields a 2-cell from k ⋆ tab1 (k* ⋆ h)›
      to h ⋆ tab0 (k* ⋆ h)›.  These results do not depend on BS3›.
    ›

    abbreviation p0
    where "p0  prj0 h k"

    abbreviation p1
    where "p1  prj1 h k"

    lemma p0_in_hom [intro]:
    shows "«p0 : src p0  src h»"
      by auto

    lemma p1_in_hom [intro]:
    shows "«p1 : src p0  src k»"
      using prj_in_hom cospan h.ide_left k_is_map by blast

    lemma p0_simps [simp]:
    shows "trg p0 = src h"
      by simp

    lemma p1_simps [simp]:
    shows "trg p1 = src k"
      using k.antipar(1) by auto

    definition φ
    where "φ  k.trnlε (h  p0) (𝖺[k*, h, p0]  tab)"

    lemma φ_in_hom [intro]:
    shows "«φ : src p0  trg h»"
    and "«φ : k  p1  h  p0»"
    proof -
      show 1: "«φ : k  p1  h  p0»"
        unfolding φ_def
        using k.antipar cospan k.adjoint_transpose_left(2) [of "h  p0" "p1"]
        by fastforce
      show "«φ : src p0  trg h»"
        using 1 k.antipar arrI cospan vconn_implies_hpar(1-2) by force
    qed

    lemma φ_simps [simp]:
    shows "arr φ"
    and "src φ = src p0" and "trg φ = trg h"
    and "dom φ = k  p1" and "cod φ = h  p0"
      using φ_in_hom by auto

    lemma transpose_φ:
    shows "tab = 𝖺-1[k*, h, p0]  k.trnlη p1 φ"
    proof -
      have "𝖺-1[k*, h, p0]  k.trnlη p1 φ = 𝖺-1[k*, h, p0]  𝖺[k*, h, p0]  tab"
        unfolding φ_def
        using k.antipar cospan
              k.adjoint_transpose_left(4)
                [of "h  p0" "p1" "𝖺[k*, h, p0]  tab"]
        by fastforce
      also have "... = (𝖺-1[k*, h, p0]  𝖺[k*, h, p0])  tab"
        using comp_assoc by presburger
      also have "... = tab"
        using k.antipar cospan comp_cod_arr comp_assoc_assoc' by simp
      finally show ?thesis by simp
    qed

    lemma transpose_triangle:
    assumes "ide w"
    and "«θ : p0  w  u»" and "«ν : v  p1  w»"
    shows "k.trnlε (h  u) (𝖺[k*, h, u]  ((k*  h)  θ)  𝖺[k*  h, p0, w]  (tab  w)  ν) =
          (h  θ)  𝖺[h, p0, w]  (φ  w)  𝖺-1[k, p1, w]  (k  ν)"
    proof -
      have u: "ide u"
        using assms(2) by auto
      have v: "ide v"
        using assms(3) by auto
      have 0: "src p0 = trg w"
        by (metis assms(2) hseqE ideD(1) src.preserves_reflects_arr u vconn_implies_hpar(3))
      have 1: "src h = trg u"
        using assms(1-2) 0 trg_dom trg_cod vconn_implies_hpar(4) by auto
      have 2: "src k = trg v"
        using assms(1,3) 0 trg_dom trg_cod hseqI'
        by (metis ideD(1) leg1_simps(2) leg1_simps(3) p1_simps trg_hcomp vconn_implies_hpar(4))
      have 3: "src u = src v  src u = src w"
        using assms 0 k.antipar src_dom src_cod hseqI'
        by (metis ideD(1) leg0_simps(2) leg1_simps(2) leg1_simps(3) src_hcomp
            vconn_implies_hpar(3))
      have 4: "src h = trg θ"
        using assms 1 k.antipar by auto
      define χ where "χ = 𝖺[k*, h, p0  w]  𝖺[k*  h, p0, w]  (tab  w)"
      have χ: "«χ : p1  w  k*  h  p0  w»"
        unfolding χ_def
        using assms 0 k.antipar cospan by (intro comp_in_homI, auto)
      have "k.trnlε (h  u) (𝖺[k*, h, u]  ((k*  h)  θ)  𝖺[k*  h, p0, w]  (tab  w)  ν) =
            k.trnlε (h  u) ((k*  h  θ)  χ  ν)"
        unfolding χ_def
        using assms 1 k.antipar cospan assoc_naturality [of "k*" h θ] comp_assoc
        by (metis "4" h.ide_left ide_char in_homE k.ide_right)
      also have "... = k.trnlε (h  u) (k*  h  θ)  (k  χ  ν)"
      proof -
        have "ide (h  u)"
          using "1" u assms h.ide_left by blast
        moreover have "seq (k*  h  θ) (χ  ν)"
          using assms 1 k.antipar cospan χ seqI'
          apply (intro seqI)
                  apply auto
            apply blast
        proof -
          have "dom (k*  h  θ) = k*  h  p0  w"
             using assms
             by (metis "4" cospan hcomp_simps(2-3) h.ide_left hseqI' ide_char in_homE
                 k.antipar(2) k.ide_right) 
          also have "... = cod χ"
             using χ by auto
          finally show "dom (k*  h  θ) = cod χ" by simp
        qed
        moreover have "src k = trg (k*  h  θ)"
          using assms k.antipar cospan calculation(2) by auto
        ultimately show ?thesis
          using k.trnlε_comp by simp
      qed
      also have "... = k.trnlε (h  u) (k*  h  θ)  (k  χ)  (k  ν)"
        using assms u χ whisker_left
        by (metis k.ide_left seqI')
      also have
        "... = (𝗅[h  u]  (k.ε  h  u)  𝖺-1[k, k*, h  u]  (k  k*  h  θ))  (k  χ)  (k  ν)"
        unfolding k.trnlε_def by simp
      also have "... = (h  θ) 
                         (𝗅[h  p0  w]  (k.ε  h  p0  w) 
                         𝖺-1[k, k*, h  p0  w]  (k  χ)) 
                         (k  ν)"
      proof -
        have "𝗅[h  u]  (k.ε  h  u)  𝖺-1[k, k*, h  u]  (k  k*  h  θ) =
              𝗅[h  u]  (k.ε  h  u)  ((k  k*)  h  θ)  𝖺-1[k, k*, h  p0  w]"
          using assms 4 k.antipar cospan assoc'_naturality [of k "k*" "h  θ"]
          by fastforce
        also have "... = 𝗅[h  u]  ((k.ε  h  u)  ((k  k*)  h  θ))  𝖺-1[k, k*, h  p0  w]"
          using comp_assoc by presburger
        also have "... = (𝗅[h  u]  (trg k  h  θ))  (k.ε  h  p0  w) 
                           𝖺-1[k, k*, h  p0  w]"
        proof -
          have "(k.ε  h  u)  ((k  k*)  h  θ) = k.ε  (k  k*)  (h  u)  (h  θ)"
            using assms 1 k.antipar cospan interchange comp_arr_dom comp_cod_arr
            by fastforce
          also have "... = k.ε  h  θ"
            using assms k.antipar cospan comp_arr_dom comp_cod_arr k.counit_in_hom
                  whisker_left
            by (metis h.ide_left in_homE)
          also have "... = (trg k  h  θ)  (k.ε  h  p0  w)"
            using assms 4 k.antipar cospan whisker_left comp_arr_dom comp_cod_arr
                  interchange [of "trg k" k.ε "h  θ" "h  p0  w"]
            by auto
          finally have "(k.ε  h  u)  ((k  k*)  h  θ) = (trg k  h  θ)  (k.ε  h  p0  w)"
            by simp
          thus ?thesis
            using comp_assoc by presburger
        qed
        also have "... = (h  θ)  𝗅[h  p0  w]  (k.ε  h  p0  w)  𝖺-1[k, k*, h  p0  w]"
        proof -
          have "𝗅[h  u]  (trg k  h  θ) = (h  θ)  𝗅[h  p0  w]"
            using assms 1 4 k.antipar cospan lunit_naturality [of "h  θ"]
            by (metis hcomp_simps(3-4) h.ide_left hseqI' ide_char in_homE trg_hcomp)
          thus ?thesis
            using comp_assoc by presburger
        qed
        finally have "𝗅[h  u]  (k.ε  h  u)  𝖺-1[k, k*, h  u]  (k  k*  h  θ) =
                      (h  θ)  𝗅[h  p0  w]  (k.ε  h  p0  w) 
                        𝖺-1[k, k*, h  p0  w]"
          by simp
        thus ?thesis
          using comp_assoc by presburger
      qed
      also have "... = (h  θ)  𝖺[h, p0, w]  (φ  w)  𝖺-1[k, p1, w]  (k  ν)"
      proof -
        have "𝗅[h  p0  w]  (k.ε  h  p0  w) 
               𝖺-1[k, k*, h  p0  w] 
               (k  𝖺[k*, h, p0  w]  𝖺[k*  h, p0, w]  (tab  w)) =
              𝖺[h, p0, w]  𝗅[(h  p0)  w] 
               (trg h  𝖺-1[h, p0, w])  (k.ε  h  p0  w) 
               𝖺-1[k, k*, h  p0  w] 
               (k  𝖺[k*, h, p0  w]  𝖺[k*  h, p0, w]  (tab  w))"
        proof -
          have "𝗅[h  p0  w] =
                𝖺[h, p0, w]  𝗅[(h  p0)  w]  (trg h  𝖺-1[h, p0, w])"
          proof -
            have "𝖺[h, p0, w]  𝗅[(h  p0)  w]  (trg h  𝖺-1[h, p0, w]) =
                  𝖺[h, p0, w]  𝔩 ((h  p0)  w)  (trg h  𝖺-1[h, p0, w])"
              using assms 0 k.antipar cospan comp_cod_arr 𝔩_ide_simp by simp
            also have "... = 𝖺[h, p0, w]  𝔩 (𝖺-1[h, p0, w])"
              using assms 0 k.antipar cospan 𝔩.is_natural_2 [of "𝖺-1[h, p0, w]"] by simp
            also have "... = 𝖺[h, p0, w]  𝖺-1[h, p0, w]  𝔩 (h  p0  w)"
              using assms 0 k.antipar cospan 𝔩.is_natural_1 [of "𝖺-1[h, p0, w]"] comp_assoc
              by simp
            also have "... = (𝖺[h, p0, w]  𝖺-1[h, p0, w])  𝔩 (h  p0  w)"
              using comp_assoc by presburger
            also have "... = 𝔩 (h  p0  w)"
              using assms 0 k.antipar cospan comp_cod_arr comp_assoc_assoc' by simp
            also have "... = 𝗅[h  p0  w]"
              using assms 0 k.antipar cospan 𝔩_ide_simp by simp
            finally show ?thesis by simp
          qed
          thus ?thesis
            using comp_assoc by presburger
        qed
        also have "... = 𝖺[h, p0, w]  (𝗅[h  p0]  w) 
               𝖺-1[trg h, h  p0, w] 
               ((trg h  𝖺-1[h, p0, w])  (k.ε  h  p0  w)) 
               𝖺-1[k, k*, h  p0  w] 
               (k  𝖺[k*, h, p0  w]  𝖺[k*  h, p0, w]  (tab  w))"
          using assms 0 k.antipar cospan lunit_hcomp comp_assoc by simp
        also have "... = 𝖺[h, p0, w]  (𝗅[h  p0]  w) 
               (𝖺-1[trg h, h  p0, w]  (k.ε  (h  p0)  w)) 
               ((k  k*)  𝖺-1[h, p0, w])  𝖺-1[k, k*, h  p0  w] 
               (k  𝖺[k*, h, p0  w]  𝖺[k*  h, p0, w]  (tab  w))"
        proof -
          have "(trg h  𝖺-1[h, p0, w])  (k.ε  h  p0  w) =
                (k.ε  (h  p0)  w)  ((k  k*)  𝖺-1[h, p0, w])"
            using assms 0 k.antipar cospan comp_arr_dom comp_cod_arr
                  interchange [of "trg h" k.ε "𝖺-1[h, p0, w]" "h  p0  w"]
                  interchange [of k.ε "k  k*" "(h  p0)  w" "𝖺-1[h, p0, w]"]
            by simp
          thus ?thesis
            using comp_assoc by presburger
        qed
        also have "... = 𝖺[h, p0, w]  (𝗅[h  p0]  w) 
               ((k.ε  (h  p0))  w)  𝖺-1[k  k*, h  p0, w]  
               ((k  k*)  𝖺-1[h, p0, w])  𝖺-1[k, k*, h  p0  w] 
               (k  𝖺[k*, h, p0  w]  𝖺[k*  h, p0, w]  (tab  w))"
          using assms 0 k.antipar cospan assoc'_naturality [of k.ε "h  p0" w] comp_assoc
          by simp
        also have "... = 𝖺[h, p0, w]  (𝗅[h  p0]  w) 
                           ((k.ε  (h  p0))  w)  𝖺-1[k  k*, h  p0, w]  
                           ((k  k*)  𝖺-1[h, p0, w])  𝖺-1[k, k*, h  p0  w] 
                           (k  𝖺[k*, h, p0  w])  (k  𝖺[k*  h, p0, w]) 
                           (k  tab  w)"
        proof -
          have "k  𝖺[k*, h, p0  w]  𝖺[k*  h, p0, w]  (tab  w) =
                (k  𝖺[k*, h, p0  w])  (k  𝖺[k*  h, p0, w]) 
                (k  tab  w)"
          proof -
            have "seq 𝖺[k*, h, p0  w] (𝖺[k*  h, p0, w]  (tab  w))"
              using χ_def assms 0 k.antipar cospan χ by blast
            thus ?thesis
              using assms 0 k.antipar cospan whisker_left by auto
          qed
          thus ?thesis
            using comp_assoc by presburger
        qed
        also have "... = 𝖺[h, p0, w]  (𝗅[h  p0]  w) 
                           ((k.ε  (h  p0))  w)  (𝖺-1[k  k*, h  p0, w]  
                           ((k  k*)  𝖺-1[h, p0, w])  𝖺-1[k, k*, h  p0  w] 
                           (k  𝖺[k*, h, p0  w])  (k  𝖺[k*  h, p0, w]) 
                           𝖺[k, (k*  h)  p0, w])  ((k  tab)  w)  𝖺-1[k, p1, w]"
        proof -
          have "k  tab  w = 𝖺[k, (k*  h)  p0, w]  𝖺-1[k, (k*  h)  p0, w]  (k  tab  w)"
          proof -
            have "𝖺[k, (k*  h)  p0, w]  𝖺-1[k, (k*  h)  p0, w]  (k  tab  w) =
                  (𝖺[k, (k*  h)  p0, w]  𝖺-1[k, (k*  h)  p0, w])  (k  tab  w)"
              using comp_assoc by presburger
            also have "... = (k  ((k*  h)  p0)  w)  (k  tab  w)"
               using assms k.antipar 0 comp_assoc_assoc' by simp
            also have "... = k  tab  w"
               using assms k.antipar 0 comp_cod_arr by simp
            finally show ?thesis by simp
          qed
          also have "... = 𝖺[k, (k*  h)  p0, w]  ((k  tab)  w)  𝖺-1[k, p1, w]"
            using assms 0 k.antipar cospan assoc'_naturality [of k tab w] by simp
          finally have "k  tab  w = 𝖺[k, (k*  h)  p0, w]  ((k  tab)  w)  𝖺-1[k, p1, w]"
            by simp
          thus ?thesis
            using comp_assoc by presburger
        qed
        also have "... = 𝖺[h, p0, w]  (𝗅[h  p0]  w) 
                           ((k.ε  h  p0)  w) 
                           (𝖺-1[k, k*, h  p0]  (k  𝖺[k*, h, p0])  w) 
                           ((k  tab)  w)  𝖺-1[k, p1, w]"
        proof -
          have "𝖺-1[k  k*, h  p0, w]  ((k  k*)  𝖺-1[h, p0, w]) 
                  𝖺-1[k, k*, h  p0  w]  (k  𝖺[k*, h, p0  w]) 
                  (k  𝖺[k*  h, p0, w])  𝖺[k, (k*  h)  p0, w] =
                𝖺-1[k, k*, h  p0]  (k  𝖺[k*, h, p0])  w"
          proof -
            have "𝖺-1[k  k*, h  p0, w]  ((k  k*)  𝖺-1[h, p0, w]) 
                    𝖺-1[k, k*, h  p0  w]  (k  𝖺[k*, h, p0  w]) 
                    (k  𝖺[k*  h, p0, w])  𝖺[k, (k*  h)  p0, w] =
                  𝖺-1[k  k*, h  p0, w]  ((k  k*)  𝖺-1[h, p0, w]) 
                    𝖺-1[k, k*, h  p0  w]  (k  𝖺[k*, h, p0  w]) 
                    (k  𝖺[k*  h, p0, w])  𝖺[k, (k*  h)  p0, w]"
              using assms 0 k.antipar cospan α_def 𝖺'_def by simp
            also have "... = 𝖺-1[k, k*, h  p0] 
                                (k  𝖺[k*, h, p0])  w"
              using assms 0 k.antipar cospan
              by (intro E.eval_eqI, simp_all)
            also have "... = 𝖺-1[k, k*, h  p0]  (k  𝖺[k*, h, p0])  w"
              using assms 0 k.antipar cospan α_def 𝖺'_def by simp
            finally show ?thesis by simp
          qed
          thus ?thesis
            using comp_assoc by presburger
        qed
        also have "... = 𝖺[h, p0, w] 
                           (𝗅[h  p0]  (k.ε  h  p0) 
                           𝖺-1[k, k*, h  p0]  (k  𝖺[k*, h, p0]) 
                           (k  tab)  w)  𝖺-1[k, p1, w]"
          using assms 0 k.antipar cospan comp_assoc whisker_right by auto
        also have "... = 𝖺[h, p0, w]  (φ  w)  𝖺-1[k, p1, w]"
          unfolding φ_def k.trnlε_def
          using assms 0 k.antipar cospan comp_assoc whisker_left by simp
        finally have "𝗅[h  p0  w]  (k.ε  h  p0  w) 
               𝖺-1[k, k*, h  p0  w] 
               (k  𝖺[k*, h, p0  w]  𝖺[k*  h, p0, w]  (tab  w)) =
                      𝖺[h, p0, w]  (φ  w)  𝖺-1[k, p1, w]"
          by blast
        thus ?thesis
          using χ_def comp_assoc by simp
      qed
      finally show ?thesis by simp
    qed

    text BS3› implies that φ› is the unique 2-cell from k ⋆ p1 to h ⋆ p0 and is an isomorphism.
    ›

    lemma φ_uniqueness:
    shows "μ. «μ : k  p1  h  p0»  μ = φ" and "iso φ"
    proof -
      have 2: "is_left_adjoint (k  p1)"
        using k.antipar cospan left_adjoints_compose by (simp add: k_is_map)
      have 3: "is_left_adjoint (h  p0)"
        using k.antipar cospan left_adjoints_compose by (simp add: h_is_map)
      show "μ. «μ : k  p1  h  p0»  μ = φ"
        using φ_in_hom 2 3 BS3 by simp
      show "iso φ"
        using φ_in_hom 2 3 BS3 by simp
    qed

    text ‹
      As a consequence, the chosen tabulation of k* ⋆ h› is the unique 2-cell from
      p1 to (k* ⋆ h) ⋆ p0, and therefore if we are given any such 2-cell we may
      conclude it yields a tabulation of k* ⋆ h›.
    ›

    lemma tab_uniqueness:
    assumes "«τ : p1  (k*  h)  p0»"
    shows "τ = tab"
    proof -
      have "«k.trnlε (h  p0) (𝖺[k*, h, p0]  τ) : k  p1  h  p0»"
        using assms k.antipar cospan k.adjoint_transpose_left(2) [of "h  p0" "p1"]
              assoc_in_hom
        by force
      hence "tab = 𝖺-1[k*, h, p0]  k.trnlη p1 (k.trnlε (h  p0) (𝖺[k*, h, p0]  τ))"
        using transpose_φ φ_uniqueness(1) by auto
      also have "... = 𝖺-1[k*, h, p0]  𝖺[k*, h, p0]  τ"
        using assms k.antipar cospan k.adjoint_transpose_left(4) assoc_in_hom by auto
      also have "... = (𝖺-1[k*, h, p0]  𝖺[k*, h, p0])  τ"
        using comp_assoc by presburger
      also have "... = τ"
        using assms k.antipar cospan comp_cod_arr comp_assoc_assoc' by auto
      finally show ?thesis by simp
    qed

    text ‹
      The following lemma reformulates the biuniversal property of the canonical tabulation
      of k* ⋆ h› as a biuniversal property of φ›, regarded as a square that commutes up to
      isomorphism.
    ›

    lemma φ_biuniversal_prop:
    assumes "ide u" and "ide v"
    shows "μ. «μ : k  v  h  u» 
                w θ ν. ide w  «θ : p0  w  u»  «ν : v  p1  w»  iso ν 
                        (h  θ)  𝖺[h, p0, w]  (φ  w)  𝖺-1[k, p1, w]  (k  ν) = μ"
    and "w w' θ θ' β.
             ide w; ide w';
              «θ : p0  w  u»; «θ' : p0  w'  u»;
              «β : p1  w  p1  w'»;
              (h  θ)  𝖺[h, p0, w]  (φ  w)  𝖺-1[k, p1, w] =
              (h  θ')  𝖺[h, p0, w']  (φ  w')  𝖺-1[k, p1, w']  (k  β) 
                     ∃!γ. «γ : w  w'»  θ = θ'  (p0  γ)  p1  γ = β"
    proof -
      fix μ
      assume μ: "«μ : k  v  h  u»"
      have 1: "src h = trg u"
        using assms μ ide_cod
        by (metis ide_def in_homE seq_if_composable)
      have 2: "src k = trg v"
        using assms μ ide_dom
        by (metis ideD(1) in_homE not_arr_null seq_if_composable)
      let  = "𝖺-1[k*, h, u]  k.trnlη v μ"
      have ω: "« : v  (k*  h)  u»"
        using assms μ 1 2 k.antipar cospan k.adjoint_transpose_left(1) [of "h  u" v]
              assoc_in_hom
        by auto
      obtain w θ ν
      where wθν: "ide w  «θ : p0  w  u»  «ν : v  p1  w»  iso ν 
                  ((k*  h)  θ)  𝖺[k*  h, p0, w]  (tab  w)  ν = "
        using assms ω T1 [of u ] comp_assoc by (metis in_homE)
      have 0: "src p0 = trg w"
        using wθν ide_dom
        by (metis hseqE ideD(1) in_homE)
      have "μ = k.trnlε (h  u) (𝖺[k*, h, u]  ((k*  h)  θ)  𝖺[k*  h, p0, w]  (tab  w)  ν)"
      proof -
        have "μ = k.trnlε (h  u) (𝖺[k*, h, u]  )"
        proof -
          have "k.trnlε (h  u) (𝖺[k*, h, u]  ) =
                k.trnlε (h  u) ((𝖺[k*, h, u]  𝖺-1[k*, h, u])  k.trnlη v μ)"
            using comp_assoc by presburger
          also have "... = k.trnlε (h  u) (k.trnlη v μ)"
          proof -
            have "(𝖺[k*, h, u]  𝖺-1[k*, h, u])  k.trnlη v μ = (k*  h  u)  k.trnlη v μ"
              using comp_assoc_assoc'
              by (simp add: "1" assms(1) cospan k.antipar(2))
            also have "... = k.trnlη v μ"
              using "1" ω assms(1) comp_ide_arr cospan k.antipar(2) by fastforce
            finally show ?thesis
              by simp
          qed
          also have "... = μ"
            using assms μ k.antipar cospan 1 2 k.adjoint_transpose_left(3) by simp
          finally show ?thesis by simp
        qed
        thus ?thesis using wθν by simp
      qed
      also have "... = (h  θ)  𝖺[h, p0, w]  (φ  w)  𝖺-1[k, p1, w]  (k  ν)"
        using assms k.antipar cospan wθν transpose_triangle [of w θ u ν] by auto
      finally have "(h  θ)  𝖺[h, p0, w]  (φ  w)  𝖺-1[k, p1, w]  (k  ν) = μ"
        by simp
      thus "w θ ν. ide w  «θ : p0  w  u» 
                    «ν : v  p1  w»  iso ν 
                    (h  θ)  𝖺[h, p0, w]  (φ  w) 
                      𝖺-1[k, p1, w]  (k  ν) = μ"
        using wθν by blast
      next
      fix w w' θ θ' β
      assume w: "ide w"
      assume w': "ide w'"
      assume θ: "«θ : p0  w  u»"
      assume θ': "«θ' : p0  w'  u»"
      assume β: "«β : p1  w  p1  w'»"
      assume eq: "(h  θ)  𝖺[h, p0, w]  (φ  w)  𝖺-1[k, p1, w] =
                  (h  θ')  𝖺[h, p0, w']  (φ  w')  𝖺-1[k, p1, w']  (k  β)"
      have 0: "src p0 = trg w"
        using θ ide_dom
        by (metis ideD(1) in_homE not_arr_null seq_if_composable)
      interpret uwθw'θ': uwθw'θ' V H 𝖺 𝗂 src trg k*  h tab p0 p1
                             u w θ w' θ'
        using w θ w' θ' apply (unfold_locales) by auto
      show "∃!γ. «γ : w  w'»  θ = θ'  (p0  γ)  p1  γ = β"
      proof -
        let ?LHS = "𝖺[k*, h, u]  ((k*  h)  θ)  𝖺[k*  h, p0, w]  (tab  w)"
        let ?RHS = "𝖺[k*, h, u]  ((k*  h)  θ')  𝖺[k*  h, p0, w']  (tab  w')  β"
        have eq': "?LHS = ?RHS"
        proof -
          have "k.trnlε (h  u) ?LHS =
                k.trnlε (h  u)
                        (𝖺[k*, h, u]  ((k*  h)  θ)  𝖺[k*  h, p0, w]  (tab  w)  (p1  w))"
            using assms 0 w θ β k.antipar cospan comp_arr_dom
            by (metis tab_simps(1) tab_simps(4) whisker_right)
          also have "... = (h  θ)  𝖺[h, p0, w]  (φ  w)  𝖺-1[k, p1, w]  (k  p1  w)"
            using assms w θ β transpose_triangle
            by (metis arr_dom ide_hcomp ide_in_hom(2) in_homE ide_leg1 not_arr_null
                seq_if_composable)
          also have "... = (h  θ)  𝖺[h, p0, w]  (φ  w)  𝖺-1[k, p1, w]"
            using assms 0 w k.antipar cospan comp_arr_dom by simp
          also have "... = (h  θ')  𝖺[h, p0, w']  (φ  w')  𝖺-1[k, p1, w']  (k  β)"
            using eq by blast
          also have "... = k.trnlε (h  u) ?RHS"
            using assms w' θ' β transpose_triangle by simp
          finally have 4: "k.trnlε (h  u) ?LHS = k.trnlε (h  u) ?RHS"
            by simp
          have "src k = trg (p1  w)"
            using assms 0 w k.antipar cospan by simp
          moreover have "src k* = trg (h  u)"
            using assms 0 w k.antipar cospan by simp
          moreover have "ide (h  u)"
            using assms 0 w k.antipar cospan by simp
          moreover have "ide (p1  w)"
            using assms 0 w k.antipar cospan by simp
          ultimately have "inj_on (k.trnlε (h  u)) (hom (p1  w) (k*  h  u))"
            using assms 0 w w' k.antipar cospan k.adjoint_transpose_left(6) bij_betw_imp_inj_on
            by blast
          moreover have "?LHS  hom (p1  w) (k*  h  u)"
          proof -
            have "«𝖺[k*, h, u]  ((k*  h)  θ)  𝖺[k*  h, p0, w]  (tab  w) :
                     p1  w  k*  h  u»"
              using k.antipar cospan
              apply (intro comp_in_homI) by auto
            thus ?thesis by simp
          qed
          moreover have "?RHS  hom (p1  w) (k*  h  u)"
          proof -
            have "«𝖺[k*, h, u]  ((k*  h)  θ')  𝖺[k*  h, p0, w'] 
                    (tab  w')  β : p1  w  k*  h  u»"
              using β k.antipar cospan
              apply (intro comp_in_homI) by auto
            thus ?thesis by blast
          qed
          ultimately show "?LHS = ?RHS"
            using assms 4 k.antipar cospan bij_betw_imp_inj_on
                  inj_on_def [of "k.trnlε (h  u)" "hom (p1  w) (k*  h  u)"]
            by simp
        qed
        moreover have "seq 𝖺[k*, h, u] (composite_cell w θ)"
          using assms k.antipar cospan tab_in_hom by fastforce
        moreover have "seq 𝖺[k*, h, u] (composite_cell w' θ'  β)"
          using assms β k.antipar cospan tab_in_hom by fastforce
        ultimately have "composite_cell w θ = composite_cell w' θ'  β"
          using assms 0 w w' β k.antipar cospan iso_assoc iso_is_section section_is_mono
                monoE [of "𝖺[k*, h, u]" "composite_cell w θ" "composite_cell w' θ'  β"]
                comp_assoc
          by simp
        thus ?thesis
          using w w' θ θ' β eq' T2 [of w w' θ u θ' β] by metis
      qed
    qed

    text ‹
      Using the uniqueness properties established for φ›, we obtain yet another reformulation
      of the biuniversal property associated with the chosen tabulation of k* ⋆ h›,
      this time as a kind of pseudo-pullback.  We will use this to show that the
      category of isomorphism classes of maps has pullbacks.
    ›

    lemma has_pseudo_pullback:
    assumes "is_left_adjoint u" and "is_left_adjoint v" and "isomorphic (k  v) (h  u)"
    shows "w. is_left_adjoint w  isomorphic (p0  w) u  isomorphic v (p1  w)"
    and "w w'.  is_left_adjoint w; is_left_adjoint w';
                  p0  w  u; v  p1  w; p0  w'  u; v  p1  w'   w  w'"
    proof -
      interpret u: map_in_bicategory V H 𝖺 𝗂 src trg u
        using assms(1) by unfold_locales auto
      interpret v: map_in_bicategory V H 𝖺 𝗂 src trg v
        using assms(2) by unfold_locales auto
      obtain μ where μ: "«μ : k  v  h  u»  iso μ"
        using assms(3) by auto
      obtain w θ ν where wθν: "ide w  «θ : p0  w  u» 
                               «ν : v  p1  w»  iso ν 
                               (h  θ)  𝖺[h, p0, w]  (φ  w)  𝖺-1[k, p1, w]  (k  ν) = μ"
        using assms μ φ_biuniversal_prop(1) [of u v μ] by auto
      have "is_left_adjoint w  isomorphic (p0  w) u  isomorphic v (p1  w)"
      proof (intro conjI)
        show 1: "is_left_adjoint w"
          using assms(2) wθν left_adjoint_preserved_by_iso' isomorphic_def BS4 leg1_is_map
          by blast
        show "v  p1  w"
          using wθν isomorphic_def by blast
        show "p0  w  u"
        proof -
          have "src p0 = trg w"
            using wθν ide_dom
            by (metis ideD(1) in_homE not_arr_null seq_if_composable)
          hence "is_left_adjoint (p0  w)"
            using 1 left_adjoints_compose by simp
          thus ?thesis
            using assms wθν 1 BS3 isomorphic_def by metis
        qed
      qed
      thus "w. is_left_adjoint w  p0  w  u  v  p1  w"
        by blast
      show "w w'.  is_left_adjoint w; is_left_adjoint w';
                     p0  w  u; v  p1  w; p0  w'  u; v  p1  w'   w  w'"
      proof -
        fix w w'
        assume w: "is_left_adjoint w" and w': "is_left_adjoint w'"
        assume 1: "p0  w  u"
        assume 2: "v  p1  w"
        assume 3: "p0  w'  u"
        assume 4: "v  p1  w'"
        obtain θ where θ: "«θ : p0  w  u»"
          using 1 by auto
        obtain θ' where θ': "«θ' : p0  w'  u»"
          using 3 by auto
        obtain ν where ν: "«ν: v  p1  w»  iso ν"
          using 2 by blast
        obtain ν' where ν': "«ν': v  p1  w'»  iso ν'"
          using 4 by blast
        let  = "ν'  inv ν"
        have β: "« : p1  w  p1  w'»"
          using ν ν' by (elim conjE in_homE, auto)
        interpret uwθ: uwθ V H 𝖺 𝗂 src trg k*  h tab p0 p1 u w θ
          using w θ left_adjoint_is_ide
          by unfold_locales auto
        interpret uw'θ': uwθ V H 𝖺 𝗂 src trg k*  h tab p0 p1
                             u w' θ'
          using w' θ' left_adjoint_is_ide
          by unfold_locales auto
        interpret uwθw'θ': uwθw'θ' V H 𝖺 𝗂 src trg k*  h tab p0 p1 u w θ w' θ'
          using w w' θ θ' left_adjoint_is_ide by unfold_locales
        have "(h  θ)  𝖺[h, p0, w]  (φ  w)  𝖺-1[k, p1, w] =
              (h  θ')  𝖺[h, p0, w']  (φ  w')  𝖺-1[k, p1, w'] 
                (k  )"
        proof -
          let ?LHS = "(h  θ)  𝖺[h, p0, w]  (φ  w)  𝖺-1[k, p1, w]"
          let ?RHS = "(h  θ')  𝖺[h, p0, w']  (φ  w')  𝖺-1[k, p1, w']  (k  )"
          have "«?LHS : k  p1  w  h  u»"
            using w k.antipar by fastforce
          moreover have "«?RHS : k  p1  w  h  u»"
            using w k.antipar β by fastforce
          moreover have "is_left_adjoint (k  p1  w)"
            using w k.is_map left_adjoints_compose by simp
          moreover have "is_left_adjoint (h  u)"
            using assms h.is_map left_adjoints_compose by auto
          ultimately show "?LHS = ?RHS"
            using BS3 by blast
        qed
        hence "∃!γ. «γ : w  w'»  θ = θ'  (p0  γ)  p1  γ = "
          using assms left_adjoint_is_ide w w' θ θ' ν ν' β
                φ_biuniversal_prop(2) [of u v w w' θ θ' ]
          by presburger
        thus "w  w'"
          using w w' BS3 isomorphic_def by metis
      qed
    qed

  end

  subsubsection "Tabulations in Maps"

  text ‹
    Here we focus our attention on the properties of tabulations in a bicategory of spans,
    in the special case in which both legs are maps.
  ›

  context tabulation_in_maps
  begin

    text ‹
      The following are the conditions under which w› is a 1-cell induced via T1› by
      a 2-cell «ω : dom ω ⇒ s ⋆ r0»›:  w› is an arrow of spans and ω› is obtained by
      composing the tabulation σ› with w› and the isomorphisms that witness w› being
      an arrow of spans.
    ›

    abbreviation is_induced_by_cell
    where "is_induced_by_cell w r0 ω 
           arrow_of_spans_of_maps V H 𝖺 𝗂 src trg r0 (dom ω) s0 s1 w 
           composite_cell w (arrow_of_spans_of_maps.the_θ V H r0 s0 w) 
             (arrow_of_spans_of_maps.the_ν V H (dom ω) s1 w) = ω"

    lemma induced_map_unique:
    assumes "is_induced_by_cell w r0 ω" and "is_induced_by_cell w' r0 ω"
    shows "isomorphic w w'"
    proof -
      interpret w: arrow_of_spans_of_maps V H 𝖺 𝗂 src trg r0 dom ω s0 s1 w
        using assms(1) by auto
      interpret w: arrow_of_spans_of_maps_to_tabulation V H 𝖺 𝗂 src trg r0 dom ω s σ s0 s1 w
        ..
      interpret w': arrow_of_spans_of_maps V H 𝖺 𝗂 src trg r0 dom ω s0 s1 w'
        using assms(2) by auto
      interpret w': arrow_of_spans_of_maps_to_tabulation V H 𝖺 𝗂 src trg r0 dom ω s σ s0 s1 w'
        ..
      let  = "w'.the_ν  inv w.the_ν"
      have β: "« : s1  w  s1  w'»"
        using w.the_ν_props w'.the_ν_props arr_iff_in_hom by auto
      have 1: "composite_cell w w.the_θ = composite_cell w' w'.the_θ  (w'.the_ν  inv w.the_ν)"
      proof -
        have "composite_cell w' w'.the_θ  (w'.the_ν  inv w.the_ν) =
              ((composite_cell w' w'.the_θ)  w'.the_ν)  inv w.the_ν"
          using comp_assoc by presburger
        also have "... = ω  inv w.the_ν"
          using assms(2) comp_assoc by simp
        also have "... = (composite_cell w w.the_θ  w.the_ν)  inv w.the_ν"
          using assms(1) comp_assoc by simp
        also have "... = composite_cell w w.the_θ  w.the_ν  inv w.the_ν"
          using comp_assoc by presburger
        also have "... = composite_cell w w.the_θ"
        proof -
          have "w.the_ν  inv w.the_ν = s1  w"
            using w.the_ν_props comp_arr_inv inv_is_inverse by auto
          thus ?thesis
            using composite_cell_in_hom w.ide_w w.the_θ_props comp_arr_dom
            by (metis composite_cell_in_hom in_homE w.w_in_hom(1))
        qed
        finally show ?thesis by auto
      qed
      have "γ. «γ : w  w'»"
        using 1 β w.is_ide w'.is_ide w.the_θ_props w'.the_θ_props
              T2 [of w w' w.the_θ r0 w'.the_θ ]
        by blast
      thus ?thesis
        using BS3 w.is_map w'.is_map by blast
    qed

    text ‹
      The object src s0 forming the apex of the tabulation satisfies the conditions for
      being a map induced via T1› by the 2-cell σ› itself.  This is ultimately required
      for the map from 2-cells to arrows of spans to be functorial with respect to vertical
      composition.
    ›

    lemma apex_is_induced_by_cell:
    shows "is_induced_by_cell (src s0) s0 σ"
    proof -
      have 1: "arrow_of_spans_of_maps V H 𝖺 𝗂 src trg s0 s1 s0 s1 (src s0)"
        using iso_runit [of s0] iso_runit [of s1] tab_in_hom
        apply unfold_locales
          apply simp
        using ide_leg0 isomorphic_def
         apply blast
        using ide_leg1 isomorphic_def leg1_simps(3) runit'_in_vhom [of s1 "src s0"] iso_runit'
        by blast
      interpret w: arrow_of_spans_of_maps V H 𝖺 𝗂 src trg s0 dom σ s0 s1 src s0
        using 1 tab_in_hom by simp
      interpret w: arrow_of_spans_of_maps_to_tabulation
                     V H 𝖺 𝗂 src trg s0 dom σ s σ s0 s1 src s0
        ..
      show "is_induced_by_cell (src s0) s0 σ"
      proof (intro conjI)
        show "arrow_of_spans_of_maps V H 𝖺 𝗂 src trg s0 (dom σ) s0 s1 (src s0)"
          using w.arrow_of_spans_of_maps_axioms by simp
        show "composite_cell (src s0) w.the_θ  w.the_ν = σ"
        proof -
          have θ: "w.the_θ = 𝗋[s0]"
            using iso_runit [of s0] w.leg0_uniquely_isomorphic w.the_θ_props
                  the1_equality [of "λθ. «θ : s0  src s0  s0»  iso θ"]
            by auto
          have ν: "w.the_ν = 𝗋-1[s1]"
            using iso_runit' w.leg1_uniquely_isomorphic w.the_ν_props leg1_simps(3)
                  the1_equality [of "λν. «ν : s1  s1  src s0»  iso ν"] tab_in_vhom'
            by auto
          have "composite_cell (src s0) 𝗋[s0]  𝗋-1[s1] = σ"
          proof -
            have "composite_cell (src s0) 𝗋[s0]  𝗋-1[s1] =
                  ((s  𝗋[s0])  𝖺[s, s0, src s0])  (σ  src s0)  𝗋-1[s1]"
              using comp_assoc by presburger
            also have "... = (𝗋[s  s0]  (σ  src s0))  𝗋-1[s1]"
              using runit_hcomp comp_assoc by simp
            also have "... = σ  𝗋[s1]  𝗋-1[s1]"
              using runit_naturality tab_in_hom
              by (metis tab_simps(1) tab_simps(2) tab_simps(4) tab_simps(5) comp_assoc)
            also have "... = σ"
              using iso_runit tab_in_hom comp_arr_dom comp_arr_inv inv_is_inverse by simp
            finally show ?thesis by simp
          qed
          thus ?thesis
            using θ ν comp_assoc by simp
        qed
      qed
    qed

  end

  subsubsection "Composing Tabulations"

  text ‹
    Given tabulations (r0, ρ, r1)› of r› and (s0, σ, s1)› of s› in a bicategory of spans,
    where (r0, r1)› and (s0, s1)› are spans of maps and 1-cells r› and s› are composable,
    we can construct a 2-cell that yields a tabulation of r ⋆ s›.
    The proof uses the fact that the 2-cell φ› associated with the cospan (r0, s1)›
    is an isomorphism, which we have proved above
    (cospan_of_maps_in_bicategory_of_spans.φ_uniqueness›) using BS3›.
    However, this is the only use of BS3› in the proof, and it seems plausible that it would be
    possible to establish that φ› is an isomorphism in more general situations in which the
    subbicategory of maps is not locally essentially discrete.  Alternatively, more general
    situations could be treated by adding the assertion that φ› is an isomorphism as part of
    a weakening of BS3›.
  ›

  locale composite_tabulation_in_maps =
    bicategory_of_spans V H 𝖺 𝗂 src trg +
    ρ: tabulation_in_maps V H 𝖺 𝗂 src trg r ρ r0 r1 +
    σ: tabulation_in_maps V H 𝖺 𝗂 src trg s σ s0 s1
  for V :: "'a comp"                 (infixr "" 55)
  and H :: "'a  'a  'a"          (infixr "" 53)
  and 𝖺 :: "'a  'a  'a  'a"     ("𝖺[_, _, _]")
  and 𝗂 :: "'a  'a"                 ("𝗂[_]")
  and src :: "'a  'a"
  and trg :: "'a  'a"
  and r :: 'a
  and ρ :: 'a
  and r0 :: 'a
  and r1 :: 'a
  and s :: 'a
  and σ :: 'a
  and s0 :: 'a
  and s1 :: 'a +
  assumes composable: "src r = trg s"
  begin

    text ‹
      Interpret (r0, s1)› as a @{locale cospan_of_maps_in_bicategory_of_spans},
      to obtain the isomorphism φ› in the central diamond, along with the assertion
      that it is unique.
    ›
    interpretation r0s1: cospan_of_maps_in_bicategory_of_spans V H 𝖺 𝗂 src trg s1 r0
      using ρ.leg0_is_map σ.leg1_is_map composable by unfold_locales auto

    text ‹
      We need access to simps, etc. in the preceding interpretation, yet trying to declare
      it as a sublocale introduces too many conflicts at the moment.
      As it confusing elsewhere to figure out exactly how, in other contexts, to express
      the particular interpretation that is needed, to make things easier we include the
      following lemma.  Then we can just recall the lemma to find out how to express
      the interpretation required in a given context.
    ›

    lemma r0s1_is_cospan:
    shows "cospan_of_maps_in_bicategory_of_spans V H 𝖺 𝗂 src trg s1 r0"
      ..

    text ‹
      The following define the projections associated with the natural tabulation of r0* ⋆ s1.
    ›

    abbreviation p0
    where "p0  r0s1.p0"

    abbreviation p1
    where "p1  r0s1.p1"

    text ‹
$$
\xymatrix{
  && {\rm src}~\phi \ar[dl]_{p_1} \ar[dr]^{p_0} \ddtwocell\omit{^\phi} \\  
  & {\rm src}~\rho \ar[dl]_{r_1} \ar[dr]^{r_0} \dtwocell\omit{^\rho}
  && {\rm src}~\sigma  \ar[dl]_{s_1} \ar[dr]^{s_0} \dtwocell\omit{^\sigma}\\
  {\rm trg}~r && {\rm src}~r = {\rm trg}~s \ar[ll]^{r} && {\rm src}~s \ar[ll]^{s}
}
$$
    ›

    text ‹
      Next, we define the 2-cell that is the composite of the tabulation σ›, the tabulation ρ›,
      and the central diamond that commutes up to unique isomorphism φ›.
    ›

    definition tab
    where "tab  𝖺-1[r, s, s0  p0]  (r  𝖺[s, s0, p0])  (r  σ  p0) 
                   (r  r0s1)  𝖺[r, r0, p1]  (ρ  p1)"

    lemma tab_in_hom [intro]:
    shows "«tab : r1  p1  (r  s)  s0  p0»"
      using ρ.T0.antipar(1) r0s1.φ_in_hom composable ρ.leg0_in_hom(1) σ.leg1_in_hom(1)
            composable tab_def
      by auto

    interpretation tabulation_data V H 𝖺 𝗂 src trg r  s tab s0  p0 r1  p1
      using composable tab_in_hom
      by unfold_locales auto

    text ‹
      In the subsequent proof we will use coherence to shortcut a few of the calculations.
    ›
    interpretation E: self_evaluation_map V H 𝖺 𝗂 src trg ..
    notation E.eval ("_")

    text ‹
      The following is applied twice in the proof of property T2› for the composite
      tabulation.  It's too long to repeat.
    ›

    lemma technical:
    assumes "ide w"
    and "«θ : (s0  p0)  w  u»"
    and "wr = p1  w"
    and "θr = (s  θ)  (s  𝖺-1[s0, p0, w])  𝖺[s, s0, p0  w] 
                (σ  p0  w)  𝖺[s1, p0, w]  (r0s1  w)  𝖺-1[r0, p1, w]"
    shows "ρ.composite_cell wr θr = 𝖺[r, s, u]  composite_cell w θ  𝖺-1[r1, p1, w]"
    text ‹
$$
\xymatrix{
  && X \ar[d]^{w} \ar@/ur20pt/[dddrr]^{u} \xtwocell[ddr]{}\omit{^{\theta}} \\
  && {\rm src}~\phi \ar[dl]_{p_1} \ar[dr]^{p_0} \ddtwocell\omit{^\phi} \\  
  & {\rm src}~\rho \ar[dl]_{r_1} \ar[dr]^{r_0} \dtwocell\omit{^\rho}
  && {\rm src}~\sigma  \ar[dl]_{s_1} \ar[dr]^{s_0} \dtwocell\omit{^\sigma}\\
  {\rm trg}~r && {\rm src}~r = {\rm trg}~s \ar[ll]^{r} && {\rm src}~s \ar[ll]^{s}
}
$$
    ›
    proof -
      interpret uwθ: uwθ V H 𝖺 𝗂 src trg r  s tab s0  p0 r1  p1 u w θ
        using assms(1-2) composable
        by unfold_locales auto
      show ?thesis
      proof -
        have "𝖺[r, s, u]  composite_cell w θ  𝖺-1[r1, p1, w] =
              (𝖺[r, s, u]  ((r  s)  θ))  𝖺[r  s, s0  p0, w]  (tab  w)  𝖺-1[r1, p1, w]"
          using comp_assoc by presburger
        also have "... = (r  s  θ)  𝖺[r, s, (s0  p0)  w]  𝖺[r  s, s0  p0, w] 
                           (tab  w)  𝖺-1[r1, p1, w]"
          using assoc_naturality [of r s θ] composable comp_assoc by simp
        also have "... = (r  s  θ)  𝖺[r, s, (s0  p0)  w]  𝖺[r  s, s0  p0, w] 
                           ((𝖺-1[r, s, s0  p0]  (r  𝖺[s, s0, p0]))  (r  σ  p0) 
                           ρ.composite_cell p1 r0s1  w)  𝖺-1[r1, p1, w]"
          unfolding tab_def
          using comp_assoc by presburger
        also have "... = (r  s  θ)  ((𝖺[r, s, (s0  p0)  w]  𝖺[r  s, s0  p0, w] 
                           (𝖺-1[r, s, s0  p0]  (r  𝖺[s, s0, p0])  w))) 
                           ((r  σ  p0)  ρ.composite_cell p1 r0s1  w) 
                           𝖺-1[r1, p1, w]"
          using composable ρ.T0.antipar(1) comp_assoc whisker_right by auto
        also have "... = (r  s  θ)  ((𝖺[r, s, (s0  p0)  w]  𝖺[r  s, s0  p0, w] 
                           (𝖺-1[r, s, s0  p0]  (r  𝖺[s, s0, p0])  w))) 
                           ((r  σ  p0)  w)  ((r  r0s1)  w) 
                            (𝖺[r, r0, p1]  w)  ((ρ  p1)  w)  𝖺-1[r1, p1, w]"
          using composable ρ.T0.antipar(1) whisker_right tab_def tab_in_hom(2)
                composable comp_assoc
          by force
        also have "... = (r  s  θ)  ((𝖺[r, s, (s0  p0)  w]  𝖺[r  s, s0  p0, w] 
                           (𝖺-1[r, s, s0  p0]  (r  𝖺[s, s0, p0])  w))) 
                           ((r  σ  p0)  w)  ((r  r0s1)  w) 
                           ((𝖺[r, r0, p1]  w)  𝖺-1[r  r0, p1, w])  (ρ  p1  w)"
          using assoc'_naturality [of ρ p1 w] ρ.T0.antipar(1) r0s1.base_simps(2) comp_assoc
          by auto
        also have "... = (r  s  θ)  ((𝖺[r, s, (s0  p0)  w]  𝖺[r  s, s0  p0, w] 
                           (𝖺-1[r, s, s0  p0]  (r  𝖺[s, s0, p0])  w))) 
                           ((r  σ  p0)  w)  (((r  r0s1)  w) 
                           𝖺-1[r, r0  p1, w])  ρ.composite_cell (p1  w) 𝖺-1[r0, p1, w]"
        proof -
          have "(𝖺[r, r0, p1]  w)  𝖺-1[r  r0, p1, w] =
                 𝖺-1[r, r0  p1, w]  (r  𝖺-1[r0, p1, w])  𝖺[r, r0, p1  w]"
          proof -
            have "(𝖺-1[r, r0, p1]  w)  𝖺-1[r, r0  p1, w]  (r  𝖺-1[r0, p1, w]) =
                  𝖺-1[r  r0, p1, w]  𝖺-1[r, r0, p1  w]"
              using pentagon' ρ.T0.antipar(1) comp_assoc by simp
            moreover have 1: "seq (𝖺-1[r, r0, p1]  w)(𝖺-1[r, r0  p1, w]  (r  𝖺-1[r0, p1, w]))"
              using ρ.T0.antipar(1)
              by (intro seqI hseqI, auto)
            ultimately
            have "𝖺-1[r  r0, p1, w] =
                   ((𝖺-1[r, r0, p1]  w)  𝖺-1[r, r0  p1, w]  (r  𝖺-1[r0, p1, w])) 
                     𝖺[r, r0, p1  w]"
              using ρ.T0.antipar(1) iso_assoc
                    invert_side_of_triangle(2)
                      [of "(𝖺-1[r, r0, p1]  w)  𝖺-1[r, r0  p1, w]  (r  𝖺-1[r0, p1, w])"
                          "𝖺-1[r  r0, p1, w]" "𝖺-1[r, r0, p1  w]"]
              by fastforce
            hence "𝖺-1[r  r0, p1, w] =
                   (𝖺-1[r, r0, p1]  w)  𝖺-1[r, r0  p1, w]  (r  𝖺-1[r0, p1, w]) 
                     𝖺[r, r0, p1  w]"
              using comp_assoc by presburger
            moreover have "seq (inv (𝖺-1[r, r0, p1]  w)) 𝖺-1[r  r0, p1, w]"
              using ρ.T0.antipar(1) 1 by fastforce
            ultimately show ?thesis
              using ρ.T0.antipar(1) iso_assoc
                    invert_side_of_triangle(1)
                      [of "𝖺-1[r  r0, p1, w]" "𝖺-1[r, r0, p1]  w"
                          "𝖺-1[r, r0  p1, w]  (r  𝖺-1[r0, p1, w])  𝖺[r, r0, p1  w]"]
              by fastforce
          qed
          thus ?thesis
            using comp_assoc by presburger
        qed
        also have "... = (r  s  θ)  ((𝖺[r, s, (s0  p0)  w]  𝖺[r  s, s0  p0, w] 
                           (𝖺-1[r, s, s0  p0]  (r  𝖺[s, s0, p0])  w))) 
                           (((r  σ  p0)  w)  𝖺-1[r, s1  p0, w]) 
                           (r  r0s1  w)  ρ.composite_cell (p1  w) 𝖺-1[r0, p1, w]"
        proof -
          have "((r  r0s1)  w)  𝖺-1[r, r0  p1, w] = 𝖺-1[r, s1  p0, w]  (r  r0s1  w)"
            using assoc'_naturality [of r r0s1 w] r0s1.cospan by auto
          thus ?thesis
            using comp_assoc by presburger
        qed
        also have "... = (r  s  θ) 
                           (𝖺[r, s, (s0  p0)  w]  𝖺[r  s, s0  p0, w] 
                             (𝖺-1[r, s, s0  p0]  (r  𝖺[s, s0, p0])  w) 
                             𝖺-1[r, (s  s0)  p0, w]) 
                           (r  (σ  p0)  w)  (r  r0s1  w) 
                             ρ.composite_cell (p1  w) 𝖺-1[r0, p1, w]"
        proof -
          have "((r  σ  p0)  w)  𝖺-1[r, s1  p0, w] =
                𝖺-1[r, (s  s0)  p0, w]  (r  (σ  p0)  w)"
            using assoc'_naturality [of r "σ  p0" w]
            by (simp add: composable)
          thus ?thesis
            using comp_assoc by presburger
        qed
        also have "... = (r  s  θ) 
                           (r  (s  𝖺-1[s0, p0, w])  𝖺[s, s0, p0  w]  𝖺[s  s0, p0, w]) 
                           ((r  (σ  p0)  w)  (r  r0s1  w)  (r  𝖺-1[r0, p1, w])) 
                           𝖺[r, r0, p1  w]  (ρ  p1  w)"
        proof -
          have "𝖺[r, s, (s0  p0)  w]  𝖺[r  s, s0  p0, w] 
                  (𝖺-1[r, s, s0  p0]  (r  𝖺[s, s0, p0])  w)  𝖺-1[r, (s  s0)  p0, w] =
                r  (s  𝖺-1[s0, p0, w])  𝖺[s, s0, p0  w]  𝖺[s  s0, p0, w]"
          proof -
            have "𝖺[r, s, (s0  p0)  w]  𝖺[r  s, s0  p0, w] 
                    (𝖺-1[r, s, s0  p0]  (r  𝖺[s, s0, p0])  w)  𝖺-1[r, (s  s0)  p0, w] =
                  𝖺[r, s, (s0  p0)  w]  𝖺[r  s, s0  p0, w] 
                    (𝖺-1[r, s, s0  p0]  (r  𝖺[s, s0, p0])  w) 
                    𝖺-1[r, (s  s0)  p0, w]"
              using α_def 𝖺'_def composable by simp
            also have "... = r  (s  𝖺-1[s0, p0, w])  𝖺[s, s0, p0  w] 
                                𝖺[s  s0, p0, w]"
              using composable
              by (intro E.eval_eqI, simp_all)
            also have "... = r  (s  𝖺-1[s0, p0, w])  𝖺[s, s0, p0  w]  𝖺[s  s0, p0, w]"
              using α_def 𝖺'_def composable by simp
            finally show ?thesis by simp
          qed
          thus ?thesis
            using comp_assoc by presburger
        qed
        also have "... = (r  s  θ) 
                           (r  (s  𝖺-1[s0, p0, w])  𝖺[s, s0, p0  w]  𝖺[s  s0, p0, w]) 
                           ρ.composite_cell (p1  w)
                             (((σ  p0)  w)  (r0s1  w)  𝖺-1[r0, p1, w])"
          using assms(3) arrI ρ.T0.antipar(1) whisker_left by auto
        also have "... = (r  (s  θ)  (s  𝖺-1[s0, p0, w])  𝖺[s, s0, p0  w] 
                              (𝖺[s  s0, p0, w]  ((σ  p0)  w)) 
                              (r0s1  w)  𝖺-1[r0, p1, w]) 
                           𝖺[r, r0, p1  w]  (ρ  p1  w)"
          using ρ.T0.antipar(1) comp_assoc whisker_left by auto
        also have "... = (r  (s  θ)  (s  𝖺-1[s0, p0, w])  𝖺[s, s0, p0  w] 
                              (σ  p0  w)  𝖺[s1, p0, w] 
                              (r0s1  w)  𝖺-1[r0, p1, w]) 
                           𝖺[r, r0, p1  w]  (ρ  p1  w)"
          using assoc_naturality [of σ p0 w] comp_assoc by simp
        finally show ?thesis
          using assms(3-4) by simp
      qed
    qed

    lemma composite_is_tabulation:
    shows "tabulation V H 𝖺 𝗂 src trg (r  s) tab (s0  p0) (r1  p1)"
    proof
      show "u ω.  ide u; «ω : dom ω  (r  s)  u»  
                   w θ ν. ide w  «θ : (s0  p0)  w  u» 
                           «ν : dom ω  (r1  p1)  w»  iso ν 
                           composite_cell w θ  ν = ω"
      proof -
        fix u ω
        assume u: "ide u"
        assume ω: "«ω : dom ω  (r  s)  u»"
        let ?v = "dom ω"
        have 1: "«𝖺[r, s, u]  ω : ?v  r  s  u»"
        proof -
          (*
           * TODO: I think this highlights the current issue with assoc_in_hom:
           * it can't be applied automatically here because there isn't any way to
           * obtain the equations src r = trg s and src s = trg u from assumption ω.
           * Maybe this can be improved with a little bit of thought, but not while
           * a lot of other stuff is being changed, too.
           *)
          have "src r = trg s  src s = trg u"
            by (metis ω arr_cod hseq_char in_homE hcomp_simps(1))
          thus ?thesis
            using u ω by fastforce
        qed

        obtain wr θr νr
        where wrθrνr: "ide wr  «θr : r0  wr  s  u» 
                               «νr : ?v  r1  wr»  iso νr 
                               ρ.composite_cell wr θr  νr = 𝖺[r, s, u]  ω"
          using u ω ρ.T1 [of "s  u" "𝖺[r, s, u]  ω"]
          by (metis 1 ρ.ide_base σ.ide_base arr_cod composable ide_hcomp in_homE
              match_1 not_arr_null seq_if_composable)
        text ‹
$$
\xymatrix{
  && X \ar@ {.>}[ddl]^{w_r} \ar@/ul20pt/[dddll]_{v} \xtwocell[dddll]{}\omit{^{<1.5>\nu_r}}
  \ar@/ur20pt/[dddrr]^{u} \xtwocell[dddr]{}\omit{^{\theta_r}} \\
  && \\  
  & {\rm src}~\rho \ar[dl]_{r_1} \ar[dr]^{r_0} \dtwocell\omit{^\rho}
  && \\
  {\rm trg}~r && {\rm src}~r = {\rm trg}~s \ar[ll]^{r} && {\rm src}~s \ar[ll]^{s}
}
$$
        ›
        text ‹We need some simps, etc., otherwise the subsequent diagram chase is very painful.›
        have wr: "ide wr"
          using wrθrνr by simp
        have [simp]: "src wr = src u"
           using wrθrνr ω 1 comp_arr_dom in_homE seqE hcomp_simps(1) vseq_implies_hpar(1)
           by (metis src_hcomp)
        have [simp]: "trg wr = src ρ"
          using wrθrνr
          by (metis 1 arrI not_arr_null seqE seq_if_composable)
        have θr_in_hom [intro]: "«θr : r0  wr  s  u»"
          using wrθrνr by simp
        have θr_in_hhom [intro]: "«θr : src u  trg s»"
          using θr_in_hom src_cod [of θr] trg_cod [of θr]
          by (metis src wr = src u σ.leg1_simps(4) arr_dom in_hhomI in_homE r0s1.cospan
              src_hcomp trg_hcomp vconn_implies_hpar(1) vconn_implies_hpar(2))
        have [simp]: "src θr = src u" using θr_in_hhom by auto
        have [simp]: "trg θr = trg s" using θr_in_hhom by auto
        have [simp]: "dom θr = r0  wr" using θr_in_hom by blast
        have [simp]: "cod θr = s  u" using θr_in_hom by blast
        have νr_in_hom [intro]: "«νr : ?v  r1  wr»" using wrθrνr by simp
        have νr_in_hhom [intro]: "«νr : src u  trg r»"
          using νr_in_hom src_dom [of νr] trg_dom [of νr]
          by (metis src wr = src u ρ.leg1_simps(4) arr_cod in_hhomI in_homE
              src_hcomp trg_hcomp vconn_implies_hpar(3) vconn_implies_hpar(4))
        have [simp]: "src νr = src u" using νr_in_hhom by auto
        have [simp]: "trg νr = trg r" using νr_in_hhom by auto
        have [simp]: "dom νr = ?v" using νr_in_hom by auto
        have [simp]: "cod νr = r1  wr" using νr_in_hom by auto
        have iso_νr: "iso νr" using wrθrνr by simp

        obtain ws θs νs
        where wsθsνs: "ide ws  «θs : s0  ws  u»  «νs : r0  wr  s1  ws»  iso νs 
                      σ.composite_cell ws θs  νs = θr"
          using u wrθrνr σ.T1 [of u θr] by auto
        text ‹
$$
\xymatrix{
  && X \ar[ddl]^{w_r} \ar@/ul20pt/[dddll]_{v} \xtwocell[dddll]{}\omit{^{<1.5>\nu_r}}
  \ar@/ur20pt/[dddrr]^{u} \ar@ {.>}[ddr]_{w_s} \xtwocell[dddrr]{}\omit{^{<-1.5>\theta_s}}
  \xtwocell[ddd]{}\omit{^{<1>\nu_s}} \\
  && \\  
  & {\rm src}~\rho \ar[dl]_{r_1} \ar[dr]^{r_0} \dtwocell\omit{^\rho}
  && {\rm src}~\sigma  \ar[dl]_{s_1} \ar[dr]^{s_0} \dtwocell\omit{^\sigma}\\
  {\rm trg}~r && {\rm src}~r = {\rm trg}~s \ar[ll]^{r} && {\rm src}~s \ar[ll]^{s}
}
$$
        ›
        have ws: "ide ws"
          using wsθsνs by simp
        have [simp]: "src ws = src u"
          using wsθsνs src_cod
          by (metis σ.leg0_simps(2) σ.tab_simps(2) θr_in_hom arrI hseqI' ideD(1) seqE
              seq_if_composable src_hcomp vconn_implies_hpar(3))
        have [simp]: "trg ws = src σ"
          using wsθsνs
          by (metis σ.tab_simps(2) arr_dom in_homE not_arr_null seq_if_composable)
        have θs_in_hom [intro]: "«θs : s0  ws  u»"
          using wsθsνs by simp
        have θs_in_hhom [intro]: "«θs : src u  src s»"
          using θs_in_hom src_cod trg_cod
          by (metis θr_in_hom arrI hseqE in_hhom_def seqE vconn_implies_hpar(1)
              vconn_implies_hpar(3) wsθsνs)
        have [simp]: "src θs = src u" using θs_in_hhom by auto
        have [simp]: "trg θs = src s" using θs_in_hhom by auto
        have [simp]: "dom θs = s0  ws" using θs_in_hom by blast
        have [simp]: "cod θs = u" using θs_in_hom by blast
        have νs_in_hom [intro]: "«νs : r0  wr  s1  ws»" using wsθsνs by simp
        have νs_in_hhom [intro]: "«νs : src u  trg s»"
          using νs_in_hom src_dom trg_cod
          by (metis src θr = src u trg θr = trg s θr_in_hom in_hhomI in_homE src_dom trg_dom)
        have [simp]: "src νs = src u" using νs_in_hhom by auto
        have [simp]: "trg νs = trg s" using νs_in_hhom by auto
        have [simp]: "dom νs = r0  wr" using νs_in_hom by auto
        have [simp]: "cod νs = s1  ws" using νs_in_hom by auto
        have iso_νs: "iso νs" using wsθsνs by simp

        obtain w θt νt
        where tνt: "ide w  «θt : p0  w  ws»  «νt : wr  p1  w»  iso νt 
                    (s1  θt)  𝖺[s1, p0, w]  (r0s1  w)  𝖺-1[r0, p1, w]  (r0  νt) = νs"
          using wrθrνr wsθsνs iso_νs r0s1.φ_biuniversal_prop(1) [of ws wr νs] by blast
        text ‹
$$
\xymatrix{
  && X \ar[ddl]_{w_r} \ar@/ul20pt/[dddll]_{v} \xtwocell[dddll]{}\omit{^{<1.5>\nu_r}}
  \ar@/ur20pt/[dddrr]^{u} \ar[ddr]^{w_s} \xtwocell[dddrr]{}\omit{^{<-1.5>\theta_s}}
  \ar@ {.>}[d]^{w} \xtwocell[ddl]{}\omit{^<-2>{\nu_t}} \xtwocell[ddr]{}\omit{^<2>{\theta_t}} \\
  && {\rm src}~\phi \ar[dl]_{p_1} \ar[dr]^{p_0} \ddtwocell\omit{^\phi} \\  
  & {\rm src}~\rho \ar[dl]_{r_1} \ar[dr]^{r_0} \dtwocell\omit{^\rho}
  && {\rm src}~\sigma  \ar[dl]_{s_1} \ar[dr]^{s_0} \dtwocell\omit{^\sigma}\\
  {\rm trg}~r && {\rm src}~r = {\rm trg}~s \ar[ll]^{r} && {\rm src}~s \ar[ll]^{s}
}
$$
        ›
        text ‹{\bf Note:} w› is not necessarily a map.›
        have w: "ide w"
          using tνt by simp
        have [simp]: "src w = src u"
          using tνt src_cod
          by (metis νs_in_hom src νs = src u arrI seqE src_hcomp src_vcomp vseq_implies_hpar(1))
        have [simp]: "trg w = src p0"
          using tνt
          by (metis νs_in_hom arrI not_arr_null r0s1.φ_simps(2) seqE seq_if_composable)
        have θt_in_hom [intro]: "«θt : p0  w  ws»"
          using tνt by simp
        have θt_in_hhom [intro]: "«θt : src u  src σ»"
          using θt_in_hom src_cod trg_cod src ws = src u trg ws = src σ by fastforce
        have [simp]: "src θt = src u" using θt_in_hhom by auto
        have [simp]: "trg θt = src σ" using θt_in_hhom by auto
        have [simp]: "dom θt = p0  w" using θt_in_hom by blast
        have (* [simp]: *) "cod θt = ws" using θt_in_hom by blast
        have νt_in_hom [intro]: "«νt : wr  p1  w»" using tνt by simp
        have νt_in_hhom [intro]: "«νt : src u  src ρ»"
          using νt_in_hom src_dom trg_dom src wr = src u trg wr = src ρ by fastforce
        have [simp]: "src νt = src u" using νt_in_hhom by auto
        have [simp]: "trg νt = src ρ" using νt_in_hhom by auto
        have (* [simp]: *) "dom νt = wr" using νt_in_hom by auto
        have [simp]: "cod νt = p1  w" using νt_in_hom by auto
        have iso_νt: "iso νt" using tνt by simp

        define θ where "θ = θs  (s0  θt)  𝖺[s0, p0, w]"
        have θ: "«θ : (s0  p0)  w  u»"
        proof (unfold θ_def, intro comp_in_homI)
          show "«𝖺[s0, p0, w] : (s0  p0)  w  s0  p0  w»"
            using tνt by auto
          show "«s0  θt : s0  p0  w  s0  ws»"
            using tνt by auto
          show "«θs : s0  ws  u»"
            using wsθsνs by simp
        qed
        define ν where "ν = 𝖺-1[r1, p1, w]  (r1  νt)  νr"
        have ν: "«ν : ?v  (r1  p1)  w»"
        proof (unfold ν_def, intro comp_in_homI)
          show "«νr : ?v  r1  wr»"
            using wrθrνr by blast
          show "«r1  νt : r1  wr  r1  p1  w»"
            using tνt by (intro hcomp_in_vhom, auto)
          show "«𝖺-1[r1, p1, w] : r1  p1  w  (r1  p1)  w»"
            using tνt assoc_in_hom by (simp add: ρ.T0.antipar(1))
        qed
        have iso_ν: "iso ν"
          using ν tνt wrθrνr ρ.T0.antipar(1)
          by (unfold ν_def, intro isos_compose) auto
        have *: "arr ((s  θs)  𝖺[s, s0, ws]  (σ  ws)  (s1  θt)  𝖺[s1, p0, w] 
                       (r0s1  w)  𝖺-1[r0, p1, w]  (r0  νt))"
          using wsθsνs tνt θr_in_hom comp_assoc by auto

        have "((r  s)  θ)  𝖺[r  s, s0  p0, w]  (tab  w)  ν = ω"
        proof -
          have "seq (r  θr) (𝖺[r, r0, wr]  (ρ  wr)  νr)"
            using wrθrνr ρ.base_simps(2) composable by fastforce
          hence "ω = 𝖺-1[r, s, u]  ρ.composite_cell wr θr  νr"
            using wrθrνr invert_side_of_triangle(1) iso_assoc
            by (metis 1 ρ.ide_base σ.ide_base arrI assoc'_eq_inv_assoc composable hseq_char'
                seqE seq_if_composable u vconn_implies_hpar(2) vconn_implies_hpar(4) wsθsνs)
          also have "... = 𝖺-1[r, s, u]  ρ.composite_cell wr (σ.composite_cell ws θs  νs)  νr"
            using wsθsνs by simp
          also have "... = 𝖺-1[r, s, u]  (r  (s  θs)  𝖺[s, s0, ws] 
                             (σ  ws)  (s1  θt)  𝖺[s1, p0, w]  (r0s1  w) 
                             𝖺-1[r0, p1, w]  (r0  νt))  𝖺[r, r0, wr]  (ρ  wr)  νr"
            using tνt comp_assoc by simp
          text ‹Rearrange to create θ› and ν›, leaving tab› in the middle.›
          also have "... = 𝖺-1[r, s, u]  (r  (s  θs)  𝖺[s, s0, ws] 
                             ((σ  ws)  (s1  θt))  𝖺[s1, p0, w]  (r0s1  w) 
                             𝖺-1[r0, p1, w]  (r0  νt))  𝖺[r, r0, wr]  (ρ  wr)  νr"
            using comp_assoc by presburger
          also have "... = 𝖺-1[r, s, u]  (r  (s  θs)  (𝖺[s, s0, ws] 
                             ((s  s0)  θt))  (σ  p0  w)  𝖺[s1, p0, w]  (r0s1  w) 
                             𝖺-1[r0, p1, w]  (r0  νt))  𝖺[r, r0, wr]  (ρ  wr)  νr"
          proof -
            have "(σ  ws)  (s1  θt) = σ  θt"
              using comp_arr_dom comp_cod_arr interchange
              by (metis cod θt = ws σ.tab_simps(1) σ.tab_simps(4) arrI tνt)
            also have "... = ((s  s0)  θt)  (σ  p0  w)"
              using comp_arr_dom comp_cod_arr interchange wsθsνs tνt σ.tab_in_hom
              by (metis dom θt = p0  w σ.tab_simps(5) arrI)
            finally have "(σ  ws)  (s1  θt) = ((s  s0)  θt)  (σ  p0  w)"
              by simp
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... = 𝖺-1[r, s, u]  (r  (s  θs)  (s  s0  θt)  𝖺[s, s0, p0  w] 
                             (σ  p0  w)  𝖺[s1, p0, w]  (r0s1  w) 
                             𝖺-1[r0, p1, w]  (r0  νt))  𝖺[r, r0, wr]  (ρ  wr)  νr"
              using assoc_naturality [of s s0 θt] tνt comp_assoc cod θt = ws arrI by force
          also have "... = 𝖺-1[r, s, u]  (r  (s  θs)  (s  s0  θt)) 
                             (r  𝖺[s, s0, p0  w]  (σ  p0  w)  𝖺[s1, p0, w] 
                             (r0s1  w)  𝖺-1[r0, p1, w]  (r0  νt))  𝖺[r, r0, wr] 
                             (ρ  wr)  νr"
          proof -
            have "r  (s  θs)  (s  s0  θt)  𝖺[s, s0, p0  w] 
                             (σ  p0  w)  𝖺[s1, p0, w]  (r0s1  w) 
                             𝖺-1[r0, p1, w]  (r0  νt) =
                  (r  (s  θs)  (s  s0  θt)) 
                    (r  𝖺[s, s0, p0  w]  (σ  p0  w)  𝖺[s1, p0, w]  (r0s1  w) 
                             𝖺-1[r0, p1, w]  (r0  νt))"
            proof -
              have "seq ((s  θs)  (s  s0  θt))
                        (𝖺[s, s0, p0  w]  (σ  p0  w)  𝖺[s1, p0, w]  (r0s1  w) 
                          𝖺-1[r0, p1, w]  (r0  νt))"
              proof -
                have "seq (s  θs)
                          ((s  s0  θt)  𝖺[s, s0, p0  w]  (σ  p0  w)  𝖺[s1, p0, w] 
                            (r0s1  w)  𝖺-1[r0, p1, w]  (r0  νt))"
                  using «𝖺[r, s, u]  ω : dom ω  r  s  u» calculation by blast
                thus ?thesis
                  using comp_assoc by presburger
              qed
              thus ?thesis
                using whisker_left [of r "(s  θs)  (s  s0  θt)"
                                         "𝖺[s, s0, p0  w]  (σ  p0  w)  𝖺[s1, p0, w] 
                                            (r0s1  w)  𝖺-1[r0, p1, w]  (r0  νt)"]
                    wsθsνs tνt comp_assoc
                by simp
            qed
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... = 𝖺-1[r, s, u]  (r  (s  θs)  (s  s0  θt)) 
                             (r  𝖺[s, s0, p0  w]  (σ  p0  w)  𝖺[s1, p0, w] 
                             (r0s1  w)  𝖺-1[r0, p1, w])  ((r  r0  νt)  𝖺[r, r0, wr]) 
                             (ρ  wr)  νr"
          proof -
            have "seq (𝖺[s, s0, p0  w]  (σ  p0  w)  𝖺[s1, p0, w]  (r0s1  w) 
                      𝖺-1[r0, p1, w]) (r0  νt)"
              using 1 r0s1.p1_simps tνt
              apply (intro seqI' comp_in_homI) by auto
            hence "r  (𝖺[s, s0, p0  w]  (σ  p0  w)  𝖺[s1, p0, w]  (r0s1  w) 
                     𝖺-1[r0, p1, w])  (r0  νt) =
                   (r  𝖺[s, s0, p0  w]  (σ  p0  w)  𝖺[s1, p0, w]  (r0s1  w) 
                     𝖺-1[r0, p1, w])  (r  r0  νt)"
              using whisker_left by simp
            thus ?thesis
              using comp_assoc by simp
          qed
          also have "... = 𝖺-1[r, s, u]  (r  (s  θs)  (s  s0  θt)) 
                             (r  𝖺[s, s0, p0  w]  (σ  p0  w)  𝖺[s1, p0, w] 
                             (r0s1  w)  𝖺-1[r0, p1, w])  𝖺[r, r0, p1  w] 
                             (((r  r0)  νt)  (ρ  wr))  νr"
          proof -
            have "(r  r0  νt)  𝖺[r, r0, wr] = 𝖺[r, r0, p1  w]  ((r  r0)  νt)"
              using assoc_naturality [of r r0 νt] νt_in_hom by auto
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... = (𝖺-1[r, s, u]  (r  (s  θs)  (s  s0  θt))) 
                             (r  𝖺[s, s0, p0  w]  (σ  p0  w)  𝖺[s1, p0, w] 
                             (r0s1  w)  𝖺-1[r0, p1, w])  𝖺[r, r0, p1  w] 
                             (ρ  p1  w)  (r1  νt)  νr"
          proof -
            have "((r  r0)  νt)  (ρ  wr) = ρ  νt"
              using comp_arr_dom comp_cod_arr interchange
              by (metis dom νt = wr ρ.tab_simps(1) ρ.tab_simps(5) arrI tνt)
            also have "... = (ρ  p1  w)  (r1  νt)"
              using comp_arr_dom comp_cod_arr interchange
              by (metis cod νt = p1  w trg νt = src ρ ρ.T0.antipar(1) ρ.tab_simps(1)
                  ρ.tab_simps(2) ρ.tab_simps(4) r0s1.base_simps(2) trg.preserves_reflects_arr
                  trg_hcomp)
            finally have "((r  r0)  νt)  (ρ  wr) = (ρ  p1  w)  (r1  νt)" by simp
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... = ((r  s)  θs  (s0  θt))  𝖺-1[r, s, s0  p0  w]  
                             (r  𝖺[s, s0, p0  w]  ((σ  p0  w)  𝖺[s1, p0, w]) 
                             (r0s1  w)  𝖺-1[r0, p1, w])  𝖺[r, r0, p1  w] 
                             (ρ  p1  w)  (r1  νt)  νr"
          proof -
            have "𝖺-1[r, s, u]  (r  (s  θs)  (s  s0  θt)) =
                  ((r  s)  θs  (s0  θt))  𝖺-1[r, s, s0  p0  w]"
            proof -
              have "seq (s  θs) (s  s0  θt)"
                using θs_in_hom θs_in_hhom θt_in_hom θt_in_hhom 1 calculation by blast
              moreover have "src r = trg (s  θs)"
                using composable hseqI by force
              ultimately
              have "𝖺-1[r, s, u]  (r  (s  θs)  (s  s0  θt)) =
                    (𝖺-1[r, s, u]  (r  s  θs))  (r  s  s0  θt)"
                using whisker_left comp_assoc by simp
              also have "... = ((r  s)  θs)  𝖺-1[r, s, s0  ws]  (r  s  s0  θt)"
                using assoc_naturality comp_assoc
                by (metis cod θs = u dom θs = s0  ws trg θs = src s
                    ρ.base_simps(2-4) σ.base_simps(2-4) arrI assoc'_naturality composable wsθsνs)
              also have "... = (((r  s)  θs)  ((r  s)  s0  θt))  𝖺-1[r, s, s0  p0  w]"
              proof -
                have "𝖺-1[r, s, s0  ws]  (r  s  s0  θt) =
                      ((r  s)  s0  θt)  𝖺-1[r, s, s0  p0  w]"
                  using arrI hseq_char assoc'_naturality [of r s "s0  θt"] cod θt = ws composable
                  by auto
                thus ?thesis
                  using comp_assoc by auto
              qed
              also have "... = ((r  s)  θs  (s0  θt))  𝖺-1[r, s, s0  p0  w]"
                using θ_def θ whisker_left
                by (metis (full_types) arrI cod_comp ide_base seqE seqI)
              finally show ?thesis by simp
            qed
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... = ((r  s)  θs  (s0  θt))  𝖺-1[r, s, s0  p0  w]  
                             ((r  𝖺[s, s0, p0  w]  𝖺[s  s0, p0, w] 
                               ((σ  p0)  w)  (r0s1  w)  𝖺-1[r0, p1, w])) 
                             𝖺[r, r0, p1  w]  (ρ  p1  w)  (r1  νt)  νr"
          proof -
            have "(σ  p0  w)  𝖺[s1, p0, w] = 𝖺[s  s0, p0, w]  ((σ  p0)  w)"
              using assoc_naturality [of σ p0 w] by (simp add: tνt)
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... = ((r  s)  θs  (s0  θt)) 
                             𝖺-1[r, s, s0  p0  w]  (r  𝖺[s, s0, p0  w]) 
                             (r  𝖺[s  s0, p0, w])  (r  (σ  p0)  w)  (r  r0s1  w) 
                             ((r  𝖺-1[r0, p1, w])  𝖺[r, r0, p1  w]) 
                             (ρ  p1  w)  (r1  νt)  νr"
            using r0s1.p1_simps tνt whisker_left comp_assoc by force
          also have "... = ((r  s)  θs  (s0  θt)) 
                             𝖺-1[r, s, s0  p0  w]  (r  𝖺[s, s0, p0  w]) 
                             (r  𝖺[s  s0, p0, w])  (r  (σ  p0)  w)  (r  r0s1  w) 
                             (𝖺[r, r0  p1, w]  (𝖺[r, r0, p1]  w)  (𝖺-1[r  r0, p1, w]) 
                             (ρ  p1  w))  (r1  νt)  νr"
          proof -
            have "(r  𝖺-1[r0, p1, w])  𝖺[r, r0, p1  w] =
                  𝖺[r, r0  p1, w]  (𝖺[r, r0, p1]  w)  𝖺-1[r  r0, p1, w]"
            proof -
              have 1: "(r  𝖺[r0, p1, w])  𝖺[r, r0  p1, w]  (𝖺[r, r0, p1]  w) =
                       𝖺[r, r0, p1  w]  𝖺[r  r0, p1, w]"
                using pentagon
                by (simp add: ρ.T0.antipar(1) w)
              moreover have 2: "seq 𝖺[r, r0, p1  w] 𝖺[r  r0, p1, w]"
                using ρ.T0.antipar(1) w by simp
              moreover have "inv (r  𝖺[r0, p1, w]) = r  𝖺-1[r0, p1, w]"
                using ρ.T0.antipar(1) w by simp
              ultimately have "𝖺[r, r0  p1, w]  (𝖺[r, r0, p1]  w) =
                                 ((r  𝖺-1[r0, p1, w])  𝖺[r, r0, p1  w])  𝖺[r  r0, p1, w]"
                using ρ.T0.antipar(1) w comp_assoc
                      invert_side_of_triangle(1)
                        [of "𝖺[r, r0, p1  w]  𝖺[r  r0, p1, w]" "r  𝖺[r0, p1, w]"
                            "𝖺[r, r0  p1, w]  (𝖺[r, r0, p1]  w)"]
                by simp
              hence "(r  𝖺-1[r0, p1, w])  𝖺[r, r0, p1  w] =
                     (𝖺[r, r0  p1, w]  (𝖺[r, r0, p1]  w))  𝖺-1[r  r0, p1, w]"
                using ρ.T0.antipar(1) w
                      invert_side_of_triangle(2)
                        [of "𝖺[r, r0  p1, w]  (𝖺[r, r0, p1]  w)"
                            "(r  𝖺-1[r0, p1, w])  𝖺[r, r0, p1  w]" "𝖺[r  r0, p1, w]"]
                using trg w = src p0 by simp
              thus ?thesis
                using comp_assoc by presburger
            qed
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... = ((r  s)  θs  (s0  θt)) 
                             𝖺-1[r, s, s0  p0  w]  (r  𝖺[s, s0, p0  w]) 
                             (r  𝖺[s  s0, p0, w])  (r  (σ  p0)  w)  ((r  r0s1  w) 
                             𝖺[r, r0  p1, w])  (𝖺[r, r0, p1]  w)  ((ρ  p1)  w) 
                             𝖺-1[r1, p1, w]  (r1  νt)  νr"
          proof -
            have "𝖺-1[r  r0, p1, w]  (ρ  p1  w) = ((ρ  p1)  w)  𝖺-1[r1, p1, w]"
              using assoc'_naturality [of ρ p1 w] by (simp add: ρ.T0.antipar(1) tνt)
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... = ((r  s)  θs  (s0  θt)) 
                             𝖺-1[r, s, s0  p0  w]  (r  𝖺[s, s0, p0  w]) 
                             (r  𝖺[s  s0, p0, w])  ((r  (σ  p0)  w)  𝖺[r, s1  p0, w]) 
                             ((r  r0s1)  w)  (𝖺[r, r0, p1]  w)  ((ρ  p1)  w) 
                             𝖺-1[r1, p1, w]  (r1  νt)  νr"
          proof -
            have "(r  r0s1  w)  𝖺[r, r0  p1, w] = 𝖺[r, s1  p0, w]  ((r  r0s1)  w)"
              using assoc_naturality [of r r0s1 w] r0s1.cospan tνt by auto
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... = ((r  s)  θs  (s0  θt)) 
                             𝖺-1[r, s, s0  p0  w]  (r  𝖺[s, s0, p0  w]) 
                             (r  𝖺[s  s0, p0, w])  𝖺[r, (s  s0)  p0, w] 
                             (((r  σ  p0)  w)  ((r  r0s1)  w)  (𝖺[r, r0, p1]  w) 
                               ((ρ  p1)  w)) 
                             𝖺-1[r1, p1, w]  (r1  νt)  νr"
          proof -
            have "(r  (σ  p0)  w)  𝖺[r, s1  p0, w] =
                  𝖺[r, (s  s0)  p0, w]  ((r  σ  p0)  w)"
            proof -
              have "arr w  dom w = w  cod w = w"
                using ide_char w by blast
              then show ?thesis
                using assoc_naturality [of r "σ  p0" w] composable by auto
            qed
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... = ((r  s)  θs  (s0  θt)) 
                             (𝖺-1[r, s, s0  p0  w]  (r  𝖺[s, s0, p0  w]) 
                             (r  𝖺[s  s0, p0, w])  𝖺[r, (s  s0)  p0, w] 
                             ((r  𝖺-1[s, s0, p0])  w)  (𝖺[r, s, s0  p0]  w)) 
                             (tab  w)  𝖺-1[r1, p1, w]  (r1  νt)  νr"
          proof -
            have "((r  σ  p0)  w)  ((r  r0s1)  w)  (𝖺[r, r0, p1]  w)  ((ρ  p1)  w) =
                  (r  σ  p0)  (r  r0s1)  𝖺[r, r0, p1]  (ρ  p1)  w"
              using w ρ.T0.antipar(1) composable whisker_right by auto
            also have "... = (((r  𝖺-1[s, s0, p0])  (𝖺[r, s, s0  p0] 
                               𝖺-1[r, s, s0  p0])  (r  𝖺[s, s0, p0]))  (r  σ  p0)) 
                               (r  r0s1)  𝖺[r, r0, p1]  (ρ  p1)  w"
            proof -
              have "((r  𝖺-1[s, s0, p0])  (𝖺[r, s, s0  p0]  𝖺-1[r, s, s0  p0])  
                      (r  𝖺[s, s0, p0]))  (r  σ  p0) =
                    r  σ  p0"
              proof -
                have "((r  𝖺-1[s, s0, p0])  (𝖺[r, s, s0  p0]  𝖺-1[r, s, s0  p0])  
                        (r  𝖺[s, s0, p0]))  (r  σ  p0) =
                      ((r  𝖺-1[s, s0, p0])  ((r  s  s0  p0)  (r  𝖺[s, s0, p0])))  (r  σ  p0)"
                  using comp_assoc_assoc' by (simp add: composable)
                also have "... = ((r  𝖺-1[s, s0, p0])  (r  𝖺[s, s0, p0]))  (r  σ  p0)"
                  using comp_cod_arr by (simp add: composable)
                also have "... = ((r  (s  s0)  p0))  (r  σ  p0)"
                  using whisker_left comp_assoc_assoc' assoc_in_hom hseqI'
                  by (metis ρ.ide_base σ.base_simps(2) σ.ide_base σ.ide_leg0
                      σ.leg0_simps(2-3) σ.leg1_simps(3) r0s1.ide_leg0
                      r0s1.leg0_simps(2) r0s1.p0_simps hcomp_simps(1))
                also have "... = r  σ  p0"
                  using comp_cod_arr
                  by (simp add: composable)
                finally show ?thesis by blast
              qed
              thus ?thesis by simp
            qed
            also have "... = (r  𝖺-1[s, s0, p0])  𝖺[r, s, s0  p0] 
                               (𝖺-1[r, s, s0  p0])  (r  𝖺[s, s0, p0])  (r  σ  p0) 
                               (r  r0s1)  𝖺[r, r0, p1]  (ρ  p1)  w"
              using comp_assoc by presburger
            also have "... = (r  𝖺-1[s, s0, p0])  𝖺[r, s, s0  p0]  tab  w"
              using tab_def by simp
            also have "... = ((r  𝖺-1[s, s0, p0])  w)  (𝖺[r, s, s0  p0]  w)  (tab  w)"
              using w ρ.T0.antipar(1) composable comp_assoc whisker_right by auto
            finally have "((r  σ  p0)  w)  ((r  r0s1)  w)  (𝖺[r, r0, p1]  w) 
                            ((ρ  p1)  w) =
                          ((r  𝖺-1[s, s0, p0])  w)  (𝖺[r, s, s0  p0]  w)  (tab  w)"
              by simp
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... = (((r  s)  θs  (s0  θt))  ((r  s)  𝖺[s0, p0, w])) 
                             𝖺[r  s, s0  p0, w]  (tab  w)  ν"
          proof -
            have "𝖺-1[r, s, s0  p0  w]  (r  𝖺[s, s0, p0  w])  (r  𝖺[s  s0, p0, w]) 
                   𝖺[r, (s  s0)  p0, w]  ((r  𝖺-1[s, s0, p0])  w)  (𝖺[r, s, s0  p0]  w) =
                  ((r  s)  𝖺[s0, p0, w])  𝖺[r  s, s0  p0, w]"
            proof -
              have "𝖺-1[r, s, s0  p0  w]  (r  𝖺[s, s0, p0  w])  (r  𝖺[s  s0, p0, w]) 
                      𝖺[r, (s  s0)  p0, w]  ((r  𝖺-1[s, s0, p0])  w) 
                      (𝖺[r, s, s0  p0]  w) =
                    𝖺-1[r, s, s0  p0  w]  (r  𝖺[s, s0, p0  w]) 
                      (r  𝖺[s  s0, p0, w])  𝖺[r, (s  s0)  p0, w] 
                      ((r  𝖺-1[s, s0, p0])  w)  (𝖺[r, s, s0  p0]  w)"
                using w comp_assoc 𝖺'_def α_def composable by simp
              also have "... = ((r  s)  𝖺[s0, p0, w])  𝖺[r  s, s0  p0, w]"
                using w composable
                apply (intro E.eval_eqI) by simp_all
              also have "... = ((r  s)  𝖺[s0, p0, w])  𝖺[r  s, s0  p0, w]"
                using w comp_assoc 𝖺'_def α_def composable by simp
              finally show ?thesis by simp
            qed
            thus ?thesis
              using ν_def comp_assoc by simp
          qed
          also have "... = ((r  s)  θ)  𝖺[r  s, s0  p0, w]  (tab  w)  ν"
          proof -
            have "((r  s)  θs  (s0  θt))  ((r  s)  𝖺[s0, p0, w]) = (r  s)  θ"
              using θ_def w whisker_left composable
              by (metis θ arrI ide_base comp_assoc)
            thus ?thesis
              using comp_assoc by presburger
          qed
          finally show "((r  s)  θ)  𝖺[r  s, s0  p0, w]  (tab  w)  ν = ω"
            by simp
        qed
        thus "w θ ν. ide w  «θ : (s0  p0)  w  u» 
                      «ν : dom ω  (r1  p1)  w»  iso ν 
                      composite_cell w θ  ν = ω"
          using tνt θ ν iso_ν comp_assoc by metis
      qed

      show "u w w' θ θ' β.
               ide w; ide w'; «θ : (s0  p0)  w  u»; «θ' : (s0  p0)  w'  u»;
                «β : (r1  p1)  w  (r1  p1)  w'»;
                composite_cell w θ = composite_cell w' θ'  β  
              ∃!γ. «γ : w  w'»  β = (r1  p1)  γ  θ = θ'  ((s0  p0)  γ)"
      proof -
        fix u w w' θ θ' β
        assume w: "ide w"
        assume w': "ide w'"
        assume θ: "«θ : (s0  p0)  w  u»"
        assume θ': "«θ' : (s0  p0)  w'  u»"
        assume β: "«β : (r1  p1)  w  (r1  p1)  w'»"
        assume eq: "composite_cell w θ = composite_cell w' θ'  β"
        interpret uwθw'θ'β: uwθw'θ'β V H 𝖺 𝗂 src trg
                              r  s tab s0  p0 r1  p1 u w θ w' θ' β
          using w w' θ θ' β eq composable tab_in_hom comp_assoc
          by unfold_locales auto
        text ‹
$$
\begin{array}{ll}
\xymatrix{
  && X \ar[d]_{w'} \xtwocell[ddl]{}\omit{^{\beta}}
  \ar@/ul20pt/[dddll]_<>(0.25){w}|<>(0.33)@ {>}_<>(0.5){p_1}|<>(0.67)@ {>}_<>(0.75){r_1}
  \ar@/ur20pt/[dddrr]^{u} \xtwocell[ddr]{}\omit{^{\theta'}}\\
  && {\rm src}~\phi \ar[dl]^{p_1} \ar[dr]_{p_0} \ddtwocell\omit{^\phi} \\  
  & {\rm src}~\rho \ar[dl]_{r_1} \ar[dr]^{r_0} \dtwocell\omit{^\rho}
  && {\rm src}~\sigma  \ar[dl]_{s_1} \ar[dr]^{s_0} \dtwocell\omit{^\sigma}\\
  {\rm trg}~r && {\rm src}~r = {\rm trg}~s \ar[ll]^{r} && {\rm src}~s \ar[ll]^{s}
}
\\
\hspace{5cm}
=
\qquad
\xy/50pt/
\xymatrix{
  && X \ar[d]_{w} \ar@/ur20pt/[dddrr]^{u} \xtwocell[ddr]{}\omit{^{\theta}}\\
  && {\rm src}~\phi \ar[dl]^{p_1} \ar[dr]_{p_0} \ddtwocell\omit{^\phi} \\  
  & {\rm src}~\rho \ar[dl]_{r_1} \ar[dr]^{r_0} \dtwocell\omit{^\rho}
  && {\rm src}~\sigma  \ar[dl]_{s_1} \ar[dr]^{s_0} \dtwocell\omit{^\sigma}\\
  {\rm trg}~r && {\rm src}~r = {\rm trg}~s \ar[ll]^{r} && {\rm src}~s \ar[ll]^{s}
}
\endxy
\end{array}
$$
        ›
        text ‹
          First apply property ρ.T2› using «βr : r1 ⋆ p1 ⋆ w ⇒ r1 ⋆ p1 ⋆ w'»›
          (obtained by composing β› with associativities) and ``everything to the right'' 
          as θr and θr'›.  This yields «γr : p1 ⋆ w ⇒ p1 ⋆ w'»›.
          Next, apply property ρ.T2› to obtain «γs : p0 ⋆ w ⇒ p0 ⋆ w'»›.
          For this use «θs : s0 ⋆ p0 ⋆ w ⇒ u»› and «θs' : s0 ⋆ p0 ⋆ w'»›
          obtained by composing θ› and θ'› with associativities.
          We also need «βs : s1 ⋆ p0 ⋆ w ⇒ s1 ⋆ p0 ⋆ w'»›.
          To get this, transport r0 ⋆ γr across φ›; we need φ› to be an isomorphism
          in order to do this.
          Finally, apply the biuniversal property of φ› to get «γ : w ⇒ w'»›
          and verify the required equation.
        ›
        let ?wr = "p1  w"
        have wr: "ide ?wr" by simp
        let ?wr' = "p1  w'"
        have wr': "ide ?wr'" by simp
        define θr where "θr = (s  θ)  (s  𝖺-1[s0, p0, w])  𝖺[s, s0, p0  w] 
                                (σ  p0  w)  𝖺[s1, p0, w]  (r0s1  w)  𝖺-1[r0, p1, w]"
        have θr: "«θr : r0  ?wr  s  u»"
          unfolding θr_def
          using ρ.T0.antipar(1) by fastforce
        define θr' where "θr' = (s  θ')  (s  𝖺-1[s0, p0, w'])  𝖺[s, s0, p0  w'] 
                                  (σ  p0  w')  𝖺[s1, p0, w']  (r0s1  w')  𝖺-1[r0, p1, w']"
        have θr': "«θr' : r0  ?wr'  s  u»"
          unfolding θr'_def
          using ρ.T0.antipar(1) by fastforce
        define βr where "βr = 𝖺[r1, p1, w']  β  𝖺-1[r1, p1, w]"
        have βr: "«βr : r1  ?wr  r1  ?wr'»"
          unfolding βr_def
          using ρ.T0.antipar(1) by force
        have eqr: "ρ.composite_cell ?wr θr = ρ.composite_cell ?wr' θr'  βr"
        text ‹
$$
\begin{array}{ll}
\xymatrix{
  && X \ar[ddl]^{w_r'} \xtwocell[dddll]{}\omit{^<2>{\beta_r}}
  \ar@/ul20pt/[dddll]_<>(0.33){w_r}|<>(0.67)@ {>}_<>(0.75){r_1}
  \ar@/ur20pt/[dddrr]^{u} \xtwocell[dddr]{}\omit{^{\theta_r'}}\\
  && \\  
  & {\rm src}~\rho \ar[dl]_{r_1} \ar[dr]^{r_0} \dtwocell\omit{^\rho}
  && \\
  {\rm trg}~r && {\rm src}~r = {\rm trg}~s \ar[ll]^{r} && {\rm src}~s \ar[ll]^{s}
}
\\
\hspace{5cm}
=\qquad
\xy/50pt/
\xymatrix{
  && X \ar[ddl]^{w_r} \ar@/ur20pt/[dddrr]^{u} \xtwocell[dddr]{}\omit{^{\theta_r}}\\
  && \\  
  & {\rm src}~\rho \ar[dl]_{r_1} \ar[dr]^{r_0} \dtwocell\omit{^\rho}
  && \\
  {\rm trg}~r && {\rm src}~r = {\rm trg}~s \ar[ll]^{r} && {\rm src}~s \ar[ll]^{s}
}
\endxy
\end{array}
$$
        ›
        proof -
          have "ρ.composite_cell ?wr θr = 𝖺[r, s, u]  composite_cell w θ  𝖺-1[r1, p1, w]"
            using θr_def technical uwθw'θ'β.uwθ.uwθ by blast
          also have "... = 𝖺[r, s, u]  (((r  s)  θ')  𝖺[r  s, s0  p0, w'] 
                             (tab  w')  β)  𝖺-1[r1, p1, w]"
            using eq comp_assoc by simp
          also have "... = (r  θr')  𝖺[r, r0, ?wr']  (ρ  ?wr')  βr"
          proof -
            have "𝖺[r, s, u]  (composite_cell w' θ'  β)  𝖺-1[r1, p1, w] =
                  𝖺[r, s, u]  composite_cell w' θ'  β  𝖺-1[r1, p1, w]"
              using comp_assoc by presburger
            also have "... = (r  (s  θ')  (s  𝖺-1[s0, p0, w'])  𝖺[s, s0, p0  w'] 
                                    (σ  p0  w')  𝖺[s1, p0, w']  (r0s1  w') 
                                    𝖺-1[r0, p1, w']) 
                                𝖺[r, r0, p1  w']  (ρ  p1  w')  𝖺[r1, p1, w']  β  𝖺-1[r1, p1, w]"
            proof -
              have "𝖺[r, s, u]  composite_cell w' θ'  β  𝖺-1[r1, p1, w] =
                    𝖺[r, s, u]  composite_cell w' θ' 
                      ((𝖺-1[r1, p1, w']  𝖺[r1, p1, w'])  β)  𝖺-1[r1, p1, w]"
              proof -
                have "(𝖺-1[r1, p1, w']  𝖺[r1, p1, w'])  β = β"
                  using comp_cod_arr ρ.T0.antipar(1) β comp_assoc_assoc' by simp
                thus ?thesis by argo
              qed
              also have "... = (𝖺[r, s, u]  ((r  s)  θ')  𝖺[r  s, s0  p0, w']  (tab  w') 
                                 𝖺-1[r1, p1, w'])  𝖺[r1, p1, w']  β  𝖺-1[r1, p1, w]"
                using comp_assoc by presburger
              also have "... = ((r  (s  θ')  (s  𝖺-1[s0, p0, w'])  𝖺[s, s0, p0  w'] 
                                 (σ  p0  w')  𝖺[s1, p0, w']  (r0s1  w') 
                                 𝖺-1[r0, p1, w'])  𝖺[r, r0, p1  w']  (ρ  p1  w')) 
                                 𝖺[r1, p1, w']  β  𝖺-1[r1, p1, w]"
                using θr'_def technical [of w' θ' u ?wr' θr'] comp_assoc by fastforce
              finally show ?thesis
                using comp_assoc by simp
            qed
            finally show ?thesis
              using θr'_def βr_def comp_assoc by auto
          qed
          finally show ?thesis
            using comp_assoc by presburger
        qed
        have 1: "∃!γ. «γ : ?wr  ?wr'»  θr = θr'  (r0  γ)  βr = r1  γ"
          using eqr ρ.T2 [of ?wr ?wr' θr "s  u" θr' βr] wr wr' θr θr' βr by blast
        obtain γr where γr: "«γr : ?wr  ?wr'»  θr = θr'  (r0  γr)  βr = r1  γr"
          using 1 by blast

        let ?ws = "p0  w"
        have ws: "ide ?ws" by simp
        let ?ws' = "p0  w'"
        have ws': "ide ?ws'" by simp
        define θs where "θs = θ  𝖺-1[s0, p0, w]"
        have θs: "«θs : s0  p0  w  u»"
          using θs_def by auto
        define θs' where "θs' = θ'  𝖺-1[s0, p0, w']"
        have θs': "«θs' : s0  p0  w'  u»"
          using θs'_def by auto
        define βs where "βs = 𝖺[s1, p0, w']  (r0s1  w')  𝖺-1[r0, p1, w']  (r0  γr) 
                                𝖺[r0, p1, w]  (inv r0s1  w)  𝖺-1[s1, p0, w]"
        have βs: "«βs : s1  ?ws  s1  ?ws'»"
          unfolding βs_def
          using γr r0s1.φ_in_hom(2) r0s1.φ_uniqueness(2) ρ.T0.antipar(1)
          apply (intro comp_in_homI)
                apply auto
          by auto
        have eqs: "σ.composite_cell (p0  w) (θ  𝖺-1[s0, p0, w]) =
                   σ.composite_cell (p0  w') (θ'  𝖺-1[s0, p0, w'])  βs"
        text ‹
$$
\begin{array}{ll}
\xy/67pt/
\xymatrix{
  && X \ar[d]^{w'} \ar@/l10pt/[dl]_{w} \ddltwocell\omit{^{\gamma_r}}
  \ar@/ur20pt/[dddrr]^{u} \xtwocell[ddr]{}\omit{^{\theta_s'}}\\
  & {\rm src}~\phi \ar[dr]_{p_1} \ar[d]_{p_0}
  & {\rm src}~\phi \ar[d]^{p_1} \ar[dr]_{p_0} \ddrtwocell\omit{^\phi} \xtwocell[ddl]{}\omit{^\;\;\;\;\phi^{-1}} \\  
  & {\rm src}~\sigma \ar[dr]_{s_1} & {\rm src}~\rho \ar[d]^{r_0}
  & {\rm src}~\sigma \ar[dl]_{s_1} \ar[dr]^{s_0} \dtwocell\omit{^\sigma}\\
  && {\rm src}~r = {\rm trg}~s && {\rm src}~s \ar[ll]^{s}
}
\endxy
\\
\hspace{5cm}=
\xy/50pt/
\xymatrix{
  & X \ar@/dl15pt/[ddr]_<>(0.5){w_s}
  \ar@/ur20pt/[dddrr]^{u} \xtwocell[ddr]{}\omit{^{\theta_s}}\\
  & \\  
  && {\rm src}~\sigma  \ar[dl]_{s_1} \ar[dr]^{s_0} \dtwocell\omit{^\sigma}\\
  & {\rm src}~r = {\rm trg}~s && {\rm src}~s \ar[ll]^{s}
}
\endxy
\end{array}
$$
        ›
        proof -
          have "σ.composite_cell (p0  w') (θ'  𝖺-1[s0, p0, w'])  βs =
                (θr'  (r0  γr))  𝖺[r0, p1, w]  (inv r0s1  w)  𝖺-1[s1, p0, w]"
            using βs_def θr'_def whisker_left comp_assoc by simp
          also have "... = θr  𝖺[r0, p1, w]  (inv r0s1  w)  𝖺-1[s1, p0, w]"
            using γr by simp
          also have "... = ((s  θ)  (s  𝖺-1[s0, p0, w]))  𝖺[s, s0, ?ws]  (σ  ?ws) 
                              𝖺[s1, p0, w]  ((r0s1  w)  (𝖺-1[r0, p1, w] 
                              𝖺[r0, p1, w])  (inv r0s1  w))  𝖺-1[s1, p0, w]"
            using θr_def comp_assoc by simp
          also have "... = (s  θ)  σ.composite_cell (p0  w) 𝖺-1[s0, p0, w]"
          proof -
            have "(σ  p0  w) 
                    𝖺[s1, p0, w]  ((r0s1  w)  (𝖺-1[r0, p1, w] 
                    𝖺[r0, p1, w])  (inv r0s1  w))  𝖺-1[s1, p0, w] =
                  σ  p0  w"
            proof -
              have "𝖺-1[r0, p1, w]  𝖺[r0, p1, w] = cod (inv r0s1  w)"
                using r0s1.φ_uniqueness(2) ρ.T0.antipar(1) comp_assoc_assoc' by simp
              text ‹Here the fact that φ› is a retraction is used.›
              moreover have "(r0s1  w)  (inv r0s1  w) = cod 𝖺-1[s1, p0, w]"
                using r0s1.φ_uniqueness(2) comp_arr_inv' whisker_right [of w r0s1 "inv r0s1"]
                by simp
              moreover have "𝖺[s1, p0, w]  𝖺-1[s1, p0, w] = dom (σ  p0  w)"
                using r0s1.base_simps(2) hseq_char comp_assoc_assoc' by auto
              moreover have "hseq (inv r0s1) w"
                using r0s1.φ_uniqueness(2)
                by (intro hseqI, auto)
              moreover have "hseq σ (p0  w)"
                by (intro hseqI, auto)
              ultimately show ?thesis
                using comp_arr_dom comp_cod_arr by simp
            qed
            thus ?thesis
              using comp_assoc by simp
          qed
          also have "... = σ.composite_cell (p0  w) (θ  𝖺-1[s0, p0, w])"
            using θs_def whisker_left
            by (metis σ.ide_base θs arrI comp_assoc)
          finally show ?thesis by simp
        qed
        hence 2: "∃!γ. «γ : ?ws  ?ws'»  θs = θs'  (s0  γ)  βs = s1  γ"
          using σ.T2 [of ?ws ?ws' θs u θs' βs] ws ws' θs θs' βs
          by (metis θs'_def θs_def)
        obtain γs where γs: "«γs : ?ws  ?ws'»  θs = θs'  (s0  γs)  βs = s1  γs"
          using 2 by blast

        have eqt: "(s1  γs)  𝖺[s1, p0, w]  (r0s1  w)  𝖺-1[r0, p1, w] =
                   (s1  ?ws')  𝖺[s1, p0, w']  (r0s1  w')  𝖺-1[r0, p1, w']  (r0  γr)"
        text ‹
$$
\xy/78pt/
\xymatrix{
  & X \ar[d]^{w'} \ar@/ul15pt/[ddl]_{w_r} \xtwocell[ddl]{}\omit{^{\gamma_r}} \\
  & {\rm src}~\phi \ar[dl]_{p_1} \ar[dr]^{p_0} \ddtwocell\omit{^\phi} \\  
  {\rm src}~\rho \ar[dr]^{r_0}
  && {\rm src}~\sigma \ar[dl]_{s_1} \\
  & {\rm src}~r = {\rm trg}~s &
}
\endxy
\qquad = \qquad
\xy/78pt/
\xymatrix{
  & X \ar[d]^{w} \ar@/ur15pt/[ddr]^{w_s'} \xtwocell[ddr]{}\omit{^{\gamma_s}} \\
  & {\rm src}~\phi \ar[dl]_{p_1} \ar[dr]^{p_0} \ddtwocell\omit{^\phi}  \\  
  {\rm src}~\rho \ar[dr]^{r_0}
  && {\rm src}~\sigma \ar[dl]_{s_1} \\
  & {\rm src}~r = {\rm trg}~s &
}
\endxy
$$
        ›
        proof -
          have "(s1  ?ws')  𝖺[s1, p0, w']  (r0s1  w')  𝖺-1[r0, p1, w']  (r0  γr) =
                βs  𝖺[s1, p0, w]  (r0s1  w)  𝖺-1[r0, p1, w]"
          proof -
            have "βs  𝖺[s1, p0, w]  (r0s1  w)  𝖺-1[r0, p1, w] =
                  (𝖺[s1, p0, w']  (r0s1  w')  𝖺-1[r0, p1, w'])  (r0  γr) 
                    𝖺[r0, p1, w]  ((inv r0s1  w)  (𝖺-1[s1, p0, w] 
                    𝖺[s1, p0, w])  (r0s1  w))  𝖺-1[r0, p1, w]"
              using βs_def comp_assoc by metis
            also have "... = (𝖺[s1, p0, w']  (r0s1  w')  𝖺-1[r0, p1, w'])  (r0  γr)"
            proof -
              have "(r0  γr)  𝖺[r0, p1, w]  ((inv r0s1  w)  (𝖺-1[s1, p0, w] 
                      𝖺[s1, p0, w])  (r0s1  w))  𝖺-1[r0, p1, w] =
                    r0  γr"
              proof -
                have "(r0  γr)  𝖺[r0, p1, w]  ((inv r0s1  w)  (𝖺-1[s1, p0, w] 
                       𝖺[s1, p0, w])  (r0s1  w))  𝖺-1[r0, p1, w] =
                      (r0  γr)  𝖺[r0, p1, w]  ((inv r0s1  w)  (r0s1  w))  𝖺-1[r0, p1, w]"
                  using r0s1.φ_uniqueness(2) comp_assoc_assoc' comp_cod_arr by simp
                  (* Used here that φ is a section. *)
                also have "... = (r0  γr)  𝖺[r0, p1, w]  𝖺-1[r0, p1, w]"
                  using r0s1.φ_uniqueness(2) comp_inv_arr' ρ.T0.antipar(1)
                        whisker_right [of w "inv r0s1" r0s1] comp_cod_arr
                  by simp
                also have "... = r0  γr"
                proof -
                  have "hseq r0 γr"
                    using βs βs_def by blast
                  thus ?thesis
                    using comp_assoc_assoc' comp_arr_dom
                      by (metis (no_types) γr ρ.ide_leg0 comp_assoc_assoc'(1) hcomp_simps(3)
                          hseq_char ide_char in_homE r0s1.ide_leg1 r0s1.p1_simps w wr)
                qed
                finally show ?thesis by blast
              qed
              thus ?thesis by simp
            qed
            also have "... = 𝖺[s1, p0, w']  (r0s1  w')  𝖺-1[r0, p1, w']  (r0  γr)"
              using comp_assoc by presburger
            also have "... = (s1  ?ws') 
                               𝖺[s1, p0, w']  (r0s1  w')  𝖺-1[r0, p1, w']  (r0  γr)"
            proof -
              have "(s1  ?ws')  𝖺[s1, p0, w'] = 𝖺[s1, p0, w']"
                using comp_cod_arr by simp
              thus ?thesis
                using comp_assoc by metis
            qed
            finally show ?thesis by simp
          qed
          also have "... = (s1  γs)  𝖺[s1, p0, w]  (r0s1  w)  𝖺-1[r0, p1, w]"
            using γs by simp
          finally show ?thesis by simp
        qed
        have 3: "∃!γ. «γ : w  w'»  γs = (p0  w')  (p0  γ)  p1  γ = γr"
          using w w' ws' wr γr γs eqt
                r0s1.φ_biuniversal_prop(2) [of ?ws' ?wr w w' γs "p0  w'" γr]
          by blast
        obtain γ where γ: "«γ : w  w'»  γs = (p0  w')  (p0  γ)  p1  γ = γr"
          using 3 by blast

        show "∃!γ. «γ : w  w'»  β = (r1  p1)  γ  θ = θ'  ((s0  p0)  γ)"
        proof -
          have "γ. «γ : w  w'»  β = (r1  p1)  γ  θ = θ'  ((s0  p0)  γ)"
          proof -
            have "θ = θ'  ((s0  p0)  γ)"
            proof -
              have "θ'  ((s0  p0)  γ) = (θs'  𝖺[s0, p0, w'])  ((s0  p0)  γ)"
                using θs'_def comp_arr_dom comp_assoc comp_assoc_assoc'(2) by auto
              also have "... = (θs'  (s0  p0  γ))  𝖺[s0, p0, w]"
                using assoc_naturality [of s0 p0 γ] comp_assoc
                by (metis γ γr σ.leg0_simps(4-5) r0s1.leg0_simps(4-5)
                    r0s1.leg1_simps(3) hseqE in_homE leg0_simps(2))
              also have "... = θs  𝖺[s0, p0, w]"
                by (metis γ γs arrI comp_ide_arr ws')
              also have "... = θ"
                using θs_def comp_assoc comp_arr_dom comp_assoc_assoc' by simp
              finally show ?thesis by simp
            qed
            moreover have "β = (r1  p1)  γ"
            proof -
              have "β = 𝖺-1[r1, p1, w']  βr  𝖺[r1, p1, w]"
              proof -
                have "𝖺-1[r1, p1, w']  βr  𝖺[r1, p1, w] =
                      (𝖺-1[r1, p1, w']  𝖺[r1, p1, w'])  β  𝖺-1[r1, p1, w]  𝖺[r1, p1, w]"
                  using βr_def comp_assoc by simp
                also have "... = β"
                  using comp_arr_dom comp_cod_arr
                  by (metis ρ.ide_leg1 r0s1.ide_leg1 comp_assoc_assoc'(2) hseqE ideD(1)
                      uwθw'θ'β.β_simps(1) uwθw'θ'β.β_simps(4-5) leg1_simps(2) w w' wr wr')
                finally show ?thesis by simp
              qed
              also have "... = 𝖺-1[r1, p1, w']  (r1  γr)  𝖺[r1, p1, w]"
                using γr by simp
              also have "... = 𝖺-1[r1, p1, w']  𝖺[r1, p1, w']  ((r1  p1)  γ)"
                using assoc_naturality [of r1 p1 γ]
                by (metis γ γr ρ.ide_leg1 r0s1.leg1_simps(5-6) hseqE
                    ide_char in_homE leg1_simps(2))
              also have "... = (𝖺-1[r1, p1, w']  𝖺[r1, p1, w'])  ((r1  p1)  γ)"
                using comp_assoc by presburger
              also have "... = (r1  p1)  γ"
                using comp_cod_arr
                by (metis ρ.ide_leg1 r0s1.ide_leg1 calculation comp_assoc_assoc'(2) comp_ide_arr
                    hseqE ideD(1) ide_cod local.uwθw'θ'β.β_simps(1) local.uwθw'θ'β.β_simps(5)
                    w' wr')
              finally show ?thesis by simp
            qed
            ultimately show "γ. «γ : w  w'»  β = (r1  p1)  γ  θ = θ'  ((s0  p0)  γ)"
               using γ by blast
          qed
          moreover have "γ'. «γ' : w  w'»  β = (r1  p1)  γ'  θ = θ'  ((s0  p0)  γ')
                                    γ' = γ"
          proof -
            fix γ'
            assume γ': "«γ' : w  w'»  β = (r1  p1)  γ'  θ = θ'  ((s0  p0)  γ')"
            show "γ' = γ"
            proof -
              let ?Pr = "λγ. «γ : ?wr  ?wr'»  θr = θr'  (r0  γ)  βr = r1  γ"
              let ?Ps = "λγ. «γ : ?ws  ?ws'»  θs = θs'  (s0  γ)  βs = s1  γ"
              let r' = "p1  γ'"
              let s' = "p0  γ'"
              let ?Pt = "λγ. «γ : w  w'»  γs = (p0  w')  (p0  γ)  p1  γ = γr"
              have "hseq p0 γ'"
              proof (intro hseqI)
                show "«γ' : src θ  src p0»"
                  using γ'
                  by (metis hseqE hseqI' in_hhom_def uwθw'θ'β.β_simps(1) src_hcomp
                      src_vcomp leg0_simps(2) leg1_simps(3)
                      uwθw'θ'β.uwθ.θ_simps(1) vseq_implies_hpar(1))
                show "«p0 : src p0  src s0»" by simp
              qed
              hence "hseq p1 γ'"
                using hseq_char by simp
              have "«r' : ?wr  ?wr'»"
                using γ' by auto
              moreover have "θr = θr'  (r0  r')"
              proof -
                text ‹
                  Note that @{term θr} is the composite of ``everything to the right''
                  of @{term "ρ  ?wr"}, and similarly for @{term θr'}.
                  We can factor @{term θr} as @{term "(s  θ)  X w"}, where @{term "X w"}
                  is a composite of @{term σ} and @{term φ}.  We can similarly factor @{term θr'}
                  as @{term "(s  θ')  X w'"}.
                  Then @{term "θr'  (r0  r') = (s  θ')  X w'  (r0  r')"},
                  which equals @{term "(s  θ')  (s  (s0  p0)  r')  X w = θr"}.
                ›
                let ?X = "λw. (s  𝖺-1[s0, p0, w])  𝖺[s, s0, p0  w]  (σ  p0  w) 
                                𝖺[s1, p0, w]  (r0s1  w)  𝖺-1[r0, p1, w]"
                have "θr'  (r0  r') = (s  θ')  ?X w'  (r0  r')"
                  using θr'_def comp_assoc by simp
                also have "... = (s  θ')  (s  (s0  p0)  γ')  ?X w"
                proof -
                  have "(s  θ')  ((s  𝖺-1[s0, p0, w'])  𝖺[s, s0, p0  w']  (σ  p0  w') 
                          𝖺[s1, p0, w']  (r0s1  w')  𝖺-1[r0, p1, w'])  (r0  p1  γ') =
                        (s  θ')  (s  𝖺-1[s0, p0, w'])  𝖺[s, s0, p0  w']  (σ  p0  w') 
                          𝖺[s1, p0, w']  (r0s1  w')  𝖺-1[r0, p1, w']  (r0  p1  γ')"
                    using comp_assoc by presburger
                  also have "... = (s  θ')  (s  𝖺-1[s0, p0, w'])  𝖺[s, s0, p0  w'] 
                                     (σ  p0  w')  𝖺[s1, p0, w']  ((r0s1  w') 
                                     ((r0  p1)  γ'))  𝖺-1[r0, p1, w]"
                    using assoc'_naturality [of r0 p1 γ'] comp_assoc
                    by (metis γ' «p1  γ' : p1  w  p1  w'» ρ.T0.antipar(1)
                        ρ.leg0_in_hom(2) r0s1.leg1_simps(4-6)
                        r0s1.base_simps(2) hcomp_in_vhomE in_homE trg_hcomp)
                  also have "... = (s  θ')  (s  𝖺-1[s0, p0, w'])  𝖺[s, s0, p0  w'] 
                                     (σ  p0  w')  (𝖺[s1, p0, w']  ((s1  p0)  γ')) 
                                     (r0s1  w)  𝖺-1[r0, p1, w]"
                  proof -
                    have "(r0s1  w')  ((r0  p1)  γ') = r0s1  γ'"
                      using γ' interchange [of r0s1 "r0  p1" w' γ'] comp_arr_dom comp_cod_arr
                      by auto
                    also have "... = ((s1  p0)  γ')  (r0s1  w)"
                      using γ' interchange hseq p0 γ' comp_arr_dom comp_cod_arr
                      by (metis comp_arr_ide r0s1.φ_simps(1,5) seqI'
                          uwθw'θ'β.uwθ.w_in_hom(2) w)
                    finally have "(r0s1  w')  ((r0  p1)  γ') = 
                                  ((s1  p0)  γ')  (r0s1  w)"
                      by simp
                    thus ?thesis
                      using comp_assoc by presburger
                  qed
                  also have "... = (s  θ')  (s  𝖺-1[s0, p0, w'])  𝖺[s, s0, p0  w'] 
                                     ((σ  p0  w')  (s1  p0  γ'))  𝖺[s1, p0, w] 
                                     (r0s1  w)  𝖺-1[r0, p1, w]"
                    using γ' assoc_naturality [of s1 p0 γ'] comp_assoc
                    by (metis σ.leg1_simps(2) σ.leg1_simps(3,5-6) r0s1.leg0_simps(4-5)
                        hcomp_in_vhomE hseqE in_homE uwθw'θ'β.β_simps(1)
                        leg0_in_hom(2) leg1_simps(3))
                  also have "... = (s  θ')  (s  𝖺-1[s0, p0, w'])  (𝖺[s, s0, p0  w'] 
                                   ((s  s0)  p0  γ'))  (σ  p0  w)  𝖺[s1, p0, w] 
                                   (r0s1  w)  𝖺-1[r0, p1, w]"
                  proof -
                    have "(σ  p0  w')  (s1  p0  γ') = σ  p0  γ'"
                      using γ' interchange [of σ s1 "p0  w'" "p0  γ'"]
                            whisker_left hseq p0 γ'comp_arr_dom comp_cod_arr
                      by (metis σ.tab_simps(1) σ.tab_simps(4) hcomp_simps(4) in_homE
                          r0s1.leg0_simps(5))
                    also have "... = ((s  s0)  p0  γ')  (σ  p0  w)"
                      using γ' interchange [of "s  s0" σ "p0  γ'" "p0  w"]
                            whisker_left comp_arr_dom comp_cod_arr hseq p0 γ'
                      by (metis σ.tab_simps(1) σ.tab_simps(5) hcomp_simps(3) in_homE
                          r0s1.leg0_simps(4))
                    finally have "(σ  p0  w')  (s1  p0  γ') =
                                  ((s  s0)  p0  γ')  (σ  p0  w)"
                      by simp
                    thus ?thesis
                      using comp_assoc by presburger
                  qed
                  also have "... = (s  θ')  (s  𝖺-1[s0, p0, w'])  ((s  s0  p0  γ') 
                                     𝖺[s, s0, p0  w])  (σ  p0  w)  𝖺[s1, p0, w] 
                                     (r0s1  w)  𝖺-1[r0, p1, w]"
                    using γ' assoc_naturality [of s s0 "p0  γ'"] hseq p0 γ' by force
                  also have "... = (s  θ')  ((s  𝖺-1[s0, p0, w'])  (s  s0  p0  γ')) 
                                     𝖺[s, s0, p0  w]  (σ  p0  w)  𝖺[s1, p0, w] 
                                     (r0s1  w)  𝖺-1[r0, p1, w]"
                    using comp_assoc by presburger
                  also have "... = (s  θ')  ((s  (s0  p0)  γ')  (s  𝖺-1[s0, p0, w])) 
                                     𝖺[s, s0, p0  w]  (σ  p0  w)  𝖺[s1, p0, w] 
                                     (r0s1  w)  𝖺-1[r0, p1, w]"
                  proof -
                    have "(s  𝖺-1[s0, p0, w'])  (s  s0  p0  γ') =
                          (s  (s0  p0)  γ')  (s  𝖺-1[s0, p0, w])"
                    proof -
                      have "(s  𝖺-1[s0, p0, w'])  (s  s0  p0  γ') =
                            s  𝖺-1[s0, p0, w']  (s0  p0  γ')"
                      proof -
                        have "seq 𝖺-1[s0, p0, w'] (s0  p0  γ')"
                        proof
                          (* It seems to be too time-consuming for auto to solve these. *)
                          show "«s0  p0  γ' : s0  p0  w  s0  p0  w'»"
                            using γ'
                            by (intro hcomp_in_vhom, auto)
                          show "«𝖺-1[s0, p0, w'] : s0  p0  w'  (s0  p0)  w'»"
                            by auto
                        qed
                        thus ?thesis
                          using w w' γ' whisker_left by simp
                      qed
                      also have "... = s  ((s0  p0)  γ')  𝖺-1[s0, p0, w]"
                        using γ' hseq p0 γ' assoc'_naturality [of s0 p0 γ'] by fastforce
                      also have "... = (s  (s0  p0)  γ')  (s  𝖺-1[s0, p0, w])"
                      proof -
                        have "seq ((s0  p0)  γ') 𝖺-1[s0, p0, w]"
                        proof
                          (* Same here. *)
                          show "«𝖺-1[s0, p0, w] : s0  p0  w  (s0  p0)  w»"
                            by auto
                          show "«(s0  p0)  γ' : (s0  p0)  w  (s0  p0)  w'»"
                            using γ' by (intro hcomp_in_vhom, auto)
                        qed
                        thus ?thesis
                          using w w' γ' whisker_left by simp
                      qed
                      finally show ?thesis by blast
                    qed
                    thus ?thesis by simp
                  qed
                  also have "... = (s  θ')  (s  (s0  p0)  γ')  ?X w"
                    using comp_assoc by presburger
                  finally show ?thesis by simp
                qed
                also have "... = θr"
                  using θr_def γ' uwθw'θ'β.uwθ.θ_simps(1) whisker_left σ.ide_base comp_assoc
                  by simp
                finally show ?thesis by simp
              qed
              moreover have "βr = r1  r'"
              proof -
                have "βr = 𝖺[r1, p1, w']  ((r1  p1)  γ')  𝖺-1[r1, p1, w]"
                  using βr_def γ' by simp
                also have "... = 𝖺[r1, p1, w']  𝖺-1[r1, p1, w']  (r1  p1  γ')"
                  using γ' assoc'_naturality
                  by (metis ρ.leg1_simps(5-6) r0s1.leg1_simps(5-6)
                      hcomp_in_vhomE hseqE in_homE uwθw'θ'β.β_simps(1) leg1_in_hom(2))
                also have "... = (𝖺[r1, p1, w']  𝖺-1[r1, p1, w'])  (r1  p1  γ')"
                  using comp_assoc by presburger
                also have "... = r1  p1  γ'"
                  using comp_cod_arr
                  by (metis (no_types, lifting) βr ρ.ide_leg1 r0s1.ide_leg1 arrI calculation
                      comp_assoc_assoc'(1) comp_ide_arr ide_hcomp hseq_char'
                      ideD(1) seq_if_composable hcomp_simps(2) leg1_simps(2) w' wr')
                finally show ?thesis by simp
              qed
              ultimately have Pr': "?Pr r'"
                by simp
              have eqr: "γr = r'"
                using 1 γr Pr' by blast
              have "«s' : ?ws  ?ws'»"
                using γ' by auto
              moreover have "θs = θs'  (s0  s')"
                using γ' hseq p0 γ' σ.leg0_simps(2,4-5)  σ.leg1_simps(3) θs'_def θs_def
                      assoc'_naturality hseqE in_homE comp_assoc r0s1.leg0_simps(4-5)
                      r0s1.p0_simps
                by metis
              moreover have "βs = s1  s'"
              proof -
                have "𝖺[s1, p0, w']  (r0s1  w')  𝖺-1[r0, p1, w']  (r0  γr) 
                        𝖺[r0, p1, w]  (inv r0s1  w)  𝖺-1[s1, p0, w] =
                      𝖺[s1, p0, w']  (r0s1  w')  (𝖺-1[r0, p1, w']  (r0  p1  γ')) 
                        𝖺[r0, p1, w]  (inv r0s1  w)  𝖺-1[s1, p0, w]"
                  using eqr comp_assoc by simp
                also have "... = 𝖺[s1, p0, w']  ((r0s1  w')  ((r0  p1)  γ')) 
                                   𝖺-1[r0, p1, w]  𝖺[r0, p1, w]  (inv r0s1  w) 
                                   𝖺-1[s1, p0, w]"
                proof -
                  have "𝖺-1[r0, p1, w']  (r0  p1  γ') = ((r0  p1)  γ')  𝖺-1[r0, p1, w]"
                    using γ' assoc'_naturality hseq p1 γ'
                    by (metis ρ.leg0_simps(2,4-5) ρ.leg1_simps(3)
                        r0s1.leg1_simps(5-6) hseqE in_homE leg1_simps(2))
                  thus ?thesis
                    using comp_assoc by presburger
                qed
                also have "... = (𝖺[s1, p0, w']  ((s1  p0)  γ'))  (r0s1  w) 
                                   𝖺-1[r0, p1, w]  𝖺[r0, p1, w]  (inv r0s1  w) 
                                   𝖺-1[s1, p0, w]"
                proof -
                  have "(r0s1  w')  ((r0  p1)  γ') = r0s1  γ'"
                    using γ' interchange [of r0s1 "r0  p1" w' γ']
                          comp_arr_dom comp_cod_arr
                    by auto
                  also have "... = ((s1  p0)  γ')  (r0s1  w)"
                    using γ' interchange hseq p0 γ' comp_arr_dom comp_cod_arr
                    by (metis in_homE r0s1.φ_simps(1,5))
                  finally have "(r0s1  w')  ((r0  p1)  γ') =
                                ((s1  p0)  γ')  (r0s1  w)"
                    by simp
                  thus ?thesis
                    using comp_assoc by presburger
                qed
                also have "... = (s1  s')  𝖺[s1, p0, w]  ((r0s1  w)  (𝖺-1[r0, p1, w] 
                                   𝖺[r0, p1, w])  (inv r0s1  w))  𝖺-1[s1, p0, w]"
                proof -
                  have "𝖺[s1, p0, w']  ((s1  p0)  γ') = (s1  s')  𝖺[s1, p0, w]"
                    using γ' assoc_naturality [of s1 p0 γ'] hseq p0 γ' by auto
                  thus ?thesis
                    using comp_assoc by presburger
                qed
                also have "... = s1  s'"
                proof -
                  have "𝖺-1[r0, p1, w]  𝖺[r0, p1, w] = cod (inv r0s1  w)"
                    using r0s1.φ_uniqueness(2) ρ.T0.antipar(1) comp_assoc_assoc'
                    by simp
                  text ‹Here the fact that φ› is a retraction is used.›
                  moreover have "(r0s1  w)  (inv r0s1  w) = cod 𝖺-1[s1, p0, w]"
                    using r0s1.φ_uniqueness(2) comp_arr_inv'
                          whisker_right [of w r0s1 "inv r0s1"]
                    by simp
                  moreover have "cod (inv r0s1  w)  (inv r0s1  w) = inv r0s1  w"
                    using βs_def βs
                    by (meson arrI comp_cod_arr seqE)
                  ultimately show ?thesis
                    using γ' hseq p0 γ' comp_arr_dom comp_cod_arr comp_assoc_assoc'
                          whisker_left [of s1 "p0  γ'" "p0  w"] whisker_left [of p0]
                    by (metis σ.ide_leg1 assoc'_simps(1) hseqE ideD(1) in_homE r0s1.ide_leg0
                        r0s1.p0_simps w ws)
                qed
                finally show ?thesis
                  using βs_def by simp
              qed
              ultimately have Ps': "?Ps s'"
                by simp
              have eqs: "γs = s'"
                using 2 γs Ps' by blast
              have "?Pt γ'"
                using γ' comp_cod_arr «p0  γ' : p0  w  p0  w'» eqr eqs by auto
              thus "γ' = γ"
                using 3 γ by blast
            qed
          qed
          ultimately show ?thesis by blast
        qed
      qed
    qed

  end

  sublocale composite_tabulation_in_maps 
              tabulation V H 𝖺 𝗂 src trg r  s tab s0  p0 r1  p1
    using composite_is_tabulation by simp

  sublocale composite_tabulation_in_maps 
              tabulation_in_maps V H 𝖺 𝗂 src trg r  s tab s0  p0 r1  p1
    using T0.is_map ρ.leg1_is_map ρ.T0.antipar(2) composable ρ.leg1_is_map ρ.T0.antipar
    apply unfold_locales
     apply simp
    apply (intro left_adjoints_compose)
    by auto

  subsection "The Classifying Category of Maps"

  text ‹
    \sloppypar
    We intend to show that if B› is a bicategory of spans, then B› is biequivalent to
    Span(Maps(B))›, for a specific category Maps(B)› derived from B›.
    The category Maps(B)› is constructed in this section as the ``classifying category'' of
    maps of B›, which has the same objects as B› and which has as 1-cells the isomorphism classes
    of maps of B›.  We show that, if B› is a bicategory of spans, then Maps(B)› has pullbacks.
  ›

  locale maps_category =
    B: bicategory_of_spans
  begin

    no_notation B.in_hhom  ("«_ : _  _»")
    no_notation B.in_hom  ("«_ : _ B _»")
    notation B.in_hhom  ("«_ : _ B _»")
    notation B.in_hom  ("«_ : _ B _»")
    notation B.isomorphic  (infix "B" 50)
    notation B.iso_class  ("_B")

    text ‹
      I attempted to modularize the construction here, by refactoring ``classifying category''
      out as a separate locale, but it ended up causing extra work because to apply it we
      first need to obtain the full sub-bicategory of 2-cells between maps, then construct its
      classifying category, and then we have to re-prove everything about it, to get rid of
      any mention of the sub-bicategory construction.  So the construction is being done
      here as a ``one-off'' special case construction, with the necessary properties proved
      just once.
     ›

     text ‹
       The ``hom-categories'' of Maps(C)› have as arrows the isomorphism classes of maps of B›.
     ›

     abbreviation Hom
     where "Hom a b  {F. f. «f : a B b»  B.is_left_adjoint f  F = fB}"

     lemma in_HomD:
     assumes "F  Hom a b"
     shows "F  {}"
     and "B.is_iso_class F"
     and "f  F  B.ide f"
     and "f  F  «f : a B b»"
     and "f  F  B.is_left_adjoint f"
     and "f  F  F = fB"
     proof -
       show "F  {}"
         using assms B.ide_in_iso_class B.left_adjoint_is_ide B.iso_class_is_nonempty by auto
       show 1: "B.is_iso_class F"
         using assms B.is_iso_classI B.left_adjoint_is_ide by fastforce
       show "f  F  B.ide f"
         using assms 1 B.iso_class_memb_is_ide by blast
       obtain f' where f': "«f' : a B b»  B.is_left_adjoint f'  F = f'B"
         using assms by auto
       show "f  F  «f : a B b»"
         using assms f' B.iso_class_def B.isomorphic_implies_hpar by auto
       show "f  F  B.is_left_adjoint f"
         using assms f' B.iso_class_def B.left_adjoint_preserved_by_iso [of f'] by auto
       show "f  F  F = fB"
         by (metis B.adjoint_pair_antipar(1) f' B.ide_in_iso_class B.is_iso_classI
             B.iso_class_elems_isomorphic B.iso_class_eqI)
     qed

     definition Comp
     where "Comp G F  {h. B.is_iso_class F  B.is_iso_class G 
                            (f g. f  F  g  G  g  f B h)}"

     lemma in_CompI [intro]:
     assumes "B.is_iso_class F" and "B.is_iso_class G"
     and "f  F" and "g  G" and "g  f B h"
     shows "h  Comp G F"
       unfolding Comp_def
       using assms by auto

     lemma in_CompE [elim]:
     assumes "h  Comp G F"
     and "f g.  B.is_iso_class F; B.is_iso_class G; f  F; g  G; g  f B h   T"
     shows T
       using assms Comp_def by auto

     lemma is_iso_class_Comp:
     assumes "Comp G F  {}"
     shows "B.is_iso_class (Comp G F)"
     proof -
       obtain h where h: "h  Comp G F"
         using assms by auto
       have ide_h: "B.ide h"
         using h Comp_def B.isomorphic_implies_hpar(2) by auto
       obtain f g where fg: "B.is_iso_class F  B.is_iso_class G  f  F  g  G  g  f B h"
         using h Comp_def by auto
       have "Comp G F = g  fB  B.ide (g  f)"
       proof (intro conjI)
         show "B.ide (g  f)"
           using fg B.iso_class_memb_is_ide B.isomorphic_implies_ide(1) by auto
         show "Comp G F = g  fB"
         proof
           show "g  fB  Comp G F"
             unfolding Comp_def B.iso_class_def
             using fg by auto
           show "Comp G F  g  fB"
           proof
             fix h'
             assume h': "h'  Comp G F"
             obtain f' g' where f'g': "f'  F  g'  G  g'  f' B h'"
               using h' Comp_def by auto
             have 1: "f' B f  g' B g"
               using f'g' fg B.iso_class_elems_isomorphic by auto
             moreover have "B.ide f  B.ide f'  B.ide g  B.ide g'"
               using 1 B.isomorphic_implies_hpar by auto
             ultimately have "g'  f' B g  f"
               using f'g' fg B.hcomp_isomorphic_ide B.hcomp_ide_isomorphic
                     B.isomorphic_transitive B.isomorphic_implies_hpar
               by (meson B.hseqE B.ideD(1))
             hence "h' B g  f"
               using f'g' B.isomorphic_symmetric B.isomorphic_transitive by blast
             thus "h'  B.iso_class (g  f)"
               using B.iso_class_def B.isomorphic_symmetric by simp
           qed
         qed
       qed
       thus ?thesis
         using assms B.is_iso_class_def B.ide_in_iso_class by auto
     qed

     lemma Comp_is_extensional:
     assumes "Comp G F  {}"
     shows "B.is_iso_class F" and "B.is_iso_class G"
     and "F  {}" and "G  {}"
       using assms Comp_def by auto

     lemma Comp_eqI [intro]:
     assumes "h  Comp G F" and "h'  Comp G' F'" and "h B h'"
     shows "Comp G F = Comp G' F'"
     proof -
       obtain f g where fg: "f  F  g  G  g  f B h"
         using assms comp_def by auto
       obtain f' g' where f'g': "f'  F'   g'  G'  g'  f' B h'"
         using assms by auto
       have "h  Comp G F  Comp G' F'"
         by (meson IntI assms in_CompE in_CompI B.isomorphic_symmetric B.isomorphic_transitive)
       hence "Comp G F  Comp G' F'  {}"
         by auto
       thus ?thesis
         using assms is_iso_class_Comp
         by (metis empty_iff B.iso_class_eq)
     qed

     lemma Comp_eq_iso_class_memb:
     assumes "h  Comp G F"
     shows "Comp G F = hB"
     proof
       show "Comp G F  hB"
       proof
         fix h'
         assume h': "h'  Comp G F"
         obtain f g where fg: "f  F  g  G  g  f B h"
           using assms by auto
         obtain f' g' where f'g': "f'  F  g'  G  g'  f' B h'"
           using h' by auto
         have "f B f'  g B g'"
           using assms fg f'g' in_HomD(6) B.iso_class_elems_isomorphic by auto
         moreover have "B.ide f  B.ide f'  B.ide g  B.ide g'"
           using assms fg f'g' in_HomD [of F] in_HomD [of G]
           by (meson calculation B.isomorphic_implies_ide(1) B.isomorphic_implies_ide(2))
         moreover have "src g = trg f  src g = trg f'  src g' = trg f  src g' = trg f'"
           using fg f'g'
           by (metis B.seq_if_composable calculation(1) B.ideD(1)
               B.isomorphic_implies_hpar(1) B.isomorphic_implies_hpar(3) B.not_arr_null)
         ultimately have "g  f B g'  f'"
           using fg f'g' B.hcomp_ide_isomorphic B.hcomp_isomorphic_ide B.isomorphic_transitive
           by metis
         thus "h'  hB"
           using fg f'g' B.isomorphic_symmetric B.isomorphic_transitive B.iso_class_def [of h]
           by blast
       qed
       show "hB  Comp G F"
       proof (unfold B.iso_class_def Comp_def)
         obtain f g where 1: "f  F  g  G  g  f B h"
           using assms in_HomD Comp_def
           by (meson in_CompE)
         show "{h'. B.isomorphic h h'} 
               {h. B.is_iso_class F  B.is_iso_class G  (f g. f  F  g  G  g  f B h)}"
           using assms 1 B.isomorphic_transitive by blast
       qed
     qed

     interpretation concrete_category Collect B.obj Hom B.iso_class λ_ _ _. Comp
     proof
       show "a. a  Collect B.obj  aB  Hom a a"
         by (metis (mono_tags, lifting) B.ide_in_hom(1) mem_Collect_eq B.objE
             B.obj_is_self_adjoint(1))
       show "a b c F G.  a  Collect B.obj; b  Collect B.obj; c  Collect B.obj;
                           F  Hom a b; G  Hom b c   Comp G F  Hom a c"
       proof -
         fix a b c F G
         assume a: "a  Collect B.obj" and b: "b  Collect B.obj" and c: "c  Collect B.obj"
         and F: "F  Hom a b" and G: "G  Hom b c"
         obtain f where f: "«f : a B b»  B.is_left_adjoint f  F = fB"
           using F by blast
         obtain g where g: "«g : b B c»  B.is_left_adjoint g  G = gB"
           using G by blast
         have "{h. B.is_iso_class F  B.is_iso_class G 
                   (f g. f  F  «f : a B b»  g  G  «g : b B c»  g  f B h)} =
               g  fB"
         proof
           show "{h. B.is_iso_class F  B.is_iso_class G 
                     (f g. f  F  «f : a B b»  g  G  «g : b B c»  g  f B h)}
                    g  fB"
           proof
             fix h
             assume "h  {h. B.is_iso_class F  B.is_iso_class G 
                             (f g. f  F  «f : a B b»  g  G  «g : b B c»  g  f B h)}"
             hence h: "B.is_iso_class F  B.is_iso_class G 
                       (f g. f  F  «f : a B b»  g  G  «g : b B c»  g  f B h)"
               by simp
             show "h  g  fB"
             proof -
               obtain f' g' where f'g': "g'  G  f'  F  g'  f' B h"
                 using h by auto
               obtain φ where φ: "«φ : f B f'»  B.iso φ"
                 using f f'g' F B.iso_class_def by auto
               obtain ψ where ψ: "«ψ : g B g'»  B.iso ψ"
                 using g f'g' G B.iso_class_def by auto
               have 1: "«ψ  φ : g  f B g'  f'»"
                 using f g φ ψ by auto
               moreover have "B.iso (ψ  φ)"
                 using f g φ ψ 1 B.iso_hcomp by auto
               ultimately show ?thesis
                  using f'g' B.iso_class_def B.isomorphic_def by auto
             qed
           qed
           show "g  fB  {h. B.is_iso_class F  B.is_iso_class G 
                               (f g. f  F  «f : a B b»  g  G  «g : b B c»  g  f B h)}"
             using f g B.iso_class_def B.isomorphic_reflexive B.left_adjoint_is_ide B.is_iso_classI
             by blast
         qed
         hence 1: "gf. gf  B.iso_class (g  f) 
                         B.is_iso_class F  B.is_iso_class G 
                         (f g. f  F  «f : a B b»  g  G  «g : b B c»  g  f B gf)"
           by blast
         show "Comp G F  Hom a c"
         proof -
           have gf: "B.is_left_adjoint (g  f)"
             by (meson f g B.hseqE B.hseqI B.left_adjoints_compose)
           obtain gf' where gf': "B.adjoint_pair (g  f) gf'"
             using gf by blast
           hence "g  fB = Comp G F"
             using 1 Comp_eq_iso_class_memb B.ide_in_iso_class B.left_adjoint_is_ide by blast
           thus ?thesis
             using f g gf' by blast
         qed
       qed
       show "a b F.  a  Collect B.obj; F  Hom a b   Comp F aB = F"
       proof -
         fix a b F
         assume a: "a  Collect B.obj"
         assume F: "F  Hom a b"
         obtain f where f: "«f : a B b»  B.is_left_adjoint f  F = fB"
           using F by auto
         have *: "f'. f'  F  «f' : a B b»  B.ide f'  f B f'"
           using f B.iso_class_def by force
         show "Comp F aB = F"
         proof
           show "Comp F aB  F"
           proof
             fix h
             assume "h  Comp F aB"
             hence h: "f' a'. f'  F  a'  aB  f'  a' B h"
               unfolding Comp_def by auto
             obtain f' a' where f'a': "f'  F  a'  aB  f'  a' B h"
               using h by auto
             have "B.isomorphic f h"
             proof -
               have "B.isomorphic f (f  a)"
                 using f B.iso_runit' [of f] B.isomorphic_def B.left_adjoint_is_ide
                 by blast
               also have "f  a B f'  a"
                 using f f'a' B.iso_class_def B.hcomp_isomorphic_ide
                 apply (elim conjE B.in_hhomE) by auto
               also have "f'  a B f'  a'"
                 using f'a' * [of f'] B.iso_class_def B.hcomp_ide_isomorphic by auto
               also have "f'  a' B h"
                 using f'a' by simp
               finally show ?thesis by blast
             qed
             thus "h  F"
               using f B.iso_class_def by simp
           qed
           show "F  Comp F aB"
           proof
             fix h
             assume h: "h  F"
             have "f  F"
               using f B.iso_class_def B.right_adjoint_determines_left_up_to_iso by auto
             moreover have "a  B.iso_class a"
               using a B.iso_class_def B.isomorphic_reflexive by auto
             moreover have "f  a B h"
             proof -
               have "f  a B f"
                 using f B.iso_runit [of f] B.isomorphic_def B.left_adjoint_is_ide by blast
               also have "f B h"
                 using h * by simp
               finally show ?thesis by blast
             qed
             ultimately show "h  Comp F aB"
               unfolding Comp_def
               using a f F B.is_iso_classI B.left_adjoint_is_ide by auto
           qed
         qed
       qed
       show "a b F.  b  Collect B.obj; F  Hom a b   Comp bB F = F"
       proof -
         fix a b F
         assume b: "b  Collect B.obj"
         assume F: "F  Hom a b"
         obtain f where f: "«f : a B b»  B.is_left_adjoint f  F = fB"
           using F by auto
         have *: "f'. f'  F  «f' : a B b»  B.ide f'  f B f'"
           using f B.iso_class_def by force
         show "Comp bB F = F"
         proof
           show "Comp bB F  F"
           proof
             fix h
             assume "h  Comp bB F"
             hence h: "b' f'. b'  bB  f'  F  b'  f' B h"
               unfolding Comp_def by auto
             obtain b' f' where b'f': "b'  bB  f'  F  b'  f' B h"
               using h by auto
             have "f B h"
             proof -
               have "f B b  f"
                 using f B.iso_lunit' [of f] B.isomorphic_def B.left_adjoint_is_ide
                 by blast
               also have "... B b  f'"
                 using f b'f' B.iso_class_def B.hcomp_ide_isomorphic
                 by (elim conjE B.in_hhomE, auto)
               also have "... B b'  f'"
                 using b'f' * [of f'] B.iso_class_def B.hcomp_isomorphic_ide by auto
               also have "... B h"
                 using b'f' by simp
               finally show ?thesis by blast
             qed
             thus "h  F"
               using f B.iso_class_def by simp
           qed
           show "F  Comp bB F"
           proof
             fix h
             assume h: "h  F"
             have "f  F"
               using f B.iso_class_def B.right_adjoint_determines_left_up_to_iso by auto
             moreover have "b  B.iso_class b"
               using b B.iso_class_def B.isomorphic_reflexive by auto
             moreover have "b  f B h"
             proof -
               have "b  f B f"
                 using f B.iso_lunit [of f] B.isomorphic_def B.left_adjoint_is_ide
                 by blast
               also have "f B h"
                 using h * by simp
               finally show ?thesis by blast
             qed
             ultimately show "h  Comp bB F"
               unfolding Comp_def
               using b f F B.is_iso_classI B.left_adjoint_is_ide by auto
           qed
         qed
       qed
       show "a b c d F G H.
                a  Collect B.obj; b  Collect B.obj; c  Collect B.obj; d  Collect B.obj;
                 F  Hom a b; G  Hom b c; H  Hom c d  
               Comp H (Comp G F) = Comp (Comp H G) F"
       proof -
         fix a b c d F G H
         assume F: "F  Hom a b" and G: "G  Hom b c" and H: "H  Hom c d"
         show "Comp H (Comp G F) = Comp (Comp H G) F"
         proof
           show "Comp H (Comp G F)  Comp (Comp H G) F"
           proof
             fix x
             assume x: "x  Comp H (Comp G F)"
             obtain f g h gf
             where 1: "B.is_iso_class F  B.is_iso_class G  B.is_iso_class H 
                       h  H  g  G  f  F  gf  Comp G F  g  f B gf  h  gf B x"
               using x unfolding Comp_def by blast
             have hgf: "B.ide f  B.ide g  B.ide h"
               using 1 F G H B.isomorphic_implies_hpar in_HomD B.left_adjoint_is_ide
               by (metis (mono_tags, lifting))
             have "h  g  f B x"
               by (metis "1" B.hcomp_ide_isomorphic B.hseqE B.ide_char'
                   B.isomorphic_implies_hpar(4) B.isomorphic_implies_ide(1)
                   B.isomorphic_transitive hgf)
             moreover have "(h  g)  f B h  g  f"
               using 1 hgf B.iso_assoc B.assoc_in_hom B.isomorphic_def
               by (metis B.hseq_char B.ideD(1-3) B.isomorphic_implies_hpar(1)
                   B.trg_hcomp calculation)
             moreover have hg: "«h  g : b B d»"
               using G H 1 in_HomD(4) by blast
             moreover have "h  g  Comp H G"
               unfolding Comp_def
               using 1 hgf F G H in_HomD [of F a b] in_HomD [of G b c] in_HomD [of H c d]
                     B.isomorphic_reflexive B.hcomp_ide_isomorphic B.hseqI'
               by (metis (no_types, lifting) B.hseqE B.hseqI mem_Collect_eq)
             ultimately show "x  Comp (Comp H G) F"
               by (metis "1" B.isomorphic_transitive emptyE in_CompI is_iso_class_Comp)
           qed
           show "Comp (Comp H G) F  Comp H (Comp G F)"
           proof
             fix x
             assume x: "x  Comp (Comp H G) F"
             obtain f g h hg
             where 1: "B.is_iso_class F  B.is_iso_class G  B.is_iso_class H 
                       h  H  g  G  f  F  hg  Comp H G  h  g B hg  hg  f B x"
               using x unfolding Comp_def by blast
             have hgf: "B.ide f  B.ide g  B.ide h  src h = trg g  src g = trg f"
               using 1 F G H in_HomD B.left_adjoint_is_ide
               by (metis (no_types, lifting) B.hseq_char' B.ideD(1) B.ide_dom
                   B.in_homE B.isomorphic_def B.isomorphic_symmetric B.seqI'
                   B.seq_if_composable B.src_dom B.src_hcomp B.vseq_implies_hpar(1))
             have 2: "(h  g)  f B x"
               by (meson "1" B.hcomp_isomorphic_ide B.hseqE B.ideD(1) B.isomorphic_implies_ide(1)
                   B.isomorphic_symmetric B.isomorphic_transitive hgf)
             moreover have "(h  g)  f B h  g  f"
               using hgf B.iso_assoc B.assoc_in_hom B.isomorphic_def by auto
             moreover have "g  f  Comp G F"
               using 1 F G hgf B.isomorphic_reflexive by blast
             ultimately show "x  Comp H (Comp G F)"
               using 1 hgf is_iso_class_Comp [of G F] B.isomorphic_reflexive [of "g  f"]
               apply (intro in_CompI)
                       apply auto[6]
                 apply simp
                apply auto[1]
               by (meson B.isomorphic_symmetric B.isomorphic_transitive)
           qed
         qed
       qed
     qed

     lemma is_concrete_category:
     shows "concrete_category (Collect B.obj) Hom B.iso_class (λ_ _ _. Comp)"
       ..

     sublocale concrete_category Collect B.obj Hom B.iso_class λ_ _ _. Comp
       using is_concrete_category by simp

     abbreviation comp  (infixr "" 55)
     where "comp  COMP"
     notation in_hom  ("«_ : _  _»")
     no_notation B.in_hom  ("«_ : _ B _»")

     lemma Map_memb_in_hhom:
     assumes "«F : A  B»" and "f  Map F"
     shows "«f : Dom A B Dom B»"
     proof -
       have "«f : Dom F B Cod F»"
         using assms arr_char [of F] in_HomD [of "Map F" "Dom F" "Cod F"] by blast
       moreover have "Dom F = Dom A"
         using assms by auto
       moreover have "Cod F = Dom B"
         using assms by auto
       ultimately show ?thesis by simp
     qed

     lemma MkArr_in_hom':
     assumes "B.is_left_adjoint f" and "«f : a B b»"
     shows "«MkArr a b fB : MkIde a  MkIde b»"
       using assms MkArr_in_hom by blast

     text ‹
       The isomorphisms in Maps(B)› are the isomorphism classes of equivalence maps in B›.
     ›

     lemma iso_char:
     shows "iso F  arr F  (f. f  Map F  B.equivalence_map f)"
     proof
       assume F: "iso F"
       have "f. f  Map F  B.equivalence_map f"
       proof -
         fix f
         assume f: "f  Map F"
         obtain G where G: "inverse_arrows F G"
           using F by auto
         obtain g where g: "g  Map G"
           using G arr_char [of G] in_HomD(1) by blast
         have f: "f  Map F  «f : Dom F B Cod F»  B.ide f  B.is_left_adjoint f"
           by (metis (mono_tags, lifting) F iso_is_arr is_concrete_category
               concrete_category.arr_char f in_HomD(2,4-5) B.iso_class_memb_is_ide)
         have g: "g  Map G  «g : Cod F B Dom F»  B.ide g  B.is_left_adjoint g"
           by (metis (no_types, lifting) F G Cod_cod Cod_dom arr_inv cod_inv dom_inv
               inverse_unique iso_is_arr is_concrete_category concrete_category.Map_in_Hom
               g in_HomD(2,4-5) B.iso_class_memb_is_ide)
         have "src (g  f) B g  f"
         proof -
           have "g  f  Map (G  F)"
           proof -
             have 1: "f  Map F  g  Map G  g  f B g  f"
               using F G f g B.isomorphic_reflexive by force
             have 2: "Dom G = Cod F  Cod G = Dom F"
               using F G arr_char
               by (metis (no_types, lifting) Cod.simps(1) Cod_dom arr_inv
                   cod_char comp_inv_arr dom_inv inverse_unique
                   iso_is_arr is_concrete_category concrete_category.Cod_comp)
             have "g  f  Comp (Map G) (Map F)"
               using 1 F iso_is_arr Map_in_Hom [of F] in_HomD(2)
               apply (intro in_CompI, auto)
             proof -
               show "B.is_iso_class (Map G)"
                 using G iso_is_arr Map_in_Hom [of G] in_HomD(2) [of "Map G"] by blast
             qed
             thus ?thesis
               using F G f g comp_char [of G F] by auto
           qed
           moreover have "Dom F  Map (G  F)"
             by (metis (no_types, lifting) F G Map_dom comp_inv_arr iso_is_arr
                 g B.ide_in_iso_class B.in_hhomE B.objE)
           moreover have "Map (G  F) = Dom FB"
             by (simp add: F G comp_inv_arr iso_is_arr)
           moreover have "Dom F = src (g  f)"
             using f g by fastforce
           ultimately show ?thesis
             using f g B.iso_class_elems_isomorphic B.is_iso_classI
             by (metis B.hseqI B.ide_src)
         qed
         moreover have "f  g B trg f"
         proof -
           have "f  g  Map (F  G)"
           proof -
             have 1: "f  Map F  g  Map G  f  g B f  g"
               using F G f g B.isomorphic_reflexive by force
             have 2: "Dom G = Cod F  Cod G = Dom F"
               using F G arr_char
               by (metis (no_types, lifting) Cod.simps(1) Cod_dom arr_inv
                   cod_char comp_inv_arr dom_inv inverse_unique
                   iso_is_arr is_concrete_category concrete_category.Cod_comp)
             hence "f  g  Comp (Map F) (Map G)"
               using 1 F iso_is_arr Map_in_Hom [of F] in_HomD(2)
               apply (intro in_CompI, auto)
             proof -
               show "B.is_iso_class (Map G)"
                 using G iso_is_arr Map_in_Hom [of G] in_HomD(2) [of "Map G"] by blast
             qed
             thus ?thesis
               using F G f g comp_char [of F G] by auto
           qed
           moreover have "Cod F  Map (F  G)"
             by (metis (no_types, lifting) F G Map_cod comp_arr_inv dom_inv
                 inverse_unique iso_is_arr g B.ide_in_iso_class B.in_hhomE B.src.preserves_ide)
           moreover have "Map (F  G) = Cod FB"
             by (metis (no_types, lifting) F G Map_cod comp_arr_inv dom_inv
                 inverse_unique iso_is_arr)
           moreover have "Cod F = trg (f  g)"
             using f g by fastforce
           ultimately show ?thesis
             using B.iso_class_elems_isomorphic
             by (metis f g B.is_iso_classI B.in_hhomE B.src.preserves_ide)
         qed
         ultimately show "B.equivalence_map f"
           using f g B.equivalence_mapI B.quasi_inversesI [of f g] by fastforce
       qed
       thus "arr F  (f. f  Map F  B.equivalence_map f)"
         using F by blast
       next
       assume F: "arr F  (f. f  Map F  B.equivalence_map f)"
       show "iso F"
       proof -
         obtain f where f: "f  Map F"
           using F arr_char in_HomD(1) by blast
         have f_in_hhom: "«f : Dom F B Cod F»"
           using f F arr_char in_HomD(4) [of "Map F" "Dom F" "Cod F" f] by simp
         have "Map F = B.iso_class f"
           using f F arr_char in_HomD(6) [of "Map F" "Dom F" "Cod F" f] by simp
         obtain g η ε' where ε': "equivalence_in_bicategory V H 𝖺 𝗂 src trg f g η ε'"
           using f F B.equivalence_map_def by auto
         interpret ε': equivalence_in_bicategory V H 𝖺 𝗂 src trg f g η ε'
           using ε' by auto
         obtain ε where ε: "adjoint_equivalence_in_bicategory V H 𝖺 𝗂 src trg f g η ε"
           using f F ε'.ide_right ε'.antipar ε'.unit_in_hom ε'.unit_is_iso B.equivalence_map_def
                 B.equivalence_refines_to_adjoint_equivalence [of f g η]
           by auto
         interpret ε: adjoint_equivalence_in_bicategory V H 𝖺 𝗂 src trg f g η ε
           using ε by simp
         have g_in_hhom: "«g : Cod F B Dom F»"
           using ε.ide_right ε.antipar f_in_hhom by auto
         have fg: "B.quasi_inverses f g"
           using B.quasi_inverses_def ε.equivalence_in_bicategory_axioms by auto
         have g: "«g : Cod F B Dom F»  B.is_left_adjoint g  gB = gB"
           using ε'.dual_equivalence B.equivalence_is_left_adjoint B.equivalence_map_def
                 g_in_hhom
           by blast
         have F_as_MkArr: "F = MkArr (Dom F) (Cod F) fB"
           using F MkArr_Map Map F = B.iso_class f by fastforce
         have F_in_hom: "in_hom F (MkIde (Dom F)) (MkIde (Cod F))"
           using F arr_char dom_char cod_char
           by (intro in_homI, auto)
         let ?G = "MkArr (Cod F) (Dom F) gB"
         have "arr ?G"
           using MkArr_in_hom' g by blast
         hence G_in_hom: "«?G : MkIde (Cod F)  MkIde (Dom F)»"
           using arr_char MkArr_in_hom by simp
         have "inverse_arrows F ?G"
         proof
           show "ide (?G  F)"
           proof -
             have "?G  F = dom F"
             proof (intro arr_eqI)
               show 1: "seq ?G F"
                 using F_in_hom G_in_hom by blast
               show "arr (dom F)"
                 using F_in_hom ide_dom by fastforce
               show "Dom (?G  F) = Dom (dom F)"
                 using F_in_hom G_in_hom 1 comp_char by auto
               show "Cod (?G  F) = Cod (dom F)"
                 using F_in_hom G_in_hom 1 comp_char by auto
               show "Map (?G  F) = Map (dom F)"
               proof -
                 have "Map (?G  F) = Comp gB fB"
                   using 1 comp_char [of ?G F] Map F = B.iso_class f by simp
                 also have "... = g  fB"
                 proof -
                   have "g  f  Comp gB fB"
                     by (metis ε.ide_left ε.ide_right ε.unit_in_vhom ε.unit_simps(3) B.arrI
                         B.ide_cod B.ide_in_iso_class in_CompI B.is_iso_classI
                         B.isomorphic_reflexive)
                   thus ?thesis
                     using Comp_eq_iso_class_memb F_in_hom G_in_hom arr_char arr_char
                           Map F = B.iso_class f
                     by auto
                 qed
                 also have "... = src fB"
                   using ε.unit_in_hom ε.unit_is_iso B.isomorphic_def B.iso_class_def
                         B.isomorphic_symmetric
                   by (intro B.iso_class_eqI, blast)
                 also have "... = Dom FB"
                   using ε.ide_left f_in_hhom B.ide_in_iso_class B.in_hhomE B.src.preserves_ide
                         B.isomorphic_reflexive
                   by (intro B.iso_class_eqI, fastforce)
                 also have "... = Map (dom F)"
                   using F_in_hom dom_char by auto
                 finally show ?thesis by blast
               qed
             qed
             thus ?thesis
               using F by simp
           qed
           show "ide (F  ?G)"
           proof -
             have "F  ?G = cod F"
             proof (intro arr_eqI)
               show 1: "seq F ?G"
                 using F_in_hom G_in_hom by blast
               show "arr (cod F)"
                 using F_in_hom ide_cod by fastforce
               show "Dom (F  ?G) = Dom (cod F)"
                 using F_in_hom G_in_hom 1 comp_char by auto
               show "Cod (F  ?G) = Cod (cod F)"
                 using F_in_hom G_in_hom 1 comp_char by auto
               show "Map (F  ?G) = Map (cod F)"
               proof -
                 have "Map (F  ?G) = Comp fB gB"
                   using 1 comp_char [of F ?G] Map F = fB by simp
                 also have "... = f  gB"
                 proof -
                   have "f  g  Comp fB gB"
                     by (metis ε.antipar(1) ε.ide_left ε.ide_right B.ide_hcomp
                         B.ide_in_iso_class in_CompI B.is_iso_classI B.isomorphic_reflexive)
                   thus ?thesis
                     using Comp_eq_iso_class_memb F_in_hom G_in_hom arr_char arr_char
                           Map F = fB
                     by auto
                 qed
                 also have "... = trg fB"
                 proof -
                   have "trg f  f  gB"
                     using ε.counit_in_hom ε.counit_is_iso B.isomorphic_def B.iso_class_def
                           B.isomorphic_symmetric
                     by blast
                   thus ?thesis
                     using B.iso_class_eqI
                     by (metis ε.antipar(1) ε.ide_left ε.ide_right B.ide_hcomp
                         B.ide_in_iso_class B.is_iso_classI B.iso_class_elems_isomorphic)
                 qed
                 also have "... = Cod FB"
                   using f_in_hhom by auto
                 also have "... = Map (cod F)"
                   using F_in_hom dom_char by auto
                 finally show ?thesis by blast
               qed
             qed
             thus ?thesis
               using F by simp
           qed
         qed
         thus ?thesis by auto
       qed
     qed

     lemma iso_char':
     shows "iso F  arr F  (f. f  Map F  B.equivalence_map f)"
     proof -
       have "arr F  (f. f  Map F  B.equivalence_map f) 
                       (f. f  Map F  B.equivalence_map f)"
       proof
         assume F: "arr F"
         show "(f. f  Map F  B.equivalence_map f) 
               (f. f  Map F  B.equivalence_map f)"
           using F arr_char in_HomD(1) by blast
         show "(f. f  Map F  B.equivalence_map f) 
               (f. f  Map F  B.equivalence_map f)"
           by (metis (no_types, lifting) F is_concrete_category concrete_category.arr_char
               B.equivalence_map_preserved_by_iso in_HomD(2) B.iso_class_elems_isomorphic)
       qed
       thus ?thesis
         using iso_char by blast
     qed

     text ‹
       The following mapping takes a map in B› to the corresponding arrow of Maps(B)›.
     ›

     abbreviation CLS  ("⟦⟦_⟧⟧")
     where "⟦⟦f⟧⟧  MkArr (src f) (trg f) fB"

     lemma ide_CLS_obj:
     assumes "B.obj a"
     shows "ide ⟦⟦a⟧⟧"
       by (simp add: assms B.obj_simps)

     lemma CLS_in_hom:
     assumes "B.is_left_adjoint f"
     shows "«⟦⟦f⟧⟧ : ⟦⟦src f⟧⟧  ⟦⟦trg f⟧⟧»"
       using assms B.left_adjoint_is_ide MkArr_in_hom MkArr_in_hom' by simp

     lemma CLS_eqI:
     assumes "B.ide f"
     shows "⟦⟦f⟧⟧ = ⟦⟦g⟧⟧  f B g"
       by (metis arr.inject assms B.ide_in_iso_class B.iso_class_def B.iso_class_eqI
           B.isomorphic_implies_hpar(3) B.isomorphic_implies_hpar(4) B.isomorphic_symmetric
           mem_Collect_eq)

     lemma CLS_hcomp:
     assumes "B.ide f" and "B.ide g" and "src f = trg g"
     shows "⟦⟦f  g⟧⟧ = MkArr (src g) (trg f) (Comp fB gB)"
     proof -
       have "⟦⟦f  g⟧⟧ = MkArr (src g) (trg f) f  gB"
         using assms B.left_adjoint_is_ide by simp
       moreover have "f  gB = Comp fB gB"
       proof
         show "f  gB  Comp fB gB"
           by (metis assms(1-2) B.ide_in_iso_class in_CompI B.is_iso_classI
               B.iso_class_def mem_Collect_eq subsetI)
         show "Comp fB gB  f  gB"
           by (metis Comp_eq_iso_class_memb f  gB  Comp fB gB
               assms(1-3) B.ide_hcomp B.ide_in_iso_class subset_iff)
       qed
       ultimately show ?thesis by simp
     qed

     lemma comp_CLS:
     assumes "B.is_left_adjoint f" and "B.is_left_adjoint g" and "f  g B h"
     shows "⟦⟦f⟧⟧  ⟦⟦g⟧⟧ = ⟦⟦h⟧⟧"
     proof -
       have hseq_fg: "B.hseq f g"
         using assms B.isomorphic_implies_hpar(1) by simp
       have "seq ⟦⟦f⟧⟧ ⟦⟦g⟧⟧"
         using assms hseq_fg CLS_in_hom [of f] CLS_in_hom [of g]  
         apply (elim B.hseqE) by auto
       hence "⟦⟦f⟧⟧  ⟦⟦g⟧⟧ = MkArr (src g) (trg f) (Comp fB gB)"
         using comp_char [of "CLS f" "CLS g"] by simp
       also have "... = ⟦⟦f  g⟧⟧"
         using assms hseq_fg CLS_hcomp
         apply (elim B.hseqE)
         using B.adjoint_pair_antipar(1) by auto
       also have "... = ⟦⟦h⟧⟧"
         using assms B.isomorphic_symmetric
         by (simp add: assms(3) B.iso_class_eqI B.isomorphic_implies_hpar(3)
             B.isomorphic_implies_hpar(4))
       finally show ?thesis by simp
     qed

     text ‹
       The following mapping that takes an arrow of Maps(B)› and chooses some
       representative map in B›.
     ›

     definition REP
     where "REP F  if arr F then SOME f. f  Map F else B.null"

     lemma REP_in_Map:
     assumes "arr A"
     shows "REP A  Map A"
     proof -
       have "Map A  {}"
         using assms arr_char in_HomD(1) by blast
       thus ?thesis
         using assms REP_def someI_ex [of "λf. f  Map A"] by auto
     qed

     lemma REP_in_hhom [intro]:
     assumes "in_hom F A B"
     shows "«REP F : src (REP A) B src (REP B)»"
     and "B.is_left_adjoint (REP F)"
     proof -
       have "Map F  {}"
         using assms in_hom_char arr_char null_char in_HomD(1) [of "Map F" "Dom F" "Cod F"]
         by blast
       hence 1: "REP F  Map F"
         using assms REP_def someI_ex [of "λf. f  Map F"] by auto
       hence 2: "B.arr (REP F)"
         using assms 1 in_hom_char [of F A B] B.iso_class_def Map_memb_in_hhom B.in_hhom_def
         by blast
       hence "«REP F : Dom A B Dom B»"
         using assms 1 in_hom_char [of F A B] Map_memb_in_hhom by auto
       thus "«REP F : src (REP A) B src (REP B)»"
         using assms
         by (metis (no_types, lifting) Map_memb_in_hhom ideD(1)
             in_homI in_hom_char REP_in_Map B.in_hhom_def)
       have "REP F  REP FB"
         using assms 1 2 arr_char [of F] in_HomD(6) B.ide_in_iso_class B.iso_class_memb_is_ide
               in_hom_char
         by blast
       thus "B.is_left_adjoint (REP F)"
         using assms 1 2 arr_char in_HomD(5) [of "Map F" "Dom F" "Cod F" "REP F"]
         by auto
     qed

     lemma ide_REP:
     assumes "arr A"
     shows "B.ide (REP A)"
       using assms REP_in_hhom(2) B.adjoint_pair_antipar(1) by blast

     lemma REP_simps [simp]:
     assumes "arr A"
     shows "B.ide (REP A)"
     and "src (REP A) = Dom A" and "trg (REP A) = Cod A"
     and "B.dom (REP A) = REP A" and "B.cod (REP A) = REP A"
     proof -
       show "B.ide (REP A)"
         using assms ide_REP by blast
       show "src (REP A) = Dom A"
         using assms REP_in_hhom
         by (metis (no_types, lifting) Map_memb_in_hhom Dom_dom in_homI
             REP_in_Map B.in_hhom_def)
       show "trg (REP A) = Cod A"
         using assms REP_in_hhom
         by (metis (no_types, lifting) Map_memb_in_hhom Dom_cod in_homI
             REP_in_Map B.in_hhom_def)
       show "B.dom (REP A) = REP A"
         using assms in_homI REP_in_hhom(2) B.adjoint_pair_antipar(1) B.ideD(2)
         by blast
       show "B.cod (REP A) = REP A"
         using assms in_homI REP_in_hhom(2) B.adjoint_pair_antipar(1) B.ideD(3)
         by blast
     qed

     lemma isomorphic_REP_src:
     assumes "ide A"
     shows "REP A B src (REP A)"
       using assms
       by (metis (no_types, lifting) ideD(1) ide_charCC REP_in_Map ide_REP
           REP_simps(2) B.is_iso_classI B.ide_in_iso_class B.iso_class_elems_isomorphic
           B.src.preserves_ide)

     lemma isomorphic_REP_trg:
     assumes "ide A"
     shows "REP A B trg (REP A)"
       using assms ide_charCC isomorphic_REP_src by auto

     lemma CLS_REP:
     assumes "arr F"
     shows "⟦⟦REP F⟧⟧ = F"
       by (metis (mono_tags, lifting) MkArr_Map
           is_concrete_category REP_in_Map REP_simps(2) REP_simps(3) assms
           concrete_category.Map_in_Hom in_HomD(6))

     lemma REP_CLS:
     assumes "B.is_left_adjoint f"
     shows "REP ⟦⟦f⟧⟧ B f"
       by (metis (mono_tags, lifting) CLS_in_hom Map.simps(1) in_homE REP_in_Map
           assms B.iso_class_def B.isomorphic_symmetric mem_Collect_eq)

     lemma isomorphic_hcomp_REP:
     assumes "seq F G"
     shows "REP F  REP G B REP (F  G)"
     proof -
       have 1: "Dom F = Cod G"
         using assms seq_char by simp
       have 2: "Map F = B.iso_class (REP F)"
         using assms seq_char arr_char REP_in_Map in_HomD(6) by meson
       have 3: "Map G = B.iso_class (REP G)"
         using assms seq_char arr_char REP_in_Map
               in_HomD(6) [of "Map G" "Dom G" "Cod G" "REP G"]
         by auto
       have "Map (F  G) = Comp REP FB REP GB"
         using assms seq_char null_char
         by (metis (no_types, lifting) CLS_REP Map.simps(1) Map_comp)
       have "Comp REP FB REP GB = REP F  REP GB"
       proof -
         have "REP F  REP G  Comp REP FB REP GB"
         proof -
           have "REP F  Map F  REP G  Map G"
             using assms seq_char REP_in_Map by auto
           moreover have "REP F  REP G B REP F  REP G"
             using assms seq_char 2 B.isomorphic_reflexive by auto                   
           ultimately show ?thesis
             using assms 1 2 3 B.iso_class_def B.is_iso_class_def
             by (intro in_CompI, auto)
         qed
         moreover have "REP FB  Hom (Cod G) (Cod F)"
           using assms 1 2 arr_char [of F] by auto
         moreover have "REP GB  Hom (Dom G) (Cod G)"
           using assms 1 3 arr_char [of G] by auto
         ultimately show ?thesis
           using Comp_eq_iso_class_memb assms arr_char arr_char REP_in_Map
           by (simp add: Comp_def)
       qed
       moreover have "REP (F  G)  Comp REP FB REP GB"
       proof -
         have "Map (F  G) = Comp (Map F) (Map G)"
           using assms 1 comp_char [of F G] by simp
         thus ?thesis
           using assms 1 2 3 REP_in_Map [of "F  G"] by simp
       qed
       ultimately show ?thesis
         using B.iso_class_def by simp
     qed

     text ‹
       We now show that Maps(B)› has pullbacks.  For this we need to exhibit
       functions PRJ0 and PRJ1 that produce the legs of the pullback of a cospan (H, K)›
       and verify that they have the required universal property.  These are obtained by
       choosing representatives h› and k› of H› and K›, respectively, and then taking
       PRJ0 = CLS (tab0 (k* ⋆ h))› and PRJ1 = CLS (tab1 (k* ⋆ h))›.  That these constitute a
       pullback in Maps(B)› follows from the fact that tab0 (k* ⋆ h)› and tab1 (k* ⋆ h)›
       form a pseudo-pullback of (h, k)› in the underlying bicategory.

       For our purposes here, it is not sufficient simply to show that Maps(B)› has pullbacks
       and then to treat it as an abstract ``category with pullbacks'' where the pullbacks
       are chosen arbitrarily.  Instead, we have to retain the connection between a pullback
       in Maps and the tabulation of k* ⋆ h› that is ultimately used to obtain it.  If we don't
       do this, then it becomes problematic to define the compositor of a pseudofunctor from
       the underlying bicategory B› to Span(Maps(B))›, because the components will go from the
       apex of a composite span (obtained by pullback) to the apex of a tabulation (chosen
       arbitrarily using BS2›) and these need not be in agreement with each other.
     ›

     definition PRJ0
     where "PRJ0  λK H. if cospan K H then ⟦⟦B.tab0 ((REP K)*  (REP H))⟧⟧ else null"
     definition PRJ1
     where "PRJ1  λK H. if cospan K H then ⟦⟦B.tab1 ((REP K)*  (REP H))⟧⟧ else null"

     interpretation elementary_category_with_pullbacks comp PRJ0 PRJ1
     proof
       show "H K. ¬ cospan K H  PRJ0 K H = null"
         unfolding PRJ0_def by auto
       show "H K. ¬ cospan K H  PRJ1 K H = null"
         unfolding PRJ1_def by auto
       show "H K. cospan K H  commutative_square K H (PRJ1 K H) (PRJ0 K H)"
       proof -
         fix H K
         assume HK: "cospan K H"
         define h where "h = REP H"
         define k where "k = REP K"
         have h: "h  Map H"
           using h_def HK REP_in_Map by blast
         have k: "k  Map K"
           using k_def HK REP_in_Map by blast
         have 1: "B.is_left_adjoint h  B.is_left_adjoint k  B.ide h  B.ide k  trg h = trg k"
           using h k h_def k_def HK arr_char cod_char B.in_hhom_def B.left_adjoint_is_ide
                 in_HomD(5) [of "Map H" "Dom H" "Cod H" h]
                 in_HomD(5) [of "Map K" "Dom K" "Cod K" k]
           apply auto
           by (metis (no_types, lifting) HK Dom_cod)
         interpret h: map_in_bicategory (⋅) (⋆) 𝖺 𝗂 src trg h
           using 1 by unfold_locales auto
         interpret k: map_in_bicategory (⋅) (⋆) 𝖺 𝗂 src trg k
           using 1 by unfold_locales auto
         interpret hk: cospan_of_maps_in_bicategory_of_spans (⋅) (⋆) 𝖺 𝗂 src trg h k
           using 1 by unfold_locales auto
         let ?f = "B.tab0 (k*  h)"
         let ?g = "B.tab1 (k*  h)"
         have span: "span ⟦⟦?f⟧⟧ ⟦⟦?g⟧⟧"
           using dom_char CLS_in_hom [of ?f] CLS_in_hom [of ?g] by auto
         have seq: "seq H ⟦⟦?f⟧⟧"
           using HK seq_char hk.leg0_is_map CLS_in_hom h_def hk.p0_simps hk.satisfies_T0
           by fastforce
         have seq': "seq K ⟦⟦?g⟧⟧"
           using HK k arr_char dom_char cod_char in_HomD(5) hk.leg1_is_map CLS_in_hom
           by (metis (no_types, lifting) Cod.simps(1) seq_char REP_simps(2)
               hk.p1_simps k_def span)
         show "commutative_square K H (PRJ1 K H) (PRJ0 K H)"
         proof
           show "cospan K H"
             using HK by simp
           show "dom K = cod (PRJ1 K H)"
             using seq' PRJ1_def HK h_def k_def by auto
           show "span (PRJ1 K H) (PRJ0 K H)"
             unfolding PRJ0_def PRJ1_def using HK span h_def k_def by simp
           show "K  PRJ1 K H = H  PRJ0 K H"
           proof -
             have iso: "h  ?f B k  ?g"
               using hk.φ_uniqueness B.isomorphic_symmetric B.isomorphic_def by blast
             have *: "Comp (Map H) ?fB = Comp (Map K) ?gB"
             proof (intro Comp_eqI)
               show "h  ?f  Comp (Map H) B.tab0 (k*  h)B"
               proof (unfold Comp_def)
                 have "B.is_iso_class ?fB"
                   by (simp add: B.is_iso_classI)
                 moreover have "B.is_iso_class (Map H)"
                   using CLS_REP HK Map.simps(1) B.is_iso_classI h.ide_left h_def
                   by (metis (no_types, lifting))
                 moreover have "?f  B.tab0 (k*  h)B"
                   by (simp add: B.ide_in_iso_class(1))
                 moreover have "«?f : src (B.tab0 (k*  h)) B Dom H»"
                   using seq seq_char by simp
                 moreover have "h  Map H"
                   by fact
                 moreover have "«h : Dom H B Cod H»"
                   by (simp add: HK h_def)
                 moreover have "h  ?f B h  ?f"
                   using B.isomorphic_reflexive by auto
                 ultimately show "h  B.tab0 (k*  h)
                                     {h'. B.is_iso_class B.tab0 (k*  h)B 
                                           B.is_iso_class (Map H) 
                                           (f g. f  B.tab0 (k*  h)B 
                                                  g  Map H  g  f B h')}"
                   by auto
               qed
               show "k  ?g  Comp (Map K) B.tab1 (k*  h)B"
               proof (unfold Comp_def)
                 have "B.is_iso_class ?gB"
                   by (simp add: B.is_iso_classI)
                 moreover have "B.is_iso_class (Map K)"
                   by (metis (no_types, lifting) CLS_REP HK Map.simps(1)
                       B.is_iso_classI k.ide_left k_def)
                 moreover have "?g  B.tab1 (k*  h)B"
                   by (simp add: B.ide_in_iso_class(1))
                 moreover have "«?g : src (B.tab1 (k*  h)) B Dom K»"
                   using seq seq_char B.in_hhom_def seq' by auto
                 moreover have "k  Map K"
                   by fact
                 moreover have "«k : Dom K B Cod K»"
                   by (simp add: HK k_def)
                 moreover have "k  ?g B k  ?g"
                   using B.isomorphic_reflexive iso B.isomorphic_implies_hpar(2) by auto
                 ultimately show "k  B.tab1 (k*  h)
                                     {h'. B.is_iso_class B.tab1 (k*  h)B 
                                           B.is_iso_class (Map K) 
                                           (f g. f  B.tab1 (k*  h)B 
                                                  g  Map K  g  f B h')}"
                   by auto
               qed
               show "h  ?f B k  ?g"
                 using iso by simp
             qed
             have "K  PRJ1 K H = K  ⟦⟦?g⟧⟧"
               unfolding PRJ1_def using HK h_def k_def by simp
             also have "... = MkArr (src ?g) (Cod K) (Comp (Map K) ?gB)"
               using seq' comp_char [of K "⟦⟦?g⟧⟧"] by simp
             also have "... = MkArr (src ?f) (Cod H) (Comp (Map H) ?fB)"
               using * HK cod_char by auto
             also have "... = comp H ⟦⟦?f⟧⟧"
               using seq comp_char [of H "⟦⟦?f⟧⟧"] by simp
             also have "... = comp H (PRJ0 K H)"
               unfolding PRJ0_def using HK h_def k_def by simp
             finally show ?thesis by simp
           qed
         qed
       qed
       show "H K U V. commutative_square K H V U 
                        ∃!E. comp (PRJ1 K H) E = V  comp (PRJ0 K H) E = U"
       proof -
         fix H K U V
         assume cs: "commutative_square K H V U"
         have HK: "cospan K H"
           using cs by auto
         (* TODO: Is there any way to avoid this repetition? *)
         define h where "h = REP H"
         define k where "k = REP K"
         have h: "h  Map H"
           using h_def HK REP_in_Map by blast
         have k: "k  Map K"
           using k_def HK REP_in_Map by blast
         have 1: "B.is_left_adjoint h  B.is_left_adjoint k  B.ide h  B.ide k  trg h = trg k"
           using h k h_def k_def HK arr_char cod_char B.in_hhom_def B.left_adjoint_is_ide
                 in_HomD(5) [of "Map H" "Dom H" "Cod H" h]
                 in_HomD(5) [of "Map K" "Dom K" "Cod K" k]
           apply auto
           by (metis (no_types, lifting) HK Dom_cod)
         interpret h: map_in_bicategory (⋅) (⋆) 𝖺 𝗂 src trg h
           using 1 by unfold_locales auto
         interpret k: map_in_bicategory (⋅) (⋆) 𝖺 𝗂 src trg k
           using 1 by unfold_locales auto
         interpret hk: cospan_of_maps_in_bicategory_of_spans (⋅) (⋆) 𝖺 𝗂 src trg h k
           using 1 by unfold_locales auto
         let ?f = "B.tab0 (k*  h)"
         let ?g = "B.tab1 (k*  h)"
         have seq_HU: "seq H U"
           using cs by auto 
         have seq_KV: "seq K V"
           using cs by auto 
         let ?u = "REP U"
         let ?v = "REP V"
         have u: "B.ide ?u"
           using ide_REP seq_HU by auto
         have v: "B.ide ?v"
           using ide_REP seq_KV by auto
         have u_is_map: "B.is_left_adjoint ?u"
           using u seq_HU REP_in_Map arr_char [of U]
                 in_HomD(5) [of "Map U" "Dom U" "Cod U" ?u]
           by auto
         have v_is_map: "B.is_left_adjoint ?v"
           using v seq_KV REP_in_Map arr_char [of V]
                 in_HomD(5) [of "Map V" "Dom V" "Cod V" ?v]
           by auto
         have *: "h  ?u B k  ?v"
         proof -
           have "h  ?u B REP (H  U)"
           proof -
             have "h  ?u B REP H  ?u"
             proof -
               have "h B REP H"
                 using h h_def HK arr_char REP_in_Map B.iso_class_elems_isomorphic
                       in_HomD(5) [of "Map H" "Dom H" "Cod H" h] B.isomorphic_reflexive
                 by auto
               thus ?thesis
                 using h_def seq_HU B.isomorphic_implies_hpar(1) B.isomorphic_reflexive
                 by (simp add: seq_char)
             qed
             also have "...  B REP (H  U)"
               using seq_HU isomorphic_hcomp_REP isomorphic_def by blast
             finally show ?thesis by blast
           qed
           also have "... B REP (K  V)"
             using seq_HU cs B.isomorphic_reflexive by auto
           also have "... B (k  ?v)"
           proof -
             have "... B REP K  ?v"
               using seq_KV isomorphic_hcomp_REP B.isomorphic_def B.isomorphic_symmetric
               by blast
             also have "... B k  ?v"
             proof -
               have "k B REP K"
                 using k k_def HK arr_char REP_in_Map B.iso_class_elems_isomorphic
                       in_HomD(5) [of "Map K" "Dom K" "Cod K" k] B.isomorphic_reflexive
                 by auto
               thus ?thesis
                 using k_def seq_KV B.isomorphic_implies_hpar(1) B.isomorphic_reflexive
                 by (simp add: seq_char)
             qed 
             finally show ?thesis by blast
           qed
           finally show ?thesis by blast
         qed
         have hseq_hu: "src h = trg ?u"
           using * B.isomorphic_implies_hpar
           by (meson B.hseqE B.ideD(1))
         have hseq_kv: "src k = trg ?v"
           using * B.isomorphic_implies_hpar
           by (meson B.hseqE B.ideD(1))

         obtain w where w: "B.is_left_adjoint w  ?f  w B ?u  ?v B (?g  w)"
           using * u_is_map v_is_map hk.has_pseudo_pullback [of ?u ?v] B.isomorphic_symmetric
           by blast
         have w_in_hom: "«w : src ?u B src ?f»  B.ide w"
           using w B.left_adjoint_is_ide B.src_cod B.trg_cod B.isomorphic_def
           by (metis B.hseqE B.ideD(1) B.in_hhom_def B.isomorphic_implies_hpar(3)
               B.isomorphic_implies_ide(1) B.src_hcomp)
         let ?W = "CLS w"
         have W: "«?W : dom U  dom ⟦⟦?f⟧⟧»"
         proof
           show "arr ?W"
             using w CLS_in_hom by blast
           thus "dom ?W = dom U"
             using w_in_hom dom_char REP_in_hhom(1) CLS_in_hom
             by (metis (no_types, lifting) Dom.simps(1) commutative_squareE
                 dom_char REP_simps(2) cs B.in_hhomE)
           show "cod ?W = dom ⟦⟦?f⟧⟧"
           proof -
             have "src ?f = trg w"
               by (metis (lifting) B.in_hhomE w_in_hom)
             thus ?thesis
               using CLS_in_hom [of ?f] CLS_in_hom [of w] hk.satisfies_T0 w by fastforce
           qed
         qed
         show "∃!E. PRJ1 K H  E = V  PRJ0 K H  E = U"
         proof -
           have "PRJ1 K H  ?W = V  PRJ0 K H  ?W = U"
           proof -
             have "⟦⟦?f⟧⟧  ?W = U"
               using w w_in_hom u CLS_in_hom comp_CLS
                     B.isomorphic_symmetric CLS_REP hk.leg0_is_map
               by (metis (mono_tags, lifting) commutative_square_def cs)
             moreover have "⟦⟦?g⟧⟧  ?W = V"
               using w w_in_hom v CLS_in_hom comp_CLS
                     B.isomorphic_symmetric CLS_REP hk.leg1_is_map
               by (metis (mono_tags, lifting) commutative_square_def cs)
             ultimately show ?thesis
               using HK h_def k_def PRJ0_def PRJ1_def by auto
           qed
           moreover have
             "W'. PRJ1 K H  W' = V  PRJ0 K H  W' = U  W' = ?W"
           proof -
              fix W'
              assume "PRJ1 K H  W' = V  PRJ0 K H  W' = U"
              hence W': "«W' : dom U  dom ⟦⟦?f⟧⟧»  ⟦⟦?f⟧⟧  W' = U  ⟦⟦?g⟧⟧  W' = V"
                using PRJ0_def PRJ1_def HK h_def k_def apply simp
                using cs arr_iff_in_hom by blast
              let ?w' = "REP W'"
              have w': "B.ide ?w'"
                using W' ide_REP by auto

              have fw'_iso_u: "?f  ?w' B ?u"
              proof -
                have "?f  ?w' B REP ⟦⟦?f⟧⟧  ?w'"
                  by (metis (no_types, lifting) Cod.simps(1) in_hom_char
                      REP_CLS REP_simps(3) W W' B.hcomp_isomorphic_ide hk.satisfies_T0
                      B.in_hhomE B.isomorphic_symmetric w' w_in_hom)
                also have "REP ⟦⟦?f⟧⟧  ?w' B ?u"
                   using W' isomorphic_hcomp_REP cs by blast
                finally show ?thesis by blast
              qed

              have gw'_iso_v: "?g  ?w' B ?v"
              proof -
                have "?g  ?w' B REP ⟦⟦?g⟧⟧  ?w'"
                proof -
                  have "?g B REP ⟦⟦?g⟧⟧"
                    using REP_CLS B.isomorphic_symmetric hk.leg1_is_map by blast
                  moreover have "B.ide (REP W')"
                    using W' by auto
                  moreover have "src ?f = trg ?w'"
                    using w_in_hom W W' in_hom_char arr_char B.in_hhom_def
                    by (meson fw'_iso_u B.hseqE B.ideD(1) B.isomorphic_implies_ide(1))
                  ultimately show ?thesis
                    using B.hcomp_isomorphic_ide by simp
                qed
                also have "... B ?v"
                   using W' isomorphic_hcomp_REP cs by blast
                finally show ?thesis by blast
              qed

              show "W' = ?W"
              proof -
                have "W' = ⟦⟦?w'⟧⟧"
                  using w' W' CLS_REP by auto
                also have "... = ?W"
                proof -
                  have "?w' B w"
                    using * w W' hk.has_pseudo_pullback(2) u_is_map v_is_map
                          B.isomorphic_symmetric fw'_iso_u gw'_iso_v
                    by blast
                  thus ?thesis
                    using CLS_eqI B.iso_class_eqI w' by blast
                qed
                finally show ?thesis by blast
              qed
            qed
            ultimately show ?thesis by auto
          qed
        qed
      qed

    lemma is_elementary_category_with_pullbacks:
    shows "elementary_category_with_pullbacks comp PRJ0 PRJ1"
       ..

    lemma is_category_with_pullbacks:
    shows "category_with_pullbacks comp"
      ..

    sublocale elementary_category_with_pullbacks comp PRJ0 PRJ1
      using is_elementary_category_with_pullbacks by simp

  end

  text ‹
    Here we relate the projections of the chosen pullbacks in Maps(B)› to the
    projections associated with the chosen tabulations in B›.
  ›

  context composite_tabulation_in_maps
  begin

    interpretation Maps: maps_category V H 𝖺 𝗂 src trg
      ..

    interpretation r0s1: cospan_of_maps_in_bicategory_of_spans V H 𝖺 𝗂 src trg s1 r0
      using ρ.leg0_is_map σ.leg1_is_map composable by unfold_locales auto

    lemma prj_char:
    shows "Maps.PRJ0 ⟦⟦r0⟧⟧ ⟦⟦s1⟧⟧ = ⟦⟦prj0 s1 r0⟧⟧"
    and "Maps.PRJ1 ⟦⟦r0⟧⟧ ⟦⟦s1⟧⟧ = ⟦⟦prj1 s1 r0⟧⟧"
    proof -
      have "Maps.arr (Maps.MkArr (src s0) (trg s) s1)"
        using σ.leg1_in_hom Maps.CLS_in_hom σ.leg1_is_map Maps.arr_char by auto
      moreover have "Maps.arr (Maps.MkArr (src r0) (trg s) r0)"
        using Maps.CLS_in_hom composable r0s1.k_is_map by fastforce
      moreover have "Maps.cod (Maps.MkArr (src r0) (trg s) r0) =
                     Maps.cod (Maps.MkArr (src s0) (trg s) s1)"
        unfolding Maps.arr_char
        using σ.leg1_in_hom ρ.leg0_in_hom
        by (simp add: Maps.cod_char calculation(1) calculation(2))
      ultimately have "Maps.PRJ0 ⟦⟦r0⟧⟧ ⟦⟦s1⟧⟧ =
                       ⟦⟦tab0 ((Maps.REP (Maps.MkArr (src r0) (trg s) r0))* 
                              Maps.REP (Maps.MkArr (src s0) (trg s) s1))⟧⟧ 
                       Maps.PRJ1 (Maps.CLS r0) (Maps.CLS s1) =
                       ⟦⟦tab1 ((Maps.REP (Maps.MkArr (src r0) (trg s) r0))* 
                              Maps.REP (Maps.MkArr (src s0) (trg s) s1))⟧⟧"
        unfolding Maps.PRJ0_def Maps.PRJ1_def
        using Maps.CLS_in_hom σ.leg1_is_map ρ.leg0_is_map composable by simp
      moreover have "r0*  s1  (Maps.REP (Maps.MkArr (src r0) (trg s) r0))* 
                                Maps.REP (Maps.MkArr (src s0) (trg s) s1)"
      proof -
        have "r0  Maps.REP (Maps.MkArr (src r0) (trg s) r0)"
          using Maps.REP_CLS composable isomorphic_symmetric r0s1.k_is_map by fastforce
        hence 3: "isomorphic r0* (Maps.REP (Maps.MkArr (src r0) (trg s) r0))*"
          using ρ.leg0_is_map
          by (simp add: isomorphic_to_left_adjoint_implies_isomorphic_right_adjoint)
        moreover have 4: "s1  Maps.REP (Maps.MkArr (src s0) (trg s) s1)"
          using Maps.REP_CLS isomorphic_symmetric r0s1.h_is_map by fastforce
        ultimately show ?thesis
        proof -
          have 1: "src r0* = trg s1"
            using ρ.T0.antipar(2) r0s1.cospan by argo
          have 2: "ide s1"
            by simp
          have "src (Maps.REP (Maps.MkArr (src r0) (trg s) r0))* = trg s1"
            by (metis 3 ρ.T0.antipar(2) isomorphic_implies_hpar(3) r0s1.cospan)
          thus ?thesis
            using 1 2
            by (meson 3 4 hcomp_ide_isomorphic hcomp_isomorphic_ide isomorphic_implies_ide(2)
                isomorphic_transitive)
        qed
      qed
      ultimately have 1: "Maps.PRJ0 ⟦⟦r0⟧⟧ ⟦⟦s1⟧⟧ = ⟦⟦prj0 s1 r0⟧⟧ 
                          Maps.PRJ1 ⟦⟦r0⟧⟧ ⟦⟦s1⟧⟧ = ⟦⟦prj1 s1 r0⟧⟧"
        using r0s1.isomorphic_implies_same_tab by simp
      show "Maps.PRJ0 ⟦⟦r0⟧⟧ ⟦⟦s1⟧⟧ = ⟦⟦prj0 s1 r0⟧⟧"
        using 1 by simp
      show "Maps.PRJ1 ⟦⟦r0⟧⟧ ⟦⟦s1⟧⟧ = ⟦⟦prj1 s1 r0⟧⟧"
        using 1 by simp
    qed

  end

  context identity_in_bicategory_of_spans
  begin

    interpretation Maps: maps_category V H 𝖺 𝗂 src trg ..
    interpretation Span: span_bicategory Maps.comp Maps.PRJ0 Maps.PRJ1 ..

    notation isomorphic (infix "" 50)

    text ‹
      A 1-cell r› in a bicategory of spans is a map if and only if the ``input leg''
      tab0 r› of the chosen tabulation of r› is an equivalence map.
      Since a tabulation of r› is unique up to equivalence, and equivalence maps compose,
      the result actually holds if ``chosen tabulation'' is replaced by ``any tabulation''.
    ›

    lemma is_map_iff_tab0_is_equivalence:
    shows "is_left_adjoint r  equivalence_map (tab0 r)"
    proof
      assume 1: "equivalence_map (tab0 r)"
      have 2: "quasi_inverses (tab0 r) (tab0 r)*"
      proof -
        obtain g' η' ε' where η'ε': "equivalence_in_bicategory V H 𝖺 𝗂 src trg (tab0 r) g' η' ε'"
          using 1 equivalence_map_def by auto
        have "adjoint_pair (tab0 r) g'"
          using η'ε' quasi_inverses_def quasi_inverses_are_adjoint_pair by blast
        moreover have "adjoint_pair (tab0 r) (tab0 r)*"
          using T0.adjunction_in_bicategory_axioms adjoint_pair_def by auto
        ultimately have "g'  (tab0 r)*"
          using left_adjoint_determines_right_up_to_iso by simp
        thus ?thesis
          using η'ε' quasi_inverses_def quasi_inverses_isomorphic_right by blast
      qed
      obtain η' ε' where η'ε': "equivalence_in_bicategory V H 𝖺 𝗂 src trg (tab0 r) (tab0 r)* η' ε'"
        using 2 quasi_inverses_def by auto
      interpret η'ε': equivalence_in_bicategory V H 𝖺 𝗂 src trg tab0 r (tab0 r)* η' ε'
        using η'ε' by auto
      have "is_left_adjoint (tab0 r)*"
        using 2 quasi_inverses_are_adjoint_pair quasi_inverses_symmetric by blast
      hence "is_left_adjoint (tab1 r  (tab0 r)*)"
        using left_adjoints_compose by simp
      thus "is_left_adjoint r"
        using yields_isomorphic_representation isomorphic_def left_adjoint_preserved_by_iso'
        by meson
      next
      assume 1: "is_left_adjoint r"
      have 2: "is_left_adjoint (tab1 r  (tab0 r)*)"
        using 1 yields_isomorphic_representation left_adjoint_preserved_by_iso'
              isomorphic_symmetric isomorphic_def
        by meson
      hence "is_left_adjoint (tab0 r)*"
        using is_ide BS4 [of "tab1 r" "(tab0 r)*"] by auto
      hence "is_left_adjoint ((tab0 r)*  tab0 r)  is_left_adjoint (tab0 r  (tab0 r)*)"
        using left_adjoints_compose T0.antipar by simp
      hence 3: "iso η  iso ε"
        using BS3 [of "src (tab0 r)" "(tab0 r)*  tab0 r" η η]
              BS3 [of "tab0 r  (tab0 r)*" "trg (tab0 r)" ε ε]
              T0.unit_in_hom T0.counit_in_hom obj_is_self_adjoint
        by auto
      hence "equivalence_in_bicategory V H 𝖺 𝗂 src trg (tab0 r) (tab0 r)* η ε"
        apply unfold_locales by auto
      thus "equivalence_map (tab0 r)"
        using equivalence_map_def by blast
    qed

    text ‹
      The chosen tabulation (and indeed, any other tabulation, which is equivalent)
      of an object is symmetric in the sense that its two legs are isomorphic.
    ›

    lemma obj_has_symmetric_tab:
    assumes "obj r"
    shows "tab0 r  tab1 r"
    proof -
      have "tab0 r  r  tab0 r"
      proof -
        have "trg (tab0 r) = r"
          using assms by auto
        moreover have "«𝗅-1[tab0 r] : tab0 r  trg (tab0 r)  tab0 r»  iso 𝗅-1[tab0 r]"
          using assms by simp
        ultimately show ?thesis
          unfolding isomorphic_def by metis
      qed
      also have "...  tab1 r"
      proof -
        have "«tab : tab1 r  r  tab0 r»"
          using tab_in_hom by simp
        moreover have "is_left_adjoint (r  tab0 r)"
          using assms left_adjoints_compose obj_is_self_adjoint by simp
        ultimately show ?thesis
          using BS3 [of "tab1 r" "r  tab0 r" tab tab] isomorphic_symmetric isomorphic_def
          by auto
      qed
      finally show ?thesis by simp
    qed

    text ‹
      The chosen tabulation of r› determines a span in Maps(B)›.
    ›

    lemma determines_span:
    assumes "ide r"
    shows "span_in_category Maps.comp Leg0 = ⟦⟦tab0 r⟧⟧, Leg1 = ⟦⟦tab1 r⟧⟧"
      using assms Maps.CLS_in_hom [of "tab0 r"] Maps.CLS_in_hom [of "tab1 r"]
            tab0_in_hom tab1_in_hom
      apply unfold_locales by fastforce

  end

  subsection "Arrows of Tabulations in Maps"

  text ‹
    Here we consider the situation of two tabulations: a tabulation ρ› of r›
    and a tabulation σ› of s›, both ``legs'' of each tabulation being maps,
    together  with an arbitrary 2-cell «μ : r ⇒ s»›.
    The 2-cell μ› at the base composes with the tabulation ρ› to yield a 2-cell
    Δ = (μ ⋆ r0) ⋅ ρ› ``over'' s.  By property T1› of tabulation σ›, this induces a map
    from the apex of ρ› to the apex of σ›, which together with the other data
    forms a triangular prism whose sides commute up to (unique) isomorphism.
  ›
  text ‹
$$
\xymatrix{
  && {\rm src}~\sigma \ar[dl]_{s_1} \ar[dr]^{s_0} \dtwocell\omit{^<-1>\sigma} & \\
  &{\rm trg}~s && {\rm src}~s \ar[ll]^{s} \\
  & \rrtwocell\omit{^\mu} &&\\
  & {\rm src}~\rho \ar[dl]_{r_1} \ar[dr]^{r_0} \ar@ {.>}[uuur]^<>(0.3){{\rm chine}} \dtwocell\omit{^\rho}& \\
  {\rm trg}~r \ar@ {=}[uuur] && {\rm src}~r \ar[ll]^{r} \ar@ {=}[uuur]
}
$$
  ›

  locale arrow_of_tabulations_in_maps =
    bicategory_of_spans V H 𝖺 𝗂 src trg +
    ρ: tabulation_in_maps V H 𝖺 𝗂 src trg r ρ r0 r1 +
    σ: tabulation_in_maps V H 𝖺 𝗂 src trg s σ s0 s1
  for V :: "'a comp"                 (infixr "" 55)
  and H :: "'a  'a  'a"          (infixr "" 53)
  and 𝖺 :: "'a  'a  'a  'a"     ("𝖺[_, _, _]")
  and 𝗂 :: "'a  'a"                 ("𝗂[_]")
  and src :: "'a  'a"
  and trg :: "'a  'a"
  and r :: 'a
  and ρ :: 'a
  and r0 :: 'a
  and r1 :: 'a
  and s :: 'a
  and σ :: 'a
  and s0 :: 'a
  and s1 :: 'a
  and μ :: 'a +
  assumes in_hom: "«μ : r  s»"
  begin

    abbreviation (input) Δ
    where "Δ  (μ  r0)  ρ"

    lemma Δ_in_hom [intro]:
    shows "«Δ : src ρ  trg σ»"
    and "«Δ : r1  s  r0»"
    proof -
      show "«Δ : r1  s  r0»"
        using in_hom ρ.leg0_in_hom(2) ρ.tab_in_vhom' by auto
      thus "«Δ : src ρ  trg σ»"
        by (metis ρ.tab_simps(3) ρ.base_in_hom(2) σ.tab_simps(3) σ.base_in_hom(2) arrI in_hom
            seqI' vcomp_in_hhom vseq_implies_hpar(1-2))
    qed

    lemma Δ_simps [simp]:
    shows "arr Δ"
    and "src Δ = src ρ" and "trg Δ = trg σ"
    and "dom Δ = r1" and "cod Δ = s  r0"
      using Δ_in_hom by auto

    abbreviation is_induced_map
    where "is_induced_map w  σ.is_induced_by_cell w r0 Δ"

    text ‹
      The following is an equivalent restatement, in elementary terms, of the conditions
      for being an induced map.
    ›

    abbreviation (input) is_induced_map'
    where "is_induced_map' w 
           ide w 
           (ν θ. «θ : s0  w  r0»  «ν : r1  s1  w»  iso ν 
                  Δ = (s  θ)  𝖺[s, s0, w]  (σ  w)  ν)"

    lemma is_induced_map_iff:
    shows "is_induced_map w  is_induced_map' w"
    proof
      assume w: "is_induced_map' w"
      show "is_induced_map w"
      proof
        have 1: "dom Δ = r1"
          by auto
        interpret w: arrow_of_spans_of_maps_to_tabulation V H 𝖺 𝗂 src trg
                       r0 dom Δ s σ s0 s1 w
        proof -
          have "arrow_of_spans_of_maps_to_tabulation V H 𝖺 𝗂 src trg r0 r1 s σ s0 s1 w"
            using w apply unfold_locales by auto
          thus "arrow_of_spans_of_maps_to_tabulation V H 𝖺 𝗂 src trg r0 (dom Δ) s σ s0 s1 w"
            using 1 by simp
        qed
        show "arrow_of_spans_of_maps V H 𝖺 𝗂 src trg r0 (dom Δ) s0 s1 w"
          using w.arrow_of_spans_of_maps_axioms by auto
        show "σ.composite_cell w w.the_θ  w.the_ν = Δ"
        proof -
          obtain θ ν
          where θν: "«θ : s0  w  r0»  «ν : r1  s1  w»  iso ν 
                     Δ = (s  θ)  𝖺[s, s0, w]  (σ  w)  ν"
            using w w.the_θ_props(1) by auto
          have "(s  θ)  𝖺[s, s0, w]  (σ  w)  ν = Δ"
            using θν by argo
          moreover have "θ = w.the_θ  ν = w.the_ν"
            using θν 1 w.the_ν_props(1) w.leg0_uniquely_isomorphic w.leg1_uniquely_isomorphic
            by auto
          ultimately show ?thesis
            using comp_assoc by simp
        qed
      qed
      next
      assume w: "is_induced_map w"
      show "is_induced_map' w"
      proof -
        interpret w: arrow_of_spans_of_maps V H 𝖺 𝗂 src trg r0 r1 s0 s1 w
          using w in_hom by auto
        interpret w: arrow_of_spans_of_maps_to_tabulation V H 𝖺 𝗂 src trg r0 r1 s σ s0 s1 w
          ..
        have "dom Δ = r1" by auto
        thus ?thesis
          using w comp_assoc w.the_ν_props(1) w.the_ν_props(2) w.uwθ by metis
      qed
    qed

    lemma exists_induced_map:
    shows "w. is_induced_map w"
    proof -
      obtain w θ ν
      where wθν: "ide w  «θ : s0  w  r0»  «ν : r1  s1  w»  iso ν 
                  Δ = (s  θ)  𝖺[s, s0, w]  (σ  w)  ν"
        using Δ_in_hom ρ.ide_leg0 σ.T1 comp_assoc
        by (metis in_homE)
      thus ?thesis
        using is_induced_map_iff by blast
    qed

    lemma induced_map_unique:
    assumes "is_induced_map w" and "is_induced_map w'"
    shows "w  w'"
      using assms σ.induced_map_unique by blast

    definition chine
    where "chine  SOME w. is_induced_map w"

    lemma chine_is_induced_map:
    shows "is_induced_map chine"
      unfolding chine_def
      using exists_induced_map someI_ex [of is_induced_map] by simp

    lemma chine_in_hom [intro]:
    shows "«chine : src r0  src s0»"
    and "«chine: chine  chine»"
    proof -
      show "«chine : src r0  src s0»"
        using chine_is_induced_map
        by (metis Δ_simps(1) Δ_simps(4) ρ.leg1_simps(3) σ.ide_base σ.ide_leg0 σ.leg0_simps(3)
            σ.tab_simps(2) arrow_of_spans_of_maps.is_ide arrow_of_spans_of_maps.the_ν_simps(2)
            assoc_simps(2) hseqE in_hhom_def seqE src_vcomp vseq_implies_hpar(1))
      show "«chine: chine  chine»"
        using chine_is_induced_map
        by (meson arrow_of_spans_of_maps.is_ide ide_in_hom(2))
    qed

    lemma chine_simps [simp]:
    shows "arr chine" and "ide chine"
    and "src chine = src r0" and "trg chine = src s0"
    and "dom chine = chine" and "cod chine = chine"
      using chine_in_hom apply auto
      by (meson arrow_of_spans_of_maps.is_ide chine_is_induced_map)

  end

  sublocale arrow_of_tabulations_in_maps 
            arrow_of_spans_of_maps V H 𝖺 𝗂 src trg r0 r1 s0 s1 chine
      using chine_is_induced_map is_induced_map_iff
      by unfold_locales auto

  sublocale arrow_of_tabulations_in_maps 
            arrow_of_spans_of_maps_to_tabulation V H 𝖺 𝗂 src trg r0 r1 s σ s0 s1 chine
      ..

  context arrow_of_tabulations_in_maps
  begin

    text ‹
      The two factorizations of the composite 2-cell Δ› amount to a naturality condition.
    ›

    lemma Δ_naturality:
    shows "(μ  r0)  ρ = (s  the_θ)  𝖺[s, s0, chine]  (σ  chine)  the_ν"
      using chine_is_induced_map is_induced_map_iff
      by (metis leg0_uniquely_isomorphic(2) leg1_uniquely_isomorphic(2) the_ν_props(1) uwθ)

    lemma induced_map_preserved_by_iso:
    assumes "is_induced_map w" and "isomorphic w w'"
    shows "is_induced_map w'"
    proof -
      interpret w: arrow_of_spans_of_maps V H 𝖺 𝗂 src trg r0 r1 s0 s1 w
        using assms in_hom by auto
      interpret w: arrow_of_spans_of_maps_to_tabulation V H 𝖺 𝗂 src trg r0 r1 s σ s0 s1 w
        ..
      obtain φ where φ: "«φ : w  w'»  iso φ"
        using assms(2) isomorphic_def by auto
      show ?thesis
      proof
        interpret w': arrow_of_spans_of_maps V H 𝖺 𝗂 src trg r0 dom Δ s0 s1 w'
        proof
          show "is_left_adjoint r0"
            by (simp add: ρ.satisfies_T0)
          show "is_left_adjoint (dom Δ)"
            by (simp add: ρ.leg1_is_map)
          show "ide w'" using assms by force
          show "θ. «θ : s0  w'  r0»"
            using φ w.the_θ_props σ.leg0_in_hom(2) assms(2) comp_in_homI hcomp_in_vhom
                  inv_in_hom isomorphic_implies_hpar(4) w.the_θ_simps(4) w.w_simps(4)
            by metis
          have "«(s1  φ)  w.the_ν : r1  s1  w'»  iso ((s1  φ)  w.the_ν)"
          proof (intro conjI)
            show "«(s1  φ)  w.the_ν : r1  s1  w'»"
              using φ w.the_ν_props
              by (intro comp_in_homI, auto)
            thus "iso ((s1  φ)  w.the_ν)"
              using φ w.the_ν_props
              by (meson σ.ide_leg1 arrI iso_hcomp hseqE ide_is_iso isos_compose seqE)
          qed
          hence "ν. «ν : r1  s1  w'»  iso ν"
            by auto
          thus "ν. «ν : dom Δ  s1  w'»  iso ν"
            by auto
        qed
        interpret w': arrow_of_spans_of_maps_to_tabulation V H 𝖺 𝗂 src trg
                        r0 dom Δ s σ s0 s1 w'
          ..
        show "arrow_of_spans_of_maps V H 𝖺 𝗂 src trg r0 (dom Δ) s0 s1 w'"
          using w'.arrow_of_spans_of_maps_axioms by auto
        show "σ.composite_cell w' w'.the_θ  w'.the_ν = Δ"
        proof -
          have 1: "w'.the_θ = w.the_θ  (s0  inv φ)" 
          proof -
            have "«w.the_θ  (s0  inv φ) : s0  w'  r0»"
              using w.the_θ_props φ
              by (intro comp_in_homI, auto)
            moreover have "«w'.the_θ : s0  w'  r0»"
              using w'.the_θ_props by simp
            ultimately show ?thesis
              using w'.leg0_uniquely_isomorphic(2) by blast
          qed
          moreover have "w'.the_ν = (s1  φ)  w.the_ν"
          proof -
            have "«(s1  φ)  w.the_ν : dom Δ  s1  w'»"
              using w.the_ν_props φ
              by (intro comp_in_homI, auto)
            moreover have "iso ((s1  φ)  w.the_ν)"
              using w.the_ν_props φ iso_hcomp
              by (meson σ.ide_leg1 arrI calculation hseqE ide_is_iso isos_compose seqE)
            ultimately show ?thesis
              using w'.the_ν_props w'.leg1_uniquely_isomorphic(2) by blast
          qed
          ultimately have "σ.composite_cell w' w'.the_θ  w'.the_ν =
                           σ.composite_cell w' (w.the_θ  (s0  inv φ))  (s1  φ)  w.the_ν"
            by simp
          also have "... = (s  w.the_θ  (s0  inv φ))  𝖺[s, s0, w'] 
                             (σ  w')  (s1  φ)  w.the_ν"
            using comp_assoc by presburger
          also have "... = (s  w.the_θ)  ((s  s0  inv φ)  𝖺[s, s0, w'] 
                             (σ  w')  (s1  φ))  w.the_ν"
            using 1 comp_assoc w'.the_θ_simps(1) whisker_left
            by auto
          also have "... = (s  w.the_θ)  (𝖺[s, s0, w]  (σ  w))  w.the_ν"
          proof -
            have "(s  s0  inv φ)  𝖺[s, s0, w']  (σ  w')  (s1  φ) =
                  𝖺[s, s0, w]  (σ  w)"
            proof -
              have "(s  s0  inv φ)  𝖺[s, s0, w']  (σ  w')  (s1  φ) =
                    𝖺[s, s0, w]  ((s  s0)  inv φ)  (σ  w')  (s1  φ)"
              proof -
                have "(s  s0  inv φ)  𝖺[s, s0, w'] = 𝖺[s, s0, w]  ((s  s0)  inv φ)"
                  using assms φ assoc_naturality [of s s0 "inv φ"] w.w_simps(4)
                  by (metis σ.leg0_simps(2-5) σ.base_simps(2-4) arr_inv cod_inv dom_inv
                      in_homE trg_cod)
                thus ?thesis using comp_assoc by metis
              qed
              also have "... = 𝖺[s, s0, w]  (σ  w)  (s1  inv φ)  (s1  φ)"
              proof -
                have "((s  s0)  inv φ)  (σ  w') = (σ  w)  (s1  inv φ)"
                  using φ comp_arr_dom comp_cod_arr in_hhom_def
                        interchange [of "s  s0" σ "inv φ" w']
                        interchange [of σ s1 w "inv φ"]
                  by auto
               thus ?thesis
                  using comp_assoc by metis
              qed
              also have "... = 𝖺[s, s0, w]  (σ  w)"
              proof -
                have "(σ  w)  (s1  inv φ)  (s1  φ) = σ  w"
                proof -
                  have "(σ  w)  (s1  inv φ)  (s1  φ) = (σ  w)  (s1  inv φ  φ)"
                    using φ whisker_left in_hhom_def by auto
                  also have "... = (σ  w)  (s1  w)"
                    using φ comp_inv_arr' by auto
                  also have "... = σ  w"
                    using whisker_right [of w σ s1] comp_arr_dom in_hhom_def by auto
                  finally show ?thesis by blast
                qed
                thus ?thesis by simp
              qed
              finally show ?thesis by simp
            qed
            thus ?thesis by simp
          qed
          also have "... = Δ"
            using assms(1) comp_assoc w.is_ide w.the_ν_props(1) w.the_θ_props(1) by simp
          finally show ?thesis
            using comp_assoc by auto
        qed
      qed
    qed

  end

  text ‹
    In the special case that μ› is an identity 2-cell, the induced map from the apex of ρ›
    to the apex of σ› is an equivalence map.
  ›

  locale identity_arrow_of_tabulations_in_maps =
    arrow_of_tabulations_in_maps +
  assumes is_ide: "ide μ"
  begin

    lemma r_eq_s:
    shows "r = s"
      using is_ide by (metis ide_char in_hom in_homE)

    lemma Δ_eq_ρ:
    shows "Δ = ρ"
      by (meson Δ_simps(1) comp_ide_arr ide_hcomp hseq_char' ide_u is_ide seqE
          seq_if_composable)

    lemma chine_is_equivalence:
    shows "equivalence_map chine"
    proof -
      obtain w w' φ ψ θ ν θ' ν'
        where e: "equivalence_in_bicategory (⋅) (⋆) 𝖺 𝗂 src trg w' w ψ φ 
                  «w : src s0  src r0»  «w' : src r0  src s0» 
                  «θ : r0  w  s0»  «ν : s1  r1  w»  iso ν 
                  σ = (s  θ)  𝖺[s, r0, w]  (ρ  w)  ν 
                  «θ' : s0  w'  r0»  «ν' : r1  s1  w'»  iso ν' 
                  ρ = (s  θ')  𝖺[s, s0, w']  (σ  w')  ν'"
        using r_eq_s σ.apex_unique_up_to_equivalence [of ρ r0 r1] ρ.tabulation_axioms by blast
      have w': "equivalence_map w'"
        using e equivalence_map_def by auto
      hence "is_induced_map w'"
        using e r_eq_s Δ_eq_ρ is_induced_map_iff comp_assoc equivalence_map_is_ide by metis
      hence "isomorphic chine w'"
        using induced_map_unique chine_is_induced_map by simp
      thus ?thesis
        using w' equivalence_map_preserved_by_iso isomorphic_symmetric by blast
    qed

  end

  text ‹
    The following gives an interpretation of @{locale arrow_of_tabulations_in_maps}
    in the special case that the tabulations are those that we have chosen for the
    domain and codomain of the underlying 2-cell «μ : r ⇒ s»›.
    In this case, we can recover μ› from Δ› via adjoint transpose.
  ›

  locale arrow_in_bicategory_of_spans =
    bicategory_of_spans V H 𝖺 𝗂 src trg +
    r: identity_in_bicategory_of_spans V H 𝖺 𝗂 src trg r +
    s: identity_in_bicategory_of_spans V H 𝖺 𝗂 src trg s
  for V :: "'a comp"                 (infixr "" 55)
  and H :: "'a  'a  'a"          (infixr "" 53)
  and 𝖺 :: "'a  'a  'a  'a"     ("𝖺[_, _, _]")
  and 𝗂 :: "'a  'a"                 ("𝗂[_]")
  and src :: "'a  'a"
  and trg :: "'a  'a"
  and r :: 'a
  and s :: 'a
  and μ :: 'a +
  assumes in_hom: "«μ : r  s»"
  begin

    abbreviation (input) r0 where "r0  tab0 r"
    abbreviation (input) r1 where "r1  tab1 r"
    abbreviation (input) s0 where "s0  tab0 s"
    abbreviation (input) s1 where "s1  tab1 s"

    lemma is_arrow_of_tabulations_in_maps:
    shows "arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg r r.tab r0 r1 s s.tab s0 s1 μ"
      using in_hom by unfold_locales auto

  end

  sublocale identity_in_bicategory_of_spans  arrow_in_bicategory_of_spans V H 𝖺 𝗂 src trg r r r
    apply unfold_locales using is_ide by auto

  context arrow_in_bicategory_of_spans
  begin

    interpretation arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg r r.tab r0 r1 s s.tab s0 s1 μ
      using is_arrow_of_tabulations_in_maps by simp
    interpretation r: arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg r r.tab r0 r1 r r.tab r0 r1 r
      using r.is_arrow_of_tabulations_in_maps by simp

    lemma μ_in_terms_of_Δ:
    shows "μ = r.T0.trnrε (cod μ) Δ  inv (r.T0.trnrε r r.tab)"
    proof -
      have μ: "arr μ"
        using in_hom by auto
      have "μ  r.T0.trnrε r r.tab = r.T0.trnrε s Δ"
      proof -
        have "μ  r.T0.trnrε r r.tab =
              (μ  𝗋[r])  (r  r.ε)  𝖺[r, tab0 r, (tab0 r)*]  (r.tab  (tab0 r)*)"
          unfolding r.T0.trnrε_def using comp_assoc by presburger
        also have "... = 𝗋[s]  ((μ  src μ)  (r  r.ε)) 
                           𝖺[r, tab0 r, (tab0 r)*]  (r.tab  (tab0 r)*)"
          using μ runit_naturality comp_assoc
          by (metis in_hom in_homE)
        also have "... = 𝗋[s]  (s  r.ε)  ((μ  tab0 r  (tab0 r)*) 
                           𝖺[r, tab0 r, (tab0 r)*])  (r.tab  (tab0 r)*)"
        proof -
          have "(μ  src μ)  (r  r.ε) = μ  r.ε"
            using μ interchange comp_arr_dom comp_cod_arr
            by (metis in_hom in_homE r.T0.counit_simps(1) r.T0.counit_simps(3) r.u_simps(3)
                src_dom)
          also have "... = (s  r.ε)  (μ  tab0 r  (tab0 r)*)"
            using in_hom interchange [of s μ r.ε "tab0 r  (tab0 r)*"]
                  comp_arr_dom comp_cod_arr r.T0.counit_simps(1) r.T0.counit_simps(2)
            by auto
          finally have "(μ  src μ)  (r  r.ε) = (s  r.ε)  (μ  tab0 r  (tab0 r)*)"
            by blast
          thus ?thesis
            using comp_assoc by presburger
        qed
        also have "... = 𝗋[s]  (s  r.ε)  𝖺[s, tab0 r, (tab0 r)*] 
                           ((μ  tab0 r)  (tab0 r)*)  (r.tab  (tab0 r)*)"
        proof -
          have "(μ  tab0 r  (tab0 r)*)  𝖺[r, tab0 r, (tab0 r)*] =
                𝖺[s, tab0 r, (tab0 r)*]  ((μ  tab0 r)  (tab0 r)*)"
            using μ assoc_naturality [of μ "tab0 r" "(tab0 r)*"]
            by (metis ide_char in_hom in_homE r.T0.antipar(1) r.T0.ide_right r.u_simps(3)
                src_dom u_simps(2) u_simps(4-5))
          thus ?thesis
            using comp_assoc by presburger
        qed
        also have "... = 𝗋[s]  (s  r.ε)  𝖺[s, tab0 r, (tab0 r)*] 
                           ((μ  tab0 r)  r.tab  (tab0 r)*)"
          using μ whisker_right Δ_simps(1) by auto
        also have "... = r.T0.trnrε s Δ"
          unfolding r.T0.trnrε_def by simp
        finally show ?thesis by blast
      qed
      thus ?thesis
        using μ r.yields_isomorphic_representation invert_side_of_triangle(2)
        by (metis in_hom in_homE seqI')
    qed

  end

  subsubsection "Vertical Composite"

  locale vertical_composite_of_arrows_of_tabulations_in_maps =
    bicategory_of_spans V H 𝖺 𝗂 src trg +
    ρ: tabulation_in_maps V H 𝖺 𝗂 src trg r ρ r0 r1 +
    σ: tabulation_in_maps V H 𝖺 𝗂 src trg s σ s0 s1 +
    τ: tabulation_in_maps V H 𝖺 𝗂 src trg t τ t0 t1 +
    μ: arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg r ρ r0 r1 s σ s0 s1 μ +
    π: arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg s σ s0 s1 t τ t0 t1 π
  for V :: "'a comp"                 (infixr "" 55)
  and H :: "'a  'a  'a"          (infixr "" 53)
  and 𝖺 :: "'a  'a  'a  'a"     ("𝖺[_, _, _]")
  and 𝗂 :: "'a  'a"                 ("𝗂[_]")
  and src :: "'a  'a"
  and trg :: "'a  'a"
  and r :: 'a
  and ρ :: 'a
  and r0 :: 'a
  and r1 :: 'a
  and s :: 'a
  and σ :: 'a
  and s0 :: 'a
  and s1 :: 'a
  and t :: 'a
  and τ :: 'a
  and t0 :: 'a
  and t1 :: 'a
  and μ :: 'a
  and π :: 'a
  begin

    text ‹
$$
\xymatrix{
  &&& {\rm src}~\tau \ar[dl]_{t_1} \ar[dr]^{t_0} \dtwocell\omit{^<-1>\tau} & \\
  &&{\rm trg}~t && {\rm src}~t \ar[ll]^{s} \\
  && \rrtwocell\omit{^\pi} && \\
  && {\rm src}~\sigma \ar[dl]_{s_1} \ar[dr]^{s_0} \ar[uuur]^<>(0.3){\pi.{\rm chine}} \dtwocell\omit{^<-1>\sigma} & \\
  &{\rm trg}~s \ar@ {=}[uuur] && {\rm src}~s \ar[ll]^{s} \ar@ {=}[uuur] \\
  & \rrtwocell\omit{^\mu} &&\\
  & {\rm src}~\rho \ar[dl]_{r_1} \ar[dr]^{r_0} \ar[uuur]^<>(0.3){\mu.{\rm chine}} \dtwocell\omit{^\rho} & \\
  {\rm trg}~r \ar@ {=}[uuur] && {\rm src}~r \ar[ll]^{r} \ar@ {=}[uuur]
}
$$
    ›

    notation isomorphic (infix "" 50)

    interpretation arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg r ρ r0 r1 t τ t0 t1 π  μ
      using μ.in_hom π.in_hom by (unfold_locales, blast)

    lemma is_arrow_of_tabulations_in_maps:
    shows "arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg r ρ r0 r1 t τ t0 t1 (π  μ)"
      ..

    lemma chine_char:
    shows "chine  π.chine  μ.chine"
    proof -
      have "is_induced_map (π.chine  μ.chine)"
      proof -
        let ?f = "μ.chine"
        have f: "«?f : src ρ  src σ»  is_left_adjoint ?f  ide ?f  μ.is_induced_map ?f"
          using μ.chine_is_induced_map μ.is_map by auto
        let ?g = "π.chine"
        have g: "«?g : src σ  src τ»  is_left_adjoint ?g  ide ?g  π.is_induced_map ?g"
          using π.chine_is_induced_map π.is_map by auto
        let  = "μ.the_θ  (π.the_θ  ?f)  𝖺-1[t0, ?g, ?f]"
        let  = "𝖺[t1, ?g, ?f]  (π.the_ν  ?f)  μ.the_ν"
        have θ: "« : t0  ?g  ?f  r0»"
          using f g π.the_θ_props μ.the_θ_props
          by (intro comp_in_homI hcomp_in_vhom, auto+)
        have ν: "« : r1  t1  ?g  ?f»"
          using f g π.the_θ_props μ.the_θ_props
          by (intro comp_in_homI hcomp_in_vhom, auto)
        interpret gf: arrow_of_spans_of_maps V H 𝖺 𝗂 src trg r0 r1 t0 t1 ?g  ?f
        proof
          show "ide (?g  ?f)" by simp
          show "θ. «θ : t0  ?g  ?f  r0»"
            using θ by blast
          show "ν. «ν : r1  t1  ?g  ?f»  iso ν"
            using ν μ.the_ν_props μ.the_θ_props π.the_ν_props π.the_θ_props
                  isos_compose [of "μ.the_ν" "π.the_ν"] μ.is_ide ν ide (π.chine  μ.chine)
                  π.uwθ π.w_simps(4) τ.ide_leg1 τ.leg1_simps(3) arrI hseq_char ideD(1)
                  ide_is_iso iso_assoc iso_hcomp isos_compose seqE
            by metis
        qed
        show ?thesis
        proof (intro conjI)
          have θ_eq: " = gf.the_θ"
            using θ gf.the_θ_props gf.leg0_uniquely_isomorphic by auto
          have ν_eq: " = gf.the_ν"
            using ν gf.the_ν_props gf.leg1_uniquely_isomorphic by auto
          have A: "src ?g = trg ?f"
            using f g by fastforce
          show "arrow_of_spans_of_maps V H 𝖺 𝗂 src trg r0 (dom Δ) t0 t1 (?g  ?f)"
            using gf.arrow_of_spans_of_maps_axioms by simp
          have "((t  gf.the_θ)  𝖺[t, t0, ?g  ?f]  (τ  ?g  ?f))  gf.the_ν = Δ"
          proof -
            have "Δ = (π  r0)  (μ  r0)  ρ"
              using whisker_right comp_assoc
              by (metis Δ_simps(1) hseqE ide_u seqE)
            also have "... = ((π  r0)  (s  μ.the_θ))  𝖺[s, s0, ?f]  (σ  ?f)  μ.the_ν"
              using μ.Δ_naturality comp_assoc by simp
            also have "... = (t  μ.the_θ)  ((π  s0  ?f)  𝖺[s, s0, ?f])  (σ  ?f)  μ.the_ν"
            proof -
              have "(π  r0)  (s  μ.the_θ) = π  μ.the_θ"
                using f comp_arr_dom comp_cod_arr μ.the_θ_props π.in_hom
                      interchange [of π s r0 μ.the_θ]
                by (metis in_homE)
              also have "... = (t  μ.the_θ)  (π  s0  ?f)"
                using f comp_arr_dom comp_cod_arr μ.the_θ_props π.in_hom
                      interchange [of t π μ.the_θ "s0  ?f"]
                by (metis in_homE)
              finally have "(π  r0)  (s  μ.the_θ) = (t  μ.the_θ)  (π  s0  ?f)"
                by simp
              thus ?thesis
                using comp_assoc by presburger
            qed
            also have "... = (t  μ.the_θ)  𝖺[t, s0, ?f]  (((π  s0)  ?f)  (σ  ?f))  μ.the_ν"
            proof -
              have "(π  s0  ?f)  𝖺[s, s0, ?f] = 𝖺[t, s0, ?f]  ((π  s0)  ?f)"
                using f assoc_naturality [of π s0 ?f] π.in_hom by auto
              thus ?thesis
                using comp_assoc by presburger
            qed
            also have "... = (t  μ.the_θ)  𝖺[t, s0, ?f]  (π.Δ  ?f)  μ.the_ν"
              using whisker_right comp_assoc by simp
            also have "... = (t  μ.the_θ)  𝖺[t, s0, ?f] 
                               ((t  π.the_θ)  𝖺[t, t0, ?g]  (τ  ?g)  π.the_ν  ?f)  μ.the_ν"
              using π.Δ_naturality by simp
            also have "... = (t  μ.the_θ)  𝖺[t, s0, ?f] 
                               (((t  π.the_θ)  ?f)  (𝖺[t, t0, ?g]  ?f)  ((τ  ?g)  ?f) 
                               (π.the_ν  ?f))  μ.the_ν"
              using f g π.the_θ_props π.the_ν_props whisker_right
              by (metis π.Δ_simps(1) π.Δ_naturality seqE)
            also have "... = (t  μ.the_θ)  (𝖺[t, s0, ?f] 
                               ((t  π.the_θ)  ?f))  (𝖺[t, t0, ?g]  ?f)  ((τ  ?g)  ?f) 
                               (π.the_ν  ?f)  μ.the_ν"
              using comp_assoc by presburger
            also have "... = (t  μ.the_θ)  (t  π.the_θ  ?f) 
                               (𝖺[t, t0  ?g, ?f]  (𝖺[t, t0, ?g]  ?f)) 
                               ((τ  ?g)  ?f)  (π.the_ν  ?f)  μ.the_ν"
              using f π.the_θ_props assoc_naturality [of t "π.the_θ" ?f] π.θ_simps(3) comp_assoc
              by auto
            also have "... = (t  μ.the_θ)  (t  π.the_θ  ?f) 
                               (t  𝖺-1[t0, ?g, ?f])  𝖺[t, t0, ?g  ?f]  (𝖺[t  t0, ?g, ?f] 
                               ((τ  ?g)  ?f))  (π.the_ν  ?f)  μ.the_ν"
            proof -
              have "seq 𝖺[t, t0, ?g  ?f] 𝖺[t  t0, ?g, ?f]"
                using f g by fastforce
              moreover have "inv (t  𝖺[t0, ?g, ?f]) = t  𝖺-1[t0, ?g, ?f]"
                using f g by simp
              moreover have "iso (t  𝖺[t0, ?g, ?f])"
                using f g by simp
              have "𝖺[t, t0  ?g, ?f]  (𝖺[t, t0, ?g]  ?f) =
                      (t  𝖺-1[t0, ?g, ?f])  𝖺[t, t0, ?g  ?f]  𝖺[t  t0, ?g, ?f]"
              proof -
                have "seq 𝖺[t, t0, ?g  ?f] 𝖺[t  t0, ?g, ?f]"
                  using f g by fastforce
                moreover have "inv (t  𝖺[t0, ?g, ?f]) = t  𝖺-1[t0, ?g, ?f]"
                  using f g by simp
                moreover have "iso (t  𝖺[t0, ?g, ?f])"
                  using f g by simp
                ultimately show ?thesis
                  using A f g pentagon invert_side_of_triangle(1)
                  by (metis π.w_simps(4) τ.ide_base τ.ide_leg0 τ.leg0_simps(3))
              qed
              thus ?thesis
                using comp_assoc by presburger
            qed
            also have "... = ((t  μ.the_θ)  (t  π.the_θ  ?f)) 
                               (t  𝖺-1[t0, ?g, ?f])  𝖺[t, t0, ?g  ?f]  (τ  ?g  ?f)  
                               𝖺[t1, ?g, ?f]  (π.the_ν  ?f)  μ.the_ν"
              using f g assoc_naturality [of τ ?g ?f] comp_assoc by simp
            also have "... = (t  μ.the_θ  (π.the_θ  ?f)  𝖺-1[t0, ?g, ?f]) 
                               𝖺[t, t0, ?g  ?f] 
                               (τ  ?g  ?f)  𝖺[t1, ?g, ?f]  (π.the_ν  ?f)  μ.the_ν"
            proof -
              have 1: "seq μ.the_θ ((π.the_θ  ?f)  𝖺-1[t0, ?g, ?f])"
                using θ_eq by auto
              hence "t  (π.the_θ  ?f)  𝖺-1[t0, ?g, ?f] =
                     (t  π.the_θ  ?f)  (t  𝖺-1[t0, ?g, ?f])"
                using whisker_left τ.ide_base by blast
              thus ?thesis
                using 1 whisker_left τ.ide_base comp_assoc by presburger
            qed
            also have "... = ((t  gf.the_θ)  𝖺[t, t0, ?g  ?f]  (τ  ?g  ?f))  gf.the_ν"
              using θ_eq ν_eq by (simp add: comp_assoc)
            finally show ?thesis
              using comp_assoc by presburger
          qed
          thus "((t  gf.the_θ)  𝖺[t, t0, ?g  ?f]  (τ  ?g  ?f)) 
                  arrow_of_spans_of_maps.the_ν (⋅) (⋆) (dom ((π  μ  r0)  ρ)) t1 (?g  ?f) =
                Δ"
            by simp
        qed
      qed
      thus ?thesis
        using chine_is_induced_map induced_map_unique by simp
    qed

  end

  sublocale vertical_composite_of_arrows_of_tabulations_in_maps 
            arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg r ρ r0 r1 t τ t0 t1 "π  μ"
    using is_arrow_of_tabulations_in_maps by simp

  subsubsection "Horizontal Composite"

  locale horizontal_composite_of_arrows_of_tabulations_in_maps =
    bicategory_of_spans V H 𝖺 𝗂 src trg +
    ρ: tabulation_in_maps V H 𝖺 𝗂 src trg r ρ r0 r1 +
    σ: tabulation_in_maps V H 𝖺 𝗂 src trg s σ s0 s1 +
    τ: tabulation_in_maps V H 𝖺 𝗂 src trg t τ t0 t1 +
    μ: tabulation_in_maps V H 𝖺 𝗂 src trg u μ u0 u1 +
    ρσ: composite_tabulation_in_maps V H 𝖺 𝗂 src trg r ρ r0 r1 s σ s0 s1 +
    τμ: composite_tabulation_in_maps V H 𝖺 𝗂 src trg t τ t0 t1 u μ u0 u1 +
    ω: arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg r ρ r0 r1 t τ t0 t1 ω +
    χ: arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg s σ s0 s1 u μ u0 u1 χ
  for V :: "'a comp"                 (infixr "" 55)
  and H :: "'a  'a  'a"          (infixr "" 53)
  and 𝖺 :: "'a  'a  'a  'a"     ("𝖺[_, _, _]")
  and 𝗂 :: "'a  'a"                 ("𝗂[_]")
  and src :: "'a  'a"
  and trg :: "'a  'a"
  and r :: 'a
  and ρ :: 'a
  and r0 :: 'a
  and r1 :: 'a
  and s :: 'a
  and σ :: 'a
  and s0 :: 'a
  and s1 :: 'a
  and t :: 'a
  and τ :: 'a
  and t0 :: 'a
  and t1 :: 'a
  and u :: 'a
  and μ :: 'a
  and u0 :: 'a
  and u1 :: 'a
  and ω :: 'a
  and χ :: 'a
  begin

    text ‹
$$
\xymatrix{
  &&& {\rm src}~t_0u_1.\phi \ar[dl]_{\tau\mu.p_1} \ar[dr]^{\tau\mu.p_0} \ddtwocell\omit{^{t_0u_1.\phi}} \\
  && {\rm src}~\tau \ar[dl]_{t_1} \ar[dr]^<>(0.4){t_0} \dtwocell\omit{^<-1>\tau}
  && {\rm src}~\mu \ar[dl]_{u_1} \ar[dr]^{u_0} \dtwocell\omit{^<-1>\mu} & \\
  & {\rm trg}~t && {\rm src}~t = {\rm trg}~u \ar[ll]^{t}
  && {\rm src}~u \ar[ll]^{u} \\
  & \xtwocell[r]{}\omit{^\omega}
  & {\rm src}~r_0s_1.\phi \ar[uuur]_<>(0.2){{\rm chine}}
    \ar[dl]^{\rho\sigma.p_1} \ar[dr]_{\rho\sigma.p_0\hspace{20pt}} \ddtwocell\omit{^{r_0s_1.\phi}}
  & \rrtwocell\omit{^\chi} && \\
  & {\rm src}~\rho \ar[dl]_{r_1} \ar[dr]^{r_0} \ar[uuur]^<>(0.4){\omega.{\rm chine}} \dtwocell\omit{^\rho}
  && {\rm src}~\sigma \ar[dl]_{s_1} \ar[dr]^{s_0} \ar[uuur]^<>(0.4){\chi.{\rm chine}} \dtwocell\omit{^<-1>\sigma} & \\
  {\rm trg}~r \ar@ {=}[uuur] && {\rm src}~r = {\rm trg}~s \ar[ll]^{r} \ar@ {=}[uuur]
  && {\rm src}~s \ar[ll]^{s} \ar@ {=}[uuur] \\
}
$$
    ›

    notation isomorphic (infix "" 50)

    interpretation arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg
                     r  s ρσ.tab s0  ρσ.p0 r1  ρσ.p1
                     t  u τμ.tab u0  τμ.p0 t1  τμ.p1 ω  χ
      using ρσ.composable ω.in_hom χ.in_hom
      by unfold_locales auto

    lemma is_arrow_of_tabulations_in_maps:
    shows "arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg
             (r  s) ρσ.tab (s0  ρσ.p0) (r1  ρσ.p1)
             (t  u) τμ.tab (u0  τμ.p0) (t1  τμ.p1) (ω  χ)"
      ..

    sublocale arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg
                r  s ρσ.tab s0  ρσ.p0 r1  ρσ.p1
                t  u τμ.tab u0  τμ.p0 t1  τμ.p1 ω  χ
      using is_arrow_of_tabulations_in_maps by simp

    interpretation Maps: maps_category V H 𝖺 𝗂 src trg ..
    notation Maps.comp  (infixr "" 55)

    interpretation r0s1: cospan_of_maps_in_bicategory_of_spans (⋅) (⋆) 𝖺 𝗂 src trg s1 r0
      using ρ.leg0_is_map σ.leg1_is_map ρσ.composable apply unfold_locales by auto
    interpretation r0s1: arrow_of_tabulations_in_maps (⋅) (⋆) 𝖺 𝗂 src trg
                           r0*  s1 r0s1.tab r0s1.p0 r0s1.p1
                           r0*  s1 r0s1.tab r0s1.p0 r0s1.p1
                           r0*  s1
      using r0s1.is_arrow_of_tabulations_in_maps by simp
    interpretation t0u1: cospan_of_maps_in_bicategory_of_spans (⋅) (⋆) 𝖺 𝗂 src trg u1 t0
      using τ.leg0_is_map μ.leg1_is_map τμ.composable apply unfold_locales by auto
    interpretation t0u1: arrow_of_tabulations_in_maps (⋅) (⋆) 𝖺 𝗂 src trg
                           t0*  u1 t0u1.tab t0u1.p0 t0u1.p1
                           t0*  u1 t0u1.tab t0u1.p0 t0u1.p1
                           t0*  u1
      using t0u1.is_arrow_of_tabulations_in_maps by simp

    interpretation E: self_evaluation_map V H 𝖺 𝗂 src trg ..
    notation E.eval ("_")

    no_notation in_hom ("«_ : _  _»")

    text ‹
      The following lemma states that the rectangular faces of the ``top prism'' commute
      up to isomorphism.  This was not already proved in @{locale composite_tabulation_in_maps},
      because there we did not consider any composite structure of the ``source'' 2-cell.
      There are common elements, though to the proof that the composite of tabulations is
      a tabulation and the present lemma.
      The proof idea is to use property T2› of the ``base'' tabulations to establish the
      existence of the desired isomorphisms.  The proofs have to be carried out in
      sequence, starting from the ``output'' side, because the arrow β› required in the
      hypotheses of T2› depends, for the ``input'' tabulation, on the isomorphism constructed
      for the ``output'' tabulation.
    ›

    lemma prj_chine:
    shows "τμ.p0  chine  χ.chine  ρσ.p0"
    and "τμ.p1  chine  ω.chine  ρσ.p1"
    proof -
      have 1: "arrow_of_spans_of_maps V H 𝖺 𝗂 src trg
                 (s0  ρσ.p0) (r1  ρσ.p1) (u0  τμ.p0) (t1  τμ.p1) chine 
               (((t  u)  the_θ)  𝖺[t  u, u0  τμ.p0, chine]  (τμ.tab  chine))  the_ν =
               ((ω  χ)  s0  ρσ.p0)  ρσ.tab"
        using chine_is_induced_map by simp
      let ?uτ = "u  s0  ρσ.p0"
      let ?wτ = "ω.chine  ρσ.p1"
      let ?wτ' = "τμ.p1  chine"
      have uτ: "ide ?uτ"
        using χ.u_simps(3) by auto
      have wτ: "ide ?wτ  is_left_adjoint ?wτ"
        by (simp add: ω.is_map ρ.T0.antipar(1) left_adjoints_compose)
      have wτ': "ide ?wτ'  is_left_adjoint ?wτ'"
        by (simp add: is_map left_adjoints_compose)
      let τ = "𝖺[u, s0, ρσ.p0]  ((χ  s0)  σ  ρσ.p0)  r0s1  (ω.the_θ  ρσ.p1) 
                   𝖺-1[t0, ω.chine, ρσ.p1]"
      let τ' = "(u  the_θ)  𝖺[u, u0  τμ.p0, chine]  (𝖺[u, u0, τμ.p0]  chine) 
                   ((μ  τμ.p0)  chine)  (t0u1  chine)  𝖺-1[t0, τμ.p1, chine]"
      let τ = "𝖺[t1, τμ.p1, chine]  the_ν  (inv ω.the_ν  ρσ.p1)  𝖺-1[t1, ω.chine, ρσ.p1]"
      have θτ: "«τ : t0  ?wτ  ?uτ»"
        using ρ.T0.antipar(1) ω.the_θ_in_hom χ.u_simps(3)
        by (intro comp_in_homI, auto)
      have θτ': "«τ' : t0  ?wτ'  ?uτ»"
      proof (intro comp_in_homI)
        show "«𝖺-1[t0, τμ.p1, chine] : t0  τμ.p1  chine  (t0  τμ.p1)  chine»"
          using t0u1.p1_simps assoc'_in_hom by simp
        show "«t0u1  chine : (t0  τμ.p1)  chine  (u1  τμ.p0)  chine»"
          using τ.T0.antipar(1)
          by (intro hcomp_in_vhom, auto)
        show "«(μ  τμ.p0)  chine : (u1  τμ.p0)  chine  ((u  u0)  τμ.p0)  chine»"
          by (intro hcomp_in_vhom, auto)
        show "«𝖺[u, u0, τμ.p0]  chine : ((u  u0)  τμ.p0)  chine  (u  u0  τμ.p0)  chine»"
          using assoc_in_hom by auto
        show "«𝖺[u, u0  τμ.p0, chine] : (u  u0  τμ.p0)  chine  u  (u0  τμ.p0)  chine»"
          by auto
        show "«u  the_θ : u  (u0  τμ.p0)  chine  ?uτ»"
          by (intro hcomp_in_vhom, auto)
      qed
      have βτ: "«τ : t1  ?wτ  t1  ?wτ'»"
      proof (intro comp_in_homI)
        show "«𝖺-1[t1, ω.chine, ρσ.p1] : t1  ?wτ  (t1  ω.chine)  ρσ.p1»"
          using ρ.T0.antipar(1) by auto
        show "«inv ω.the_ν  ρσ.p1 : (t1  ω.chine)  ρσ.p1  r1  ρσ.p1»"
          using ω.the_ν_props ρ.T0.antipar(1)
          by (intro hcomp_in_vhom, auto)
        show "«the_ν : r1  ρσ.p1  (t1  τμ.p1)  chine»"
          using the_ν_in_hom(2) by simp
        show "«𝖺[t1, τμ.p1, chine] : (t1  τμ.p1)  chine  t1  ?wτ'»"
          using t0u1.p1_simps assoc_in_hom by simp
      qed
      define LHS where "LHS = (t  τ)  𝖺[t, t0, ?wτ]  (τ  ?wτ)"
      have LHS: "«LHS : t1  ?wτ  t  ?uτ»"
      proof (unfold LHS_def, intro comp_in_homI)
        show "«τ  ?wτ : t1  ω.chine  ρσ.p1  (t  t0)  ?wτ»"
          using ρ.T0.antipar(1)
          by (intro hcomp_in_vhom, auto)
        show "«𝖺[t, t0, ?wτ] : (t  t0)  ?wτ  t  t0  ?wτ»"
          using ρ.T0.antipar(1) by auto
        show "«t  τ : t  t0  ?wτ  t  ?uτ»"
        proof -
          have "src t = trg (t0  ω.chine  r0s1.p1)"
            by (metis χ.u_simps(3) μ.ide_base σ.ide_leg0 σ.leg1_simps(3) τμ.composable
                θτ arrI assoc_simps(3) r0s1.ide_u r0s1.p0_simps trg_vcomp vconn_implies_hpar(2))
          thus ?thesis
            using θτ by blast
        qed
      qed
      define RHS where "RHS = ((t  τ')  𝖺[t, t0, ?wτ']  (τ  ?wτ'))  τ"
      have RHS: "«RHS : t1  ?wτ  t  ?uτ»"
        unfolding RHS_def
      proof
        show "«τ : t1  ?wτ  t1  ?wτ'»"
          using βτ by simp
        show "«(t  τ')  𝖺[t, t0, ?wτ']  (τ  ?wτ') : t1  ?wτ'  t  ?uτ»"
        proof
          show "«𝖺[t, t0, ?wτ']  (τ  ?wτ') : t1  ?wτ'  t  t0  ?wτ'»"
            using τ.T0.antipar(1) by fastforce
          show "«t  τ' : t  t0  ?wτ'  t  ?uτ»"
            using wτ' θτ' τ.leg0_simps(2) τ.leg0_simps(3) hseqI' ideD(1) t0u1.p1_simps
                  trg_hcomp τ.base_in_hom(2) hcomp_in_vhom
            by presburger
        qed
      qed
      have eq: "LHS = RHS"
      proof -
        have "𝖺-1[t, u, s0  ρσ.p0]  LHS  𝖺[t1, ω.chine, ρσ.p1]  (ω.the_ν  ρσ.p1) = Δ"
        proof -
          text ‹
            Here we use ω.Δ_naturality› to replace @{term ω.chine}
            in favor of @{term ω}.
            We have to bring @{term ω.the_ν}, @{term τ}, and @{term ω.the_θ} together, 
            with @{term ρσ.p1} on the right.
          ›
          have "𝖺-1[t, u, s0  ρσ.p0]  LHS  𝖺[t1, ω.chine, ρσ.p1]  (ω.the_ν  ρσ.p1) =
                𝖺-1[t, u, s0  ρσ.p0] 
                  (t  𝖺[u, s0, ρσ.p0]  ((χ  s0)  σ  ρσ.p0)  r0s1 
                       (ω.the_θ  ρσ.p1)  𝖺-1[t0, ω.chine, ρσ.p1]) 
                  𝖺[t, t0, ω.chine  ρσ.p1]  (τ  ω.chine  ρσ.p1)  𝖺[t1, ω.chine, ρσ.p1] 
                  (ω.the_ν  ρσ.p1)"
            unfolding LHS_def
            using comp_assoc by presburger
          also have "... = 𝖺-1[t, u, s0  ρσ.p0]  (t  𝖺[u, s0, ρσ.p0]) 
                             (t  (χ  s0)  σ  ρσ.p0)  (t  r0s1) 
                             (t  ω.the_θ  ρσ.p1)  (t  𝖺-1[t0, ω.chine, ρσ.p1]) 
                             𝖺[t, t0, ω.chine  ρσ.p1]  ((τ  ω.chine  ρσ.p1) 
                             𝖺[t1, ω.chine, ρσ.p1])  (ω.the_ν  ρσ.p1)"
          proof -
            have "t  𝖺[u, s0, ρσ.p0]  ((χ  s0)  σ  ρσ.p0)  r0s1 
                       (ω.the_θ  ρσ.p1)  𝖺-1[t0, ω.chine, ρσ.p1] =
                  (t  𝖺[u, s0, ρσ.p0])  (t  (χ  s0)  σ  ρσ.p0)  (t  r0s1) 
                       (t  ω.the_θ  ρσ.p1)  (t  𝖺-1[t0, ω.chine, ρσ.p1])"
              using whisker_left τ.ide_base θτ arrI seqE
              by (metis (full_types))
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... = 𝖺-1[t, u, s0  ρσ.p0]  (t  𝖺[u, s0, ρσ.p0]) 
                             (t  (χ  s0)  σ  ρσ.p0)  (t  r0s1) 
                             (t  ω.the_θ  ρσ.p1)  ((t  𝖺-1[t0, ω.chine, ρσ.p1]) 
                             𝖺[t, t0, ω.chine  ρσ.p1]  𝖺[t  t0, ω.chine, ρσ.p1]) 
                             ((τ  ω.chine)  ρσ.p1)  (ω.the_ν  ρσ.p1)"
          proof -
            have "(τ  ω.chine  ρσ.p1)  𝖺[t1, ω.chine, ρσ.p1] =
                  𝖺[t  t0, ω.chine, ρσ.p1]  ((τ  ω.chine)  ρσ.p1)"
              using assoc_naturality
              by (metis ω.w_simps(2-6) ρ.leg1_simps(3) ρσ.leg1_simps(2) τ.tab_simps(1)
                  τ.tab_simps(2,4-5) hseqE r0s1.leg1_simps(5) r0s1.leg1_simps(6))
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... = 𝖺-1[t, u, s0  ρσ.p0]  (t  𝖺[u, s0, ρσ.p0]) 
                             (t  (χ  s0)  σ  ρσ.p0)  (t  r0s1) 
                             ((t  ω.the_θ  ρσ.p1)  𝖺[t, t0  ω.chine, ρσ.p1]) 
                             (𝖺[t, t0, ω.chine]  ρσ.p1)  ((τ  ω.chine)  ρσ.p1) 
                             (ω.the_ν  ρσ.p1)"
          proof -
            have "(t  𝖺-1[t0, ω.chine, ρσ.p1])  𝖺[t, t0, ω.chine  ρσ.p1] 
                    𝖺[t  t0, ω.chine, ρσ.p1] =
                  𝖺[t, t0  ω.chine, ρσ.p1]  (𝖺[t, t0, ω.chine]  ρσ.p1)"
            proof -
              have "seq 𝖺[t, t0, ω.chine  ρσ.p1] 𝖺[t  t0, ω.chine, ρσ.p1]"
                by (simp add: ρ.T0.antipar(1))
              moreover have "inv (t  𝖺[t0, ω.chine, ρσ.p1]) = t  𝖺-1[t0, ω.chine, ρσ.p1]"
                using ρ.T0.antipar(1) by simp
              ultimately show ?thesis
                using pentagon ρ.T0.antipar(1) iso_hcomp
                      invert_side_of_triangle(1)
                        [of "𝖺[t, t0, ω.chine  ρσ.p1]  𝖺[t  t0, ω.chine, ρσ.p1]"
                            "t  𝖺[t0, ω.chine, ρσ.p1]"
                            "𝖺[t, t0  ω.chine, ρσ.p1]  (𝖺[t, t0, ω.chine]  ρσ.p1)"]
                by simp
            qed
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... = 𝖺-1[t, u, s0  ρσ.p0]  (t  𝖺[u, s0, ρσ.p0]) 
                             (t  (χ  s0)  σ  ρσ.p0)  (t  r0s1) 
                             𝖺[t, r0, ρσ.p1]  (((t  ω.the_θ)  ρσ.p1) 
                             (𝖺[t, t0, ω.chine]  ρσ.p1)  ((τ  ω.chine)  ρσ.p1)) 
                             (ω.the_ν  ρσ.p1)"
          proof -
            have "(t  ω.the_θ  ρσ.p1)  𝖺[t, t0  ω.chine, ρσ.p1] =
                  𝖺[t, r0, ρσ.p1]  ((t  ω.the_θ)  ρσ.p1)"
              using assoc_naturality [of t ω.the_θ ρσ.p1] ω.θ_simps(3) ρσ.leg1_simps(2) hseq_char
              by auto
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... = 𝖺-1[t, u, s0  ρσ.p0]  (t  𝖺[u, s0, ρσ.p0]) 
                             (t  (χ  s0)  σ  ρσ.p0)  (t  r0s1) 
                             𝖺[t, r0, ρσ.p1]  ((ω  r0)  ρ  ρσ.p1)"
            using whisker_right ρ.T0.antipar(1) ω.Δ_simps(1) ω.Δ_naturality comp_assoc
            by fastforce
          also have "... = 𝖺-1[t, u, s0  ρσ.p0]  ((t  𝖺[u, s0, ρσ.p0]) 
                             (t  (χ  s0)  ρσ.p0))  (t  σ  ρσ.p0)  (t  r0s1) 
                             𝖺[t, r0, ρσ.p1]  ((ω  r0)  ρ  ρσ.p1)"
          proof -
            have "t  (χ  s0)  σ  ρσ.p0 = (t  (χ  s0)  ρσ.p0)  (t  σ  ρσ.p0)"
              using whisker_left whisker_right ρ.T0.antipar(1)
              by (metis (full_types) χ.Δ_simps(1) τ.ide_base θτ arrI r0s1.ide_u seqE)
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... = 𝖺-1[t, u, s0  ρσ.p0]  (t  χ  s0  ρσ.p0) 
                             (t  𝖺[s, s0, ρσ.p0])  (t  σ  ρσ.p0)  (t  r0s1) 
                             𝖺[t, r0, ρσ.p1]  ((ω  r0)  ρ  ρσ.p1)"
          proof -
            have "(t  𝖺[u, s0, ρσ.p0])  (t  (χ  s0)  ρσ.p0) =
                  t  𝖺[u, s0, ρσ.p0]  ((χ  s0)  ρσ.p0)"
              using χ.in_hom whisker_left by auto
            also have "... = t  (χ  s0  ρσ.p0)  𝖺[s, s0, ρσ.p0]"
              using assoc_naturality [of χ s0 ρσ.p0] χ.in_hom by auto
            also have "... = (t  χ  s0  ρσ.p0)  (t  𝖺[s, s0, ρσ.p0])"
            proof -
              have "seq (χ  s0  ρσ.p0) 𝖺[s, s0, ρσ.p0]"
                using χ.in_hom
                apply (intro seqI hseqI)
                   apply auto
              proof -
                show "«χ : src u  trg χ»"
                  by (metis χ.Δ_simps(1) χ.u_simps(3) hseqE in_hhom_def seqE)
                show "dom (χ  s0  ρσ.p0) = s  s0  ρσ.p0"
                  by (metis Δ_simps(1) χ.in_hom hcomp_simps(1,3) hseq_char in_homE seqE
                      u_simps(4))
              qed
              thus ?thesis
                using whisker_left by simp
            qed
            finally show ?thesis
              using comp_assoc by presburger
          qed
          also have "... = 𝖺-1[t, u, s0  ρσ.p0]  (t  χ  s0  ρσ.p0) 
                             (t  𝖺[s, s0, ρσ.p0])  (t  σ  ρσ.p0)  (t  r0s1) 
                             (𝖺[t, r0, ρσ.p1]  ((ω  r0)  ρσ.p1))  (ρ  ρσ.p1)"
            using whisker_right comp_assoc by simp
          also have "... = 𝖺-1[t, u, s0  ρσ.p0]  (t  χ  s0  ρσ.p0) 
                             (t  𝖺[s, s0, ρσ.p0])  (t  σ  ρσ.p0)  ((t  r0s1) 
                             (ω  r0  ρσ.p1))  𝖺[r, r0, ρσ.p1]  (ρ  ρσ.p1)"
            using assoc_naturality [of ω r0 ρσ.p1] ω.in_hom ρ.T0.antipar(1) comp_assoc
            by fastforce
          also have "... = 𝖺-1[t, u, s0  ρσ.p0]  ((t  χ  s0  ρσ.p0) 
                             (t  𝖺[s, s0, ρσ.p0])  (t  σ  ρσ.p0))  (ω  s1  ρσ.p0) 
                             (r  r0s1)  𝖺[r, r0, ρσ.p1]  (ρ  ρσ.p1)"
          proof -
            have "(t  r0s1)  (ω  r0  ρσ.p1) = ω  r0s1"
              using comp_cod_arr comp_arr_dom ω.in_hom interchange comp_ide_arr
              by (metis τ.base_in_hom(2) τ.ide_base r0s1.φ_simps(1) r0s1.φ_simps(4) seqI')
            also have "... = (ω  s1  ρσ.p0)  (r  r0s1)"
              using r0s1.φ_in_hom comp_cod_arr comp_arr_dom ω.in_hom interchange
              by (metis in_homE)
            finally have "(t  r0s1)  (ω  r0  ρσ.p1) = (ω  s1  ρσ.p0)  (r  r0s1)"
              by simp
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... = 𝖺-1[t, u, s0  ρσ.p0] 
                             ((t  (χ  s0  ρσ.p0)  𝖺[s, s0, ρσ.p0]  (σ  ρσ.p0)) 
                             (ω  s1  ρσ.p0)) 
                             (r  r0s1)  𝖺[r, r0, ρσ.p1]  (ρ  ρσ.p1)"
            using whisker_left ρ.T0.antipar(1) ρσ.composable χ.in_hom comp_assoc by auto
          also have "... = 𝖺-1[t, u, s0  ρσ.p0] 
                             (ω  (χ  s0  ρσ.p0)  𝖺[s, s0, ρσ.p0]  (σ  ρσ.p0)) 
                             (r  r0s1)  𝖺[r, r0, ρσ.p1]  (ρ  ρσ.p1)"
          proof -
            have "(t  (χ  s0  ρσ.p0)  𝖺[s, s0, ρσ.p0]  (σ  ρσ.p0))  (ω  s1  ρσ.p0) =
                   ω  (χ  s0  ρσ.p0)  𝖺[s, s0, ρσ.p0]  (σ  ρσ.p0)"
            proof -
              have "«(χ  s0  ρσ.p0)  𝖺[s, s0, ρσ.p0]  (σ  ρσ.p0) : s1  ρσ.p0  u  s0  ρσ.p0»"
                using ω.in_hom χ.in_hom by force
              thus ?thesis
                  by (metis (no_types) ω.in_hom comp_arr_dom comp_cod_arr in_homE
                      interchange)
            qed
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... = (𝖺-1[t, u, s0  ρσ.p0] 
                             (ω  χ  s0  ρσ.p0))  (r  𝖺[s, s0, ρσ.p0]  (σ  ρσ.p0)) 
                             (r  r0s1)  𝖺[r, r0, ρσ.p1]  (ρ  ρσ.p1)"
          proof -
            have "ω  (χ  s0  ρσ.p0)  𝖺[s, s0, ρσ.p0]  (σ  ρσ.p0) =
                  (ω  χ  s0  ρσ.p0)  (r  𝖺[s, s0, ρσ.p0]  (σ  ρσ.p0))"
            proof -
              have "seq (χ  s0  ρσ.p0) (𝖺[s, s0, ρσ.p0]  (σ  ρσ.p0))"
                using χ.in_hom by force
              thus ?thesis
               using comp_arr_dom comp_cod_arr ω.in_hom χ.in_hom interchange
               by (metis in_homE)
            qed
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... = ((ω  χ)  s0  ρσ.p0) 
                             𝖺-1[r, s, s0  ρσ.p0]  (r  𝖺[s, s0, ρσ.p0]  (σ  ρσ.p0)) 
                             (r  r0s1)  𝖺[r, r0, ρσ.p1]  (ρ  ρσ.p1)"
          proof -
            have "𝖺-1[t, u, s0  ρσ.p0]  (ω  χ  s0  ρσ.p0) =
                  ((ω  χ)  s0  ρσ.p0)  𝖺-1[r, s, s0  ρσ.p0]"
              using assoc_naturality ω.in_hom χ.in_hom
              by (metis ρσ.leg0_simps(3) assoc'_naturality hcomp_in_vhomE in_hom in_homE
                  u_simps(2) u_simps(4) u_simps(5))
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... = Δ"
            using whisker_left ρσ.tab_def comp_assoc by simp
          finally show ?thesis by auto
        qed
        also have "... = 𝖺-1[t, u, s0  ρσ.p0]  RHS  𝖺[t1, ω.chine, ρσ.p1]  (ω.the_ν  ρσ.p1)"
        proof -
          text ‹Now cancel @{term ω.the_ν} and its inverse.›
          have "𝖺-1[t, u, s0  ρσ.p0]  RHS  𝖺[t1, ω.chine, ρσ.p1]  (ω.the_ν  ρσ.p1) =
                𝖺-1[t, u, s0  ρσ.p0] 
                  (t  (u  the_θ)  𝖺[u, u0  τμ.p0, chine] 
                         (𝖺[u, u0, τμ.p0]  chine)  ((μ  τμ.p0)  chine) 
                         (t0u1  chine)  𝖺-1[t0, τμ.p1, chine]) 
                    𝖺[t, t0, τμ.p1  chine]  (τ  τμ.p1  chine) 
                    𝖺[t1, τμ.p1, chine]  the_ν  (inv ω.the_ν  ρσ.p1) 
                    ((𝖺-1[t1, ω.chine, ρσ.p1])  𝖺[t1, ω.chine, ρσ.p1])  (ω.the_ν  ρσ.p1)"
            unfolding RHS_def
            using comp_assoc by presburger
          also have "... = 𝖺-1[t, u, s0  ρσ.p0] 
                             (t  (u  the_θ)  𝖺[u, u0  τμ.p0, chine] 
                                     (𝖺[u, u0, τμ.p0]  chine)  ((μ  τμ.p0)  chine) 
                                     (t0u1  chine)  𝖺-1[t0, τμ.p1, chine]) 
                               𝖺[t, t0, τμ.p1  chine]  (τ  τμ.p1  chine) 
                               𝖺[t1, τμ.p1, chine]  the_ν"
          proof -
            have "the_ν  (inv ω.the_ν  ρσ.p1) 
                    ((𝖺-1[t1, ω.chine, ρσ.p1])  𝖺[t1, ω.chine, ρσ.p1])  (ω.the_ν  ρσ.p1) =
                  the_ν  (inv ω.the_ν  ρσ.p1) 
                    ((t1  ω.chine)  ρσ.p1)  (ω.the_ν  ρσ.p1)"
              using comp_inv_arr ρ.T0.antipar(1) comp_assoc_assoc' by simp
            also have "... = the_ν  (inv ω.the_ν  ρσ.p1)  (ω.the_ν  ρσ.p1)"
              using comp_cod_arr ρ.T0.antipar(1) by simp
            also have "... = the_ν  (r1  ρσ.p1)"
              using whisker_right [of ρσ.p1] r0s1.ide_leg1 ω.the_ν_props(2) ω.the_ν_simps(4)
                    ρ.leg1_simps(2) comp_inv_arr'
              by metis
            also have "... = the_ν"
              using comp_arr_dom by simp
            finally show ?thesis
              using comp_assoc by simp
          qed
          text ‹
            Now reassociate to move @{term the_θ} to the left and get other terms composed
            with @{term chine}, where they can be reduced to @{term τμ.tab}.
          ›
          also have "... = (𝖺-1[t, u, s0  ρσ.p0] 
                             (t  u  the_θ))  (t  𝖺[u, u0  τμ.p0, chine]) 
                             (t  𝖺[u, u0, τμ.p0]  chine)  (t  (μ  τμ.p0)  chine) 
                             (t  t0u1  chine)  (t  𝖺-1[t0, τμ.p1, chine]) 
                             𝖺[t, t0, τμ.p1  chine]  (τ  τμ.p1  chine) 
                             𝖺[t1, τμ.p1, chine]  the_ν"
          proof -
            have "arr ((u  the_θ)  𝖺[u, u0  τμ.p0, chine]  (𝖺[u, u0, τμ.p0]  chine) 
                       ((μ  τμ.p0)  chine)  (t0u1  chine)  𝖺-1[t0, τμ.p1, chine])"
              using θτ' by blast
            moreover have "arr (𝖺[u, u0  τμ.p0, chine]  (𝖺[u, u0, τμ.p0]  chine) 
                                 ((μ  τμ.p0)  chine)  (t0u1  chine)  𝖺-1[t0, τμ.p1, chine])"
              using calculation by blast
            moreover have "arr ((𝖺[u, u0, τμ.p0]  chine) 
                                 ((μ  τμ.p0)  chine)  (t0u1  chine)  𝖺-1[t0, τμ.p1, chine])"
              using calculation by blast
            moreover have "arr (((μ  τμ.p0)  chine)  (t0u1  chine)  𝖺-1[t0, τμ.p1, chine])"
              using calculation by blast
            moreover have "arr ((t0u1  chine)  𝖺-1[t0, τμ.p1, chine])"
              using calculation by blast
            ultimately
            have "t  (u  the_θ)  𝖺[u, u0  τμ.p0, chine]  (𝖺[u, u0, τμ.p0]  chine) 
                    ((μ  τμ.p0)  chine)  (t0u1  chine)  𝖺-1[t0, τμ.p1, chine] =
                  (t  u  the_θ)  (t  𝖺[u, u0  τμ.p0, chine]) 
                    (t  𝖺[u, u0, τμ.p0]  chine)  (t  (μ  τμ.p0)  chine) 
                    (t  t0u1  chine)  (t  𝖺-1[t0, τμ.p1, chine])"
              using whisker_left ρ.T0.antipar(1) τ.ide_base by presburger
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... = ((t  u)  the_θ)  𝖺-1[t, u, (u0  τμ.p0)  chine] 
                             (t  𝖺[u, u0  τμ.p0, chine])  (t  𝖺[u, u0, τμ.p0]  chine) 
                             (t  (μ  τμ.p0)  chine) 
                             (t  t0u1  chine) 
                             (t  𝖺-1[t0, τμ.p1, chine])  𝖺[t, t0, τμ.p1  chine] 
                             ((τ  τμ.p1  chine)  𝖺[t1, τμ.p1, chine])  the_ν"
            using assoc'_naturality [of t u the_θ] τμ.composable θ_simps(3) comp_assoc by auto
          also have "... = ((t  u)  the_θ)  𝖺-1[t, u, (u0  τμ.p0)  chine] 
                             (t  𝖺[u, u0  τμ.p0, chine])  (t  𝖺[u, u0, τμ.p0]  chine) 
                             (t  (μ  τμ.p0)  chine) 
                             (t  t0u1  chine) 
                             ((t  𝖺-1[t0, τμ.p1, chine])  𝖺[t, t0, τμ.p1  chine] 
                             𝖺[t  t0, τμ.p1, chine])  ((τ  τμ.p1)  chine)  the_ν"
          proof -
            have "(τ  τμ.p1  chine)  𝖺[t1, τμ.p1, chine] =
                  𝖺[t  t0, τμ.p1, chine]  ((τ  τμ.p1)  chine)"
              using assoc_naturality
              by (metis τ.leg1_simps(3) τ.tab_simps(1,2,4) τ.tab_simps(5) τμ.leg0_simps(2)
                  τμ.leg1_simps(2) hseqE src_hcomp t0u1.leg1_simps(3,5-6) w_simps(2)
                  w_simps(4-6))
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... = ((t  u)  the_θ)  𝖺-1[t, u, (u0  τμ.p0)  chine] 
                             (t  𝖺[u, u0  τμ.p0, chine])  (t  𝖺[u, u0, τμ.p0]  chine) 
                             (t  (μ  τμ.p0)  chine) 
                             ((t  t0u1  chine)  𝖺[t, t0  τμ.p1, chine]) 
                             (𝖺[t, t0, τμ.p1]  chine)  ((τ  τμ.p1)  chine)  the_ν"
          proof -
            have "(t  𝖺-1[t0, τμ.p1, chine])  𝖺[t, t0, τμ.p1  chine] 
                    𝖺[t  t0, τμ.p1, chine] =
                  𝖺[t, t0  τμ.p1, chine]  (𝖺[t, t0, τμ.p1]  chine)"
              using pentagon t0u1.p1_simps uwθ τ.T0.antipar(1) iso_hcomp
                    comp_assoc_assoc'
                    invert_side_of_triangle(1)
                      [of "𝖺[t, t0, τμ.p1  chine]  𝖺[t  t0, τμ.p1, chine]"
                          "t  𝖺[t0, τμ.p1, chine]"
                          "𝖺[t, t0  τμ.p1, chine]  (𝖺[t, t0, τμ.p1]  chine)"]
              by auto
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... = ((t  u)  the_θ)  𝖺-1[t, u, (u0  τμ.p0)  chine] 
                             (t  𝖺[u, u0  τμ.p0, chine])  (t  𝖺[u, u0, τμ.p0]  chine) 
                             ((t  (μ  τμ.p0)  chine)  𝖺[t, u1  τμ.p0, chine]) 
                             ((t  t0u1)  chine)  
                             (𝖺[t, t0, τμ.p1]  chine)  ((τ  τμ.p1)  chine)  the_ν"
          proof -
            have "(t  t0u1  chine)  𝖺[t, t0  τμ.p1, chine] =
                  𝖺[t, u1  τμ.p0, chine]  ((t  t0u1)  chine)"
              using assoc_naturality [of t t0u1 chine] t0u1.cospan by auto
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... = ((t  u)  the_θ)  𝖺-1[t, u, (u0  τμ.p0)  chine] 
                             (t  𝖺[u, u0  τμ.p0, chine])  (t  𝖺[u, u0, τμ.p0]  chine) 
                             𝖺[t, (u  u0)  τμ.p0, chine] 
                             ((t  μ  τμ.p0)  chine)  ((t  t0u1)  chine)  
                             (𝖺[t, t0, τμ.p1]  chine)  ((τ  τμ.p1)  chine)  the_ν"
          proof -
            have "(t  (μ  τμ.p0)  chine)  𝖺[t, u1  τμ.p0, chine] =
                  𝖺[t, (u  u0)  τμ.p0, chine]  ((t  (μ  τμ.p0))  chine)"
              using assoc_naturality [of t "μ  τμ.p0" chine]
              by (simp add: τμ.composable)
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... = ((t  u)  the_θ)  𝖺-1[t, u, (u0  τμ.p0)  chine] 
                             (t  𝖺[u, u0  τμ.p0, chine])  (t  𝖺[u, u0, τμ.p0]  chine) 
                             𝖺[t, (u  u0)  τμ.p0, chine] 
                             ((t  𝖺-1[u, u0, τμ.p0])  chine)  ((t  𝖺[u, u0, τμ.p0])  chine) 
                             ((t  μ  τμ.p0)  chine)  ((t  t0u1)  chine)  
                             (𝖺[t, t0, τμ.p1]  chine)  ((τ  τμ.p1)  chine)  the_ν"
          proof -
            have "(((t  𝖺-1[u, u0, τμ.p0])  chine)  ((t  𝖺[u, u0, τμ.p0])  chine)) 
                    ((t  μ  τμ.p0)  chine) =
                  ((t  ((u  u0)  τμ.p0))  chine)  ((t  μ  τμ.p0)  chine)"
              using whisker_right whisker_left [of t "𝖺-1[u, u0, τμ.p0]" "𝖺[u, u0, τμ.p0]"]
                    τμ.composable comp_assoc_assoc'
              by simp
            also have "... = (t  μ  τμ.p0)  chine"
              using comp_cod_arr τμ.composable by simp
            finally have "(((t  𝖺-1[u, u0, τμ.p0])  chine)  ((t  𝖺[u, u0, τμ.p0])  chine)) 
                            ((t  μ  τμ.p0)  chine) =
                          (t  μ  τμ.p0)  chine"
              by simp
            thus ?thesis
              using comp_assoc by metis
          qed
          also have "... = ((t  u)  the_θ)  𝖺-1[t, u, (u0  τμ.p0)  chine] 
                             (t  𝖺[u, u0  τμ.p0, chine])  (t  𝖺[u, u0, τμ.p0]  chine) 
                             𝖺[t, (u  u0)  τμ.p0, chine] 
                             ((t  𝖺-1[u, u0, τμ.p0])  chine)  (((𝖺[t, u, u0  τμ.p0]  chine) 
                             (𝖺-1[t, u, u0  τμ.p0]  chine))  ((t  𝖺[u, u0, τμ.p0])  chine))  
                             ((t  μ  τμ.p0)  chine)  ((t  t0u1)  chine)  
                             (𝖺[t, t0, τμ.p1]  chine)  ((τ  τμ.p1)  chine)  the_ν"
          proof -
            have "((𝖺[t, u, u0  τμ.p0]  chine)  (𝖺-1[t, u, u0  τμ.p0]  chine)) 
                    ((t  𝖺[u, u0, τμ.p0])  chine) =
                  ((t  𝖺[u, u0, τμ.p0])  chine)"
              using comp_inv_arr' comp_cod_arr τμ.composable comp_assoc_assoc'
                    whisker_right [of chine "𝖺[t, u, u0  τμ.p0]" "𝖺-1[t, u, u0  τμ.p0]"]
              by simp
            thus ?thesis
              using comp_assoc by simp
          qed
          also have "... = ((t  u)  the_θ)  𝖺-1[t, u, (u0  τμ.p0)  chine] 
                             (t  𝖺[u, u0  τμ.p0, chine])  (t  𝖺[u, u0, τμ.p0]  chine) 
                             𝖺[t, (u  u0)  τμ.p0, chine] 
                             ((t  𝖺-1[u, u0, τμ.p0])  chine)  (𝖺[t, u, u0  τμ.p0]  chine) 
                             ((𝖺-1[t, u, u0  τμ.p0]  chine)  ((t  𝖺[u, u0, τμ.p0])  chine)  
                             ((t  μ  τμ.p0)  chine)  ((t  t0u1)  chine)  
                             (𝖺[t, t0, τμ.p1]  chine)  ((τ  τμ.p1)  chine))  the_ν"
            using comp_assoc by presburger
          also have "... = ((t  u)  the_θ) 
                             (𝖺-1[t, u, (u0  τμ.p0)  chine] 
                             (t  𝖺[u, u0  τμ.p0, chine])  (t  𝖺[u, u0, τμ.p0]  chine) 
                             𝖺[t, (u  u0)  τμ.p0, chine] 
                             ((t  𝖺-1[u, u0, τμ.p0])  chine)  (𝖺[t, u, u0  τμ.p0]  chine)) 
                             (τμ.tab  chine)  the_ν"
          proof -
            have "(𝖺-1[t, u, u0  τμ.p0]  chine)  ((t  𝖺[u, u0, τμ.p0])  chine)  
                             ((t  μ  τμ.p0)  chine)  ((t  t0u1)  chine)  
                             (𝖺[t, t0, τμ.p1]  chine)  ((τ  τμ.p1)  chine) =
                  τμ.tab  chine"
              using uwθ whisker_right [of chine]
              by (metis τμ.tab_def τμ.tab_in_vhom' arrI seqE)
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... = ((t  u)  the_θ)  𝖺[t  u, u0  τμ.p0, chine]  (τμ.tab  chine)  the_ν"
          proof -
            have "𝖺-1[t, u, (u0  τμ.p0)  chine]  (t  𝖺[u, u0  τμ.p0, chine]) 
                    (t  𝖺[u, u0, τμ.p0]  chine)  𝖺[t, (u  u0)  τμ.p0, chine] 
                    ((t  𝖺-1[u, u0, τμ.p0])  chine)  (𝖺[t, u, u0  τμ.p0]  chine) =
                  𝖺-1[t, u, (u0  τμ.p0)  chine]  (t  𝖺[u, u0  τμ.p0, chine]) 
                   (t  𝖺[u, u0, τμ.p0]  chine)  𝖺[t, (u  u0)  τμ.p0, chine] 
                   ((t  𝖺-1[u, u0, τμ.p0])  chine) 
                   (𝖺[t, u, u0  τμ.p0]  chine)"
              using 𝖺'_def α_def τμ.composable by simp
            also have "... = 𝖺[t  u, u0  τμ.p0, chine]"
              using τμ.composable
              apply (intro E.eval_eqI) by simp_all
            also have "... = 𝖺[t  u, u0  τμ.p0, chine]"
              using 𝖺'_def α_def τμ.composable by simp
            finally show ?thesis by simp
          qed
          also have "... = Δ"
            using Δ_naturality by simp
          finally show ?thesis by simp
        qed
        finally have "𝖺-1[t, u, s0  ρσ.p0]  LHS  𝖺[t1, ω.chine, ρσ.p1]  (ω.the_ν  ρσ.p1) =
                      𝖺-1[t, u, s0  ρσ.p0]  RHS  𝖺[t1, ω.chine, ρσ.p1]  (ω.the_ν  ρσ.p1)"
          by blast
        (*
         * TODO: This is common enough that there should be "cancel_iso_left" and
         * "cancel_iso_right" rules for doing it.
         *)
        hence "(LHS  𝖺[t1, ω.chine, ρσ.p1])  (ω.the_ν  ρσ.p1) =
               (RHS  𝖺[t1, ω.chine, ρσ.p1])  (ω.the_ν  ρσ.p1)"
          using uτ r0s1.ide_u LHS RHS iso_is_section [of "𝖺-1[t, u, s0  ρσ.p0]"] section_is_mono
                monoE τμ.composable comp_assoc
          by (metis (no_types, lifting) Δ_simps(1) μ.ide_base
              𝖺-1[t, u, s0  r0s1.p0]  LHS  𝖺[t1, ω.chine, r0s1.p1]  (ω.the_ν  r0s1.p1) =
               ((ω  χ)  s0  r0s1.p0)  ρσ.tab
              τ.ide_base hseq_char ideD(1) ide_u iso_assoc')
        hence 1: "LHS  𝖺[t1, ω.chine, ρσ.p1] = RHS  𝖺[t1, ω.chine, ρσ.p1]"
          using epiE LHS RHS iso_is_retraction retraction_is_epi τμ.composable
                ω.the_ν_props iso_hcomp
          by (metis Δ_simps(1) ω.the_ν_simps(2)
              ((ω  χ)  s0  r0s1.p0)  ρσ.tab =
               𝖺-1[t, u, s0  r0s1.p0]  RHS  𝖺[t1, ω.chine, r0s1.p1]  (ω.the_ν  r0s1.p1)
              ρ.leg1_simps(3) ide_is_iso local.comp_assoc r0s1.ide_leg1 r0s1.p1_simps seqE)
        show "LHS = RHS"
        proof -
          have "epi 𝖺[t1, ω.chine, ρσ.p1]"
            using iso_is_retraction retraction_is_epi ρ.T0.antipar(1) by simp
          moreover have "seq LHS 𝖺[t1, ω.chine, ρσ.p1]"
            using LHS ρ.T0.antipar(1) by auto
          moreover have "seq RHS 𝖺[t1, ω.chine, ρσ.p1]"
            using RHS ρ.T0.antipar(1) by auto
          ultimately show ?thesis
            using epiE 1 by blast
        qed
      qed
      have 1: "∃!γ. «γ : ?wτ  ?wτ'»  τ = t1  γ  τ = τ'  (t0  γ)"
        using LHS_def RHS_def uτ wτ wτ' βτ θτ θτ' eq τ.T2 [of ?wτ ?wτ' τ ?uτ τ' τ]
        by fastforce
      obtain γτ where γτ: "«γτ : ?wτ  ?wτ'»  τ = t1  γτ  τ = τ'  (t0  γτ)"
        using 1 by auto
      text ‹
        At this point we could show that @{term γτ} is invertible using BS3›,
        but we want to avoid using BS3› if possible and we also want to
        establish a characterization of @{term "inv γτ"}.  So we show the invertibility of
        @{term γτ} directly, using a few more applications of T2›.
      ›
      have iso_βτ: "iso τ"
        using uwθ βτ the_ν_props ω.the_ν_props hseqI' iso_assoc' ω.hseq_leg0
        apply (intro isos_compose)
              apply (metis ω.is_ide ρσ.leg1_simps(2) τ.ide_leg1 τ.leg1_simps(2)
                           τ.leg1_simps(3) hseqE r0s1.ide_leg1 hcomp_simps(1) vconn_implies_hpar(3))
             apply (metis ρσ.leg1_simps(2) hseqE ide_is_iso r0s1.ide_leg1 src_inv iso_inv_iso
                          iso_hcomp vconn_implies_hpar(1))
            apply blast
           apply blast
          apply blast
         apply (metis τ.ide_leg1 τ.leg1_simps(3) hseqE ide_char iso_assoc t0u1.ide_leg1
                      t0u1.p1_simps wτ')
        by blast
      hence eq': "((t  τ')  𝖺[t, t0, ?wτ']  (τ  ?wτ')) =
                  ((t  τ)  𝖺[t, t0, ?wτ]  (τ  ?wτ))  inv τ"
      proof -
        have "seq ((t  τ')  𝖺[t, t0, ?wτ']  (τ  ?wτ')) τ"
          using LHS RHS_def eq by blast
        hence "(t  τ')  𝖺[t, t0, ?wτ']  (τ  ?wτ') =
               (((t  τ')  𝖺[t, t0, ?wτ']  (τ  ?wτ'))  τ)  inv τ"
          by (meson invert_side_of_triangle(2) iso_βτ)
        thus ?thesis
          using LHS_def RHS_def eq by argo
      qed
      have 2: "∃!γ. «γ : ?wτ'  ?wτ»  inv τ = t1  γ  τ' = τ  (t0  γ)"
        using uτ wτ wτ' βτ θτ θτ' eq' τ.T2 [of ?wτ' ?wτ τ'?uτ τ "inv τ"] iso_βτ comp_assoc
        by blast
      obtain γτ' where
        γτ': "«γτ' : ?wτ'  ?wτ»  inv τ = t1  γτ'  τ' = τ  (t0  γτ')"
        using 2 by auto
      have "inverse_arrows γτ γτ'"
      proof
        have "«γτ'  γτ : ?wτ  ?wτ»"
          using γτ γτ' by auto
        moreover have "t1  γτ'  γτ = t1  ?wτ"
          using γτ γτ' whisker_left βτ iso_βτ comp_inv_arr'
          by (metis (no_types, lifting) τ.ide_leg1 calculation in_homE)
        moreover have "τ = τ  (t0  γτ'  γτ)"
        proof -
          have "τ = τ'  (t0  γτ)"
            using γτ by simp
          also have "... = τ  (t0  γτ')  (t0  γτ)"
            using γτ' comp_assoc by simp
          also have "... = τ  (t0  γτ'  γτ)"
            using γτ γτ' whisker_left
            by (metis (full_types) τ.ide_leg0 seqI')
          finally show ?thesis by simp
        qed
        moreover have
          "γ. «γ : ?wτ  ?wτ»  t1  γ = t1  ?wτ  τ = τ  (t0  γ)  γ = ?wτ"
        proof -
          have "∃!γ. «γ : ?wτ  ?wτ»  t1  ?wτ = t1  γ  τ = τ  (t0  γ)"
          proof -
            have "(t  τ)  𝖺[t, t0, ?wτ]  (τ  ?wτ) =
                  ((t  τ)  𝖺[t, t0, ?wτ]  (τ  ?wτ))  (t1  ?wτ)"
              by (metis LHS LHS_def comp_arr_dom in_homE)
            thus ?thesis
              using wτ θτ ω.w_simps(4) τ.leg1_in_hom(2) τ.leg1_simps(3) hcomp_in_vhom ideD(1)
                    trg_hcomp ide_in_hom(2) τ.T2
              by presburger
          qed
          thus "γ. «γ : ?wτ  ?wτ»  t1  γ = t1  ?wτ  τ = τ  (t0  γ)  γ = ?wτ"
            by (metis θτ comp_arr_dom ide_in_hom(2) in_homE wτ)
        qed
        ultimately have "γτ'  γτ = ?wτ"
          by simp
        thus "ide (γτ'  γτ)"
          using wτ by simp
        have "«γτ  γτ' : ?wτ'  ?wτ'»"
          using γτ γτ' by auto
        moreover have "t1  γτ  γτ' = t1  ?wτ'"
          by (metis βτ γτ γτ' τ.ide_leg1 calculation comp_arr_inv' in_homE iso_βτ whisker_left)
        moreover have "τ' = τ'  (t0  γτ  γτ')"
        proof -
          have "τ' = τ  (t0  γτ')"
            using γτ' by simp
          also have "... = τ'  (t0  γτ)  (t0  γτ')"
            using γτ comp_assoc by simp
          also have "... = τ'  (t0  γτ  γτ')"
            using γτ γτ' whisker_left τ.ide_leg0 by fastforce
          finally show ?thesis by simp
        qed
        moreover have "γ. «γ : ?wτ'  ?wτ'»  t1  γ = t1  ?wτ'  τ' = τ'  (t0  γ)
                               γ = ?wτ'"
        proof -
          have "∃!γ. «γ : ?wτ'  ?wτ'»  t1  ?wτ' = t1  γ  τ' = τ'  (t0  γ)"
          proof -
            have "(t  τ')  𝖺[t, t0, ?wτ']  (τ  ?wτ') =
                  ((t  τ')  𝖺[t, t0, ?wτ']  (τ  ?wτ'))  (t1  ?wτ')"
            proof -
              have 1: "t1  γτ  γτ' = (t1  γτ)  (t1  γτ')"
                by (meson γτ γτ' τ.ide_leg1 seqI' whisker_left)
              have "((LHS  inv τ)  (t1  γτ))  (t1  γτ') = LHS  inv τ"
                using LHS_def RHS_def γτ γτ' eq eq' by argo
              thus ?thesis
                unfolding LHS_def
                using 1 by (simp add: calculation(2) eq' comp_assoc)
            qed
            thus ?thesis
              using wτ' θτ' ω.w_simps(4) τ.leg1_in_hom(2) τ.leg1_simps(3) hcomp_in_vhom ideD(1)
                    trg_hcomp ide_in_hom(2) τ.T2 τ.T0.antipar(1) t0u1.base_simps(2)
                    t0u1.leg1_simps(4)
              by presburger
          qed
          thus "γ. «γ : ?wτ'  ?wτ'»  t1  γ = t1  ?wτ'  τ' = τ'  (t0  γ)
                          γ = ?wτ'"
            by (metis θτ' comp_arr_dom ide_in_hom(2) in_homE wτ')
        qed
        ultimately have "γτ  γτ' = ?wτ'"
          by simp
        thus "ide (γτ  γτ')"
          using wτ' by simp
      qed
      thus "τμ.p1  chine  ω.chine  ρσ.p1"
        using wτ wτ' γτ isomorphic_symmetric isomorphic_def by blast
      have iso_γτ: "iso γτ"
        using inverse_arrows γτ γτ' by auto
      have γτ'_eq: "γτ' = inv γτ"
        using inverse_arrows γτ γτ' inverse_unique by blast

      let ?wμ = "τμ.p0  chine"
      let ?wμ' = "χ.chine  ρσ.p0"
      let ?uμ = "s0  ρσ.p0"
      let μ = "the_θ  𝖺-1[u0, τμ.p0, chine]"
      let μ' = "(χ.the_θ  ρσ.p0)  𝖺-1[u0, χ.chine, ρσ.p0]"
      let μ = "𝖺[u1, χ.chine, ρσ.p0]  (χ.the_ν  ρσ.p0)  r0s1  (ω.the_θ  ρσ.p1) 
                  𝖺-1[t0, ω.chine, ρσ.p1]  (t0  inv γτ)  𝖺[t0, τμ.p1, chine] 
                  (inv t0u1  chine)  𝖺-1[u1, τμ.p0, chine]"
      have wμ: "ide ?wμ  is_left_adjoint ?wμ"
        using is_map left_adjoints_compose by simp
      have wμ': "ide ?wμ'  is_left_adjoint ?wμ'"
        using χ.is_map left_adjoints_compose
        by (simp add: is_map left_adjoints_compose)
      have 1: "∃!γ. «γ : ?wμ  ?wμ'»  μ = u1  γ  μ = μ'  (u0  γ)"
      proof -
        have θμ: "«μ : u0  ?wμ  ?uμ»"
          by auto
        have θμ': "«μ' : u0  ?wμ'  ?uμ»"
          by fastforce
        have βμ: "«μ : u1  ?wμ  u1  ?wμ'»"
        proof (intro comp_in_homI)
          show "«𝖺-1[u1, τμ.p0, chine] : u1  τμ.p0  chine  (u1  τμ.p0)  chine»"
            by auto
          show "«inv t0u1  chine : (u1  τμ.p0)  chine  (t0  τμ.p1)  chine»"
            using t0u1.φ_uniqueness(2) hcomp_in_vhom
            by (simp add: t0u1.φ_in_hom(2) w_in_hom(2))
          show "«𝖺[t0, τμ.p1, chine] : (t0  τμ.p1)  chine  t0  τμ.p1  chine»"
            using τ.T0.antipar(1) by auto
          show "«t0  inv γτ : t0  τμ.p1  chine  t0  ω.chine  ρσ.p1»"
            using γτ iso_γτ using τ.T0.antipar(1) by auto
          show "«𝖺-1[t0, ω.chine, ρσ.p1] : t0  ω.chine  ρσ.p1  (t0  ω.chine)  ρσ.p1»"
            using ρ.T0.antipar(1) by auto
          show "«ω.the_θ  ρσ.p1 : (t0  ω.chine)  ρσ.p1  r0  ρσ.p1»"
            using ρ.T0.antipar(1) by auto
          show "«r0s1 : r0  ρσ.p1  s1  ρσ.p0»"
            by auto
          show "«χ.the_ν  ρσ.p0 : s1  ρσ.p0  (u1  χ.chine)  ρσ.p0»"
            by auto
          show "«𝖺[u1, χ.chine, ρσ.p0] : (u1  χ.chine)  ρσ.p0  u1  χ.chine  ρσ.p0»"
            by auto
        qed
        text ‹
          The proof of the equation below needs to make use of the equation
          θτ' = θτ ⋅ (t0 ⋆ γτ')› from the previous section.  So the overall strategy is to
          work toward an expression of the form θτ ⋅ (t0 ⋆ γτ')› and perform the replacement
          to eliminate γτ'›.
        ›
        have eqμ: "(u  μ)  𝖺[u, u0, ?wμ]  (μ  ?wμ) =
                   ((u  μ')  𝖺[u, u0, ?wμ']  (μ  ?wμ'))  μ"
        proof -
          let ?LHS = "(u  μ)  𝖺[u, u0, ?wμ]  (μ  ?wμ)"
          let ?RHS = "((u  μ')  𝖺[u, u0, ?wμ']  (μ  ?wμ'))  μ"
          have "?RHS = (u  (χ.the_θ  ρσ.p0)  𝖺-1[u0, χ.chine, ρσ.p0]) 
                         𝖺[u, u0, χ.chine  ρσ.p0]  (μ  χ.chine  ρσ.p0) 
                         𝖺[u1, χ.chine, ρσ.p0]  (χ.the_ν  ρσ.p0)  r0s1 
                         (ω.the_θ  ρσ.p1)  𝖺-1[t0, ω.chine, ρσ.p1]  (t0  inv γτ) 
                         𝖺[t0, τμ.p1, chine]  (inv t0u1  chine)  𝖺-1[u1, τμ.p0, chine]"
            using comp_assoc by presburger
          also have "... = (u  χ.the_θ  ρσ.p0)  ((u  𝖺-1[u0, χ.chine, ρσ.p0]) 
                             𝖺[u, u0, χ.chine  ρσ.p0])  (μ  χ.chine  ρσ.p0) 
                             𝖺[u1, χ.chine, ρσ.p0]  (χ.the_ν  ρσ.p0)  r0s1 
                             (ω.the_θ  ρσ.p1)  𝖺-1[t0, ω.chine, ρσ.p1]  (t0  inv γτ) 
                             𝖺[t0, τμ.p1, chine]  (inv t0u1  chine) 
                             𝖺-1[u1, τμ.p0, chine]"
          proof -
            have "u  (χ.the_θ  ρσ.p0)  𝖺-1[u0, χ.chine, ρσ.p0] =
                  (u  χ.the_θ  ρσ.p0)  (u  𝖺-1[u0, χ.chine, ρσ.p0])"
              using whisker_left μ.ide_base θμ' by blast
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... = ((u  χ.the_θ  ρσ.p0)  𝖺[u, u0  χ.chine, ρσ.p0]) 
                             (𝖺[u, u0, χ.chine]  ρσ.p0)  𝖺-1[u  u0, χ.chine, ρσ.p0] 
                             (μ  χ.chine  ρσ.p0) 
                             𝖺[u1, χ.chine, ρσ.p0]  (χ.the_ν  ρσ.p0)  r0s1 
                             (ω.the_θ  ρσ.p1)  𝖺-1[t0, ω.chine, ρσ.p1]  (t0  inv γτ) 
                             𝖺[t0, τμ.p1, chine]  (inv t0u1  chine) 
                             𝖺-1[u1, τμ.p0, chine]"
          proof -
            have "seq (u  𝖺[u0, χ.chine, ρσ.p0])
                      (𝖺[u, u0  χ.chine, ρσ.p0]  (𝖺[u, u0, χ.chine]  ρσ.p0))"
              by auto
            moreover have "src u = trg 𝖺[u0, χ.chine, ρσ.p0]"
              by simp
            moreover have "inv (u  𝖺[u0, χ.chine, ρσ.p0]) = u  𝖺-1[u0, χ.chine, ρσ.p0]"
              by simp
            moreover have "iso (u  𝖺[u0, χ.chine, ρσ.p0])"
              by simp
            moreover have "iso 𝖺[u  u0, χ.chine, ρσ.p0]"
              by simp
            ultimately have "(u  𝖺-1[u0, χ.chine, ρσ.p0])  𝖺[u, u0, χ.chine  ρσ.p0] =
                             𝖺[u, u0  χ.chine, ρσ.p0]  (𝖺[u, u0, χ.chine]  ρσ.p0) 
                               𝖺-1[u  u0, χ.chine, ρσ.p0]"
              using pentagon hseqI' comp_assoc
                    invert_opposite_sides_of_square
                      [of "u  𝖺[u0, χ.chine, ρσ.p0]"
                          "𝖺[u, u0  χ.chine, ρσ.p0]  (𝖺[u, u0, χ.chine]  ρσ.p0)"
                          "𝖺[u, u0, χ.chine  ρσ.p0]" "𝖺[u  u0, χ.chine, ρσ.p0]"]
                    inv_hcomp χ.is_ide χ.w_simps(3) χ.w_simps(4) μ.base_simps(2) μ.ide_base
                    μ.ide_leg0 μ.leg0_simps(2) μ.leg0_simps(3) σ.leg1_simps(3)
                    assoc'_eq_inv_assoc ide_hcomp r0s1.ide_u r0s1.p0_simps hcomp_simps(1)
              by presburger
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... = 𝖺[u, s0, ρσ.p0]  ((u  χ.the_θ)  ρσ.p0) 
                             (𝖺[u, u0, χ.chine]  ρσ.p0)  (𝖺-1[u  u0, χ.chine, ρσ.p0] 
                             (μ  χ.chine  ρσ.p0)) 
                             𝖺[u1, χ.chine, ρσ.p0]  (χ.the_ν  ρσ.p0)  r0s1 
                             (ω.the_θ  ρσ.p1)  𝖺-1[t0, ω.chine, ρσ.p1]  (t0  inv γτ) 
                             𝖺[t0, τμ.p1, chine]  (inv t0u1  chine) 
                             𝖺-1[u1, τμ.p0, chine]"
          proof -
            have "(u  χ.the_θ  ρσ.p0)  𝖺[u, u0  χ.chine, ρσ.p0] =
                  𝖺[u, s0, ρσ.p0]  ((u  χ.the_θ)  ρσ.p0)"
              using assoc_naturality [of u χ.the_θ ρσ.p0] χ.θ_simps(3) by auto
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... = 𝖺[u, s0, ρσ.p0]  ((u  χ.the_θ)  ρσ.p0) 
                             (𝖺[u, u0, χ.chine]  ρσ.p0)  ((μ  χ.chine)  ρσ.p0) 
                             𝖺-1[u1, χ.chine, ρσ.p0]  𝖺[u1, χ.chine, ρσ.p0] 
                             (χ.the_ν  ρσ.p0)  r0s1 
                             (ω.the_θ  ρσ.p1)  𝖺-1[t0, ω.chine, ρσ.p1]  (t0  inv γτ) 
                             𝖺[t0, τμ.p1, chine]  (inv t0u1  chine) 
                             𝖺-1[u1, τμ.p0, chine]"
          proof -
            have "𝖺-1[u  u0, χ.chine, ρσ.p0]  (μ  χ.chine  ρσ.p0) =
                  ((μ  χ.chine)  ρσ.p0)  𝖺-1[u1, χ.chine, ρσ.p0]"
              using assoc'_naturality [of μ χ.chine ρσ.p0] by simp
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... = 𝖺[u, s0, ρσ.p0]  ((u  χ.the_θ)  ρσ.p0) 
                             (𝖺[u, u0, χ.chine]  ρσ.p0)  ((μ  χ.chine)  ρσ.p0) 
                             ((𝖺-1[u1, χ.chine, ρσ.p0]  𝖺[u1, χ.chine, ρσ.p0]) 
                             (χ.the_ν  ρσ.p0))  r0s1 
                             (ω.the_θ  ρσ.p1)  𝖺-1[t0, ω.chine, ρσ.p1]  (t0  inv γτ) 
                             𝖺[t0, τμ.p1, chine]  (inv t0u1  chine) 
                             𝖺-1[u1, τμ.p0, chine]"
            using comp_assoc by metis
          also have "... = 𝖺[u, s0, ρσ.p0]  (((u  χ.the_θ)  ρσ.p0) 
                             (𝖺[u, u0, χ.chine]  ρσ.p0)  ((μ  χ.chine)  ρσ.p0) 
                             (χ.the_ν  ρσ.p0))  r0s1 
                             (ω.the_θ  ρσ.p1)  𝖺-1[t0, ω.chine, ρσ.p1]  (t0  inv γτ) 
                             𝖺[t0, τμ.p1, chine]  (inv t0u1  chine) 
                             𝖺-1[u1, τμ.p0, chine]"
          proof -
            have "(𝖺-1[u1, χ.chine, ρσ.p0]  𝖺[u1, χ.chine, ρσ.p0])  (χ.the_ν  ρσ.p0) =
                  χ.the_ν  ρσ.p0"
              using comp_inv_arr' comp_cod_arr by auto
            thus ?thesis
              using comp_assoc by simp
          qed
          also have "... = (𝖺[u, s0, ρσ.p0]  ((χ  s0)  σ  ρσ.p0)  r0s1 
                             (ω.the_θ  ρσ.p1)  𝖺-1[t0, ω.chine, ρσ.p1])  (t0  inv γτ) 
                             𝖺[t0, τμ.p1, chine]  (inv t0u1  chine) 
                             𝖺-1[u1, τμ.p0, chine]"
          proof -
            have "arr ((u  χ.the_θ)  𝖺[u, u0, χ.chine]  (μ  χ.chine)  χ.the_ν)"
              using χ.θ_simps(3) by simp
            hence "((u  χ.the_θ)  ρσ.p0)  (𝖺[u, u0, χ.chine]  ρσ.p0) 
                     ((μ  χ.chine)  ρσ.p0)  (χ.the_ν  ρσ.p0) =
                   (u  χ.the_θ)  𝖺[u, u0, χ.chine]  (μ  χ.chine)  χ.the_ν  ρσ.p0"
              using whisker_right by simp
            also have "... = (χ  s0)  σ  ρσ.p0"
              using χ.Δ_naturality by simp
            finally have "((u  χ.the_θ)  ρσ.p0)  (𝖺[u, u0, χ.chine]  ρσ.p0) 
                            ((μ  χ.chine)  ρσ.p0)  (χ.the_ν  ρσ.p0) =
                          (χ  s0)  σ  ρσ.p0"
              by simp
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... = (τ  (t0  inv γτ)) 
                             𝖺[t0, τμ.p1, chine]  (inv t0u1  chine) 
                             𝖺-1[u1, τμ.p0, chine]"
            using comp_assoc by presburger
          also have "... = τ'  𝖺[t0, τμ.p1, chine]  (inv t0u1  chine) 
                             𝖺-1[u1, τμ.p0, chine]"
            using γτ' γτ'_eq by simp
          also have "... = (u  the_θ)  𝖺[u, u0  τμ.p0, chine]  (𝖺[u, u0, τμ.p0]  chine) 
                             ((μ  τμ.p0)  chine)  ((t0u1  chine) 
                             ((𝖺-1[t0, τμ.p1, chine])  𝖺[t0, τμ.p1, chine]) 
                             (inv t0u1  chine))  𝖺-1[u1, τμ.p0, chine]"
            using comp_assoc by presburger
          also have "... = (u  the_θ)  𝖺[u, u0  τμ.p0, chine]  (𝖺[u, u0, τμ.p0]  chine) 
                             ((μ  τμ.p0)  chine)  𝖺-1[u1, τμ.p0, chine]"
          proof -
            have "((t0u1  chine)  ((𝖺-1[t0, τμ.p1, chine])  𝖺[t0, τμ.p1, chine]) 
                    (inv t0u1  chine))  𝖺-1[u1, τμ.p0, chine] =
                  𝖺-1[u1, τμ.p0, chine]"
            proof -
              have "((t0u1  chine)  ((𝖺-1[t0, τμ.p1, chine])  𝖺[t0, τμ.p1, chine]) 
                      (inv t0u1  chine))  𝖺-1[u1, τμ.p0, chine] =
                    ((t0u1  chine)  ((t0  τμ.p1)  chine) 
                      (inv t0u1  chine))  𝖺-1[u1, τμ.p0, chine]"
                using comp_inv_arr' τ.T0.antipar(1) by auto
              also have "... = ((t0u1  chine)  (inv t0u1  chine))  𝖺-1[u1, τμ.p0, chine]"
                using comp_cod_arr t0u1.φ_uniqueness by simp
              also have "... = (t0u1  inv t0u1  chine)  𝖺-1[u1, τμ.p0, chine]"
                using whisker_right t0u1.φ_uniqueness by simp
              also have "... = ((u1  τμ.p0)  chine)  𝖺-1[u1, τμ.p0, chine]"
                using comp_arr_inv' τ.T0.antipar(1) t0u1.φ_uniqueness by simp
              also have "... = 𝖺-1[u1, τμ.p0, chine]"
                using comp_cod_arr τ.T0.antipar(1) by simp
              finally show ?thesis by simp
            qed
            thus ?thesis
              using comp_assoc by simp
          qed
          also have "... = (u  the_θ)  (𝖺[u, u0  τμ.p0, chine]  (𝖺[u, u0, τμ.p0]  chine) 
                             𝖺-1[u  u0, τμ.p0, chine])  (μ  ?wμ)"
            using assoc'_naturality [of μ τμ.p0 chine] comp_assoc by auto
          also have "... = ((u  the_θ)  (u  𝖺-1[u0, τμ.p0, chine]))  𝖺[u, u0, ?wμ]  (μ  ?wμ)"
            using uwθ pentagon comp_assoc
                  invert_opposite_sides_of_square
                    [of "u  𝖺[u0, τμ.p0, chine]"
                        "𝖺[u, u0  τμ.p0, chine]  (𝖺[u, u0, τμ.p0]  chine)" "𝖺[u, u0, ?wμ]"
                        "𝖺[u  u0, τμ.p0, chine]"]
                  μ.base_simps(2) μ.ide_base μ.ide_leg0 μ.leg0_simps(2) assoc'_eq_inv_assoc
                  ide_hcomp hcomp_simps(1) t0u1.ide_u
            by force
          also have "... = (u  the_θ  𝖺-1[u0, τμ.p0, chine])  𝖺[u, u0, ?wμ]  (μ  ?wμ)"
            using whisker_left comp_assoc by simp
          finally show ?thesis by simp
        qed
        show "∃!γ. «γ : ?wμ  ?wμ'»  μ = u1  γ  μ = μ'  (u0  γ)"
          using wμ wμ' θμ θμ' βμ eqμ μ.T2 [of ?wμ ?wμ' μ ?uμ μ' μ] by fast
      qed
      obtain γμ where γμ: "«γμ : ?wμ  ?wμ'»  μ = u1  γμ  μ = μ'  (u0  γμ)"
        using 1 by auto
      show "?wμ  ?wμ'"
        using wμ wμ' γμ BS3 [of ?wμ ?wμ' γμ γμ] isomorphic_def by auto
    qed

    lemma comp_L:
    shows "Maps.seq ⟦⟦t0⟧⟧ ⟦⟦ω.chine  ρσ.p1⟧⟧"
    and "⟦⟦t0⟧⟧  ⟦⟦ω.chine  ρσ.p1⟧⟧ =
         Maps.MkArr (src (ω.chine  ρσ.p1)) (src t) (Maps.Comp t0 ω.chine  ρσ.p1)"
    proof -
      show "Maps.seq ⟦⟦t0⟧⟧ ⟦⟦ω.chine  ρσ.p1⟧⟧"
      proof -
        have "is_left_adjoint (ω.chine  ρσ.p1)"
          using ω.is_map r0s1.leg1_is_map left_adjoints_compose r0s1.p1_simps by auto
        thus ?thesis
          using Maps.CLS_in_hom r0s1.leg1_is_map
          apply (intro Maps.seqI')
           apply blast
          using Maps.CLS_in_hom [of t0] τ.leg0_is_map ρσ.leg1_in_hom by auto
      qed
      thus "⟦⟦t0⟧⟧  ⟦⟦ω.chine  ρσ.p1⟧⟧ =
            Maps.MkArr (src (ω.chine  ρσ.p1)) (src t) (Maps.Comp t0 ω.chine  ρσ.p1)"
        using Maps.comp_char by auto
    qed

    lemma comp_R:
    shows "Maps.seq ⟦⟦u1⟧⟧ ⟦⟦χ.chine  ρσ.p0⟧⟧"
    and "⟦⟦u1⟧⟧  ⟦⟦χ.chine  ρσ.p0⟧⟧ =
         Maps.MkArr (src r0s1.p0) (trg u) (Maps.Comp u1 χ.chine  r0s1.p0)"
    proof -
      show "Maps.seq ⟦⟦u1⟧⟧ ⟦⟦χ.chine  ρσ.p0⟧⟧"
      proof -
        have "is_left_adjoint (χ.chine  ρσ.p0)"
          using χ.is_map r0s1.leg0_is_map left_adjoints_compose [of χ.chine ρσ.p0] by simp
        thus ?thesis
          using Maps.CLS_in_hom μ.leg1_is_map
          apply (intro Maps.seqI')
           apply blast
          using Maps.CLS_in_hom [of u1] μ.leg1_is_map by simp
      qed
      thus "⟦⟦u1⟧⟧  ⟦⟦χ.chine  ρσ.p0⟧⟧ =
            Maps.MkArr (src r0s1.p0) (trg u) (Maps.Comp u1 χ.chine  r0s1.p0)"
        using Maps.comp_char by auto
    qed

    lemma comp_L_eq_comp_R:
    shows "⟦⟦t0⟧⟧  ⟦⟦ω.chine  ρσ.p1⟧⟧ = ⟦⟦u1⟧⟧  ⟦⟦χ.chine  ρσ.p0⟧⟧"
    proof (intro Maps.arr_eqI)
      show "Maps.seq ⟦⟦t0⟧⟧ ⟦⟦ω.chine  ρσ.p1⟧⟧"
        using comp_L(1) by simp
      show "Maps.seq ⟦⟦u1⟧⟧ ⟦⟦χ.chine  ρσ.p0⟧⟧"
        using comp_R(1) by simp
      show "Maps.Dom (⟦⟦t0⟧⟧  ⟦⟦ω.chine  ρσ.p1⟧⟧) = Maps.Dom (⟦⟦u1⟧⟧  ⟦⟦χ.chine  ρσ.p0⟧⟧)"
        by (metis (no_types, lifting) Maps.Dom.simps(1) ω.w_simps(2) ω.w_simps(3)
            ρ.leg1_simps(3) ρσ.leg1_in_hom(2) comp_L(2) comp_R(2) hcomp_in_vhomE hseqI'
            r0s1.leg1_simps(3) hcomp_simps(1))
      show "Maps.Cod (⟦⟦t0⟧⟧  ⟦⟦ω.chine  ρσ.p1⟧⟧) = Maps.Cod (⟦⟦u1⟧⟧  ⟦⟦χ.chine  ρσ.p0⟧⟧)"
        by (metis Maps.Cod.simps(1) τμ.composable comp_L(2) comp_R(2))
      have A: "Maps.Map (⟦⟦t0⟧⟧  ⟦⟦ω.chine  ρσ.p1⟧⟧) = Maps.Comp t0 ω.chine  ρσ.p1"
        using comp_L(1) Maps.comp_char by auto
      have B: "Maps.Map (⟦⟦u1⟧⟧  ⟦⟦χ.chine  ρσ.p0⟧⟧) = Maps.Comp u1 χ.chine  ρσ.p0"
        using comp_R(1) Maps.comp_char by auto
      have C: "Maps.Comp t0 ω.chine  ρσ.p1 = Maps.Comp u1 χ.chine  ρσ.p0"
      proof (intro Maps.Comp_eqI)
        show "t0  ω.chine  ρσ.p1  Maps.Comp t0 ω.chine  ρσ.p1"
        proof (intro Maps.in_CompI)
          show "is_iso_class ω.chine  ρσ.p1"
            using prj_chine(2) is_iso_classI isomorphic_implies_hpar(2) by blast
          show "is_iso_class t0"
            using is_iso_classI by auto
          show "ω.chine  ρσ.p1  ω.chine  ρσ.p1"
            using ide_in_iso_class prj_chine(2) isomorphic_implies_hpar(2) by blast
          show "t0  t0"
            using ide_in_iso_class by simp
          show "t0  ω.chine  ρσ.p1  t0  ω.chine  ρσ.p1"
            using isomorphic_reflexive prj_chine(2) isomorphic_implies_hpar(2) by auto
        qed
        show "u1  χ.chine  ρσ.p0  Maps.Comp u1 χ.chine  ρσ.p0"
        proof (intro Maps.in_CompI)
          show "is_iso_class χ.chine  ρσ.p0"
            using is_iso_classI by simp
          show "is_iso_class u1"
            using is_iso_classI by simp
          show "χ.chine  ρσ.p0  χ.chine  ρσ.p0"
            using ide_in_iso_class by simp
          show "u1  iso_class u1"
            using ide_in_iso_class by simp
          show "u1  χ.chine  ρσ.p0  u1  χ.chine  ρσ.p0"
            using isomorphic_reflexive isomorphic_implies_hpar(2) by auto
        qed
        show "t0  ω.chine  ρσ.p1  u1  χ.chine  ρσ.p0"
        proof -
          have "t0  ω.chine  ρσ.p1  (t0  ω.chine)  ρσ.p1"
            using assoc'_in_hom [of t0 ω.chine ρσ.p1] iso_assoc' isomorphic_def r0s1.p1_simps
            by auto
          also have "...  r0  ρσ.p1"
            using ω.leg0_uniquely_isomorphic hcomp_isomorphic_ide
            by (simp add: ρ.T0.antipar(1))
          also have "...  s1  ρσ.p0"
            using isomorphic_def r0s1.φ_uniqueness(2) by blast
          also have "...  (u1  χ.chine)  ρσ.p0"
            using χ.leg1_uniquely_isomorphic hcomp_isomorphic_ide by auto
          also have "...  u1  χ.chine  ρσ.p0"
            using assoc_in_hom [of u1 χ.chine ρσ.p0] iso_assoc isomorphic_def by auto
          finally show ?thesis by simp
        qed
      qed
      show "Maps.Map (⟦⟦t0⟧⟧  ⟦⟦ω.chine  ρσ.p1⟧⟧) = Maps.Map (⟦⟦u1⟧⟧  ⟦⟦χ.chine  ρσ.p0⟧⟧)"
        using A B C by simp
    qed

    lemma csq:
    shows "Maps.commutative_square ⟦⟦t0⟧⟧ ⟦⟦u1⟧⟧ ⟦⟦ω.chine  ρσ.p1⟧⟧ ⟦⟦χ.chine  ρσ.p0⟧⟧"
    proof
      show "Maps.cospan ⟦⟦t0⟧⟧ ⟦⟦u1⟧⟧"
        using comp_L(1) comp_R(1) comp_L_eq_comp_R
        by (metis (no_types, lifting) Maps.cod_comp Maps.seq_char)
      show "Maps.span ⟦⟦ω.chine  ρσ.p1⟧⟧ ⟦⟦χ.chine  ρσ.p0⟧⟧"
        using comp_L(1) comp_R(1) comp_L_eq_comp_R
        by (metis (no_types, lifting) Maps.dom_comp Maps.seq_char)
      show "Maps.dom ⟦⟦t0⟧⟧ = Maps.cod ⟦⟦ω.chine  ρσ.p1⟧⟧"
        using comp_L(1) by auto
      show "⟦⟦t0⟧⟧  ⟦⟦ω.chine  ρσ.p1⟧⟧ = ⟦⟦u1⟧⟧  ⟦⟦χ.chine  ρσ.p0⟧⟧"
        using comp_L_eq_comp_R by simp
    qed

    lemma CLS_chine:
    shows "⟦⟦chine⟧⟧ = Maps.tuple ⟦⟦ω.chine  ρσ.p1⟧⟧ ⟦⟦t0⟧⟧ ⟦⟦u1⟧⟧ ⟦⟦χ.chine  ρσ.p0⟧⟧"
    proof -
      let ?T = "Maps.tuple ⟦⟦ω.chine  ρσ.p1⟧⟧ ⟦⟦t0⟧⟧ ⟦⟦u1⟧⟧ ⟦⟦χ.chine  ρσ.p0⟧⟧"
      have "∃!l. ⟦⟦t0u1.p1⟧⟧  l = ⟦⟦ω.chine  ρσ.p1⟧⟧  ⟦⟦t0u1.p0⟧⟧  l = ⟦⟦χ.chine  ρσ.p0⟧⟧"
        using csq τμ.prj_char
              Maps.universal [of "⟦⟦t0⟧⟧" "⟦⟦u1⟧⟧" "⟦⟦ω.chine  ρσ.p1⟧⟧" "⟦⟦χ.chine  ρσ.p0⟧⟧"]
        by simp
      moreover have "⟦⟦τμ.p1⟧⟧  ?T = ⟦⟦ω.chine  ρσ.p1⟧⟧ 
                     ⟦⟦τμ.p0⟧⟧  ?T = ⟦⟦χ.chine  ρσ.p0⟧⟧"
        using csq τμ.prj_char
              Maps.prj_tuple [of "⟦⟦t0⟧⟧" "⟦⟦u1⟧⟧" "⟦⟦ω.chine  ρσ.p1⟧⟧" "⟦⟦χ.chine  ρσ.p0⟧⟧"]
        by simp
      moreover have "⟦⟦t0u1.p1⟧⟧  ⟦⟦chine⟧⟧ = ⟦⟦ω.chine  ρσ.p1⟧⟧ 
                     ⟦⟦t0u1.p0⟧⟧  ⟦⟦chine⟧⟧ = ⟦⟦χ.chine  ρσ.p0⟧⟧"
        using prj_chine τμ.leg0_is_map τμ.leg1_is_map is_map t0u1.leg1_is_map
              t0u1.satisfies_T0 Maps.comp_CLS
        by blast
      ultimately show "⟦⟦chine⟧⟧ = ?T" by auto
    qed

  end

  subsection "Equivalence of B and Span(Maps(B))"

  subsubsection "The Functor SPN"
 
  text ‹
    We now define a function SPN› on arrows and will ultimately show that it extends to a
    biequivalence from the underlying bicategory B› to Span(Maps(B))›.
    The idea is that SPN› takes «μ : r ⇒ s»› to the isomorphism class of an induced arrow
    of spans from the chosen tabulation of r› to the chosen tabulation of s›.
    To obtain this, we first use isomorphisms r.tab1 ⋆ r.tab0* ≅ r› and s.tab1 ⋆ s.tab0* ≅ s›
    to transform μ› to «μ' : r.tab1 ⋆ r.tab0* ⇒ s.tab1 ⋆ s.tab0*»›.
    We then take the adjoint transpose of μ'› to obtain
    «ω : r.tab1 ⇒ (s.tab1 ⋆ s.tab0*) ⋆ r.tab0»›.  The 2-cell ω› induces a map w›
    which is an arrow of spans from (r.tab0, r.tab1)› to (s.tab0, s.tab1)›.
    We take the arrow of Span(Maps(B))› defined by w› as the value of SPN μ›.

    Ensuring that SPN› is functorial is a somewhat delicate point, which requires that all
    the underlying definitions that have been set up are ``just so'', with no extra choices
    other than those that are forced, and with the tabulation assigned to each 1-cell r› in
    the proper relationship with the canonical tabulation assigned to its chosen factorization
    r = g ⋆ f*.
  ›

  context bicategory_of_spans
  begin

    interpretation Maps: maps_category V H 𝖺 𝗂 src trg ..
    interpretation Span: span_bicategory Maps.comp Maps.PRJ0 Maps.PRJ1 ..

    no_notation Fun.comp (infixl "" 55)
    notation Span.vcomp (infixr "" 55)
    notation Span.hcomp (infixr "" 53)
    notation Maps.comp  (infixr "" 55)
    notation isomorphic (infix "" 50)

    definition spn
    where "spn μ 
           arrow_of_tabulations_in_maps.chine V H 𝖺 𝗂 src trg
             (tab_of_ide (dom μ)) (tab0 (dom μ)) (cod μ)
             (tab_of_ide (cod μ)) (tab0 (cod μ)) (tab1 (cod μ)) μ"

    lemma is_induced_map_spn:
    assumes "arr μ"
    shows "arrow_of_tabulations_in_maps.is_induced_map V H 𝖺 𝗂 src trg
             (tab_of_ide (dom μ)) (tab0 (dom μ)) (cod μ)
             (tab_of_ide (cod μ)) (tab0 (cod μ)) (tab1 (cod μ))
             μ (spn μ)"
    proof -
      interpret μ: arrow_in_bicategory_of_spans V H 𝖺 𝗂 src trg dom μ cod μ μ
        using assms by unfold_locales auto
      interpret μ: arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg
                     dom μ μ.r.tab tab0 (dom μ) tab1 (dom μ)
                     cod μ μ.s.tab tab0 (cod μ) tab1 (cod μ)
                     μ
        using μ.is_arrow_of_tabulations_in_maps by simp
      show ?thesis
        unfolding spn_def
        using μ.chine_is_induced_map by blast
    qed

    lemma spn_props:
    assumes "arr μ"
    shows "«spn μ : src (tab0 (dom μ))  src (tab0 (cod μ))»"
    and "is_left_adjoint (spn μ)"
    and "tab0 (cod μ)  spn μ  tab0 (dom μ)"
    and "tab1 (cod μ)  spn μ  tab1 (dom μ)"
    proof -
      interpret μ: arrow_in_bicategory_of_spans V H 𝖺 𝗂 src trg dom μ cod μ μ
        using assms by unfold_locales auto
      interpret μ: arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg
                     dom μ μ.r.tab tab0 (dom μ) tab1 (dom μ)
                     cod μ μ.s.tab tab0 (cod μ) tab1 (cod μ)
                     μ
        using μ.is_arrow_of_tabulations_in_maps by simp
      show "«spn μ : src (tab0 (dom μ))  src (tab0 (cod μ))»"
        using spn_def by simp
      show "is_left_adjoint (spn μ)"
        using spn_def by (simp add: μ.is_map)
      show "tab0 (cod μ)  spn μ  tab0 (dom μ)"
        using spn_def isomorphic_def μ.leg0_uniquely_isomorphic(1) by auto
      show "tab1 (cod μ)  spn μ  tab1 (dom μ)"
        using spn_def isomorphic_def isomorphic_symmetric
              μ.leg1_uniquely_isomorphic(1)
        by auto
    qed

    lemma spn_in_hom [intro]:
    assumes "arr μ"
    shows "«spn μ : src (tab0 (dom μ))  src (tab0 (cod μ))»"
    and "«spn μ : spn μ  spn μ»"
      using assms spn_props left_adjoint_is_ide by auto

    lemma spn_simps [simp]:
    assumes "arr μ"
    shows "is_left_adjoint (spn μ)"
    and "ide (spn μ)"
    and "src (spn μ) = src (tab0 (dom μ))"
    and "trg (spn μ) = src (tab0 (cod μ))"
      using assms spn_props left_adjoint_is_ide by auto

    text ‹
      We need the next result to show that SPN› is functorial; in particular,
      that it takes «r : r ⇒ r»› in the underlying bicategory to a 1-cell
      in Span(Maps(B))›.  The 1-cells in Span(Maps(B))› have objects of Maps(B)›
      as their chines, and objects of Maps(B)› are isomorphism classes of objects in the
      underlying bicategory B›. So we need the induced map associated with r› to be isomorphic
      to an object.
    ›

    lemma spn_ide:
    assumes "ide r"
    shows "spn r  src (tab0 r)"
    proof -
      interpret r: identity_in_bicategory_of_spans V H 𝖺 𝗂 src trg r
        using assms by unfold_locales auto
      interpret r: arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg
                     r r.tab tab0 r tab1 r r r.tab tab0 r tab1 r r
        using r.is_arrow_of_tabulations_in_maps by simp
      interpret tab: tabulation V H 𝖺 𝗂 src trg r r.tab tab0 r dom r.tab
        using assms r.tab_is_tabulation by simp
      interpret tab: tabulation_in_maps V H 𝖺 𝗂 src trg r r.tab tab0 r dom r.tab
        by (unfold_locales, simp_all)
      have "tab.is_induced_by_cell (spn r) (tab0 r) r.tab"
        using spn_def comp_ide_arr r.chine_is_induced_map by auto
      thus ?thesis
        using tab.induced_map_unique [of "tab0 r" "r.tab" "spn r" "src r.s0"]
              tab.apex_is_induced_by_cell
        by (simp add: comp_assoc)
    qed

    text ‹
      The other key result we need to show that SPN› is functorial is to show
      that the induced map of a composite is isomorphic to the composite of
      induced maps.
    ›

    lemma spn_hcomp:
    assumes "seq τ μ" and "g  spn τ" and "f  spn μ"
    shows "spn (τ  μ)  g  f"
    proof -
      interpret τ: arrow_in_bicategory_of_spans V H 𝖺 𝗂 src trg dom τ cod τ τ
        using assms by unfold_locales auto
      interpret τ: arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg
                     dom τ τ.r.tab tab0 (dom τ) tab1 (dom τ)
                     cod τ τ.s.tab tab0 (cod τ) tab1 (cod τ)
                     τ
        using τ.is_arrow_of_tabulations_in_maps by simp
      interpret μ: arrow_in_bicategory_of_spans V H 𝖺 𝗂 src trg dom μ dom τ μ
        using assms apply unfold_locales
         apply auto[1]
        by (elim seqE, auto)
      interpret μ: arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg
                     dom μ μ.r.tab tab0 (dom μ) tab1 (dom μ)
                     dom τ τ.r.tab tab0 (dom τ) tab1 (dom τ)
                     μ
        using μ.is_arrow_of_tabulations_in_maps by simp
      interpret τμ: vertical_composite_of_arrows_of_tabulations_in_maps V H 𝖺 𝗂 src trg
                      dom μ μ.r.tab tab0 (dom μ) tab1 (dom μ)
                      dom τ τ.r.tab tab0 (dom τ) tab1 (dom τ)
                      cod τ τ.s.tab tab0 (cod τ) tab1 (cod τ)
                      μ τ
        ..
      have "g  τ.chine"
        using assms(2) spn_def by auto
      moreover have "f  μ.chine"
        using assms(1) assms(3) spn_def by auto
      moreover have "src g = trg f"
        using calculation(1-2) isomorphic_implies_hpar(3-4) by auto
      moreover have "src g = trg μ.chine"
        using calculation(1) isomorphic_implies_hpar(3) by auto
      ultimately have "g  f  τ.chine  μ.chine"
        using hcomp_ide_isomorphic hcomp_isomorphic_ide isomorphic_transitive
        by (meson μ.is_ide isomorphic_implies_ide(1))
      also have "...  spn (τ  μ)"
        using spn_def τμ.chine_char isomorphic_symmetric
        by (metis τμ.in_hom in_homE)
      finally show ?thesis
        using isomorphic_symmetric by simp
    qed

    abbreviation (input) SPN0
    where "SPN0 r  Span.mkIde ⟦⟦tab0 r⟧⟧ ⟦⟦tab1 r⟧⟧"

    definition SPN
    where "SPN μ  if arr μ then
                      Chn = ⟦⟦spn μ⟧⟧,
                       Dom = Leg0 = ⟦⟦tab0 (dom μ)⟧⟧, Leg1 = ⟦⟦tab1 (dom μ)⟧⟧,
                       Cod = Leg0 = ⟦⟦tab0 (cod μ)⟧⟧, Leg1 = ⟦⟦tab1 (cod μ)⟧⟧
                    else Span.null"

    lemma Dom_SPN [simp]:
    assumes "arr μ"
    shows "Dom (SPN μ) = Leg0 = ⟦⟦tab0 (dom μ)⟧⟧, Leg1 = ⟦⟦tab1 (dom μ)⟧⟧"
      using assms SPN_def by simp

    lemma Cod_SPN [simp]:
    assumes "arr μ"
    shows "Cod (SPN μ) = Leg0 = ⟦⟦tab0 (cod μ)⟧⟧, Leg1 = ⟦⟦tab1 (cod μ)⟧⟧"
      using assms SPN_def by simp

    text ‹Now we have to show this does the right thing for us.›

    lemma SPN_in_hom:
    assumes "arr μ"
    shows "Span.in_hom (SPN μ) (SPN0 (dom μ)) (SPN0 (cod μ))"
    proof
      interpret Dom: span_in_category Maps.comp Dom (SPN μ)
      proof
        interpret r: identity_in_bicategory_of_spans V H 𝖺 𝗂 src trg dom μ
          using assms by unfold_locales auto
        show "Maps.span (Leg0 (Dom (SPN μ))) (Leg1 (Dom (SPN μ)))"
          using assms Maps.CLS_in_hom SPN_def
          by (metis (no_types, lifting) Maps.in_homE bicategory_of_spans.Dom_SPN
              bicategory_of_spans_axioms r.leg1_is_map r.leg1_simps(3) r.satisfies_T0
              span_data.simps(1) span_data.simps(2))
      qed
      interpret Dom': span_in_category Maps.comp 
                        Leg0 = ⟦⟦tab0 (dom μ)⟧⟧, Leg1 = ⟦⟦tab1 (dom μ)⟧⟧
        using assms Dom.span_in_category_axioms SPN_def by simp
      interpret Cod: span_in_category Maps.comp Cod (SPN μ)
      proof
        interpret s: identity_in_bicategory_of_spans V H 𝖺 𝗂 src trg cod μ
          using assms by unfold_locales auto
        show "Maps.span (Leg0 (Cod (SPN μ))) (Leg1 (Cod (SPN μ)))"
          using assms Maps.CLS_in_hom SPN_def
          by (metis (no_types, lifting) bicategory_of_spans.Cod_SPN bicategory_of_spans_axioms
              ide_dom s.base_simps(2) s.base_simps(3) s.determines_span span_in_category.is_span)
      qed
      interpret Cod': span_in_category Maps.comp 
                        Leg0 = ⟦⟦tab0 (cod μ)⟧⟧, Leg1 = ⟦⟦tab1 (cod μ)⟧⟧
        using assms Cod.span_in_category_axioms SPN_def by simp
      show 1: "Span.arr (SPN μ)"
      proof (unfold Span.arr_char)
        show "arrow_of_spans Maps.comp (SPN μ)"
        proof (unfold_locales)
          show "Maps.in_hom (Chn (SPN μ)) Dom.apex Cod.apex"
            unfolding SPN_def Maps.in_hom_char
            using assms Dom'.apex_def Cod'.apex_def Dom'.is_span Cod'.is_span Maps.arr_char
            by auto
          show "Cod.leg0  Chn (SPN μ) = Dom.leg0"
            unfolding SPN_def
            using assms spn_props [of μ] Maps.comp_CLS [of "tab0 (cod μ)" "spn μ"] by simp
          show "Cod.leg1  Chn (SPN μ) = Dom.leg1"
            unfolding SPN_def
            using assms spn_props [of μ] Maps.comp_CLS [of "tab1 (cod μ)" "spn μ"] by simp
        qed
      qed
      show "Span.dom (SPN μ) = SPN0 (dom μ)"
        using assms 1 Span.dom_char Dom'.apex_def SPN_def by simp
      show "Span.cod (SPN μ) = SPN0 (cod μ)"
        using assms 1 Span.cod_char Cod'.apex_def SPN_def by simp
    qed

    interpretation SPN: "functor" V Span.vcomp SPN
    proof
      show "μ. ¬ arr μ  SPN μ = Span.null"
        unfolding SPN_def by simp
      show 1: "μ. arr μ  Span.arr (SPN μ)"
        using SPN_in_hom by auto
      show "μ. arr μ  Span.dom (SPN μ) = SPN (dom μ)"
      proof -
        fix μ
        assume μ: "arr μ"
        have 1: "Maps.arr (Maps.MkArr (src (tab0 (dom μ))) (src μ) tab0 (dom μ))"
        proof -
          have "src (tab0 (dom μ))  Collect obj"
            using μ by simp
          moreover have "src μ  Collect obj"
            using μ by simp
          moreover have "tab0 (dom μ)  Maps.Hom (src (tab0 (local.dom μ))) (src μ)"
          proof -
            have "«tab0 (dom μ) : src (tab0 (dom μ))  src μ»"
              using μ by simp
            moreover have "is_left_adjoint (tab0 (dom μ))"
              using μ tab0_simps [of "dom μ"] by auto
            ultimately show ?thesis by auto
          qed
          ultimately show ?thesis by simp
        qed
        have "spn (dom μ) = src (tab0 (dom μ))"
          using μ spn_ide iso_class_eqI by auto
        hence "SPN (dom μ) = SPN0 (dom μ)"
          unfolding SPN_def
          using μ 1 Maps.dom_char by simp
        thus "Span.dom (SPN μ) = SPN (dom μ)"
          using μ SPN_in_hom by auto
      qed
      show "μ. arr μ  Span.cod (SPN μ) = SPN (cod μ)"
      proof -
        fix μ
        assume μ: "arr μ"
        have 1: "Maps.arr (Maps.MkArr (src (tab0 (cod μ))) (src μ) tab0 (cod μ))"
        proof -
          have "src (tab0 (cod μ))  Collect obj"
            using μ by simp
          moreover have "src μ  Collect obj"
            using μ by simp
          moreover have "tab0 (cod μ)  Maps.Hom (src (tab0 (cod μ))) (src μ)"
          proof -
            have "«tab0 (cod μ) : src (tab0 (cod μ))  src μ»"
              using μ by simp
            moreover have "is_left_adjoint (tab0 (cod μ))"
              using μ by simp
            ultimately show ?thesis by auto
          qed
          ultimately show ?thesis by simp
        qed
        have "spn (cod μ) = src (tab0 (cod μ))"
          using μ spn_ide iso_class_eqI by auto
        hence "SPN (cod μ) = SPN0 (cod μ)"
          unfolding SPN_def
          using μ 1 Maps.dom_char by simp
        thus "Span.cod (SPN μ) = SPN (cod μ)"
          using μ SPN_in_hom by auto
      qed
      show "ν μ. seq ν μ  SPN (ν  μ) = SPN ν  SPN μ"
      proof -
        fix μ ν
        assume seq: "seq ν μ"
        have "Dom (SPN (ν  μ)) = Dom (SPN ν  SPN μ)"
          using seq 1 Span.vcomp_def Span.arr_char
          by (elim seqE, simp)
        moreover have "Cod (SPN (ν  μ)) = Cod (SPN ν  SPN μ)"
          using seq 1 Span.vcomp_def Span.arr_char
          by (elim seqE, simp)
        moreover have "Chn (SPN (ν  μ)) = Chn (SPN ν  SPN μ)"
        proof -
          have *: "spn (ν  μ) = Maps.Comp spn ν spn μ"
          proof
            show "spn (ν  μ)  Maps.Comp spn ν spn μ"
            proof
              fix h
              assume h: "h  spn (ν  μ)"
              show "h  Maps.Comp spn ν spn μ"
              proof -
                have 1: "spn ν  spn ν"
                  using seq ide_in_iso_class by auto
                moreover have 2: "spn μ  spn μ"
                  using seq ide_in_iso_class by auto
                moreover have "spn ν  spn μ  h"
                proof -
                  have "spn ν  spn μ  spn (ν  μ)"
                    using seq spn_hcomp 1 2 iso_class_def isomorphic_reflexive
                          isomorphic_symmetric
                    by simp
                  thus ?thesis
                    using h isomorphic_transitive iso_class_def by simp
                qed
                ultimately show ?thesis
                  unfolding Maps.Comp_def
                  by (metis (mono_tags, lifting) is_iso_classI spn_simps(2)
                      mem_Collect_eq seq seqE)
              qed
            qed
            show "Maps.Comp spn ν spn μ  spn (ν  μ)"
            proof
              fix h
              assume h: "h  Maps.Comp spn ν spn μ"
              show "h  spn (ν  μ)"
              proof -
                obtain f g where 1: "g  spn ν  f  spn μ  g  f  h"
                  using h Maps.Comp_def [of "iso_class (spn ν)" "iso_class (spn μ)"]
                        iso_class_def iso_class_elems_isomorphic
                  by blast
                have fg: "g  spn ν  f  spn μ  g  f  h"
                proof -
                  have "spn ν  spn ν  spn μ  spn μ"
                    using seq ide_in_iso_class by auto
                  thus ?thesis
                    using 1 iso_class_elems_isomorphic isomorphic_symmetric is_iso_classI
                    by (meson spn_simps(2) seq seqE)
                qed
                have "g  f  spn (ν  μ)"
                  using seq fg 1 spn_hcomp iso_class_def isomorphic_symmetric by simp
                thus ?thesis
                  using fg isomorphic_transitive iso_class_def by blast
              qed
            qed
          qed
          have "Chn (SPN ν  SPN μ) =
                Maps.MkArr (src (tab0 (cod μ))) (src (tab0 (cod ν))) spn ν 
                  Maps.MkArr (src (tab0 (dom μ))) (src (tab0 (cod μ))) spn μ"
            using 1 seq SPN_def Span.vcomp_def Span.arr_char
            apply (elim seqE)
            apply simp
            by (metis (no_types, lifting) seq vseq_implies_hpar(1) vseq_implies_hpar(2))
          also have "... = Maps.MkArr (src (tab0 (dom μ))) (src (tab0 (cod ν)))
                                      (Maps.Comp spn ν spn μ)"
          proof -
            have "Maps.seq (Maps.MkArr (src (tab0 (cod μ))) (src (tab0 (cod ν))) spn ν)
                           (Maps.MkArr (src (tab0 (dom μ))) (src (tab0 (cod μ))) spn μ)"
            proof
              show "Maps.in_hom (Maps.MkArr (src (tab0 (local.dom μ))) (src (tab0 (cod μ)))
                                            spn μ)
                                (Maps.MkIde (src (tab0 (dom μ))))
                                (Maps.MkIde (src (tab0 (cod μ))))"
              proof -
                have "src (tab0 (dom μ))  Collect obj"
                  using in_hhom_def seq by auto
                moreover have "src (tab0 (cod μ))  Collect obj"
                  using seq by auto
                moreover have "spn μ  Maps.Hom (src (tab0 (dom μ))) (src (tab0 (cod μ)))"
                  using spn_props
                  by (metis (mono_tags, lifting) mem_Collect_eq seq seqE)
                ultimately show ?thesis
                  using Maps.MkArr_in_hom by simp
              qed
              show "Maps.in_hom (Maps.MkArr (src (tab0 (cod μ))) (src (tab0 (cod ν))) spn ν)
                                (Maps.MkIde (src (tab0 (cod μ))))
                                (Maps.MkIde (src (tab0 (cod ν))))"
              proof -
                have "src (tab0 (cod μ))  Collect obj"
                  using in_hhom_def seq by auto
                moreover have "src (tab0 (cod ν))  Collect obj"
                  using seq by auto
                moreover have "spn ν  Maps.Hom (src (tab0 (cod μ))) (src (tab0 (cod ν)))"
                  using spn_props
                  by (metis (mono_tags, lifting) mem_Collect_eq seq seqE)
                ultimately show ?thesis
                  using Maps.MkArr_in_hom by simp
              qed
            qed
            thus ?thesis
              using Maps.comp_char
                     [of "Maps.MkArr (src (tab0 (cod μ))) (src (tab0 (cod ν))) spn ν"
                         "Maps.MkArr (src (tab0 (dom μ))) (src (tab0 (cod μ))) spn μ"]
              by simp
          qed
          also have "... = Maps.MkArr (src (tab0 (dom μ))) (src (tab0 (cod ν))) spn (ν  μ)"
            using * by simp
          also have "... = Chn (SPN (ν  μ))"
            using seq SPN_def Span.vcomp_def
            by (elim seqE, simp)
          finally show ?thesis by simp
        qed
        ultimately show "SPN (ν  μ) = SPN ν  SPN μ" by simp
      qed
    qed

    lemma SPN_is_functor:
    shows "functor V Span.vcomp SPN"
      ..

    interpretation SPN: weak_arrow_of_homs V src trg Span.vcomp Span.src Span.trg SPN
    proof
      show "μ. arr μ  Span.isomorphic (SPN (src μ)) (Span.src (SPN μ))"
      proof -
        fix μ
        assume μ: "arr μ"
        let ?src = "Maps.MkIde (src μ)"
        have src: "Maps.ide ?src"
          using μ by simp
        interpret src: identity_in_bicategory_of_spans V H 𝖺 𝗂 src trg src μ
          using μ by unfold_locales auto
        interpret src: arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg
                         src μ src.tab tab0 (src μ) tab1 (src μ)
                         src μ src.tab tab0 (src μ) tab1 (src μ)
                         src μ
          using src.is_arrow_of_tabulations_in_maps by simp
        interpret src: span_in_category Maps.comp Leg0 = ?src, Leg1 = ?src
          using src by (unfold_locales, simp)

        let ?tab0 = "Maps.MkArr (src (tab0 (src μ))) (src μ) tab0 (src μ)"
        have tab0_src: "«tab0 (src μ) : src (tab0 (src μ))  src μ» 
                        is_left_adjoint (tab0 (src μ))  tab0 (src μ) = tab0 (src μ)"
          using μ by simp
        have tab0: "Maps.arr ?tab0"
          using μ Maps.arr_MkArr tab0_src by blast
        let ?tab1 = "Maps.MkArr (src (tab0 (src μ))) (src μ) tab1 (src μ)"
        have tab1_src: "«tab1 (src μ) : src (tab0 (src μ))  src μ» 
                        is_left_adjoint (tab1 (src μ))  tab1 (src μ) = tab1 (src μ)"
          using μ by simp
        have tab1: "Maps.arr ?tab1"
          using μ Maps.arr_MkArr tab1_src by blast
        interpret tab: span_in_category Maps.comp Leg0 = ?tab0, Leg1 = ?tab1
          using tab0 tab1 Maps.dom_char Maps.cod_char by (unfold_locales, simp)

        have "src μ  tab0 (src μ)  tab0 (src μ)"
          using μ iso_lunit isomorphic_def
          by (metis lunit_in_hom(2) src.ide_u src.u_simps(3) src_src)
        hence "src μ  tab0 (src μ)  tab1 (src μ)"
          using μ src.obj_has_symmetric_tab isomorphic_transitive by blast

        have "tab0 (src μ)  Maps.Hom (src (tab0 (src μ))) (src μ)"
          using μ tab0_src by blast
        have "src μ  Maps.Hom (src μ) (src μ)"
        proof -
          have "«src μ : src μ  src μ»  is_left_adjoint (src μ)  src μ = src μ"
            using μ obj_is_self_adjoint by simp
          thus ?thesis by auto
        qed

        interpret SPN_src: arrow_of_spans Maps.comp SPN (src μ)
          using μ SPN.preserves_reflects_arr Span.arr_char by blast
        have SPN_src: "SPN (src μ) =
                       Chn = Maps.MkArr (src (tab0 (src μ))) (src (tab0 (src μ))) spn (src μ),
                        Dom = Leg0 = ?tab0, Leg1 = ?tab1,
                        Cod = Leg0 = ?tab0, Leg1 = ?tab1"
          unfolding SPN_def using μ by simp

        interpret src_SPN: arrow_of_spans Maps.comp Span.src (SPN μ)
          using μ SPN.preserves_reflects_arr Span.arr_char by blast
        have src_SPN: "Span.src (SPN μ) =
                       Chn = ?src,
                        Dom = Leg0 = ?src, Leg1 = ?src,
                        Cod = Leg0 = ?src, Leg1 = ?src"
        proof -
          let ?tab0_dom = "Maps.MkArr (src (tab0 (dom μ))) (src μ) tab0 (dom μ)"
          have "Maps.arr ?tab0_dom"
          proof -
            have "«tab0 (dom μ) : src (tab0 (dom μ))  src μ» 
                  is_left_adjoint (tab0 (dom μ))  tab0 (dom μ) = tab0 (dom μ)"
              using μ by simp
            thus ?thesis
              using μ Maps.arr_MkArr by blast
          qed
          thus ?thesis
            using μ Maps.cod_char Span.src_def by simp
        qed

        text ‹
          The idea of the proof is that @{term "iso_class (tab0 (src μ))"} is invertible
          in Maps(B)› and determines an invertible arrow of spans from @{term "SPN (src μ)"}
          to @{term "Span.src (SPN μ)"}.
        ›

        let  = "Chn = ?tab0, Dom = Dom (SPN (src μ)), Cod = Cod (Span.src (SPN μ))"
        interpret φ: arrow_of_spans Maps.comp 
          apply (unfold_locales, simp_all)
        proof -
          show "Maps.in_hom ?tab0 SPN_src.dom.apex src_SPN.cod.apex"
            using μ tab0 Maps.dom_char Maps.cod_char SPN_src src_SPN
                  tab.apex_def src_SPN.cod.apex_def
            apply (intro Maps.in_homI) by simp_all
          show "src_SPN.cod.leg0  ?tab0 = SPN_src.dom.leg0"
          proof -
            have "Maps.seq src_SPN.cod.leg0 ?tab0"
              using μ src_SPN tab0 Maps.dom_char Maps.cod_char
              by (intro Maps.seqI, auto)
            moreover have "Maps.Comp src μ tab0 (src μ) = tab0 (src μ)"
            proof -
              have "tab0 (src μ)  Maps.Comp src μ tab0 (src μ)"
                using μ is_iso_classI ide_in_iso_class [of "src μ"]
                      ide_in_iso_class [of "tab0 (src μ)"] src μ  tab0 (src μ)  tab0 (src μ)
                by auto
              thus ?thesis
                using Maps.Comp_eq_iso_class_memb
                      tab0 (src μ)  Maps.Hom (src (tab0 (src μ))) (src μ)
                      src μ  Maps.Hom (src μ) (src μ)
                by meson
            qed
            ultimately show ?thesis
              using μ Maps.comp_char [of src_SPN.cod.leg0 ?tab0] src_SPN by simp
          qed
          show "src_SPN.cod.leg1  ?tab0 = SPN_src.dom.leg1"
          proof -
            have "Maps.seq src_SPN.cod.leg1 ?tab0"
              using μ src_SPN tab0 Maps.dom_char Maps.cod_char
              by (intro Maps.seqI, auto)
            moreover have "Maps.Comp src μ tab0 (src μ) = tab1 (src μ)"
            proof -
              have "tab1 (src μ)  Maps.Comp src μ tab0 (src μ)"
                using μ is_iso_classI ide_in_iso_class [of "src μ"]
                      ide_in_iso_class [of "tab0 (src μ)"]
                      isomorphic (src μ  tab0 (src μ)) (tab1 (src μ))
                by auto
              thus ?thesis
                using Maps.Comp_eq_iso_class_memb
                      tab0 (src μ)  Maps.Hom (src (tab0 (src μ))) (src μ)
                      src μ  Maps.Hom (src μ) (src μ)
                by meson
            qed
            ultimately show ?thesis
              using μ Maps.comp_char [of src_SPN.cod.leg1 ?tab0] src_SPN by simp
          qed
        qed
        have "Span.in_hom  (SPN (src μ)) (Span.src (SPN μ))"
          using μ tab0 spn_ide [of "src μ"] iso_class_eqI
                Span.arr_char Span.dom_char Span.cod_char φ.arrow_of_spans_axioms
                SPN_src src_SPN src.apex_def tab.apex_def Maps.dom_char
          apply (intro Span.in_homI) by auto
          (* The preceding cannot be written "by (intro Span.in_homI, auto)", why? *)
        moreover have "Maps.iso ?tab0"
          using μ tab0 ide_in_iso_class src.is_map_iff_tab0_is_equivalence obj_is_self_adjoint
                Maps.iso_char' [of ?tab0]
          by auto
        ultimately show "Span.isomorphic (SPN (src μ)) (Span.src (SPN μ))"
          using Span.isomorphic_def Span.iso_char by auto
      qed
      show "μ. arr μ  Span.isomorphic (SPN (trg μ)) (Span.trg (SPN μ))"
      proof -
        fix μ
        assume μ: "arr μ"
        let ?trg = "Maps.MkIde (trg μ)"
        have trg: "Maps.ide ?trg"
          using μ by simp
        interpret trg: identity_in_bicategory_of_spans V H 𝖺 𝗂 src trg trg μ
          using μ by unfold_locales auto
        interpret trg: span_in_category Maps.comp Leg0 = ?trg, Leg1 = ?trg
          using trg by (unfold_locales, simp)

        let ?tab0 = "Maps.MkArr (src (tab0 (trg μ))) (trg μ) tab0 (trg μ)"
        have tab0_trg: "«tab0 (trg μ) : src (tab0 (trg μ))  trg μ» 
                        is_left_adjoint (tab0 (trg μ))  tab0 (trg μ) = tab0 (trg μ)"
          using μ by simp
        have tab0: "Maps.arr ?tab0"
          using μ Maps.arr_MkArr tab0_trg by blast
        let ?tab1 = "Maps.MkArr (src (tab0 (trg μ))) (trg μ) tab1 (trg μ)"
        have tab1_trg: "«tab1 (trg μ) : src (tab0 (trg μ))  trg μ» 
                        is_left_adjoint (tab1 (trg μ))  tab1 (trg μ) = tab1 (trg μ)"
          using μ by simp
        have tab1: "Maps.arr ?tab1"
          using μ Maps.arr_MkArr tab1_trg by blast
        interpret tab: span_in_category Maps.comp Leg0 = ?tab0, Leg1 = ?tab1
          using tab0 tab1 Maps.dom_char Maps.cod_char by (unfold_locales, simp)

        have "trg μ  tab1 (trg μ)  tab0 (trg μ)"
        proof -
          have "«𝗅[tab1 (trg μ)] : trg μ  tab1 (trg μ)  tab1 (trg μ)»"
            using μ by simp
          moreover have "iso 𝗅[tab1 (trg μ)]"
            using μ iso_lunit by simp
          ultimately have "trg μ  tab1 (trg μ)  tab1 (trg μ)"
            using isomorphic_def by auto
          also have "tab1 (trg μ)  tab0 (trg μ)"
            using μ trg.obj_has_symmetric_tab isomorphic_symmetric by auto
          finally show ?thesis by blast
        qed
        hence "trg μ  tab1 (trg μ)  tab1 (trg μ)"
          using μ trg.obj_has_symmetric_tab isomorphic_transitive by blast

        have "tab1 (trg μ)  Maps.Hom (src (tab0 (trg μ))) (trg μ)"
        proof -
          have "«tab1 (trg μ) : src (tab0 (trg μ))  trg μ»  is_left_adjoint (tab0 (trg μ)) 
                tab0 (trg μ) = tab0 (trg μ)"
            using μ by simp
          thus ?thesis by auto
        qed
        have "trg μ  Maps.Hom (trg μ) (trg μ)"
        proof -
          have "«trg μ : trg μ  trg μ»  is_left_adjoint (trg μ)  trg μ = trg μ"
            using μ obj_is_self_adjoint by simp
          thus ?thesis by auto
        qed

        interpret SPN_trg: arrow_of_spans Maps.comp SPN (trg μ)
          using μ SPN.preserves_reflects_arr Span.arr_char by blast
        have SPN_trg: "SPN (trg μ) =
                       Chn = Maps.MkArr (src (tab1 (trg μ))) (src (tab1 (trg μ))) spn (trg μ),
                        Dom = Leg0 = ?tab0, Leg1 = ?tab1,
                        Cod = Leg0 = ?tab0, Leg1 = ?tab1"
          unfolding SPN_def using μ by simp

        interpret trg_SPN: arrow_of_spans Maps.comp Span.trg (SPN μ)
          using μ SPN.preserves_reflects_arr Span.arr_char by blast
        have trg_SPN: "Span.trg (SPN μ) = Chn = ?trg,
                                           Dom = Leg0 = ?trg, Leg1 = ?trg,
                                           Cod = Leg0 = ?trg, Leg1 = ?trg"
        proof -
          let ?tab1_dom = "Maps.MkArr (src (tab1 (dom μ))) (trg μ) tab1 (dom μ)"
          have "Maps.arr ?tab1_dom"
          proof -
            have "«tab1 (dom μ) : src (tab1 (dom μ))  trg μ» 
                  is_left_adjoint (tab1 (dom μ))  tab1 (dom μ) = tab1 (dom μ)"
              using μ by simp
            thus ?thesis
              using μ Maps.arr_MkArr by blast
          qed
          thus ?thesis
            using μ Maps.cod_char Span.trg_def by simp
        qed

        let  = "Chn = ?tab1, Dom = Dom (SPN (trg μ)), Cod = Cod (Span.trg (SPN μ))"
        interpret φ: arrow_of_spans Maps.comp 
          apply (unfold_locales, simp_all)
        proof -
          show "Maps.in_hom ?tab1 SPN_trg.dom.apex trg_SPN.cod.apex"
            using μ tab0 tab1 Maps.dom_char Maps.cod_char SPN_trg trg_SPN
                  tab.apex_def trg_SPN.cod.apex_def
            apply (intro Maps.in_homI) by simp_all
            (* The preceding cannot be written "by (intro Maps.in_homI, simp_all)", why? *)
          show "Maps.comp trg_SPN.cod.leg0 ?tab1 = SPN_trg.dom.leg0"
          proof -
            have "Maps.seq trg_SPN.cod.leg0 ?tab1"
              using μ trg_SPN tab1 Maps.dom_char Maps.cod_char
              by (intro Maps.seqI, auto)
            moreover have "Maps.Comp trg μ tab1 (trg μ) = tab1 (trg μ)"
            proof -
              have "tab1 (trg μ)  Maps.Comp trg μ tab1 (trg μ)"
                using μ is_iso_classI ide_in_iso_class [of "trg μ"]
                      ide_in_iso_class [of "tab1 (trg μ)"] trg μ  tab1 (trg μ)  tab1 (trg μ)
                by auto
              thus ?thesis
                using Maps.Comp_eq_iso_class_memb
                      iso_class (tab1 (trg μ))  Maps.Hom (src (tab0 (trg μ))) (trg μ)
                      iso_class (trg μ)  Maps.Hom (trg μ) (trg μ)
                by meson
            qed
            moreover have "tab1 (trg μ) = tab0 (trg μ)"
              using μ iso_class_eqI trg.obj_has_symmetric_tab by auto
            ultimately show ?thesis
              using μ Maps.comp_char [of trg_SPN.cod.leg0 ?tab1] trg_SPN
              by simp
          qed
          show "trg_SPN.cod.leg1  ?tab1 = SPN_trg.dom.leg1"
          proof -
            have "Maps.seq trg_SPN.cod.leg1 ?tab1"
              using μ trg_SPN tab1 Maps.dom_char Maps.cod_char
              by (intro Maps.seqI, auto)
            moreover have "Maps.Comp trg μ tab1 (trg μ) = tab1 (trg μ)"
            proof -
              have "tab1 (trg μ)  Maps.Comp trg μ tab1 (trg μ)"
                using μ is_iso_classI ide_in_iso_class [of "trg μ"]
                      ide_in_iso_class [of "tab1 (trg μ)"] trg μ  tab1 (trg μ)  tab1 (trg μ)
                by auto
              thus ?thesis
                using Maps.Comp_eq_iso_class_memb
                      tab1 (trg μ)  Maps.Hom (src (tab0 (trg μ))) (trg μ)
                      trg μ  Maps.Hom (trg μ) (trg μ)
                by meson
            qed
            ultimately show ?thesis
              using μ Maps.comp_char [of trg_SPN.cod.leg1 ?tab1] trg_SPN by simp
          qed
        qed
        have φ: "Span.in_hom  (SPN (trg μ)) (Span.trg (SPN μ))"
          using μ tab0 spn_ide [of "trg μ"] iso_class_eqI
                Span.arr_char Span.dom_char Span.cod_char φ.arrow_of_spans_axioms
                SPN_trg trg_SPN trg.apex_def tab.apex_def Maps.dom_char
          apply (intro Span.in_homI) by auto
        have "Maps.iso ?tab1"
        proof -
          have "Maps.iso ?tab0"
            using μ tab0 ide_in_iso_class trg.is_map_iff_tab0_is_equivalence obj_is_self_adjoint
                  Maps.iso_char' [of ?tab0]
            by auto
          moreover have "?tab0 = ?tab1"
          proof -
            have "tab0 (trg μ) = tab1 (trg μ)"
              using μ iso_class_eqI trg.obj_has_symmetric_tab by auto
            thus ?thesis by simp
          qed
          ultimately show ?thesis by simp
        qed
        thus "Span.isomorphic (SPN (trg μ)) (Span.trg (SPN μ))"
          using φ Span.isomorphic_def Span.iso_char by auto
      qed
    qed

    lemma SPN_is_weak_arrow_of_homs:
    shows "weak_arrow_of_homs V src trg Span.vcomp Span.src Span.trg SPN"
      ..

  end

  subsubsection "Compositors"

  text ‹
    To complete the proof that SPN› is a pseudofunctor, we need to obtain a natural
    isomorphism Φ›, whose component at (r, s)› is an isomorphism Φ (r, s)›
    from the horizontal composite SPN r ∘ SPN s› to SPN (r ⋆ s)› in Span(Maps(B))›,
    and we need to prove that the coherence conditions are satisfied.

    We have shown that the tabulations of r› and s› compose to yield a tabulation of r ⋆ s›.
    Since tabulations of the same arrow are equivalent, this tabulation must be equivalent
    to the chosen tabulation of r ⋆ s›.  We therefore obtain an equivalence map from the
    apex of SPN r ∘ SPN s› to the apex of SPN (r ⋆ s)› which commutes with the
    legs of these spans up to isomorphism.  This equivalence map determines an invertible
    arrow in Span(Maps(B))›.  Moreover, by property T2›, any two such equivalence maps are
    connected by a unique 2-cell, which is consequently an isomorphism.  This shows that
    the arrow in Span(Maps(B))› is uniquely determined, which fact we can exploit to establish
    the required coherence conditions.
  ›

  locale two_composable_identities_in_bicategory_of_spans =
    bicategory_of_spans V H 𝖺 𝗂 src trg +
    r: identity_in_bicategory_of_spans V H 𝖺 𝗂 src trg r +
    s: identity_in_bicategory_of_spans V H 𝖺 𝗂 src trg s
  for V :: "'a comp"                 (infixr "" 55)
  and H :: "'a  'a  'a"          (infixr "" 53)
  and 𝖺 :: "'a  'a  'a  'a"     ("𝖺[_, _, _]")
  and 𝗂 :: "'a  'a"                 ("𝗂[_]")
  and src :: "'a  'a"
  and trg :: "'a  'a"
  and r :: 'a
  and s :: 'a +
  assumes composable: "src r = trg s"
  begin

    notation isomorphic (infix "" 50)

    interpretation r: arrow_in_bicategory_of_spans V H 𝖺 𝗂 src trg r r r
      by unfold_locales auto
    interpretation r: arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg
                        r r.tab tab0 r tab1 r
                        r r.tab tab0 r tab1 r
                        r
      using r.is_arrow_of_tabulations_in_maps by simp
    interpretation s: arrow_in_bicategory_of_spans V H 𝖺 𝗂 src trg s s s
      by unfold_locales auto
    interpretation s: arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg
                        s s.tab tab0 s tab1 s
                        s s.tab tab0 s tab1 s
                        s
      using s.is_arrow_of_tabulations_in_maps by simp

    sublocale identity_in_bicategory_of_spans V H 𝖺 𝗂 src trg r  s
      apply unfold_locales by (simp add: composable)
    sublocale horizontal_composite_of_arrows_of_tabulations_in_maps V H 𝖺 𝗂 src trg
                r r.tab tab0 r tab1 r s s.tab tab0 s tab1 s
                r r.tab tab0 r tab1 r s s.tab tab0 s tab1 s
                r s
      using composable by unfold_locales auto

    abbreviation p0 where "p0  ρσ.p0"
    abbreviation p1 where "p1  ρσ.p1"

    text ‹
      We will take as the composition isomorphism from SPN r ∘ SPN s› to SPN (r ⋆ s)›
      the arrow of tabulations, induced by the identity r ⋆ s›, from the composite of
      the chosen tabulations of r› and s› to the chosen tabulation of r ⋆ s›.
      As this arrow of tabulations is induced by an identity, it is an equivalence map.
    ›

    interpretation cmp: identity_arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg
                          r  s ρσ.tab tab0 s  ρσ.p0 tab1 r  ρσ.p1
                          r  s tab tab0 (r  s) tab1 (r  s)
                          r  s
      using composable
      by unfold_locales auto

    lemma cmp_interpretation:
    shows "identity_arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg
                          (r  s) ρσ.tab (tab0 s  ρσ.p0) (tab1 r  ρσ.p1)
                          (r  s) tab (tab0 (r  s)) (tab1 (r  s))
                          (r  s)"
      ..

    definition cmp
    where "cmp = cmp.chine"

    lemma cmp_props:
    shows "«cmp : src ρσ.tab  src tab»"
    and "«cmp : cmp  cmp»"
    and "equivalence_map cmp"
    and "tab0 (r  s)  cmp  tab0 s  ρσ.p0"
    and "tab1 (r  s)  cmp  tab1 r  ρσ.p1"
      using cmp_def cmp.leg0_uniquely_isomorphic(1) cmp.leg1_uniquely_isomorphic(1)
            isomorphic_symmetric cmp.chine_is_equivalence
      by auto

    lemma cmp_in_hom [intro]:
    shows "«cmp : src ρσ.tab  src tab»"
    and "«cmp : cmp  cmp»"
      using cmp_props by auto

    lemma cmp_simps [simp]:
    shows "arr cmp" and "ide cmp"
    and "src cmp = src ρσ.tab" and "trg cmp = src tab"
    and "dom cmp = cmp" and "cod cmp = cmp"
      using cmp_props equivalence_map_is_ide by auto

    text ‹
      Now we have to use the above properties of the underlying bicategory to
      exhibit the composition isomorphisms as actual arrows in Span(Maps(B))›
      and to prove the required naturality and coherence conditions.
    ›

    interpretation Maps: maps_category V H 𝖺 𝗂 src trg ..
    interpretation Span: span_bicategory Maps.comp Maps.PRJ0 Maps.PRJ1 ..

    no_notation Fun.comp (infixl "" 55)
    notation Span.vcomp (infixr "" 55)
    notation Span.hcomp (infixr "" 53)
    notation Maps.comp (infixr "" 55)

    interpretation SPN: "functor" V Span.vcomp SPN
      using SPN_is_functor by simp
    interpretation SPN: weak_arrow_of_homs V src trg Span.vcomp Span.src Span.trg SPN
      using SPN_is_weak_arrow_of_homs by simp

    interpretation SPN_r_SPN_s: arrow_of_spans Maps.comp SPN r  SPN s
      using composable Span.ide_char [of "SPN r  SPN s"] by simp
    interpretation SPN_r_SPN_s: identity_arrow_of_spans Maps.comp SPN r  SPN s
      using composable Span.ide_char [of "SPN r  SPN s"]
      by (unfold_locales, simp)
    interpretation SPN_rs: arrow_of_spans Maps.comp SPN (r  s)
      using composable Span.arr_char r.base_simps(2) s.base_simps(2) by blast
    interpretation SPN_rs: identity_arrow_of_spans Maps.comp SPN (r  s)
      using composable Span.ide_char SPN.preserves_ide r.is_ide s.is_ide
      by (unfold_locales, simp)

    text ‹
      The following are the legs (as arrows of Maps›) of the spans SPN r› and SPN s›.
    ›

    definition R0 where "R0 = ⟦⟦tab0 r⟧⟧"
    definition R1 where "R1 = ⟦⟦tab1 r⟧⟧"
    definition S0 where "S0 = ⟦⟦tab0 s⟧⟧"
    definition S1 where "S1 = ⟦⟦tab1 s⟧⟧"

    lemma span_legs_eq:
    shows "Leg0 (Dom (SPN r)) = R0" and "Leg1 (Dom (SPN r)) = R1"
    and "Leg0 (Dom (SPN s)) = S0" and "Leg1 (Dom (SPN s)) = S1"
      using SPN_def R0_def R1_def S0_def S1_def composable by auto

    lemma R0_in_hom [intro]:
    shows "Maps.in_hom R0 (Maps.MkIde (src r.s0)) (Maps.MkIde (src r))"
      by (simp add: Maps.MkArr_in_hom' R0_def)

    lemma R1_in_hom [intro]:
    shows "Maps.in_hom R1 (Maps.MkIde (src r.s0)) (Maps.MkIde (trg r))"
      by (simp add: Maps.MkArr_in_hom' R1_def)

    lemma S0_in_hom [intro]:
    shows "Maps.in_hom S0 (Maps.MkIde (src s.s0)) (Maps.MkIde (src s))"
      by (simp add: Maps.MkArr_in_hom' S0_def)

    lemma S1_in_hom [intro]:
    shows "Maps.in_hom S1 (Maps.MkIde (src s.s0)) (Maps.MkIde (trg s))"
      by (simp add: Maps.MkArr_in_hom' S1_def)

    lemma RS_simps [simp]:
    shows "Maps.arr R0" and "Maps.dom R0 = Maps.MkIde (src r.s0)"
    and "Maps.cod R0 = Maps.MkIde (src r)"
    and "Maps.Dom R0 = src r.s0" and "Maps.Cod R0 = src r"
    and "Maps.arr R1" and "Maps.dom R1 = Maps.MkIde (src r.s0)"
    and "Maps.cod R1 = Maps.MkIde (trg r)"
    and "Maps.Dom R1 = src r.s0" and "Maps.Cod R1 = trg r"
    and "Maps.arr S0" and "Maps.dom S0 = Maps.MkIde (src s.s0)"
    and "Maps.cod S0 = Maps.MkIde (src s)"
    and "Maps.Dom S0 = src s.s0" and "Maps.Cod S0 = src s"
    and "Maps.arr S1" and "Maps.dom S1 = Maps.MkIde (src s.s0)"
    and "Maps.cod S1 = Maps.MkIde (trg s)"
    and "Maps.Dom S1 = src s.s0" and "Maps.Cod S1 = trg s"
      using R0_in_hom R1_in_hom S0_in_hom S1_in_hom composable
      by (auto simp add: R0_def R1_def S0_def S1_def)

    text ‹
      The apex of the composite span @{term "SPN r  SPN s"} (defined in terms of pullback)
      coincides with the apex of the composite tabulation ρσ› (defined using
      the chosen tabulation of (tab0 r)* ⋆ tab1 s)›).  We need this to be true in order
      to define the compositor of a pseudofunctor from the underlying bicategory B›
      to Span(Maps(B))›.  It is only true if we have carefully chosen pullbacks in Maps(B)›
      in order to ensure the relationship with the chosen tabulations.
    ›

    lemma SPN_r_SPN_s_apex_eq:
    shows "SPN_r_SPN_s.apex = Maps.MkIde (src ρσ.tab)"
    proof -
      have "obj (Maps.Cod SPN_r_SPN_s.leg0)"
        using Maps.arr_char [of "SPN_r_SPN_s.leg0"] by simp
      moreover have "obj (Maps.Dom SPN_r_SPN_s.leg0)"
        using Maps.arr_char [of "SPN_r_SPN_s.leg0"] by simp
      moreover have "SPN_r_SPN_s.leg0  Maps.Null"
        using Maps.arr_char [of "SPN_r_SPN_s.leg0"] by simp
      moreover have "Maps.Dom SPN_r_SPN_s.leg0 = src ρσ.tab"
      proof -
        interpret REP_S1: map_in_bicategory V H 𝖺 𝗂 src trg Maps.REP S1
          using Maps.REP_in_Map Maps.arr_char Maps.in_HomD S1_def
          apply unfold_locales
          by (meson Maps.REP_in_hhom(2) S1_in_hom)
        interpret REP_R0: map_in_bicategory V H 𝖺 𝗂 src trg Maps.REP R0
          using Maps.REP_in_Map Maps.arr_char Maps.in_HomD R0_def
          apply unfold_locales
          by (meson Maps.REP_in_hhom(2) R0_in_hom)
        have "Maps.Dom SPN_r_SPN_s.leg0 = Maps.Dom (S0  Maps.PRJ0 R0 S1)"
          using composable Span.hcomp_def S0_def R0_def S1_def by simp
        also have "... = Maps.Dom ⟦⟦tab0 ((Maps.REP R0)*  Maps.REP S1)⟧⟧"
        proof -
          have "is_left_adjoint (tab0 ((Maps.REP R0)*  Maps.REP S1))"
          proof -
            have "ide ((Maps.REP R0)*  Maps.REP S1)"
            proof -
              have "src (Maps.REP R0)* = trg (Maps.REP S1)"
                using REP_R0.is_map REP_S1.is_map left_adjoint_is_ide R0_def S1_def
                by (metis (no_types, lifting) Maps.REP_CLS REP_R0.antipar(2)
                    isomorphic_implies_hpar(4) composable r.leg0_simps(3)
                    r.satisfies_T0 s.leg1_is_map s.leg1_simps(3) s.leg1_simps(4))
              thus ?thesis by simp
            qed
            thus ?thesis by simp
          qed
          moreover have "Maps.Dom (S0  ⟦⟦tab0 ((Maps.REP R0)*  Maps.REP S1)⟧⟧) =
                         src (tab0 ((Maps.REP R0)*  Maps.REP S1))"
          proof -
            have "Maps.arr (⟦⟦prj0 (Maps.REP S1) (Maps.REP R0)⟧⟧)"
              using Maps.CLS_in_hom Maps.prj0_simps(1) Maps.PRJ0_def composable by fastforce
            moreover have "Maps.Dom S0 = Maps.Cod ⟦⟦prj0 (Maps.REP S1) (Maps.REP R0)⟧⟧"
            proof -
              have "Maps.Cod ⟦⟦prj0 (Maps.REP S1) (Maps.REP R0)⟧⟧ =
                    trg (prj0 (Maps.REP S1) (Maps.REP R0))"
                by simp
              also have "... = src (Maps.REP S1)"
              proof -
                have "ide (Maps.REP S1)"
                  by simp
                moreover have "is_left_adjoint (Maps.REP R0)"
                  by auto
                moreover have "trg (Maps.REP S1) = trg (Maps.REP R0)"
                  by (simp add: composable)
                ultimately show ?thesis
                  using S1_def Maps.REP_CLS r.leg0_is_map s.leg1_is_map by simp
              qed
              also have "... = src (tab0 s)"
                using tab0_in_hom(1) by simp
              also have "... = Maps.Dom S0"
                using S0_def by simp
              finally show ?thesis by simp
            qed
            ultimately have
              "Maps.Dom (S0  ⟦⟦tab0 ((Maps.REP R0)*  Maps.REP S1)⟧⟧) =
               Maps.Dom ⟦⟦tab0 ((Maps.REP R0)*  Maps.REP S1)⟧⟧"
              using Maps.CLS_in_hom by simp
            thus ?thesis by simp
          qed
          ultimately show ?thesis
            using Maps.PRJ0_def composable Maps.Dom.simps(1) RS_simps(1) RS_simps(16)
                  RS_simps(18) RS_simps(3)
            by presburger
        qed
        also have "... = src (tab0 ((Maps.REP R0)*  Maps.REP S1))"
          by simp
        finally have
          "Maps.Dom SPN_r_SPN_s.leg0 = src (tab0 ((Maps.REP R0)*  Maps.REP S1))"
          by simp
        also have "... = src (tab0 (r.s0*  s.s1))"
        proof -
          interpret r0's1: identity_in_bicategory_of_spans V H 𝖺 𝗂 src trg r.s0*  s.s1
            using composable by (unfold_locales, simp)
          have "(Maps.REP R0)*  Maps.REP S1  r.s0*  s.s1"
          proof -
            have "(Maps.REP R0)*  r.s0*"
            proof -
              have 1: "adjoint_pair (Maps.REP R0) (Maps.REP R0)*"
                using REP_R0.is_map left_adjoint_extends_to_adjoint_pair by blast
              moreover have "adjoint_pair r.s0 (Maps.REP R0)*"
              proof -
                have "Maps.REP R0  r.s0"
                  unfolding R0_def
                  using Maps.REP_CLS r.leg0_is_map composable by force
                thus ?thesis
                  using 1 adjoint_pair_preserved_by_iso isomorphic_def
                        REP_R0.triangle_in_hom(4) REP_R0.triangle_right'
                  by auto
              qed
              ultimately show ?thesis
                using r.leg0_is_map left_adjoint_determines_right_up_to_iso
                      left_adjoint_extends_to_adjoint_pair
                by auto
            qed
            moreover have "Maps.REP S1  s.s1"
              unfolding S1_def
              using Maps.REP_CLS s.leg1_is_map composable by force
            moreover have "a. a  (tab0 r)*  tab1 s  (Maps.REP R0)*  Maps.REP S1  a"
              using calculation composable isomorphic_implies_hpar(3)
                    hcomp_ide_isomorphic hcomp_isomorphic_ide [of "(Maps.REP R0)*" "r.s0*" s.s1]
              by auto
            ultimately show ?thesis
              using isomorphic_transitive by blast
          qed
          thus ?thesis
            using r0's1.isomorphic_implies_same_tab isomorphic_symmetric by metis
        qed
        also have "... = src ρσ.tab"
          using VV.ide_charSbC VV.arr_charSbC composable Span.hcomp_def ρσ.tab_simps(2) by auto
        finally show ?thesis by simp
      qed
      ultimately show ?thesis
        using composable Maps.arr_char Maps.dom_char SPN_r_SPN_s.dom.apex_def
        apply auto
        by (metis (no_types, lifting) Maps.not_arr_null SPN_r_SPN_s.chine_eq_apex
            SPN_r_SPN_s.chine_simps(1))
    qed

    text ‹
      We will be taking the arrow @{term "CLS cmp"} of Maps› as the composition isomorphism from
      @{term "SPN r  SPN s"} to @{term "SPN (r  s)"}.  The following result shows that it
      has the right domain and codomain for that purpose.
    ›

    lemma iso_class_cmp_in_hom:
    shows "Maps.in_hom (Maps.MkArr (src ρσ.tab) (src tab) cmp)
                       SPN_r_SPN_s.apex SPN_rs.apex"
    and "Maps.in_hom ⟦⟦cmp⟧⟧ SPN_r_SPN_s.apex SPN_rs.apex"
    proof -
      show "Maps.in_hom (Maps.MkArr (src ρσ.tab) (src tab) cmp)
                        SPN_r_SPN_s.apex SPN_rs.apex"
      proof -
        have "obj (src ρσ.tab)"
          using obj_src ρσ.tab_in_hom by blast
        moreover have "obj (src tab)"
          using obj_src by simp
        moreover have "cmp  Maps.Hom (src ρσ.tab) (src tab)"
          by (metis (mono_tags, lifting) cmp.is_map cmp_def cmp_props(1) mem_Collect_eq)
        moreover have "SPN_r_SPN_s.apex = Maps.MkIde (src ρσ.tab)"
          using SPN_r_SPN_s_apex_eq by simp
        moreover have "SPN_rs.apex = Maps.MkIde (src tab)"
          using SPN_def composable SPN_rs.cod.apex_def Maps.arr_char Maps.dom_char
                SPN_rs.dom.leg_simps(1)
          by fastforce
        ultimately show ?thesis
          using Maps.MkArr_in_hom by simp
      qed
      thus "Maps.in_hom ⟦⟦cmp⟧⟧ SPN_r_SPN_s.apex SPN_rs.apex" by simp
    qed

    interpretation r0's1: two_composable_identities_in_bicategory_of_spans
                            V H 𝖺 𝗂 src trg (Maps.REP R0)* Maps.REP S1
    proof
      show "ide (Maps.REP S1)"
        using Maps.REP_in_Map Maps.arr_char left_adjoint_is_ide
        by (meson Maps.REP_in_hhom(2) S1_in_hom)
      show "ide (Maps.REP R0)*"
        using Maps.REP_in_Map Maps.arr_char left_adjoint_is_ide
              Maps.REP_in_hhom(2) R0_in_hom by auto
      show "src (Maps.REP R0)* = trg (Maps.REP S1)"
        using Maps.REP_in_hhom(2) R0_in_hom composable by auto
    qed

    interpretation R0'S1: identity_in_bicategory_of_spans V H 𝖺 𝗂 src trg (tab0 r)*  tab1 s
      by (unfold_locales, simp add: composable)

    lemma prj_tab_agreement:
    shows "(tab0 r)*  tab1 s  (Maps.REP R0)*  Maps.REP S1"
    and "ρσ.p0  prj0 (Maps.REP S1) (Maps.REP R0)"
    and "ρσ.p1  prj1 (Maps.REP S1) (Maps.REP R0)"
    proof -
      show 1: "(tab0 r)*  tab1 s  (Maps.REP R0)*  Maps.REP S1"
      proof -
        have "(tab0 r)*  (Maps.REP R0)*"
          using Maps.REP_CLS isomorphic_symmetric R0_def composable r.satisfies_T0
                isomorphic_to_left_adjoint_implies_isomorphic_right_adjoint
          by fastforce
        moreover have "tab1 s  Maps.REP S1"
          by (metis Maps.REP_CLS isomorphic_symmetric S1_def s.leg1_is_map s.leg1_simps(3-4))
        moreover have "src (Maps.REP R0)* = trg (tab1 s)"
          using composable r.T0.antipar right_adjoint_simps(2) by fastforce
        ultimately show ?thesis
          using hcomp_isomorphic_ide [of "(tab0 r)*" "(Maps.REP R0)*" "tab1 s"]
                hcomp_ide_isomorphic isomorphic_transitive composable
          by auto
      qed
      show "ρσ.p0  tab0 ((Maps.REP R0)*  Maps.REP S1)"
        using 1 R0'S1.isomorphic_implies_same_tab isomorphic_reflexive by auto
      show "ρσ.p1  tab1 ((Maps.REP R0)*  Maps.REP S1)"
        using 1 R0'S1.isomorphic_implies_same_tab isomorphic_reflexive by auto
    qed

    lemma chine_hcomp_SPN_SPN:
    shows "Span.chine_hcomp (SPN r) (SPN s) = Maps.MkIde (src ρσ.p0)"
    proof -
      have "Span.chine_hcomp (SPN r) (SPN s) =
            Maps.MkIde (src (tab0 ((Maps.REP R0)*  Maps.REP S1)))"
        using Span.chine_hcomp_ide_ide [of "SPN r" "SPN s"] composable
              Maps.pbdom_def Maps.PRJ0_def Maps.CLS_in_hom Maps.dom_char R0_def S1_def
        apply simp
        using Maps.prj0_simps(1) RS_simps(1) RS_simps(16) RS_simps(18) RS_simps(3)
        by presburger
      also have "... = Maps.MkIde (src ρσ.p0)"
        using prj_tab_agreement isomorphic_implies_hpar(3) by force
      finally show ?thesis by simp
    qed

  end

  text ‹
    The development above focused on two specific composable 1-cells in bicategory B›.
    Here we reformulate those results as statements about the entire bicategory.
  ›

  context bicategory_of_spans
  begin

    interpretation Maps: maps_category V H 𝖺 𝗂 src trg ..
    interpretation Span: span_bicategory Maps.comp Maps.PRJ0 Maps.PRJ1 ..

    no_notation Fun.comp (infixl "" 55)
    notation Span.vcomp (infixr "" 55)
    notation Span.hcomp (infixr "" 53)
    notation Maps.comp (infixr "" 55)
    notation isomorphic (infix "" 50)

    interpretation SPN: "functor" V Span.vcomp SPN
      using SPN_is_functor by simp
    interpretation SPN: weak_arrow_of_homs V src trg Span.vcomp Span.src Span.trg SPN
      using SPN_is_weak_arrow_of_homs by simp

    interpretation HoSPN_SPN: composite_functor VV.comp Span.VV.comp Span.vcomp
                                SPN.FF λμν. fst μν  snd μν
      ..
    interpretation SPNoH: composite_functor VV.comp V
                            Span.vcomp λμν. fst μν  snd μν SPN
      ..

    text ‹
      Given arbitrary composable 1-cells r› and s›, obtain an arrow of spans in Maps›
      having the isomorphism class of rs.cmp› as its chine.
    ›

    definition CMP
    where "CMP r s 
           Chn = ⟦⟦two_composable_identities_in_bicategory_of_spans.cmp V H 𝖺 𝗂 src trg r s⟧⟧,
            Dom = Dom (SPN r  SPN s), Cod = Dom (SPN (r  s))"

    lemma compositor_in_hom [intro]:
    assumes "ide r" and "ide s" and "src r = trg s"
    shows "Span.in_hhom (CMP r s) (SPN.map0 (src s)) (SPN.map0 (trg r))"
    and "Span.in_hom (CMP r s) (HoSPN_SPN.map (r, s)) (SPNoH.map (r, s))"
    proof -
      have rs: "VV.ide (r, s)"
        using assms VV.ide_charSbC VV.arr_charSbC by simp
      interpret rs: two_composable_identities_in_bicategory_of_spans V H 𝖺 𝗂 src trg r s
        using rs VV.ide_charSbC VV.arr_charSbC by unfold_locales auto
      interpret cmp: arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg
                     r  s rs.ρσ.tab tab0 s  rs.ρσ.p0 tab1 r  rs.ρσ.p1
                     r  s rs.tab tab0 (r  s) tab1 (r  s)
                     r  s
        by unfold_locales auto
      have "rs.cmp = cmp.chine"
        using rs.cmp_def by simp

      interpret SPN_r_SPN_s: arrow_of_spans Maps.comp SPN r  SPN s
        using rs.composable Span.ide_char [of "SPN r  SPN s"] by simp
      interpret SPN_r_SPN_s: identity_arrow_of_spans Maps.comp SPN r  SPN s
        using rs.composable Span.ide_char [of "SPN r  SPN s"]
        by (unfold_locales, simp)
      interpret SPN_rs: arrow_of_spans Maps.comp SPN (r  s)
        using Span.arr_char rs.is_ide SPN.preserves_arr by blast
      interpret SPN_rs: identity_arrow_of_spans Maps.comp SPN (r  s)
        using Span.ide_char rs.is_ide SPN.preserves_ide
        by (unfold_locales, simp)

      interpret Dom: span_in_category Maps.comp Dom (CMP r s)
        by (unfold_locales, simp add: CMP_def)
      interpret Cod: span_in_category Maps.comp Cod (CMP r s)
      proof -
        (* TODO: I don't understand what makes this so difficult. *)
        have "«tab0 (r  s) : src (tab0 (r  s))  src s»  is_left_adjoint (tab0 (r  s)) 
              tab0 (r  s) = tab0 (r  s)"
          by simp
        hence "f. «f : src (tab0 (r  s))  src s»  is_left_adjoint f  tab0 (r  s) = f"
          by blast
        moreover have "f. «f : src (tab0 (r  s))  trg r»  is_left_adjoint f 
                           tab1 (r  s) = f"
          by (metis rs.base_simps(2) rs.leg1_in_hom(1) rs.leg1_is_map trg_hcomp)
        ultimately show "span_in_category Maps.comp (Cod (CMP r s))"
          using assms Maps.arr_char Maps.dom_char CMP_def
          by unfold_locales auto
      qed

      interpret r0's1: two_composable_identities_in_bicategory_of_spans
                         V H 𝖺 𝗂 src trg (Maps.REP rs.R0)* Maps.REP rs.S1
      proof
        show "ide (Maps.REP rs.S1)"
          using Maps.REP_in_Map Maps.arr_char left_adjoint_is_ide
          by (meson Maps.REP_in_hhom(2) rs.S1_in_hom)
        show "ide (Maps.REP rs.R0)*"
          using Maps.REP_in_Map Maps.arr_char left_adjoint_is_ide
                Maps.REP_in_hhom(2) rs.R0_in_hom by auto
        show "src (Maps.REP rs.R0)* = trg (Maps.REP rs.S1)"
          using Maps.REP_in_hhom(2) rs.R0_in_hom rs.composable by auto
      qed

      interpret R0'S1: identity_in_bicategory_of_spans V H 𝖺 𝗂 src trg (tab0 r)*  tab1 s
        by (unfold_locales, simp add: rs.composable)

      text ‹
        Here we obtain explicit formulas for the legs and apex of SPN_r_SPN_s› and SPN_rs›.
      ›

      have SPN_r_SPN_s_leg0_eq:
             "SPN_r_SPN_s.leg0 = Maps.comp rs.S0 (Maps.PRJ0 rs.R0 rs.S1)"
        using rs.composable Span.hcomp_def rs.S0_def rs.R0_def rs.S1_def by simp
      have SPN_r_SPN_s_leg1_eq:
             "SPN_r_SPN_s.leg1 = Maps.comp rs.R1 (Maps.PRJ1 rs.R0 rs.S1)"
        using rs.composable Span.hcomp_def rs.R1_def rs.R0_def rs.S1_def by simp
      have "SPN_r_SPN_s.apex = Maps.MkIde (src rs.ρσ.tab)"
        using rs.SPN_r_SPN_s_apex_eq by auto

      have SPN_rs_leg0_eq: "SPN_rs.leg0 = ⟦⟦tab0 (r  s)⟧⟧"
        unfolding SPN_def using rs by simp
      have SPN_rs_leg1_eq: "SPN_rs.leg1 = ⟦⟦tab1 (r  s)⟧⟧"
        unfolding SPN_def using rs by simp
      have "SPN_rs.apex = Maps.MkIde (src (tab_of_ide (r  s)))"
        using SPN_rs.dom.apex_def Maps.dom_char SPN_rs_leg0_eq SPN_rs.dom.leg_simps(1)
        by simp

      text ‹
        The composition isomorphism @{term "CMP r s"} is an arrow of spans in Maps(B)›.
      ›

      interpret arrow_of_spans Maps.comp CMP r s
      proof
        show 1: "Maps.in_hom (Chn (CMP r s)) Dom.apex Cod.apex"
          using rs.iso_class_cmp_in_hom rs.composable CMP_def by simp
        show "Cod.leg0  Chn (CMP r s) = Dom.leg0"
        proof (intro Maps.arr_eqI)
          show 2: "Maps.seq Cod.leg0 (Chn (CMP r s))"
            using 1 Maps.dom_char Maps.cod_char by blast
          show 3: "Maps.arr Dom.leg0" by simp
          show "Maps.Dom (Cod.leg0  Chn (CMP r s)) = Maps.Dom Dom.leg0"
            using 1 2 Maps.dom_char Maps.cod_char Maps.comp_char
                  Dom.leg_in_hom Maps.in_hom_char Maps.seq_char
            by auto
          show "Maps.Cod (Cod.leg0  Chn (CMP r s)) = Maps.Cod Dom.leg0"
            using 2 3 Maps.comp_char [of Cod.leg0 "Chn (CMP r s)"]
                  Dom.leg_simps Dom.apex_def Maps.dom_char SPN_r_SPN_s_leg0_eq
                  Maps.comp_char [of rs.S0 "Maps.PRJ0 rs.R0 rs.S1"] CMP_def
            by simp
          show "Maps.Map (Cod.leg0  Chn (CMP r s)) = Maps.Map Dom.leg0"
          proof -
            have "Maps.Map (Cod.leg0  Chn (CMP r s)) = Maps.Comp tab0 (r  s) rs.cmp"
              using 1 2 Maps.dom_char Maps.cod_char
                    Maps.comp_char [of Cod.leg0 "Chn (CMP r s)"] CMP_def
              by simp
            also have "... = Maps.Comp tab0 s rs.ρσ.p0"
            proof -
              have "Maps.Comp tab0 (r  s) rs.cmp = tab0 (r  s)  rs.cmp"
                using Maps.Comp_eq_iso_class_memb Maps.CLS_hcomp cmp.is_map rs.cmp_def
                by auto
              also have "... = Maps.Comp tab0 s rs.ρσ.p0"
                using Maps.Comp_eq_iso_class_memb Maps.CLS_hcomp iso_class_eqI rs.cmp_props(4)
                by auto
              finally show ?thesis by simp
            qed
            also have "... = Maps.Map Dom.leg0"
            proof -
              have "Maps.seq rs.S0 (Maps.PRJ0 rs.R0 rs.S1)"
                by (intro Maps.seqI, simp_all add: rs.composable)
              moreover have "prj0 (Maps.REP rs.S1) (Maps.REP rs.R0) = rs.ρσ.p0"
                using "rs.prj_tab_agreement" iso_class_eqI by auto
              moreover have "Maps.Dom (Maps.PRJ0 rs.R0 rs.S1) = src rs.ρσ.p0"
                using rs.prj_tab_agreement Maps.PRJ0_def rs.composable
                      isomorphic_implies_hpar(3)
                by auto
              ultimately show ?thesis
                using SPN_r_SPN_s_leg0_eq Maps.comp_char [of rs.S0 "Maps.PRJ0 rs.R0 rs.S1"]
                      rs.S0_def Maps.PRJ0_def rs.composable CMP_def
                by simp
            qed
            finally show ?thesis by simp
          qed
        qed
        show "Cod.leg1  Chn (CMP r s) = Dom.leg1"
        proof (intro Maps.arr_eqI)
          show 2: "Maps.seq Cod.leg1 (Chn (CMP r s))"
            using 1 Maps.dom_char Maps.cod_char by blast
          show 3: "Maps.arr Dom.leg1" by simp
          show "Maps.Dom (Cod.leg1  Chn (CMP r s)) = Maps.Dom Dom.leg1"
            using 1 2 Maps.dom_char Maps.cod_char Maps.comp_char
                  Dom.leg_in_hom Maps.in_hom_char Maps.seq_char
            by auto
          show "Maps.Cod (Cod.leg1  Chn (CMP r s)) = Maps.Cod Dom.leg1"
            using 2 3 Maps.comp_char [of Cod.leg1 "Chn (CMP r s)"]
                  Dom.apex_def Maps.dom_char SPN_r_SPN_s_leg1_eq
                  Maps.comp_char [of rs.R1 "Maps.PRJ1 rs.R0 rs.S1"] CMP_def
            by simp
          show "Maps.Map (Cod.leg1  Chn (CMP r s)) = Maps.Map Dom.leg1"
          proof -
            have "Maps.Map (Cod.leg1  Chn (CMP r s)) = Maps.Comp tab1 (r  s) rs.cmp"
              using 1 2 Maps.dom_char Maps.cod_char
                    Maps.comp_char [of Cod.leg1 "Chn (CMP r s)"] CMP_def
              by simp
            also have "... = Maps.Comp tab1 r rs.ρσ.p1"
            proof -
              have "Maps.Comp tab1 (r  s) rs.cmp = tab1 (r  s)  rs.cmp"
                using Maps.Comp_eq_iso_class_memb Maps.CLS_hcomp cmp.is_map rs.cmp_def
                by auto
              also have "... = Maps.Comp tab1 r rs.ρσ.p1"
                using Maps.Comp_eq_iso_class_memb
                      Maps.CLS_hcomp [of "tab1 r" rs.ρσ.p1] iso_class_eqI rs.cmp_props(5)
                by auto
              finally show ?thesis by simp
            qed
            also have "... = Maps.Map Dom.leg1"
            proof -
              have "Maps.seq rs.R1 (Maps.PRJ1 rs.R0 rs.S1)"
                by (intro Maps.seqI, simp_all add: rs.composable)
              moreover have "prj1 (Maps.REP rs.S1) (Maps.REP rs.R0) = rs.ρσ.p1"
                using rs.prj_tab_agreement iso_class_eqI by auto
              moreover have "Maps.Dom (Maps.PRJ1 rs.R0 rs.S1) = src rs.ρσ.p1"
                using rs.prj_tab_agreement Maps.PRJ1_def rs.composable
                      isomorphic_implies_hpar(3)
                by auto
              ultimately show ?thesis
                using SPN_r_SPN_s_leg1_eq Maps.comp_char [of rs.R1 "Maps.PRJ1 rs.R0 rs.S1"]
                      rs.R1_def Maps.PRJ1_def rs.composable CMP_def
                by simp
            qed
            finally show ?thesis by simp
            (*
             * Very simple, right?  Yeah, once you sort through the notational morass and
             * figure out what equals what.
             *)
          qed
        qed
      qed
      show "Span.in_hom (CMP r s) (HoSPN_SPN.map (r, s)) (SPNoH.map (r, s))"
          using Span.arr_char arrow_of_spans_axioms Span.dom_char Span.cod_char
                CMP_def SPN.FF_def VV.arr_charSbC rs.composable
          by auto
      thus "Span.in_hhom (CMP r s) (SPN.map0 (src s)) (SPN.map0 (trg r))"
        using assms VV.ide_charSbC VV.arr_charSbC VV.in_hom_charSbC SPN.FF_def
        apply (intro Span.in_hhomI)
          apply auto
        using Span.src_dom [of "CMP r s"] Span.trg_dom [of "CMP r s"]
         apply (elim Span.in_homE)
         apply auto
        using Span.src_dom [of "CMP r s"] Span.trg_dom [of "CMP r s"]
        apply (elim Span.in_homE)
        by auto
    qed

    lemma compositor_simps [simp]:
    assumes "ide r" and "ide s" and "src r = trg s"
    shows "Span.arr (CMP r s)"
    and "Span.src (CMP r s) = SPN.map0 (src s)" and "Span.trg (CMP r s) = SPN.map0 (trg r)"
    and "Span.dom (CMP r s) = HoSPN_SPN.map (r, s)"
    and "Span.cod (CMP r s) = SPNoH.map (r, s)"
      using assms compositor_in_hom [of r s] by auto

    lemma compositor_is_iso:
    assumes "ide r" and "ide s" and "src r = trg s"
    shows "Span.iso (CMP r s)"
    proof -
      interpret rs: two_composable_identities_in_bicategory_of_spans V H 𝖺 𝗂 src trg r s
        using assms by unfold_locales auto
      have "Span.arr (CMP r s)"
        using assms compositor_in_hom by blast
      moreover have "Maps.iso ⟦⟦rs.cmp⟧⟧"
        using assms Maps.iso_char'
        by (metis (mono_tags, lifting) Maps.CLS_in_hom Maps.Map.simps(1) Maps.in_homE
            equivalence_is_left_adjoint ide_in_iso_class rs.cmp_props(3) rs.cmp_simps(2))
      ultimately show ?thesis
        unfolding CMP_def
        using assms Span.iso_char by simp
    qed

    interpretation Ξ: transformation_by_components VV.comp Span.vcomp
                        HoSPN_SPN.map SPNoH.map λrs. CMP (fst rs) (snd rs)
    proof
      fix rs
      assume rs: "VV.ide rs"
      let ?r = "fst rs"
      let ?s = "snd rs"
      show "Span.in_hom (CMP ?r ?s) (HoSPN_SPN.map rs) (SPNoH.map rs)"
        using rs compositor_in_hom [of ?r ?s] VV.ide_charSbC VV.arr_charSbC by simp
      next
      fix μν
      assume μν: "VV.arr μν"
      let  = "fst μν"
      let  = "snd μν"
      show "CMP (fst (VV.cod μν)) (snd (VV.cod μν))  HoSPN_SPN.map μν =
            SPNoH.map μν  CMP (fst (VV.dom μν)) (snd (VV.dom μν))"
      proof -
        let ?LHS = "CMP (fst (VV.cod μν)) (snd (VV.cod μν))  HoSPN_SPN.map μν"
        let ?RHS = "SPNoH.map μν  CMP (fst (VV.dom μν)) (snd (VV.dom μν))"
        have LHS:
          "Span.in_hom ?LHS (HoSPN_SPN.map (VV.dom μν)) (SPNoH.map (VV.cod μν))"
        proof
          show "Span.in_hom (HoSPN_SPN.map μν) (HoSPN_SPN.map (VV.dom μν))
                            (HoSPN_SPN.map (VV.cod μν))"
            using μν by blast
          show "Span.in_hom (CMP (fst (VV.cod μν)) (snd (VV.cod μν)))
                            (HoSPN_SPN.map (VV.cod μν)) (SPNoH.map (VV.cod μν))"
            using μν VV.cod_simp by (auto simp add: VV.arr_charSbC)
        qed
        have RHS:
          "Span.in_hom ?RHS (HoSPN_SPN.map (VV.dom μν)) (SPNoH.map (VV.cod μν))"
          using μν VV.dom_simp VV.cod_simp by (auto simp add: VV.arr_charSbC)
        show "?LHS = ?RHS"
        proof (intro Span.arr_eqI)
          show "Span.par ?LHS ?RHS"
            apply (intro conjI)
            using LHS RHS apply auto[2]
          proof -
            show "Span.dom ?LHS = Span.dom ?RHS"
            proof -
              have "Span.dom ?LHS = HoSPN_SPN.map (VV.dom μν)"
                using LHS by auto
              also have "... = Span.dom ?RHS"
                using RHS by auto
              finally show ?thesis by simp
            qed
            show "Span.cod ?LHS = Span.cod ?RHS"
            proof -
              have "Span.cod ?LHS = SPNoH.map (VV.cod μν)"
                using LHS by auto
              also have "... = Span.cod ?RHS"
                using RHS by auto
              finally show ?thesis by simp
            qed
          qed
          show "Chn ?LHS = Chn ?RHS"
          proof -
            interpret dom_μ_ν: two_composable_identities_in_bicategory_of_spans V H 𝖺 𝗂 src trg
                                 dom  dom 
              using μν VV.ide_charSbC VV.arr_charSbC by unfold_locales auto
            interpret cod_μ_ν: two_composable_identities_in_bicategory_of_spans V H 𝖺 𝗂 src trg
                                 cod  cod 
              using μν VV.ide_charSbC VV.arr_charSbC by unfold_locales auto
            interpret μ_ν: horizontal_composite_of_arrows_of_tabulations_in_maps
                             V H 𝖺 𝗂 src trg
                             dom  tab_of_ide (dom ) tab0 (dom ) tab1 (dom )
                             dom  tab_of_ide (dom ) tab0 (dom ) tab1 (dom )
                             cod  tab_of_ide (cod ) tab0 (cod ) tab1 (cod )
                             cod  tab_of_ide (cod ) tab0 (cod ) tab1 (cod )
                              
              using μν VV.arr_charSbC by unfold_locales auto

            let ?μν = "  "
            interpret dom_μν: identity_in_bicategory_of_spans V H 𝖺 𝗂 src trg dom ?μν
              using μν by (unfold_locales, simp)
            interpret cod_μν: identity_in_bicategory_of_spans V H 𝖺 𝗂 src trg cod ?μν
              using μν by (unfold_locales, simp)
            interpret μν: arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg
                            dom ?μν tab_of_ide (dom ?μν) tab0 (dom ?μν) tab1 (dom ?μν)
                            cod ?μν tab_of_ide (cod ?μν) tab0 (cod ?μν) tab1 (cod ?μν)
                            ?μν
              using μν by unfold_locales auto

            have Chn_LHS_eq:
              "Chn ?LHS = ⟦⟦cod_μ_ν.cmp⟧⟧  Span.chine_hcomp (SPN (fst μν)) (SPN (snd μν))"
            proof -
              have "Chn ?LHS = Chn (CMP (fst (VV.cod μν)) (snd (VV.cod μν))) 
                                 Chn (HoSPN_SPN.map μν)"
                using μν LHS Span.Chn_vcomp by blast
              also have "... = ⟦⟦cod_μ_ν.cmp⟧⟧  Chn (HoSPN_SPN.map μν)"
                using μν VV.arr_charSbC VV.cod_charSbC CMP_def by simp
              also have "... = ⟦⟦cod_μ_ν.cmp⟧⟧ 
                                 Span.chine_hcomp (SPN (fst μν)) (SPN (snd μν))"
                using μν VV.arr_charSbC SPN.FF_def Span.hcomp_def by simp
              finally show ?thesis by blast
            qed
            have Chn_RHS_eq:
               "Chn ?RHS = Maps.MkArr (src (tab0 (dom   dom )))
                                      (src (tab0 (cod   cod )))
                                      μν.chine 
                           Maps.MkArr (src dom_μ_ν.ρσ.p0) (src (tab_of_ide (dom   dom )))
                                      dom_μ_ν.cmp"
            proof -
              have "Chn ?RHS = Chn (SPN (  )) 
                               Maps.MkArr (src dom_μ_ν.ρσ.p0) (src (tab_of_ide (dom   dom )))
                                          dom_μ_ν.cmp"
                using μν RHS Span.vcomp_def VV.arr_charSbC CMP_def Span.arr_char Span.not_arr_Null
                      VV.dom_simp
                by auto
              moreover have "Chn (SPN (  )) =
                             Maps.MkArr (src (tab0 (dom   dom )))
                                        (src (tab0 (cod   cod )))
                                        μν.chine"
              proof -
                have "Chn (SPN (  )) =
                      Maps.MkArr (src (tab0 (dom   dom )))
                                 (src (tab0 (cod   cod )))
                                 spn ?μν"
                  using μν SPN_def by simp
                also have "... = Maps.MkArr (src (tab0 (dom   dom )))
                                            (src (tab0 (cod   cod )))
                                            μν.chine"
                  using spn_def by simp
                finally show ?thesis by simp
              qed
              ultimately show ?thesis by simp
            qed

            let ?Chn_LHS =
                "Maps.MkArr (src cod_μ_ν.ρσ.p0) (src (tab_of_ide (cod   cod )))
                   cod_μ_ν.cmp 
                 Span.chine_hcomp (SPN ) (SPN )"
            let ?Chn_RHS =
                "Maps.MkArr (src (tab0 (dom   dom ))) (src (tab0 (cod   cod )))
                   μν.chine 
                 Maps.MkArr (src dom_μ_ν.ρσ.p0) (src (tab_of_ide (dom   dom )))
                   dom_μ_ν.cmp"

            have "?Chn_LHS = ?Chn_RHS"
            proof (intro Maps.arr_eqI)
              interpret LHS: arrow_of_spans Maps.comp ?LHS
                using LHS Span.arr_char by auto
              interpret RHS: arrow_of_spans Maps.comp ?RHS
                using RHS Span.arr_char by auto
              show 1: "Maps.arr ?Chn_LHS"
                using LHS.chine_in_hom Chn_LHS_eq by auto
              show 2: "Maps.arr ?Chn_RHS"
                using RHS.chine_in_hom Chn_RHS_eq by auto
              text ‹
                Here is where we use dom_μ_ν.chine_hcomp_SPN_SPN›,
                which depends on our having chosen the ``right'' pullbacks for Maps(B)›.
                The map Chn_LHS› has as its domain the apex of the
                horizontal composite of the components of @{term "VV.dom μν"},
                whereas Chn_RHS› has as its
                domain the apex of the chosen tabulation of r0* ⋆ s1.
                We need these to be equal in order for Chn_LHS› and Chn_RHS› to be equal.
              ›
              show "Maps.Dom ?Chn_LHS = Maps.Dom ?Chn_RHS"
              proof -
                have "Maps.Dom ?Chn_LHS = Maps.Dom (Maps.dom ?Chn_LHS)"
                   using μν 1 Maps.Dom_dom by presburger
                also have
                  "... = Maps.Dom (Span.chine_hcomp (SPN (dom )) (SPN (dom )))"
                proof -
                  have "... = Maps.Dom (Maps.dom (Span.chine_hcomp (SPN ) (SPN )))"
                    using 1 Maps.seq_char Maps.Dom_comp by auto
                  also have "... = Maps.Dom (Maps.pbdom (Leg0 (Dom (SPN )))
                                                        (Leg1 (Dom (SPN ))))"
                    using μν VV.arr_charSbC Span.chine_hcomp_in_hom [of "SPN " "SPN "]
                    by auto
                  also have "... = Maps.Dom (Maps.dom (Maps.pbdom (Leg0 (Dom (SPN )))
                                                                  (Leg1 (Dom (SPN )))))"
                  proof -
                    have "Maps.cospan (Leg0 (Dom (SPN (fst μν)))) (Leg1 (Dom (SPN (snd μν))))"
                      using μν VV.arr_charSbC SPN_in_hom Span.arr_char Span.dom_char SPN_def
                            Maps.CLS_in_hom Maps.arr_char Maps.cod_char dom_μ_ν.composable
                            dom_μ_ν.RS_simps(16) dom_μ_ν.S1_def dom_μ_ν.RS_simps(1)
                            dom_μ_ν.R0_def Maps.pbdom_in_hom
                      by simp
                    thus ?thesis
                      using μν VV.arr_charSbC Maps.pbdom_in_hom by simp
                  qed
                  also have "... = Maps.Dom
                                     (Maps.dom (Maps.pbdom (Leg0 (Dom (SPN (dom ))))
                                                           (Leg1 (Dom (SPN (dom ))))))"
                    using μν SPN_def VV.arr_charSbC by simp
                  also have "... = Maps.Dom
                                     (Maps.dom (Span.chine_hcomp (SPN (dom )) (SPN (dom ))))"
                    using μν VV.arr_charSbC ide_dom
                    by (simp add: Span.chine_hcomp_ide_ide)
                  also have "... = Maps.Dom (Span.chine_hcomp (SPN (dom )) (SPN (dom )))"
                    using Maps.Dom_dom Maps.in_homE SPN.preserves_reflects_arr SPN.preserves_src
                          SPN.preserves_trg Span.chine_hcomp_in_hom dom_μ_ν.composable
                          dom_μ_ν.r.base_simps(2) dom_μ_ν.s.base_simps(2)
                    by (metis (no_types, lifting))
                  finally show ?thesis by simp
                qed
                also have "... = src dom_μ_ν.ρσ.p0"
                  using "dom_μ_ν.chine_hcomp_SPN_SPN" by simp
                also have "... = Maps.Dom ?Chn_RHS"
                  using 2 Maps.seq_char Maps.Dom_comp by auto
                finally show ?thesis by simp
              qed
              show "Maps.Cod ?Chn_LHS = Maps.Cod ?Chn_RHS"
              proof -
                have "Maps.Cod ?Chn_LHS = src (tab_of_ide (cod   cod ))"
                  using μν 1 VV.arr_charSbC Maps.seq_char by auto
                also have "... = src (tab0 (cod   cod ))"
                  using μν VV.arr_charSbC cod_μν.tab_simps(2) by auto
                also have "... = Maps.Cod ?Chn_RHS"
                  by (metis (no_types, lifting) "2" Maps.Cod.simps(1) Maps.Cod_comp Maps.seq_char)
                finally show ?thesis by simp
              qed
              show "Maps.Map ?Chn_LHS = Maps.Map ?Chn_RHS"
              proof -
                have RHS: "Maps.Map ?Chn_RHS = iso_class (μν.chine  dom_μ_ν.cmp)"
                proof -
                  have "Maps.Map ?Chn_RHS = Maps.Comp μν.chine dom_μ_ν.cmp"
                    using μν 2 VV.arr_charSbC Maps.Map_comp
                          Maps.comp_char
                            [of "Maps.MkArr (src (tab0 (dom   dom )))
                                            (src (tab0 (cod   cod )))
                                            μν.chine"
                                "Maps.MkArr (src dom_μ_ν.ρσ.p0)
                                            (src (tab_of_ide (dom   dom )))
                                            dom_μ_ν.cmp"]
                    by simp
                  also have "... = μν.chine  dom_μ_ν.cmp"
                  proof -
                    have "dom_μ_ν.cmp 
                          Maps.Hom (src dom_μ_ν.ρσ.p0) (src (tab0 (dom   dom )))"
                    proof -
                      have "dom_μ_ν.cmp 
                            Maps.Hom (src dom_μ_ν.ρσ.tab) (src (tab_of_ide (dom   dom )))"
                        using μν VV.arr_charSbC dom_μ_ν.cmp_props(1-3)
                        by (metis (mono_tags, lifting) equivalence_is_left_adjoint mem_Collect_eq)
                      thus ?thesis
                        using μν VV.arr_charSbC dom_μν.tab_simps(2) by simp
                    qed
                    moreover have "μν.chine 
                                   Maps.Hom (src (tab0 (dom   dom )))
                                            (src (tab0 (cod   cod )))"
                      using μν VV.arr_charSbC μν.chine_in_hom μν.is_map by auto
                    moreover have
                      "μν.chine  dom_μ_ν.cmp  Maps.Comp μν.chine dom_μ_ν.cmp"
                    proof
                      show "is_iso_class dom_μ_ν.cmp"
                        using is_iso_classI by simp
                      show "is_iso_class μν.chine"
                        using is_iso_classI by simp
                      show "dom_μ_ν.cmp  dom_μ_ν.cmp"
                        using ide_in_iso_class [of dom_μ_ν.cmp] by simp
                      show "μν.chine  μν.chine" 
                        using ide_in_iso_class by simp
                      show "μν.chine  dom_μ_ν.cmp  μν.chine  dom_μ_ν.cmp"
                        using μν VV.arr_charSbC μν.chine_simps dom_μ_ν.cmp_simps dom_μν.tab_simps(2)
                              isomorphic_reflexive
                        by auto
                    qed
                    ultimately show ?thesis
                      using μν dom_μ_ν.cmp_props μν.chine_in_hom μν.chine_is_induced_map
                            Maps.Comp_eq_iso_class_memb
                      by blast
                  qed
                  finally show ?thesis by simp
                qed

                have LHS: "Maps.Map ?Chn_LHS = cod_μ_ν.cmp  μ_ν.chine"
                proof -
                  have "Maps.Map ?Chn_LHS =
                           Maps.Comp cod_μ_ν.cmp
                                     (Maps.Map
                                       (Maps.tuple (Maps.CLS (spn   dom_μ_ν.ρσ.p1))
                                                   (Maps.CLS (tab0 (cod )))
                                                   (Maps.CLS (tab1 (cod )))
                                                   (Maps.CLS (spn   dom_μ_ν.ρσ.p0))))"
                  proof -
                    have "Maps.Map ?Chn_LHS =
                          Maps.Comp cod_μ_ν.cmp
                                    (Maps.Map (Span.chine_hcomp (SPN ) (SPN )))"
                      using μν 1 VV.arr_charSbC Maps.Map_comp cod_μν.tab_simps(2)
                            Maps.comp_char
                              [of "Maps.MkArr (src cod_μ_ν.ρσ.p0)
                                              (src (tab_of_ide (cod   cod )))
                                              cod_μ_ν.cmp"
                                  "Span.chine_hcomp (SPN ) (SPN )"]
                      by simp
                    moreover have "Span.chine_hcomp (SPN ) (SPN ) =
                                   Maps.tuple
                                     (Maps.CLS (spn   dom_μ_ν.ρσ.p1))
                                      (Maps.CLS (tab0 (cod )))
                                      (Maps.CLS (tab1 (cod )))
                                     (Maps.CLS (spn   dom_μ_ν.ρσ.p0))"
                    proof -
                      have "Maps.PRJ0
                              (Maps.MkArr (src (tab0 (dom ))) (trg ) tab0 (dom ))
                              (Maps.MkArr (src (tab0 (dom ))) (trg ) tab1 (dom )) =
                            ⟦⟦dom_μ_ν.ρσ.p0⟧⟧ 
                            Maps.PRJ1
                              (Maps.MkArr (src (tab0 (dom ))) (trg ) tab0 (dom ))
                              (Maps.MkArr (src (tab0 (dom ))) (trg ) tab1 (dom )) =
                            ⟦⟦dom_μ_ν.ρσ.p1⟧⟧"
                      proof -
                        interpret X: identity_in_bicategory_of_spans V H 𝖺 𝗂 src trg
                                       (tab0 (dom ))*  tab1 (dom )
                          using μν VV.arr_charSbC
                          by (unfold_locales, simp)
                        have "Maps.PRJ0
                                (Maps.MkArr (src (tab0 (dom ))) (trg ) tab0 (dom ))
                                (Maps.MkArr (src (tab0 (dom ))) (trg ) tab1 (dom )) =
                              ⟦⟦tab0 ((Maps.REP (Maps.MkArr (src (tab0 (dom ))) (trg (snd μν))
                                                           tab0 (dom )))* 
                                      Maps.REP (Maps.MkArr (src (tab0 (dom ))) (trg )
                                                           tab1 (dom )))⟧⟧"
                          unfolding Maps.PRJ0_def
                          using μν VV.arr_charSbC dom_μ_ν.RS_simps(1) dom_μ_ν.RS_simps(16)
                                dom_μ_ν.RS_simps(18) dom_μ_ν.RS_simps(3) dom_μ_ν.R0_def
                                dom_μ_ν.S1_def
                          by auto
                        moreover
                        have "Maps.PRJ1
                                (Maps.MkArr (src (tab0 (dom ))) (trg ) tab0 (dom ))
                                (Maps.MkArr (src (tab0 (dom ))) (trg ) tab1 (dom )) =
                              ⟦⟦tab1 ((Maps.REP (Maps.MkArr (src (tab0 (dom ))) (trg (snd μν))
                                                           tab0 (dom )))* 
                                     Maps.REP (Maps.MkArr (src (tab0 (dom ))) (trg )
                                                          tab1 (dom )))⟧⟧"
                          unfolding Maps.PRJ1_def
                          using μν VV.arr_charSbC dom_μ_ν.RS_simps(1) dom_μ_ν.RS_simps(16)
                                dom_μ_ν.RS_simps(18) dom_μ_ν.RS_simps(3) dom_μ_ν.R0_def
                                dom_μ_ν.S1_def
                          by auto
                        moreover
                        have "(Maps.REP (Maps.MkArr (src (tab0 (dom ))) (trg (snd μν))
                                                    tab0 (dom )))* 
                              Maps.REP (Maps.MkArr (src (tab0 (dom ))) (trg )
                                                   tab1 (dom )) 
                              (tab0 (dom ))*  tab1 (dom )"
                          using VV.arr_charSbC μν dom_μ_ν.S1_def dom_μ_ν.s.leg1_simps(3)
                                dom_μ_ν.s.leg1_simps(4) trg_dom dom_μ_ν.R0_def
                                dom_μ_ν.prj_tab_agreement(1) isomorphic_symmetric
                          by simp
                        ultimately show ?thesis
                           using X.isomorphic_implies_same_tab isomorphic_symmetric by metis
                      qed
                      thus ?thesis
                        unfolding Span.chine_hcomp_def
                        using μν VV.arr_charSbC SPN_def isomorphic_reflexive
                              Maps.comp_CLS [of "spn " dom_μ_ν.ρσ.p1 "spn   dom_μ_ν.ρσ.p1"]
                              Maps.comp_CLS [of "spn " dom_μ_ν.ρσ.p0 "spn   dom_μ_ν.ρσ.p0"]
                        by simp
                    qed
                    moreover have "Maps.Dom (Span.chine_hcomp (SPN ) (SPN )) =
                                   src dom_μ_ν.ρσ.p0"
                      by (metis (no_types, lifting) "1" "2" Maps.Dom.simps(1) Maps.comp_char
                          Maps.Dom ?Chn_LHS = Maps.Dom ?Chn_RHS)
                    ultimately show ?thesis by simp
                  qed
                  also have "... = Maps.Comp cod_μ_ν.cmp μ_ν.chine"
                  proof -
                    let ?tuple = "Maps.tuple ⟦⟦spn (fst μν)  dom_μ_ν.ρσ.p1⟧⟧
                                               ⟦⟦tab0 (cod )⟧⟧ ⟦⟦tab1 (cod )⟧⟧
                                             ⟦⟦spn (snd μν)  dom_μ_ν.ρσ.p0⟧⟧"
                    have "iso_class μ_ν.chine = Maps.Map ?tuple"
                      using μ_ν.CLS_chine spn_def Maps.Map.simps(1)
                      by (metis (no_types, lifting))
                    thus ?thesis by simp
                  qed
                  also have "... = cod_μ_ν.cmp  μ_ν.chine"
                  proof -
                    have "μ_ν.chine  Maps.Hom (src dom_μ_ν.ρσ.p0) (src cod_μ_ν.ρσ.p0)"
                    proof -
                      have "«μ_ν.chine : src dom_μ_ν.ρσ.p0  src cod_μ_ν.ρσ.p0»"
                        using μν VV.arr_charSbC by simp
                      thus ?thesis
                        using μ_ν.is_map ide_in_iso_class left_adjoint_is_ide by blast
                    qed
                    moreover have "cod_μ_ν.cmp 
                                   Maps.Hom (src cod_μ_ν.ρσ.p0) (src (tab0 (cod   cod )))"
                    proof -
                      have "«cod_μ_ν.cmp : src cod_μ_ν.ρσ.p0  src (tab0 (cod   cod ))»"
                        using μν VV.arr_charSbC cod_μ_ν.cmp_in_hom cod_μν.tab_simps(2)
                        by simp
                      thus ?thesis
                        using cod_μ_ν.cmp_props equivalence_is_left_adjoint left_adjoint_is_ide
                              ide_in_iso_class [of cod_μ_ν.cmp]
                        by (metis (mono_tags, lifting) mem_Collect_eq)
                    qed
                    moreover have
                      "cod_μ_ν.cmp  μ_ν.chine  Maps.Comp cod_μ_ν.cmp μ_ν.chine"
                    proof
                      show "is_iso_class μ_ν.chine"
                        using μ_ν.w_simps(1) is_iso_classI by blast
                      show "is_iso_class cod_μ_ν.cmp"
                        using cod_μ_ν.cmp_simps(2) is_iso_classI by blast
                      show "μ_ν.chine  μ_ν.chine"
                        using ide_in_iso_class by simp
                      show "cod_μ_ν.cmp  cod_μ_ν.cmp"
                        using ide_in_iso_class by simp
                      show "cod_μ_ν.cmp  μ_ν.chine  cod_μ_ν.cmp  μ_ν.chine"
                        by (simp add: isomorphic_reflexive)
                    qed
                    ultimately show ?thesis
                      using μν cod_μ_ν.cmp_props μ_ν.chine_in_hom μ_ν.chine_is_induced_map
                            Maps.Comp_eq_iso_class_memb
                      by simp
                  qed
                  finally show ?thesis by simp
                qed

                have EQ: "μν.chine  dom_μ_ν.cmp = cod_μ_ν.cmp  μ_ν.chine"
                proof (intro iso_class_eqI)
                  show "μν.chine  dom_μ_ν.cmp  cod_μ_ν.cmp  μ_ν.chine"
                  proof -
                    interpret dom_cmp: identity_arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg
                                         dom ?μν
                                         dom_μ_ν.ρσ.tab
                                         tab0 (dom )  dom_μ_ν.ρσ.p0
                                         tab1 (dom )  dom_μ_ν.ρσ.p1
                                         dom ?μν
                                         tab_of_ide (dom   dom )
                                         tab0 (dom   dom )
                                         tab1 (dom   dom )
                                         dom ?μν
                      using μν VV.arr_charSbC dom_μ_ν.cmp_interpretation by simp
                    interpret cod_cmp: identity_arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg
                                        cod ?μν
                                        cod_μ_ν.ρσ.tab
                                        tab0 (cod )  cod_μ_ν.ρσ.p0
                                        tab1 (cod )  cod_μ_ν.ρσ.p1
                                        cod ?μν
                                        tab_of_ide (cod   cod )
                                        tab0 (cod   cod )
                                        tab1 (cod   cod )
                                        cod ?μν
                      using μν VV.arr_charSbC cod_μ_ν.cmp_interpretation by simp
                    interpret L: vertical_composite_of_arrows_of_tabulations_in_maps
                                   V H 𝖺 𝗂 src trg
                                   dom ?μν
                                   dom_μ_ν.ρσ.tab
                                   tab0 (dom )  dom_μ_ν.ρσ.p0
                                   tab1 (dom )  dom_μ_ν.ρσ.p1
                                   dom ?μν
                                   tab_of_ide (dom ?μν)
                                   tab0 (dom ?μν)
                                   tab1 (dom ?μν)
                                   cod ?μν
                                   cod_μν.tab
                                   tab0 (cod ?μν)
                                   tab1 (cod ?μν)
                                   dom ?μν
                                     
                      using μν VV.arr_charSbC dom_μ_ν.cmp_in_hom
                      by unfold_locales auto
                    interpret R: vertical_composite_of_arrows_of_tabulations_in_maps
                                   V H 𝖺 𝗂 src trg
                                   dom ?μν
                                   dom_μ_ν.ρσ.tab
                                   tab0 (dom )  dom_μ_ν.ρσ.p0
                                   tab1 (dom )  dom_μ_ν.ρσ.p1
                                   cod ?μν
                                   cod_μ_ν.ρσ.tab
                                   tab0 (cod )  cod_μ_ν.ρσ.p0
                                   tab1 (cod )  cod_μ_ν.ρσ.p1
                                   cod ?μν
                                   cod_μν.tab
                                   tab0 (cod ?μν)
                                   tab1 (cod ?μν)
                                     
                                   cod ?μν
                       using μν VV.arr_charSbC cod_μ_ν.cmp_in_hom
                       by unfold_locales auto
                    have "μν.chine  dom_μ_ν.cmp  L.chine"
                      using μν VV.arr_charSbC L.chine_char dom_μ_ν.cmp_def isomorphic_symmetric
                      by simp
                    also have "... = R.chine"
                      using L.is_ide μν comp_arr_dom comp_cod_arr isomorphic_reflexive by force
                    also have "...  cod_μ_ν.cmp  μ_ν.chine"
                      using μν VV.arr_charSbC R.chine_char cod_μ_ν.cmp_def by simp
                    finally show ?thesis by simp
                  qed
                qed
                show ?thesis
                  using LHS RHS EQ by simp
              qed
            qed
            thus ?thesis
              using Chn_LHS_eq Chn_RHS_eq by simp
          qed
        qed
      qed
    qed

    interpretation Ξ: natural_isomorphism VV.comp Span.vcomp
                        HoSPN_SPN.map SPNoH.map Ξ.map
      using VV.ide_charSbC VV.arr_charSbC Ξ.map_simp_ide compositor_is_iso
      by (unfold_locales, simp)

    lemma compositor_is_natural_transformation:
    shows "transformation_by_components VV.comp Span.vcomp HoSPN_SPN.map SPNoH.map
             (λrs. CMP (fst rs) (snd rs))"
      ..

    lemma compositor_is_natural_isomorphism:
    shows "natural_isomorphism VV.comp Span.vcomp HoSPN_SPN.map SPNoH.map Ξ.map"
      ..

  end
  
  subsubsection "Associativity Coherence"

  locale three_composable_identities_in_bicategory_of_spans =
    bicategory_of_spans V H 𝖺 𝗂 src trg +
    f: identity_in_bicategory_of_spans V H 𝖺 𝗂 src trg f +
    g: identity_in_bicategory_of_spans V H 𝖺 𝗂 src trg g +
    h: identity_in_bicategory_of_spans V H 𝖺 𝗂 src trg h
  for V :: "'a comp"                 (infixr "" 55)
  and H :: "'a  'a  'a"          (infixr "" 53)
  and 𝖺 :: "'a  'a  'a  'a"     ("𝖺[_, _, _]")
  and 𝗂 :: "'a  'a"                 ("𝗂[_]")
  and src :: "'a  'a"
  and trg :: "'a  'a"
  and f :: 'a
  and g :: 'a
  and h :: 'a +
  assumes fg: "src f = trg g"
  and gh: "src g = trg h"
  begin

    interpretation f: arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg
                        f f.tab tab0 f tab1 f f f.tab tab0 f tab1 f f
      using f.is_arrow_of_tabulations_in_maps by simp
    interpretation h: arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg
                        h h.tab tab0 h tab1 h h h.tab tab0 h tab1 h h
      using h.is_arrow_of_tabulations_in_maps by simp

    interpretation E: self_evaluation_map V H 𝖺 𝗂 src trg ..
    notation E.eval ("_")

    interpretation Maps: maps_category V H 𝖺 𝗂 src trg ..
    interpretation Span: span_bicategory Maps.comp Maps.PRJ0 Maps.PRJ1 ..

    no_notation Fun.comp (infixl "" 55)
    notation Span.vcomp (infixr "" 55)
    notation Span.hcomp (infixr "" 53)
    notation Maps.comp (infixr "" 55)
    notation isomorphic (infix "" 50)

    interpretation SPN: "functor" V Span.vcomp SPN
      using SPN_is_functor by simp
    interpretation SPN: weak_arrow_of_homs V src trg Span.vcomp Span.src Span.trg SPN
      using SPN_is_weak_arrow_of_homs by simp
    interpretation SPN_SPN: "functor" VV.comp Span.VV.comp SPN.FF
      using SPN.functor_FF by auto
    interpretation HoSPN_SPN: composite_functor VV.comp Span.VV.comp Span.vcomp
                                SPN.FF λμν. fst μν  snd μν
      ..
    interpretation SPNoH: composite_functor VV.comp V Span.vcomp λμν. fst μν  snd μν SPN
      ..

    text ‹
      Here come a lot of interpretations for ``composite things''.
      We need these in order to have relatively short, systematic names for entities that will
      appear in the lemmas to follow.
      The names of the interpretations use a prefix notation, where H› refers to horizontal
      composition of 1-cells and T› refers to composite of tabulations.
      So, for example, THfgh› refers to the composite of the tabulation associated with the
      horizontal composition f ⋆ g› with the tabulation associated with h›.
    ›
    interpretation HHfgh: identity_in_bicategory_of_spans V H 𝖺 𝗂 src trg (f  g)  h
      using fg gh by unfold_locales auto
    interpretation HfHgh: identity_in_bicategory_of_spans V H 𝖺 𝗂 src trg f  g  h
      using fg gh by unfold_locales auto
    interpretation Tfg: two_composable_identities_in_bicategory_of_spans V H 𝖺 𝗂 src trg f g
      using fg gh by unfold_locales auto
    interpretation Tgh: two_composable_identities_in_bicategory_of_spans V H 𝖺 𝗂 src trg g h
      using fg gh by unfold_locales auto
    interpretation THfgh: two_composable_identities_in_bicategory_of_spans V H 𝖺 𝗂 src trg
                            f  g h
      using fg gh by unfold_locales auto
    interpretation THfgh: tabulation V H 𝖺 𝗂 src trg (f  g)  h THfgh.ρσ.tab
                            tab0 h  THfgh.ρσ.p0 tab1 (f  g)  THfgh.ρσ.p1
      using THfgh.ρσ.composite_is_tabulation by simp
    interpretation TfHgh: two_composable_identities_in_bicategory_of_spans V H 𝖺 𝗂 src trg
                            f g  h
      using fg gh by unfold_locales auto
    interpretation TfHgh: tabulation V H 𝖺 𝗂 src trg f  g  h TfHgh.ρσ.tab
                            tab0 (g  h)  TfHgh.ρσ.p0 tab1 f  TfHgh.ρσ.p1
      using TfHgh.ρσ.composite_is_tabulation by simp

    interpretation Tfg_Hfg: arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg
                              f  g Tfg.ρσ.tab tab0 g  Tfg.ρσ.p0 tab1 f  Tfg.ρσ.p1
                              f  g tab_of_ide (f  g) tab0 (f  g) tab1 (f  g)
                              f  g
      by unfold_locales auto
    interpretation Tgh_Hgh: arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg
                              g  h Tgh.ρσ.tab tab0 h  Tgh.ρσ.p0 tab1 g  Tgh.ρσ.p1
                              g  h tab_of_ide (g  h) tab0 (g  h) tab1 (g  h)
                              g  h
      by unfold_locales auto
    interpretation THfgh_HHfgh:
        arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg
          (f  g)  h THfgh.ρσ.tab tab0 h  THfgh.ρσ.p0 tab1 (f  g)  THfgh.ρσ.p1
          (f  g)  h tab_of_ide ((f  g)  h) tab0 ((f  g)  h) tab1 ((f  g)  h)
          (f  g)  h
      using fg gh by unfold_locales auto
    interpretation TfHgh_HfHgh:
        arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg
          f  g  h TfHgh.ρσ.tab tab0 (g  h)  TfHgh.ρσ.p0 tab1 f  TfHgh.ρσ.p1
          f  g  h tab_of_ide (f  g  h) tab0 (f  g  h) tab1 (f  g  h)
          f  g  h
      using fg gh by unfold_locales auto
    interpretation TTfgh: composite_tabulation_in_maps V H 𝖺 𝗂 src trg
                            f  g Tfg.ρσ.tab tab0 g  Tfg.ρσ.p0 tab1 f  Tfg.ρσ.p1
                            h tab_of_ide h tab0 h tab1 h
      using fg gh by unfold_locales auto
    interpretation TTfgh_THfgh:
        horizontal_composite_of_arrows_of_tabulations_in_maps V H 𝖺 𝗂 src trg
          f  g Tfg.ρσ.tab tab0 g  Tfg.ρσ.p0 tab1 f  Tfg.ρσ.p1
          h tab_of_ide h tab0 h tab1 h
          f  g tab_of_ide (f  g) tab0 (f  g) tab1 (f  g)
          h tab_of_ide h tab0 h tab1 h
          f  g h
      ..
    interpretation TfTgh: composite_tabulation_in_maps V H 𝖺 𝗂 src trg
                            f tab_of_ide f tab0 f tab1 f
                            g  h Tgh.ρσ.tab tab0 h  Tgh.ρσ.p0 tab1 g  Tgh.ρσ.p1
      using fg gh by unfold_locales auto
    interpretation TfTgh_TfHgh:
        horizontal_composite_of_arrows_of_tabulations_in_maps V H 𝖺 𝗂 src trg
          f tab_of_ide f tab0 f tab1 f
          g  h Tgh.ρσ.tab tab0 h  Tgh.ρσ.p0 tab1 g  Tgh.ρσ.p1
          f tab_of_ide f tab0 f tab1 f
          g  h tab_of_ide (g  h) tab0 (g  h) tab1 (g  h)
          f g  h
      ..
    interpretation TfTgh_TfTgh:
        horizontal_composite_of_arrows_of_tabulations_in_maps V H 𝖺 𝗂 src trg
          f tab_of_ide f tab0 f tab1 f
          g  h Tgh.ρσ.tab tab0 h  Tgh.ρσ.p0 tab1 g  Tgh.ρσ.p1
          f tab_of_ide f tab0 f tab1 f
          g  h Tgh.ρσ.tab tab0 h  Tgh.ρσ.p0 tab1 g  Tgh.ρσ.p1
          f g  h
      ..
    text ‹
      The following interpretation defines the associativity between the peaks
      of the two composite tabulations TTfgh› (associated to the left) and TfTgh›
      (associated to the right).
    ›
    (* TODO: Try to get rid of the .ρσ in, e.g., Tfg.ρσ.p1. *)
    interpretation TTfgh_TfTgh:
        arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg
          (f  g)  h TTfgh.tab tab0 h  TTfgh.p0 (tab1 f  Tfg.ρσ.p1)  TTfgh.p1
          f  g  h TfTgh.tab (tab0 h  Tgh.ρσ.p0)  TfTgh.p0 tab1 f  TfTgh.p1
          𝖺[f, g, h]
      using fg gh by unfold_locales auto

    text ‹
      This interpretation defines the map, from the apex of the tabulation associated
      with the horizontal composite (f ⋆ g) ⋆ h› to the apex of the tabulation associated
      with the horizontal composite f ⋆ g ⋆ h›, induced by the associativity isomorphism
      𝖺[f, g, h]› from (f ⋆ g) ⋆ h› to f ⋆ g ⋆ h›.
    ›

    interpretation HHfgh_HfHgh: arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg
                                  dom (α (f, g, h)) tab_of_ide (dom (α (f, g, h)))
                                  tab0 (dom (α (f, g, h))) tab1 (dom (α (f, g, h)))
                                  cod (α (f, g, h)) tab_of_ide (cod (α (f, g, h)))
                                  tab0 (cod (α (f, g, h))) tab1 (cod (α (f, g, h)))
                                  α (f, g, h)
    proof -
      have "arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg
              ((f  g)  h) (tab_of_ide ((f  g)  h)) (tab0 ((f  g)  h)) (tab1 ((f  g)  h))
              (f  g  h) (tab_of_ide (f  g  h)) (tab0 (f  g  h)) (tab1 (f  g  h))
              𝖺[f, g, h]"
        using fg gh by unfold_locales auto
      thus "arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg
              (dom (α (f, g, h))) (tab_of_ide (dom (α (f, g, h))))
              (tab0 (dom (α (f, g, h)))) (tab1 (dom (α (f, g, h))))
              (cod (α (f, g, h))) (tab_of_ide (cod (α (f, g, h))))
              (tab0 (cod (α (f, g, h)))) (tab1 (cod (α (f, g, h))))
              (α (f, g, h))"
        using fg gh α_def by auto
    qed

    interpretation SPN_f: arrow_of_spans Maps.comp SPN f
      using SPN_in_hom Span.arr_char [of "SPN f"] by simp
    interpretation SPN_g: arrow_of_spans Maps.comp SPN g
      using SPN_in_hom Span.arr_char [of "SPN g"] by simp
    interpretation SPN_h: arrow_of_spans Maps.comp SPN h
      using SPN_in_hom Span.arr_char [of "SPN h"] by simp
    interpretation SPN_fgh: three_composable_identity_arrows_of_spans Maps.comp
                              Maps.PRJ0 Maps.PRJ1 SPN f SPN g SPN h
      using fg gh Span.arr_char SPN_in_hom SPN.preserves_ide Span.ide_char
      apply unfold_locales by auto

    text ‹
      The following relates the projections associated with the composite span SPN_fgh›
      with tabulations in the underlying bicategory.
    ›

    lemma prj_char:
    shows "SPN_fgh.Prj11 = ⟦⟦Tfg.ρσ.p1  TTfgh.p1⟧⟧"
    and "SPN_fgh.Prj01 = ⟦⟦Tfg.ρσ.p0  TTfgh.p1⟧⟧"
    and "SPN_fgh.Prj0 = ⟦⟦TTfgh.p0⟧⟧"
    and "SPN_fgh.Prj1 = ⟦⟦TfTgh.p1⟧⟧"
    and "SPN_fgh.Prj10 = ⟦⟦Tgh.ρσ.p1  TfTgh.p0⟧⟧"
    and "SPN_fgh.Prj00 = ⟦⟦Tgh.ρσ.p0  TfTgh.p0⟧⟧"
    proof -
      show "SPN_fgh.Prj11 = ⟦⟦Tfg.ρσ.p1  TTfgh.p1⟧⟧"
      proof -
        have "ide (Tfg.ρσ.p1  TTfgh.p1)"
          by (metis TTfgh.composable TTfgh.leg1_simps(2) Tfg.ρσ.T0.antipar(2)
              Tfg.ρσ.T0.ide_right Tfg_Hfg.u_simps(3) f.T0.antipar(2) f.T0.ide_right
              f.u_simps(3) fg g.ide_leg1 g.leg1_simps(4) h.ide_leg1 h.leg1_simps(4)
              ide_hcomp hseqE hcomp_simps(1) tab1_simps(1))
        thus ?thesis
          using fg gh Tfg.ρσ.prj_char TTfgh.prj_char isomorphic_reflexive
                Maps.comp_CLS [of "tab0 g" Tfg.ρσ.p0 "tab0 g  Tfg.ρσ.p0"]
                Maps.comp_CLS [of Tfg.ρσ.p1 TTfgh.p1 "Tfg.ρσ.p1  TTfgh.p1"]
          by (simp add: TTfgh.composable Tfg.ρσ.T0.antipar(2))
      qed
      show "SPN_fgh.Prj01 = ⟦⟦Tfg.ρσ.p0  TTfgh.p1⟧⟧"
      proof -
        have "ide (Tfg.ρσ.p0  TTfgh.p1)"
          by (metis TTfgh.leg1_simps(2) bicategory_of_spans.tab0_simps(1)
              bicategory_of_spans.tab1_simps(1) bicategory_of_spans_axioms
              Tfg.ρσ.T0.antipar(2) Tfg.ρσ.T0.ide_right Tfg.composable f.T0.antipar(2)
              f.T0.ide_right f.u_simps(3) g.ide_leg1 g.leg1_simps(4)
              Tfg.u_simps(3) THfgh.composable h.ide_leg1 h.leg1_simps(4)
              ide_hcomp hseqE hcomp_simps(1) tab1_simps(3))
        thus ?thesis
          using fg gh Tfg.ρσ.prj_char TTfgh.prj_char isomorphic_reflexive
                Maps.comp_CLS [of "tab0 g" Tfg.ρσ.p0 "tab0 g  Tfg.ρσ.p0"]
                Maps.comp_CLS [of Tfg.ρσ.p0 TTfgh.p1 "Tfg.ρσ.p0  TTfgh.p1"]
          by (simp add: Tfg.ρσ.T0.antipar(2) THfgh.composable)
      qed
      show "SPN_fgh.Prj0 = ⟦⟦TTfgh.p0⟧⟧"
        using isomorphic_reflexive TTfgh.prj_char Tfg.ρσ.prj_char
              Maps.comp_CLS [of "tab0 g" Tfg.ρσ.p0 "tab0 g  Tfg.ρσ.p0"]
        by (simp add: Tfg.composable)
      show "SPN_fgh.Prj1 = ⟦⟦TfTgh.p1⟧⟧"
        using Tgh.ρσ.prj_char isomorphic_reflexive Tgh.composable
              Maps.comp_CLS [of "tab1 g" Tgh.ρσ.p1 "tab1 g  Tgh.ρσ.p1"]
              TfTgh.prj_char
        by simp
      show "SPN_fgh.Prj10 = ⟦⟦Tgh.ρσ.p1  TfTgh.p0⟧⟧"
       using fg gh Tgh.ρσ.prj_char TfTgh.prj_char isomorphic_reflexive
              Maps.comp_CLS [of "tab1 g" "prj1 (tab1 h) (tab0 g)" "tab1 g  Tgh.ρσ.p1"]
              Maps.comp_CLS [of Tgh.ρσ.p1 TfTgh.p0 "Tgh.ρσ.p1  TfTgh.p0"]
        by simp
      show "SPN_fgh.Prj00 = ⟦⟦Tgh.ρσ.p0  TfTgh.p0⟧⟧"
        using fg gh Tgh.ρσ.prj_char TfTgh.prj_char isomorphic_reflexive
              Maps.comp_CLS [of "tab1 g" "Tgh.ρσ.p1" "tab1 g  Tgh.ρσ.p1"]
              Maps.comp_CLS [of Tgh.ρσ.p0 TfTgh.p0 "Tgh.ρσ.p0  TfTgh.p0"]
        by simp
    qed

    interpretation Φ: transformation_by_components VV.comp Span.vcomp
                        HoSPN_SPN.map SPNoH.map λrs. CMP (fst rs) (snd rs)
      using compositor_is_natural_transformation by simp
    interpretation Φ: natural_isomorphism VV.comp Span.vcomp
                        HoSPN_SPN.map SPNoH.map Φ.map
      using compositor_is_natural_isomorphism by simp

    (*
     * TODO: Figure out how this subcategory gets introduced.
     * The simps in the locale are used in the subsequent proofs.
     *)
    interpretation VVV': subcategory VxVxV.comp
                           λτμν. arr (fst τμν)  arr (fst (snd τμν))  arr (snd (snd τμν)) 
                                  src (fst (snd τμν)) = trg (snd (snd τμν)) 
                                  src (fst τμν) = trg (fst (snd τμν))
      using fg gh VVV.arr_charSbC VV.arr_charSbC VVV.subcategory_axioms by simp

    text ‹
      We define abbreviations for the left and right-hand sides of the equation for
      associativity coherence.
    ›
    (*
     * TODO: Φ doesn't really belong in this locale.  Replace it with CMP and rearrange
     * material so that this locale comes first and the definition of Φ comes later
     * in bicategory_of_spans.
     *)
    abbreviation LHS
    where "LHS  SPN 𝖺[f, g, h]  Φ.map (f  g, h)  (Φ.map (f, g)  SPN h)"

    abbreviation RHS
    where "RHS  Φ.map (f, g  h)  (SPN f  Φ.map (g, h)) 
                    Span.assoc (SPN f) (SPN g) (SPN h)"

    lemma arr_LHS:
    shows "Span.arr LHS"
      using fg gh VVV.arr_charSbC VVV.ide_charSbC VV.arr_charSbC VV.ide_charSbC Span.hseqI'
            HoHV_def compositor_in_hom α_def
      apply (intro Span.seqI)
          apply simp_all
      using SPN.FF_def
       apply simp
    proof -
      have "SPN ((f  g)  h) = Span.cod (CMP (f  g) h)"
        using fg gh compositor_in_hom by simp
      also have "... = Span.cod (CMP (f  g) h  (CMP f g  SPN h))"
      proof -
        have "Span.seq (CMP (f  g) h) (CMP f g  SPN h)"
        proof (intro Span.seqI Span.hseqI)
          show 1: "Span.in_hhom (SPN h) (SPN.map0 (src h)) (SPN.map0 (trg h))"
            using SPN.preserves_src SPN.preserves_trg by simp
          show 2: "Span.in_hhom (CMP f g) (SPN.map0 (trg h)) (SPN.map0 (trg f))"
            using compositor_in_hom SPN_fgh.νπ.composable fg by auto
          show 3: "Span.arr (CMP (f  g) h)"
            using TTfgh.composable Tfg.ρσ.ide_base compositor_simps(1) h.is_ide by auto
          show "Span.dom (CMP (f  g) h) = Span.cod (CMP f g  SPN h)"
            using 1 2 3 fg gh compositor_in_hom SPN_fgh.νπ.composable SPN_in_hom SPN.FF_def
            by auto
        qed
        thus ?thesis by simp
      qed
      finally show "SPN ((f  g)  h) = Span.cod (CMP (f  g) h  (CMP f g  SPN h))"
        by blast
    qed

    lemma arr_RHS:
    shows "Span.arr RHS"
      using fg gh VV.ide_charSbC VV.arr_charSbC Φ.map_simp_ide SPN.FF_def Span.hseqI'
      by (intro Span.seqI, simp_all)

    lemma par_LHS_RHS:
    shows "Span.par LHS RHS"
    proof (intro conjI)
      show "Span.arr LHS"
        using arr_LHS by simp
      show "Span.arr RHS"
        using arr_RHS by simp
      show "Span.dom LHS = Span.dom RHS"
      proof -
        have "Span.dom LHS = Span.dom (Φ.map (f, g)  SPN h)"
          using arr_LHS by auto
        also have "... = Span.dom (Φ.map (f, g))  Span.dom (SPN h)"
          using arr_LHS Span.dom_hcomp [of "SPN h" "Φ.map (f, g)"] by blast
        also have "... = (SPN f  SPN g)  SPN h"
          using fg Φ.map_simp_ide VV.ide_charSbC VV.arr_charSbC SPN.FF_def by simp
        also have "... = Span.dom (Span.assoc (SPN f) (SPN g) (SPN h))"
          using fg gh VVV.arr_charSbC VVV.ide_charSbC VV.arr_charSbC VV.ide_charSbC by simp
        also have "... = Span.dom RHS"
          using Span.arr RHS by auto
        finally show ?thesis by blast
      qed
      show "Span.cod LHS = Span.cod RHS"
      proof -
        have "Span.cod LHS = Span.cod (SPN 𝖺[f, g, h])"
          using arr_LHS by simp
        also have "... = SPN (f  g  h)"
          unfolding α_def
          using fg gh VVV.ide_charSbC VVV.arr_charSbC VV.ide_charSbC VV.arr_charSbC HoVH_def
          by simp
        also have "... = Span.cod RHS"
          using arr_RHS fg gh Φ.map_simp_ide VV.ide_charSbC VV.arr_charSbC SPN.FF_def
                compositor_in_hom
          by simp
        finally show ?thesis by blast
      qed
    qed

    lemma Chn_LHS_eq:
    shows "Chn LHS =
           ⟦⟦HHfgh_HfHgh.chine⟧⟧  ⟦⟦THfgh_HHfgh.chine⟧⟧  ⟦⟦TTfgh_THfgh.chine⟧⟧"
    proof -
      have "Chn LHS = ⟦⟦HHfgh_HfHgh.chine⟧⟧  ⟦⟦THfgh_HHfgh.chine⟧⟧ 
                        Span.chine_hcomp (CMP f g) (SPN h)"
      proof -
        have "Chn LHS = Chn (SPN 𝖺[f, g, h])  Chn (CMP (f  g) h) 
                          Chn (CMP f g  SPN h)"
          using fg gh arr_LHS Φ.map_simp_ide VV.ide_charSbC VV.arr_charSbC Span.Chn_vcomp
          by auto
        moreover have "Chn (SPN 𝖺[f, g, h]) = Maps.CLS HHfgh_HfHgh.chine"
          using fg gh SPN_def VVV.arr_charSbC VV.arr_charSbC spn_def α_def by simp
        moreover have "Chn (CMP (f  g) h) = Maps.CLS THfgh_HHfgh.chine"
          using fg gh CMP_def THfgh.cmp_def by simp
        moreover have "Chn (CMP f g  SPN h) = Span.chine_hcomp (CMP f g) (SPN h)"
          using fg gh Span.hcomp_def by simp
        ultimately show ?thesis by simp
      qed
      moreover have "Span.chine_hcomp (CMP f g) (SPN h) = ⟦⟦TTfgh_THfgh.chine⟧⟧"
      proof -
        have "Span.chine_hcomp (CMP f g) (SPN h) =
              Maps.tuple
                (⟦⟦Tfg.cmp⟧⟧  Maps.PRJ1 ⟦⟦tab0 g  Tfg.ρσ.p0⟧⟧ ⟦⟦tab1 h⟧⟧)
                   ⟦⟦tab0 (f  g)⟧⟧ ⟦⟦tab1 h⟧⟧
                (⟦⟦spn h⟧⟧  Maps.PRJ0 ⟦⟦tab0 g  Tfg.ρσ.p0⟧⟧ ⟦⟦tab1 h⟧⟧)"
        proof -
          have "⟦⟦tab0 g⟧⟧  ⟦⟦Tfg.ρσ.p0⟧⟧ = ⟦⟦tab0 g  Tfg.ρσ.p0⟧⟧"
            using fg isomorphic_reflexive
                  Maps.comp_CLS [of "tab0 g" "Tfg.ρσ.p0" "tab0 g  Tfg.ρσ.p0"]
            by simp
          moreover have "span_in_category.apex Maps.comp
                           Leg0 = ⟦⟦tab0 h⟧⟧, Leg1 = ⟦⟦tab1 h⟧⟧ =
                         ⟦⟦spn h⟧⟧"
          proof -
            interpret h: span_in_category Maps.comp Leg0 = ⟦⟦tab0 h⟧⟧, Leg1 = ⟦⟦tab1 h⟧⟧
              using h.determines_span by simp
            interpret dom_h: identity_arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg
                               dom h tab_of_ide (dom h) tab0 (dom h) tab1 (dom h)
                               cod h tab_of_ide (cod h) tab0 (cod h) tab1 (cod h)
                               h
              by (simp add: h.is_arrow_of_tabulations_in_maps
                  identity_arrow_of_tabulations_in_maps.intro
                  identity_arrow_of_tabulations_in_maps_axioms.intro)
            have "Maps.arr h.leg0"
              using h.leg_simps(1) by simp
            hence "Maps.dom h.leg0 = ⟦⟦dom_h.chine⟧⟧"
              using Maps.dom_char Maps.CLS_in_hom
              apply simp
            proof -
              have "h.is_induced_map (src (tab0 h))"
                using h.is_induced_map_iff dom_h.Δ_eq_ρ h.apex_is_induced_by_cell by force
              hence "src (tab0 h)  h.chine"
                using h.chine_is_induced_map h.induced_map_unique by simp
              thus "src (tab0 h) = h.chine"
                using iso_class_eqI by simp
            qed
            thus ?thesis
              using h.apex_def spn_def by simp
          qed
          ultimately show ?thesis
            unfolding Span.chine_hcomp_def
            using fg gh CMP_def Tfg.ρσ.prj_char Span.hcomp_def by simp
        qed
        also have "... = ⟦⟦TTfgh_THfgh.chine⟧⟧"
        proof -
          have "⟦⟦TTfgh_THfgh.chine⟧⟧ =
                Maps.tuple ⟦⟦Tfg_Hfg.chine  TTfgh.p1⟧⟧
                             ⟦⟦tab0 (f  g)⟧⟧ ⟦⟦tab1 h⟧⟧
                           ⟦⟦h.chine  TTfgh.p0⟧⟧"
            using TTfgh_THfgh.CLS_chine by simp
          also have "... =
                     Maps.tuple (⟦⟦Tfg_Hfg.chine⟧⟧  ⟦⟦TTfgh.p1⟧⟧)
                                  ⟦⟦tab0 (f  g)⟧⟧ ⟦⟦tab1 h⟧⟧
                                (⟦⟦h.chine⟧⟧  ⟦⟦TTfgh.p0⟧⟧)"
          proof -
            have "⟦⟦Tfg_Hfg.chine  TTfgh.p1⟧⟧ = ⟦⟦Tfg_Hfg.chine⟧⟧  ⟦⟦TTfgh.p1⟧⟧"
            proof -
              have "is_left_adjoint TTfgh.p1"
                using Tfg.ρσ.T0.antipar(2) THfgh.composable by simp
              moreover have "Tfg_Hfg.chine  TTfgh.p1  Tfg_Hfg.chine  TTfgh.p1"
                using TTfgh_THfgh.prj_chine(2) isomorphic_reflexive isomorphic_implies_hpar(2)
                by blast
              ultimately show ?thesis
                using Tfg_Hfg.is_map
                      Maps.comp_CLS [of Tfg_Hfg.chine TTfgh.p1 "Tfg_Hfg.chine  TTfgh.p1"]
                by simp
            qed
            moreover have "⟦⟦h.chine  TTfgh.p0⟧⟧ = ⟦⟦h.chine⟧⟧  ⟦⟦TTfgh.p0⟧⟧"
            proof -
              have "is_left_adjoint TTfgh.p0"
                by (simp add: Tfg.ρσ.T0.antipar(2) THfgh.composable)
              moreover have "h.chine  TTfgh.p0  h.chine  TTfgh.p0"
                using TTfgh_THfgh.prj_chine(1) isomorphic_reflexive isomorphic_implies_hpar(2)
                by blast
              ultimately show ?thesis
                using h.is_map Maps.comp_CLS [of h.chine TTfgh.p0 "h.chine  TTfgh.p0"]
                by simp
            qed
            ultimately show ?thesis by argo
          qed
          also have "... =
                     Maps.tuple (⟦⟦Tfg.cmp⟧⟧  Maps.PRJ1 ⟦⟦tab0 g  Tfg.ρσ.p0⟧⟧ ⟦⟦tab1 h⟧⟧)
                                  ⟦⟦tab0 (f  g)⟧⟧ ⟦⟦tab1 h⟧⟧
                                (⟦⟦spn h⟧⟧  Maps.PRJ0 ⟦⟦tab0 g  Tfg.ρσ.p0⟧⟧ ⟦⟦tab1 h⟧⟧)"
            using Tfg.cmp_def spn_def TTfgh.prj_char by simp
          finally show ?thesis by simp
        qed
        finally show ?thesis by blast
      qed
      ultimately show ?thesis by simp
    qed

    abbreviation tuple_BC
    where "tuple_BC  Maps.tuple SPN_fgh.Prj01 SPN_fgh.ν.leg0 SPN_fgh.π.leg1 SPN_fgh.Prj0"

    abbreviation tuple_ABC
    where "tuple_ABC  Maps.tuple SPN_fgh.Prj11
                                    SPN_fgh.μ.leg0
                                    (SPN_fgh.ν.leg1  SPN_fgh.νπ.prj1)
                                  tuple_BC"

    abbreviation tuple_BC'
    where "tuple_BC'  Maps.tuple ⟦⟦Tfg.ρσ.p0  TTfgh.p1⟧⟧ ⟦⟦tab0 g⟧⟧ ⟦⟦tab1 h⟧⟧ ⟦⟦TTfgh.p0⟧⟧"

    abbreviation tuple_ABC'
    where "tuple_ABC'  Maps.tuple ⟦⟦Tfg.ρσ.p1  TTfgh.p1⟧⟧
                                      ⟦⟦tab0 f⟧⟧ ⟦⟦tab1 g  Tgh.ρσ.p1⟧⟧
                                   tuple_BC'"

    lemma csq:
    shows "Maps.commutative_square SPN_fgh.ν.leg0 SPN_fgh.π.leg1
                                    SPN_fgh.Prj01 SPN_fgh.Prj0"
    and "Maps.commutative_square SPN_fgh.μ.leg0 (SPN_fgh.ν.leg1  SPN_fgh.νπ.prj1)
                                 SPN_fgh.Prj11 tuple_BC"
    proof -
      show 1: "Maps.commutative_square SPN_fgh.ν.leg0 SPN_fgh.π.leg1
                                       SPN_fgh.Prj01 SPN_fgh.Prj0"
      proof
        show "Maps.cospan SPN_fgh.ν.leg0 SPN_fgh.π.leg1"
          using SPN_fgh.νπ.legs_form_cospan(1) by simp
        show "Maps.span SPN_fgh.Prj01 SPN_fgh.Prj0"
          using SPN_fgh.prj_simps(2-3,5-6) by presburger
        show "Maps.dom SPN_fgh.ν.leg0 = Maps.cod SPN_fgh.Prj01"
          using SPN_fgh.prj_simps(8) SPN_g.dom.is_span SPN_g.dom.leg_simps(2)
          by auto
        show "SPN_fgh.ν.leg0  SPN_fgh.Prj01 = SPN_fgh.π.leg1  SPN_fgh.Prj0"
          by (metis (no_types, lifting) Maps.cod_comp Maps.comp_assoc
              Maps.pullback_commutes' SPN_fgh.μν.dom.leg_simps(1)
              SPN_fgh.μν.leg0_composite SPN_fgh.cospan_νπ)
      qed
      show "Maps.commutative_square
              SPN_fgh.μ.leg0 (Maps.comp SPN_fgh.ν.leg1 SPN_fgh.νπ.prj1)
              SPN_fgh.Prj11 tuple_BC"
      proof
        show "Maps.cospan SPN_fgh.μ.leg0 (SPN_fgh.ν.leg1  SPN_fgh.νπ.prj1)"
          using fg gh SPN_fgh.prj_simps(10) by blast
        show "Maps.span SPN_fgh.Prj11 tuple_BC"
          using fg gh 1 Maps.tuple_simps(1) Maps.tuple_simps(2) SPN_fgh.prj_simps(1)
                SPN_fgh.prj_simps(4) SPN_fgh.prj_simps(5)
          by presburger
        show "Maps.dom SPN_fgh.μ.leg0 = Maps.cod SPN_fgh.Prj11"
          using fg gh SPN_f.dom.leg_simps(2) SPN_fgh.prj_simps(7) by auto
        show "SPN_fgh.μ.leg0  SPN_fgh.Prj11 = (SPN_fgh.ν.leg1  SPN_fgh.νπ.prj1)  tuple_BC"
          using 1 fg gh Maps.comp_assoc Maps.prj_tuple
          by (metis (no_types, lifting) Maps.pullback_commutes' SPN_fgh.cospan_μν)
      qed
    qed

    lemma tuple_ABC_eq_ABC':
    shows "tuple_BC = tuple_BC'"
    and "tuple_ABC = tuple_ABC'"
    proof -
      have "SPN_fgh.Prj11 = ⟦⟦Tfg.ρσ.p1  TTfgh.p1⟧⟧"
        using prj_char by simp
      moreover have "SPN_fgh.μ.leg0 = ⟦⟦tab0 f⟧⟧"
        by simp
      moreover have "SPN_fgh.ν.leg1  SPN_fgh.νπ.prj1 = ⟦⟦tab1 g  Tgh.ρσ.p1⟧⟧"
        using Tgh.ρσ.prj_char isomorphic_reflexive Tgh.composable
              Maps.comp_CLS [of "tab1 g" Tgh.ρσ.p1 "tab1 g  Tgh.ρσ.p1"]
        by (simp add: g.T0.antipar(2))
      moreover show "tuple_BC = tuple_BC'"
        using prj_char Tfg.ρσ.prj_char by simp
      ultimately show "tuple_ABC = tuple_ABC'"
        by argo
    qed

    lemma tuple_BC_in_hom:
    shows "Maps.in_hom tuple_BC (Maps.MkIde (src TTfgh.p0)) (Maps.MkIde (src Tgh.ρσ.p0))"
    proof
      show 1: "Maps.arr tuple_BC"
        using csq(1) by simp
      have 2: "Maps.commutative_square
                 ⟦⟦tab0 g⟧⟧ ⟦⟦tab1 h⟧⟧ ⟦⟦Tfg.ρσ.p0  TTfgh.p1⟧⟧ ⟦⟦TTfgh.p0⟧⟧"
        by (metis Tfg.S0_def Tfg.span_legs_eq(3) Tgh.S1_def Tgh.span_legs_eq(4) csq(1)
                  prj_char(2) prj_char(3))
      show "Maps.dom tuple_BC = Maps.MkIde (src TTfgh.p0)"
      proof -
        have "Maps.dom tuple_BC' = Maps.dom ⟦⟦Tfg.ρσ.p0  TTfgh.p1⟧⟧"
          using 2 Maps.tuple_simps by simp
        also have "... = Chn (Span.hcomp (Span.hcomp (SPN f) (SPN g)) (SPN h))"
          using Maps.dom_char
          by (metis SPN_fgh.prj_simps(5) prj_char(2))
        also have "... = Maps.MkIde (src TTfgh.p0)"
          using 1 fg gh Maps.dom_char csq(1) prj_char(3) tuple_ABC_eq_ABC'(1)
                Maps.Dom.simps(1) Maps.tuple_simps(2) SPN_fgh.prj_simps(3,5-6)
          by presburger
        finally have "Maps.dom tuple_BC' = Maps.MkIde (src TTfgh.p0)"
          by blast
        thus ?thesis
          using tuple_ABC_eq_ABC' by simp
      qed
      show "Maps.cod tuple_BC = Maps.MkIde (src Tgh.ρσ.p0)"
      proof -
        have "Maps.cod tuple_BC' = Maps.pbdom ⟦⟦tab0 g⟧⟧ ⟦⟦tab1 h⟧⟧"
          using 1 2 fg gh Maps.tuple_in_hom by blast
        also have "... = Maps.MkIde (src Tgh.ρσ.p0)"
          using 1 2 fg gh Maps.pbdom_def
          by (metis (no_types, lifting) SPN.preserves_ide SPN_fgh.νπ.are_identities(2)
              SPN_fgh.νπ.composable Span.chine_hcomp_ide_ide Tfg.S0_def Tfg.span_legs_eq(3)
              Tgh.S1_def Tgh.chine_hcomp_SPN_SPN Tgh.span_legs_eq(4) g.is_ide)
        finally show ?thesis
          using tuple_ABC_eq_ABC' by simp
      qed
    qed

    lemma tuple_ABC_in_hom:
    shows "Maps.in_hom tuple_ABC (Maps.MkIde (src TTfgh.p0)) (Maps.MkIde (src TfTgh.p0))"
    proof
      show 1: "Maps.arr tuple_ABC"
        using SPN_fgh.chine_assoc_def SPN_fgh.chine_assoc_in_hom by auto
      show "Maps.dom tuple_ABC = Maps.MkIde (src TTfgh.p0)"
      proof -
        have "Maps.dom tuple_ABC = Maps.dom SPN_fgh.chine_assoc"
          by (simp add: SPN_fgh.chine_assoc_def)
        also have "... = Chn ((SPN f  SPN g)  SPN h)"
          using SPN_fgh.chine_assoc_in_hom by blast
        also have "... = Maps.MkIde (src TTfgh.p0)"
          by (metis (lifting) Maps.Dom.simps(1) Maps.dom_char SPN_fgh.prj_simps(3)
              SPN_fgh.prj_simps(6) prj_char(3))
        finally show ?thesis by blast
      qed
      show "Maps.cod tuple_ABC = Maps.MkIde (src TfTgh.p0)"
      proof -
        have "Maps.cod tuple_ABC = Maps.cod SPN_fgh.chine_assoc"
          by (simp add: SPN_fgh.chine_assoc_def)
        also have 1: "... = Chn (SPN f  SPN g  SPN h)"
          using SPN_fgh.chine_assoc_in_hom by blast
        also have "... = Maps.MkIde (src TfTgh.p0)"
          by (metis (lifting) Maps.Dom.simps(1) Maps.cod_char Maps.seq_char
              SPN_fgh.prj_chine_assoc(1) SPN_fgh.prj_simps(1) TfTgh.leg1_in_hom(1)
              TfTgh_TfTgh.u_in_hom 1 in_hhomE prj_char(4) src_hcomp)
        finally show ?thesis by argo
      qed
    qed

    lemma Chn_RHS_eq:
    shows "Chn RHS = ⟦⟦TfHgh_HfHgh.chine⟧⟧  ⟦⟦TfTgh_TfHgh.chine⟧⟧  tuple_ABC'"
    proof -
      have "Chn RHS =
            Chn (Φ.map (f, g  h))  Chn (SPN f  Φ.map (g, h)) 
              Chn (Span.assoc (SPN f) (SPN g) (SPN h))"
      proof -
        have "Chn RHS = Chn (Φ.map (f, g  h)) 
                          Chn ((SPN f  Φ.map (g, h))  Span.assoc (SPN f) (SPN g) (SPN h))"
          using arr_RHS Span.vcomp_eq Span.Chn_vcomp by blast
        also have "... = Chn (Φ.map (f, g  h))  Chn (SPN f  Φ.map (g, h)) 
                           Chn (Span.assoc (SPN f) (SPN g) (SPN h))"
        proof -
          have "Span.seq (SPN f  Φ.map (g, h)) (Span.assoc (SPN f) (SPN g) (SPN h))"
            using arr_RHS by auto
          thus ?thesis
            using fg gh Span.vcomp_eq [of "SPN f  Φ.map (g, h)"
                                          "Span.assoc (SPN f) (SPN g) (SPN h)"]
            by simp
        qed
        finally show ?thesis by blast
      qed
      moreover have "Chn (Φ.map (f, g  h)) = ⟦⟦TfHgh_HfHgh.chine⟧⟧"
        using arr_RHS fg gh Φ.map_simp_ide VV.ide_charSbC VV.arr_charSbC CMP_def TfHgh.cmp_def
        by simp
      moreover have "Chn (SPN f  Φ.map (g, h)) = Span.chine_hcomp (SPN f) (CMP g h)"
        using fg gh Span.hcomp_def Φ.map_simp_ide VV.ide_charSbC VV.arr_charSbC SPN.FF_def
        by simp
      moreover have "Chn (Span.assoc (SPN f) (SPN g) (SPN h)) = tuple_ABC"
        using fg gh Span.α_ide VVV.ide_charSbC VVV.arr_charSbC VV.ide_charSbC VV.arr_charSbC
              SPN_fgh.chine_assoc_def Span.α_def
        by simp
      moreover have "Span.chine_hcomp (SPN f) (CMP g h) = ⟦⟦TfTgh_TfHgh.chine⟧⟧"
      proof -
        have "Span.chine_hcomp (SPN f) (CMP g h) =
              Maps.tuple
                (⟦⟦f.chine⟧⟧  ⟦⟦TfTgh.p1⟧⟧)
                   ⟦⟦tab0 f⟧⟧ ⟦⟦tab1 (g  h)⟧⟧
                (⟦⟦Tgh_Hgh.chine⟧⟧  ⟦⟦TfTgh.p0⟧⟧)"
        proof -
          interpret f: span_in_category Maps.comp
                         Leg0 = Maps.MkArr (src (tab0 f)) (trg g) tab0 f,
                           Leg1 = Maps.MkArr (src (tab0 f)) (trg f) tab1 f
            using f.determines_span
            by (simp add: Tfg.composable)
          interpret f: arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg
                        f f.tab tab0 f tab1 f f f.tab tab0 f tab1 f f
            using f.is_arrow_of_tabulations_in_maps by simp
          have "f.apex = Maps.CLS f.chine"
          proof (intro Maps.arr_eqI)
            show "Maps.arr f.apex" by simp
            show "Maps.arr ⟦⟦f.chine⟧⟧"
              using Maps.CLS_in_hom f.is_map by blast
            show "Maps.Dom f.apex = Maps.Dom ⟦⟦f.chine⟧⟧"
              using f.apex_def Tfg.RS_simps(2) Tfg.R0_def Tfg.composable by auto
            show "Maps.Cod f.apex = Maps.Cod ⟦⟦f.chine⟧⟧"
              using f.apex_def Tfg.RS_simps(2) Tfg.R0_def Tfg.composable by auto
            show "Maps.Map f.apex = Maps.Map ⟦⟦f.chine⟧⟧"
            proof -
              have "Maps.Map f.apex = src (tab0 f)"
                using f.apex_def Maps.dom_char Tfg.RS_simps(2) Tfg.R0_def Tfg.composable
                by auto
              also have "... = f.chine"
              proof (intro iso_class_eqI)
                have "f.is_induced_map (src (tab0 f))"
                  using f.apex_is_induced_by_cell comp_cod_arr by auto
                thus "src (tab0 f)  f.chine"
                  using f.induced_map_unique f.chine_is_induced_map by simp
              qed
              also have "... = Maps.Map ⟦⟦f.chine⟧⟧"
                by simp
              finally show ?thesis by simp
            qed
          qed
          thus ?thesis
            unfolding Span.chine_hcomp_def
            using fg gh CMP_def Tgh.ρσ.prj_char Span.hcomp_def isomorphic_reflexive
                  Maps.comp_CLS [of "tab1 g" Tgh.ρσ.p1 "tab1 g  Tgh.ρσ.p1"]
                  Tgh.cmp_def TfTgh.prj_char
            by simp
        qed
        also have "... = Maps.tuple ⟦⟦f.chine  TfTgh.p1⟧⟧
                                      ⟦⟦tab0 f⟧⟧ ⟦⟦tab1 (g  h)⟧⟧
                                    ⟦⟦Tgh_Hgh.chine  TfTgh.p0⟧⟧"
          using isomorphic_reflexive TfHgh.composable f.is_map TfHgh.composable Tgh_Hgh.is_map
                Maps.comp_CLS [of f.chine TfTgh.p1 "f.chine  TfTgh.p1"]
                Maps.comp_CLS [of Tgh_Hgh.chine TfTgh.p0 "Tgh_Hgh.chine  TfTgh.p0"]
          by auto
        also have "... = ⟦⟦TfTgh_TfHgh.chine⟧⟧"
          using TfTgh_TfHgh.CLS_chine by simp
        finally show ?thesis by blast
      qed
      ultimately have "Chn RHS =⟦⟦TfHgh_HfHgh.chine⟧⟧  ⟦⟦TfTgh_TfHgh.chine⟧⟧  tuple_ABC"
        by simp
      also have "... = ⟦⟦TfHgh_HfHgh.chine⟧⟧  ⟦⟦TfTgh_TfHgh.chine⟧⟧  tuple_ABC'"
        using tuple_ABC_eq_ABC' by simp
      finally show ?thesis by simp
    qed

    interpretation g0h1: cospan_of_maps_in_bicategory_of_spans V H 𝖺 𝗂 src trg tab1 h tab0 g
      using gh by unfold_locales auto
    interpretation f0g1: cospan_of_maps_in_bicategory_of_spans V H 𝖺 𝗂 src trg tab1 g tab0 f
      using fg by unfold_locales auto
    interpretation f0gh1: cospan_of_maps_in_bicategory_of_spans V H 𝖺 𝗂 src trg
                            tab1 g  Tgh.ρσ.p1 tab0 f
      using fg gh Tgh.ρσ.leg1_is_map
      by unfold_locales auto
    interpretation fg0h1: cospan_of_maps_in_bicategory_of_spans V H 𝖺 𝗂 src trg
                            tab1 h tab0 g  Tfg.p0
      using TTfgh.r0s1_is_cospan by simp

    lemma src_tab_eq:
    shows "(𝖺-1[f, g, h]  tab0 h  TTfgh.p0) 
              TfTgh.composite_cell TTfgh_TfTgh.chine TTfgh_TfTgh.the_θ  TTfgh_TfTgh.the_ν =
           TTfgh.tab"
    proof -
      have "TfTgh.composite_cell TTfgh_TfTgh.chine TTfgh_TfTgh.the_θ  TTfgh_TfTgh.the_ν =
            (𝖺[f, g, h]  tab0 h  TTfgh.p0)  TTfgh.tab"
        unfolding TTfgh.tab_def
        using TTfgh_TfTgh.chine_is_induced_map TTfgh.tab_def TTfgh_TfTgh.Δ_simps(4)
        by auto
      moreover have "iso (𝖺[f, g, h]  tab0 h  TTfgh.p0)"
        by (simp add: fg gh)
      moreover have "inv (𝖺[f, g, h]  tab0 h  TTfgh.p0) = 𝖺-1[f, g, h]  tab0 h  TTfgh.p0"
        using fg gh by simp
      ultimately show ?thesis
        using TTfgh_TfTgh.Δ_simps(1)
              invert_side_of_triangle(1)
                [of "TfTgh.composite_cell TTfgh_TfTgh.chine TTfgh_TfTgh.the_θ  TTfgh_TfTgh.the_ν"
                    "𝖺[f, g, h]  tab0 h  TTfgh.p0" TTfgh.tab]
        by argo
    qed

    text ‹
      We need to show that the associativity isomorphism (defined in terms of tupling) coincides
      with TTfgh_TfTgh.chine› (defined in terms of tabulations).  In order to do this,
      we need to know how the latter commutes with projections.  That is the purpose of
      the following lemma.  Unfortunately, it requires some lengthy calculations,
      which I haven't seen any way to avoid.
    ›

    lemma prj_chine:
    shows "⟦⟦TfTgh.p1  TTfgh_TfTgh.chine⟧⟧ = ⟦⟦Tfg.p1  TTfgh.p1⟧⟧"
    and "⟦⟦Tgh.p1  TfTgh.p0  TTfgh_TfTgh.chine⟧⟧ = ⟦⟦Tfg.p0  TTfgh.p1⟧⟧"
    and "⟦⟦Tgh.p0  TfTgh.p0  TTfgh_TfTgh.chine⟧⟧ = ⟦⟦TTfgh.p0⟧⟧"
    proof -
      have 1: "ide TfTgh.p1"
        by (simp add: TfTgh.composable)
      have 2: "ide TTfgh_TfTgh.chine"
        by simp
      have 3: "src TfTgh.p1 = trg TTfgh_TfTgh.chine"
        using TTfgh_TfTgh.chine_in_hom(1) by simp
      have 4: "src (tab1 f) = trg TfTgh.p1"
        using TfTgh.leg1_simps(2) by blast
      text ‹
        The required isomorphisms will each be established via T2›, using the equation
        src_tab_eq› (associativities omitted from diagram):
$$
\begin{array}{l}
\xymatrix{
  && \xtwocell[dddd]{}\omit{^{\rm the\_}\nu}
  & \scriptstyle{{\rm TTfgh}.{\rm apex}} \ar[dd]^{{\rm chine}} \ar[dddlll]_{{\rm TfTgh}.p_1} \ar[dddrrr]^{{\rm TfTgh}.p_0}
  & \xtwocell[dddd]{}\omit{^{\rm the\_}\theta} \\
  &&&&& \\
  &&& \scriptstyle{{\rm TfTgh.apex}} \ar[ddll]_{{\rm TfTgh}.p_1} \ar[dr]^{{\rm TfTgh}.p_0} && \\
  \scriptstyle{f.{\rm apex}} \ar[dd]_{f.{\rm tab}_1}
  && \dtwocell\omit{^<-7>{f_0gh_1.\phi}}
  && \scriptstyle{{\rm Tgh.apex}} \ar[dl]_{{\rm Tgh}.p_1} \ar[dr]^{{\rm Tgh}.p_0} \ddtwocell\omit{^{g_0h_1.\phi}}
  && \scriptstyle{h.{\rm apex}} \ar[dd]^{h.{\rm tab}_0} \\
  & \scriptstyle{f.{\rm apex}} \ar[dl]_{f.{\rm tab}_1} \ar[dr]^{f.{\rm tab}_0} \dtwocell\omit{^f.{\rm tab}}
  && \scriptstyle{g.{\rm apex}} \ar[dl]_{g.{\rm tab}_1} \ar[dr]^{g.{\rm tab}_0} \dtwocell\omit{^g.{\rm tab}}
  && \scriptstyle{h.{\rm apex}} \ar[dl]_{h.{\rm tab}_1} \ar[dr]^{h.{\rm tab}_0}  \dtwocell\omit{^h.{\rm tab}} \\  
  \scriptstyle{{\rm trg}~f} && \scriptstyle{{\rm src}~f = {\rm trg}~g} \ar[ll]^{f}
  && \scriptstyle{{\rm src}~g = {\rm trg}~h} \ar[ll]^{g} && \scriptstyle{{\rm src}~h} \ar[ll]^{h}
}
\\
\\
\hspace{7cm}=
\\
\\
\xymatrix{
  &&& \scriptstyle{{\rm TTfgh.apex}} \ar[dl]_{{\rm TTfgh}.p_1} \ar[ddrr]^{{\rm TTfgh}.p_0} && \\
  && \scriptstyle{{\rm Tfg.apex}} \ar[dl]_{{\rm Tfg}.p_1} \ar[dr]^{{\rm Tfg}.p_0} \ddtwocell\omit{^{f_0g_1.\phi}}
  & \dtwocell\omit{^<-7>{fg_0h_1.\phi}} &&& \\
  & \scriptstyle{f.{\rm apex}} \ar[dl]_{f.{\rm tab}_1} \ar[dr]^{f.{\rm tab}_0} \dtwocell\omit{^f.{\rm tab}}
  && \scriptstyle{g.{\rm apex}} \ar[dl]_{g.{\rm tab}_1} \ar[dr]^{g.{\rm tab}_0} \dtwocell\omit{^g.{\rm tab}}
  && \scriptstyle{h.{\rm apex}} \ar[dl]_{h.{\rm tab}_1} \ar[dr]^{h.{\rm tab}_0}  \dtwocell\omit{^h.{\rm tab}} \\  
  \scriptstyle{{\rm trg}~f} && \scriptstyle{{\rm src}~f = {\rm trg}~g} \ar[ll]^{f}
  && \scriptstyle{{\rm src}~g = {\rm trg}~h} \ar[ll]^{g} && \scriptstyle{{\rm src}~h} \ar[ll]^{h}
}
\end{array}
$$
        There is a sequential dependence between the proofs, such as we have already
        seen for horizontal_composite_of_arrows_of_tabulations_in_maps.prj_chine›.
      ›
      define uf where "uf = g  h  tab0 h  TTfgh.p0"
      define wf where "wf = Tfg.p1  TTfgh.p1"
      define wf' where "wf' = TfTgh.p1  TTfgh_TfTgh.chine"
      define θf
      where "θf = (g  𝖺[h, tab0 h, TTfgh.p0])  (g  h.tab  TTfgh.p0)  (g  fg0h1) 
                   𝖺[g, tab0 g  Tfg.p0, TTfgh.p1]  (𝖺[g, tab0 g, Tfg.p0]  TTfgh.p1) 
                   ((g.tab  Tfg.p0)  TTfgh.p1)  (f0g1  TTfgh.p1) 
                   𝖺-1[tab0 f, Tfg.p1, TTfgh.p1]"
      define θf'
      where "θf' = (g  h  TTfgh_TfTgh.the_θ) 
                   can (g  h  ((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine)
                       (((g  (h  tab0 h)  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                   (((g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                   (((g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                   ((𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                   (((g.tab  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                   (f0gh1  TTfgh_TfTgh.chine) 
                   𝖺-1[tab0 f, TfTgh.p1, TTfgh_TfTgh.chine]"
      define βf
      where "βf = 𝖺[tab1 f, TfTgh.p1, TTfgh_TfTgh.chine] 
                  TTfgh_TfTgh.the_ν 
                  𝖺-1[tab1 f, Tfg.p1, TTfgh.p1]"
      have wf: "ide wf"
        using wf_def fg0h1.p1_simps by auto
      have wf_is_map: "is_left_adjoint wf"
        using wf_def fg0h1.p1_simps
          by (simp add: left_adjoints_compose)
      have wf': "ide wf'"
        unfolding wf'_def by simp
      have wf'_is_map: "is_left_adjoint wf'"
        unfolding wf'_def
        using 3 TTfgh_TfTgh.is_map f0gh1.leg1_is_map
        by (simp add: left_adjoints_compose)
      have θf: "«θf : tab0 f  wf  uf»"
      proof (unfold θf_def wf_def uf_def, intro comp_in_homI)
        show "«𝖺-1[tab0 f, Tfg.p1, TTfgh.p1] :
                 tab0 f  Tfg.p1  TTfgh.p1  (tab0 f  Tfg.p1)  TTfgh.p1»"
          using f0g1.leg1_simps fg0h1.p1_simps f0g1.cospan g0h1.cospan by auto
        show "«f0g1  TTfgh.p1 :
                 (tab0 f  Tfg.p1)  TTfgh.p1  (tab1 g  Tfg.p0)  TTfgh.p1»"
          using f0g1.φ_in_hom(2) Tfg.ρσ.T0.antipar(1)
          by (intro hcomp_in_vhom, auto)
        show "«(g.tab  Tfg.p0)  TTfgh.p1 :
                 (tab1 g  Tfg.p0)  TTfgh.p1  ((g  tab0 g)  Tfg.p0)  TTfgh.p1»"
          using Tfg.ρσ.T0.antipar(1)
          by (intro hcomp_in_vhom, auto)
        show "«𝖺[g, tab0 g, Tfg.p0]  TTfgh.p1 :
                 ((g  tab0 g)  Tfg.p0)  TTfgh.p1  (g  tab0 g  Tfg.p0)  TTfgh.p1»"
          using fg0h1.p1_simps
          by (intro hcomp_in_vhom, auto)
        show "«𝖺[g, tab0 g  Tfg.p0, TTfgh.p1] :
                (g  tab0 g  Tfg.p0)  TTfgh.p1  g  (tab0 g  Tfg.p0)  TTfgh.p1»"
          using fg0h1.p1_simps by auto
        show "«g  fg0h1 : g  (tab0 g  Tfg.p0)  TTfgh.p1  g  tab1 h  TTfgh.p0»"
          using fg0h1.φ_in_hom fg0h1.p1_simps
          by (intro hcomp_in_vhom, auto)
        show "«g  h.tab  TTfgh.p0 : g  tab1 h  TTfgh.p0  g  (h  tab0 h)  TTfgh.p0»"
          using gh fg0h1.φ_in_hom fg0h1.p1_simps
          by (intro hcomp_in_vhom, auto)
        show "«g  𝖺[h, tab0 h, TTfgh.p0] :
                g  (h  tab0 h)  TTfgh.p0  g  h  tab0 h  TTfgh.p0»"
          using gh fg0h1.φ_in_hom fg0h1.p1_simps
          by (intro hcomp_in_vhom, auto)
      qed
      have θf': "«θf' : tab0 f  wf'  uf»"
      proof (unfold θf'_def wf'_def uf_def, intro comp_in_homI)
        show "«𝖺-1[tab0 f, TfTgh.p1, TTfgh_TfTgh.chine] :
                 tab0 f  TfTgh.p1  TTfgh_TfTgh.chine  (tab0 f  TfTgh.p1)  TTfgh_TfTgh.chine»"
          using "1" "2" "3" "4" assoc'_in_hom(2) f.ide_u f.leg1_simps(3) by auto
        show "«f0gh1  TTfgh_TfTgh.chine :
                 (tab0 f  TfTgh.p1)  TTfgh_TfTgh.chine 
                     ((tab1 g  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine»"
          using f0gh1.φ_in_hom(2)
          by (intro hcomp_in_vhom, auto)
        show "«((g.tab  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine :
                 ((tab1 g  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine
                          (((g  tab0 g)  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine»"
          using f0gh1.cospan g0h1.cospan
          by (intro hcomp_in_vhom, auto)
        show "«(𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine :
                 (((g  tab0 g)  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine
                     ((g  tab0 g  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine»"
          using f0gh1.cospan g0h1.cospan
          by (intro hcomp_in_vhom, auto)
        show "«((g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine :
                ((g  tab0 g  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine
                    ((g  tab1 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine»"
          using f0gh1.cospan g0h1.cospan g0h1.φ_in_hom(2)
          by (intro hcomp_in_vhom, auto)
        show "«((g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine :
                ((g  tab1 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine
                    ((g  (h  tab0 h)  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine»"
          using f0gh1.cospan g0h1.cospan
          by (intro hcomp_in_vhom, auto)
        show "«can (g  h  ((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine)
                   (((g  (h  tab0 h)  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) :
                   ((g  (h  tab0 h)  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine
                      g  h  ((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine»"
          using f0gh1.cospan g0h1.cospan by auto
        show "«g  h  TTfgh_TfTgh.the_θ :
                 g  h  ((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine
                     g  h  tab0 h  TTfgh.p0»"
          using f0gh1.cospan g0h1.cospan TTfgh_TfTgh.the_θ_in_hom
          by (intro hcomp_in_vhom, auto)
      qed
      have βf: "«βf : tab1 f  wf  tab1 f  wf'»"
      proof (unfold βf_def wf_def wf'_def, intro comp_in_homI)
        show "«𝖺-1[tab1 f, Tfg.p1, TTfgh.p1] :
                 tab1 f  Tfg.p1  TTfgh.p1  (tab1 f  Tfg.p1)  TTfgh.p1»"
          using TTfgh.leg1_in_hom(2) assoc'_in_hom by auto
        show "«TTfgh_TfTgh.the_ν :
                 (tab1 f  Tfg.p1)  TTfgh.p1  (tab1 f  TfTgh.p1)  TTfgh_TfTgh.chine»"
          using TTfgh_TfTgh.the_ν_in_hom TTfgh_TfTgh.the_ν_props by simp
        show "«𝖺[tab1 f, TfTgh.p1, TTfgh_TfTgh.chine] :
                (tab1 f  TfTgh.p1)  TTfgh_TfTgh.chine  tab1 f  TfTgh.p1  TTfgh_TfTgh.chine»"
          using 1 2 3 4 by auto
      qed
      have iso_βf: "iso βf"
        unfolding βf_def
        using 1 2 3 4 βf βf_def isos_compose
        apply (intro isos_compose)
            apply (metis TTfgh.composable TTfgh.leg1_in_hom(2) Tfg.ρσ.T0.antipar(2)
                         Tfg.ρσ.T0.ide_right Tfg.ρσ.leg1_in_hom(2) Tfg_Hfg.u_simps(3)
                         f.T0.antipar(2) f.T0.ide_right f.ide_leg1 f0g1.cospan g.ide_leg1
                         h.ide_leg1 h.leg1_simps(4) hcomp_in_vhomE ide_hcomp
                         iso_assoc' tab1_simps(1))
        using TTfgh_TfTgh.the_ν_props(2) f.ide_leg1 iso_assoc by blast+
      have uf: "ide uf"
        using θf ide_cod by blast
      have wf_in_hhom: "in_hhom wf (src uf) (src (tab0 f))"
        using uf wf uf_def wf_def by simp
      have wf'_in_hhom: "in_hhom wf' (src uf) (src (tab0 f))"
        using uf wf' wf'_def uf_def by simp
      have 5: "∃!γ. «γ : wf  wf'»  βf = tab1 f  γ  θf = θf'  (tab0 f  γ)"
      proof -
        have eqf: "f.composite_cell wf θf = f.composite_cell wf' θf'  βf"
        proof -
          text ‹
            I don't see any alternative here to just grinding out the calculation.
            The idea is to bring f.composite_cell wf θf into a form in which
            src_tab_eq› can be applied to eliminate θf in favor of θf'›.
          ›
          have "f.composite_cell wf θf =
                  ((f  g  𝖺[h, tab0 h, fg0h1.p0]) 
                    (f  g  h.tab  fg0h1.p0) 
                    (f  g  fg0h1) 
                    (f  𝖺[g, tab0 g  f0g1.p0, fg0h1.p1]) 
                    (f  𝖺[g, tab0 g, f0g1.p0]  fg0h1.p1) 
                    (f  (g.tab  f0g1.p0)  fg0h1.p1) 
                    (f  f0g1  fg0h1.p1) 
                    (f  𝖺-1[tab0 f, f0g1.p1, fg0h1.p1])) 
                    𝖺[f, tab0 f, f0g1.p1  fg0h1.p1] 
                    (f.tab  f0g1.p1  fg0h1.p1)"
            unfolding wf_def θf_def
            using fg0h1.p1_simps Tgh.composable whisker_left by simp  (* 12 sec, 30 sec cpu *)
          also have "... =
                  (f  g  𝖺[h, tab0 h, TTfgh.p0]) 
                    (f  g  h.tab  TTfgh.p0) 
                    (f  g  fg0h1) 
                    (f  𝖺[g, tab0 g  Tfg.p0, TTfgh.p1]) 
                    (f  𝖺[g, tab0 g, Tfg.p0]  TTfgh.p1) 
                    (f  (g.tab  Tfg.p0)  TTfgh.p1) 
                    (f  f0g1  TTfgh.p1) 
                    (f  𝖺-1[tab0 f, Tfg.p1, TTfgh.p1]) 
                    𝖺[f, tab0 f, Tfg.p1  TTfgh.p1] 
                    (f.tab  Tfg.p1  TTfgh.p1)"
            using comp_assoc by simp
          also have "... =
                       (𝖺[f, g, h  tab0 h  TTfgh.p0] 
                         𝖺-1[f, g, h  tab0 h  TTfgh.p0] 
                         (f  g  𝖺[h, tab0 h, TTfgh.p0])) 
                       (f  g  h.tab  TTfgh.p0) 
                       (f  g  fg0h1) 
                       (f  𝖺[g, tab0 g  Tfg.p0, TTfgh.p1]) 
                       (f  𝖺[g, tab0 g, Tfg.p0]  TTfgh.p1) 
                       (f  (g.tab  Tfg.p0)  TTfgh.p1) 
                       (f  f0g1  TTfgh.p1) 
                       (f  𝖺-1[tab0 f, Tfg.p1, TTfgh.p1]) 
                       𝖺[f, tab0 f, Tfg.p1  TTfgh.p1] 
                       (f.tab  Tfg.p1  TTfgh.p1)"
          proof -
            have "(𝖺[f, g, h  tab0 h  TTfgh.p0] 
                    𝖺-1[f, g, h  tab0 h  TTfgh.p0]) 
                    (f  g  𝖺[h, tab0 h, TTfgh.p0]) =
                  f  g  𝖺[h, tab0 h, TTfgh.p0]"
              using fg gh fg0h1.p1_simps comp_cod_arr comp_assoc_assoc' by simp
            thus ?thesis
              using comp_assoc by simp
          qed
          also have "... =
                       𝖺[f, g, h  tab0 h  TTfgh.p0] 
                       (𝖺-1[f, g, h  tab0 h  TTfgh.p0] 
                         (f  g  𝖺[h, tab0 h, TTfgh.p0])) 
                       (f  g  h.tab  TTfgh.p0) 
                       (f  g  fg0h1) 
                       (f  𝖺[g, tab0 g  Tfg.p0, TTfgh.p1]) 
                       (f  𝖺[g, tab0 g, Tfg.p0]  TTfgh.p1) 
                       (f  (g.tab  Tfg.p0)  TTfgh.p1) 
                       (f  f0g1  TTfgh.p1) 
                       (f  𝖺-1[tab0 f, Tfg.p1, TTfgh.p1]) 
                       𝖺[f, tab0 f, Tfg.p1  TTfgh.p1] 
                       (f.tab  Tfg.p1  TTfgh.p1)"
            using comp_assoc by presburger
          also have "... =
                       𝖺[f, g, h  tab0 h  TTfgh.p0] 
                       ((f  g)  𝖺[h, tab0 h, TTfgh.p0]) 
                       (𝖺-1[f, g, (h  tab0 h)  TTfgh.p0] 
                       (f  g  h.tab  TTfgh.p0)) 
                       (f  g  fg0h1) 
                       (f  𝖺[g, tab0 g  Tfg.p0, TTfgh.p1]) 
                       (f  𝖺[g, tab0 g, Tfg.p0]  TTfgh.p1) 
                       (f  (g.tab  Tfg.p0)  TTfgh.p1) 
                       (f  f0g1  TTfgh.p1) 
                       (f  𝖺-1[tab0 f, Tfg.p1, TTfgh.p1]) 
                       𝖺[f, tab0 f, Tfg.p1  TTfgh.p1] 
                       (f.tab  Tfg.p1  TTfgh.p1)"
            using fg gh fg0h1.p0_simps fg0h1.p1_simps comp_assoc
                  assoc'_naturality [of f g "𝖺[h, tab0 h, TTfgh.p0]"]
            by simp
          also have "... =
                       𝖺[f, g, h  tab0 h  TTfgh.p0] 
                       ((f  g)  𝖺[h, tab0 h, TTfgh.p0]) 
                       ((f  g)  h.tab  TTfgh.p0) 
                       (𝖺-1[f, g, tab1 h  TTfgh.p0] 
                       (f  g  fg0h1)) 
                       (f  𝖺[g, tab0 g  Tfg.p0, TTfgh.p1]) 
                       (f  𝖺[g, tab0 g, Tfg.p0]  TTfgh.p1) 
                       (f  (g.tab  Tfg.p0)  TTfgh.p1) 
                       (f  f0g1  TTfgh.p1) 
                       (f  𝖺-1[tab0 f, Tfg.p1, TTfgh.p1]) 
                       𝖺[f, tab0 f, Tfg.p1  TTfgh.p1] 
                       (f.tab  Tfg.p1  TTfgh.p1)"
            using fg gh fg0h1.p0_simps fg0h1.p1_simps comp_assoc
                  assoc'_naturality [of f g "h.tab  TTfgh.p0"]
            by simp
          also have "... =
                       𝖺[f, g, h  tab0 h  TTfgh.p0] 
                       ((f  g)  𝖺[h, tab0 h, TTfgh.p0]) 
                       ((f  g)  h.tab  TTfgh.p0) 
                       ((f  g)  fg0h1) 
                       𝖺-1[f, g, (tab0 g  Tfg.p0)  TTfgh.p1] 
                       (f  𝖺[g, tab0 g  Tfg.p0, TTfgh.p1]) 
                       (f  𝖺[g, tab0 g, Tfg.p0]  TTfgh.p1) 
                       (f  (g.tab  Tfg.p0)  TTfgh.p1) 
                       (f  f0g1  TTfgh.p1) 
                       ((f  𝖺-1[tab0 f, Tfg.p1, TTfgh.p1]) 
                       𝖺[f, tab0 f, Tfg.p1  TTfgh.p1]) 
                       (f.tab  Tfg.p1  TTfgh.p1)"
            using fg gh fg0h1.p0_simps fg0h1.p1_simps comp_assoc
                  assoc'_naturality [of f g fg0h1]
            by simp
          also have "... =
                       𝖺[f, g, h  tab0 h  TTfgh.p0] 
                       ((f  g)  𝖺[h, tab0 h, TTfgh.p0]) 
                       ((f  g)  h.tab  TTfgh.p0) 
                       ((f  g)  fg0h1) 
                       𝖺-1[f, g, (tab0 g  Tfg.p0)  TTfgh.p1] 
                       (f  𝖺[g, tab0 g  Tfg.p0, TTfgh.p1]) 
                       (f  𝖺[g, tab0 g, Tfg.p0]  TTfgh.p1) 
                       (f  (g.tab  Tfg.p0)  TTfgh.p1) 
                       ((f  f0g1  TTfgh.p1) 
                       𝖺[f, tab0 f  Tfg.p1, TTfgh.p1]) 
                       (𝖺[f, tab0 f, Tfg.p1]  TTfgh.p1) 
                       𝖺-1[f  tab0 f, Tfg.p1, TTfgh.p1] 
                       (f.tab  Tfg.p1  TTfgh.p1)"
          proof -
            have "(f  𝖺-1[tab0 f, Tfg.p1, TTfgh.p1])  𝖺[f, tab0 f, Tfg.p1  TTfgh.p1] =
                  𝖺[f, tab0 f  Tfg.p1, TTfgh.p1]  (𝖺[f, tab0 f, Tfg.p1]  TTfgh.p1) 
                    𝖺-1[f  tab0 f, Tfg.p1, TTfgh.p1]"
            proof -
              have "(f  𝖺-1[tab0 f, Tfg.p1, TTfgh.p1])  𝖺[f, tab0 f, Tfg.p1  TTfgh.p1] =
                    (f  𝖺-1[tab0 f, Tfg.p1, TTfgh.p1]) 
                              𝖺[f, tab0 f, Tfg.p1  TTfgh.p1]"
                using fg gh fg0h1.p0_simps fg0h1.p1_simps f0g1.p0_simps f0g1.p1_simps
                      𝖺'_def α_def
                by simp
              also have "... =
                         𝖺[f, tab0 f  Tfg.p1, TTfgh.p1] 
                          (𝖺[f, tab0 f, Tfg.p1]  TTfgh.p1) 
                          𝖺-1[f  tab0 f, Tfg.p1, TTfgh.p1]"
                using fg gh fg0h1.p0_simps fg0h1.p1_simps f0g1.p0_simps f0g1.p1_simps
                by (intro E.eval_eqI, simp_all)
              also have "... = 𝖺[f, tab0 f  Tfg.p1, TTfgh.p1]  (𝖺[f, tab0 f, Tfg.p1]  TTfgh.p1) 
                                 𝖺-1[f  tab0 f, Tfg.p1, TTfgh.p1]"
                using fg gh fg0h1.p0_simps fg0h1.p1_simps f0g1.p0_simps f0g1.p1_simps
                      𝖺'_def α_def
                by simp
              finally show ?thesis by blast
            qed
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... =
                       𝖺[f, g, h  tab0 h  TTfgh.p0] 
                       ((f  g)  𝖺[h, tab0 h, TTfgh.p0]) 
                       ((f  g)  h.tab  TTfgh.p0) 
                       ((f  g)  fg0h1) 
                       𝖺-1[f, g, (tab0 g  Tfg.p0)  TTfgh.p1] 
                       (f  𝖺[g, tab0 g  Tfg.p0, TTfgh.p1]) 
                       (f  𝖺[g, tab0 g, Tfg.p0]  TTfgh.p1) 
                       ((f  (g.tab  Tfg.p0)  TTfgh.p1) 
                       𝖺[f, tab1 g  Tfg.p0, TTfgh.p1]) 
                       ((f  f0g1)  TTfgh.p1) 
                       (𝖺[f, tab0 f, Tfg.p1]  TTfgh.p1) 
                       𝖺-1[f  tab0 f, Tfg.p1, TTfgh.p1] 
                       (f.tab  Tfg.p1  TTfgh.p1)"
          proof -
            have "(f  f0g1  TTfgh.p1)  𝖺[f, tab0 f  Tfg.p1, TTfgh.p1] =
                  𝖺[f, tab1 g  Tfg.p0, TTfgh.p1]  ((f  f0g1)  TTfgh.p1)"
              using fg gh fg0h1.p0_simps fg0h1.p1_simps f0g1.φ_in_hom
                    assoc_naturality [of f f0g1 TTfgh.p1]
              by simp
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... =
                       𝖺[f, g, h  tab0 h  TTfgh.p0] 
                       ((f  g)  𝖺[h, tab0 h, TTfgh.p0]) 
                       ((f  g)  h.tab  TTfgh.p0) 
                       ((f  g)  fg0h1) 
                       𝖺-1[f, g, (tab0 g  Tfg.p0)  TTfgh.p1] 
                       (f  𝖺[g, tab0 g  Tfg.p0, TTfgh.p1]) 
                       (f  𝖺[g, tab0 g, Tfg.p0]  TTfgh.p1) 
                       𝖺[f, (g  tab0 g)  Tfg.p0, TTfgh.p1] 
                       ((f  (g.tab  Tfg.p0))  TTfgh.p1) 
                       ((f  f0g1)  TTfgh.p1) 
                       (𝖺[f, tab0 f, Tfg.p1]  TTfgh.p1) 
                       𝖺-1[f  tab0 f, Tfg.p1, TTfgh.p1] 
                       (f.tab  Tfg.p1  TTfgh.p1)"
          proof -
            have "(f  (g.tab  Tfg.p0)  TTfgh.p1)  𝖺[f, tab1 g  Tfg.p0, TTfgh.p1] =
                  𝖺[f, (g  tab0 g)  Tfg.p0, TTfgh.p1]  ((f  (g.tab  Tfg.p0))  TTfgh.p1)"
              using fg gh fg0h1.p0_simps fg0h1.p1_simps f0g1.φ_in_hom
                    assoc_naturality [of f "g.tab  Tfg.p0" TTfgh.p1]
              by simp
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... =
                       𝖺[f, g, h  tab0 h  TTfgh.p0] 
                       ((f  g)  𝖺[h, tab0 h, TTfgh.p0]) 
                       ((f  g)  h.tab  TTfgh.p0) 
                       ((f  g)  fg0h1) 
                       𝖺-1[f, g, (tab0 g  Tfg.p0)  TTfgh.p1] 
                       (f  𝖺[g, tab0 g  Tfg.p0, TTfgh.p1]) 
                       ((f  𝖺[g, tab0 g, Tfg.p0]  TTfgh.p1) 
                       𝖺[f, (g  tab0 g)  Tfg.p0, TTfgh.p1]) 
                       ((f  (g.tab  Tfg.p0))  TTfgh.p1) 
                       ((f  f0g1)  TTfgh.p1) 
                       (𝖺[f, tab0 f, Tfg.p1]  TTfgh.p1) 
                       ((f.tab  Tfg.p1)  TTfgh.p1) 
                       𝖺-1[tab1 f, Tfg.p1, TTfgh.p1]"
            using fg0h1.p1_simps assoc'_naturality [of f.tab Tfg.p1 TTfgh.p1] comp_assoc
            by simp
          also have "... =
                       𝖺[f, g, h  tab0 h  TTfgh.p0] 
                       ((f  g)  𝖺[h, tab0 h, TTfgh.p0]) 
                       ((f  g)  h.tab  TTfgh.p0) 
                       ((f  g)  fg0h1) 
                       𝖺-1[f, g, (tab0 g  Tfg.p0)  TTfgh.p1] 
                       (f  𝖺[g, tab0 g  Tfg.p0, TTfgh.p1]) 
                       (f  𝖺[g, tab0 g, Tfg.p0]  TTfgh.p1) 
                       𝖺[f, (g  tab0 g)  Tfg.p0, TTfgh.p1] 
                       ((((f  𝖺-1[g, tab0 g, Tfg.p0])  TTfgh.p1) 
                       ((f  𝖺[g, tab0 g, Tfg.p0])  TTfgh.p1)) 
                       ((f  (g.tab  Tfg.p0))  TTfgh.p1)) 
                       ((f  f0g1)  TTfgh.p1) 
                       (𝖺[f, tab0 f, Tfg.p1]  TTfgh.p1) 
                       ((f.tab  Tfg.p1)  TTfgh.p1) 
                       𝖺-1[tab1 f, Tfg.p1, TTfgh.p1]"
          proof -
            have "(((f  𝖺-1[g, tab0 g, Tfg.p0])  TTfgh.p1) 
                     ((f  𝖺[g, tab0 g, Tfg.p0])  TTfgh.p1)) 
                     ((f  (g.tab  Tfg.p0))  TTfgh.p1) =
                  (f  (g.tab  Tfg.p0))  TTfgh.p1"
              using fg gh fg0h1.p1_simps comp_cod_arr whisker_right comp_assoc_assoc'
                    whisker_left [of f "𝖺-1[g, tab0 g, Tfg.p0]" "𝖺[g, tab0 g, Tfg.p0]"]
              by simp
            thus ?thesis
              using comp_assoc by simp
          qed
          also have "... =
                       𝖺[f, g, h  tab0 h  TTfgh.p0] 
                       ((f  g)  𝖺[h, tab0 h, TTfgh.p0]) 
                       ((f  g)  h.tab  TTfgh.p0) 
                       ((f  g)  fg0h1) 
                       𝖺-1[f, g, (tab0 g  Tfg.p0)  TTfgh.p1] 
                       (f  𝖺[g, tab0 g  Tfg.p0, TTfgh.p1]) 
                       (f  𝖺[g, tab0 g, Tfg.p0]  TTfgh.p1) 
                       𝖺[f, (g  tab0 g)  Tfg.p0, TTfgh.p1] 
                       ((f  𝖺-1[g, tab0 g, Tfg.p0])  TTfgh.p1) 
                       ((f  𝖺[g, tab0 g, Tfg.p0])  TTfgh.p1) 
                       ((f  (g.tab  Tfg.p0))  TTfgh.p1) 
                       ((f  f0g1)  TTfgh.p1) 
                       (𝖺[f, tab0 f, Tfg.p1]  TTfgh.p1) 
                       ((f.tab  Tfg.p1)  TTfgh.p1) 
                       𝖺-1[tab1 f, Tfg.p1, TTfgh.p1]"
            using comp_assoc by presburger
          also have "... =
                       𝖺[f, g, h  tab0 h  TTfgh.p0] 
                       ((f  g)  𝖺[h, tab0 h, TTfgh.p0]) 
                       ((f  g)  h.tab  TTfgh.p0) 
                       ((f  g)  fg0h1) 
                       𝖺-1[f, g, (tab0 g  Tfg.p0)  TTfgh.p1] 
                       (f  𝖺[g, tab0 g  Tfg.p0, TTfgh.p1]) 
                       (f  𝖺[g, tab0 g, Tfg.p0]  TTfgh.p1) 
                       𝖺[f, (g  tab0 g)  Tfg.p0, TTfgh.p1] 
                       ((f  𝖺-1[g, tab0 g, Tfg.p0])  TTfgh.p1) 
                       (((𝖺[f, g, tab0 g  Tfg.p0]  TTfgh.p1) 
                       (𝖺-1[f, g, tab0 g  Tfg.p0]  TTfgh.p1)) 
                       ((f  𝖺[g, tab0 g, Tfg.p0])  TTfgh.p1)) 
                       ((f  (g.tab  Tfg.p0))  TTfgh.p1) 
                       ((f  f0g1)  TTfgh.p1) 
                       (𝖺[f, tab0 f, Tfg.p1]  TTfgh.p1) 
                       ((f.tab  Tfg.p1)  TTfgh.p1) 
                       𝖺-1[tab1 f, Tfg.p1, TTfgh.p1]"
          proof -
            have "((𝖺[f, g, tab0 g  Tfg.p0]  TTfgh.p1) 
                     (𝖺-1[f, g, tab0 g  Tfg.p0]  TTfgh.p1)) 
                     ((f  𝖺[g, tab0 g, Tfg.p0])  TTfgh.p1) =
                  (f  𝖺[g, tab0 g, Tfg.p0])  TTfgh.p1"
              using fg fg0h1.p1_simps comp_cod_arr comp_assoc_assoc'
                    whisker_right
                      [of TTfgh.p1 "𝖺[f, g, tab0 g  Tfg.p0]" "𝖺-1[f, g, tab0 g  Tfg.p0]"]
              by simp
            thus ?thesis
              using comp_assoc by simp
          qed
          also have "... =
                       𝖺[f, g, h  tab0 h  TTfgh.p0] 
                       ((f  g)  𝖺[h, tab0 h, TTfgh.p0]) 
                       ((f  g)  h.tab  TTfgh.p0) 
                       ((f  g)  fg0h1) 
                       𝖺-1[f, g, (tab0 g  Tfg.p0)  TTfgh.p1] 
                       (f  𝖺[g, tab0 g  Tfg.p0, TTfgh.p1]) 
                       (f  𝖺[g, tab0 g, Tfg.p0]  TTfgh.p1) 
                       𝖺[f, (g  tab0 g)  Tfg.p0, TTfgh.p1] 
                       ((f  𝖺-1[g, tab0 g, Tfg.p0])  TTfgh.p1) 
                       (𝖺[f, g, tab0 g  Tfg.p0]  TTfgh.p1) 
                       ((𝖺-1[f, g, tab0 g  Tfg.p0]  TTfgh.p1) 
                       ((f  𝖺[g, tab0 g, Tfg.p0])  TTfgh.p1) 
                       ((f  (g.tab  Tfg.p0))  TTfgh.p1) 
                       ((f  f0g1)  TTfgh.p1) 
                       (𝖺[f, tab0 f, Tfg.p1]  TTfgh.p1) 
                       ((f.tab  Tfg.p1)  TTfgh.p1)) 
                       𝖺-1[tab1 f, Tfg.p1, TTfgh.p1]"
            using comp_assoc by presburger
          also have "... =
                       𝖺[f, g, h  tab0 h  TTfgh.p0] 
                       ((f  g)  𝖺[h, tab0 h, TTfgh.p0]) 
                       ((f  g)  h.tab  TTfgh.p0) 
                       ((f  g)  fg0h1) 
                       (𝖺-1[f, g, (tab0 g  Tfg.p0)  TTfgh.p1] 
                       (f  𝖺[g, tab0 g  Tfg.p0, TTfgh.p1]) 
                       (f  𝖺[g, tab0 g, Tfg.p0]  TTfgh.p1) 
                       𝖺[f, (g  tab0 g)  Tfg.p0, TTfgh.p1] 
                       ((f  𝖺-1[g, tab0 g, Tfg.p0])  TTfgh.p1) 
                       (𝖺[f, g, tab0 g  Tfg.p0]  TTfgh.p1)) 
                       (𝖺-1[f, g, tab0 g  Tfg.p0] 
                         (f  𝖺[g, tab0 g, Tfg.p0]) 
                         (f  (g.tab  Tfg.p0)) 
                         (f  f0g1) 
                         𝖺[f, tab0 f, Tfg.p1] 
                         (f.tab  Tfg.p1)
                         TTfgh.p1) 
                       𝖺-1[tab1 f, Tfg.p1, TTfgh.p1]"
            using fg gh fg0h1.p0_simps fg0h1.p1_simps f0g1.p0_simps f0g1.p1_simps
                  whisker_right comp_assoc
            by simp
          also have "... =
                       𝖺[f, g, h  tab0 h  TTfgh.p0] 
                       ((f  g)  𝖺[h, tab0 h, TTfgh.p0]) 
                       ((f  g)  h.tab  TTfgh.p0) 
                       ((f  g)  fg0h1) 
                       𝖺[f  g, tab0 g  Tfg.p0, TTfgh.p1] 
                       (𝖺-1[f, g, tab0 g  Tfg.p0] 
                         (f  𝖺[g, tab0 g, Tfg.p0]) 
                         (f  (g.tab  Tfg.p0)) 
                         (f  f0g1) 
                         𝖺[f, tab0 f, Tfg.p1] 
                         (f.tab  Tfg.p1)
                         TTfgh.p1) 
                       𝖺-1[tab1 f, Tfg.p1, TTfgh.p1]"
          proof -
            have "𝖺-1[f, g, (tab0 g  Tfg.p0)  TTfgh.p1] 
                       (f  𝖺[g, tab0 g  Tfg.p0, TTfgh.p1]) 
                       (f  𝖺[g, tab0 g, Tfg.p0]  TTfgh.p1) 
                       𝖺[f, (g  tab0 g)  Tfg.p0, TTfgh.p1] 
                       ((f  𝖺-1[g, tab0 g, Tfg.p0])  TTfgh.p1) 
                       (𝖺[f, g, tab0 g  Tfg.p0]  TTfgh.p1) =
                      𝖺-1[f, g, (tab0 g  Tfg.p0)  TTfgh.p1] 
                       (f  𝖺[g, tab0 g  Tfg.p0, TTfgh.p1]) 
                       (f  𝖺[g, tab0 g, Tfg.p0]  TTfgh.p1) 
                        𝖺[f, (g  tab0 g)  Tfg.p0, TTfgh.p1] 
                       ((f  𝖺-1[g, tab0 g, Tfg.p0])  TTfgh.p1) 
                       (𝖺[f, g, tab0 g  Tfg.p0]  TTfgh.p1)"
              using fg gh fg0h1.p0_simps fg0h1.p1_simps f0g1.p0_simps f0g1.p1_simps
                    𝖺'_def α_def
              by simp
            also have "... = 𝖺[f  g, tab0 g  Tfg.p0, TTfgh.p1]"
              using fg gh fg0h1.p0_simps fg0h1.p1_simps f0g1.p0_simps f0g1.p1_simps
              by (intro E.eval_eqI, auto)
            also have "... = 𝖺[f  g, tab0 g  Tfg.p0, TTfgh.p1]"
              using fg gh fg0h1.p0_simps fg0h1.p1_simps f0g1.p0_simps f0g1.p1_simps
                    𝖺'_def α_def
              by simp
            finally show ?thesis
              using comp_assoc by presburger
          qed
          also have "... =
                       𝖺[f, g, h  tab0 h  TTfgh.p0] 
                       ((𝖺[f  g, h, tab0 h  TTfgh.p0] 
                       𝖺-1[f  g, h, tab0 h  TTfgh.p0]) 
                       ((f  g)  𝖺[h, tab0 h, TTfgh.p0])) 
                       ((f  g)  h.tab  TTfgh.p0) 
                       ((f  g)  fg0h1) 
                       𝖺[f  g, tab0 g  Tfg.p0, TTfgh.p1] 
                       (𝖺-1[f, g, tab0 g  Tfg.p0] 
                         (f  𝖺[g, tab0 g, Tfg.p0]) 
                         (f  (g.tab  Tfg.p0)) 
                         (f  f0g1) 
                         𝖺[f, tab0 f, Tfg.p1] 
                         (f.tab  Tfg.p1)
                         TTfgh.p1) 
                       𝖺-1[tab1 f, Tfg.p1, TTfgh.p1]"
          proof -
            have "(𝖺[f  g, h, tab0 h  TTfgh.p0] 
                       𝖺-1[f  g, h, tab0 h  TTfgh.p0]) 
                       ((f  g)  𝖺[h, tab0 h, TTfgh.p0]) =
                    ((f  g)  𝖺[h, tab0 h, TTfgh.p0])"
              using fg gh fg0h1.p0_simps fg0h1.p1_simps f0g1.p0_simps f0g1.p1_simps
                    comp_cod_arr comp_assoc_assoc'
              by simp
            thus ?thesis by simp
          qed
          also have "... =
                       𝖺[f, g, h  tab0 h  TTfgh.p0] 
                       𝖺[f  g, h, tab0 h  TTfgh.p0] 
                       (𝖺-1[f  g, h, tab0 h  TTfgh.p0] 
                       ((f  g)  𝖺[h, tab0 h, TTfgh.p0]) 
                       ((f  g)  h.tab  TTfgh.p0) 
                       ((f  g)  fg0h1) 
                       𝖺[f  g, tab0 g  Tfg.p0, TTfgh.p1] 
                       (𝖺-1[f, g, tab0 g  Tfg.p0] 
                         (f  𝖺[g, tab0 g, Tfg.p0]) 
                         (f  (g.tab  Tfg.p0)) 
                         (f  f0g1) 
                         𝖺[f, tab0 f, Tfg.p1] 
                         (f.tab  Tfg.p1)
                         TTfgh.p1)) 
                       𝖺-1[tab1 f, Tfg.p1, TTfgh.p1]"
            using comp_assoc by presburger
          also have "... =
                       𝖺[f, g, h  tab0 h  TTfgh.p0] 
                       𝖺[f  g, h, tab0 h  TTfgh.p0] 
                       TTfgh.tab 
                       𝖺-1[tab1 f, Tfg.p1, TTfgh.p1]"
            using TTfgh.tab_def Tfg.ρσ.tab_def by simp
          also have "... =
                       𝖺[f, g, h  tab0 h  TTfgh.p0] 
                         𝖺[f  g, h, tab0 h  TTfgh.p0] 
                         ((𝖺-1[f, g, h]  tab0 h  TTfgh.p0) 
                         ((f  g  h)  TTfgh_TfTgh.the_θ) 
                         𝖺[f  g  h, (tab0 h  Tgh.p0)  TfTgh.p0, TTfgh_TfTgh.chine] 
                         (𝖺-1[f, g  h, (tab0 h  Tgh.p0)  TfTgh.p0] 
                           (f  𝖺[g  h, tab0 h  Tgh.p0, TfTgh.p0]) 
                           (f  𝖺-1[g, h, tab0 h  Tgh.p0] 
                                 (g  𝖺[h, tab0 h, Tgh.p0]) 
                                 (g  h.tab  Tgh.p0) 
                                 (g  g0h1) 
                                 𝖺[g, tab0 g, Tgh.p1] 
                                 (g.tab  Tgh.p1)
                               TfTgh.p0) 
                           (f  f0gh1) 
                           𝖺[f, tab0 f, TfTgh.p1] 
                           (f.tab  TfTgh.p1)
                           TTfgh_TfTgh.chine) 
                         TTfgh_TfTgh.the_ν) 
                         𝖺-1[tab1 f, Tfg.p1, TTfgh.p1]"
            using src_tab_eq TfTgh.tab_def Tgh.ρσ.tab_def comp_assoc by simp
          text ‹Now we have to make this look like f.composite_cell wf' θf' ⋅ βf.›
          also have "... =
                       𝖺[f, g, h  tab0 h  TTfgh.p0] 
                         𝖺[f  g, h, tab0 h  TTfgh.p0] 
                         ((𝖺-1[f, g, h]  tab0 h  TTfgh.p0) 
                         ((f  g  h)  TTfgh_TfTgh.the_θ) 
                         𝖺[f  g  h, (tab0 h  Tgh.p0)  TfTgh.p0, TTfgh_TfTgh.chine] 
                         (𝖺-1[f, g  h, (tab0 h  Tgh.p0)  TfTgh.p0] 
                           (f  𝖺[g  h, tab0 h  Tgh.p0, TfTgh.p0]) 
                           ((f  𝖺-1[g, h, tab0 h  Tgh.p0]  TfTgh.p0) 
                             (f  (g  𝖺[h, tab0 h, Tgh.p0])  TfTgh.p0) 
                             (f  (g  h.tab  Tgh.p0)  TfTgh.p0) 
                             (f  (g  g0h1)  TfTgh.p0) 
                             (f  𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0) 
                             (f  (g.tab  Tgh.p1)  TfTgh.p0) 
                             (f  f0gh1)) 
                             𝖺[f, tab0 f, TfTgh.p1] 
                             (f.tab  TfTgh.p1)
                             TTfgh_TfTgh.chine) 
                         TTfgh_TfTgh.the_ν) 
                         𝖺-1[tab1 f, Tfg.p1, TTfgh.p1]"
          proof -
            have "f  𝖺-1[g, h, tab0 h  Tgh.p0] 
                          (g  𝖺[h, tab0 h, Tgh.p0]) 
                          (g  h.tab  Tgh.p0) 
                          (g  g0h1) 
                          𝖺[g, tab0 g, Tgh.p1] 
                          (g.tab  Tgh.p1)
                       TfTgh.p0 =
                    (f  𝖺-1[g, h, tab0 h  Tgh.p0]  TfTgh.p0) 
                      (f  (g  𝖺[h, tab0 h, Tgh.p0])  TfTgh.p0) 
                      (f  (g  h.tab  Tgh.p0)  TfTgh.p0) 
                      (f  (g  g0h1)  TfTgh.p0) 
                      (f  𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0) 
                      (f  (g.tab  Tgh.p1)  TfTgh.p0)"
              using fg gh whisker_right whisker_left by simp
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... =
                       𝖺[f, g, h  tab0 h  TTfgh.p0] 
                       𝖺[f  g, h, tab0 h  TTfgh.p0] 
                       (𝖺-1[f, g, h]  tab0 h  TTfgh.p0) 
                       ((f  g  h)  TTfgh_TfTgh.the_θ) 
                       𝖺[f  g  h, (tab0 h  Tgh.p0)  TfTgh.p0, TTfgh_TfTgh.chine] 
                       (𝖺-1[f, g  h, (tab0 h  Tgh.p0)  TfTgh.p0]  TTfgh_TfTgh.chine) 
                       ((f  𝖺[g  h, tab0 h  Tgh.p0, TfTgh.p0])  TTfgh_TfTgh.chine) 
                       ((f  𝖺-1[g, h, tab0 h  Tgh.p0]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  (g  𝖺[h, tab0 h, Tgh.p0])  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  (g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  (g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  (g.tab  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  f0gh1)  TTfgh_TfTgh.chine) 
                       (𝖺[f, tab0 f, TfTgh.p1]  TTfgh_TfTgh.chine) 
                       ((f.tab  TfTgh.p1)  TTfgh_TfTgh.chine) 
                       TTfgh_TfTgh.the_ν 
                       𝖺-1[tab1 f, Tfg.p1, TTfgh.p1]"
          proof -
            have "𝖺-1[f, g  h, (tab0 h  Tgh.p0)  TfTgh.p0] 
                      (f  𝖺[g  h, tab0 h  Tgh.p0, TfTgh.p0]) 
                      (f  𝖺-1[g, h, tab0 h  Tgh.p0]  TfTgh.p0) 
                      (f  (g  𝖺[h, tab0 h, Tgh.p0])  TfTgh.p0) 
                      (f  (g  h.tab  Tgh.p0)  TfTgh.p0) 
                      (f  (g  g0h1)  TfTgh.p0) 
                      (f  𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0) 
                      (f  (g.tab  Tgh.p1)  TfTgh.p0) 
                      (f  f0gh1) 
                      𝖺[f, tab0 f, TfTgh.p1] 
                      (f.tab  TfTgh.p1)
                      TTfgh_TfTgh.chine =
                    (𝖺-1[f, g  h, (tab0 h  Tgh.p0)  TfTgh.p0]  TTfgh_TfTgh.chine) 
                    ((f  𝖺[g  h, tab0 h  Tgh.p0, TfTgh.p0])  TTfgh_TfTgh.chine) 
                    ((f  𝖺-1[g, h, tab0 h  Tgh.p0]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                    ((f  (g  𝖺[h, tab0 h, Tgh.p0])  TfTgh.p0)  TTfgh_TfTgh.chine) 
                    ((f  (g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                    ((f  (g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                    ((f  𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                    ((f  (g.tab  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                    ((f  f0gh1)  TTfgh_TfTgh.chine) 
                    (𝖺[f, tab0 f, TfTgh.p1]  TTfgh_TfTgh.chine) 
                    ((f.tab  TfTgh.p1)  TTfgh_TfTgh.chine)"
              (* using fg gh whisker_right [of TTfgh_TfTgh.chine] by auto (* 2 min *) *)
            proof -
              have "arr (𝖺-1[f, g  h, (tab0 h  Tgh.p0)  TfTgh.p0] 
                      (f  𝖺[g  h, tab0 h  Tgh.p0, TfTgh.p0]) 
                      (f  𝖺-1[g, h, tab0 h  Tgh.p0]  TfTgh.p0) 
                      (f  (g  𝖺[h, tab0 h, Tgh.p0])  TfTgh.p0) 
                      (f  (g  h.tab  Tgh.p0)  TfTgh.p0) 
                      (f  (g  g0h1)  TfTgh.p0) 
                      (f  𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0) 
                      (f  (g.tab  Tgh.p1)  TfTgh.p0) 
                      (f  f0gh1) 
                      𝖺[f, tab0 f, TfTgh.p1] 
                      (f.tab  TfTgh.p1))"
                using fg gh
                by (intro seqI' comp_in_homI) auto
              moreover
              have "arr ((f  𝖺[g  h, tab0 h  Tgh.p0, TfTgh.p0]) 
                      (f  𝖺-1[g, h, tab0 h  Tgh.p0]  TfTgh.p0) 
                      (f  (g  𝖺[h, tab0 h, Tgh.p0])  TfTgh.p0) 
                      (f  (g  h.tab  Tgh.p0)  TfTgh.p0) 
                      (f  (g  g0h1)  TfTgh.p0) 
                      (f  𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0) 
                      (f  (g.tab  Tgh.p1)  TfTgh.p0) 
                      (f  f0gh1) 
                      𝖺[f, tab0 f, TfTgh.p1] 
                      (f.tab  TfTgh.p1))"
                using calculation by blast
              moreover
              have "arr ((f  𝖺-1[g, h, tab0 h  Tgh.p0]  TfTgh.p0) 
                      (f  (g  𝖺[h, tab0 h, Tgh.p0])  TfTgh.p0) 
                      (f  (g  h.tab  Tgh.p0)  TfTgh.p0) 
                      (f  (g  g0h1)  TfTgh.p0) 
                      (f  𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0) 
                      (f  (g.tab  Tgh.p1)  TfTgh.p0) 
                      (f  f0gh1) 
                      𝖺[f, tab0 f, TfTgh.p1] 
                      (f.tab  TfTgh.p1))"
                using calculation by blast
              moreover
              have "arr ((f  (g  𝖺[h, tab0 h, Tgh.p0])  TfTgh.p0) 
                      (f  (g  h.tab  Tgh.p0)  TfTgh.p0) 
                      (f  (g  g0h1)  TfTgh.p0) 
                      (f  𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0) 
                      (f  (g.tab  Tgh.p1)  TfTgh.p0) 
                      (f  f0gh1) 
                      𝖺[f, tab0 f, TfTgh.p1] 
                      (f.tab  TfTgh.p1))"
                using calculation by blast
              moreover
              have "arr ((f  (g  h.tab  Tgh.p0)  TfTgh.p0) 
                      (f  (g  g0h1)  TfTgh.p0) 
                      (f  𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0) 
                      (f  (g.tab  Tgh.p1)  TfTgh.p0) 
                      (f  f0gh1) 
                      𝖺[f, tab0 f, TfTgh.p1] 
                      (f.tab  TfTgh.p1))"
                using calculation by blast
              moreover
              have "arr ((f  (g  g0h1)  TfTgh.p0) 
                      (f  𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0) 
                      (f  (g.tab  Tgh.p1)  TfTgh.p0) 
                      (f  f0gh1) 
                      𝖺[f, tab0 f, TfTgh.p1] 
                      (f.tab  TfTgh.p1))"
                using calculation by blast
              moreover
              have "arr ((f  𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0) 
                      (f  (g.tab  Tgh.p1)  TfTgh.p0) 
                      (f  f0gh1) 
                      𝖺[f, tab0 f, TfTgh.p1] 
                      (f.tab  TfTgh.p1))"
                using calculation by blast
              moreover
              have "arr ((f  (g.tab  Tgh.p1)  TfTgh.p0) 
                      (f  f0gh1) 
                      𝖺[f, tab0 f, TfTgh.p1] 
                      (f.tab  TfTgh.p1))"
                using calculation by blast
              moreover
              have "arr ((f  f0gh1) 
                      𝖺[f, tab0 f, TfTgh.p1] 
                      (f.tab  TfTgh.p1))"
                using calculation by blast
              moreover
              have "arr (𝖺[f, tab0 f, TfTgh.p1] 
                      (f.tab  TfTgh.p1))"
                using calculation by blast
              ultimately show ?thesis
                using whisker_right [of TTfgh_TfTgh.chine] TTfgh_TfTgh.is_ide by presburger
            qed
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... =
                       𝖺[f, g, h  tab0 h  TTfgh.p0] 
                       𝖺[f  g, h, tab0 h  TTfgh.p0] 
                       (𝖺-1[f, g, h]  tab0 h  TTfgh.p0) 
                       ((f  g  h)  TTfgh_TfTgh.the_θ) 
                       𝖺[f  g  h, (tab0 h  Tgh.p0)  TfTgh.p0, TTfgh_TfTgh.chine] 
                       (𝖺-1[f, g  h, (tab0 h  Tgh.p0)  TfTgh.p0]  TTfgh_TfTgh.chine) 
                       ((f  𝖺[g  h, tab0 h  Tgh.p0, TfTgh.p0])  TTfgh_TfTgh.chine) 
                       ((f  𝖺-1[g, h, tab0 h  Tgh.p0]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  (g  𝖺[h, tab0 h, Tgh.p0])  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  (g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  (g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  (g.tab  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  f0gh1)  TTfgh_TfTgh.chine) 
                       (𝖺[f, tab0 f, TfTgh.p1]  TTfgh_TfTgh.chine) 
                       (((f.tab  TfTgh.p1)  TTfgh_TfTgh.chine) 
                       𝖺-1[tab1 f, TfTgh.p1, TTfgh_TfTgh.chine] 
                       𝖺[tab1 f, TfTgh.p1, TTfgh_TfTgh.chine]) 
                       TTfgh_TfTgh.the_ν 
                       𝖺-1[tab1 f, Tfg.p1, TTfgh.p1]"
         proof -
           have "((f.tab  TfTgh.p1)  TTfgh_TfTgh.chine) 
                       𝖺-1[tab1 f, TfTgh.p1, TTfgh_TfTgh.chine] 
                       𝖺[tab1 f, TfTgh.p1, TTfgh_TfTgh.chine] =
                    (f.tab  TfTgh.p1)  TTfgh_TfTgh.chine"
             using fg gh fg0h1.p0_simps fg0h1.p1_simps f0g1.p0_simps f0g1.p1_simps
                   comp_arr_dom comp_assoc_assoc'
             by simp
           thus ?thesis by simp
         qed
         also have "... =
                       𝖺[f, g, h  tab0 h  TTfgh.p0] 
                       𝖺[f  g, h, tab0 h  TTfgh.p0] 
                       (𝖺-1[f, g, h]  tab0 h  TTfgh.p0) 
                       ((f  g  h)  TTfgh_TfTgh.the_θ) 
                       𝖺[f  g  h, (tab0 h  Tgh.p0)  TfTgh.p0, TTfgh_TfTgh.chine] 
                       (𝖺-1[f, g  h, (tab0 h  Tgh.p0)  TfTgh.p0]  TTfgh_TfTgh.chine) 
                       ((f  𝖺[g  h, tab0 h  Tgh.p0, TfTgh.p0])  TTfgh_TfTgh.chine) 
                       ((f  𝖺-1[g, h, tab0 h  Tgh.p0]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  (g  𝖺[h, tab0 h, Tgh.p0])  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  (g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  (g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  (g.tab  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  f0gh1)  TTfgh_TfTgh.chine) 
                       (𝖺[f, tab0 f, TfTgh.p1]  TTfgh_TfTgh.chine) 
                       (((f.tab  TfTgh.p1)  TTfgh_TfTgh.chine) 
                       𝖺-1[tab1 f, TfTgh.p1, TTfgh_TfTgh.chine]) 
                       𝖺[tab1 f, TfTgh.p1, TTfgh_TfTgh.chine] 
                       TTfgh_TfTgh.the_ν 
                       𝖺-1[tab1 f, Tfg.p1, TTfgh.p1]"
            using comp_assoc by presburger
          also have "... =
                       𝖺[f, g, h  tab0 h  TTfgh.p0] 
                       𝖺[f  g, h, tab0 h  TTfgh.p0] 
                       (𝖺-1[f, g, h]  tab0 h  TTfgh.p0) 
                       ((f  g  h)  TTfgh_TfTgh.the_θ) 
                       𝖺[f  g  h, (tab0 h  Tgh.p0)  TfTgh.p0, TTfgh_TfTgh.chine] 
                       (𝖺-1[f, g  h, (tab0 h  Tgh.p0)  TfTgh.p0]  TTfgh_TfTgh.chine) 
                       ((f  𝖺[g  h, tab0 h  Tgh.p0, TfTgh.p0])  TTfgh_TfTgh.chine) 
                       ((f  𝖺-1[g, h, tab0 h  Tgh.p0]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  (g  𝖺[h, tab0 h, Tgh.p0])  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  (g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  (g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  (g.tab  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  f0gh1)  TTfgh_TfTgh.chine) 
                       (𝖺[f, tab0 f, TfTgh.p1]  TTfgh_TfTgh.chine) 
                       𝖺-1[f  tab0 f, TfTgh.p1, TTfgh_TfTgh.chine] 
                       (f.tab  TfTgh.p1  TTfgh_TfTgh.chine) 
                       𝖺[tab1 f, TfTgh.p1, TTfgh_TfTgh.chine] 
                       TTfgh_TfTgh.the_ν 
                       𝖺-1[tab1 f, Tfg.p1, TTfgh.p1]"
          proof -
            have "((f.tab  TfTgh.p1)  TTfgh_TfTgh.chine) 
                      𝖺-1[tab1 f, TfTgh.p1, TTfgh_TfTgh.chine] =
                    𝖺-1[f  tab0 f, TfTgh.p1, TTfgh_TfTgh.chine] 
                      (f.tab  TfTgh.p1  TTfgh_TfTgh.chine)"
              using fg gh fg0h1.p0_simps fg0h1.p1_simps f0g1.p0_simps f0g1.p1_simps
                    assoc'_naturality [of f.tab TfTgh.p1 TTfgh_TfTgh.chine]
              by simp
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... =
                       𝖺[f, g, h  tab0 h  TTfgh.p0] 
                       𝖺[f  g, h, tab0 h  TTfgh.p0] 
                       (𝖺-1[f, g, h]  tab0 h  TTfgh.p0) 
                       ((f  g  h)  TTfgh_TfTgh.the_θ) 
                       𝖺[f  g  h, (tab0 h  Tgh.p0)  TfTgh.p0, TTfgh_TfTgh.chine] 
                       (𝖺-1[f, g  h, (tab0 h  Tgh.p0)  TfTgh.p0]  TTfgh_TfTgh.chine) 
                       ((f  𝖺[g  h, tab0 h  Tgh.p0, TfTgh.p0])  TTfgh_TfTgh.chine) 
                       ((f  𝖺-1[g, h, tab0 h  Tgh.p0]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  (g  𝖺[h, tab0 h, Tgh.p0])  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  (g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  (g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  (g.tab  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  f0gh1)  TTfgh_TfTgh.chine) 
                       (𝖺[f, tab0 f, TfTgh.p1]  TTfgh_TfTgh.chine) 
                       𝖺-1[f  tab0 f, TfTgh.p1, TTfgh_TfTgh.chine] 
                       ((𝖺-1[f, tab0 f, TfTgh.p1  TTfgh_TfTgh.chine] 
                       𝖺[f, tab0 f, TfTgh.p1  TTfgh_TfTgh.chine]) 
                       (f.tab  TfTgh.p1  TTfgh_TfTgh.chine)) 
                       𝖺[tab1 f, TfTgh.p1, TTfgh_TfTgh.chine] 
                       TTfgh_TfTgh.the_ν 
                       𝖺-1[tab1 f, Tfg.p1, TTfgh.p1]"
          proof -
            have "(𝖺-1[f, tab0 f, TfTgh.p1  TTfgh_TfTgh.chine] 
                      𝖺[f, tab0 f, TfTgh.p1  TTfgh_TfTgh.chine]) 
                      (f.tab  TfTgh.p1  TTfgh_TfTgh.chine) =
                    f.tab  TfTgh.p1  TTfgh_TfTgh.chine"
              using fg gh fg0h1.p0_simps fg0h1.p1_simps f0g1.p0_simps f0g1.p1_simps
                    comp_cod_arr comp_assoc_assoc'
              by simp
            thus ?thesis
              using comp_assoc by simp
          qed
          also have "... =
                       𝖺[f, g, h  tab0 h  TTfgh.p0] 
                       𝖺[f  g, h, tab0 h  TTfgh.p0] 
                       (𝖺-1[f, g, h]  tab0 h  TTfgh.p0) 
                       ((f  g  h)  TTfgh_TfTgh.the_θ) 
                       𝖺[f  g  h, (tab0 h  Tgh.p0)  TfTgh.p0, TTfgh_TfTgh.chine] 
                       (𝖺-1[f, g  h, (tab0 h  Tgh.p0)  TfTgh.p0]  TTfgh_TfTgh.chine) 
                       ((f  𝖺[g  h, tab0 h  Tgh.p0, TfTgh.p0])  TTfgh_TfTgh.chine) 
                       ((f  𝖺-1[g, h, tab0 h  Tgh.p0]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  (g  𝖺[h, tab0 h, Tgh.p0])  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  (g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  (g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  (g.tab  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  f0gh1)  TTfgh_TfTgh.chine) 
                       ((𝖺[f, tab0 f, TfTgh.p1]  TTfgh_TfTgh.chine) 
                       𝖺-1[f  tab0 f, TfTgh.p1, TTfgh_TfTgh.chine] 
                       𝖺-1[f, tab0 f, TfTgh.p1  TTfgh_TfTgh.chine]) 
                       𝖺[f, tab0 f, TfTgh.p1  TTfgh_TfTgh.chine] 
                       (f.tab  TfTgh.p1  TTfgh_TfTgh.chine) 
                       𝖺[tab1 f, TfTgh.p1, TTfgh_TfTgh.chine] 
                       TTfgh_TfTgh.the_ν 
                       𝖺-1[tab1 f, Tfg.p1, TTfgh.p1]"
            using comp_assoc by presburger
          also have "... =
                       𝖺[f, g, h  tab0 h  TTfgh.p0] 
                       𝖺[f  g, h, tab0 h  TTfgh.p0] 
                       (𝖺-1[f, g, h]  tab0 h  TTfgh.p0) 
                       ((f  g  h)  TTfgh_TfTgh.the_θ) 
                       𝖺[f  g  h, (tab0 h  Tgh.p0)  TfTgh.p0, TTfgh_TfTgh.chine] 
                       (𝖺-1[f, g  h, (tab0 h  Tgh.p0)  TfTgh.p0]  TTfgh_TfTgh.chine) 
                       (((f  𝖺[g  h, tab0 h  Tgh.p0, TfTgh.p0])  TTfgh_TfTgh.chine) 
                       ((f  𝖺-1[g, h, tab0 h  Tgh.p0]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  (g  𝖺[h, tab0 h, Tgh.p0])  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  (g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  (g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  (g.tab  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  f0gh1)  TTfgh_TfTgh.chine) 
                       𝖺-1[f, tab0 f  TfTgh.p1, TTfgh_TfTgh.chine]) 
                       (f  𝖺-1[tab0 f, TfTgh.p1, TTfgh_TfTgh.chine]) 
                       𝖺[f, tab0 f, TfTgh.p1  TTfgh_TfTgh.chine] 
                       (f.tab  TfTgh.p1  TTfgh_TfTgh.chine) 
                       𝖺[tab1 f, TfTgh.p1, TTfgh_TfTgh.chine] 
                       TTfgh_TfTgh.the_ν 
                       𝖺-1[tab1 f, Tfg.p1, TTfgh.p1]"
          proof -
            have "(𝖺[f, tab0 f, TfTgh.p1]  TTfgh_TfTgh.chine) 
                       𝖺-1[f  tab0 f, TfTgh.p1, TTfgh_TfTgh.chine] 
                       𝖺-1[f, tab0 f, TfTgh.p1  TTfgh_TfTgh.chine] =
                    𝖺-1[f, tab0 f  TfTgh.p1, TTfgh_TfTgh.chine] 
                      (f  𝖺-1[tab0 f, TfTgh.p1, TTfgh_TfTgh.chine])"
            proof -
              have "(𝖺[f, tab0 f, TfTgh.p1]  TTfgh_TfTgh.chine) 
                        𝖺-1[f  tab0 f, TfTgh.p1, TTfgh_TfTgh.chine] 
                        𝖺-1[f, tab0 f, TfTgh.p1  TTfgh_TfTgh.chine] =
                      (𝖺[f, tab0 f, TfTgh.p1]  TTfgh_TfTgh.chine) 
                         𝖺-1[f  tab0 f, TfTgh.p1, TTfgh_TfTgh.chine] 
                         𝖺-1[f, tab0 f, TfTgh.p1  TTfgh_TfTgh.chine]"
                using 𝖺'_def α_def by simp
              also have "... = 𝖺-1[f, tab0 f  TfTgh.p1, TTfgh_TfTgh.chine] 
                                  (f  𝖺-1[tab0 f, TfTgh.p1, TTfgh_TfTgh.chine])"
                using fg gh fg0h1.p0_simps fg0h1.p1_simps f0g1.p0_simps f0g1.p1_simps
                by (intro E.eval_eqI, auto)
              also have "... = 𝖺-1[f, tab0 f  TfTgh.p1, TTfgh_TfTgh.chine] 
                                   (f  𝖺-1[tab0 f, TfTgh.p1, TTfgh_TfTgh.chine])"
                using 𝖺'_def α_def by simp
              finally show ?thesis by simp
            qed
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... =
                       𝖺[f, g, h  tab0 h  TTfgh.p0] 
                       𝖺[f  g, h, tab0 h  TTfgh.p0] 
                       (𝖺-1[f, g, h]  tab0 h  TTfgh.p0) 
                       ((f  g  h)  TTfgh_TfTgh.the_θ) 
                       𝖺[f  g  h, (tab0 h  Tgh.p0)  TfTgh.p0, TTfgh_TfTgh.chine] 
                       (𝖺-1[f, g  h, (tab0 h  Tgh.p0)  TfTgh.p0]  TTfgh_TfTgh.chine) 
                       𝖺-1[f, (g  h)  (tab0 h  Tgh.p0)  TfTgh.p0, TTfgh_TfTgh.chine] 
                       (f  𝖺[g  h, tab0 h  Tgh.p0, TfTgh.p0]  TTfgh_TfTgh.chine) 
                       (f  (𝖺-1[g, h, tab0 h  Tgh.p0]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       (f  ((g  𝖺[h, tab0 h, Tgh.p0])  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  ((g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       (f  ((g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       (f  (𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       (f  ((g.tab  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       (f  f0gh1  TTfgh_TfTgh.chine) 
                       (f  𝖺-1[tab0 f, TfTgh.p1, TTfgh_TfTgh.chine])) 
                       𝖺[f, tab0 f, TfTgh.p1  TTfgh_TfTgh.chine] 
                       (f.tab  TfTgh.p1  TTfgh_TfTgh.chine) 
                       𝖺[tab1 f, TfTgh.p1, TTfgh_TfTgh.chine] 
                       TTfgh_TfTgh.the_ν 
                       𝖺-1[tab1 f, Tfg.p1, TTfgh.p1]"
          proof -
            have "((f  𝖺[g  h, tab0 h  Tgh.p0, TfTgh.p0])  TTfgh_TfTgh.chine) 
                       ((f  𝖺-1[g, h, tab0 h  Tgh.p0]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  (g  𝖺[h, tab0 h, Tgh.p0])  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  (g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  (g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  (g.tab  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       ((f  f0gh1)  TTfgh_TfTgh.chine) 
                       𝖺-1[f, tab0 f  TfTgh.p1, TTfgh_TfTgh.chine] =
                    𝖺-1[f, (g  h)  (tab0 h  Tgh.p0)  TfTgh.p0, TTfgh_TfTgh.chine] 
                       (f  𝖺[g  h, tab0 h  Tgh.p0, TfTgh.p0]  TTfgh_TfTgh.chine) 
                       (f  (𝖺-1[g, h, tab0 h  Tgh.p0]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       (f  ((g  𝖺[h, tab0 h, Tgh.p0])  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       (f  ((g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       (f  ((g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       (f  (𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       (f  ((g.tab  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       (f  f0gh1  TTfgh_TfTgh.chine)"
            proof -
              (*
               * This one can't be shortcut with a straight coherence-based proof,
               * due to the presence of f0gh1.φ, g0h1.φ, h.tab, with associativities that
               * do not respect their domain and codomain.
               *
               * I also tried to avoid distributing the "f ⋆" in advance, in order to
               * reduce the number of associativity proof steps, but it then becomes
               * less automatic to prove the necessary "arr" facts to do the proof.
               * So unfortunately the mindless grind seems to be the path of least
               * resistance.
               *)
              have "((f  𝖺[g  h, tab0 h  Tgh.p0, TfTgh.p0])  TTfgh_TfTgh.chine) 
                        ((f  𝖺-1[g, h, tab0 h  Tgh.p0]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                        ((f  (g  𝖺[h, tab0 h, Tgh.p0])  TfTgh.p0)  TTfgh_TfTgh.chine) 
                        ((f  (g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                        ((f  (g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                        ((f  𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                        ((f  (g.tab  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                        ((f  f0gh1)  TTfgh_TfTgh.chine) 
                        𝖺-1[f, tab0 f  TfTgh.p1, TTfgh_TfTgh.chine] =
                      ((f  𝖺[g  h, tab0 h  Tgh.p0, TfTgh.p0])  TTfgh_TfTgh.chine) 
                        ((f  𝖺-1[g, h, tab0 h  Tgh.p0]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                        ((f  (g  𝖺[h, tab0 h, Tgh.p0])  TfTgh.p0)  TTfgh_TfTgh.chine) 
                        ((f  (g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                        ((f  (g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                        ((f  𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                        (((f  (g.tab  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                        𝖺-1[f, (tab1 g  Tgh.p1)  TfTgh.p0, TTfgh_TfTgh.chine]) 
                        (f  f0gh1  TTfgh_TfTgh.chine)"
                using fg gh fg0h1.p0_simps fg0h1.p1_simps f0g1.p0_simps f0g1.p1_simps
                      assoc'_naturality [of f f0gh1 TTfgh_TfTgh.chine] comp_assoc
                by simp
              also have "... =
                           ((f  𝖺[g  h, tab0 h  Tgh.p0, TfTgh.p0])  TTfgh_TfTgh.chine) 
                             ((f  𝖺-1[g, h, tab0 h  Tgh.p0]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                             ((f  (g  𝖺[h, tab0 h, Tgh.p0])  TfTgh.p0)  TTfgh_TfTgh.chine) 
                             ((f  (g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                             ((f  (g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                             (((f  𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                             𝖺-1[f, ((g  tab0 g)  Tgh.p1)  TfTgh.p0, TTfgh_TfTgh.chine]) 
                             (f  ((g.tab  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                             (f  f0gh1  TTfgh_TfTgh.chine)"
              proof -
                have "((f  (g.tab  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                          𝖺-1[f, (tab1 g  Tgh.p1)  TfTgh.p0, TTfgh_TfTgh.chine] =
                        𝖺-1[f, ((g  tab0 g)  Tgh.p1)  TfTgh.p0, TTfgh_TfTgh.chine] 
                          (f  ((g.tab  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine)"
                  using fg gh fg0h1.p0_simps fg0h1.p1_simps f0g1.p0_simps f0g1.p1_simps
                        assoc'_naturality [of f "(g.tab  Tgh.p1)  TfTgh.p0" TTfgh_TfTgh.chine]
                  by simp
                thus ?thesis
                  using comp_assoc by presburger
              qed
              also have "... =
                           ((f  𝖺[g  h, tab0 h  Tgh.p0, TfTgh.p0])  TTfgh_TfTgh.chine) 
                             ((f  𝖺-1[g, h, tab0 h  Tgh.p0]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                             ((f  (g  𝖺[h, tab0 h, Tgh.p0])  TfTgh.p0)  TTfgh_TfTgh.chine) 
                             ((f  (g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                             (((f  (g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                             𝖺-1[f, (g  tab0 g  Tgh.p1)  TfTgh.p0, TTfgh_TfTgh.chine]) 
                             (f  (𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                             (f  ((g.tab  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                             (f  f0gh1  TTfgh_TfTgh.chine)"
              proof -
                have "((f  𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                          𝖺-1[f, ((g  tab0 g)  Tgh.p1)  TfTgh.p0, TTfgh_TfTgh.chine] =
                        𝖺-1[f, (g  tab0 g  Tgh.p1)  TfTgh.p0, TTfgh_TfTgh.chine] 
                          (f  (𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine)"
                  using fg gh fg0h1.p0_simps fg0h1.p1_simps f0g1.p0_simps f0g1.p1_simps
                          assoc'_naturality
                            [of f "𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0" TTfgh_TfTgh.chine]
                  by simp
                thus ?thesis
                  using comp_assoc by presburger
              qed
              also have "... =
                           ((f  𝖺[g  h, tab0 h  Tgh.p0, TfTgh.p0])  TTfgh_TfTgh.chine) 
                             ((f  𝖺-1[g, h, tab0 h  Tgh.p0]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                             ((f  (g  𝖺[h, tab0 h, Tgh.p0])  TfTgh.p0)  TTfgh_TfTgh.chine) 
                             (((f  (g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                             𝖺-1[f, (g  tab1 h  Tgh.p0)  TfTgh.p0, TTfgh_TfTgh.chine]) 
                             (f  ((g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                             (f  (𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                             (f  ((g.tab  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                             (f  f0gh1  TTfgh_TfTgh.chine)"
              proof -
                have "((f  (g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                          𝖺-1[f, (g  tab0 g  Tgh.p1)  TfTgh.p0, TTfgh_TfTgh.chine] =
                        𝖺-1[f, (g  tab1 h  Tgh.p0)  TfTgh.p0, TTfgh_TfTgh.chine] 
                          (f  ((g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine)"
                  using fg gh fg0h1.p0_simps fg0h1.p1_simps f0g1.p0_simps f0g1.p1_simps
                        assoc'_naturality [of f "(g  g0h1)  TfTgh.p0" TTfgh_TfTgh.chine]
                  by simp
                thus ?thesis
                  using comp_assoc by presburger
              qed
              also have "... =
                           (((f  𝖺[g  h, tab0 h  Tgh.p0, TfTgh.p0])  TTfgh_TfTgh.chine) 
                             ((f  𝖺-1[g, h, tab0 h  Tgh.p0]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                             ((f  (g  𝖺[h, tab0 h, Tgh.p0])  TfTgh.p0)  TTfgh_TfTgh.chine) 
                             𝖺-1[f, (g  (h  tab0 h)  Tgh.p0)  TfTgh.p0, TTfgh_TfTgh.chine]) 
                             (f  ((g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                             (f  ((g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                             (f  (𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                             (f  ((g.tab  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                             (f  f0gh1  TTfgh_TfTgh.chine)"
              proof -
                have "((f  (g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                          𝖺-1[f, (g  tab1 h  Tgh.p0)  TfTgh.p0, TTfgh_TfTgh.chine] =
                        𝖺-1[f, (g  (h  tab0 h)  Tgh.p0)  TfTgh.p0, TTfgh_TfTgh.chine] 
                          (f  ((g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine)"
                  using fg gh fg0h1.p0_simps fg0h1.p1_simps f0g1.p0_simps f0g1.p1_simps
                        assoc'_naturality
                          [of f "(g  h.tab  Tgh.p0)  TfTgh.p0" TTfgh_TfTgh.chine]
                  by simp
                thus ?thesis
                  using comp_assoc by presburger
              qed
              also have "... =
                           𝖺-1[f, (g  h)  (tab0 h  Tgh.p0)  TfTgh.p0, TTfgh_TfTgh.chine] 
                             (f  𝖺[g  h, tab0 h  Tgh.p0, TfTgh.p0]  TTfgh_TfTgh.chine) 
                             (f  (𝖺-1[g, h, tab0 h  Tgh.p0]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                             (f  ((g  𝖺[h, tab0 h, Tgh.p0])  TfTgh.p0)  TTfgh_TfTgh.chine) 
                             (f  ((g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                             (f  ((g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                             (f  (𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                             (f  ((g.tab  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                             (f  f0gh1  TTfgh_TfTgh.chine)"
              proof -
                (* OK, we can perhaps shortcut the last few steps... *)
                have "((f  𝖺[g  h, tab0 h  Tgh.p0, TfTgh.p0])  TTfgh_TfTgh.chine) 
                          ((f  𝖺-1[g, h, tab0 h  Tgh.p0]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                          ((f  (g  𝖺[h, tab0 h, Tgh.p0])  TfTgh.p0)  TTfgh_TfTgh.chine) 
                          𝖺-1[f, (g  (h  tab0 h)  Tgh.p0)  TfTgh.p0, TTfgh_TfTgh.chine] =
                        ((f  𝖺[g  h, tab0 h  Tgh.p0, TfTgh.p0])
                               TTfgh_TfTgh.chine) 
                           ((f  𝖺-1[g, h, tab0 h  Tgh.p0]  TfTgh.p0)
                                 TTfgh_TfTgh.chine) 
                           ((f  (g  𝖺[h, tab0 h, Tgh.p0])  TfTgh.p0)
                                 TTfgh_TfTgh.chine) 
                           𝖺-1[f, (g  (h  tab0 h)  Tgh.p0)  TfTgh.p0,
                               TTfgh_TfTgh.chine]"
                  using fg gh fg0h1.p0_simps fg0h1.p1_simps f0g1.p0_simps f0g1.p1_simps
                        𝖺'_def α_def
                  by simp
                also have "... =
                             𝖺-1[f, (g  h)  (tab0 h  Tgh.p0)  TfTgh.p0,
                                  TTfgh_TfTgh.chine] 
                                (f  (𝖺[g  h, tab0 h  Tgh.p0, TfTgh.p0])
                                     TTfgh_TfTgh.chine) 
                                (f  (𝖺-1[g, h, tab0 h  Tgh.p0]  TfTgh.p0)
                                      TTfgh_TfTgh.chine) 
                                (f  ((g  𝖺[h, tab0 h, Tgh.p0])  TfTgh.p0)
                                      TTfgh_TfTgh.chine)"
                    using fg gh fg0h1.p0_simps fg0h1.p1_simps f0g1.p0_simps f0g1.p1_simps
                    apply (intro E.eval_eqI) by simp_all
                  also have "... =
                             𝖺-1[f, (g  h)  (tab0 h  Tgh.p0)  TfTgh.p0, TTfgh_TfTgh.chine] 
                               (f  𝖺[g  h, tab0 h  Tgh.p0, TfTgh.p0]  TTfgh_TfTgh.chine) 
                               (f  (𝖺-1[g, h, tab0 h  Tgh.p0]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                               (f  ((g  𝖺[h, tab0 h, Tgh.p0])  TfTgh.p0)  TTfgh_TfTgh.chine)"
                  using fg gh fg0h1.p0_simps fg0h1.p1_simps f0g1.p0_simps f0g1.p1_simps
                        𝖺'_def α_def
                  by simp
                finally show ?thesis
                  using comp_assoc by presburger
              qed
              finally show ?thesis
                using comp_assoc by presburger
            qed
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... =
                       (𝖺[f, g, h  tab0 h  TTfgh.p0] 
                       𝖺[f  g, h, tab0 h  TTfgh.p0] 
                       (𝖺-1[f, g, h]  tab0 h  TTfgh.p0) 
                       ((f  g  h)  TTfgh_TfTgh.the_θ)) 
                       𝖺[f  g  h, (tab0 h  Tgh.p0)  TfTgh.p0, TTfgh_TfTgh.chine] 
                       (𝖺-1[f, g  h, (tab0 h  Tgh.p0)  TfTgh.p0]  TTfgh_TfTgh.chine) 
                       𝖺-1[f, (g  h)  (tab0 h  Tgh.p0)  TfTgh.p0, TTfgh_TfTgh.chine] 
                       (f  𝖺[g  h, tab0 h  Tgh.p0, TfTgh.p0]  TTfgh_TfTgh.chine) 
                       (f  (𝖺-1[g, h, tab0 h  Tgh.p0]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       (f  ((g  𝖺[h, tab0 h, Tgh.p0])  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       (f  (((g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                            (((g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                            ((𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                            (((g.tab  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                            (f0gh1  TTfgh_TfTgh.chine) 
                            𝖺-1[tab0 f, TfTgh.p1, TTfgh_TfTgh.chine]) 
                       𝖺[f, tab0 f, TfTgh.p1  TTfgh_TfTgh.chine] 
                       (f.tab  TfTgh.p1  TTfgh_TfTgh.chine) 
                       𝖺[tab1 f, TfTgh.p1, TTfgh_TfTgh.chine] 
                       TTfgh_TfTgh.the_ν 
                       𝖺-1[tab1 f, Tfg.p1, TTfgh.p1]"
          proof -
            have "(f  ((g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                      (f  ((g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                      (f  (𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                      (f  ((g.tab  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                      (f  f0gh1  TTfgh_TfTgh.chine) 
                      (f  𝖺-1[tab0 f, TfTgh.p1, TTfgh_TfTgh.chine]) =
                    (f  (((g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           (((g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           ((𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           (((g.tab  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           (f0gh1  TTfgh_TfTgh.chine) 
                           𝖺-1[tab0 f, TfTgh.p1, TTfgh_TfTgh.chine])"
              using fg gh whisker_left by simp  (* 15 sec elapsed, 30 sec cpu *)
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... =
                       (f  g  h  TTfgh_TfTgh.the_θ) 
                       (𝖺[f, g, h  ((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine] 
                       𝖺[f  g, h, ((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine] 
                       (𝖺-1[f, g, h]  ((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       𝖺[f  g  h, (tab0 h  Tgh.p0)  TfTgh.p0, TTfgh_TfTgh.chine] 
                       (𝖺-1[f, g  h, (tab0 h  Tgh.p0)  TfTgh.p0]  TTfgh_TfTgh.chine) 
                       𝖺-1[f, (g  h)  (tab0 h  Tgh.p0)  TfTgh.p0, TTfgh_TfTgh.chine] 
                       (f  𝖺[g  h, tab0 h  Tgh.p0, TfTgh.p0]  TTfgh_TfTgh.chine) 
                       (f  (𝖺-1[g, h, tab0 h  Tgh.p0]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       (f  ((g  𝖺[h, tab0 h, Tgh.p0])  TfTgh.p0)  TTfgh_TfTgh.chine)) 
                       (f  (((g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                            (((g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                            ((𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                            (((g.tab  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                            (f0gh1  TTfgh_TfTgh.chine) 
                            𝖺-1[tab0 f, TfTgh.p1, TTfgh_TfTgh.chine]) 
                       𝖺[f, tab0 f, TfTgh.p1  TTfgh_TfTgh.chine] 
                       (f.tab  TfTgh.p1  TTfgh_TfTgh.chine) 
                       𝖺[tab1 f, TfTgh.p1, TTfgh_TfTgh.chine] 
                       TTfgh_TfTgh.the_ν 
                       𝖺-1[tab1 f, Tfg.p1, TTfgh.p1]"
          proof -
            have "𝖺[f, g, h  tab0 h  TTfgh.p0] 
                       𝖺[f  g, h, tab0 h  TTfgh.p0] 
                       (𝖺-1[f, g, h]  tab0 h  TTfgh.p0) 
                       ((f  g  h)  TTfgh_TfTgh.the_θ) =
                    𝖺[f, g, h  tab0 h  TTfgh.p0] 
                       (𝖺[f  g, h, tab0 h  TTfgh.p0] 
                       (((f  g)  h)  TTfgh_TfTgh.the_θ)) 
                       (𝖺-1[f, g, h]  ((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine)"
            proof -
              have "(𝖺-1[f, g, h]  tab0 h  TTfgh.p0)  ((f  g  h)  TTfgh_TfTgh.the_θ) =
                      𝖺-1[f, g, h]  TTfgh_TfTgh.the_θ"
                using fg gh comp_arr_dom comp_cod_arr
                      interchange [of "𝖺-1[f, g, h]" "f  g  h"
                                      "tab0 h  TTfgh.p0" TTfgh_TfTgh.the_θ]
                by simp
              also have "... = (((f  g)  h)  TTfgh_TfTgh.the_θ) 
                               (𝖺-1[f, g, h]  ((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine)"
                using fg gh comp_arr_dom comp_cod_arr
                      interchange [of "(f  g)  h" "𝖺-1[f, g, h]" TTfgh_TfTgh.the_θ
                                      "((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine"]
                by simp
              finally show ?thesis
                using comp_assoc by presburger
            qed
            also have "... =
                         (𝖺[f, g, h  tab0 h  TTfgh.p0] 
                           ((f  g)  h  TTfgh_TfTgh.the_θ)) 
                           𝖺[f  g, h, ((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine] 
                           (𝖺-1[f, g, h]  ((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine)"
              using fg gh assoc_naturality [of "f  g" h TTfgh_TfTgh.the_θ] comp_assoc
              by simp
            also have "... =
                         (f  g  h  TTfgh_TfTgh.the_θ) 
                           𝖺[f, g, h  ((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine] 
                           𝖺[f  g, h, ((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine] 
                           (𝖺-1[f, g, h]  ((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine)"
              using fg gh assoc_naturality [of f g "h  TTfgh_TfTgh.the_θ"] comp_assoc
              by simp
            finally show ?thesis
              using comp_assoc by presburger
          qed
          also have "... =
                       ((f  g  h  TTfgh_TfTgh.the_θ) 
                       (f  can (g  h  ((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine)
                                (((g  (h  tab0 h)  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine)) 
                       (f  (((g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                            (((g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                            ((𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                            (((g.tab  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                            (f0gh1  TTfgh_TfTgh.chine) 
                            𝖺-1[tab0 f, TfTgh.p1, TTfgh_TfTgh.chine])) 
                       𝖺[f, tab0 f, TfTgh.p1  TTfgh_TfTgh.chine] 
                       (f.tab  TfTgh.p1  TTfgh_TfTgh.chine) 
                       𝖺[tab1 f, TfTgh.p1, TTfgh_TfTgh.chine] 
                       TTfgh_TfTgh.the_ν 
                       𝖺-1[tab1 f, Tfg.p1, TTfgh.p1]"
          proof -
            have "𝖺[f, g, h  ((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine] 
                       𝖺[f  g, h, ((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine] 
                       (𝖺-1[f, g, h]  ((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       𝖺[f  g  h, (tab0 h  Tgh.p0)  TfTgh.p0, TTfgh_TfTgh.chine] 
                       (𝖺-1[f, g  h, (tab0 h  Tgh.p0)  TfTgh.p0]  TTfgh_TfTgh.chine) 
                       𝖺-1[f, (g  h)  (tab0 h  Tgh.p0)  TfTgh.p0, TTfgh_TfTgh.chine] 
                       (f  𝖺[g  h, tab0 h  Tgh.p0, TfTgh.p0]  TTfgh_TfTgh.chine) 
                       (f  (𝖺-1[g, h, tab0 h  Tgh.p0]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                       (f  ((g  𝖺[h, tab0 h, Tgh.p0])  TfTgh.p0)  TTfgh_TfTgh.chine) =
                    f  can
                          (g  h  ((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine)
                          (((g  (h  tab0 h)  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine)"
            proof -
              have "𝖺[f, g, h  ((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine] 
                        𝖺[f  g, h, ((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine] 
                        (𝖺-1[f, g, h]  ((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                        𝖺[f  g  h, (tab0 h  Tgh.p0)  TfTgh.p0, TTfgh_TfTgh.chine] 
                        (𝖺-1[f, g  h, (tab0 h  Tgh.p0)  TfTgh.p0]  TTfgh_TfTgh.chine) 
                        𝖺-1[f, (g  h)  (tab0 h  Tgh.p0)  TfTgh.p0, TTfgh_TfTgh.chine] 
                        (f  𝖺[g  h, tab0 h  Tgh.p0, TfTgh.p0]  TTfgh_TfTgh.chine) 
                        (f  (𝖺-1[g, h, tab0 h  Tgh.p0]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                        (f  ((g  𝖺[h, tab0 h, Tgh.p0])  TfTgh.p0)  TTfgh_TfTgh.chine) =
                      𝖺[f, g,
                         h  ((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine] 
                        𝖺[f  g, h,
                         ((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine] 
                        (𝖺-1[f, g, h]
                            ((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                        𝖺[f  g  h, (tab0 h  Tgh.p0)  TfTgh.p0,
                          TTfgh_TfTgh.chine] 
                        (𝖺-1[f, g  h,
                             (tab0 h  Tgh.p0)  TfTgh.p0]  TTfgh_TfTgh.chine) 
                        𝖺-1[f, (g  h)  (tab0 h  Tgh.p0)  TfTgh.p0,
                            TTfgh_TfTgh.chine] 
                        (f  𝖺[g  h, tab0 h  Tgh.p0, TfTgh.p0] 
                          TTfgh_TfTgh.chine) 
                        (f  (𝖺-1[g, h, tab0 h  Tgh.p0]  TfTgh.p0)
                            TTfgh_TfTgh.chine) 
                        (f  ((g  𝖺[h, tab0 h, Tgh.p0])  TfTgh.p0)
                            TTfgh_TfTgh.chine)"
                using 𝖺'_def α_def by simp
              also have "... =
                         can (f  (g  h  ((tab0 h  Tgh.p0)  TfTgh.p0)
                                 TTfgh_TfTgh.chine))
                           (f  (((g  (h  tab0 h)  Tgh.p0)  TfTgh.p0)
                                 TTfgh_TfTgh.chine))"
                using fg gh
                apply (unfold can_def)
                apply (intro E.eval_eqI)
                by simp_all
              also have "... =
                         f  can (g  h  ((tab0 h  Tgh.p0)  TfTgh.p0)
                                     TTfgh_TfTgh.chine)
                                 (((g  (h  tab0 h)  Tgh.p0)  TfTgh.p0)
                                     TTfgh_TfTgh.chine)"
                using fg gh whisker_can_left_0 by simp
              finally show ?thesis by blast
            qed
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... =
                       (f  (g  h  TTfgh_TfTgh.the_θ) 
                          can (g  h  ((tab0 h  Tgh.p0)  TfTgh.p0)
                                  TTfgh_TfTgh.chine)
                              (((g  (h  tab0 h)  Tgh.p0)  TfTgh.p0)
                                  TTfgh_TfTgh.chine) 
                          (((g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                          (((g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                          ((𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                          (((g.tab  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                          (f0gh1  TTfgh_TfTgh.chine) 
                          𝖺-1[tab0 f, TfTgh.p1, TTfgh_TfTgh.chine]) 
                       𝖺[f, tab0 f, TfTgh.p1  TTfgh_TfTgh.chine] 
                       (f.tab  TfTgh.p1  TTfgh_TfTgh.chine) 
                       𝖺[tab1 f, TfTgh.p1, TTfgh_TfTgh.chine] 
                       TTfgh_TfTgh.the_ν 
                       𝖺-1[tab1 f, Tfg.p1, TTfgh.p1]"
          proof -
            have "((f  g  h  TTfgh_TfTgh.the_θ) 
                       (f  can (g  h  ((tab0 h  Tgh.p0)  TfTgh.p0)
                                    TTfgh_TfTgh.chine)
                                (((g  (h  tab0 h)  Tgh.p0)  TfTgh.p0)
                                    TTfgh_TfTgh.chine)) 
                       (f  (((g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                            (((g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                            ((𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                            (((g.tab  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                            (f0gh1  TTfgh_TfTgh.chine) 
                            𝖺-1[tab0 f, TfTgh.p1, TTfgh_TfTgh.chine])) =
                    f  (g  h  TTfgh_TfTgh.the_θ) 
                          can (g  h  ((tab0 h  Tgh.p0)  TfTgh.p0)
                                  TTfgh_TfTgh.chine)
                              (((g  (h  tab0 h)  Tgh.p0)  TfTgh.p0)
                                  TTfgh_TfTgh.chine) 
                          (((g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                          (((g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                          ((𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                          (((g.tab  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                          (f0gh1  TTfgh_TfTgh.chine) 
                          𝖺-1[tab0 f, TfTgh.p1, TTfgh_TfTgh.chine]"
            proof -
              have 1: "arr ((g  h  TTfgh_TfTgh.the_θ) 
                           can (g  h  ((tab0 h  Tgh.p0)  TfTgh.p0)
                                   TTfgh_TfTgh.chine)
                               (((g  (h  tab0 h)  Tgh.p0)  TfTgh.p0)
                                   TTfgh_TfTgh.chine) 
                           (((g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           (((g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           ((𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           (((g.tab  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           (f0gh1  TTfgh_TfTgh.chine) 
                           𝖺-1[tab0 f, TfTgh.p1, TTfgh_TfTgh.chine])"
                using fg gh
                apply (intro seqI' comp_in_homI) by auto
              moreover
              have 2: "arr (can (g  h  ((tab0 h  Tgh.p0)  TfTgh.p0)
                                    TTfgh_TfTgh.chine)
                               (((g  (h  tab0 h)  Tgh.p0)  TfTgh.p0)
                                    TTfgh_TfTgh.chine) 
                           (((g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           (((g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           ((𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           (((g.tab  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           (f0gh1  TTfgh_TfTgh.chine) 
                           𝖺-1[tab0 f, TfTgh.p1, TTfgh_TfTgh.chine])"
                using calculation by blast
              ultimately show ?thesis
                using whisker_left f.ide_base by presburger
            qed
            thus ?thesis
              using comp_assoc by presburger
          qed
          also have "... = f.composite_cell wf' θf'  βf"
            unfolding wf'_def θf'_def βf_def
            using comp_assoc by presburger
          finally show ?thesis by blast
        qed
        show ?thesis
          using wf wf' θf θf' βf f.T2 [of wf wf' θf uf θf' βf] eqf by fast
      qed
      obtain γf where γf: "«γf : wf  wf'»  βf = tab1 f  γf  θf = θf'  (tab0 f  γf)"
        using 5 by auto
      show "⟦⟦TfTgh.p1  TTfgh_TfTgh.chine⟧⟧ = ⟦⟦Tfg.p1  TTfgh.p1⟧⟧"
      proof -
        have "iso γf"
          using γf BS3 wf_is_map wf'_is_map by blast
        hence "isomorphic wf wf'"
          using γf isomorphic_def isomorphic_symmetric by auto
        thus ?thesis
          using wf wf_def wf'_def Maps.CLS_eqI isomorphic_symmetric by auto
      qed
      text ‹
        On to the next equation:
        \[
           ⟦⟦Tgh.p1 ⋆ TfTgh.p0 ⋆ TTfgh_TfTgh.chine⟧⟧ = ⟦⟦Tfg.p0 ⋆ TTfgh.p1⟧⟧›.
        \]
        We have to make use of the equation θf = θf' ⋅ (tab0 f ⋆ γf)› in this part,
        similarly to how the equation src_tab_eq› was used to replace
        TTfgh.tab› in the first part.
      ›
      define ug where "ug = h  tab0 h  TTfgh.p0"
      define wg where "wg = Tfg.p0  TTfgh.p1"
      define wg' where "wg' = Tgh.p1  TfTgh.p0  TTfgh_TfTgh.chine"
      define θg
      where "θg = 𝖺[h, tab0 h, TTfgh.p0]  (h.tab  TTfgh.p0)  fg0h1 
                    𝖺-1[tab0 g, Tfg.p0, TTfgh.p1]"
      define θg'
      where "θg' = (h  TTfgh_TfTgh.the_θ) 
                   can (h  ((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine)
                       (((h  tab0 h)  Tgh.p0)  TfTgh.p0  TTfgh_TfTgh.chine) 
                   ((h.tab  Tgh.p0)  TfTgh.p0  TTfgh_TfTgh.chine) 
                   (g0h1  TfTgh.p0  TTfgh_TfTgh.chine) 
                   𝖺-1[tab0 g, Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine]"
      define βg
      where "βg = 𝖺[tab1 g, Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine] 
                  𝖺[tab1 g  Tgh.p1, TfTgh.p0, TTfgh_TfTgh.chine] 
                  (f0gh1  TTfgh_TfTgh.chine)  𝖺-1[tab0 f, TfTgh.p1, TTfgh_TfTgh.chine] 
                  (tab0 f  γf)  𝖺[tab0 f, Tfg.p1, TTfgh.p1]  (inv f0g1  TTfgh.p1) 
                  𝖺-1[tab1 g, Tfg.p0, TTfgh.p1]"
      have ug: "ide ug"
        unfolding ug_def by simp
      have wg: "ide wg"
        unfolding wg_def using fg0h1.p1_simps by auto
      have wg_is_map: "is_left_adjoint wg"
        unfolding wg_def
        using fg0h1.p1_simps left_adjoints_compose by simp
      have wg': "ide wg'"
        unfolding wg'_def by simp
      have wg'_is_map: "is_left_adjoint wg'"
        unfolding wg'_def
        using TTfgh_TfTgh.is_map left_adjoints_compose by simp
      have θg: "«θg : tab0 g  wg  ug»"
        using wg_def ug_def θg_def fg0h1.p1_simps fg0h1.φ_in_hom by auto
      have θg': "«θg' : tab0 g  wg'  ug»"
        unfolding wg'_def ug_def θg'_def
        by (intro comp_in_homI) auto
      have wg_in_hhom: "in_hhom wg (src ug) (src (tab0 g))"
        unfolding wg_def ug_def by auto
      have wg'_in_hhom: "in_hhom wg' (src ug) (src (tab0 g))"
        unfolding wg'_def ug_def by auto
      have βg: "«βg : tab1 g  wg  tab1 g  wg'»"
      proof (unfold βg_def wg_def, intro comp_in_homI)
        (* auto can solve this, but it's too slow *)
        show "«𝖺-1[tab1 g, Tfg.p0, TTfgh.p1] :
                 tab1 g  Tfg.p0  TTfgh.p1  (tab1 g  Tfg.p0)  TTfgh.p1»"
          using fg0h1.p1_simps by auto
        show "«inv f0g1  TTfgh.p1 :
                 (tab1 g  Tfg.p0)  TTfgh.p1  (tab0 f  Tfg.p1)  TTfgh.p1»"
          using fg0h1.p1_simps f0g1.φ_in_hom f0g1.φ_uniqueness(2)
          by (intro hcomp_in_vhom) auto
        show "«𝖺[tab0 f, Tfg.p1, TTfgh.p1] :
                 (tab0 f  Tfg.p1)  TTfgh.p1  tab0 f  Tfg.p1  TTfgh.p1»"
          using fg0h1.p1_simps γf wf_def wf'_def by auto
        show "«tab0 f  γf : tab0 f  Tfg.p1  TTfgh.p1  tab0 f  TfTgh.p1  TTfgh_TfTgh.chine»"
          using fg0h1.p1_simps γf wf_def wf'_def by auto
        show "«𝖺-1[tab0 f, TfTgh.p1, TTfgh_TfTgh.chine] :
                 tab0 f  TfTgh.p1  TTfgh_TfTgh.chine  (tab0 f  TfTgh.p1)  TTfgh_TfTgh.chine»"
          by auto
        show "«f0gh1  TTfgh_TfTgh.chine :
                 (tab0 f  TfTgh.p1)  TTfgh_TfTgh.chine
                     ((tab1 g  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine»"
          using f0gh1.φ_in_hom
          by (intro hcomp_in_vhom) auto
        show "«𝖺[tab1 g  Tgh.p1, TfTgh.p0, TTfgh_TfTgh.chine] :
                 ((tab1 g  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine
                     (tab1 g  Tgh.p1)  TfTgh.p0  TTfgh_TfTgh.chine»"
          by auto
        show "«𝖺[tab1 g, Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine] :
                 (tab1 g  Tgh.p1)  TfTgh.p0  TTfgh_TfTgh.chine  tab1 g  wg'»"
          using wg'_def by auto
      qed
      have eqg: "g.composite_cell wg θg = g.composite_cell wg' θg'  βg"
      proof -
        have "g.composite_cell wg θg =
              (g  𝖺[h, tab0 h, TTfgh.p0] 
                     (h.tab  TTfgh.p0) 
                     fg0h1 
                     𝖺-1[tab0 g, Tfg.p0, TTfgh.p1]) 
                𝖺[g, tab0 g, Tfg.p0  TTfgh.p1] 
                (g.tab  Tfg.p0  TTfgh.p1)"
          unfolding wg_def θg_def by simp
        also have "... =
                   (g  𝖺[h, tab0 h, TTfgh.p0]) 
                     (g  h.tab  TTfgh.p0) 
                     (g  fg0h1) 
                     ((g  𝖺-1[tab0 g, Tfg.p0, TTfgh.p1]) 
                     𝖺[g, tab0 g, Tfg.p0  TTfgh.p1]) 
                     (g.tab  Tfg.p0  TTfgh.p1)"
          using fg gh f0g1.p0_simps fg0h1.p0_simps fg0h1.p1_simps whisker_left
                comp_assoc
          by simp
        also have "... =
                   (g  𝖺[h, tab0 h, TTfgh.p0]) 
                     (g  h.tab  TTfgh.p0) 
                     (g  fg0h1) 
                     (𝖺[g, tab0 g  Tfg.p0, TTfgh.p1] 
                     (𝖺-1[g, tab0 g  Tfg.p0, TTfgh.p1] 
                     (g  𝖺-1[tab0 g, Tfg.p0, TTfgh.p1]))) 
                     𝖺[g, tab0 g, Tfg.p0  TTfgh.p1] 
                     (g.tab  Tfg.p0  TTfgh.p1)"
        proof -
          have "(𝖺[g, tab0 g  Tfg.p0, TTfgh.p1] 
                   𝖺-1[g, tab0 g  Tfg.p0, TTfgh.p1]) 
                   (g  𝖺-1[tab0 g, Tfg.p0, TTfgh.p1]) =
                g  𝖺-1[tab0 g, Tfg.p0, TTfgh.p1]"
            using fg0h1.p1_simps comp_cod_arr comp_assoc_assoc' by simp
          thus ?thesis
            by (simp add: comp_assoc)
        qed
        also have "... =
                   (g  𝖺[h, tab0 h, TTfgh.p0]) 
                     (g  h.tab  TTfgh.p0) 
                     (g  fg0h1) 
                     𝖺[g, tab0 g  Tfg.p0, TTfgh.p1] 
                     (𝖺-1[g, tab0 g  Tfg.p0, TTfgh.p1] 
                     (g  𝖺-1[tab0 g, Tfg.p0, TTfgh.p1]) 
                     𝖺[g, tab0 g, Tfg.p0  TTfgh.p1]) 
                     (g.tab  Tfg.p0  TTfgh.p1)"
          using comp_assoc by presburger
        also have "... =
                   (g  𝖺[h, tab0 h, TTfgh.p0]) 
                     (g  h.tab  TTfgh.p0) 
                     (g  fg0h1) 
                     𝖺[g, tab0 g  Tfg.p0, TTfgh.p1] 
                     (𝖺[g, tab0 g, Tfg.p0]  TTfgh.p1) 
                     (𝖺-1[g  tab0 g, Tfg.p0, TTfgh.p1] 
                     (g.tab  Tfg.p0  TTfgh.p1))"
          using fg gh f0g1.p0_simps fg0h1.p0_simps fg0h1.p1_simps fg0h1.p1_simps comp_assoc pentagon'
                invert_opposite_sides_of_square
                  [of "𝖺-1[g, tab0 g, Tfg.p0]  TTfgh.p1"
                      "(𝖺-1[g, tab0 g  Tfg.p0, TTfgh.p1])  (g  𝖺-1[tab0 g, Tfg.p0, TTfgh.p1])"
                      "𝖺-1[g  tab0 g, Tfg.p0, TTfgh.p1]" "𝖺-1[g, tab0 g, Tfg.p0  TTfgh.p1]"]
          by simp
        also have "... =
                   (g  𝖺[h, tab0 h, TTfgh.p0]) 
                     (g  h.tab  TTfgh.p0) 
                     (g  fg0h1) 
                     𝖺[g, tab0 g  Tfg.p0, TTfgh.p1] 
                     (𝖺[g, tab0 g, Tfg.p0]  TTfgh.p1) 
                     ((g.tab  Tfg.p0)  TTfgh.p1) 
                     𝖺-1[tab1 g, Tfg.p0, TTfgh.p1]"
          using fg0h1.p1_simps assoc'_naturality [of g.tab Tfg.p0 TTfgh.p1] by simp
        also have "... =
                   (g  𝖺[h, tab0 h, TTfgh.p0]) 
                     (g  h.tab  TTfgh.p0) 
                     (g  fg0h1) 
                     𝖺[g, tab0 g  Tfg.p0, TTfgh.p1] 
                     (𝖺[g, tab0 g, Tfg.p0]  TTfgh.p1) 
                     ((g.tab  Tfg.p0)  TTfgh.p1) 
                     (f0g1  TTfgh.p1) 
                     𝖺-1[tab0 f, Tfg.p1, TTfgh.p1] 
                     𝖺[tab0 f, Tfg.p1, TTfgh.p1] 
                     (inv f0g1  TTfgh.p1) 
                     𝖺-1[tab1 g, Tfg.p0, TTfgh.p1]"
        proof -
          have "(f0g1  TTfgh.p1) 
                     𝖺-1[tab0 f, Tfg.p1, TTfgh.p1] 
                     𝖺[tab0 f, Tfg.p1, TTfgh.p1] 
                     (inv f0g1  TTfgh.p1) 
                     𝖺-1[tab1 g, Tfg.p0, TTfgh.p1] =
                ((f0g1  TTfgh.p1) 
                     (𝖺-1[tab0 f, Tfg.p1, TTfgh.p1] 
                     𝖺[tab0 f, Tfg.p1, TTfgh.p1]) 
                     (inv f0g1  TTfgh.p1)) 
                     𝖺-1[tab1 g, Tfg.p0, TTfgh.p1]"
            using comp_assoc by presburger
          also have "... = ((f0g1  TTfgh.p1) 
                             ((tab0 f  Tfg.p1)  TTfgh.p1) 
                             (inv f0g1  TTfgh.p1)) 
                             𝖺-1[tab1 g, Tfg.p0, TTfgh.p1]"
            using fg0h1.p1_simps whisker_right comp_assoc_assoc' by simp
          also have "... = ((f0g1  TTfgh.p1) 
                             (inv f0g1  TTfgh.p1)) 
                             𝖺-1[tab1 g, Tfg.p0, TTfgh.p1]"
            using fg0h1.p1_simps f0g1.φ_uniqueness comp_cod_arr by simp
          also have "... = ((tab1 g  Tfg.p0)  TTfgh.p1)  𝖺-1[tab1 g, Tfg.p0, TTfgh.p1]"
          proof -
            have "(f0g1  TTfgh.p1)  (inv f0g1  TTfgh.p1) =
                  f0g1  inv f0g1  TTfgh.p1"
              using f0g1.φ_uniqueness whisker_right by simp
            also have "... = (tab1 g  Tfg.p0)  TTfgh.p1"
              using f0g1.φ_uniqueness comp_arr_inv' by simp
            finally show ?thesis by simp
          qed
          also have "... = 𝖺-1[tab1 g, Tfg.p0, TTfgh.p1]"
            using fg0h1.p1_simps comp_cod_arr by simp
          finally have "(f0g1  TTfgh.p1)  𝖺-1[tab0 f, Tfg.p1, TTfgh.p1] 
                          𝖺[tab0 f, Tfg.p1, TTfgh.p1]  (inv f0g1  TTfgh.p1) 
                     𝖺-1[tab1 g, Tfg.p0, TTfgh.p1] = 𝖺-1[tab1 g, Tfg.p0, TTfgh.p1]"
            by simp
          thus ?thesis
            using comp_assoc by presburger
        qed
        also have "... = θf 
                           𝖺[tab0 f, Tfg.p1, TTfgh.p1] 
                           (inv f0g1  TTfgh.p1) 
                           𝖺-1[tab1 g, Tfg.p0, TTfgh.p1]"
          unfolding θf_def using comp_assoc by presburger
        also have "... = θf'  (tab0 f  γf) 
                           𝖺[tab0 f, Tfg.p1, TTfgh.p1] 
                           (inv f0g1  TTfgh.p1) 
                           𝖺-1[tab1 g, Tfg.p0, TTfgh.p1]"
          using γf comp_assoc by simp
        also have "... = (g  h  TTfgh_TfTgh.the_θ) 
                           can (g  h  ((tab0 h  Tgh.p0)  TfTgh.p0)
                                   TTfgh_TfTgh.chine)
                               (((g  (h  tab0 h)  Tgh.p0)  TfTgh.p0)
                                   TTfgh_TfTgh.chine) 
                           (((g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           (((g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           ((𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           (((g.tab  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           (f0gh1  TTfgh_TfTgh.chine) 
                           𝖺-1[tab0 f, TfTgh.p1, TTfgh_TfTgh.chine] 
                           (tab0 f  γf) 
                           𝖺[tab0 f, Tfg.p1, TTfgh.p1] 
                           (inv f0g1  TTfgh.p1) 
                           𝖺-1[tab1 g, Tfg.p0, TTfgh.p1]"
          unfolding θf'_def using comp_assoc by presburger
        also have "... = (g  h  TTfgh_TfTgh.the_θ) 
                           can (g  h  ((tab0 h  Tgh.p0)  TfTgh.p0)
                                   TTfgh_TfTgh.chine)
                               (((g  (h  tab0 h)  Tgh.p0)  TfTgh.p0)
                                   TTfgh_TfTgh.chine) 
                           (((g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           (((g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           ((𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           (((g.tab  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           (𝖺-1[tab1 g  Tgh.p1, TfTgh.p0, TTfgh_TfTgh.chine] 
                           (𝖺-1[tab1 g, Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine] 
                           𝖺[tab1 g, Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine]) 
                           𝖺[tab1 g  Tgh.p1, TfTgh.p0, TTfgh_TfTgh.chine] 
                           (f0gh1  TTfgh_TfTgh.chine)) 
                           𝖺-1[tab0 f, TfTgh.p1, TTfgh_TfTgh.chine] 
                           (tab0 f  γf) 
                           𝖺[tab0 f, Tfg.p1, TTfgh.p1] 
                           (inv f0g1  TTfgh.p1) 
                           𝖺-1[tab1 g, Tfg.p0, TTfgh.p1]"
        proof -
          have "(𝖺-1[tab1 g  Tgh.p1, TfTgh.p0, TTfgh_TfTgh.chine] 
                           (𝖺-1[tab1 g, Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine] 
                           𝖺[tab1 g, Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine]) 
                           𝖺[tab1 g  Tgh.p1, TfTgh.p0, TTfgh_TfTgh.chine]) 
                           (f0gh1  TTfgh_TfTgh.chine) =
                f0gh1  TTfgh_TfTgh.chine"
            using f0gh1.p0_simps comp_cod_arr comp_arr_dom comp_assoc_assoc' by simp
          thus ?thesis
            using comp_assoc by fastforce
        qed
        also have "... = (g  h  TTfgh_TfTgh.the_θ) 
                           can (g  h  ((tab0 h  Tgh.p0)  TfTgh.p0)
                                   TTfgh_TfTgh.chine)
                               (((g  (h  tab0 h)  Tgh.p0)  TfTgh.p0)
                                   TTfgh_TfTgh.chine) 
                           (((g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           (((g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           ((𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           ((((g.tab  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           𝖺-1[tab1 g  Tgh.p1, TfTgh.p0, TTfgh_TfTgh.chine]) 
                           𝖺-1[tab1 g, Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine] 
                           βg"
          unfolding βg_def using comp_assoc by presburger
        also have "... = (g  h  TTfgh_TfTgh.the_θ) 
                           can (g  h  ((tab0 h  Tgh.p0)  TfTgh.p0)
                                   TTfgh_TfTgh.chine)
                               (((g  (h  tab0 h)  Tgh.p0)  TfTgh.p0)
                                   TTfgh_TfTgh.chine) 
                           (((g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           (((g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           ((𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           𝖺-1[(g  tab0 g)  Tgh.p1, TfTgh.p0, TTfgh_TfTgh.chine] 
                           (((g.tab  Tgh.p1)  TfTgh.p0  TTfgh_TfTgh.chine) 
                           𝖺-1[tab1 g, Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine]) 
                           βg"
        proof -
          have "(((g.tab  Tgh.p1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                  𝖺-1[tab1 g  Tgh.p1, TfTgh.p0, TTfgh_TfTgh.chine] =
                𝖺-1[(g  tab0 g)  Tgh.p1, TfTgh.p0, TTfgh_TfTgh.chine] 
                  ((g.tab  Tgh.p1)  TfTgh.p0  TTfgh_TfTgh.chine)"
            using f0gh1.p0_simps
                  assoc'_naturality [of "(g.tab  Tgh.p1)" TfTgh.p0 TTfgh_TfTgh.chine]
            by simp
          thus ?thesis
            using comp_assoc by presburger
        qed
        also have "... = (g  h  TTfgh_TfTgh.the_θ) 
                           can (g  h  ((tab0 h  Tgh.p0)  TfTgh.p0)
                                   TTfgh_TfTgh.chine)
                               (((g  (h  tab0 h)  Tgh.p0)  TfTgh.p0)
                                   TTfgh_TfTgh.chine) 
                           (((g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           (((g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           ((𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           𝖺-1[(g  tab0 g)  Tgh.p1, TfTgh.p0, TTfgh_TfTgh.chine] 
                           𝖺-1[g  tab0 g, Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine] 
                           (g.tab  Tgh.p1  TfTgh.p0  TTfgh_TfTgh.chine) 
                           βg"
        proof -
          have "((g.tab  Tgh.p1)  TfTgh.p0  TTfgh_TfTgh.chine) 
                  𝖺-1[tab1 g, Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine] =
                𝖺-1[g  tab0 g, Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine] 
                  (g.tab  Tgh.p1  TfTgh.p0  TTfgh_TfTgh.chine)"
            using g0h1.p1_simps f0gh1.p0_simps
                  assoc'_naturality [of g.tab Tgh.p1 "TfTgh.p0  TTfgh_TfTgh.chine"]
            by simp
          thus ?thesis
            using comp_assoc by presburger
        qed
        also have "... = (g  h  TTfgh_TfTgh.the_θ) 
                           can (g  h  ((tab0 h  Tgh.p0)  TfTgh.p0)
                                   TTfgh_TfTgh.chine)
                               (((g  (h  tab0 h)  Tgh.p0)  TfTgh.p0)
                                   TTfgh_TfTgh.chine) 
                           (((g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           (((g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           ((𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           𝖺-1[(g  tab0 g)  Tgh.p1, TfTgh.p0, TTfgh_TfTgh.chine] 
                           𝖺-1[g  tab0 g, Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine] 
                           ((𝖺-1[g, tab0 g, Tgh.p1  TfTgh.p0  TTfgh_TfTgh.chine] 
                           𝖺[g, tab0 g, Tgh.p1  TfTgh.p0  TTfgh_TfTgh.chine]) 
                           (g.tab  Tgh.p1  TfTgh.p0  TTfgh_TfTgh.chine)) 
                           βg"
        proof -
          have "(𝖺-1[g, tab0 g, Tgh.p1  TfTgh.p0  TTfgh_TfTgh.chine] 
                  𝖺[g, tab0 g, Tgh.p1  TfTgh.p0  TTfgh_TfTgh.chine]) 
                  (g.tab  Tgh.p1  TfTgh.p0  TTfgh_TfTgh.chine) =
                g.tab  Tgh.p1  TfTgh.p0  TTfgh_TfTgh.chine"
            using comp_cod_arr comp_assoc_assoc' by simp
          thus ?thesis
            using comp_assoc g0h1.φ_in_hom by simp
        qed
        also have "... = (g  h  TTfgh_TfTgh.the_θ) 
                           can (g  h  ((tab0 h  Tgh.p0)  TfTgh.p0)
                                   TTfgh_TfTgh.chine)
                               (((g  (h  tab0 h)  Tgh.p0)  TfTgh.p0)
                                   TTfgh_TfTgh.chine) 
                           (((g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           (((g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           ((𝖺-1[g  tab0 g  Tgh.p1, TfTgh.p0, TTfgh_TfTgh.chine] 
                           (𝖺-1[g, tab0 g  Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine] 
                           𝖺[g, tab0 g  Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine]) 
                           𝖺[g  tab0 g  Tgh.p1, TfTgh.p0, TTfgh_TfTgh.chine]) 
                           ((𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine)) 
                           𝖺-1[(g  tab0 g)  Tgh.p1, TfTgh.p0, TTfgh_TfTgh.chine] 
                           𝖺-1[g  tab0 g, Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine] 
                           ((𝖺-1[g, tab0 g, Tgh.p1  TfTgh.p0  TTfgh_TfTgh.chine] 
                           𝖺[g, tab0 g, Tgh.p1  TfTgh.p0  TTfgh_TfTgh.chine]) 
                           (g.tab  Tgh.p1  TfTgh.p0  TTfgh_TfTgh.chine)) 
                           βg"
        proof -
          have "(𝖺-1[g  tab0 g  Tgh.p1, TfTgh.p0, TTfgh_TfTgh.chine] 
                           (𝖺-1[g, tab0 g  Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine] 
                           𝖺[g, tab0 g  Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine]) 
                           𝖺[g  tab0 g  Tgh.p1, TfTgh.p0, TTfgh_TfTgh.chine]) 
                           ((𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) =
                (𝖺-1[g  tab0 g  Tgh.p1, TfTgh.p0, TTfgh_TfTgh.chine] 
                           ((g  (tab0 g  Tgh.p1))  TfTgh.p0  TTfgh_TfTgh.chine) 
                           𝖺[g  tab0 g  Tgh.p1, TfTgh.p0, TTfgh_TfTgh.chine]) 
                           ((𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine)"
            using g0h1.p1_simps f0gh1.p0_simps comp_assoc comp_assoc_assoc' by simp
          also have "... = (𝖺-1[g  tab0 g  Tgh.p1, TfTgh.p0, TTfgh_TfTgh.chine] 
                           𝖺[g  tab0 g  Tgh.p1, TfTgh.p0, TTfgh_TfTgh.chine]) 
                           ((𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine)"
            using g0h1.p1_simps f0gh1.p0_simps comp_cod_arr comp_assoc_assoc' by simp
          also have "... = (((g  (tab0 g  Tgh.p1))  TfTgh.p0)  TTfgh_TfTgh.chine) 
                             ((𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine)"
            using g0h1.p1_simps f0gh1.p0_simps whisker_right comp_assoc_assoc' by simp
          also have "... = (𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine"
            using g0h1.p1_simps f0gh1.p0_simps comp_cod_arr by simp
          finally show ?thesis by presburger
        qed
        also have "... = (g  h  TTfgh_TfTgh.the_θ) 
                           can (g  h  ((tab0 h  Tgh.p0)  TfTgh.p0)
                                   TTfgh_TfTgh.chine)
                               (((g  (h  tab0 h)  Tgh.p0)  TfTgh.p0)
                                   TTfgh_TfTgh.chine) 
                           ((((g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           (((g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           𝖺-1[g  tab0 g  Tgh.p1, TfTgh.p0, TTfgh_TfTgh.chine] 
                           𝖺-1[g, tab0 g  Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine]) 
                           𝖺[g, tab0 g  Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine] 
                           𝖺[g  tab0 g  Tgh.p1, TfTgh.p0, TTfgh_TfTgh.chine] 
                           ((𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           𝖺-1[(g  tab0 g)  Tgh.p1, TfTgh.p0, TTfgh_TfTgh.chine] 
                           𝖺-1[g  tab0 g, Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine] 
                           (𝖺-1[g, tab0 g, Tgh.p1  TfTgh.p0  TTfgh_TfTgh.chine] 
                           𝖺[g, tab0 g, Tgh.p1  TfTgh.p0  TTfgh_TfTgh.chine] 
                           (g.tab  Tgh.p1  TfTgh.p0  TTfgh_TfTgh.chine)) 
                           βg"
          using comp_assoc by presburger
        also have "... = (g  h  TTfgh_TfTgh.the_θ) 
                           (can (g  h  ((tab0 h  Tgh.p0)  TfTgh.p0)
                                    TTfgh_TfTgh.chine)
                               (((g  (h  tab0 h)  Tgh.p0)  TfTgh.p0)
                                    TTfgh_TfTgh.chine) 
                𝖺-1[g  (h  tab0 h)  Tgh.p0, TfTgh.p0, TTfgh_TfTgh.chine] 
                    𝖺-1[g, (h  tab0 h)  Tgh.p0, TfTgh.p0  TTfgh_TfTgh.chine]) 
                      (g  (h.tab  Tgh.p0)  TfTgh.p0  TTfgh_TfTgh.chine) 
                      (g  g0h1  TfTgh.p0  TTfgh_TfTgh.chine) 
                           (𝖺[g, tab0 g  Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine] 
                           𝖺[g  tab0 g  Tgh.p1, TfTgh.p0, TTfgh_TfTgh.chine] 
                           ((𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           𝖺-1[(g  tab0 g)  Tgh.p1, TfTgh.p0, TTfgh_TfTgh.chine] 
                           𝖺-1[g  tab0 g, Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine] 
                           𝖺-1[g, tab0 g, Tgh.p1  TfTgh.p0  TTfgh_TfTgh.chine]) 
                           𝖺[g, tab0 g, Tgh.p1  TfTgh.p0  TTfgh_TfTgh.chine] 
                           (g.tab  Tgh.p1  TfTgh.p0  TTfgh_TfTgh.chine) 
                           βg"
        proof -
          have "(((g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           (((g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           𝖺-1[g  tab0 g  Tgh.p1, TfTgh.p0, TTfgh_TfTgh.chine] 
                           𝖺-1[g, tab0 g  Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine] =
                𝖺-1[g  (h  tab0 h)  Tgh.p0, TfTgh.p0, TTfgh_TfTgh.chine] 
                    𝖺-1[g, (h  tab0 h)  Tgh.p0, TfTgh.p0  TTfgh_TfTgh.chine] 
                      (g  (h.tab  Tgh.p0)  TfTgh.p0  TTfgh_TfTgh.chine) 
                      (g  g0h1  TfTgh.p0  TTfgh_TfTgh.chine)"
          proof -
            have "(((g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           (((g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           𝖺-1[g  tab0 g  Tgh.p1, TfTgh.p0, TTfgh_TfTgh.chine] 
                           𝖺-1[g, tab0 g  Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine] =
                  (((g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           ((((g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           𝖺-1[g  tab0 g  Tgh.p1, TfTgh.p0, TTfgh_TfTgh.chine]) 
                           𝖺-1[g, tab0 g  Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine]"
              using comp_assoc by presburger
            also have "... = ((((g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                               𝖺-1[g  tab1 h  Tgh.p0, TfTgh.p0, TTfgh_TfTgh.chine]) 
                               ((g  g0h1)  TfTgh.p0  TTfgh_TfTgh.chine) 
                               𝖺-1[g, tab0 g  Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine]"
            proof -
              have "(((g  g0h1)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           𝖺-1[g  tab0 g  Tgh.p1, TfTgh.p0, TTfgh_TfTgh.chine] =
                    𝖺-1[g  tab1 h  Tgh.p0, TfTgh.p0, TTfgh_TfTgh.chine] 
                      ((g  g0h1)  TfTgh.p0  TTfgh_TfTgh.chine)"
                using gh f0gh1.p0_simps
                      assoc'_naturality [of "g  g0h1" TfTgh.p0 TTfgh_TfTgh.chine]
                by simp
              thus ?thesis
                using comp_assoc by presburger
            qed
            also have "... = 𝖺-1[g  (h  tab0 h)  Tgh.p0, TfTgh.p0, TTfgh_TfTgh.chine] 
                      ((g  h.tab  Tgh.p0)  TfTgh.p0  TTfgh_TfTgh.chine) 
                               ((g  g0h1)  TfTgh.p0  TTfgh_TfTgh.chine) 
                               𝖺-1[g, tab0 g  Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine]"
            proof -
              have "(((g  h.tab  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                               𝖺-1[g  tab1 h  Tgh.p0, TfTgh.p0, TTfgh_TfTgh.chine] =
                    𝖺-1[g  (h  tab0 h)  Tgh.p0, TfTgh.p0, TTfgh_TfTgh.chine] 
                      ((g  h.tab  Tgh.p0)  TfTgh.p0  TTfgh_TfTgh.chine)"
                using gh f0gh1.p0_simps
                      assoc'_naturality [of "g  h.tab  Tgh.p0" TfTgh.p0 TTfgh_TfTgh.chine]
                by simp
              thus ?thesis
                using comp_assoc by presburger
            qed
            also have "... = 𝖺-1[g  (h  tab0 h)  Tgh.p0, TfTgh.p0, TTfgh_TfTgh.chine] 
                      (((g  h.tab  Tgh.p0)  TfTgh.p0  TTfgh_TfTgh.chine) 
                    𝖺-1[g, tab1 h  Tgh.p0, TfTgh.p0  TTfgh_TfTgh.chine]) 
                      (g  g0h1  TfTgh.p0  TTfgh_TfTgh.chine)"
            proof -
              have "((g  g0h1)  TfTgh.p0  TTfgh_TfTgh.chine) 
                      𝖺-1[g, tab0 g  Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine] =
                    𝖺-1[g, tab1 h  Tgh.p0, TfTgh.p0  TTfgh_TfTgh.chine] 
                      (g  g0h1  TfTgh.p0  TTfgh_TfTgh.chine)"
                using gh f0gh1.p0_simps
                      assoc'_naturality [of g g0h1 "TfTgh.p0  TTfgh_TfTgh.chine"]
                by simp
              thus ?thesis
                using comp_assoc by presburger
            qed
            also have "... = 𝖺-1[g  (h  tab0 h)  Tgh.p0, TfTgh.p0, TTfgh_TfTgh.chine] 
                    𝖺-1[g, (h  tab0 h)  Tgh.p0, TfTgh.p0  TTfgh_TfTgh.chine] 
                      (g  (h.tab  Tgh.p0)  TfTgh.p0  TTfgh_TfTgh.chine) 
                      (g  g0h1  TfTgh.p0  TTfgh_TfTgh.chine)"
            proof -
              have "((g  h.tab  Tgh.p0)  TfTgh.p0  TTfgh_TfTgh.chine) 
                    𝖺-1[g, tab1 h  Tgh.p0, TfTgh.p0  TTfgh_TfTgh.chine] =
                    𝖺-1[g, (h  tab0 h)  Tgh.p0, TfTgh.p0  TTfgh_TfTgh.chine] 
                      (g  (h.tab  Tgh.p0)  TfTgh.p0  TTfgh_TfTgh.chine)"
                using gh f0gh1.p0_simps
                      assoc'_naturality [of g "h.tab  Tgh.p0" "TfTgh.p0  TTfgh_TfTgh.chine"]
                by simp
              thus ?thesis
                using comp_assoc by presburger
            qed
            finally show ?thesis by simp
          qed
          thus ?thesis
            using comp_assoc by presburger
        qed
        also have "... = ((g  h  TTfgh_TfTgh.the_θ) 
                           (g  can (h  ((tab0 h  Tgh.p0)  TfTgh.p0)
                                        TTfgh_TfTgh.chine)
                                    (((h  tab0 h)  Tgh.p0)  TfTgh.p0
                                        TTfgh_TfTgh.chine)) 
                           (g  (h.tab  Tgh.p0)  TfTgh.p0  TTfgh_TfTgh.chine) 
                           (g  g0h1  TfTgh.p0  TTfgh_TfTgh.chine) 
                           (g  𝖺-1[tab0 g, Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine])) 
                           𝖺[g, tab0 g, Tgh.p1  TfTgh.p0  TTfgh_TfTgh.chine] 
                           (g.tab  Tgh.p1  TfTgh.p0  TTfgh_TfTgh.chine) 
                           βg"
        proof -
          have "can (g  h  ((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine)
                    (((g  (h  tab0 h)  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine) 
                𝖺-1[g  (h  tab0 h)  Tgh.p0, TfTgh.p0, TTfgh_TfTgh.chine] 
                    𝖺-1[g, (h  tab0 h)  Tgh.p0, TfTgh.p0  TTfgh_TfTgh.chine] =
                g  can (h  ((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine)
                        (((h  tab0 h)  Tgh.p0)  TfTgh.p0  TTfgh_TfTgh.chine)"
          proof -
            have "𝖺-1[g  (h  tab0 h)  Tgh.p0, TfTgh.p0, TTfgh_TfTgh.chine] =
                  can (((g  (h  tab0 h)  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine)
                      ((g  (h  tab0 h)  Tgh.p0)  TfTgh.p0  TTfgh_TfTgh.chine)"
            proof -
              have "𝖺-1[g  (h  tab0 h)  Tgh.p0, TfTgh.p0, TTfgh_TfTgh.chine] =
                    𝖺-1[g  (h  tab0 h)  Tgh.p0, TfTgh.p0, TTfgh_TfTgh.chine]"
                using gh f0gh1.p0_simps canI_associator_0 𝖺'_def α_def by simp
              also have "... = can (((g  (h  tab0 h)  Tgh.p0)  TfTgh.p0)
                                         TTfgh_TfTgh.chine)
                                   ((g  (h  tab0 h)  Tgh.p0)  TfTgh.p0
                                         TTfgh_TfTgh.chine)"
                unfolding can_def
                using gh
                apply (intro E.eval_eqI) by simp_all
              finally show ?thesis by blast
            qed
            moreover
            have "𝖺-1[g, (h  tab0 h)  Tgh.p0, TfTgh.p0  TTfgh_TfTgh.chine] =
                  can ((g  (h  tab0 h)  Tgh.p0)  TfTgh.p0  TTfgh_TfTgh.chine)
                      (g  ((h  tab0 h)  Tgh.p0)  TfTgh.p0  TTfgh_TfTgh.chine)"
            proof -
              have "𝖺-1[g, (h  tab0 h)  Tgh.p0, TfTgh.p0  TTfgh_TfTgh.chine] =
                      𝖺-1[g, (h  tab0 h)  Tgh.p0, TfTgh.p0  TTfgh_TfTgh.chine]"
                using gh f0gh1.p0_simps canI_associator_0 𝖺'_def α_def by simp
              also have "... = can ((g  (h  tab0 h)  Tgh.p0)  TfTgh.p0
                                        TTfgh_TfTgh.chine)
                                   (g  ((h  tab0 h)  Tgh.p0)  TfTgh.p0
                                        TTfgh_TfTgh.chine)"
                unfolding can_def
                using gh
                apply (intro E.eval_eqI) by simp_all
              finally show ?thesis by blast
            qed
            ultimately show ?thesis
              using gh whisker_can_left_0 by simp
          qed
          moreover have "𝖺[g, tab0 g  Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine] 
                           𝖺[g  tab0 g  Tgh.p1, TfTgh.p0, TTfgh_TfTgh.chine] 
                           ((𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           𝖺-1[(g  tab0 g)  Tgh.p1, TfTgh.p0, TTfgh_TfTgh.chine] 
                           𝖺-1[g  tab0 g, Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine] 
                           𝖺-1[g, tab0 g, Tgh.p1  TfTgh.p0  TTfgh_TfTgh.chine] =
                         g  𝖺-1[tab0 g, Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine]"
          proof -
            have "𝖺[g, tab0 g  Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine] 
                           𝖺[g  tab0 g  Tgh.p1, TfTgh.p0, TTfgh_TfTgh.chine] 
                           ((𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                           𝖺-1[(g  tab0 g)  Tgh.p1, TfTgh.p0, TTfgh_TfTgh.chine] 
                           𝖺-1[g  tab0 g, Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine] 
                           𝖺-1[g, tab0 g, Tgh.p1  TfTgh.p0  TTfgh_TfTgh.chine] =
                  𝖺[g, tab0 g  Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine] 
                    𝖺[g  tab0 g  Tgh.p1, TfTgh.p0, TTfgh_TfTgh.chine] 
                    ((𝖺[g, tab0 g, Tgh.p1]  TfTgh.p0)  TTfgh_TfTgh.chine) 
                    𝖺-1[(g  tab0 g)  Tgh.p1, TfTgh.p0, TTfgh_TfTgh.chine] 
                    𝖺-1[g  tab0 g, Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine] 
                    𝖺-1[g, tab0 g, Tgh.p1  TfTgh.p0  TTfgh_TfTgh.chine]"
              using gh g0h1.p1_simps f0gh1.p0_simps 𝖺'_def α_def by simp
            also have "... = g  𝖺-1[tab0 g, Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine]"
              apply (intro E.eval_eqI) by simp_all
            also have "... = g  𝖺-1[tab0 g, Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine]"
              using gh g0h1.p1_simps f0gh1.p0_simps 𝖺'_def α_def by simp
            finally show ?thesis by simp
          qed
          ultimately show ?thesis
            using comp_assoc by presburger
        qed
        also have "... = (g 
                           (h  TTfgh_TfTgh.the_θ) 
                             (can (h  ((tab0 h  Tgh.p0)  TfTgh.p0)
                                      TTfgh_TfTgh.chine)
                                  (((h  tab0 h)  Tgh.p0)  TfTgh.p0
                                      TTfgh_TfTgh.chine)) 
                             ((h.tab  Tgh.p0)  TfTgh.p0  TTfgh_TfTgh.chine) 
                             (g0h1  TfTgh.p0  TTfgh_TfTgh.chine) 
                             𝖺-1[tab0 g, Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine]) 
                           𝖺[g, tab0 g, Tgh.p1  TfTgh.p0  TTfgh_TfTgh.chine] 
                           (g.tab  Tgh.p1  TfTgh.p0  TTfgh_TfTgh.chine) 
                           βg"
          using gh whisker_left by auto (* 11 sec *)
        also have "... = g.composite_cell wg' θg'  βg"
          unfolding wg'_def θg'_def
          using comp_assoc by presburger
        finally show ?thesis by blast
      qed
      have 6: "∃!γ. «γ : wg  wg'»  βg = tab1 g  γ  θg = θg'  (tab0 g  γ)"
        using wg wg' θg θg' βg eqg g.T2 [of wg wg' θg ug θg' βg] by blast
      obtain γg where γg: "«γg : wg  wg'»  βg = tab1 g  γg  θg = θg'  (tab0 g  γg)"
        using 6 by auto
      show "⟦⟦Tgh.p1  TfTgh.p0  TTfgh_TfTgh.chine⟧⟧ = ⟦⟦Tfg.p0  TTfgh.p1⟧⟧"
      proof -
        have "iso γg"
          using γg BS3 wg_is_map wg'_is_map by blast
        hence "isomorphic wg wg'"
          using γg isomorphic_def isomorphic_symmetric by auto
        thus ?thesis
          using wg wg' wg_def wg'_def Maps.CLS_eqI by auto
      qed

      text ‹Now the last equation: similar, but somewhat simpler.›
      define uh where "uh = tab0 h  TTfgh.p0"
      define wh where "wh = TTfgh.p0"
      define wh' where "wh' = Tgh.p0  TfTgh.p0  TTfgh_TfTgh.chine"
      define θh
      where "θh = tab0 h  TTfgh.p0"
      define θh'
      where "θh' = TTfgh_TfTgh.the_θ  𝖺-1[tab0 h  Tgh.p0, TfTgh.p0, TTfgh_TfTgh.chine] 
                 𝖺-1[tab0 h, Tgh.p0, TfTgh.p0  TTfgh_TfTgh.chine]"
      define βh
      where "βh = 𝖺[tab1 h, Tgh.p0, TfTgh.p0  TTfgh_TfTgh.chine] 
                   (g0h1  TfTgh.p0  TTfgh_TfTgh.chine) 
                   𝖺-1[tab0 g, Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine]  (tab0 g  γg) 
                   𝖺[tab0 g, Tfg.p0, TTfgh.p1]  inv fg0h1"
      have uh: "ide uh"
        unfolding uh_def by simp
      have wh: "ide wh"
        unfolding wh_def by simp
      have wh_is_map: "is_left_adjoint wh"
        unfolding wh_def by simp
      have wh': "ide wh'"
        unfolding wh'_def by simp
      have wh'_is_map: "is_left_adjoint wh'"
        unfolding wh'_def
        using g0h1.p0_simps f0gh1.p0_simps TTfgh_TfTgh.is_map left_adjoints_compose by simp
      have θh: "«θh : tab0 h  wh  uh»"
        unfolding θh_def wh_def uh_def by auto
      have θh': "«θh' : tab0 h  wh'  uh»"
        unfolding θh'_def wh'_def uh_def
        using g0h1.p0_simps f0gh1.p0_simps
        by (intro comp_in_homI) auto
      have βh: "«βh : tab1 h  wh  tab1 h  wh'»"
      proof (unfold βh_def wh_def wh'_def, intro comp_in_homI)
        (* auto can solve this, but it's too slow *)
        show "«inv fg0h1 : tab1 h  TTfgh.p0  (tab0 g  Tfg.p0)  TTfgh.p1»"
          using fg0h1.φ_uniqueness by blast
        show "«𝖺[tab0 g, Tfg.p0, TTfgh.p1] :
                 (tab0 g  Tfg.p0)  TTfgh.p1  tab0 g  Tfg.p0  TTfgh.p1»"
          using fg0h1.p1_simps by auto
        show "«tab0 g  γg :
                 tab0 g  Tfg.p0  TTfgh.p1  tab0 g  Tgh.p1  TfTgh.p0  TTfgh_TfTgh.chine»"
          using γg wg_def wg'_def fg0h1.p1_simps by auto
        show "«𝖺-1[tab0 g, Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine] :
                 tab0 g  Tgh.p1  TfTgh.p0  TTfgh_TfTgh.chine
                     (tab0 g  Tgh.p1)  TfTgh.p0  TTfgh_TfTgh.chine»"
          using fg0h1.p1_simps by auto
        show "«g0h1  TfTgh.p0  TTfgh_TfTgh.chine :
                (tab0 g  Tgh.p1)  TfTgh.p0  TTfgh_TfTgh.chine
                    (tab1 h  Tgh.p0)  TfTgh.p0  TTfgh_TfTgh.chine»"
          using fg0h1.p1_simps by force
        show "«𝖺[tab1 h, Tgh.p0, TfTgh.p0  TTfgh_TfTgh.chine] :
                (tab1 h  Tgh.p0)  TfTgh.p0  TTfgh_TfTgh.chine
                    tab1 h  Tgh.p0  TfTgh.p0  TTfgh_TfTgh.chine»"
          using fg0h1.p1_simps by auto
      qed
      have eqh: "h.composite_cell wh θh = h.composite_cell wh' θh'  βh"
      proof -
        text ‹
          Once again, the strategy is to form the subexpression
          \[
             𝖺[h, tab0 h, TTfgh.p0] ⋅ (h.tab ⋆ TTfgh.p0) ⋅ fg0h1.φ ⋅ 𝖺-1[tab0 g, Tfg.p0, TTfgh.p1]›
          \]
          which is equal to θg, so that we can make use of the equation θg = θg' ⋅ (tab0 g ⋆ γg)›.
        ›
        have "h.composite_cell wh θh =
              (h  tab0 h  TTfgh.p0)  𝖺[h, tab0 h, TTfgh.p0]  (h.tab  TTfgh.p0)"
          unfolding wh_def θh_def by simp
        also have "... = 𝖺[h, tab0 h, TTfgh.p0]  (h.tab  TTfgh.p0)"
        proof -
          have "(h  tab0 h  TTfgh.p0)  𝖺[h, tab0 h, TTfgh.p0] = 𝖺[h, tab0 h, TTfgh.p0]"
            using comp_cod_arr by simp
          thus ?thesis
            using comp_assoc by metis
        qed
        also have "... = (𝖺[h, tab0 h, TTfgh.p0]  (h.tab  TTfgh.p0) 
                           fg0h1  𝖺-1[tab0 g, Tfg.p0, TTfgh.p1]) 
                           𝖺[tab0 g, Tfg.p0, TTfgh.p1]  inv fg0h1"
        proof -
          have "(h.tab  TTfgh.p0)  fg0h1  (𝖺-1[tab0 g, Tfg.p0, TTfgh.p1] 
                           𝖺[tab0 g, Tfg.p0, TTfgh.p1])  inv fg0h1 =
                (h.tab  TTfgh.p0)  fg0h1  ((tab0 g  Tfg.p0)  TTfgh.p1)  inv fg0h1"
           using fg0h1.p1_simps comp_assoc_assoc' by simp
          also have "... = (h.tab  TTfgh.p0)  fg0h1  inv fg0h1"
           using fg0h1.p1_simps fg0h1.φ_uniqueness comp_cod_arr by simp
          also have "... = (h.tab  TTfgh.p0)  (tab1 h  TTfgh.p0)"
           using comp_arr_inv' fg0h1.φ_uniqueness by simp
          also have "... = h.tab  TTfgh.p0"
           using comp_arr_dom fg0h1.p0_simps by simp
          finally have "(h.tab  TTfgh.p0)  fg0h1  (𝖺-1[tab0 g, Tfg.p0, TTfgh.p1] 
                           𝖺[tab0 g, Tfg.p0, TTfgh.p1])  inv fg0h1 =
                        h.tab  TTfgh.p0"
            by simp
          thus ?thesis
            using comp_assoc by simp
        qed
        also have "... = θg  𝖺[tab0 g, Tfg.p0, TTfgh.p1]  inv fg0h1"
          unfolding θg_def by simp
        also have "... = (θg'  (tab0 g  γg))  𝖺[tab0 g, Tfg.p0, TTfgh.p1]  inv fg0h1"
          using γg by simp
        also have "... = (h  TTfgh_TfTgh.the_θ) 
                           can (h  ((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine)
                               (((h  tab0 h)  Tgh.p0)  TfTgh.p0  TTfgh_TfTgh.chine) 
                           ((h.tab  Tgh.p0)  TfTgh.p0  TTfgh_TfTgh.chine) 
                           (g0h1  TfTgh.p0  TTfgh_TfTgh.chine) 
                           𝖺-1[tab0 g, Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine] 
                           (tab0 g  γg) 
                           𝖺[tab0 g, Tfg.p0, TTfgh.p1] 
                           inv fg0h1"
          unfolding θg'_def
          using comp_assoc by presburger
        also have "... = (h  TTfgh_TfTgh.the_θ) 
                           can (h  ((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine)
                               (((h  tab0 h)  Tgh.p0)  TfTgh.p0  TTfgh_TfTgh.chine) 
                           ((𝖺-1[h  tab0 h, Tgh.p0, TfTgh.p0  TTfgh_TfTgh.chine] 
                           𝖺[h  tab0 h, Tgh.p0, TfTgh.p0  TTfgh_TfTgh.chine]) 
                           ((h.tab  Tgh.p0)  TfTgh.p0  TTfgh_TfTgh.chine)) 
                           (g0h1  TfTgh.p0  TTfgh_TfTgh.chine) 
                           𝖺-1[tab0 g, Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine] 
                           (tab0 g  γg) 
                           𝖺[tab0 g, Tfg.p0, TTfgh.p1] 
                           inv fg0h1"
        proof -
          have "(𝖺-1[h  tab0 h, Tgh.p0, TfTgh.p0  TTfgh_TfTgh.chine] 
                           𝖺[h  tab0 h, Tgh.p0, TfTgh.p0  TTfgh_TfTgh.chine]) 
                           ((h.tab  Tgh.p0)  TfTgh.p0  TTfgh_TfTgh.chine) =
                 (h.tab  Tgh.p0)  TfTgh.p0  TTfgh_TfTgh.chine"
            using comp_cod_arr comp_assoc_assoc' by simp
          thus ?thesis
            using comp_assoc by simp
        qed
        also have "... = (h  TTfgh_TfTgh.the_θ) 
                           can (h  ((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine)
                               (((h  tab0 h)  Tgh.p0)  TfTgh.p0  TTfgh_TfTgh.chine) 
                           𝖺-1[h  tab0 h, Tgh.p0, TfTgh.p0  TTfgh_TfTgh.chine] 
                           (𝖺[h  tab0 h, Tgh.p0, TfTgh.p0  TTfgh_TfTgh.chine] 
                           ((h.tab  Tgh.p0)  TfTgh.p0  TTfgh_TfTgh.chine)) 
                           (g0h1  TfTgh.p0  TTfgh_TfTgh.chine) 
                           𝖺-1[tab0 g, Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine] 
                           (tab0 g  γg) 
                           𝖺[tab0 g, Tfg.p0, TTfgh.p1] 
                           inv fg0h1"
          using comp_assoc by presburger
        also have "... = (h  TTfgh_TfTgh.the_θ) 
                           can (h  ((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine)
                                 (((h  tab0 h)  Tgh.p0)  TfTgh.p0  TTfgh_TfTgh.chine) 
                           𝖺-1[h  tab0 h, Tgh.p0, TfTgh.p0  TTfgh_TfTgh.chine] 
                           (h.tab  Tgh.p0  TfTgh.p0  TTfgh_TfTgh.chine) 
                           𝖺[tab1 h, Tgh.p0, TfTgh.p0  TTfgh_TfTgh.chine] 
                           (g0h1  TfTgh.p0  TTfgh_TfTgh.chine) 
                           𝖺-1[tab0 g, Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine] 
                           (tab0 g  γg) 
                           𝖺[tab0 g, Tfg.p0, TTfgh.p1] 
                           inv fg0h1"
          using assoc_naturality [of h.tab Tgh.p0 "TfTgh.p0  TTfgh_TfTgh.chine"] comp_assoc
          by simp
        also have "... = (h  TTfgh_TfTgh.the_θ) 
                           can (h  ((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine)
                                 (((h  tab0 h)  Tgh.p0)  TfTgh.p0  TTfgh_TfTgh.chine) 
                           𝖺-1[h  tab0 h, Tgh.p0, TfTgh.p0  TTfgh_TfTgh.chine] 
                           ((𝖺-1[h, tab0 h, Tgh.p0  TfTgh.p0  TTfgh_TfTgh.chine] 
                           𝖺[h, tab0 h, Tgh.p0  TfTgh.p0  TTfgh_TfTgh.chine]) 
                           (h.tab  Tgh.p0  TfTgh.p0  TTfgh_TfTgh.chine)) 
                           𝖺[tab1 h, Tgh.p0, TfTgh.p0  TTfgh_TfTgh.chine] 
                           (g0h1  TfTgh.p0  TTfgh_TfTgh.chine) 
                           𝖺-1[tab0 g, Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine] 
                           (tab0 g  γg) 
                           𝖺[tab0 g, Tfg.p0, TTfgh.p1] 
                           inv fg0h1"
        proof -
          have "(𝖺-1[h, tab0 h, Tgh.p0  TfTgh.p0  TTfgh_TfTgh.chine] 
                           𝖺[h, tab0 h, Tgh.p0  TfTgh.p0  TTfgh_TfTgh.chine]) 
                           (h.tab  Tgh.p0  TfTgh.p0  TTfgh_TfTgh.chine) =
                h.tab  Tgh.p0  TfTgh.p0  TTfgh_TfTgh.chine"
            using comp_cod_arr comp_assoc_assoc' by simp
          thus ?thesis
            using comp_assoc by simp
        qed
        also have "... = (h  TTfgh_TfTgh.the_θ) 
                           (can (h  ((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine)
                                 (((h  tab0 h)  Tgh.p0)  TfTgh.p0  TTfgh_TfTgh.chine) 
                           (𝖺-1[h  tab0 h, Tgh.p0, TfTgh.p0  TTfgh_TfTgh.chine] 
                           𝖺-1[h, tab0 h, Tgh.p0  TfTgh.p0  TTfgh_TfTgh.chine])) 
                           𝖺[h, tab0 h, Tgh.p0  TfTgh.p0  TTfgh_TfTgh.chine] 
                           (h.tab  Tgh.p0  TfTgh.p0  TTfgh_TfTgh.chine) 
                           𝖺[tab1 h, Tgh.p0, TfTgh.p0  TTfgh_TfTgh.chine] 
                           (g0h1  TfTgh.p0  TTfgh_TfTgh.chine) 
                           𝖺-1[tab0 g, Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine] 
                           (tab0 g  γg) 
                           𝖺[tab0 g, Tfg.p0, TTfgh.p1] 
                           inv fg0h1"
          using comp_assoc by presburger
        also have "... = ((h  TTfgh_TfTgh.the_θ) 
                           (h  can (((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine)
                                    (tab0 h  Tgh.p0  TfTgh.p0  TTfgh_TfTgh.chine))) 
                           𝖺[h, tab0 h, Tgh.p0  TfTgh.p0  TTfgh_TfTgh.chine] 
                           (h.tab  Tgh.p0  TfTgh.p0  TTfgh_TfTgh.chine) 
                           𝖺[tab1 h, Tgh.p0, TfTgh.p0  TTfgh_TfTgh.chine] 
                           (g0h1  TfTgh.p0  TTfgh_TfTgh.chine) 
                           𝖺-1[tab0 g, Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine] 
                           (tab0 g  γg) 
                           𝖺[tab0 g, Tfg.p0, TTfgh.p1] 
                           inv fg0h1"
        proof -
          have "can (h  ((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine)
                    (((h  tab0 h)  Tgh.p0)  TfTgh.p0  TTfgh_TfTgh.chine) 
                  (𝖺-1[h  tab0 h, Tgh.p0, TfTgh.p0  TTfgh_TfTgh.chine] 
                  𝖺-1[h, tab0 h, Tgh.p0  TfTgh.p0  TTfgh_TfTgh.chine]) =
                can (h  ((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine)
                    (((h  tab0 h)  Tgh.p0)  TfTgh.p0  TTfgh_TfTgh.chine) 
                  𝖺-1[h  tab0 h, Tgh.p0, TfTgh.p0  TTfgh_TfTgh.chine] 
                   𝖺-1[h, tab0 h, Tgh.p0  TfTgh.p0  TTfgh_TfTgh.chine]"
            using 𝖺'_def α_def by simp
          also have "... = can (h  ((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine)
                               (((h  tab0 h)  Tgh.p0)  TfTgh.p0  TTfgh_TfTgh.chine) 
                           can (((h  tab0 h)  Tgh.p0)  TfTgh.p0  TTfgh_TfTgh.chine)
                               (h  tab0 h  Tgh.p0  TfTgh.p0  TTfgh_TfTgh.chine)"
          proof -
            have "𝖺-1[h  tab0 h, Tgh.p0, TfTgh.p0  TTfgh_TfTgh.chine] 
                   𝖺-1[h, tab0 h, Tgh.p0  TfTgh.p0  TTfgh_TfTgh.chine] =
                  can (((h  tab0 h)  Tgh.p0)  TfTgh.p0  TTfgh_TfTgh.chine)
                               (h  tab0 h  Tgh.p0  TfTgh.p0  TTfgh_TfTgh.chine)"
              unfolding can_def
              apply (intro E.eval_eqI) by simp_all
            thus ?thesis by simp
          qed
          also have "... = can (h  ((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine)
                               (h  tab0 h  Tgh.p0  TfTgh.p0  TTfgh_TfTgh.chine)"
            by simp
          also have "... = h  can (((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine)
                                   (tab0 h  Tgh.p0  TfTgh.p0  TTfgh_TfTgh.chine)"
            using whisker_can_left_0 by simp
          finally have "can (h  ((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine)
                            (((h  tab0 h)  Tgh.p0)  TfTgh.p0  TTfgh_TfTgh.chine) 
                          (𝖺-1[h  tab0 h, Tgh.p0, TfTgh.p0  TTfgh_TfTgh.chine] 
                          𝖺-1[h, tab0 h, Tgh.p0  TfTgh.p0  TTfgh_TfTgh.chine]) =
                        h  can (((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine)
                                (tab0 h  Tgh.p0  TfTgh.p0  TTfgh_TfTgh.chine)"
            by simp
          thus ?thesis
            using comp_assoc by presburger
        qed
        also have "... = (h  TTfgh_TfTgh.the_θ 
                              can (((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine)
                                  (tab0 h  Tgh.p0  TfTgh.p0  TTfgh_TfTgh.chine)) 
                           𝖺[h, tab0 h, Tgh.p0  TfTgh.p0  TTfgh_TfTgh.chine] 
                           (h.tab  Tgh.p0  TfTgh.p0  TTfgh_TfTgh.chine) 
                           𝖺[tab1 h, Tgh.p0, TfTgh.p0  TTfgh_TfTgh.chine] 
                           (g0h1  TfTgh.p0  TTfgh_TfTgh.chine) 
                           𝖺-1[tab0 g, Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine] 
                           (tab0 g  γg) 
                           𝖺[tab0 g, Tfg.p0, TTfgh.p1] 
                           inv fg0h1"
          using whisker_left [of h] comp_assoc by simp
        also have "... = (h  TTfgh_TfTgh.the_θ 
                              𝖺-1[tab0 h  Tgh.p0, TfTgh.p0, TTfgh_TfTgh.chine] 
                              𝖺-1[tab0 h, Tgh.p0, TfTgh.p0  TTfgh_TfTgh.chine]) 
                           𝖺[h, tab0 h, Tgh.p0  TfTgh.p0  TTfgh_TfTgh.chine] 
                           (h.tab  Tgh.p0  TfTgh.p0  TTfgh_TfTgh.chine) 
                           𝖺[tab1 h, Tgh.p0, TfTgh.p0  TTfgh_TfTgh.chine] 
                           (g0h1  TfTgh.p0  TTfgh_TfTgh.chine) 
                           𝖺-1[tab0 g, Tgh.p1, TfTgh.p0  TTfgh_TfTgh.chine] 
                           (tab0 g  γg) 
                           𝖺[tab0 g, Tfg.p0, TTfgh.p1] 
                           inv fg0h1"
        proof -
          have "can (((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine)
                    (tab0 h  Tgh.p0  TfTgh.p0  TTfgh_TfTgh.chine) =
                𝖺-1[tab0 h  Tgh.p0, TfTgh.p0, TTfgh_TfTgh.chine] 
                 𝖺-1[tab0 h, Tgh.p0, TfTgh.p0  TTfgh_TfTgh.chine]"
            unfolding can_def
            apply (intro E.eval_eqI) by auto
          also have "... = 𝖺-1[tab0 h  Tgh.p0, TfTgh.p0, TTfgh_TfTgh.chine] 
                             𝖺-1[tab0 h, Tgh.p0, TfTgh.p0  TTfgh_TfTgh.chine]"
            using 𝖺'_def α_def by simp
          finally have "can (((tab0 h  Tgh.p0)  TfTgh.p0)  TTfgh_TfTgh.chine)
                            (tab0 h  Tgh.p0  TfTgh.p0  TTfgh_TfTgh.chine) =
                        𝖺-1[tab0 h  Tgh.p0, TfTgh.p0, TTfgh_TfTgh.chine] 
                          𝖺-1[tab0 h, Tgh.p0, TfTgh.p0  TTfgh_TfTgh.chine]"
            by simp
          thus ?thesis by simp
        qed
        also have "... = h.composite_cell wh' θh'  βh"
          unfolding wh'_def θh'_def βh_def
          using comp_assoc by presburger
        finally show ?thesis by simp
      qed
      have 7: "∃!γ. «γ : wh  wh'»  βh = tab1 h  γ  θh = θh'  (tab0 h  γ)"
        using wh wh' θh θh' βh eqh h.T2 [of wh wh' θh uh θh' βh] by blast
      obtain γh where γh: "«γh : wh  wh'»  βh = tab1 h  γh  θh = θh'  (tab0 h  γh)"
        using 7 by auto
      show "⟦⟦Tgh.p0  TfTgh.p0  TTfgh_TfTgh.chine⟧⟧ = ⟦⟦TTfgh.p0⟧⟧"
      proof -
        have "iso γh"
          using γh BS3 wh_is_map wh'_is_map by blast
        hence "isomorphic wh wh'"
          using γh isomorphic_def isomorphic_symmetric by auto
        thus ?thesis
          using wh wh' wh_def wh'_def Maps.CLS_eqI [of wh wh'] by simp
      qed
    qed

    text ‹
      Finally, we can show that @{term TTfgh_TfTgh.chine} is given by tupling.
    ›

    lemma CLS_chine:
    shows "⟦⟦TTfgh_TfTgh.chine⟧⟧ = tuple_ABC"
    proof -
      have "tuple_ABC = SPN_fgh.chine_assoc"
        using SPN_fgh.chine_assoc_def by simp
      also have "... = ⟦⟦TTfgh_TfTgh.chine⟧⟧"
      proof (intro Maps.arr_eqI)
        show "Maps.arr SPN_fgh.chine_assoc"
          using SPN_fgh.chine_assoc_in_hom by auto
        show "Maps.arr ⟦⟦TTfgh_TfTgh.chine⟧⟧"
          using Maps.CLS_in_hom TTfgh_TfTgh.is_map by blast
        show "Maps.Dom SPN_fgh.chine_assoc = Maps.Dom ⟦⟦TTfgh_TfTgh.chine⟧⟧"
          using SPN_fgh.chine_assoc_def Maps.dom_char tuple_ABC_in_hom TTfgh_TfTgh.chine_in_hom
          by fastforce
        show "Maps.Cod SPN_fgh.chine_assoc = Maps.Cod ⟦⟦TTfgh_TfTgh.chine⟧⟧"
        proof - 
          have "Maps.Cod SPN_fgh.chine_assoc = Maps.Cod tuple_ABC"
            using SPN_fgh.chine_assoc_def by simp
          also have "... = src (prj0 (tab1 g  prj1 (tab1 h) (tab0 g)) (tab0 f))"
            by (metis (lifting) Maps.Dom.simps(1) Maps.seq_char SPN_fgh.prj_chine_assoc(1)
                SPN_fgh.prj_simps(1) calculation f0gh1.leg1_simps(3) prj_char(4))
          also have "... = Maps.Cod ⟦⟦TTfgh_TfTgh.chine⟧⟧"
            using Maps.cod_char TTfgh_TfTgh.chine_in_hom by simp
          finally show ?thesis by blast
        qed
        show "Maps.Map SPN_fgh.chine_assoc = Maps.Map ⟦⟦TTfgh_TfTgh.chine⟧⟧"
        proof -
          have 0: "Chn (Span.hcomp (SPN f) (Span.hcomp (SPN g) (SPN h))) =
                   Maps.MkIde (src TfTgh.p0)"
            using fg gh
            by (metis (mono_tags, lifting) Maps.in_homE Maps.seqE SPN_fgh.prj_chine_assoc(1)
                SPN_fgh.prj_simps(1) SPN_fgh.prj_simps(13) calculation tuple_ABC_in_hom)
          have "tuple_ABC = ⟦⟦TTfgh_TfTgh.chine⟧⟧"
          proof (intro Maps.prj_joint_monic
                         [of SPN_fgh.μ.leg0 "SPN_fgh.ν.leg1  SPN_fgh.νπ.prj1"
                             tuple_ABC "⟦⟦TTfgh_TfTgh.chine⟧⟧"])
            show "Maps.cospan SPN_fgh.μ.leg0 (SPN_fgh.ν.leg1  SPN_fgh.νπ.prj1)"
              using SPN_fgh.νπ.dom.is_span SPN_fgh.νπ.leg1_composite SPN_fgh.cospan_μν
              by auto
            show "Maps.seq SPN_fgh.Prj1 tuple_ABC"
              using 0 tuple_ABC_in_hom SPN_fgh.prj_in_hom(4) by auto
            show "Maps.seq SPN_fgh.Prj1 ⟦⟦TTfgh_TfTgh.chine⟧⟧"
            proof
              show "Maps.in_hom ⟦⟦TTfgh_TfTgh.chine⟧⟧ ⟦⟦src TTfgh_TfTgh.chine⟧⟧
                                ⟦⟦trg TTfgh_TfTgh.chine⟧⟧"
                using fg gh TTfgh_TfTgh.chine_in_hom Maps.CLS_in_hom TTfgh_TfTgh.is_map
                by blast
              show "Maps.in_hom SPN_fgh.Prj1 ⟦⟦trg TTfgh_TfTgh.chine⟧⟧ SPN_fgh.μ.apex"
              proof
                show "Maps.cospan SPN_fgh.μ.leg0 (SPN_fgh.ν.leg1  SPN_fgh.νπ.prj1)"
                  using SPN_fgh.prj_in_hom(4) by blast
                show "⟦⟦trg TTfgh_TfTgh.chine⟧⟧ =
                      Maps.pbdom SPN_fgh.μ.leg0 (SPN_fgh.ν.leg1  SPN_fgh.νπ.prj1)"
                proof -
                  have "⟦⟦trg TTfgh_TfTgh.chine⟧⟧ = Maps.MkIde (src TfTgh.p0)"
                    by simp
                  also have "... = Maps.pbdom SPN_fgh.μ.leg0 (SPN_fgh.ν.leg1  SPN_fgh.νπ.prj1)"
                    using 0 Maps.pbdom_def SPN_fgh.chine_composite(2) by presburger
                  finally show ?thesis by blast
                qed
                show "SPN_fgh.μ.apex = Maps.dom SPN_fgh.μ.leg0"
                  using SPN_f.dom.apex_def by blast
              qed
            qed
            have 2: "Maps.commutative_square SPN_fgh.ν.leg0 SPN_fgh.π.leg1
                                             SPN_fgh.Prj01 SPN_fgh.Prj0"
            proof
              show "Maps.cospan SPN_fgh.ν.leg0 SPN_fgh.π.leg1"
                using SPN_fgh.νπ.legs_form_cospan(1) by simp
              show "Maps.span SPN_fgh.Prj01 SPN_fgh.Prj0"
                using SPN_fgh.prj_simps(2-3,5-6) by presburger
              show "Maps.dom SPN_fgh.ν.leg0 = Maps.cod SPN_fgh.Prj01"
                using SPN_fgh.prj_simps(8) SPN_g.dom.is_span SPN_g.dom.leg_simps(2)
                by auto
              show "SPN_fgh.ν.leg0  SPN_fgh.Prj01 = SPN_fgh.π.leg1  SPN_fgh.Prj0"
                by (metis (no_types, lifting) Maps.cod_comp Maps.comp_assoc
                    Maps.pullback_commutes' SPN_fgh.μν.dom.leg_simps(1)
                    SPN_fgh.μν.leg0_composite SPN_fgh.cospan_νπ)
            qed
            have 1: "Maps.commutative_square
                       SPN_fgh.μ.leg0 (Maps.comp SPN_fgh.ν.leg1 SPN_fgh.νπ.prj1)
                       SPN_fgh.Prj11 tuple_BC"
            proof
              show "Maps.cospan SPN_fgh.μ.leg0 (Maps.comp SPN_fgh.ν.leg1 SPN_fgh.νπ.prj1)"
                using fg gh SPN_fgh.prj_simps(10) by blast
              show "Maps.span SPN_fgh.Prj11 tuple_BC"
                using fg gh csq(2) by blast
              show "Maps.dom SPN_fgh.μ.leg0 = Maps.cod SPN_fgh.Prj11"
                using fg gh SPN_f.dom.leg_simps(2) SPN_fgh.prj_simps(7) by auto
              show "SPN_fgh.μ.leg0  SPN_fgh.Prj11 =
                    (SPN_fgh.ν.leg1  SPN_fgh.νπ.prj1)  tuple_BC"
                using 2 fg gh Maps.comp_assoc csq(2)
                      Maps.prj_tuple [of SPN_fgh.ν.leg0 SPN_fgh.π.leg1 SPN_fgh.Prj01 SPN_fgh.Prj0]
                by blast
            qed
            show "SPN_fgh.Prj1  tuple_ABC = SPN_fgh.Prj1  Maps.CLS TTfgh_TfTgh.chine"
            proof -
              have "SPN_fgh.Prj1  tuple_ABC = SPN_fgh.Prj11"
                using csq(2) by simp
              also have "... = ⟦⟦Tfg.p1  TTfgh.p1⟧⟧"
                using prj_char by simp
              also have "... = ⟦⟦TfTgh.p1  TTfgh_TfTgh.chine⟧⟧"
                using prj_chine(1) by simp
              also have "... = ⟦⟦TfTgh.p1⟧⟧  ⟦⟦TTfgh_TfTgh.chine⟧⟧"
              proof -
                have "is_left_adjoint TfTgh.p1"
                  by (simp add: fg)
                moreover have "is_left_adjoint TTfgh_TfTgh.chine"
                  using TTfgh_TfTgh.is_map by simp
                moreover have "TfTgh.p1  TTfgh_TfTgh.chine  TfTgh.p1  TTfgh_TfTgh.chine"
                  using fg gh isomorphic_reflexive by simp
                ultimately show ?thesis
                  using Maps.comp_CLS
                          [of TfTgh.p1 TTfgh_TfTgh.chine "TfTgh.p1  TTfgh_TfTgh.chine"]
                  by simp
              qed
              also have "... = SPN_fgh.Prj1  Maps.CLS TTfgh_TfTgh.chine"
                using prj_char by simp
              finally show ?thesis by blast
            qed
            show
              "Maps.PRJ0 SPN_fgh.μ.leg0 (SPN_fgh.ν.leg1  SPN_fgh.νπ.prj1)  tuple_ABC =
               Maps.PRJ0 SPN_fgh.μ.leg0 (SPN_fgh.ν.leg1  SPN_fgh.νπ.prj1) 
                 ⟦⟦TTfgh_TfTgh.chine⟧⟧"
            proof -
              have
                "Maps.PRJ0 SPN_fgh.μ.leg0 (SPN_fgh.ν.leg1  SPN_fgh.νπ.prj1)  tuple_ABC =
                 tuple_BC"
                using csq(2)
                      Maps.prj_tuple [of SPN_fgh.μ.leg0 "SPN_fgh.ν.leg1  SPN_fgh.νπ.prj1"
                                         SPN_fgh.Prj11 tuple_BC]
                by simp
              also have "... =
                         Maps.comp
                           (Maps.PRJ0 SPN_fgh.μ.leg0 (Maps.comp SPN_fgh.ν.leg1 SPN_fgh.νπ.prj1))
                           ⟦⟦TTfgh_TfTgh.chine⟧⟧"
              proof (intro Maps.prj_joint_monic
                             [of SPN_fgh.ν.leg0 SPN_fgh.π.leg1 tuple_BC
                                 "Maps.PRJ0 SPN_fgh.μ.leg0 (SPN_fgh.ν.leg1  SPN_fgh.νπ.prj1) 
                                    ⟦⟦TTfgh_TfTgh.chine⟧⟧"])
                show "Maps.cospan SPN_fgh.ν.leg0 SPN_fgh.π.leg1"
                  using SPN_fgh.νπ.legs_form_cospan(1) by simp
                show "Maps.seq SPN_fgh.νπ.prj1 tuple_BC"
                proof
                  show "Maps.in_hom tuple_BC
                                    (Maps.MkIde (src TTfgh.p0)) (Maps.MkIde (src Tgh.p0))"
                    using tuple_BC_in_hom by simp
                  show "Maps.in_hom SPN_fgh.νπ.prj1 (Maps.MkIde (src Tgh.p0)) SPN_fgh.ν.apex"
                  proof -
                    have "Maps.pbdom SPN_fgh.ν.leg0 SPN_fgh.π.leg1 = Maps.MkIde (src Tgh.p0)"
                      using fg gh Maps.pbdom_def
                      by (metis (no_types, lifting) SPN.preserves_ide SPN_fgh.νπ.are_identities(2)
                          SPN_fgh.νπ.composable Span.chine_hcomp_ide_ide Tgh.chine_hcomp_SPN_SPN
                          g.is_ide)
                    thus ?thesis
                      using SPN_fgh.νπ.prj_in_hom(1) by simp
                  qed
                qed
                show "Maps.seq SPN_fgh.νπ.prj1
                               (Maps.PRJ0 SPN_fgh.μ.leg0 (SPN_fgh.ν.leg1  SPN_fgh.νπ.prj1) 
                                  ⟦⟦TTfgh_TfTgh.chine⟧⟧)"
                proof
                  show "Maps.in_hom SPN_fgh.νπ.prj1
                                    (Maps.pbdom SPN_fgh.ν.leg0 SPN_fgh.π.leg1) SPN_fgh.ν.apex"
                    using SPN_fgh.νπ.prj_in_hom(1) by simp
                  show "Maps.in_hom
                         (Maps.PRJ0 SPN_fgh.μ.leg0 (SPN_fgh.ν.leg1  SPN_fgh.νπ.prj1) 
                             ⟦⟦TTfgh_TfTgh.chine⟧⟧)
                         ⟦⟦src TTfgh_TfTgh.chine⟧⟧
                         (Maps.pbdom SPN_fgh.ν.leg0 SPN_fgh.π.leg1)"
                  proof
                    show "Maps.in_hom ⟦⟦TTfgh_TfTgh.chine⟧⟧ ⟦⟦src TTfgh_TfTgh.chine⟧⟧
                                      ⟦⟦trg TTfgh_TfTgh.chine⟧⟧"
                      using fg gh TTfgh_TfTgh.chine_in_hom Maps.CLS_in_hom TTfgh_TfTgh.is_map
                      by blast
                    show "Maps.in_hom
                           (Maps.PRJ0 SPN_fgh.μ.leg0 (SPN_fgh.ν.leg1  SPN_fgh.νπ.prj1))
                           ⟦⟦trg TTfgh_TfTgh.chine⟧⟧
                           (Maps.pbdom SPN_fgh.ν.leg0 SPN_fgh.π.leg1)"
                    proof
                      show "Maps.cospan SPN_fgh.μ.leg0 (SPN_fgh.ν.leg1  SPN_fgh.νπ.prj1)"
                        using SPN_fgh.prj_in_hom(4) by blast
                      show "⟦⟦trg TTfgh_TfTgh.chine⟧⟧ =
                            Maps.pbdom SPN_fgh.μ.leg0 (SPN_fgh.ν.leg1  SPN_fgh.νπ.prj1)"
                      proof -
                        have "⟦⟦trg TTfgh_TfTgh.chine⟧⟧ = Maps.MkIde (src TfTgh.p0)"
                          by simp
                        also have "... = Maps.pbdom SPN_fgh.μ.leg0
                                                    (SPN_fgh.ν.leg1  SPN_fgh.νπ.prj1)"
                          using 0 Maps.pbdom_def SPN_fgh.chine_composite(2) by presburger
                        finally show ?thesis by blast
                      qed
                      show "Maps.pbdom SPN_fgh.ν.leg0 SPN_fgh.π.leg1 =
                            Maps.dom (SPN_fgh.ν.leg1  SPN_fgh.νπ.prj1)"
                        using fg gh Maps.pbdom_def SPN_fgh.νπ.apex_composite
                              SPN_fgh.νπ.dom.apex_def SPN_fgh.νπ.dom.is_span
                              SPN_fgh.νπ.leg1_composite
                        by presburger
                    qed
                  qed
                qed
                show "SPN_fgh.νπ.prj0  tuple_BC =
                      SPN_fgh.νπ.prj0 
                        Maps.PRJ0 SPN_fgh.μ.leg0 (SPN_fgh.ν.leg1  SPN_fgh.νπ.prj1) 
                        ⟦⟦TTfgh_TfTgh.chine⟧⟧"
                proof - 
                  have "SPN_fgh.νπ.prj0  tuple_BC = SPN_fgh.Prj0"
                    using csq(1) by simp
                  also have "... = SPN_fgh.νπ.prj0 
                                     Maps.PRJ0 SPN_fgh.μ.leg0 (SPN_fgh.ν.leg1  SPN_fgh.νπ.prj1) 
                                     ⟦⟦TTfgh_TfTgh.chine⟧⟧"
                  proof -
                    have "SPN_fgh.νπ.prj0 
                            Maps.PRJ0 SPN_fgh.μ.leg0 (SPN_fgh.ν.leg1  SPN_fgh.νπ.prj1) 
                            ⟦⟦TTfgh_TfTgh.chine⟧⟧ =
                          ⟦⟦Tgh.p0⟧⟧  ⟦⟦TfTgh.p0⟧⟧  ⟦⟦TTfgh_TfTgh.chine⟧⟧"
                      using fg gh Tgh.ρσ.prj_char TfTgh.prj_char(1) isomorphic_reflexive
                            Maps.comp_CLS [of "tab1 g" "prj1 (tab1 h) (tab0 g)" "tab1 g  Tgh.p1"]
                      by simp
                    also have "... = ⟦⟦Tgh.p0⟧⟧  ⟦⟦TfTgh.p0  TTfgh_TfTgh.chine⟧⟧"
                      using fg gh TTfgh_TfTgh.is_map isomorphic_reflexive
                            Maps.comp_CLS
                              [of TfTgh.p0 TTfgh_TfTgh.chine "TfTgh.p0  TTfgh_TfTgh.chine"]
                      by simp
                    also have "... = ⟦⟦Tgh.p0  TfTgh.p0  TTfgh_TfTgh.chine⟧⟧"
                      using fg gh TTfgh_TfTgh.is_map left_adjoints_compose isomorphic_reflexive
                            Maps.comp_CLS [of Tgh.p0 "TfTgh.p0  TTfgh_TfTgh.chine"
                                           "Tgh.p0  TfTgh.p0  TTfgh_TfTgh.chine"]
                      by simp
                    also have "... = ⟦⟦TTfgh.p0⟧⟧"
                      using prj_chine(3) by simp
                    also have "... = SPN_fgh.Prj0"
                      using prj_char by simp
                    finally show ?thesis by argo
                  qed
                  finally show ?thesis by blast
                qed
                show "SPN_fgh.νπ.prj1  tuple_BC =
                      SPN_fgh.νπ.prj1 
                        Maps.PRJ0 SPN_fgh.μ.leg0 (SPN_fgh.ν.leg1  SPN_fgh.νπ.prj1) 
                        ⟦⟦TTfgh_TfTgh.chine⟧⟧"
                proof -
                  have "SPN_fgh.νπ.prj1  tuple_BC = SPN_fgh.Prj01"
                    using csq(1) by simp
                  also have "... = SPN_fgh.νπ.prj1 
                                     Maps.PRJ0 SPN_fgh.μ.leg0 (SPN_fgh.ν.leg1  SPN_fgh.νπ.prj1) 
                                      ⟦⟦TTfgh_TfTgh.chine⟧⟧"
                  proof -
                    have "SPN_fgh.νπ.prj1 
                            Maps.PRJ0 SPN_fgh.μ.leg0 (SPN_fgh.ν.leg1  SPN_fgh.νπ.prj1) 
                             ⟦⟦TTfgh_TfTgh.chine⟧⟧ =
                          ⟦⟦Tgh.p1⟧⟧  ⟦⟦TfTgh.p0⟧⟧  ⟦⟦TTfgh_TfTgh.chine⟧⟧"
                      using fg gh Tgh.ρσ.prj_char TfTgh.prj_char(1) isomorphic_reflexive
                            Maps.comp_CLS [of "tab1 g" "prj1 (tab1 h) (tab0 g)" "tab1 g  Tgh.p1"]
                      by simp
                    also have "... = ⟦⟦Tgh.p1⟧⟧  ⟦⟦TfTgh.p0  TTfgh_TfTgh.chine⟧⟧"
                      using fg gh TTfgh_TfTgh.is_map isomorphic_reflexive
                            Maps.comp_CLS
                              [of TfTgh.p0 TTfgh_TfTgh.chine "TfTgh.p0  TTfgh_TfTgh.chine"]
                      by simp
                    also have "... = ⟦⟦Tgh.p1  TfTgh.p0  TTfgh_TfTgh.chine⟧⟧"
                      using fg gh TTfgh_TfTgh.is_map left_adjoints_compose isomorphic_reflexive
                            Maps.comp_CLS [of Tgh.p1 "TfTgh.p0  TTfgh_TfTgh.chine"
                                         "Tgh.p1  TfTgh.p0  TTfgh_TfTgh.chine"]
                      by simp
                    also have "... = ⟦⟦Tfg.p0  TTfgh.p1⟧⟧"
                      using prj_chine(2) by simp
                    also have "... = SPN_fgh.Prj01"
                      using prj_char by simp
                    finally show ?thesis by argo
                  qed
                  finally show ?thesis by blast
                qed
              qed
              finally show ?thesis by simp
            qed
          qed
          thus ?thesis
            using SPN_fgh.chine_assoc_def by simp
        qed
      qed
      finally show ?thesis by simp
    qed

    text ‹
      At long last, we can show associativity coherence for SPN›.
    ›

    lemma assoc_coherence:
    shows "LHS = RHS"
    proof (intro Span.arr_eqI)
      show "Span.par LHS RHS"
        using par_LHS_RHS by blast
      show "Chn LHS = Chn RHS"
      proof -
        have "Chn LHS = ⟦⟦HHfgh_HfHgh.chine⟧⟧  ⟦⟦THfgh_HHfgh.chine⟧⟧  ⟦⟦TTfgh_THfgh.chine⟧⟧"
          using Chn_LHS_eq by simp
        also have "... = ⟦⟦HHfgh_HfHgh.chine  THfgh_HHfgh.chine  TTfgh_THfgh.chine⟧⟧"
        proof -
          have "⟦⟦THfgh_HHfgh.chine⟧⟧  ⟦⟦TTfgh_THfgh.chine⟧⟧ =
                ⟦⟦THfgh_HHfgh.chine  TTfgh_THfgh.chine⟧⟧"
            using fg gh isomorphic_reflexive HHfgh_HfHgh.is_map THfgh_HHfgh.is_map
                  TTfgh_THfgh.is_map left_adjoints_compose
                  Maps.comp_CLS
                    [of THfgh_HHfgh.chine TTfgh_THfgh.chine "THfgh_HHfgh.chine  TTfgh_THfgh.chine"]
            by simp
          moreover
          have "⟦⟦HHfgh_HfHgh.chine⟧⟧  ⟦⟦THfgh_HHfgh.chine  TTfgh_THfgh.chine⟧⟧ =
                ⟦⟦HHfgh_HfHgh.chine  THfgh_HHfgh.chine  TTfgh_THfgh.chine⟧⟧"
          proof -
            have "ide (HHfgh_HfHgh.chine  THfgh_HHfgh.chine  TTfgh_THfgh.chine)"
            proof -
              have "ide (THfgh_HHfgh.chine  TTfgh_THfgh.chine)"  
                using fg gh HHfgh_HfHgh.is_map THfgh_HHfgh.is_map TTfgh_THfgh.is_map
                      left_adjoint_is_ide left_adjoints_compose
                by auto
              moreover have "src HHfgh_HfHgh.chine = trg (THfgh_HHfgh.chine  TTfgh_THfgh.chine)"
                using fg gh HHfgh_HfHgh.chine_in_hom α_def by auto
              ultimately show ?thesis by simp
            qed
            thus ?thesis
              using fg gh isomorphic_reflexive HHfgh_HfHgh.is_map THfgh_HHfgh.is_map
                    TTfgh_THfgh.is_map left_adjoints_compose
                    Maps.comp_CLS
                      [of HHfgh_HfHgh.chine "THfgh_HHfgh.chine  TTfgh_THfgh.chine"
                          "HHfgh_HfHgh.chine  THfgh_HHfgh.chine  TTfgh_THfgh.chine"]
              by auto
          qed
          ultimately show ?thesis by argo
        qed
        also have "... = ⟦⟦TfHgh_HfHgh.chine  TfTgh_TfHgh.chine  TTfgh_TfTgh.chine⟧⟧"
        proof -
          interpret A: vertical_composite_of_arrows_of_tabulations_in_maps V H 𝖺 𝗂 src trg
                         (f  g)  h TTfgh.tab tab0 h  TTfgh.p0 (tab1 f  Tfg.p1)  TTfgh.p1
                         f  g  h TfTgh.tab (tab0 h  Tgh.p0)  TfTgh.p0 tab1 f  TfTgh.p1
                         f  g  h TfHgh.ρσ.tab tab0 (g  h)  TfHgh.ρσ.p0 tab1 f  TfHgh.ρσ.p1
                         𝖺[f, g, h] f  g  h
            ..
          interpret B: vertical_composite_of_arrows_of_tabulations_in_maps V H 𝖺 𝗂 src trg
                         (f  g)  h TTfgh.tab tab0 h  TTfgh.p0 (tab1 f  Tfg.p1)  TTfgh.p1
                         f  g  h TfHgh.ρσ.tab tab0 (g  h)  TfHgh.ρσ.p0 tab1 f  TfHgh.ρσ.p1
                         f  g  h HfHgh.tab tab0 (f  g  h) tab1 (f  g  h)
                         𝖺[f, g, h] f  g  h
            using fg gh by unfold_locales auto
          interpret C: vertical_composite_of_arrows_of_tabulations_in_maps V H 𝖺 𝗂 src trg
                         (f  g)  h THfgh.ρσ.tab tab0 h  THfgh.ρσ.p0 tab1 (f  g)  THfgh.ρσ.p1
                         (f  g)  h HHfgh.tab tab0 ((f  g)  h) tab1 ((f  g)  h)
                         f  g  h HfHgh.tab tab0 (f  g  h) tab1 (f  g  h)
                         (f  g)  h 𝖺[f, g, h]
            using fg gh by unfold_locales auto
          interpret D: vertical_composite_of_arrows_of_tabulations_in_maps V H 𝖺 𝗂 src trg
                         (f  g)  h TTfgh.tab tab0 h  TTfgh.p0 (tab1 f  Tfg.p1)  TTfgh.p1
                         (f  g)  h THfgh.ρσ.tab tab0 h  THfgh.ρσ.p0 tab1 (f  g)  THfgh.ρσ.p1
                         f  g  h HfHgh.tab tab0 (f  g  h) tab1 (f  g  h)
                         (f  g)  h 𝖺[f, g, h]
            using fg gh by unfold_locales auto
          have "HHfgh_HfHgh.chine  THfgh_HHfgh.chine  TTfgh_THfgh.chine  D.chine"
          proof -
            have "D.chine  D.π.chine  TTfgh_THfgh.chine"
              using D.chine_char by simp
            also have "...  C.chine  TTfgh_THfgh.chine"
              using fg gh comp_arr_dom isomorphic_reflexive by simp
            also have "...  (C.π.chine  THfgh_HHfgh.chine)  TTfgh_THfgh.chine"
              using C.chine_char hcomp_isomorphic_ide by simp
            also have "...  (HHfgh_HfHgh.chine  THfgh_HHfgh.chine)  TTfgh_THfgh.chine"
            proof -
              have "C.π.chine = HHfgh_HfHgh.chine"
                using fg gh comp_arr_dom comp_cod_arr α_def by simp
              hence "isomorphic C.π.chine HHfgh_HfHgh.chine"
                using isomorphic_reflexive by simp
              thus ?thesis
                using hcomp_isomorphic_ide by simp
            qed
            also have "...  HHfgh_HfHgh.chine  THfgh_HHfgh.chine  TTfgh_THfgh.chine"
            proof -
              have "ide HHfgh_HfHgh.chine  ide THfgh_HHfgh.chine  ide TTfgh_THfgh.chine"
                by simp
              moreover have "src HHfgh_HfHgh.chine = trg THfgh_HHfgh.chine 
                             src THfgh_HHfgh.chine = trg TTfgh_THfgh.chine"
                using fg gh HHfgh_HfHgh.chine_in_hom THfgh_HHfgh.chine_in_hom 
                      TTfgh_THfgh.chine_in_hom α_def
                 by auto
              ultimately show ?thesis
                using fg gh iso_assoc isomorphic_def
                       assoc_in_hom [of HHfgh_HfHgh.chine THfgh_HHfgh.chine TTfgh_THfgh.chine]
                by auto
            qed
            finally show ?thesis
              using isomorphic_symmetric by blast
          qed
          also have "...  B.chine"
          proof -
            have "D.chine = B.chine"
              using fg gh comp_arr_dom comp_cod_arr by simp
            thus ?thesis
              using isomorphic_reflexive by simp
          qed
          also have "...  TfHgh_HfHgh.chine  TfTgh_TfHgh.chine  TTfgh_TfTgh.chine"
          proof -
            have "B.chine  TfHgh_HfHgh.chine  B.μ.chine"
              using B.chine_char by simp
            also have "...  TfHgh_HfHgh.chine  A.chine"
                using fg gh comp_cod_arr isomorphic_reflexive by simp
            also have "...  TfHgh_HfHgh.chine  TfTgh_TfHgh.chine  TTfgh_TfTgh.chine"
              using A.chine_char hcomp_ide_isomorphic by simp
            finally show ?thesis by blast
          qed
          finally have "HHfgh_HfHgh.chine  THfgh_HHfgh.chine  TTfgh_THfgh.chine 
                        TfHgh_HfHgh.chine  TfTgh_TfHgh.chine  TTfgh_TfTgh.chine"
            by blast
          thus ?thesis
            using fg gh Maps.CLS_eqI isomorphic_implies_hpar(1) by blast
        qed
        also have "... = ⟦⟦TfHgh_HfHgh.chine⟧⟧  ⟦⟦TfTgh_TfHgh.chine⟧⟧  ⟦⟦TTfgh_TfTgh.chine⟧⟧"
          using fg gh isomorphic_reflexive TfTgh_TfHgh.is_map TfHgh_HfHgh.is_map TTfgh_TfTgh.is_map
                left_adjoints_compose
                Maps.comp_CLS
                  [of TfTgh_TfHgh.chine TTfgh_TfTgh.chine "TfTgh_TfHgh.chine  TTfgh_TfTgh.chine"]
                Maps.comp_CLS
                  [of TfHgh_HfHgh.chine "TfTgh_TfHgh.chine  TTfgh_TfTgh.chine"
                      "TfHgh_HfHgh.chine  TfTgh_TfHgh.chine  TTfgh_TfTgh.chine"]
          by simp
        also have "... = Chn RHS"
          using Chn_RHS_eq CLS_chine tuple_ABC_eq_ABC'(2) by simp
        finally show ?thesis
          by blast
      qed
    qed

  end

  subsubsection "SPN is an Equivalence Pseudofunctor"

  context bicategory_of_spans
  begin

    interpretation Maps: maps_category V H 𝖺 𝗂 src trg ..
    interpretation Span: span_bicategory Maps.comp Maps.PRJ0 Maps.PRJ1 ..

    no_notation Fun.comp (infixl "" 55)
    notation Span.vcomp (infixr "" 55)
    notation Span.hcomp (infixr "" 53)
    notation Maps.comp (infixr "" 55)
    notation isomorphic (infix "" 50)

    interpretation SPN: "functor" V Span.vcomp SPN
      using SPN_is_functor by simp
    interpretation SPN: weak_arrow_of_homs V src trg Span.vcomp Span.src Span.trg SPN
      using SPN_is_weak_arrow_of_homs by simp
    interpretation HoSPN_SPN: composite_functor VV.comp Span.VV.comp Span.vcomp
                                SPN.FF λμν. Span.hcomp (fst μν) (snd μν)
      ..
    interpretation SPNoH: composite_functor VV.comp V Span.vcomp
                            λμν. fst μν  snd μν SPN
      ..

    interpretation Φ: transformation_by_components VV.comp Span.vcomp
                        HoSPN_SPN.map SPNoH.map λrs. CMP (fst rs) (snd rs)
      using compositor_is_natural_transformation by simp
    interpretation Φ: natural_isomorphism VV.comp Span.vcomp
                        HoSPN_SPN.map SPNoH.map Φ.map
      using compositor_is_natural_isomorphism by simp

    abbreviation Φ
    where "Φ  Φ.map"

    interpretation SPN: pseudofunctor V H 𝖺 𝗂 src trg
                        Span.vcomp Span.hcomp Span.assoc Span.unit Span.src Span.trg SPN Φ
    proof
      show "f g h.  ide f; ide g; ide h; src f = trg g; src g = trg h  
                    SPN 𝖺[f, g, h]  Φ (f  g, h)  (Φ (f, g)  SPN h) =
                    Φ (f, g  h)  (SPN f  Φ (g, h))  Span.assoc (SPN f) (SPN g) (SPN h)"
      proof -
        fix f g h
        assume f: "ide f" and g: "ide g" and h: "ide h"
        assume fg: "src f = trg g" and gh: "src g = trg h"
        interpret fgh: three_composable_identities_in_bicategory_of_spans V H 𝖺 𝗂 src trg f g h
          using f g h fg gh
          by (unfold_locales, simp)
        show "fgh.LHS = fgh.RHS"
          using fgh.assoc_coherence by simp
      qed
    qed

    lemma SPN_is_pseudofunctor:
    shows "pseudofunctor V H 𝖺 𝗂 src trg
             Span.vcomp Span.hcomp Span.assoc Span.unit Span.src Span.trg SPN Φ"
      ..

    interpretation SPN: equivalence_pseudofunctor V H 𝖺 𝗂 src trg
                        Span.vcomp Span.hcomp Span.assoc Span.unit Span.src Span.trg SPN Φ
    proof
      show "μ μ'. par μ μ'; SPN μ = SPN μ'  μ = μ'"
      proof -
        fix μ μ'
        assume par: "par μ μ'"
        assume eq: "SPN μ = SPN μ'"
        interpret dom_μ: identity_in_bicategory_of_spans V H 𝖺 𝗂 src trg dom μ
          using par apply unfold_locales by auto
        interpret cod_μ: identity_in_bicategory_of_spans V H 𝖺 𝗂 src trg cod μ
          using par apply unfold_locales by auto
        interpret μ: arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg
                       dom μ tab_of_ide (dom μ) tab0 (dom μ) tab1 (dom μ)
                       cod μ tab_of_ide (cod μ) tab0 (cod μ) tab1 (cod μ)
                       μ
          using par apply unfold_locales by auto
        interpret μ: arrow_in_bicategory_of_spans V H 𝖺 𝗂 src trg dom μ cod μ μ
          using par apply unfold_locales by auto
        interpret μ': arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg
                        dom μ tab_of_ide (dom μ) tab0 (dom μ) tab1 (dom μ)
                        cod μ tab_of_ide (cod μ) tab0 (cod μ) tab1 (cod μ)
                        μ'
          using par apply unfold_locales by auto
        interpret μ': arrow_in_bicategory_of_spans V H 𝖺 𝗂 src trg dom μ cod μ μ'
          using par apply unfold_locales by auto
        have "μ.chine  μ'.chine"
          using par eq SPN_def spn_def Maps.CLS_eqI μ.is_ide by auto
        hence "μ.Δ = μ'.Δ"
          using μ.Δ_naturality μ'.Δ_naturality
          by (metis μ.Δ_simps(4) μ'.Δ_simps(4) μ.chine_is_induced_map μ'.chine_is_induced_map
              μ.induced_map_preserved_by_iso)
        thus "μ = μ'"
          using par μ.μ_in_terms_of_Δ μ'.μ_in_terms_of_Δ by metis
      qed
      show "a'. Span.obj a'  a. obj a  Span.equivalent_objects (SPN.map0 a) a'"
      proof -
        fix a'
        assume a': "Span.obj a'"
        let ?a = "Maps.Dom (Chn a')"
        have a: "obj ?a"
          using a' Span.obj_char Span.ide_char Maps.ide_charCC by blast
        moreover have "Span.equivalent_objects (SPN.map0 ?a) a'"
        proof -
          have "SPN.map0 ?a = a'"
          proof (intro Span.arr_eqI)
            have "Chn (SPN.map0 ?a) = Chn (Span.src (SPN ?a))"
              using a a' by auto
            also have "... = Maps.MkIde (Maps.Dom (Chn a'))"
            proof -
              have "Maps.arr ⟦⟦tab0 (dom (Maps.Dom (Chn a')))⟧⟧"
              proof -
                have "is_left_adjoint (tab0 (dom (Maps.Dom (Chn a'))))"
                  using a by auto
                thus ?thesis
                  using Maps.CLS_in_hom by auto
              qed
              moreover have "arr (Maps.Dom (Chn a'))"
                using a by auto
              moreover have "Span.arr (SPN (Maps.Dom (Chn a')))"
                using a a' SPN_in_hom by auto
              ultimately show ?thesis
                using a a' SPN_def Span.src_def Maps.cod_char obj_simps(2) by simp
            qed
            also have "... = Chn a'"
              using a' Maps.MkIde_Dom Span.obj_char Span.ide_char by simp
            finally show "Chn (SPN.map0 ?a) = Chn a'" by simp
            show "Span.par (SPN.map0 (Maps.Dom (Chn a'))) a'"
              using a a' Span.obj_char
              apply (intro conjI)
              using SPN.map0_simps(1) Span.obj_def
                 apply blast
                apply simp
               apply (metis (no_types, lifting) SPN.map0_def SPN.preserves_arr Span.obj_src
                  Chn (SPN.map0 (Maps.Dom (Chn a'))) = Chn a' obj_def)
              by (metis (no_types, lifting) SPN.map0_def SPN.preserves_arr Span.obj_src
                  Chn (SPN.map0 (Maps.Dom (Chn a'))) = Chn a' obj_def)
          qed
          thus ?thesis
            using Span.equivalent_objects_reflexive
            by (simp add: a')
        qed
        ultimately show "a. obj a  Span.equivalent_objects (SPN.map0 a) a'"
          by auto
      qed
      show "a b g. obj a; obj b; Span.in_hhom g (SPN.map0 a) (SPN.map0 b); Span.ide g
                        f. «f : a  b»  ide f  Span.isomorphic (SPN f) g"
      proof -
        fix a b g
        assume a: "obj a" and b: "obj b"
        and g_in_hhom: "Span.in_hhom g (SPN.map0 a) (SPN.map0 b)"
        and ide_g: "Span.ide g"
        have arr_a: "arr a"
          using a by auto
        have arr_b: "arr b"
         using b by auto
        show "f. «f : a  b»  ide f  Span.isomorphic (SPN f) g"
        proof -
          interpret g: arrow_of_spans Maps.comp g
            using ide_g Span.ide_char by blast
          interpret g: identity_arrow_of_spans Maps.comp g
            using ide_g Span.ide_char
            by unfold_locales auto
          interpret REP_leg0: map_in_bicategory V H 𝖺 𝗂 src trg Maps.REP g.leg0
            using Maps.REP_in_Map [of g.leg0]
            by unfold_locales auto
          have 0: "«Maps.REP g.leg0 : src (Maps.REP g.apex)  Maps.Cod g.leg0»"
            using g.dom.leg_in_hom Maps.REP_in_hhom
            by (metis (no_types, lifting) Maps.Dom_cod Maps.REP_simps(2) Maps.arr_cod
                g.dom.leg_simps(1))
          have 1: "«Maps.REP g.leg1 : src (Maps.REP g.apex)  Maps.Cod g.leg1»"
            using g.dom.leg_in_hom Maps.REP_in_hhom
            by (metis (no_types, lifting) Maps.Dom_cod Maps.REP_simps(2) Maps.arr_cod
                g.dom.leg_simps(3))
          let ?f = "Maps.REP g.leg1  (Maps.REP g.leg0)*"
          have f_in_hhom: "«?f : a  b»"
          proof
            show "«Maps.REP g.leg1 : src (Maps.REP g.apex)  b»"
            proof -
              have "«Maps.REP g.leg1 : src (Maps.REP g.apex)  Maps.Cod g.leg1»"
                using 1 by simp
              moreover have "Maps.Cod g.leg1 = b"
              proof -
                have "src (Maps.REP g.dtrg) = src (Maps.REP (Leg0 (Dom (SPN.map0 b))))"
                  using g_in_hhom Span.trg_def [of g] by auto
                also have "... = src (Maps.REP (Maps.cod ⟦⟦tab0 b⟧⟧))"
                  using b arr_b SPN.map0_def Span.src_def SPN_in_hom by auto
                also have "... = src (Maps.REP ⟦⟦trg (tab0 b)⟧⟧)"
                  using b Maps.CLS_in_hom [of "tab0 b"] by force
                also have "... = src (Maps.REP ⟦⟦b⟧⟧)"
                  using b by fastforce
                also have "... = b"
                  using b obj_simps by auto
                finally show ?thesis by simp
              qed
              ultimately show ?thesis by argo
            qed
            show "«(Maps.REP g.leg0)* : a  src (Maps.REP g.apex)»"
            proof -
              have "«Maps.REP g.leg0 : src (Maps.REP g.apex)  a»"
              proof -
                have "src (Maps.REP g.dsrc) = src (Maps.REP (Leg0 (Dom (SPN.map0 a))))"
                  using g_in_hhom Span.src_def [of g] by auto
                also have "... = src (Maps.REP (Maps.cod ⟦⟦tab0 a⟧⟧))"
                  using a arr_a SPN.map0_def Span.src_def SPN_in_hom by auto
                also have "... = src (Maps.REP ⟦⟦trg (tab0 a)⟧⟧)"
                  using a Maps.CLS_in_hom [of "tab0 a"] by force
                also have "... = src (Maps.REP ⟦⟦a⟧⟧)"
                  using a by fastforce
                also have "... = a"
                  using a obj_simps by auto
                finally show ?thesis by fast
              qed
              thus ?thesis
                using REP_leg0.antipar REP_leg0.ide_right
                apply (intro in_hhomI) by auto
            qed
          qed
          moreover have ide_f: "ide ?f"
            using REP_leg0.antipar f_in_hhom by fastforce
          moreover have "Span.isomorphic (SPN ?f) g"
          proof -
            have SPN_f_eq: "SPN ?f = Chn = ⟦⟦spn ?f⟧⟧,
                                      Dom = Leg0 = ⟦⟦tab0 ?f⟧⟧, Leg1 = ⟦⟦tab1 ?f⟧⟧,
                                      Cod = Leg0 = ⟦⟦tab0 ?f⟧⟧, Leg1 = ⟦⟦tab1 ?f⟧⟧"
              using calculation(1) SPN_def [of ?f] REP_leg0.antipar by auto
            text ‹
              We need an invertible arrow of spans from SPN f› to g›.
              There exists a tabulation (REP g.leg0, ρ, REP g.leg1)› of f›.
              There is also a tabulation (tab0 f, ρ', tab1 f) of f›.
              As these are tabulations of the same arrow, they are equivalent.
              This yields an equivalence map which is an arrow of spans from
              (tab0 f, tab1 f)› to (REP g.leg0, ρ, REP g.leg1)›.
              Its isomorphism class is an invertible arrow of spans in maps
              from (CLS (tab0 f), CLS (tab1 f))› to (g.leg0, g.leg1)›.
            ›
            interpret f: identity_in_bicategory_of_spans V H 𝖺 𝗂 src trg ?f
              using ide_f apply unfold_locales by auto
            interpret f: arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg
                           ?f f.tab tab0 ?f tab1 ?f ?f f.tab tab0 ?f tab1 ?f ?f
              using f.is_arrow_of_tabulations_in_maps by simp
            interpret g: span_of_maps V H 𝖺 𝗂 src trg Maps.REP g.leg0 Maps.REP g.leg1
              using Span.arr_char
              by (unfold_locales, blast+)

            have 2: "src (Maps.REP g.leg0) = src (Maps.REP g.leg1)"
              using 0 1 by fastforce
            hence "ρ. tabulation (⋅) (⋆) 𝖺 𝗂 src trg ?f ρ (Maps.REP g.leg0) (Maps.REP g.leg1)"
              using BS2' [of "Maps.REP g.leg0" "Maps.REP g.leg1" ?f] isomorphic_reflexive
                    Span.arr_char
              by auto
            hence "tabulation V H 𝖺 𝗂 src trg ?f
                     (REP_leg0.trnrη (Maps.REP g.leg1) ?f) (Maps.REP g.leg0) (Maps.REP g.leg1)"
              using 2 REP_leg0.canonical_tabulation [of "Maps.REP g.leg1"] by auto
            hence 3: "w w' φ ψ θ ν θ' ν'.
                        equivalence_in_bicategory (⋅) (⋆) 𝖺 𝗂 src trg w' w ψ φ 
                        «w : src (tab0 ?f)  src (Maps.REP g.leg0)» 
                        «w' : src (Maps.REP g.leg0)  src (tab0 ?f)» 
                        «θ : Maps.REP g.leg0  w  tab0 ?f» 
                        «ν : tab1 ?f  Maps.REP g.leg1  w»  iso ν 
                        f.tab = (?f  θ)  𝖺[?f, Maps.REP g.leg0, w] 
                                  (REP_leg0.trnrη (Maps.REP g.leg1) ?f  w)  ν 
                        «θ' : tab0 ?f  w'  Maps.REP g.leg0» 
                        «ν' : Maps.REP g.leg1  tab1 ?f  w'»  iso ν' 
                        REP_leg0.trnrη (Maps.REP g.leg1) ?f =
                        (?f  θ')  𝖺[?f, tab0 ?f, w']  (f.tab  w')  ν'"
                          using f.apex_unique_up_to_equivalence
                                  [of "REP_leg0.trnrη (Maps.REP g.leg1) ?f"
                                      "Maps.REP g.leg0" "Maps.REP g.leg1"]
              by simp
            obtain w w' φ ψ θ ν θ' ν'
              where 4: "equivalence_in_bicategory (⋅) (⋆) 𝖺 𝗂 src trg w' w ψ φ 
                        «w : src (tab0 ?f)  src (Maps.REP g.leg0)» 
                        «w' : src (Maps.REP g.leg0)  src (tab0 ?f)» 
                        «θ : Maps.REP g.leg0  w  tab0 ?f» 
                        «ν : tab1 ?f  Maps.REP g.leg1  w»  iso ν 
                        f.tab = (?f  θ)  𝖺[?f, Maps.REP g.leg0, w] 
                                  (REP_leg0.trnrη (Maps.REP g.leg1) ?f  w)  ν 
                        «θ' : tab0 ?f  w'  Maps.REP g.leg0» 
                        «ν' : Maps.REP g.leg1  tab1 ?f  w'»  iso ν' 
                        REP_leg0.trnrη (Maps.REP g.leg1) ?f =
                        (?f  θ')  𝖺[?f, tab0 ?f, w']  (f.tab  w')  ν'"
              using 3 by meson
            hence wθν: "equivalence_map w  «w : src (tab0 ?f)  src (Maps.REP g.leg0)» 
                        «θ : Maps.REP g.leg0  w  tab0 ?f» 
                        «ν : tab1 ?f  Maps.REP g.leg1  w»  iso ν"
              using equivalence_map_def quasi_inverses_def quasi_inverses_symmetric
              by meson
            let ?W = "Chn = ⟦⟦w⟧⟧, Dom = Dom (SPN ?f), Cod = Dom g"
            have W_in_hom: "Span.in_hom ?W (SPN ?f) g"
            proof
              have "arrow_of_spans Maps.comp ?W"
              proof
                interpret Dom_W: span_in_category Maps.comp Dom ?W
                proof (unfold_locales, intro conjI)
                  show "Maps.arr (Leg0 (Dom ?W))"
                    apply (intro Maps.arrICC)
                       apply auto
                    by (metis f.base_simps(2) f.satisfies_T0 f.u_in_hom src_hcomp)
                  show "Maps.arr (Leg1 (Dom ?W))"
                    using 1
                    apply (intro Maps.arrICC)
                       apply auto
                  proof -
                    let ?f = "tab1 (Maps.REP g.leg1  (Maps.REP g.leg0)*)"
                    assume 1: "«Maps.REP g.leg1 : Maps.Dom g.apex  Maps.Cod g.leg1»"
                    have "«?f : src (tab0 (Maps.REP g.leg1  (Maps.REP g.leg0)*))
                                    Maps.Cod g.leg1» 
                          is_left_adjoint ?f  tab1 (Maps.REP g.leg1  (Maps.REP g.leg0)*) = ?f"
                      using 1 by simp
                    thus "f. «f : src (tab0 (Maps.REP g.leg1  (Maps.REP g.leg0)*))
                                    Maps.Cod g.leg1» 
                              is_left_adjoint f 
                              tab1 (Maps.REP g.leg1  (Maps.REP g.leg0)*) = f"
                      by blast
                  qed
                  show "Maps.dom (Leg0 (Dom ?W)) = Maps.dom (Leg1 (Dom ?W))"
                  proof -
                    have "Maps.dom (Leg0 (Dom ?W)) =
                          Maps.MkIde (src (tab0 (Maps.REP g.leg1  (Maps.REP g.leg0)*)))"
                      using Maps.dom_char
                      apply simp
                      by (metis (no_types, lifting) Maps.CLS_in_hom Maps.in_homE f.base_simps(2)
                          f.satisfies_T0 f.u_simps(3) hcomp_simps(1))
                    also have "... = Maps.dom (Leg1 (Dom ?W))"
                      using Maps.dom_char Maps.CLS_in_hom f.leg1_is_map f_in_hhom
                      apply simp
                      by (metis (no_types, lifting) Maps.in_homE Maps.REP_simps(3) f.base_simps(2)
                          f.leg1_is_map f.leg1_simps(3) f.leg1_simps(4) g.dom.leg_simps(3)
                          trg_hcomp)
                    finally show ?thesis by blast
                  qed
                qed
                show "Maps.span (Leg0 (Dom ?W)) (Leg1 (Dom ?W))"
                  using Dom_W.span_in_category_axioms Dom_W.is_span by blast
                interpret Cod_W: span_in_category Maps.comp Cod ?W
                  apply unfold_locales by fastforce
                show "Maps.span (Leg0 (Cod ?W)) (Leg1 (Cod ?W))"
                  by fastforce
                show "Maps.in_hom (Chn ?W) Dom_W.apex Cod_W.apex"
                proof
                  show 1: "Maps.arr (Chn ?W)"
                    using wθν Maps.CLS_in_hom [of w] equivalence_is_adjoint by auto
                  show "Maps.dom (Chn ?W) = Dom_W.apex"
                  proof -
                    have "Maps.dom (Chn ?W) = Maps.MkIde (src w)"
                      using 1 wθν Maps.dom_char by simp
                    also have "... = Dom_W.apex"
                    proof -
                      have "src w = src (tab0 ?f)"
                        using wθν by blast
                      thus ?thesis
                        using Dom_W.apex_def Maps.arr_char Maps.dom_char
                        apply simp
                        by (metis (no_types, lifting) f.base_simps(2) f.satisfies_T0
                            f.u_in_hom hcomp_simps(1))
                    qed
                    finally show ?thesis by fastforce
                  qed
                  show "Maps.cod (Chn ?W) = Cod_W.apex"
                  proof -
                    have "Maps.cod (Chn ?W) = Maps.MkIde (trg w)"
                      using 1 wθν Maps.cod_char by simp
                    also have "... = Cod_W.apex"
                    proof -
                      have "trg w = src (Maps.REP g.leg0)"
                        using wθν by blast
                      thus ?thesis
                        using Cod_W.apex_def Maps.arr_char Maps.cod_char
                        apply simp
                        using g.dom.apex_def Maps.dom_char Maps.REP_simps(2) g.dom.is_span
                        by presburger
                    qed
                    finally show ?thesis by fastforce
                  qed
                qed
                show "Cod_W.leg0  Chn ?W = Dom_W.leg0"
                proof -
                  have "Cod_W.leg0  Chn ?W = g.leg0  ⟦⟦w⟧⟧"
                    by simp
                  also have "... = ⟦⟦Maps.REP g.leg0⟧⟧  ⟦⟦w⟧⟧"
                    using g.dom.leg_simps(1) Maps.CLS_REP [of g.leg0]
                    by simp
                  also have "... = ⟦⟦Maps.REP g.leg0  w⟧⟧"
                  proof -
                    have "is_left_adjoint (Maps.REP g.leg0)"
                      by fast
                    moreover have "is_left_adjoint w"
                      using wθν equivalence_is_adjoint by simp
                    moreover have "Maps.REP g.leg0  w  Maps.REP g.leg0  w"
                      using wθν isomorphic_reflexive Maps.REP_in_hhom
                      by (metis (no_types, lifting) REP_leg0.ide_left adjoint_pair_antipar(1)
                          calculation(2) ide_hcomp in_hhomE)
                    ultimately show ?thesis
                      using wθν Maps.comp_CLS isomorphic_reflexive equivalence_is_adjoint
                      by blast
                  qed
                  also have "... = ⟦⟦tab0 ?f⟧⟧"
                  proof -
                    have "iso θ"
                    proof -
                      have "is_left_adjoint (Maps.REP g.leg0  w)"
                        using wθν equivalence_is_adjoint Maps.REP_in_hhom
                        by (simp add: g.leg0_is_map in_hhom_def left_adjoints_compose)
                      moreover have "is_left_adjoint (tab0 ?f)"
                        by simp
                      ultimately show ?thesis
                        using wθν BS3 by blast
                    qed
                    thus ?thesis
                      using wθν Maps.CLS_eqI equivalence_is_adjoint
                      by (meson isomorphic_def isomorphic_implies_hpar(1))
                  qed
                  finally show ?thesis by fastforce
                qed
                show "Cod_W.leg1  Chn ?W = Dom_W.leg1"
                proof -
                  have "Cod_W.leg1  Chn ?W = g.leg1  ⟦⟦w⟧⟧"
                    by simp
                  also have "... = ⟦⟦Maps.REP g.leg1⟧⟧  ⟦⟦w⟧⟧"
                    using g.dom.leg_simps(3) Maps.CLS_REP by presburger
                  also have "... = ⟦⟦Maps.REP g.leg1  w⟧⟧"
                  proof -
                    have "is_left_adjoint (Maps.REP g.leg1)"
                      by fast
                    moreover have "is_left_adjoint w"
                      using wθν equivalence_is_adjoint by simp
                    moreover have "Maps.REP g.leg1  w  Maps.REP g.leg1  w"
                      using wθν isomorphic_reflexive Maps.REP_in_hhom
                      by (metis (no_types, lifting) "2" calculation(2) g.dom.is_span
                          hcomp_ide_isomorphic Maps.ide_REP in_hhomE
                          right_adjoint_determines_left_up_to_iso)
                    ultimately show ?thesis
                      using wθν Maps.comp_CLS isomorphic_reflexive equivalence_is_adjoint
                      by blast
                  qed
                  also have "... = ⟦⟦tab1 ?f⟧⟧"
                  proof -
                    have "ide (Maps.REP g.leg1  w)"
                      using 2 wθν equivalence_map_is_ide by auto
                    moreover have "Maps.REP g.leg1  w 
                                   tab1 (Maps.REP g.leg1  (Maps.REP g.leg0)*)"
                      using wθν equivalence_is_adjoint f.leg1_is_map
                            right_adjoint_determines_left_up_to_iso adjoint_pair_preserved_by_iso
                      by (meson adjoint_pair_antipar(2) ide_in_hom(2) ide_is_iso)
                    ultimately show ?thesis
                      using Maps.CLS_eqI by blast
                  qed
                  finally show ?thesis by fastforce
                qed
              qed
              thus W: "Span.arr ?W"
                using Span.arr_char by blast
              interpret Dom_W:
                span_in_category Maps.comp
                  Leg0 = Maps.MkArr (src (tab0 (Maps.REP g.leg1  (Maps.REP g.leg0)*)))
                                      (src (Maps.REP g.leg0)*)
                                      (iso_class (tab0 (Maps.REP g.leg1  (Maps.REP g.leg0)*))),
                    Leg1 = Maps.MkArr (src (tab0 (Maps.REP g.leg1  (Maps.REP g.leg0)*)))
                                      (Maps.Cod g.leg1)
                                      (iso_class (tab1 (Maps.REP g.leg1  (Maps.REP g.leg0)*)))
                using W Span.arr_char
                by (simp add: arrow_of_spans_def)
              interpret Cod_W: span_in_category Maps.comp Cod ?W
                using W Span.arr_char
                by (simp add: arrow_of_spans_def)
              show "Span.dom ?W = SPN ?f"
              proof -
                have "Span.dom ?W =
                      Chn = Dom_W.apex,
                       Dom = Leg0 = ⟦⟦tab0 (Maps.REP g.leg1  (Maps.REP g.leg0)*)⟧⟧,
                              Leg1 = ⟦⟦tab1 (Maps.REP g.leg1  (Maps.REP g.leg0)*)⟧⟧,
                       Cod = Leg0 = ⟦⟦tab0 (Maps.REP g.leg1  (Maps.REP g.leg0)*)⟧⟧,
                              Leg1 = ⟦⟦tab1 (Maps.REP g.leg1  (Maps.REP g.leg0)*)⟧⟧"
                  using 0 W Span.dom_char by simp
                also have "... = SPN ?f"
                  using SPN_def Dom_W.apex_def Maps.dom_char Dom_W.is_span iso_class_eqI
                        spn_ide
                  apply simp
                  using ide_f by blast
                finally show ?thesis by blast
              qed
              show "Span.cod ?W = g"
                using 0 W Span.cod_char Cod_W.apex_def by simp
            qed
            moreover have "Span.iso ?W"
            proof -
              have "Maps.iso ⟦⟦w⟧⟧"
              proof -
                have "Maps.arr ⟦⟦w⟧⟧  w  Maps.Map ⟦⟦w⟧⟧  equivalence_map w"
                proof (intro conjI)
                  show 1: "Maps.arr ⟦⟦w⟧⟧"
                    using wθν Maps.CLS_in_hom equivalence_is_adjoint by blast
                  show "equivalence_map w"
                    using wθν by blast
                  show "w  Maps.Map ⟦⟦w⟧⟧"
                    using 1 wθν equivalence_is_adjoint Maps.arr_char
                    by (simp add: equivalence_map_is_ide ide_in_iso_class)
                qed
                thus ?thesis
                  using Maps.iso_char' by blast
              qed
              thus ?thesis
                using wθν W_in_hom Span.iso_char by auto
            qed
            ultimately show ?thesis
              using Span.isomorphic_def by blast
          qed
          ultimately show ?thesis by blast
        qed
      qed
      show "r s τ. ide r; ide s; src r = src s; trg r = trg s; Span.in_hom τ (SPN r) (SPN s)
                        μ. «μ : r  s»  SPN μ = τ"
      proof -
        fix r s τ
        assume r: "ide r" and s: "ide s"
        assume src_eq: "src r = src s" and trg_eq: "trg r = trg s"
        assume τ: "Span.in_hom τ (SPN r) (SPN s)"
        interpret τ: arrow_of_spans Maps.comp τ
          using τ Span.arr_char by auto
        interpret r: identity_in_bicategory_of_spans V H 𝖺 𝗂 src trg r
          using r by unfold_locales auto
        interpret s: identity_in_bicategory_of_spans V H 𝖺 𝗂 src trg s
          using s by unfold_locales auto
        interpret s: arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg
                       s s.tab tab0 s tab1 s s s.tab tab0 s tab1 s s
          using s.is_arrow_of_tabulations_in_maps by simp
        have τ_dom_leg0_eq: "τ.dom.leg0 = ⟦⟦tab0 r⟧⟧"
          using τ Span.dom_char SPN_def [of r] by auto
        have τ_dom_leg1_eq: "τ.dom.leg1 = ⟦⟦tab1 r⟧⟧"
          using τ Span.dom_char SPN_def [of r] by auto
        have τ_cod_leg0_eq: "τ.cod.leg0 = ⟦⟦tab0 s⟧⟧"
          using τ Span.cod_char SPN_def [of s] by auto
        have τ_cod_leg1_eq: "τ.cod.leg1 = ⟦⟦tab1 s⟧⟧"
          using τ Span.cod_char SPN_def [of s] by auto

        have 1: "tab0 s  Maps.REP τ.chine  tab0 r"
        proof -
          have "tab0 s  Maps.REP τ.chine  Maps.REP τ.cod.leg0  Maps.REP τ.chine"
          proof -
            have "Maps.REP τ.cod.leg0  tab0 s"
              using τ_cod_leg0_eq Maps.CLS_REP Maps.CLS_eqI Maps.REP_CLS s.satisfies_T0
              by presburger
            moreover have "src (tab0 s) = trg (Maps.REP τ.chine)"
              using Maps.seq_char [of "⟦⟦tab0 s⟧⟧" τ.chine] τ.cod.leg_simps(1)
                    τ.leg0_commutes τ_cod_leg0_eq
              by auto
            ultimately show ?thesis
              using hcomp_isomorphic_ide [of "tab0 s" "Maps.REP τ.cod.leg0" "Maps.REP τ.chine"]
                    isomorphic_symmetric
              by fastforce
          qed
          also have "...  Maps.REP τ.dom.leg0"
          proof -
            have "⟦⟦Maps.REP τ.cod.leg0⟧⟧  ⟦⟦Maps.REP τ.chine⟧⟧ = ⟦⟦Maps.REP τ.dom.leg0⟧⟧"
              using τ.leg0_commutes Maps.CLS_REP τ.chine_simps(1)
                    τ.cod.leg_simps(1) τ.dom.leg_simps(1)
              by presburger
            hence "⟦⟦Maps.REP τ.cod.leg0  Maps.REP τ.chine⟧⟧ = ⟦⟦Maps.REP τ.dom.leg0⟧⟧"
              using Maps.comp_CLS [of "Maps.REP τ.cod.leg0" "Maps.REP τ.chine"
                                      "Maps.REP τ.cod.leg0  Maps.REP τ.chine"]
                    isomorphic_reflexive
              by (metis (no_types, lifting) Maps.seq_char Maps.REP_in_hhom(2) Maps.REP_simps(2-3)
                  τ.chine_in_hom τ.cod.leg_in_hom(1) τ.dom.leg_simps(1) τ.leg0_commutes
                  ide_hcomp Maps.ide_REP)
            thus ?thesis
              using Maps.CLS_eqI Maps.seq_char Maps.ide_REP
              by (meson calculation isomorphic_implies_ide(2))
          qed
          also have "...  tab0 r"
            using τ_dom_leg0_eq Maps.CLS_REP Maps.CLS_eqI Maps.REP_CLS r.satisfies_T0
            by presburger
          finally show ?thesis by blast
        qed
        obtain θ where θ: "«θ : tab0 s  Maps.REP τ.chine  tab0 r»  iso θ"
          using 1 by blast
        have 2: "tab1 s  Maps.REP τ.chine  tab1 r"
        proof -
          have "tab1 s  Maps.REP τ.chine  Maps.REP τ.cod.leg1  Maps.REP τ.chine"
          proof -
            have "Maps.REP τ.cod.leg1  tab1 s"
              using τ_cod_leg1_eq Maps.CLS_REP Maps.CLS_eqI Maps.REP_CLS s.leg1_is_map
              by presburger
            moreover have "src (Maps.REP τ.cod.leg1) = trg (Maps.REP τ.chine)"
              using Maps.seq_char by auto
            ultimately show ?thesis
              using hcomp_isomorphic_ide [of "Maps.REP τ.cod.leg1" "tab1 s" "Maps.REP τ.chine"]
                    isomorphic_symmetric
              by fastforce
          qed
          also have "...  Maps.REP τ.dom.leg1"
          proof -
            have "⟦⟦Maps.REP τ.cod.leg1⟧⟧  ⟦⟦Maps.REP τ.chine⟧⟧ = ⟦⟦Maps.REP τ.dom.leg1⟧⟧"
              using τ.leg1_commutes Maps.CLS_REP τ.chine_simps(1)
                    τ.cod.leg_simps(3) τ.dom.leg_simps(3)
              by presburger
            hence "⟦⟦Maps.REP τ.cod.leg1  Maps.REP τ.chine⟧⟧ = ⟦⟦Maps.REP τ.dom.leg1⟧⟧"
              using Maps.comp_CLS [of "Maps.REP τ.cod.leg1" "Maps.REP τ.chine"
                                      "Maps.REP τ.cod.leg1  Maps.REP τ.chine"]
                    isomorphic_reflexive
              by (metis (no_types, lifting) Maps.seq_char Maps.REP_in_hhom(2)
                  Maps.REP_simps(2) Maps.REP_simps(3) τ.chine_in_hom τ.cod.leg_in_hom(2)
                  τ.dom.leg_simps(3) τ.leg1_commutes ide_hcomp Maps.ide_REP)
            thus ?thesis
              using Maps.CLS_eqI Maps.seq_char Maps.ide_REP
              by (meson calculation isomorphic_implies_ide(2))
          qed
          also have "...  tab1 r"
            using τ_dom_leg1_eq Maps.CLS_REP Maps.CLS_eqI Maps.REP_CLS r.leg1_is_map
            by presburger
          finally show ?thesis by blast
        qed
        obtain ν where ν: "«ν : tab1 r  tab1 s  Maps.REP τ.chine»  iso ν"
          using 2 isomorphic_symmetric by blast

        define Δ
          where "Δ  (s  θ)  𝖺[s, tab0 s, Maps.REP τ.chine]  (s.tab  Maps.REP τ.chine)  ν"
        have Δ: "«Δ : tab1 r  s  tab0 r»"
        proof (unfold Δ_def, intro comp_in_homI)
          show "«ν : tab1 r  tab1 s  Maps.REP τ.chine»"
            using ν by simp
          show 3: "«s.tab  Maps.REP τ.chine :
                      tab1 s  Maps.REP τ.chine  (s  tab0 s)  Maps.REP τ.chine»"
            apply (intro hcomp_in_vhom)
              apply auto
            using "1" by fastforce
          show "«𝖺[s, tab0 s, Maps.REP τ.chine] :
                   (s  tab0 s)  Maps.REP τ.chine  s  tab0 s  Maps.REP τ.chine»"
            using s assoc_in_hom [of s "tab0 s" "Maps.REP τ.chine"]
            by (metis (no_types, lifting) Maps.ide_REP 3 τ.chine_simps(1) hcomp_in_vhomE
                ideD(2) ideD(3) s.ide_u s.tab_simps(2) s.u_simps(3))
          show "«s  θ : s  tab0 s  Maps.REP τ.chine  s  tab0 r»"
            using 1 s θ isomorphic_implies_hpar(4) src_eq by auto
        qed
        define μ where "μ  r.T0.trnrε s Δ  inv (r.T0.trnrε r r.tab)"
        have μ: "«μ : r  s»"
        proof (unfold μ_def, intro comp_in_homI)
          show "«inv (r.T0.trnrε r r.tab) : r  tab1 r  (tab0 r)*»"
            using r.yields_isomorphic_representation by fastforce
          show "«r.T0.trnrε s Δ : tab1 r  (tab0 r)*  s»"
            using s Δ src_eq r.T0.adjoint_transpose_right(2) [of s "tab1 r"] by auto
        qed
        interpret μ: arrow_in_bicategory_of_spans V H 𝖺 𝗂 src trg r s μ
          using μ by unfold_locales auto
        interpret μ: arrow_of_tabulations_in_maps V H 𝖺 𝗂 src trg
                       r r.tab tab0 r tab1 r s s.tab tab0 s tab1 s μ
          using μ.is_arrow_of_tabulations_in_maps by simp
        have Δ_eq: "Δ = μ.Δ"
        proof -
          have "r.T0.trnrε s Δ  inv (r.T0.trnrε r r.tab) =
                r.T0.trnrε s μ.Δ  inv (r.T0.trnrε r r.tab)"
            using μ μ_def μ.μ_in_terms_of_Δ by auto
          hence "r.T0.trnrε s Δ = r.T0.trnrε s μ.Δ"
            using r s Δ r.T0.adjoint_transpose_right(2) r.yields_isomorphic_representation
                  iso_inv_iso iso_is_retraction retraction_is_epi epiE
            by (metis μ.in_hom μ_def arrI)
          thus ?thesis
            using Δ μ.Δ_in_hom(2) src_eq r.T0.adjoint_transpose_right(6)
                  bij_betw_imp_inj_on
                    [of "r.T0.trnrε s" "hom (tab1 r) (s  tab0 r)" "hom (tab1 r  (tab0 r)*) s"]
                  inj_on_def [of "r.T0.trnrε s" "hom (tab1 r) (s  tab0 r)"]
            by simp
        qed
        have "μ.is_induced_map (Maps.REP τ.chine)"
          using θ ν Δ_eq Δ_def μ.is_induced_map_iff τ.chine_simps(1) Maps.ide_REP by blast
        hence 3: "Maps.REP τ.chine  μ.chine"
          using μ.chine_is_induced_map μ.induced_map_unique by simp
        have "SPN μ = τ"
        proof (intro Span.arr_eqI)
          show "Span.par (SPN μ) τ"
            using μ τ SPN_in_hom
            by (metis (no_types, lifting) SPN.preserves_cod SPN.preserves_dom Span.in_homE
                in_homE)
          show "Chn (SPN μ) = τ.chine"
          proof -
            have "Chn (SPN μ) = ⟦⟦spn μ⟧⟧"
              using μ SPN_def spn_def by auto
            also have "... = ⟦⟦μ.chine⟧⟧"
              using μ spn_def by fastforce
            also have "... = ⟦⟦Maps.REP τ.chine⟧⟧"
              using 3 isomorphic_symmetric Maps.CLS_eqI iso_class_eqI isomorphic_implies_hpar(3)
                    isomorphic_implies_hpar(4)
              by auto
            also have "... = τ.chine"
              using Maps.CLS_REP τ.chine_simps(1) by blast
            finally show ?thesis by blast
          qed
        qed
        thus "μ. «μ : r  s»  SPN μ = τ"
          using μ by auto
      qed
    qed

    theorem SPN_is_equivalence_pseudofunctor:
    shows "equivalence_pseudofunctor V H 𝖺 𝗂 src trg
             Span.vcomp Span.hcomp Span.assoc Span.unit Span.src Span.trg SPN Φ"
      ..

    text ‹
      We have completed the proof of the second half of the main result (CKS Theorem 4):
      B› is biequivalent (via SPN›) to Span(Maps(B))›.
    ›

    corollary
    shows "equivalent_bicategories V H 𝖺 𝗂 src trg
             Span.vcomp Span.hcomp Span.assoc Span.unit Span.src Span.trg"
      using SPN_is_equivalence_pseudofunctor equivalent_bicategories_def by blast

  end

end