diff --git a/thys/Bicategory/BicategoryOfSpans.thy b/thys/Bicategory/BicategoryOfSpans.thy --- a/thys/Bicategory/BicategoryOfSpans.thy +++ b/thys/Bicategory/BicategoryOfSpans.thy @@ -1,14758 +1,14758 @@ (* Title: BicategoryOfSpans Author: Eugene W. Stark , 2019 Maintainer: Eugene W. Stark *) section "Bicategories of Spans" theory BicategoryOfSpans imports CanonicalIsos SpanBicategory Category3.ConcreteCategory IsomorphismClass 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\<^sup>*\, 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: "\r. ide r \ \f g. is_left_adjoint f \ is_left_adjoint g \ isomorphic r (g \ f\<^sup>*)" and BS2: "\f g. \ is_left_adjoint f; is_left_adjoint g; src f = src g \ \ \r \. tabulation V H \ \ src trg r \ f g" and BS3: "\f f' \ \'. \ 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 V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D" and "bicategory_of_spans V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C" shows "bicategory_of_spans V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D" proof - interpret C: bicategory_of_spans V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C using assms by simp interpret C: chosen_right_adjoints V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C .. interpret D: bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D using assms equivalent_bicategories_def equivalence_pseudofunctor.axioms(1) pseudofunctor.axioms(2) by fast interpret D: chosen_right_adjoints V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D .. obtain F \ where F: "equivalence_pseudofunctor V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D F \" using assms equivalent_bicategories_def by blast interpret F: equivalence_pseudofunctor V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D F \ using F by simp interpret G: converse_equivalence_pseudofunctor V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D F \ .. write V\<^sub>C (infixr "\\<^sub>C" 55) write V\<^sub>D (infixr "\\<^sub>D" 55) write H\<^sub>C (infixr "\\<^sub>C" 53) write H\<^sub>D (infixr "\\<^sub>D" 53) write \\<^sub>C ("\\<^sub>C[_, _, _]") write \\<^sub>D ("\\<^sub>D[_, _, _]") write C.in_hhom ("\_ : _ \\<^sub>C _\") write C.in_hom ("\_ : _ \\<^sub>C _\") write D.in_hhom ("\_ : _ \\<^sub>D _\") write D.in_hom ("\_ : _ \\<^sub>D _\") write C.isomorphic (infix "\\<^sub>C" 50) write D.isomorphic (infix "\\<^sub>D" 50) write C.some_right_adjoint ("_\<^sup>*\<^sup>C" [1000] 1000) write D.some_right_adjoint ("_\<^sup>*\<^sup>D" [1000] 1000) show "bicategory_of_spans V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D" proof show "\r'. D.ide r' \ \f' g'. D.is_left_adjoint f' \ D.is_left_adjoint g' \ r' \\<^sub>D g' \\<^sub>D (f')\<^sup>*\<^sup>D" proof - fix r' assume r': "D.ide r'" obtain f g where fg: "C.is_left_adjoint f \ C.is_left_adjoint g \ G.G r' \\<^sub>C g \\<^sub>C f\<^sup>*\<^sup>C" using r' C.BS1 G.G\<^sub>1_props(1) G.G_ide by presburger have trg_g: "trg\<^sub>C g = G.G\<^sub>0 (trg\<^sub>D r')" using fg r' C.isomorphic_def C.hcomp_simps(2) by (metis C.ideD(1) C.isomorphic_implies_hpar(4) C.isomorphic_implies_ide(2) C.trg_hcomp D.ideD(1) G.G_props(1) C.in_hhomE) have trg_f: "trg\<^sub>C f = G.G\<^sub>0 (src\<^sub>D r')" using fg r' C.isomorphic_def C.hcomp_simps(1) by (metis C.ideD(1) C.isomorphic_implies_hpar(2-3) C.src_hcomp D.ideD(1) G.G_props(1) C.right_adjoint_simps(2) C.in_hhomE) interpret e_src: equivalence_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D \G.e (src\<^sub>D r')\ \G.d (src\<^sub>D r')\ \G.\ (src\<^sub>D r')\ \G.\ (src\<^sub>D r')\ using r' G.G\<^sub>0_props [of "src\<^sub>D r'"] by simp interpret e_trg: equivalence_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D \G.e (trg\<^sub>D r')\ \G.d (trg\<^sub>D r')\ \G.\ (trg\<^sub>D r')\ \G.\ (trg\<^sub>D r')\ using r' G.G\<^sub>0_props [of "trg\<^sub>D r'"] by simp interpret e: two_equivalences_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D \G.e (src\<^sub>D r')\ \G.d (src\<^sub>D r')\ \G.\ (src\<^sub>D r')\ \G.\ (src\<^sub>D r')\ \G.e (trg\<^sub>D r')\ \G.d (trg\<^sub>D r')\ \G.\ (trg\<^sub>D r')\ \G.\ (trg\<^sub>D r')\ .. interpret hom: subcategory V\<^sub>D \\\. \\ : src\<^sub>D (G.e (src\<^sub>D r')) \\<^sub>D src\<^sub>D (G.e (trg\<^sub>D r'))\\ using D.hhom_is_subcategory by simp interpret hom': subcategory V\<^sub>D \\\. \\ : trg\<^sub>D (G.e (src\<^sub>D r')) \\<^sub>D trg\<^sub>D (G.e (trg\<^sub>D r'))\\ using D.hhom_is_subcategory by simp interpret e: equivalence_of_categories hom'.comp hom.comp e.F e.G e.\ e.\ using e.induces_equivalence_of_hom_categories by simp define g' where "g' = G.d (trg\<^sub>D r') \\<^sub>D F g" have g': "D.is_left_adjoint g'" proof - have "D.equivalence_map (G.d (trg\<^sub>D r'))" using D.equivalence_map_def e_trg.dual_equivalence by blast hence "D.is_left_adjoint (G.d (trg\<^sub>D r'))" using r' D.equivalence_is_adjoint by simp moreover have "src\<^sub>D (G.d (trg\<^sub>D r')) = trg\<^sub>D (F g)" using fg r' G.G\<^sub>0_props trg_g by (simp add: C.left_adjoint_is_ide) ultimately show ?thesis unfolding g'_def using fg r' D.left_adjoints_compose F.preserves_left_adjoint by blast qed have 1: "D.is_right_adjoint (F f\<^sup>*\<^sup>C \\<^sub>D G.e (src\<^sub>D r'))" proof - have "D.equivalence_map (G.e (src\<^sub>D r'))" using D.equivalence_map_def e_src.equivalence_in_bicategory_axioms by blast hence "D.is_right_adjoint (G.e (src\<^sub>D r'))" using r' D.equivalence_is_adjoint by simp moreover have "D.is_right_adjoint (F f\<^sup>*\<^sup>C)" using fg C.left_adjoint_extends_to_adjoint_pair F.preserves_adjoint_pair by blast moreover have "src\<^sub>D (F f\<^sup>*\<^sup>C) = trg\<^sub>D (G.e (src\<^sub>D r'))" using fg r' G.G\<^sub>0_props trg_f by (simp add: C.right_adjoint_is_ide) 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\<^sup>*\<^sup>C \\<^sub>D G.e (src\<^sub>D r'))" using 1 by auto have f': "D.is_left_adjoint f' \ F f\<^sup>*\<^sup>C \\<^sub>D G.e (src\<^sub>D r') \\<^sub>D (f')\<^sup>*\<^sup>D" using f' D.left_adjoint_determines_right_up_to_iso D.left_adjoint_extends_to_adjoint_pair by blast have "r' \\<^sub>D G.d (trg\<^sub>D r') \\<^sub>D (G.e (trg\<^sub>D r') \\<^sub>D r' \\<^sub>D G.d (src\<^sub>D r')) \\<^sub>D G.e (src\<^sub>D r')" using r' e.\.components_are_iso e.\_in_hom [of r'] D.isomorphic_def hom.ide_char hom.arr_char hom.iso_char by auto also have 1: "... \\<^sub>D (G.d (trg\<^sub>D r') \\<^sub>D F (G.G r') \\<^sub>D G.e (src\<^sub>D r'))" proof - have "G.e (trg\<^sub>D r') \\<^sub>D r' \\<^sub>D G.d (src\<^sub>D r') \\<^sub>D F (G.G r')" by (simp add: D.isomorphic_symmetric G.G\<^sub>1_props(3) G.G_ide r') thus ?thesis using r' D.hcomp_isomorphic_ide D.hcomp_ide_isomorphic by simp qed also have 2: "... \\<^sub>D G.d (trg\<^sub>D r') \\<^sub>D (F g \\<^sub>D F f\<^sup>*\<^sup>C) \\<^sub>D G.e (src\<^sub>D r')" proof - have "F (G.G r') \\<^sub>D F (g \\<^sub>C f\<^sup>*\<^sup>C)" using fg F.preserves_iso C.isomorphic_def D.isomorphic_def by auto also have "F (g \\<^sub>C f\<^sup>*\<^sup>C) \\<^sub>D F g \\<^sub>D F f\<^sup>*\<^sup>C" using fg by (meson C.adjoint_pair_antipar(1) C.hseqE C.ideD(1) C.isomorphic_implies_hpar(2) C.right_adjoint_simps(1) D.isomorphic_symmetric F.weakly_preserves_hcomp) finally have "D.isomorphic (F (G.G r')) (F g \\<^sub>D F f\<^sup>*\<^sup>C)" by simp moreover have "\f g. D.ide (H\<^sub>D f g) \ src\<^sub>D f = trg\<^sub>D g" by (metis (no_types) D.hseqE D.ideD(1)) ultimately show ?thesis by (meson 1 D.hcomp_ide_isomorphic D.hcomp_isomorphic_ide D.hseqE D.ideD(1) D.isomorphic_implies_hpar(2) e_src.ide_left e_trg.ide_right) qed also have 3: "... \\<^sub>D (G.d (trg\<^sub>D r') \\<^sub>D F g) \\<^sub>D F f\<^sup>*\<^sup>C \\<^sub>D G.e (src\<^sub>D r')" proof - let ?a = "\\<^sub>D\<^sup>-\<^sup>1[G.d (trg\<^sub>D r'), F g, F f\<^sup>*\<^sup>C \\<^sub>D G.e (src\<^sub>D r')] \\<^sub>D (G.d (trg\<^sub>D r') \\<^sub>D \\<^sub>D[F g, F f\<^sup>*\<^sup>C, G.e (src\<^sub>D r')])" have "\?a : G.d (trg\<^sub>D r') \\<^sub>D (F g \\<^sub>D F f\<^sup>*\<^sup>C) \\<^sub>D G.e (src\<^sub>D r') \\<^sub>D (G.d (trg\<^sub>D r') \\<^sub>D F g) \\<^sub>D F f\<^sup>*\<^sup>C \\<^sub>D G.e (src\<^sub>D r')\" proof (intro D.comp_in_homI) show "\G.d (trg\<^sub>D r') \\<^sub>D \\<^sub>D[F g, F f\<^sup>*\<^sup>C, G.e (src\<^sub>D r')] : G.d (trg\<^sub>D r') \\<^sub>D (F g \\<^sub>D F f\<^sup>*\<^sup>C) \\<^sub>D G.e (src\<^sub>D r') \\<^sub>D G.d (trg\<^sub>D r') \\<^sub>D F g \\<^sub>D F f\<^sup>*\<^sup>C \\<^sub>D G.e (src\<^sub>D r')\" using fg r' 2 C.left_adjoint_is_ide D.hseqE D.ideD(1) D.isomorphic_implies_ide(2) D.src_hcomp D.assoc_in_hom [of "F g" "F f\<^sup>*\<^sup>C" "G.e (src\<^sub>D r')"] apply auto by (metis C.hseqE C.ideD(1) C.isomorphic_implies_hpar(2) C.right_adjoint_simps(3) D.hcomp_in_vhom D.ideD(1) D.ide_in_hom(2) e_trg.ide_right trg_f) show "\\\<^sub>D\<^sup>-\<^sup>1[G.d (trg\<^sub>D r'), F g, F f\<^sup>*\<^sup>C \\<^sub>D G.e (src\<^sub>D r')] : G.d (trg\<^sub>D r') \\<^sub>D F g \\<^sub>D F f\<^sup>*\<^sup>C \\<^sub>D G.e (src\<^sub>D r') \\<^sub>D (G.d (trg\<^sub>D r') \\<^sub>D F g) \\<^sub>D F f\<^sup>*\<^sup>C \\<^sub>D G.e (src\<^sub>D r')\" proof - have "D.ide (F f\<^sup>*\<^sup>C \\<^sub>D G.e (src\<^sub>D r'))" using f' D.isomorphic_implies_ide by auto moreover have "src\<^sub>D (F g) = trg\<^sub>D (F f\<^sup>*\<^sup>C \\<^sub>D G.e (src\<^sub>D r'))" using fg r' f' 2 D.hseqE D.ideD(1) D.isomorphic_implies_ide D.trg_hcomp by metis ultimately show ?thesis using fg r' trg_g C.left_adjoint_is_ide D.assoc'_in_hom(2) by simp qed qed moreover have "D.iso ?a" using fg r' D.isos_compose by (metis 2 C.left_adjoint_is_ide C.right_adjoint_simps(1) D.arrI D.assoc_simps(3) D.iso_hcomp D.hseqE D.ideD(1) D.ide_is_iso D.iso_assoc D.iso_assoc' D.isomorphic_implies_ide(1) D.isomorphic_implies_ide(2) D.trg_hcomp F.preserves_ide calculation e_src.ide_left e_trg.ide_right f') ultimately show ?thesis using D.isomorphic_def by auto qed also have "(G.d (trg\<^sub>D r') \\<^sub>D F g) \\<^sub>D F f\<^sup>*\<^sup>C \\<^sub>D G.e (src\<^sub>D r') \\<^sub>D g' \\<^sub>D f'\<^sup>*\<^sup>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' \\<^sub>D f'\<^sup>*\<^sup>D)" by simp thus "\f' g'. D.is_left_adjoint f' \ D.is_left_adjoint g' \ r' \\<^sub>D g' \\<^sub>D f'\<^sup>*\<^sup>D" using f' g' by auto qed show "\f f' \ \'. \ D.is_left_adjoint f; D.is_left_adjoint f'; \\ : f \\<^sub>D f'\; \\' : f \\<^sub>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 \\<^sub>D f'\" and \': "\\' : f \\<^sub>D f'\" have "C.is_left_adjoint (G.G f) \ C.is_left_adjoint (G.G f')" using f f' G.preserves_left_adjoint by simp moreover have "\G.G \ : G.G f \\<^sub>C G.G f'\ \ \G.G \' : G.G f \\<^sub>C G.G f'\" using \ \' G.preserves_hom by simp ultimately have "C.iso (G.G \) \ C.iso (G.G \') \ G.G \ = G.G \'" 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; src\<^sub>D f = src\<^sub>D g \ \ \r \. tabulation V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D r \ f g" proof - fix f g assume f: "D.is_left_adjoint f" assume g: "D.is_left_adjoint g" assume fg: "src\<^sub>D f = src\<^sub>D g" have "C.is_left_adjoint (G.G f)" using f G.preserves_left_adjoint by blast moreover have "C.is_left_adjoint (G.G g)" using g G.preserves_left_adjoint by blast moreover have "src\<^sub>C (G.G f) = src\<^sub>C (G.G g)" using f g D.left_adjoint_is_ide fg by simp ultimately have 1: "\r \. tabulation V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C r \ (G.G f) (G.G g)" using C.BS2 by simp obtain r \ where \: "tabulation V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C r \ (G.G f) (G.G g)" using 1 by auto interpret \: tabulation V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C r \ \G.G f\ \G.G g\ using \ by simp obtain r' where r': "D.ide r' \ D.in_hhom r' (trg\<^sub>D f) (trg\<^sub>D g) \ C.isomorphic (G.G r') r" using f g \.ide_base \.tab_in_hom G.locally_essentially_surjective by (metis D.obj_trg G.preserves_reflects_arr G.preserves_trg \.leg0_simps(2-3) \.leg1_simps(2,4) \.base_in_hom(1)) obtain \ where \: "\\ : r \\<^sub>C G.G r'\ \ C.iso \" using r' C.isomorphic_symmetric by blast have \: "tabulation V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C (G.G r') (V\<^sub>C (H\<^sub>C \ (G.G f)) \) (G.G f) (G.G g)" using \ \.is_preserved_by_base_iso by simp have 1: "\\'. \\' : g \\<^sub>D H\<^sub>D r' f\ \ G.G \' = V\<^sub>C (G.\ (r', f)) (V\<^sub>C (H\<^sub>C \ (G.G f)) \)" proof - have "D.ide g" by (simp add: D.left_adjoint_is_ide g) moreover have "D.ide (H\<^sub>D r' f)" using f r' D.left_adjoint_is_ide by auto moreover have "src\<^sub>D g = src\<^sub>D (H\<^sub>D r' f)" using fg by (simp add: calculation(2)) moreover have "trg\<^sub>D g = trg\<^sub>D (H\<^sub>D r' f)" using calculation(2) r' by auto moreover have "\V\<^sub>C (G.\ (r', f)) (V\<^sub>C (H\<^sub>C \ (G.G f)) \) : G.G g \\<^sub>C G.G (H\<^sub>D r' f)\" using f g r' \ G.\_in_hom [of r' f] 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 \\<^sub>D H\<^sub>D r' f\ \ G.G \' = V\<^sub>C (G.\ (r', f)) (V\<^sub>C (H\<^sub>C \ (G.G f)) \)" using 1 by auto have "tabulation V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D r' \' f g" proof - have "V\<^sub>C (C.inv (G.\ (r', f))) (G.G \') = V\<^sub>C (H\<^sub>C \ (G.G f)) \" using r' f \' C.comp_assoc C.comp_cod_arr G.\_components_are_iso C.invert_side_of_triangle(1) [of "G.G \'" "G.\ (r', f)" "V\<^sub>C (H\<^sub>C \ (G.G f)) \"] by (metis (no_types, lifting) D.arrI D.in_hhom_def D.left_adjoint_is_ide G.preserves_arr) thus ?thesis using \ \' G.reflects_tabulation by (simp add: D.left_adjoint_is_ide f r') qed thus "\r' \'. tabulation V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D 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 = span_bicategory C prj0 prj1 + r: identity_arrow_of_spans C r for C :: "'a comp" (infixr "\" 55) and prj0 :: "'a \ 'a \ 'a" ("\

\<^sub>0[_, _]") and prj1 :: "'a \ 'a \ 'a" ("\

\<^sub>1[_, _]") and r :: "'a arrow_of_spans_data" begin text \ CKS say: ``Suppose \r = (r\<^sub>0, R, r\<^sub>1): A \ B\ and put \f = (1, R, r\<^sub>0)\, \g = (1, R, r\<^sub>1)\. Let \k\<^sub>0, k\<^sub>1\ form a kernel pair for \r\<^sub>0\ and define \\\ by \k\<^sub>0\ = k\<^sub>1\ = 1\<^sub>R\.'' \ 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 \ \

\<^sub>0[r0, r0]" abbreviation k1 where "k1 \ \

\<^sub>1[r0, r0]" text \ Here \ra\ is the apex \R\ of the span \(r\<^sub>0, R, r\<^sub>1)\, 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, r\<^sub>1)\ to \r \ f\, which is the span \(k\<^sub>0, r\<^sub>1 \ k\<^sub>1)\. \ 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\" proof show 1: "arr \" using arr_char \.arrow_of_spans_axioms by simp show "dom \ = g" using 1 dom_char ideD(2) ide_g by fastforce show "cod \ = r \ f" using 1 cod_char rf Cod_rf.apex_def by simp qed 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 = (u\<^sub>0, U, u\<^sub>1): X \ A\, \v = (v\<^sub>0, V, v\<^sub>1): X \ B\, \\: v \ ru\ as in \T1\. Let \P\ be the pullback of \u\<^sub>1, r\<^sub>0\. Let \w = (v\<^sub>0, V, p\<^sub>1\): X \ R\, \\ = p\<^sub>0\: 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 = "\

\<^sub>0[r0, ?u1]" let ?p1 = "\

\<^sub>1[r0, ?u1]" let ?w1 = "?p1 \ ?\" define w where "w = mkIde ?v0 ?w1" let ?Q = "?R \\ ?w1" let ?q1 = "\

\<^sub>0[?R, ?w1]\. What CKS say is only correct if the projection \\

\<^sub>0[?R, ?w1]\ is an identity, which can't be guaranteed for an arbitrary choice of pullbacks. \ define \ where "\ = \Chn = ?p0 \ ?\ \ \

\<^sub>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 \ \

\<^sub>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 \ ?\ \ \

\<^sub>0[?R, ?w1]" using \_def by simp interpret \: arrow_of_spans C \ proof show 1: "\Chn \ : Dom_\.apex \\<^sub>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 \ ?\ \ \

\<^sub>0[?R, ?w1]" using Cod_\_leg0_eq Chn_\_eq by simp also have "... = ?v0 \ \

\<^sub>0[?R, ?w1]" proof - have "?u0 \ ?p0 \ ?\ \ \

\<^sub>0[?R, ?w1] = (?u0 \ ?p0 \ ?\) \ \

\<^sub>0[?R, ?w1]" using C.comp_assoc by simp also have "... = ?v0 \ \

\<^sub>0[?R, ?w1]" proof - have "?u0 \ ?p0 \ ?\ = (?u0 \ ?p0) \ ?\" using C.comp_assoc by simp also have "... = \.cod.leg0 \ ?\" proof - have "\.cod.leg0 = ru.leg0" using \ cod_char hcomp_def hseq_ru by auto also have "... = ?u0 \ ?p0" using hcomp_def hseq_ru by auto finally show ?thesis by simp qed 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 \ ?\ \ \

\<^sub>0[?R, ?w1]" using Cod_\_leg1_eq Chn_\_eq by simp also have "... = r0 \ ?q1" proof - have "?u1 \ ?p0 \ ?\ \ \

\<^sub>0[?R, ?w1] = ((?u1 \ ?p0) \ ?\) \ \

\<^sub>0[?R, ?w1]" using C.comp_assoc by fastforce also have "... = ((r0 \ ?p1) \ ?\) \ \

\<^sub>0[?R, ?w1]" using C.pullback_commutes' ru.legs_form_cospan by simp also have "... = r0 \ ?w1 \ \

\<^sub>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 \

\<^sub>0[?R, ?w1]\, noting that \\

\<^sub>0[?R, ?w1]\ is the pullback of an identity, hence is an isomorphism. Then \?v0 \ \

\<^sub>0[?R, ?w1] \ Chn \ = ?v0\ and \?v1 \ \

\<^sub>1[?R, ?w1] \ Chn \ = ?v1 \ ?w1\, showing that \\\ is an arrow of spans. \ let ?\' = "\

\<^sub>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_\ \ w1 r.dom.apex_def r.cod.apex_def r.chine_eq_apex 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 \\<^sub>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) gw.chine_simps(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 apply (intro in_homI) 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 \ \ : 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 r\ by auto show "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 \" proof - have "Chn (T.composite_cell w \ \ \) = Chn (T.composite_cell w \) \ Chn \" using 1 Chn_vcomp by blast also have "... = (Chn (r \ \) \ Chn \[r, f, w] \ Chn (\ \ w)) \ Chn \" proof - have "seq (r \ \) (\[r, f, w] \ (\ \ w)) \ 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 "... = ?\" proof - let ?LHS = "Chn (r \ \) \ Chn \[r, f, w] \ Chn (\ \ w) \ Chn \" have Chn_r\: "Chn (r \ \) = \r.chine \ \

\<^sub>1[r0, r0 \ ?q1] \r0, ?u1\ \.chine \ \

\<^sub>0[r0, r0 \ ?q1]\" using r\ 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 \\<^sub>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 \\<^sub>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 \ \

\<^sub>0[r0, r0 \ ?q1]) \ Chn \[r, f, w] \ Chn (\ \ w) \ Chn \" proof - have "?p0 \ Chn (r \ \) = \.chine \ \

\<^sub>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) r\ u.cod_simps(3)) thus ?thesis by argo qed also have "... = ?p0 \ ?\ \ (rfw.Prj\<^sub>0\<^sub>0 \ Chn \[r, f, w]) \ Chn (\ \ w) \ Chn \" using w_def \_def C.comp_assoc by simp also have "... = ?p0 \ ?\ \ (rfw.Prj\<^sub>0 \ Chn (\ \ w)) \ Chn \" using Chn_arfw rfw.prj_chine_assoc C.comp_assoc by simp also have "... = ?p0 \ ?\ \ ?\' \ Chn \" proof - have "rfw.Prj\<^sub>0 \ Chn (\ \ w) = \

\<^sub>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.Prj\<^sub>0 \ 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 \ \

\<^sub>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.Prj\<^sub>1 \ Chn \[r, f, w]) \ Chn (\ \ w) \ Chn \" using w_def Dom_\_leg1_eq C.comp_assoc by simp also have "... = r.chine \ (rfw.Prj\<^sub>1\<^sub>1 \ 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.Prj\<^sub>1\<^sub>1 \ Chn (\ \ w) = (k1 \ \

\<^sub>1[k0, ?w1]) \ \?\ \ ?q1 \k0, ?w1\ ?\'\" using w_def Chn_\w C.comp_cod_arr by simp also have "... = k1 \ \

\<^sub>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.Prj\<^sub>1\<^sub>1 \ 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 = "\

\<^sub>0[?R, ?w1]" let ?p0' = "\

\<^sub>0[?R, ?w1']" let ?p1 = "\

\<^sub>1[?R, ?w1]" let ?p1' = "\

\<^sub>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 \ \ : 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 by auto interpret r\: two_composable_arrows_of_spans C prj0 prj1 r \ using r\ 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 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 r\ 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 = (w\<^sub>0, W, r\<^sub>0 w\<^sub>1)\, \gw = (w\<^sub>0, W, r\<^sub>1 w\<^sub>1)\, \emph{etc}. So \\: W \ W'\ satisfies \w\<^sub>0 = w\<^sub>0' \\, \r\<^sub>1 w\<^sub>1 = r\<^sub>1 w\<^sub>1' \\. But the equation \(r\)(\w) = (r\')(\w')\\ gives \w\<^sub>1 = w\<^sub>1'\. 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 \\<^sub>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 \\<^sub>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 = \\

\<^sub>1[r0, r0 \ ?p1] \r0, ?u1\ \.chine \ \

\<^sub>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 \\<^sub>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 = \\

\<^sub>1[r0, r0 \ ?p1'] \r0, ?u1\ \'.chine \ \

\<^sub>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 \\<^sub>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' \\<^sub>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 apply (intro C.commutative_squareI) by auto have "C.commutative_square r0 ?u1 (?p1' \ \.chine) (\'.chine \ \.chine)" proof have 1: "r0 \ ?p1' = ?u1 \ \'.chine" using \'.leg1_commutes Cod_\'_1 Dom_\'_1 fw'.leg1_composite by simp show "C.cospan r0 ?u1" using ru.legs_form_cospan(1) by blast show "C.span (?p1' \ \.chine) (\'.chine \ \.chine)" using \.chine_in_hom \'.chine_in_hom by (metis "1" C.dom_comp C.in_homE C.prj1_simps(1) C.prj1_simps(2) C.seqI Cod_\'_1 Dom_\'.leg_simps(3) Chn_\ \'.leg1_commutes cospan') show "C.dom r0 = C.cod (?p1' \ \.chine)" using \.chine_in_hom by (metis C.cod_comp C.prj1_simps(3) \C.span (?p1' \ \.chine) (\'.chine \ \.chine)\ cospan' r.dom.apex_def r.chine_eq_apex r.chine_simps(2)) show "r0 \ ?p1' \ \.chine = ?u1 \ \'.chine \ \.chine" using 1 \.chine_in_hom C.comp_assoc by metis qed have "C.commutative_square r0 ?u1 \

\<^sub>1[r0, r0 \ ?p1] (\.chine \ \

\<^sub>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 \

\<^sub>1[r0, r0 \ ?p1] (\.chine \ \

\<^sub>0[r0, r0 \ ?p1])" using fw.leg1_composite by auto have "C.commutative_square r0 ?u1 \

\<^sub>1[r0, r0 \ ?p1'] (\'.chine \ \

\<^sub>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.Prj\<^sub>0\<^sub>1 rfw.Prj\<^sub>0" using C.pullback_commutes' gw.legs_form_cospan(1) rfw.prj_simps(2) C.comp_assoc C.comp_cod_arr apply (intro C.commutative_squareI) by auto have "C.commutative_square ?R ?w1' rfw'.Prj\<^sub>0\<^sub>1 rfw'.Prj\<^sub>0" using cospan' apply (intro C.commutative_squareI) apply simp_all by (metis C.comp_assoc C.prj0_simps_arr C.pullback_commutes' arrow_of_spans_data.select_convs(2) rfw'.prj_simps(3) span_data.select_convs(1-2)) have "C.commutative_square r0 (r0 \ ?p1) rfw.Prj\<^sub>1\<^sub>1 \rfw.Prj\<^sub>0\<^sub>1 \ra, ?w1\ rfw.Prj\<^sub>0\" 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'.Prj\<^sub>1\<^sub>1 \rfw'.Prj\<^sub>0\<^sub>1 \ra, ?w1'\ rfw'.Prj\<^sub>0\" 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 \\<^sub>C r\'.cod.apex\" proof (intro C.comp_in_homI) show "\\.chine : ?R \\ ?w1 \\<^sub>C ?R \\ ?w1'\" using Chn_\ by simp show "\\w'.chine : ?R \\ ?w1' \\<^sub>C Cod_\.leg0 \\ w'.cod.leg1\" using Chn_\w' by simp show "\Chn \[r, f, w'] : Cod_\.leg0 \\ w'.cod.leg1 \\<^sub>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 \\<^sub>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 \\<^sub>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.prj\<^sub>1 ?LHS" using LHS r\'_cod_apex_eq by auto show "C.seq ru.prj\<^sub>1 ?RHS" using RHS r\'_cod_apex_eq by auto show "ru.prj\<^sub>0 \ ?LHS = ru.prj\<^sub>0 \ ?RHS" proof - have "ru.prj\<^sub>0 \ ?LHS = (ru.prj\<^sub>0 \ r\'.chine) \ Chn \[r, f, w'] \ \w'.chine \ \.chine" using C.comp_assoc by simp also have "... = ((\'.chine \ \

\<^sub>0[r0, r0 \ ?p1']) \ Chn \[r, f, w']) \ \w'.chine \ \.chine" using Chn_r\'_eq C.comp_assoc fw' \C.commutative_square r0 ?u1 \

\<^sub>1[r0, r0 \ ?p1'] (\'.chine \ \

\<^sub>0[r0, r0 \ ?p1'])\ by simp also have "... = \'.chine \ (\

\<^sub>0[r0, r0 \ ?p1'] \ Chn \[r, f, w']) \ \w'.chine \ \.chine" using C.comp_assoc by simp also have "... = \'.chine \ (\rfw'.Prj\<^sub>0\<^sub>1 \?R, ?w1'\ rfw'.Prj\<^sub>0\ \ \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'.Prj\<^sub>1\<^sub>1 \rfw'.Prj\<^sub>0\<^sub>1 \?R, ?w1'\ rfw'.Prj\<^sub>0\\ by simp also have "... = \'.chine \ \.chine" proof - have "\rfw'.Prj\<^sub>0\<^sub>1 \?R, ?w1'\ rfw'.Prj\<^sub>0\ \ \w'.chine = gw'.apex" proof (intro C.prj_joint_monic [of ?R ?w1' "\rfw'.Prj\<^sub>0\<^sub>1 \?R, ?w1'\ rfw'.Prj\<^sub>0\ \ \w'.chine" gw'.apex]) show "C.cospan ?R ?w1'" using fw'.legs_form_cospan(1) by simp show "C.seq ?p1' (\rfw'.Prj\<^sub>0\<^sub>1 \?R, ?w1'\ rfw'.Prj\<^sub>0\ \ \w'.chine)" proof (intro C.seqI' C.comp_in_homI) show "\\w'.chine : Dom_\.leg0 \\ w'.leg1 \\<^sub>C Cod_\.leg0 \\ w'.cod.leg1\" using \w'.chine_composite_in_hom by simp show "\\rfw'.Prj\<^sub>0\<^sub>1 \?R, w'.leg1\ rfw'.Prj\<^sub>0\ : Cod_\.leg0 \\ w'.cod.leg1 \\<^sub>C ?R \\ w'.leg1\" using \C.commutative_square ?R ?w1' rfw'.Prj\<^sub>0\<^sub>1 rfw'.Prj\<^sub>0\ C.tuple_in_hom [of ?R ?w1' rfw'.Prj\<^sub>0\<^sub>1 rfw'.Prj\<^sub>0] rf rf.leg0_composite by auto show "\?p1' : ?R \\ w'.leg1 \\<^sub>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'.Prj\<^sub>0\<^sub>1 \?R, ?w1'\ rfw'.Prj\<^sub>0\ \ \w'.chine = ?p0' \ gw'.apex" proof - have "?p0' \ \rfw'.Prj\<^sub>0\<^sub>1 \?R, ?w1'\ rfw'.Prj\<^sub>0\ \ \w'.chine = (?p0' \ \rfw'.Prj\<^sub>0\<^sub>1 \?R, ?w1'\ rfw'.Prj\<^sub>0\) \ \w'.chine" using C.comp_assoc by simp also have "... = rfw'.Prj\<^sub>0 \ \w'.chine" using \C.commutative_square ?R ?w1' rfw'.Prj\<^sub>0\<^sub>1 rfw'.Prj\<^sub>0\ by auto also have "... = \

\<^sub>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'.Prj\<^sub>0\<^sub>1 \ra, ?w1'\ rfw'.Prj\<^sub>0\ \ \w'.chine = ?p1' \ gw'.apex" proof - have "?p1' \ \rfw'.Prj\<^sub>0\<^sub>1 \ra, ?w1'\ rfw'.Prj\<^sub>0\ \ \w'.chine = (?p1' \ \rfw'.Prj\<^sub>0\<^sub>1 \ra, ?w1'\ rfw'.Prj\<^sub>0\) \ \w'.chine" using C.comp_assoc by simp also have "... = rfw'.Prj\<^sub>0\<^sub>1 \ \w'.chine" using \C.commutative_square ?R ?w1' rfw'.Prj\<^sub>0\<^sub>1 rfw'.Prj\<^sub>0\ by simp also have "... = k0 \ \

\<^sub>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 "... = \

\<^sub>0[r0, ?u1] \ ?RHS" using RHS 2 C.prj_tuple [of r0 ?u1] by simp finally show ?thesis by simp qed show "ru.prj\<^sub>1 \ ?LHS = ru.prj\<^sub>1 \ ?RHS" proof - have "ru.prj\<^sub>1 \ ?LHS = (ru.prj\<^sub>1 \ r\'.chine) \ Chn \[r, f, w'] \ \w'.chine \ \.chine" using C.comp_assoc by simp also have "... = \

\<^sub>1[r0, fw'.leg1] \ Chn \[r, f, w'] \ \w'.chine \ \.chine" using Chn_r\' Chn_r\'_eq fw' \C.commutative_square r0 ?u1 \

\<^sub>1[r0, r0 \ ?p1'] (\'.chine \ \

\<^sub>0[r0, r0 \ ?p1'])\ by simp also have "... = (rfw'.Prj\<^sub>1 \ 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'.Prj\<^sub>1\<^sub>1 \ \w'.chine) \ \.chine" using rfw'.prj_chine_assoc C.comp_assoc by simp also have "... = ((k1 \ \

\<^sub>1[k0, ?w1']) \ \w'.chine) \ \.chine" using C.comp_cod_arr by simp also have "... = (k1 \ \

\<^sub>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.prj\<^sub>1 \ ?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 \\<^sub>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 \\<^sub>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.prj\<^sub>1 ?LHS" using LHS r\_cod_apex_eq by auto show "C.seq ru.prj\<^sub>1 ?RHS" using RHS r\_cod_apex_eq by auto show "ru.prj\<^sub>0 \ ?LHS = ru.prj\<^sub>0 \ ?RHS" proof - have "ru.prj\<^sub>0 \ ?LHS = (ru.prj\<^sub>0 \ r\.chine) \ Chn \[r, f, w] \ \w.chine" using C.comp_assoc by simp also have "... = (\.chine \ \

\<^sub>0[r0, f.leg1 \ fw.prj\<^sub>1]) \ Chn \[r, f, w] \ \w.chine" using Chn_r\_eq Dom_\_1 Cod_\_1 fw.leg1_composite \C.commutative_square r0 ?u1 \

\<^sub>1[r0, r0 \ ?p1] (\.chine \ \

\<^sub>0[r0, r0 \ ?p1])\ by simp also have "... = \.chine \ (\

\<^sub>0[r0, r0 \ ?p1] \ Chn \[r, f, w]) \ \w.chine" using C.comp_assoc by simp also have "... = \.chine \ \rfw.Prj\<^sub>0\<^sub>1 \?R, ?w1\ rfw.Prj\<^sub>0\ \ \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 "\

\<^sub>0[r0, r0 \ ?p1] \ rfw.chine_assoc = \rfw.Prj\<^sub>0\<^sub>1 \?R, ?w1\ rfw.Prj\<^sub>0\" using rfw.chine_assoc_def \C.commutative_square r0 (r0 \ ?p1) rfw.Prj\<^sub>1\<^sub>1 \rfw.Prj\<^sub>0\<^sub>1 \?R, ?w1\ rfw.Prj\<^sub>0\\ by simp ultimately show ?thesis by simp qed also have "... = \.chine \ (?R \\ ?w1)" proof - have "\rfw.Prj\<^sub>0\<^sub>1 \?R, ?w1\ rfw.Prj\<^sub>0\ \ \w.chine = ?R \\ ?w1" proof (intro C.prj_joint_monic [of ?R ?w1 "\rfw.Prj\<^sub>0\<^sub>1 \?R, ?w1\ rfw.Prj\<^sub>0\ \ \w.chine" "?R \\ ?w1"]) show "C.cospan ?R ?w1" by fact show "C.seq ?p1 (\rfw.Prj\<^sub>0\<^sub>1 \?R, ?w1\ rfw.Prj\<^sub>0\ \ \w.chine)" proof - have "C.seq rfw.Prj\<^sub>0\<^sub>1 \w.chine" by (meson C.seqI' Chn_in_hom \w rfw.prj_in_hom(2) \C.commutative_square ?R ?w1 rfw.Prj\<^sub>0\<^sub>1 rfw.Prj\<^sub>0\) thus ?thesis using \C.commutative_square ?R ?w1 rfw.Prj\<^sub>0\<^sub>1 rfw.Prj\<^sub>0\ 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.Prj\<^sub>0\<^sub>1 \?R, ?w1\ rfw.Prj\<^sub>0\ \ \w.chine = ?p0 \ (?R \\ ?w1)" proof - have "?p0 \ \rfw.Prj\<^sub>0\<^sub>1 \?R, ?w1\ rfw.Prj\<^sub>0\ \ \w.chine = (?p0 \ \rfw.Prj\<^sub>0\<^sub>1 \?R, ?w1\ rfw.Prj\<^sub>0\) \ \w.chine" using C.comp_assoc by simp also have "... = rfw.Prj\<^sub>0 \ \w.chine" using \C.commutative_square ?R ?w1 rfw.Prj\<^sub>0\<^sub>1 rfw.Prj\<^sub>0\ by simp also have "... = \

\<^sub>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.Prj\<^sub>0\<^sub>1 \?R, ?w1\ rfw.Prj\<^sub>0\ \ \w.chine = ?p1 \ (?R \\ ?w1)" proof - have "?p1 \ \rfw.Prj\<^sub>0\<^sub>1 \?R, ?w1\ rfw.Prj\<^sub>0\ \ \w.chine = (?p1 \ \rfw.Prj\<^sub>0\<^sub>1 \?R, ?w1\ rfw.Prj\<^sub>0\) \ \w.chine" using C.comp_assoc by simp also have "... = rfw.Prj\<^sub>0\<^sub>1 \ \w.chine" using \C.commutative_square ?R ?w1 rfw.Prj\<^sub>0\<^sub>1 rfw.Prj\<^sub>0\ by simp also have "... = (k0 \ \

\<^sub>1[k0, ?w1]) \ \\.chine \ ?p1 \k0, ?w1\ ?p0\" using Chn_\w_eq C.comp_cod_arr by simp also have "... = k0 \ \

\<^sub>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.prj\<^sub>0 \ ?RHS" using 2 by simp finally show ?thesis by blast qed show "ru.prj\<^sub>1 \ ?LHS = ru.prj\<^sub>1 \ ?RHS" proof - have "ru.prj\<^sub>1 \ ?LHS = (ru.prj\<^sub>1 \ r\.chine) \ Chn \[r, f, w] \ \w.chine" using C.comp_assoc by simp also have "... = (r.chine \ \

\<^sub>1[r0, r0 \ ?p1]) \ Chn \[r, f, w] \ \w.chine" proof - have "r\.chine \ C.null \ \

\<^sub>1[r.cod.leg0, Cod_\.leg1] \ r\.chine = r.chine \ \

\<^sub>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.Prj\<^sub>1 \ Chn \[r, f, w]) \ \w.chine" using C.comp_assoc fw.leg1_composite by simp also have "... = r.chine \ rfw.Prj\<^sub>1\<^sub>1 \ \w.chine" using ide_f ide_w hseq_rf hseq_char \_ide rfw.prj_chine_assoc(1) by auto also have "... = r.chine \ k1 \ \

\<^sub>1[k0, ?w1] \ \w.chine" using C.comp_cod_arr C.comp_assoc by simp also have "... = r.chine \ k1 \ \.chine \ \

\<^sub>1[Dom_\.leg0, ?w1]" using Chn_\w_eq \C.commutative_square k0 ?w1 (\.chine \ \

\<^sub>1[ra, w.leg1]) \

\<^sub>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.prj\<^sub>1 \ ?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 (\

\<^sub>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 \\<^sub>C Cod_\.apex\" proof - have "\Chn \: gw.apex \\<^sub>C gw'.apex\" using Chn_in_hom \ gw'.chine_eq_apex gw.chine_eq_apex by force moreover have "\?p0' : gw'.apex \\<^sub>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 \\<^sub>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 \w\<^sub>1 = w\<^sub>1'\ 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 f\: two_composable_arrows_of_spans C prj0 prj1 f ?\ using hseq_f\ hseq_char by (unfold_locales, simp) interpret f\: arrow_of_spans C \f \ ?\\ using f\.composite_is_arrow arr_char by simp interpret g\: two_composable_arrows_of_spans C prj0 prj1 g ?\ using hseq_g\ hseq_char by (unfold_locales, simp) interpret g\: 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 \ ?\)" proof - have "Chn (g \ ?\) = \?p1 \?R, ?w1'\ ?p0' \ Chn \\" using Chn_g\ by simp also have "... = \.chine" text \Here was another score by sledgehammer while I was still trying to understand it.\ using ** C.prj_joint_monic by (metis C.prj1_simps(1) C.tuple_prj cospan cospan') finally show ?thesis by simp qed 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 \ ?\ : f \ w \ f \ w'\" using \ ide_f by auto show par: "par \ (\' \ (f \ ?\))" using \ \' f\ 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 ("_\<^sup>*" [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\<^sup>*\. \ show "\r. ide r \ \f g. is_left_adjoint f \ is_left_adjoint g \ r \ g \ f\<^sup>*" 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\<^sup>*" using \ \.yields_isomorphic_representation' \.T0.is_map left_adjoint_extends_to_adjoint_pair isomorphic_def [of "r.g \ r.f\<^sup>*" r] isomorphic_symmetric by auto thus "\f g. is_left_adjoint f \ is_left_adjoint g \ r \ g \ f\<^sup>*" 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 = (f\<^sub>1, f\<^sub>0)\ and \g = (g\<^sub>1, g\<^sub>0)\ be given. Let \f' = (f\<^sub>1 \ C.inv f\<^sub>0, C.cod f\<^sub>0)\ and \g' = (g\<^sub>1 \ C.inv g\<^sub>0, C.cod g\<^sub>0)\; 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 = (f\<^sub>1 \ C.inv f\<^sub>0, g\<^sub>1 \ C.inv g\<^sub>0)\. 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 f\<^sub>a \\<^sub>f \\<^sub>f where ff\<^sub>a: "adjunction_in_bicategory vcomp hcomp assoc unit src trg f f\<^sub>a \\<^sub>f \\<^sub>f" using f adjoint_pair_def by auto interpret ff\<^sub>a: adjunction_in_bicategory vcomp hcomp assoc unit src trg f f\<^sub>a \\<^sub>f \\<^sub>f using ff\<^sub>a 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 g\<^sub>a \\<^sub>g \\<^sub>g where G: "adjunction_in_bicategory vcomp hcomp assoc unit src trg g g\<^sub>a \\<^sub>g \\<^sub>g" using g adjoint_pair_def by auto interpret gg\<^sub>a: adjunction_in_bicategory vcomp hcomp assoc unit src trg g g\<^sub>a \\<^sub>g \\<^sub>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 = \

\<^sub>0[f'.leg1, f'.leg1], Leg1 = g'.leg1 \ \

\<^sub>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 \\<^sub>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 C.iso_inv_iso 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 \\<^sub>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 C.iso_inv_iso 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'` iso_inv_iso 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 f\<^sub>a \ \ where f\<^sub>a: "adjunction_in_bicategory vcomp hcomp assoc unit src trg f f\<^sub>a \ \" using f adjoint_pair_def by auto obtain f'\<^sub>a \' \' where f'\<^sub>a: "adjunction_in_bicategory vcomp hcomp assoc unit src trg f' f'\<^sub>a \' \'" using f' adjoint_pair_def adjunction_def by auto interpret f\<^sub>a: adjunction_in_bicategory vcomp hcomp assoc unit src trg f f\<^sub>a \ \ using f\<^sub>a by simp interpret f'\<^sub>a: adjunction_in_bicategory vcomp hcomp assoc unit src trg f' f'\<^sub>a \' \' using f'\<^sub>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 C.iso_inv_iso \ dom_char cod_char iso_char arr_char \.arrow_of_spans_axioms 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 V\<^sub>1 H\<^sub>1 \\<^sub>1 \\<^sub>1 src\<^sub>1 trg\<^sub>1" shows "bicategory_of_spans V\<^sub>1 H\<^sub>1 \\<^sub>1 \\<^sub>1 src\<^sub>1 trg\<^sub>1" 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\<^sup>*\, not a tabulation of some \r\ isomorphic to \g \ f\<^sup>*\. \ lemma BS2': assumes "is_left_adjoint f" and "is_left_adjoint g" and "src f = src g" and "isomorphic (g \ f\<^sup>*) r" shows "\\. tabulation V H \ \ src trg r \ f g" proof - have 1: "is_left_adjoint f \ is_left_adjoint g \ g \ f\<^sup>* \ r" using assms BS1 by simp obtain \ where \: "\\ : g \ f\<^sup>* \ 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\<^sub>\ r' \'" have \: "\?\ : g \ f\<^sup>* \ r'\ \ iso ?\" using \'.yields_isomorphic_representation by blast have "\\ \ inv ?\ : r' \ r\ \ iso (\ \ inv ?\)" using \ \ iso_inv_iso isos_compose inv_in_hom 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\<^sup>*\, 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\<^sup>*\ 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 "isomorphic ((g \ w) \ (src w)\<^sup>*) (g \ w)" proof - have "(g \ w) \ src (g \ w) \ g \ w" using calculation isomorphic_implies_ide(2) by auto moreover have "isomorphic (src (g \ w)) (src w)\<^sup>*" 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)\<^sup>*\ 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)\<^sup>*" using adjoint_pair_def src_w.adjunction_in_bicategory_axioms by auto ultimately have "src w \ (src w)\<^sup>*" 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\<^sup>*)" using assms BS1 [of w] by auto have m\<^sub>a: "adjoint_pair m m\<^sup>* \ isomorphic w (n \ m\<^sup>*)" 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\<^sup>*" using mn isomorphic_implies_ide by auto have "trg (n \ m\<^sup>*) = trg w" using mn m\<^sub>a isomorphic_def by (metis (no_types, lifting) arr_inv 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 blast have "hseq g w" using assms left_adjoint_is_ide by simp have "src m = src n" using mn m\<^sub>a `hseq n m\<^sup>*` adjoint_pair_antipar [of m "m\<^sup>*"] 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\<^sup>* \ g \ w" proof - have 1: "src g = trg (n \ m\<^sup>*)" using assms `trg (n \ m\<^sup>*) = trg w` `hseq g w` by fastforce hence "(g \ n) \ m\<^sup>* \ g \ n \ m\<^sup>*" using assms mn m\<^sub>a assoc_in_hom iso_assoc `hseq g n` `hseq n m\<^sup>*` 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 m\<^sub>a 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 m\<^sub>a 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\<^sup>*\ is also an equivalence and therefore a map. But \w = n \ m\<^sub>a\, 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 \] apply (elim conjE in_homE) by simp have "hseq (src w) e'" using * ide_dom [of \'] apply (elim conjE in_homE) by simp have "e'e.trnr\<^sub>\ 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\<^sub>\ m \ : m \ src w \ e'\" by simp hence "\\' \ e'e.trnr\<^sub>\ m \ : m \ m\ \ \e'e.trnr\<^sub>\ 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\<^sub>\ 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\<^sub>\ m \ = m \ e'e.trnr\<^sub>\ m \ \ \' = src w \ e'" using mn BS3 [of m m "\' \ e'e.trnr\<^sub>\ m \" m] BS3 [of "src w \ e'" "src w \ e'" "e'e.trnr\<^sub>\ m \ \ \'" "src w \ e'"] by auto hence "inverse_arrows \' (e'e.trnr\<^sub>\ 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\<^sup>*" using mn m\<^sub>a right_adjoint_to_equivalence_is_equivalence by simp hence "is_left_adjoint m\<^sup>*" using equivalence_is_left_adjoint by simp moreover have "hseq n m\<^sup>*" using mn isomorphic_implies_ide by auto ultimately have "is_left_adjoint (n \ m\<^sup>*)" 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 \ tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>*\ 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\<^sup>* \ r" definition tab\<^sub>0 where "tab\<^sub>0 r \ SOME f. \g. isomorphic_rep (iso_class_rep \r\) f g" definition tab\<^sub>1 where "tab\<^sub>1 r \ SOME g. isomorphic_rep (iso_class_rep \r\) (tab\<^sub>0 r) g" definition rep where "rep r \ SOME \. \\ : tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>* \ r\ \ iso \" lemma rep_props: assumes "ide r" shows "\rep r : tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>* \ r\" and "iso (rep r)" and "r \ iso_class_rep \r\" and "isomorphic_rep r (tab\<^sub>0 r) (tab\<^sub>1 r)" and "tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>* \ r" proof - have 1: "isomorphic_rep r (tab\<^sub>0 r) (tab\<^sub>1 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\) (tab\<^sub>0 r) (tab\<^sub>1 r)" using assms tab\<^sub>0_def tab\<^sub>1_def someI_ex [of "\f. \g. isomorphic_rep (iso_class_rep \r\) f g"] someI_ex [of "\g. isomorphic_rep (iso_class_rep \r\) (tab\<^sub>0 r) g"] by simp thus ?thesis using assms isomorphic_symmetric isomorphic_transitive rep_iso_class by blast qed hence "\\. \\ : tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>* \ r\ \ iso \" using isomorphic_def by blast hence 2: "\rep r : tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>* \ r\ \ iso (rep r)" using someI_ex [of "\\. \\ : tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>* \ r\ \ iso \"] rep_def by auto show "\rep r : tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>* \ 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 (tab\<^sub>0 r) (tab\<^sub>1 r)" using 1 isomorphic_transitive by blast thus "tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>* \ r" by simp qed lemma tab\<^sub>0_in_hom [intro]: assumes "ide r" shows "\tab\<^sub>0 r : src (tab\<^sub>0 r) \ src r\" and "\tab\<^sub>0 r : tab\<^sub>0 r \ tab\<^sub>0 r\" proof - show "\tab\<^sub>0 r : tab\<^sub>0 r \ tab\<^sub>0 r\" using assms rep_props left_adjoint_is_ide by auto have "trg (tab\<^sub>0 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 "\tab\<^sub>0 r : src (tab\<^sub>0 r) \ src r\" using assms rep_props left_adjoint_is_ide by (intro in_hhomI, auto) qed lemma tab\<^sub>0_simps [simp]: assumes "ide r" shows "ide (tab\<^sub>0 r)" and "is_left_adjoint (tab\<^sub>0 r)" and "trg (tab\<^sub>0 r) = src r" and "dom (tab\<^sub>0 r) = tab\<^sub>0 r" and "cod (tab\<^sub>0 r) = tab\<^sub>0 r" using assms tab\<^sub>0_in_hom rep_props ide_dom left_adjoint_is_ide by auto lemma tab\<^sub>1_in_hom [intro]: assumes "ide r" shows "\tab\<^sub>1 r : src (tab\<^sub>0 r) \ trg r\" and "\tab\<^sub>1 r : tab\<^sub>1 r \ tab\<^sub>1 r\" proof - show "\tab\<^sub>1 r : tab\<^sub>1 r \ tab\<^sub>1 r\" using assms rep_props left_adjoint_is_ide by auto have "trg (tab\<^sub>1 r) = trg r" using assms rep_props by (metis ideD(1) isomorphic_implies_hpar(1) isomorphic_implies_hpar(4) trg_hcomp) moreover have "src (tab\<^sub>0 r) = src (tab\<^sub>1 r)" using assms rep_props by fastforce ultimately show "\tab\<^sub>1 r : src (tab\<^sub>0 r) \ trg r\" using assms rep_props left_adjoint_is_ide by (intro in_hhomI, auto) qed lemma tab\<^sub>1_simps [simp]: assumes "ide r" shows "ide (tab\<^sub>1 r)" and "is_left_adjoint (tab\<^sub>1 r)" and "src (tab\<^sub>1 r) = src (tab\<^sub>0 r)" and "trg (tab\<^sub>1 r) = trg r" and "dom (tab\<^sub>1 r) = tab\<^sub>1 r" and "cod (tab\<^sub>1 r) = tab\<^sub>1 r" using assms tab\<^sub>1_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 : tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>* \ r\" proof - show "\rep r : tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>* \ r\" using assms rep_props by auto thus "\rep r : src r \ trg r\" using src_cod trg_cod by fastforce 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) = tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>*" 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 \tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>*\ 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 : tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>* \ r\\. \ locale identity_in_bicategory_of_spans = bicategory_of_spans + fixes r :: 'a assumes is_ide: "ide r" begin interpretation tab\<^sub>0: map_in_bicategory V H \ \ src trg \tab\<^sub>0 r\ using is_ide rep_props by unfold_locales auto interpretation tab\<^sub>1: map_in_bicategory V H \ \ src trg \tab\<^sub>1 r\ using is_ide rep_props by unfold_locales auto text \ A tabulation \(tab\<^sub>0 r, tab, tab\<^sub>1 r)\ of \r\ can be obtained as the adjoint transpose of the isomorphism \\rep r : (tab\<^sub>1 r) \ (tab\<^sub>0 r)\<^sup>* \ 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 \ tab\<^sub>0.trnr\<^sub>\ (tab\<^sub>1 r) (rep r)" text \ In view of \BS2'\, the 1-cell \(tab\<^sub>1 r) \ (tab\<^sub>0 r)\<^sup>*\ 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 : (tab\<^sub>1 r) \ (tab\<^sub>0 r)\<^sup>* \ 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 ((tab\<^sub>1 r) \ (tab\<^sub>0 r)\<^sup>*) (tab\<^sub>0.trnr\<^sub>\ (tab\<^sub>1 r) ((tab\<^sub>1 r) \ (tab\<^sub>0 r)\<^sup>*)) (tab\<^sub>0 r) (tab\<^sub>1 r)" proof - have "\\. tabulation V H \ \ src trg ((tab\<^sub>1 r) \ (tab\<^sub>0 r)\<^sup>*) \ (tab\<^sub>0 r) (tab\<^sub>1 r)" by (simp add: bicategory_of_spans.BS2' bicategory_of_spans_axioms is_ide isomorphic_reflexive) thus ?thesis using is_ide tab\<^sub>0.canonical_tabulation by simp qed lemma tab_def_alt: shows "tab = (rep r \ tab\<^sub>0 r) \ tab\<^sub>0.trnr\<^sub>\ (tab\<^sub>1 r) ((tab\<^sub>1 r) \ (tab\<^sub>0 r)\<^sup>*)" and "(inv (rep r) \ tab\<^sub>0 r) \ tab = tab\<^sub>0.trnr\<^sub>\ (tab\<^sub>1 r) ((tab\<^sub>1 r) \ (tab\<^sub>0 r)\<^sup>*)" proof - have "tab = tab\<^sub>0.trnr\<^sub>\ (tab\<^sub>1 r) (rep r \ ((tab\<^sub>1 r) \ (tab\<^sub>0 r)\<^sup>*))" using tab_def is_ide rep_in_hom [of r] comp_arr_dom by auto also have "... = (rep r \ tab\<^sub>0 r) \ tab\<^sub>0.trnr\<^sub>\ (tab\<^sub>1 r) ((tab\<^sub>1 r) \ (tab\<^sub>0 r)\<^sup>*)" using is_ide tab\<^sub>0.trnr\<^sub>\_comp by auto finally show 1: "tab = (rep r \ tab\<^sub>0 r) \ tab\<^sub>0.trnr\<^sub>\ (tab\<^sub>1 r) ((tab\<^sub>1 r) \ (tab\<^sub>0 r)\<^sup>*)" by simp have "(inv (rep r) \ tab\<^sub>0 r) \ tab = ((inv (rep r) \ tab\<^sub>0 r) \ (rep r \ tab\<^sub>0 r)) \ tab\<^sub>0.trnr\<^sub>\ (tab\<^sub>1 r) ((tab\<^sub>1 r) \ (tab\<^sub>0 r)\<^sup>*)" unfolding 1 using comp_assoc by presburger also have "... = tab\<^sub>0.trnr\<^sub>\ (tab\<^sub>1 r) ((tab\<^sub>1 r) \ (tab\<^sub>0 r)\<^sup>*)" proof - have 1: "(inv (rep r) \ tab\<^sub>0 r) \ (rep r \ tab\<^sub>0 r) = ((tab\<^sub>1 r) \ (tab\<^sub>0 r)\<^sup>*) \ tab\<^sub>0 r" using whisker_right [of "tab\<^sub>0 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 "\tab\<^sub>0.trnr\<^sub>\ (tab\<^sub>1 r) ((tab\<^sub>1 r) \ (tab\<^sub>0 r)\<^sup>*) : tab\<^sub>1 r \ (tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>*) \ tab\<^sub>0 r\" by (meson canonical_tabulation tabulation_data.tab_in_hom(2) tabulation_def) hence "((tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>*) \ tab\<^sub>0 r) \ tab\<^sub>0.trnr\<^sub>\ (tab\<^sub>1 r) ((tab\<^sub>1 r) \ (tab\<^sub>0 r)\<^sup>*) = tab\<^sub>0.trnr\<^sub>\ (tab\<^sub>1 r) ((tab\<^sub>1 r) \ (tab\<^sub>0 r)\<^sup>*)" using 1 comp_cod_arr by blast thus ?thesis using 1 by simp qed qed finally show "(inv (rep r) \ tab\<^sub>0 r) \ tab = tab\<^sub>0.trnr\<^sub>\ (tab\<^sub>1 r) ((tab\<^sub>1 r) \ (tab\<^sub>0 r)\<^sup>*)" by blast qed lemma tab_is_tabulation: shows "tabulation V H \ \ src trg r tab (tab\<^sub>0 r) (tab\<^sub>1 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 (tab\<^sub>0 r) \ trg r\" and "\tab : tab\<^sub>1 r \ r \ tab\<^sub>0 r\" proof - interpret tab: tabulation V H \ \ src trg r tab \tab\<^sub>0 r\ \tab\<^sub>1 r\ using tab_is_tabulation by simp show "\tab : src (tab\<^sub>0 r) \ trg r\" using tab.tab_in_hom by auto show "\tab : tab\<^sub>1 r \ r \ tab\<^sub>0 r\" using tab.tab_in_hom by auto qed lemma tab_simps [simp]: shows "arr tab" and "src tab = src (tab\<^sub>0 r)" and "trg tab = trg r" and "dom tab = tab\<^sub>1 r" and "cod tab = r \ tab\<^sub>0 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 \tab\<^sub>0 r\ \tab\<^sub>1 r\ using is_ide tab_is_tabulation by simp context identity_in_bicategory_of_spans begin interpretation tab\<^sub>0: map_in_bicategory V H \ \ src trg \tab\<^sub>0 r\ using is_ide rep_props by unfold_locales auto interpretation tab\<^sub>1: map_in_bicategory V H \ \ src trg \tab\<^sub>1 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\<^sub>\ r tab" using is_ide T0.adjoint_transpose_right(3) [of r "tab\<^sub>1 r" "rep r"] tab_def by fastforce lemma isomorphic_implies_same_tab: assumes "isomorphic r r'" shows "tab\<^sub>0 r = tab\<^sub>0 r'" and "tab\<^sub>1 r = tab\<^sub>1 r'" using assms tab\<^sub>0_def tab\<^sub>1_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 prj\<^sub>0 where "prj\<^sub>0 h k \ tab\<^sub>0 (k\<^sup>* \ h)" abbreviation prj\<^sub>1 where "prj\<^sub>1 h k \ tab\<^sub>1 (k\<^sup>* \ h)" lemma prj_in_hom [intro]: assumes "ide h" and "is_left_adjoint k" and "trg h = trg k" shows "\prj\<^sub>0 h k : src (prj\<^sub>0 h k) \ src h\" and "\prj\<^sub>1 h k : src (prj\<^sub>0 h k) \ src k\" and "\prj\<^sub>0 h k : prj\<^sub>0 h k \ prj\<^sub>0 h k\" and "\prj\<^sub>1 h k : prj\<^sub>1 h k \ prj\<^sub>1 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 (prj\<^sub>0 h k) = src h" and "src (prj\<^sub>1 h k) = src (prj\<^sub>0 h k)" and "trg (prj\<^sub>1 h k) = src k" and "dom (prj\<^sub>0 h k) = prj\<^sub>0 h k" and "cod (prj\<^sub>0 h k) = prj\<^sub>0 h k" and "dom (prj\<^sub>1 h k) = prj\<^sub>1 h k" and "cod (prj\<^sub>1 h k) = prj\<^sub>1 h k" and "is_left_adjoint (prj\<^sub>0 h k)" and "is_left_adjoint (prj\<^sub>1 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 leg\<^sub>0 :: 'a and leg\<^sub>1 :: 'a assumes leg0_is_map: "is_left_adjoint leg\<^sub>0" and leg1_is_map : "is_left_adjoint leg\<^sub>1" 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 r\<^sub>0 r\<^sub>1 + s: span_of_maps V H \ \ src trg s\<^sub>0 s\<^sub>1 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\<^sub>0 :: 'a and r\<^sub>1 :: 'a and s\<^sub>0 :: 'a and s\<^sub>1 :: 'a and w :: 'a + assumes is_ide: "ide w" and leg0_lax: "\\. \\ : s\<^sub>0 \ w \ r\<^sub>0\" and leg1_iso: "\\. \\ : r\<^sub>1 \ s\<^sub>1 \ w\ \ iso \" begin notation isomorphic (infix "\" 50) lemma composite_leg1_is_map: shows "is_left_adjoint (s\<^sub>1 \ 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 s\<^sub>1 w] by auto lemma hseq_leg\<^sub>0: shows "hseq s\<^sub>0 w" by (metis ideD(1) ide_dom in_homE leg0_lax) lemma composite_with_leg0_is_map: shows "is_left_adjoint (s\<^sub>0 \ w)" using left_adjoints_compose is_map s.leg0_is_map hseq_leg\<^sub>0 by blast lemma leg0_uniquely_isomorphic: shows "s\<^sub>0 \ w \ r\<^sub>0" and "\!\. \\ : s\<^sub>0 \ w \ r\<^sub>0\" proof - show 1: "s\<^sub>0 \ w \ r\<^sub>0" using leg0_lax composite_with_leg0_is_map r.leg0_is_map BS3 [of "s\<^sub>0 \ w" r\<^sub>0] isomorphic_def by auto have "\\. \\ : s\<^sub>0 \ w \ r\<^sub>0\ \ iso \" using 1 isomorphic_def by simp moreover have "\\ \'. \\ : s\<^sub>0 \ w \ r\<^sub>0\ \ \\' : s\<^sub>0 \ w \ r\<^sub>0\ \ \ = \'" using BS3 r.leg0_is_map composite_with_leg0_is_map by blast ultimately show "\!\. \\ : s\<^sub>0 \ w \ r\<^sub>0\" by blast qed lemma leg1_uniquely_isomorphic: shows "r\<^sub>1 \ s\<^sub>1 \ w" and "\!\. \\ : r\<^sub>1 \ s\<^sub>1 \ w\" proof - show 1: "r\<^sub>1 \ s\<^sub>1 \ w" using leg1_iso isomorphic_def by auto have "\\. \\ : r\<^sub>1 \ s\<^sub>1 \ w\ \ iso \" using leg1_iso isomorphic_def isomorphic_symmetric by simp moreover have "\\ \'. \\ : r\<^sub>1 \ s\<^sub>1 \ w\ \ \\' : r\<^sub>1 \ s\<^sub>1 \ w\ \ \ = \'" using BS3 r.leg1_is_map composite_leg1_is_map by blast ultimately show "\!\. \\ : r\<^sub>1 \ s\<^sub>1 \ w\" by blast qed definition the_\ where "the_\ \ THE \. \\ : s\<^sub>0 \ w \ r\<^sub>0\" definition the_\ where "the_\ \ THE \. \\ : r\<^sub>1 \ s\<^sub>1 \ w\" lemma the_\_props: shows "\the_\ : s\<^sub>0 \ w \ r\<^sub>0\" and "iso the_\" proof - show "\the_\ : s\<^sub>0 \ w \ r\<^sub>0\" unfolding the_\_def using the1I2 [of "\\. \\ : s\<^sub>0 \ w \ r\<^sub>0\" "\\. \\ : s\<^sub>0 \ w \ r\<^sub>0\"] 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 r\<^sub>0 \ trg r\<^sub>0\" and "\the_\ : s\<^sub>0 \ w \ r\<^sub>0\" 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 r\<^sub>0" and "trg the_\ = trg r\<^sub>0" and "dom the_\ = s\<^sub>0 \ w" and "cod the_\ = r\<^sub>0" using the_\_in_hom by auto lemma the_\_props: shows "\the_\ : r\<^sub>1 \ s\<^sub>1 \ w\" and "iso the_\" proof - show "\the_\ : r\<^sub>1 \ s\<^sub>1 \ w\" unfolding the_\_def using the1I2 [of "\\. \\ : r\<^sub>1 \ s\<^sub>1 \ w\" "\\. \\ : r\<^sub>1 \ s\<^sub>1 \ 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 r\<^sub>1 \ trg r\<^sub>1\" and "\the_\ : r\<^sub>1 \ s\<^sub>1 \ 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 r\<^sub>1" and "trg the_\ = trg r\<^sub>1" and "dom the_\ = r\<^sub>1" and "cod the_\ = s\<^sub>1 \ 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 r\<^sub>0 r\<^sub>1 s\<^sub>0 s\<^sub>1 w + \: tabulation_data V H \ \ src trg s \ s\<^sub>0 s\<^sub>1 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\<^sub>0 :: 'a and r\<^sub>1 :: 'a and s :: 'a and \ :: 'a and s\<^sub>0 :: 'a and s\<^sub>1 :: '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 \ s\<^sub>0 s\<^sub>1 r\<^sub>0 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 \ s\<^sub>0 s\<^sub>1 locale tabulation_in_maps = span_of_maps V H \ \ src trg s\<^sub>0 s\<^sub>1 + tabulation V H \ \ src trg s \ s\<^sub>0 s\<^sub>1 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 s\<^sub>0 :: 'a and s\<^sub>1 :: 'a sublocale tabulation_in_maps \ tabulation V H \ \ src trg s \ s\<^sub>0 s\<^sub>1 .. sublocale identity_in_bicategory_of_spans \ tabulation_in_maps V H \ \ src trg r tab \tab\<^sub>0 r\ \tab\<^sub>1 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\<^sup>* \ h\ and its chosen tabulation. \ sublocale identity_in_bicategory_of_spans V H \ \ src trg \k\<^sup>* \ 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\<^sup>* \ h\ in terms of its transpose, which yields a 2-cell from \k \ tab\<^sub>1 (k\<^sup>* \ h)\ to \h \ tab\<^sub>0 (k\<^sup>* \ h)\. These results do not depend on \BS3\. \ abbreviation p\<^sub>0 where "p\<^sub>0 \ prj\<^sub>0 h k" abbreviation p\<^sub>1 where "p\<^sub>1 \ prj\<^sub>1 h k" lemma p\<^sub>0_in_hom [intro]: shows "\p\<^sub>0 : src p\<^sub>0 \ src h\" by auto lemma p\<^sub>1_in_hom [intro]: shows "\p\<^sub>1 : src p\<^sub>0 \ src k\" using prj_in_hom cospan h.ide_left k_is_map by blast lemma p\<^sub>0_simps [simp]: shows "trg p\<^sub>0 = src h" by simp lemma p\<^sub>1_simps [simp]: shows "trg p\<^sub>1 = src k" using k.antipar(1) by auto definition \ where "\ \ k.trnl\<^sub>\ (h \ p\<^sub>0) (\[k\<^sup>*, h, p\<^sub>0] \ tab)" lemma \_in_hom [intro]: shows "\\ : src p\<^sub>0 \ trg h\" and "\\ : k \ p\<^sub>1 \ h \ p\<^sub>0\" proof - show 1: "\\ : k \ p\<^sub>1 \ h \ p\<^sub>0\" unfolding \_def using k.antipar cospan k.adjoint_transpose_left(2) [of "h \ p\<^sub>0" "p\<^sub>1"] by fastforce show "\\ : src p\<^sub>0 \ trg h\" using 1 k.antipar src_dom trg_cod by fastforce qed lemma \_simps [simp]: shows "arr \" and "src \ = src p\<^sub>0" and "trg \ = trg h" and "dom \ = k \ p\<^sub>1" and "cod \ = h \ p\<^sub>0" using \_in_hom by auto lemma transpose_\: shows "tab = \\<^sup>-\<^sup>1[k\<^sup>*, h, p\<^sub>0] \ k.trnl\<^sub>\ p\<^sub>1 \" proof - have "\\<^sup>-\<^sup>1[k\<^sup>*, h, p\<^sub>0] \ k.trnl\<^sub>\ p\<^sub>1 \ = \\<^sup>-\<^sup>1[k\<^sup>*, h, p\<^sub>0] \ \[k\<^sup>*, h, p\<^sub>0] \ tab" unfolding \_def using k.antipar cospan k.adjoint_transpose_left(4) [of "h \ p\<^sub>0" "p\<^sub>1" "\[k\<^sup>*, h, p\<^sub>0] \ tab"] by fastforce also have "... = (\\<^sup>-\<^sup>1[k\<^sup>*, h, p\<^sub>0] \ \[k\<^sup>*, h, p\<^sub>0]) \ 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 "\\ : p\<^sub>0 \ w \ u\" and "\\ : v \ p\<^sub>1 \ w\" shows "k.trnl\<^sub>\ (h \ u) (\[k\<^sup>*, h, u] \ ((k\<^sup>* \ h) \ \) \ \[k\<^sup>* \ h, p\<^sub>0, w] \ (tab \ w) \ \) = (h \ \) \ \[h, p\<^sub>0, w] \ (\ \ w) \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w] \ (k \ \)" proof - have u: "ide u" using assms(2) by auto have v: "ide v" using assms(3) by auto have 0: "src p\<^sub>0 = 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) p\<^sub>1_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\<^sup>*, h, p\<^sub>0 \ w] \ \[k\<^sup>* \ h, p\<^sub>0, w] \ (tab \ w)" have \: "\\ : p\<^sub>1 \ w \ k\<^sup>* \ h \ p\<^sub>0 \ w\" unfolding \_def using assms 0 k.antipar cospan by (intro comp_in_homI, auto) have "k.trnl\<^sub>\ (h \ u) (\[k\<^sup>*, h, u] \ ((k\<^sup>* \ h) \ \) \ \[k\<^sup>* \ h, p\<^sub>0, w] \ (tab \ w) \ \) = k.trnl\<^sub>\ (h \ u) ((k\<^sup>* \ h \ \) \ \ \ \)" unfolding \_def using assms 1 k.antipar cospan assoc_naturality [of "k\<^sup>*" h \] comp_assoc by (metis "4" h.ide_left ide_char in_homE k.ide_right) also have "... = k.trnl\<^sub>\ (h \ u) (k\<^sup>* \ h \ \) \ (k \ \ \ \)" proof - have "ide (h \ u)" using "1" u assms h.ide_left by blast moreover have "seq (k\<^sup>* \ h \ \) (\ \ \)" using assms 1 k.antipar cospan \ seqI' apply (intro seqI) apply auto apply blast proof - have "dom (k\<^sup>* \ h \ \) = k\<^sup>* \ h \ p\<^sub>0 \ 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\<^sup>* \ h \ \) = cod \" by simp qed moreover have "src k = trg (k\<^sup>* \ h \ \)" using assms k.antipar cospan calculation(2) by auto ultimately show ?thesis using k.trnl\<^sub>\_comp by simp qed also have "... = k.trnl\<^sub>\ (h \ u) (k\<^sup>* \ h \ \) \ (k \ \) \ (k \ \)" using assms u \ whisker_left by (metis k.ide_left seqI') also have "... = (\[h \ u] \ (k.\ \ h \ u) \ \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ u] \ (k \ k\<^sup>* \ h \ \)) \ (k \ \) \ (k \ \)" unfolding k.trnl\<^sub>\_def by simp also have "... = (h \ \) \ (\[h \ p\<^sub>0 \ w] \ (k.\ \ h \ p\<^sub>0 \ w) \ \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ p\<^sub>0 \ w] \ (k \ \)) \ (k \ \)" proof - have "\[h \ u] \ (k.\ \ h \ u) \ \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ u] \ (k \ k\<^sup>* \ h \ \) = \[h \ u] \ (k.\ \ h \ u) \ ((k \ k\<^sup>*) \ h \ \) \ \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ p\<^sub>0 \ w]" using assms 4 k.antipar cospan assoc'_naturality [of k "k\<^sup>*" "h \ \"] by fastforce also have "... = \[h \ u] \ ((k.\ \ h \ u) \ ((k \ k\<^sup>*) \ h \ \)) \ \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ p\<^sub>0 \ w]" using comp_assoc by presburger also have "... = (\[h \ u] \ (trg k \ h \ \)) \ (k.\ \ h \ p\<^sub>0 \ w) \ \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ p\<^sub>0 \ w]" proof - have "(k.\ \ h \ u) \ ((k \ k\<^sup>*) \ h \ \) = k.\ \ (k \ k\<^sup>*) \ (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 \ p\<^sub>0 \ w)" using assms 4 k.antipar cospan whisker_left comp_arr_dom comp_cod_arr interchange [of "trg k" k.\ "h \ \" "h \ p\<^sub>0 \ w"] by auto finally have "(k.\ \ h \ u) \ ((k \ k\<^sup>*) \ h \ \) = (trg k \ h \ \) \ (k.\ \ h \ p\<^sub>0 \ w)" by simp thus ?thesis using comp_assoc by presburger qed also have "... = (h \ \) \ \[h \ p\<^sub>0 \ w] \ (k.\ \ h \ p\<^sub>0 \ w) \ \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ p\<^sub>0 \ w]" proof - have "\[h \ u] \ (trg k \ h \ \) = (h \ \) \ \[h \ p\<^sub>0 \ 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) \ \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ u] \ (k \ k\<^sup>* \ h \ \) = (h \ \) \ \[h \ p\<^sub>0 \ w] \ (k.\ \ h \ p\<^sub>0 \ w) \ \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ p\<^sub>0 \ w]" by simp thus ?thesis using comp_assoc by presburger qed also have "... = (h \ \) \ \[h, p\<^sub>0, w] \ (\ \ w) \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w] \ (k \ \)" proof - have "\[h \ p\<^sub>0 \ w] \ (k.\ \ h \ p\<^sub>0 \ w) \ \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ p\<^sub>0 \ w] \ (k \ \[k\<^sup>*, h, p\<^sub>0 \ w] \ \[k\<^sup>* \ h, p\<^sub>0, w] \ (tab \ w)) = \[h, p\<^sub>0, w] \ \[(h \ p\<^sub>0) \ w] \ (trg h \ \\<^sup>-\<^sup>1[h, p\<^sub>0, w]) \ (k.\ \ h \ p\<^sub>0 \ w) \ \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ p\<^sub>0 \ w] \ (k \ \[k\<^sup>*, h, p\<^sub>0 \ w] \ \[k\<^sup>* \ h, p\<^sub>0, w] \ (tab \ w))" proof - have "\[h \ p\<^sub>0 \ w] = \[h, p\<^sub>0, w] \ \[(h \ p\<^sub>0) \ w] \ (trg h \ \\<^sup>-\<^sup>1[h, p\<^sub>0, w])" proof - have "\[h, p\<^sub>0, w] \ \[(h \ p\<^sub>0) \ w] \ (trg h \ \\<^sup>-\<^sup>1[h, p\<^sub>0, w]) = \[h, p\<^sub>0, w] \ \ ((h \ p\<^sub>0) \ w) \ (trg h \ \\<^sup>-\<^sup>1[h, p\<^sub>0, w])" using assms 0 k.antipar cospan comp_cod_arr \_ide_simp by simp also have "... = \[h, p\<^sub>0, w] \ \ (\\<^sup>-\<^sup>1[h, p\<^sub>0, w])" using assms 0 k.antipar cospan \.is_natural_2 [of "\\<^sup>-\<^sup>1[h, p\<^sub>0, w]"] by simp also have "... = \[h, p\<^sub>0, w] \ \\<^sup>-\<^sup>1[h, p\<^sub>0, w] \ \ (h \ p\<^sub>0 \ w)" using assms 0 k.antipar cospan \.is_natural_1 [of "\\<^sup>-\<^sup>1[h, p\<^sub>0, w]"] comp_assoc by simp also have "... = (\[h, p\<^sub>0, w] \ \\<^sup>-\<^sup>1[h, p\<^sub>0, w]) \ \ (h \ p\<^sub>0 \ w)" using comp_assoc by presburger also have "... = \ (h \ p\<^sub>0 \ w)" using assms 0 k.antipar cospan comp_cod_arr comp_assoc_assoc' by simp also have "... = \[h \ p\<^sub>0 \ 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, p\<^sub>0, w] \ (\[h \ p\<^sub>0] \ w) \ \\<^sup>-\<^sup>1[trg h, h \ p\<^sub>0, w] \ ((trg h \ \\<^sup>-\<^sup>1[h, p\<^sub>0, w]) \ (k.\ \ h \ p\<^sub>0 \ w)) \ \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ p\<^sub>0 \ w] \ (k \ \[k\<^sup>*, h, p\<^sub>0 \ w] \ \[k\<^sup>* \ h, p\<^sub>0, w] \ (tab \ w))" using assms 0 k.antipar cospan lunit_hcomp comp_assoc by simp also have "... = \[h, p\<^sub>0, w] \ (\[h \ p\<^sub>0] \ w) \ (\\<^sup>-\<^sup>1[trg h, h \ p\<^sub>0, w] \ (k.\ \ (h \ p\<^sub>0) \ w)) \ ((k \ k\<^sup>*) \ \\<^sup>-\<^sup>1[h, p\<^sub>0, w]) \ \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ p\<^sub>0 \ w] \ (k \ \[k\<^sup>*, h, p\<^sub>0 \ w] \ \[k\<^sup>* \ h, p\<^sub>0, w] \ (tab \ w))" proof - have "(trg h \ \\<^sup>-\<^sup>1[h, p\<^sub>0, w]) \ (k.\ \ h \ p\<^sub>0 \ w) = (k.\ \ (h \ p\<^sub>0) \ w) \ ((k \ k\<^sup>*) \ \\<^sup>-\<^sup>1[h, p\<^sub>0, w])" using assms 0 k.antipar cospan comp_arr_dom comp_cod_arr interchange [of "trg h" k.\ "\\<^sup>-\<^sup>1[h, p\<^sub>0, w]" "h \ p\<^sub>0 \ w"] interchange [of k.\ "k \ k\<^sup>*" "(h \ p\<^sub>0) \ w" "\\<^sup>-\<^sup>1[h, p\<^sub>0, w]"] by simp thus ?thesis using comp_assoc by presburger qed also have "... = \[h, p\<^sub>0, w] \ (\[h \ p\<^sub>0] \ w) \ ((k.\ \ (h \ p\<^sub>0)) \ w) \ \\<^sup>-\<^sup>1[k \ k\<^sup>*, h \ p\<^sub>0, w] \ ((k \ k\<^sup>*) \ \\<^sup>-\<^sup>1[h, p\<^sub>0, w]) \ \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ p\<^sub>0 \ w] \ (k \ \[k\<^sup>*, h, p\<^sub>0 \ w] \ \[k\<^sup>* \ h, p\<^sub>0, w] \ (tab \ w))" using assms 0 k.antipar cospan assoc'_naturality [of k.\ "h \ p\<^sub>0" w] comp_assoc by simp also have "... = \[h, p\<^sub>0, w] \ (\[h \ p\<^sub>0] \ w) \ ((k.\ \ (h \ p\<^sub>0)) \ w) \ \\<^sup>-\<^sup>1[k \ k\<^sup>*, h \ p\<^sub>0, w] \ ((k \ k\<^sup>*) \ \\<^sup>-\<^sup>1[h, p\<^sub>0, w]) \ \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ p\<^sub>0 \ w] \ (k \ \[k\<^sup>*, h, p\<^sub>0 \ w]) \ (k \ \[k\<^sup>* \ h, p\<^sub>0, w]) \ (k \ tab \ w)" proof - have "k \ \[k\<^sup>*, h, p\<^sub>0 \ w] \ \[k\<^sup>* \ h, p\<^sub>0, w] \ (tab \ w) = (k \ \[k\<^sup>*, h, p\<^sub>0 \ w]) \ (k \ \[k\<^sup>* \ h, p\<^sub>0, w]) \ (k \ tab \ w)" proof - have "seq \[k\<^sup>*, h, p\<^sub>0 \ w] (\[k\<^sup>* \ h, p\<^sub>0, 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, p\<^sub>0, w] \ (\[h \ p\<^sub>0] \ w) \ ((k.\ \ (h \ p\<^sub>0)) \ w) \ (\\<^sup>-\<^sup>1[k \ k\<^sup>*, h \ p\<^sub>0, w] \ ((k \ k\<^sup>*) \ \\<^sup>-\<^sup>1[h, p\<^sub>0, w]) \ \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ p\<^sub>0 \ w] \ (k \ \[k\<^sup>*, h, p\<^sub>0 \ w]) \ (k \ \[k\<^sup>* \ h, p\<^sub>0, w]) \ \[k, (k\<^sup>* \ h) \ p\<^sub>0, w]) \ ((k \ tab) \ w) \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w]" proof - have "k \ tab \ w = \[k, (k\<^sup>* \ h) \ p\<^sub>0, w] \ \\<^sup>-\<^sup>1[k, (k\<^sup>* \ h) \ p\<^sub>0, w] \ (k \ tab \ w)" proof - have "\[k, (k\<^sup>* \ h) \ p\<^sub>0, w] \ \\<^sup>-\<^sup>1[k, (k\<^sup>* \ h) \ p\<^sub>0, w] \ (k \ tab \ w) = (\[k, (k\<^sup>* \ h) \ p\<^sub>0, w] \ \\<^sup>-\<^sup>1[k, (k\<^sup>* \ h) \ p\<^sub>0, w]) \ (k \ tab \ w)" using comp_assoc by presburger also have "... = (k \ ((k\<^sup>* \ h) \ p\<^sub>0) \ 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\<^sup>* \ h) \ p\<^sub>0, w] \ ((k \ tab) \ w) \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w]" using assms 0 k.antipar cospan assoc'_naturality [of k tab w] by simp finally have "k \ tab \ w = \[k, (k\<^sup>* \ h) \ p\<^sub>0, w] \ ((k \ tab) \ w) \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w]" by simp thus ?thesis using comp_assoc by presburger qed also have "... = \[h, p\<^sub>0, w] \ (\[h \ p\<^sub>0] \ w) \ ((k.\ \ h \ p\<^sub>0) \ w) \ (\\<^sup>-\<^sup>1[k, k\<^sup>*, h \ p\<^sub>0] \ (k \ \[k\<^sup>*, h, p\<^sub>0]) \ w) \ ((k \ tab) \ w) \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w]" proof - have "\\<^sup>-\<^sup>1[k \ k\<^sup>*, h \ p\<^sub>0, w] \ ((k \ k\<^sup>*) \ \\<^sup>-\<^sup>1[h, p\<^sub>0, w]) \ \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ p\<^sub>0 \ w] \ (k \ \[k\<^sup>*, h, p\<^sub>0 \ w]) \ (k \ \[k\<^sup>* \ h, p\<^sub>0, w]) \ \[k, (k\<^sup>* \ h) \ p\<^sub>0, w] = \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ p\<^sub>0] \ (k \ \[k\<^sup>*, h, p\<^sub>0]) \ w" proof - have "\\<^sup>-\<^sup>1[k \ k\<^sup>*, h \ p\<^sub>0, w] \ ((k \ k\<^sup>*) \ \\<^sup>-\<^sup>1[h, p\<^sub>0, w]) \ \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ p\<^sub>0 \ w] \ (k \ \[k\<^sup>*, h, p\<^sub>0 \ w]) \ (k \ \[k\<^sup>* \ h, p\<^sub>0, w]) \ \[k, (k\<^sup>* \ h) \ p\<^sub>0, w] = \\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\k\<^bold>\ \<^bold>\ \<^bold>\k\<^sup>*\<^bold>\, \<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\p\<^sub>0\<^bold>\, \<^bold>\w\<^bold>\\<^bold>] \<^bold>\ ((\<^bold>\k\<^bold>\ \<^bold>\ \<^bold>\k\<^sup>*\<^bold>\) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\h\<^bold>\, \<^bold>\p\<^sub>0\<^bold>\, \<^bold>\w\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\k\<^bold>\, \<^bold>\k\<^sup>*\<^bold>\, \<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\w\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\k\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\k\<^sup>*\<^bold>\, \<^bold>\h\<^bold>\, \<^bold>\p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\w\<^bold>\\<^bold>]) \<^bold>\ (\<^bold>\k\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\k\<^sup>*\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\, \<^bold>\p\<^sub>0\<^bold>\, \<^bold>\w\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\k\<^bold>\, (\<^bold>\k\<^sup>*\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) \<^bold>\ \<^bold>\p\<^sub>0\<^bold>\, \<^bold>\w\<^bold>\\<^bold>]\" using assms 0 k.antipar cospan \_def \'_def by simp also have "... = \\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\k\<^bold>\, \<^bold>\k\<^sup>*\<^bold>\, \<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\p\<^sub>0\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\k\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\k\<^sup>*\<^bold>\, \<^bold>\h\<^bold>\, \<^bold>\p\<^sub>0\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\w\<^bold>\\" using assms 0 k.antipar cospan by (intro E.eval_eqI, simp_all) also have "... = \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ p\<^sub>0] \ (k \ \[k\<^sup>*, h, p\<^sub>0]) \ 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, p\<^sub>0, w] \ (\[h \ p\<^sub>0] \ (k.\ \ h \ p\<^sub>0) \ \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ p\<^sub>0] \ (k \ \[k\<^sup>*, h, p\<^sub>0]) \ (k \ tab) \ w) \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w]" using assms 0 k.antipar cospan comp_assoc whisker_right by auto also have "... = \[h, p\<^sub>0, w] \ (\ \ w) \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w]" unfolding \_def k.trnl\<^sub>\_def using assms 0 k.antipar cospan comp_assoc whisker_left by simp finally have "\[h \ p\<^sub>0 \ w] \ (k.\ \ h \ p\<^sub>0 \ w) \ \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ p\<^sub>0 \ w] \ (k \ \[k\<^sup>*, h, p\<^sub>0 \ w] \ \[k\<^sup>* \ h, p\<^sub>0, w] \ (tab \ w)) = \[h, p\<^sub>0, w] \ (\ \ w) \ \\<^sup>-\<^sup>1[k, p\<^sub>1, 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 \ p\<^sub>1\ to \h \ p\<^sub>0\ and is an isomorphism. \ lemma \_uniqueness: shows "\\. \\ : k \ p\<^sub>1 \ h \ p\<^sub>0\ \ \ = \" and "iso \" proof - have 2: "is_left_adjoint (k \ p\<^sub>1)" using k.antipar cospan left_adjoints_compose by (simp add: k_is_map) have 3: "is_left_adjoint (h \ p\<^sub>0)" using k.antipar cospan left_adjoints_compose by (simp add: h_is_map) show "\\. \\ : k \ p\<^sub>1 \ h \ p\<^sub>0\ \ \ = \" 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\<^sup>* \ h\ is the unique 2-cell from \p\<^sub>1\ to \(k\<^sup>* \ h) \ p\<^sub>0\, and therefore if we are given any such 2-cell we may conclude it yields a tabulation of \k\<^sup>* \ h\. \ lemma tab_uniqueness: assumes "\\ : p\<^sub>1 \ (k\<^sup>* \ h) \ p\<^sub>0\" shows "\ = tab" proof - have "\k.trnl\<^sub>\ (h \ p\<^sub>0) (\[k\<^sup>*, h, p\<^sub>0] \ \) : k \ p\<^sub>1 \ h \ p\<^sub>0\" using assms k.antipar cospan k.adjoint_transpose_left(2) [of "h \ p\<^sub>0" "p\<^sub>1"] assoc_in_hom by force hence "tab = \\<^sup>-\<^sup>1[k\<^sup>*, h, p\<^sub>0] \ k.trnl\<^sub>\ p\<^sub>1 (k.trnl\<^sub>\ (h \ p\<^sub>0) (\[k\<^sup>*, h, p\<^sub>0] \ \))" using transpose_\ \_uniqueness(1) by auto also have "... = \\<^sup>-\<^sup>1[k\<^sup>*, h, p\<^sub>0] \ \[k\<^sup>*, h, p\<^sub>0] \ \" using assms k.antipar cospan k.adjoint_transpose_left(4) assoc_in_hom by auto also have "... = (\\<^sup>-\<^sup>1[k\<^sup>*, h, p\<^sub>0] \ \[k\<^sup>*, h, p\<^sub>0]) \ \" 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\<^sup>* \ 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 \ \\ : p\<^sub>0 \ w \ u\ \ \\ : v \ p\<^sub>1 \ w\ \ iso \ \ (h \ \) \ \[h, p\<^sub>0, w] \ (\ \ w) \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w] \ (k \ \) = \" and "\w w' \ \' \. \ ide w; ide w'; \\ : p\<^sub>0 \ w \ u\; \\' : p\<^sub>0 \ w' \ u\; \\ : p\<^sub>1 \ w \ p\<^sub>1 \ w'\; (h \ \) \ \[h, p\<^sub>0, w] \ (\ \ w) \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w] = (h \ \') \ \[h, p\<^sub>0, w'] \ (\ \ w') \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w'] \ (k \ \) \ \ \!\. \\ : w \ w'\ \ \ = \' \ (p\<^sub>0 \ \) \ p\<^sub>1 \ \ = \" 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 ?\ = "\\<^sup>-\<^sup>1[k\<^sup>*, h, u] \ k.trnl\<^sub>\ v \" have \: "\?\ : v \ (k\<^sup>* \ 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 \ \\ : p\<^sub>0 \ w \ u\ \ \\ : v \ p\<^sub>1 \ w\ \ iso \ \ ((k\<^sup>* \ h) \ \) \ \[k\<^sup>* \ h, p\<^sub>0, w] \ (tab \ w) \ \ = ?\" using assms \ T1 [of u ?\] comp_assoc by (metis in_homE) have 0: "src p\<^sub>0 = trg w" using w\\ ide_dom by (metis hseqE ideD(1) in_homE) have "\ = k.trnl\<^sub>\ (h \ u) (\[k\<^sup>*, h, u] \ ((k\<^sup>* \ h) \ \) \ \[k\<^sup>* \ h, p\<^sub>0, w] \ (tab \ w) \ \)" proof - have "\ = k.trnl\<^sub>\ (h \ u) (\[k\<^sup>*, h, u] \ ?\)" proof - have "k.trnl\<^sub>\ (h \ u) (\[k\<^sup>*, h, u] \ ?\) = k.trnl\<^sub>\ (h \ u) ((\[k\<^sup>*, h, u] \ \\<^sup>-\<^sup>1[k\<^sup>*, h, u]) \ k.trnl\<^sub>\ v \)" using comp_assoc by presburger also have "... = k.trnl\<^sub>\ (h \ u) (k.trnl\<^sub>\ v \)" proof - have "(\[k\<^sup>*, h, u] \ \\<^sup>-\<^sup>1[k\<^sup>*, h, u]) \ k.trnl\<^sub>\ v \ = (k\<^sup>* \ h \ u) \ k.trnl\<^sub>\ v \" using comp_assoc_assoc' by (simp add: "1" assms(1) cospan k.antipar(2)) also have "... = k.trnl\<^sub>\ 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, p\<^sub>0, w] \ (\ \ w) \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w] \ (k \ \)" using assms k.antipar cospan w\\ transpose_triangle [of w \ u \] by auto finally have "(h \ \) \ \[h, p\<^sub>0, w] \ (\ \ w) \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w] \ (k \ \) = \" by simp thus "\w \ \. ide w \ \\ : p\<^sub>0 \ w \ u\ \ \\ : v \ p\<^sub>1 \ w\ \ iso \ \ (h \ \) \ \[h, p\<^sub>0, w] \ (\ \ w) \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w] \ (k \ \) = \" using w\\ by blast next fix w w' \ \' \ assume w: "ide w" assume w': "ide w'" assume \: "\\ : p\<^sub>0 \ w \ u\" assume \': "\\' : p\<^sub>0 \ w' \ u\" assume \: "\\ : p\<^sub>1 \ w \ p\<^sub>1 \ w'\" assume eq: "(h \ \) \ \[h, p\<^sub>0, w] \ (\ \ w) \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w] = (h \ \') \ \[h, p\<^sub>0, w'] \ (\ \ w') \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w'] \ (k \ \)" have 0: "src p\<^sub>0 = 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\<^sup>* \ h\ tab \p\<^sub>0\ \p\<^sub>1\ u w \ w' \' using w \ w' \' apply (unfold_locales) by auto show "\!\. \\ : w \ w'\ \ \ = \' \ (p\<^sub>0 \ \) \ p\<^sub>1 \ \ = \" proof - let ?LHS = "\[k\<^sup>*, h, u] \ ((k\<^sup>* \ h) \ \) \ \[k\<^sup>* \ h, p\<^sub>0, w] \ (tab \ w)" let ?RHS = "\[k\<^sup>*, h, u] \ ((k\<^sup>* \ h) \ \') \ \[k\<^sup>* \ h, p\<^sub>0, w'] \ (tab \ w') \ \" have eq': "?LHS = ?RHS" proof - have "k.trnl\<^sub>\ (h \ u) ?LHS = k.trnl\<^sub>\ (h \ u) (\[k\<^sup>*, h, u] \ ((k\<^sup>* \ h) \ \) \ \[k\<^sup>* \ h, p\<^sub>0, w] \ (tab \ w) \ (p\<^sub>1 \ 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, p\<^sub>0, w] \ (\ \ w) \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w] \ (k \ p\<^sub>1 \ 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, p\<^sub>0, w] \ (\ \ w) \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w]" using assms 0 w k.antipar cospan comp_arr_dom by simp also have "... = (h \ \') \ \[h, p\<^sub>0, w'] \ (\ \ w') \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w'] \ (k \ \)" using eq by blast also have "... = k.trnl\<^sub>\ (h \ u) ?RHS" using assms w' \' \ transpose_triangle by simp finally have 4: "k.trnl\<^sub>\ (h \ u) ?LHS = k.trnl\<^sub>\ (h \ u) ?RHS" by simp have "src k = trg (p\<^sub>1 \ w)" using assms 0 w k.antipar cospan by simp moreover have "src k\<^sup>* = 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 (p\<^sub>1 \ w)" using assms 0 w k.antipar cospan by simp ultimately have "inj_on (k.trnl\<^sub>\ (h \ u)) (hom (p\<^sub>1 \ w) (k\<^sup>* \ 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 (p\<^sub>1 \ w) (k\<^sup>* \ h \ u)" proof - have "\\[k\<^sup>*, h, u] \ ((k\<^sup>* \ h) \ \) \ \[k\<^sup>* \ h, p\<^sub>0, w] \ (tab \ w) : p\<^sub>1 \ w \ k\<^sup>* \ h \ u\" using k.antipar cospan apply (intro comp_in_homI) by auto thus ?thesis by simp qed moreover have "?RHS \ hom (p\<^sub>1 \ w) (k\<^sup>* \ h \ u)" proof - have "\\[k\<^sup>*, h, u] \ ((k\<^sup>* \ h) \ \') \ \[k\<^sup>* \ h, p\<^sub>0, w'] \ (tab \ w') \ \ : p\<^sub>1 \ w \ k\<^sup>* \ 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\<^sub>\ (h \ u)" "hom (p\<^sub>1 \ w) (k\<^sup>* \ h \ u)"] by simp qed moreover have "seq \[k\<^sup>*, h, u] (composite_cell w \)" using assms k.antipar cospan tab_in_hom by fastforce moreover have "seq \[k\<^sup>*, 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\<^sup>*, 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\<^sup>* \ 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 (p\<^sub>0 \ w) u \ isomorphic v (p\<^sub>1 \ w)" and "\w w'. \ is_left_adjoint w; is_left_adjoint w'; p\<^sub>0 \ w \ u; v \ p\<^sub>1 \ w; p\<^sub>0 \ w' \ u; v \ p\<^sub>1 \ 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 \ \\ : p\<^sub>0 \ w \ u\ \ \\ : v \ p\<^sub>1 \ w\ \ iso \ \ (h \ \) \ \[h, p\<^sub>0, w] \ (\ \ w) \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w] \ (k \ \) = \" using assms \ \_biuniversal_prop(1) [of u v \] by auto have "is_left_adjoint w \ isomorphic (p\<^sub>0 \ w) u \ isomorphic v (p\<^sub>1 \ 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 \ p\<^sub>1 \ w" using w\\ isomorphic_def by blast show "p\<^sub>0 \ w \ u" proof - have "src p\<^sub>0 = trg w" using w\\ ide_dom by (metis ideD(1) in_homE not_arr_null seq_if_composable) hence "is_left_adjoint (p\<^sub>0 \ 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 \ p\<^sub>0 \ w \ u \ v \ p\<^sub>1 \ w" by blast show "\w w'. \ is_left_adjoint w; is_left_adjoint w'; p\<^sub>0 \ w \ u; v \ p\<^sub>1 \ w; p\<^sub>0 \ w' \ u; v \ p\<^sub>1 \ w' \ \ w \ w'" proof - fix w w' assume w: "is_left_adjoint w" and w': "is_left_adjoint w'" assume 1: "p\<^sub>0 \ w \ u" assume 2: "v \ p\<^sub>1 \ w" assume 3: "p\<^sub>0 \ w' \ u" assume 4: "v \ p\<^sub>1 \ w'" obtain \ where \: "\\ : p\<^sub>0 \ w \ u\" using 1 by auto obtain \' where \': "\\' : p\<^sub>0 \ w' \ u\" using 3 by auto obtain \ where \: "\\: v \ p\<^sub>1 \ w\ \ iso \" using 2 by blast obtain \' where \': "\\': v \ p\<^sub>1 \ w'\ \ iso \'" using 4 by blast let ?\ = "\' \ inv \" have \: "\?\ : p\<^sub>1 \ w \ p\<^sub>1 \ w'\" using \ \' by (elim conjE in_homE, auto) interpret uw\: uw\ V H \ \ src trg \k\<^sup>* \ h\ tab \p\<^sub>0\ \p\<^sub>1\ u w \ using w \ left_adjoint_is_ide by unfold_locales auto interpret uw'\': uw\ V H \ \ src trg \k\<^sup>* \ h\ tab \p\<^sub>0\ \p\<^sub>1\ u w' \' using w' \' left_adjoint_is_ide by unfold_locales auto interpret uw\w'\': uw\w'\' V H \ \ src trg \k\<^sup>* \ h\ tab \p\<^sub>0\ \p\<^sub>1\ u w \ w' \' using w w' \ \' left_adjoint_is_ide by unfold_locales have "(h \ \) \ \[h, p\<^sub>0, w] \ (\ \ w) \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w] = (h \ \') \ \[h, p\<^sub>0, w'] \ (\ \ w') \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w'] \ (k \ ?\)" proof - let ?LHS = "(h \ \) \ \[h, p\<^sub>0, w] \ (\ \ w) \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w]" let ?RHS = "(h \ \') \ \[h, p\<^sub>0, w'] \ (\ \ w') \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w'] \ (k \ ?\)" have "\?LHS : k \ p\<^sub>1 \ w \ h \ u\" using w k.antipar by fastforce moreover have "\?RHS : k \ p\<^sub>1 \ w \ h \ u\" using w k.antipar \ by fastforce moreover have "is_left_adjoint (k \ p\<^sub>1 \ 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'\ \ \ = \' \ (p\<^sub>0 \ \) \ p\<^sub>1 \ \ = ?\" 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 \ r\<^sub>0\\: \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 r\<^sub>0 \ \ arrow_of_spans_of_maps V H \ \ src trg r\<^sub>0 (dom \) s\<^sub>0 s\<^sub>1 w \ composite_cell w (arrow_of_spans_of_maps.the_\ V H r\<^sub>0 s\<^sub>0 w) \ (arrow_of_spans_of_maps.the_\ V H (dom \) s\<^sub>1 w) = \" lemma induced_map_unique: assumes "is_induced_by_cell w r\<^sub>0 \" and "is_induced_by_cell w' r\<^sub>0 \" shows "isomorphic w w'" proof - interpret w: arrow_of_spans_of_maps V H \ \ src trg r\<^sub>0 \dom \\ s\<^sub>0 s\<^sub>1 w using assms(1) by auto interpret w: arrow_of_spans_of_maps_to_tabulation V H \ \ src trg r\<^sub>0 \dom \\ s \ s\<^sub>0 s\<^sub>1 w .. interpret w': arrow_of_spans_of_maps V H \ \ src trg r\<^sub>0 \dom \\ s\<^sub>0 s\<^sub>1 w' using assms(2) by auto interpret w': arrow_of_spans_of_maps_to_tabulation V H \ \ src trg r\<^sub>0 \dom \\ s \ s\<^sub>0 s\<^sub>1 w' .. let ?\ = "w'.the_\ \ inv w.the_\" have \: "\?\ : s\<^sub>1 \ w \ s\<^sub>1 \ 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_\ = s\<^sub>1 \ 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_\ r\<^sub>0 w'.the_\ ?\] by blast thus ?thesis using BS3 w.is_map w'.is_map by blast qed text \ The object src \s\<^sub>0\ 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 s\<^sub>0) s\<^sub>0 \" proof - have 1: "arrow_of_spans_of_maps V H \ \ src trg s\<^sub>0 s\<^sub>1 s\<^sub>0 s\<^sub>1 (src s\<^sub>0)" using iso_runit [of s\<^sub>0] iso_runit [of s\<^sub>1] 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 s\<^sub>1 "src s\<^sub>0"] iso_runit' by blast interpret w: arrow_of_spans_of_maps V H \ \ src trg s\<^sub>0 \dom \\ s\<^sub>0 s\<^sub>1 \src s\<^sub>0\ using 1 tab_in_hom by simp interpret w: arrow_of_spans_of_maps_to_tabulation V H \ \ src trg s\<^sub>0 \dom \\ s \ s\<^sub>0 s\<^sub>1 \src s\<^sub>0\ .. show "is_induced_by_cell (src s\<^sub>0) s\<^sub>0 \" proof (intro conjI) show "arrow_of_spans_of_maps V H \ \ src trg s\<^sub>0 (dom \) s\<^sub>0 s\<^sub>1 (src s\<^sub>0)" using w.arrow_of_spans_of_maps_axioms by simp show "composite_cell (src s\<^sub>0) w.the_\ \ w.the_\ = \" proof - have \: "w.the_\ = \[s\<^sub>0]" using iso_runit [of s\<^sub>0] w.leg0_uniquely_isomorphic w.the_\_props the1_equality [of "\\. \\ : s\<^sub>0 \ src s\<^sub>0 \ s\<^sub>0\ \ iso \"] by auto have \: "w.the_\ = \\<^sup>-\<^sup>1[s\<^sub>1]" using iso_runit' w.leg1_uniquely_isomorphic w.the_\_props leg1_simps(3) the1_equality [of "\\. \\ : s\<^sub>1 \ s\<^sub>1 \ src s\<^sub>0\ \ iso \"] tab_in_vhom' by auto have "composite_cell (src s\<^sub>0) \[s\<^sub>0] \ \\<^sup>-\<^sup>1[s\<^sub>1] = \" proof - have "composite_cell (src s\<^sub>0) \[s\<^sub>0] \ \\<^sup>-\<^sup>1[s\<^sub>1] = ((s \ \[s\<^sub>0]) \ \[s, s\<^sub>0, src s\<^sub>0]) \ (\ \ src s\<^sub>0) \ \\<^sup>-\<^sup>1[s\<^sub>1]" using comp_assoc by presburger also have "... = (\[s \ s\<^sub>0] \ (\ \ src s\<^sub>0)) \ \\<^sup>-\<^sup>1[s\<^sub>1]" using runit_hcomp comp_assoc by simp also have "... = \ \ \[s\<^sub>1] \ \\<^sup>-\<^sup>1[s\<^sub>1]" 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 \(r\<^sub>0, \, r\<^sub>1)\ of \r\ and \(s\<^sub>0, \, s\<^sub>1)\ of \s\ in a bicategory of spans, where \(r\<^sub>0, r\<^sub>1)\ and \(s\<^sub>0, s\<^sub>1)\ 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 \(r\<^sub>0, s\<^sub>1)\ 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 \ r\<^sub>0 r\<^sub>1 + \: tabulation_in_maps V H \ \ src trg s \ s\<^sub>0 s\<^sub>1 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 r\<^sub>0 :: 'a and r\<^sub>1 :: 'a and s :: 'a and \ :: 'a and s\<^sub>0 :: 'a and s\<^sub>1 :: 'a + assumes composable: "src r = trg s" begin text \ Interpret \(r\<^sub>0, s\<^sub>1)\ 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 r\<^sub>0s\<^sub>1: cospan_of_maps_in_bicategory_of_spans V H \ \ src trg s\<^sub>1 r\<^sub>0 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 r\<^sub>0s\<^sub>1_is_cospan: shows "cospan_of_maps_in_bicategory_of_spans V H \ \ src trg s\<^sub>1 r\<^sub>0" .. text \ The following define the projections associated with the natural tabulation of \r\<^sub>0\<^sup>* \ s\<^sub>1\. \ abbreviation p\<^sub>0 where "p\<^sub>0 \ r\<^sub>0s\<^sub>1.p\<^sub>0" abbreviation p\<^sub>1 where "p\<^sub>1 \ r\<^sub>0s\<^sub>1.p\<^sub>1" 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 \ \\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0] \ (r \ \[s, s\<^sub>0, p\<^sub>0]) \ (r \ \ \ p\<^sub>0) \ (r \ r\<^sub>0s\<^sub>1.\) \ \[r, r\<^sub>0, p\<^sub>1] \ (\ \ p\<^sub>1)" lemma tab_in_hom [intro]: shows "\tab : r\<^sub>1 \ p\<^sub>1 \ (r \ s) \ s\<^sub>0 \ p\<^sub>0\" using \.T0.antipar(1) r\<^sub>0s\<^sub>1.\_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 \s\<^sub>0 \ p\<^sub>0\ \r\<^sub>1 \ p\<^sub>1\ 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 "\\ : (s\<^sub>0 \ p\<^sub>0) \ w \ u\" and "w\<^sub>r = p\<^sub>1 \ w" and "\\<^sub>r = (s \ \) \ (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]) \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]" shows "\.composite_cell w\<^sub>r \\<^sub>r = \[r, s, u] \ composite_cell w \ \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, 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 \s\<^sub>0 \ p\<^sub>0\ \r\<^sub>1 \ p\<^sub>1\ u w \ using assms(1-2) composable by unfold_locales auto show ?thesis proof - have "\[r, s, u] \ composite_cell w \ \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w] = (\[r, s, u] \ ((r \ s) \ \)) \ \[r \ s, s\<^sub>0 \ p\<^sub>0, w] \ (tab \ w) \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]" using comp_assoc by presburger also have "... = (r \ s \ \) \ \[r, s, (s\<^sub>0 \ p\<^sub>0) \ w] \ \[r \ s, s\<^sub>0 \ p\<^sub>0, w] \ (tab \ w) \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]" using assoc_naturality [of r s \] composable comp_assoc by simp also have "... = (r \ s \ \) \ \[r, s, (s\<^sub>0 \ p\<^sub>0) \ w] \ \[r \ s, s\<^sub>0 \ p\<^sub>0, w] \ ((\\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0] \ (r \ \[s, s\<^sub>0, p\<^sub>0])) \ (r \ \ \ p\<^sub>0) \ \.composite_cell p\<^sub>1 r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]" unfolding tab_def using comp_assoc by presburger also have "... = (r \ s \ \) \ ((\[r, s, (s\<^sub>0 \ p\<^sub>0) \ w] \ \[r \ s, s\<^sub>0 \ p\<^sub>0, w] \ (\\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0] \ (r \ \[s, s\<^sub>0, p\<^sub>0]) \ w))) \ ((r \ \ \ p\<^sub>0) \ \.composite_cell p\<^sub>1 r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]" using composable \.T0.antipar(1) comp_assoc whisker_right by auto also have "... = (r \ s \ \) \ ((\[r, s, (s\<^sub>0 \ p\<^sub>0) \ w] \ \[r \ s, s\<^sub>0 \ p\<^sub>0, w] \ (\\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0] \ (r \ \[s, s\<^sub>0, p\<^sub>0]) \ w))) \ ((r \ \ \ p\<^sub>0) \ w) \ ((r \ r\<^sub>0s\<^sub>1.\) \ w) \ (\[r, r\<^sub>0, p\<^sub>1] \ w) \ ((\ \ p\<^sub>1) \ w) \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, 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, (s\<^sub>0 \ p\<^sub>0) \ w] \ \[r \ s, s\<^sub>0 \ p\<^sub>0, w] \ (\\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0] \ (r \ \[s, s\<^sub>0, p\<^sub>0]) \ w))) \ ((r \ \ \ p\<^sub>0) \ w) \ ((r \ r\<^sub>0s\<^sub>1.\) \ w) \ ((\[r, r\<^sub>0, p\<^sub>1] \ w) \ \\<^sup>-\<^sup>1[r \ r\<^sub>0, p\<^sub>1, w]) \ (\ \ p\<^sub>1 \ w)" using assoc'_naturality [of \ p\<^sub>1 w] \.T0.antipar(1) r\<^sub>0s\<^sub>1.base_simps(2) comp_assoc by auto also have "... = (r \ s \ \) \ ((\[r, s, (s\<^sub>0 \ p\<^sub>0) \ w] \ \[r \ s, s\<^sub>0 \ p\<^sub>0, w] \ (\\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0] \ (r \ \[s, s\<^sub>0, p\<^sub>0]) \ w))) \ ((r \ \ \ p\<^sub>0) \ w) \ (((r \ r\<^sub>0s\<^sub>1.\) \ w) \ \\<^sup>-\<^sup>1[r, r\<^sub>0 \ p\<^sub>1, w]) \ \.composite_cell (p\<^sub>1 \ w) \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]" proof - have "(\[r, r\<^sub>0, p\<^sub>1] \ w) \ \\<^sup>-\<^sup>1[r \ r\<^sub>0, p\<^sub>1, w] = \\<^sup>-\<^sup>1[r, r\<^sub>0 \ p\<^sub>1, w] \ (r \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \ \[r, r\<^sub>0, p\<^sub>1 \ w]" proof - have "(\\<^sup>-\<^sup>1[r, r\<^sub>0, p\<^sub>1] \ w) \ \\<^sup>-\<^sup>1[r, r\<^sub>0 \ p\<^sub>1, w] \ (r \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) = \\<^sup>-\<^sup>1[r \ r\<^sub>0, p\<^sub>1, w] \ \\<^sup>-\<^sup>1[r, r\<^sub>0, p\<^sub>1 \ w]" using pentagon' \.T0.antipar(1) comp_assoc by simp moreover have 1: "seq (\\<^sup>-\<^sup>1[r, r\<^sub>0, p\<^sub>1] \ w)(\\<^sup>-\<^sup>1[r, r\<^sub>0 \ p\<^sub>1, w] \ (r \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]))" using \.T0.antipar(1) by (intro seqI hseqI, auto) ultimately have "\\<^sup>-\<^sup>1[r \ r\<^sub>0, p\<^sub>1, w] = ((\\<^sup>-\<^sup>1[r, r\<^sub>0, p\<^sub>1] \ w) \ \\<^sup>-\<^sup>1[r, r\<^sub>0 \ p\<^sub>1, w] \ (r \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w])) \ \[r, r\<^sub>0, p\<^sub>1 \ w]" using \.T0.antipar(1) iso_inv_iso iso_assoc inv_inv invert_side_of_triangle(2) [of "(\\<^sup>-\<^sup>1[r, r\<^sub>0, p\<^sub>1] \ w) \ \\<^sup>-\<^sup>1[r, r\<^sub>0 \ p\<^sub>1, w] \ (r \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w])" "\\<^sup>-\<^sup>1[r \ r\<^sub>0, p\<^sub>1, w]" "\\<^sup>-\<^sup>1[r, r\<^sub>0, p\<^sub>1 \ w]"] by fastforce hence "\\<^sup>-\<^sup>1[r \ r\<^sub>0, p\<^sub>1, w] = (\\<^sup>-\<^sup>1[r, r\<^sub>0, p\<^sub>1] \ w) \ \\<^sup>-\<^sup>1[r, r\<^sub>0 \ p\<^sub>1, w] \ (r \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \ \[r, r\<^sub>0, p\<^sub>1 \ w]" using comp_assoc by presburger moreover have "seq (inv (\\<^sup>-\<^sup>1[r, r\<^sub>0, p\<^sub>1] \ w)) \\<^sup>-\<^sup>1[r \ r\<^sub>0, p\<^sub>1, w]" using \.T0.antipar(1) iso_inv_iso 1 by fastforce ultimately show ?thesis using \.T0.antipar(1) iso_inv_iso iso_assoc inv_inv inv_hcomp invert_side_of_triangle(1) [of "\\<^sup>-\<^sup>1[r \ r\<^sub>0, p\<^sub>1, w]" "\\<^sup>-\<^sup>1[r, r\<^sub>0, p\<^sub>1] \ w" "\\<^sup>-\<^sup>1[r, r\<^sub>0 \ p\<^sub>1, w] \ (r \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \ \[r, r\<^sub>0, p\<^sub>1 \ w]"] by fastforce qed thus ?thesis using comp_assoc by presburger qed also have "... = (r \ s \ \) \ ((\[r, s, (s\<^sub>0 \ p\<^sub>0) \ w] \ \[r \ s, s\<^sub>0 \ p\<^sub>0, w] \ (\\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0] \ (r \ \[s, s\<^sub>0, p\<^sub>0]) \ w))) \ (((r \ \ \ p\<^sub>0) \ w) \ \\<^sup>-\<^sup>1[r, s\<^sub>1 \ p\<^sub>0, w]) \ (r \ r\<^sub>0s\<^sub>1.\ \ w) \ \.composite_cell (p\<^sub>1 \ w) \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]" proof - have "((r \ r\<^sub>0s\<^sub>1.\) \ w) \ \\<^sup>-\<^sup>1[r, r\<^sub>0 \ p\<^sub>1, w] = \\<^sup>-\<^sup>1[r, s\<^sub>1 \ p\<^sub>0, w] \ (r \ r\<^sub>0s\<^sub>1.\ \ w)" using assoc'_naturality [of r r\<^sub>0s\<^sub>1.\ w] r\<^sub>0s\<^sub>1.cospan by auto thus ?thesis using comp_assoc by presburger qed also have "... = (r \ s \ \) \ (\[r, s, (s\<^sub>0 \ p\<^sub>0) \ w] \ \[r \ s, s\<^sub>0 \ p\<^sub>0, w] \ (\\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0] \ (r \ \[s, s\<^sub>0, p\<^sub>0]) \ w) \ \\<^sup>-\<^sup>1[r, (s \ s\<^sub>0) \ p\<^sub>0, w]) \ (r \ (\ \ p\<^sub>0) \ w) \ (r \ r\<^sub>0s\<^sub>1.\ \ w) \ \.composite_cell (p\<^sub>1 \ w) \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]" proof - have "((r \ \ \ p\<^sub>0) \ w) \ \\<^sup>-\<^sup>1[r, s\<^sub>1 \ p\<^sub>0, w] = \\<^sup>-\<^sup>1[r, (s \ s\<^sub>0) \ p\<^sub>0, w] \ (r \ (\ \ p\<^sub>0) \ w)" using assoc'_naturality [of r "\ \ p\<^sub>0" w] by (simp add: composable) thus ?thesis using comp_assoc by presburger qed also have "... = (r \ s \ \) \ (r \ (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]) \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ \[s \ s\<^sub>0, p\<^sub>0, w]) \ ((r \ (\ \ p\<^sub>0) \ w) \ (r \ r\<^sub>0s\<^sub>1.\ \ w) \ (r \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w])) \ \[r, r\<^sub>0, p\<^sub>1 \ w] \ (\ \ p\<^sub>1 \ w)" proof - have "\[r, s, (s\<^sub>0 \ p\<^sub>0) \ w] \ \[r \ s, s\<^sub>0 \ p\<^sub>0, w] \ (\\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0] \ (r \ \[s, s\<^sub>0, p\<^sub>0]) \ w) \ \\<^sup>-\<^sup>1[r, (s \ s\<^sub>0) \ p\<^sub>0, w] = r \ (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]) \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ \[s \ s\<^sub>0, p\<^sub>0, w]" proof - have "\[r, s, (s\<^sub>0 \ p\<^sub>0) \ w] \ \[r \ s, s\<^sub>0 \ p\<^sub>0, w] \ (\\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0] \ (r \ \[s, s\<^sub>0, p\<^sub>0]) \ w) \ \\<^sup>-\<^sup>1[r, (s \ s\<^sub>0) \ p\<^sub>0, w] = \\<^bold>\\<^bold>[\<^bold>\r\<^bold>\, \<^bold>\s\<^bold>\, (\<^bold>\s\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\w\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\<^bold>[\<^bold>\r\<^bold>\ \<^bold>\ \<^bold>\s\<^bold>\, \<^bold>\s\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\p\<^sub>0\<^bold>\, \<^bold>\w\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\r\<^bold>\, \<^bold>\s\<^bold>\, \<^bold>\s\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\p\<^sub>0\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\r\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\s\<^bold>\, \<^bold>\s\<^sub>0\<^bold>\, \<^bold>\p\<^sub>0\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\w\<^bold>\) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\r\<^bold>\, (\<^bold>\s\<^bold>\ \<^bold>\ \<^bold>\s\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\p\<^sub>0\<^bold>\, \<^bold>\w\<^bold>\\<^bold>]\" using \_def \'_def composable by simp also have "... = \\<^bold>\r\<^bold>\ \<^bold>\ (\<^bold>\s\<^bold>\ \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\s\<^sub>0\<^bold>\, \<^bold>\p\<^sub>0\<^bold>\, \<^bold>\w\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\s\<^bold>\, \<^bold>\s\<^sub>0\<^bold>\, \<^bold>\p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\w\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\<^bold>[\<^bold>\s\<^bold>\ \<^bold>\ \<^bold>\s\<^sub>0\<^bold>\, \<^bold>\p\<^sub>0\<^bold>\, \<^bold>\w\<^bold>\\<^bold>]\" using composable by (intro E.eval_eqI, simp_all) also have "... = r \ (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]) \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ \[s \ s\<^sub>0, p\<^sub>0, 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 \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]) \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ \[s \ s\<^sub>0, p\<^sub>0, w]) \ \.composite_cell (p\<^sub>1 \ w) (((\ \ p\<^sub>0) \ w) \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w])" using assms(3) arrI \.T0.antipar(1) whisker_left by auto also have "... = (r \ (s \ \) \ (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]) \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ (\[s \ s\<^sub>0, p\<^sub>0, w] \ ((\ \ p\<^sub>0) \ w)) \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \ \[r, r\<^sub>0, p\<^sub>1 \ w] \ (\ \ p\<^sub>1 \ w)" using \.T0.antipar(1) comp_assoc whisker_left by auto also have "... = (r \ (s \ \) \ (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]) \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \ \[r, r\<^sub>0, p\<^sub>1 \ w] \ (\ \ p\<^sub>1 \ w)" using assoc_naturality [of \ p\<^sub>0 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 (s\<^sub>0 \ p\<^sub>0) (r\<^sub>1 \ p\<^sub>1)" proof show "\u \. \ ide u; \\ : dom \ \ (r \ s) \ u\ \ \ \w \ \. ide w \ \\ : (s\<^sub>0 \ p\<^sub>0) \ w \ u\ \ \\ : dom \ \ (r\<^sub>1 \ p\<^sub>1) \ 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 w\<^sub>r \\<^sub>r \\<^sub>r where w\<^sub>r\\<^sub>r\\<^sub>r: "ide w\<^sub>r \ \\\<^sub>r : r\<^sub>0 \ w\<^sub>r \ s \ u\ \ \\\<^sub>r : ?v \ r\<^sub>1 \ w\<^sub>r\ \ iso \\<^sub>r \ \.composite_cell w\<^sub>r \\<^sub>r \ \\<^sub>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 w\<^sub>r: "ide w\<^sub>r" using w\<^sub>r\\<^sub>r\\<^sub>r by simp have [simp]: "src w\<^sub>r = src u" using w\<^sub>r\\<^sub>r\\<^sub>r \ 1 comp_arr_dom in_homE seqE hcomp_simps(1) vseq_implies_hpar(1) by (metis src_hcomp) have [simp]: "trg w\<^sub>r = src \" using w\<^sub>r\\<^sub>r\\<^sub>r by (metis 1 arrI not_arr_null seqE seq_if_composable) have \\<^sub>r_in_hom [intro]: "\\\<^sub>r : r\<^sub>0 \ w\<^sub>r \ s \ u\" using w\<^sub>r\\<^sub>r\\<^sub>r by simp have \\<^sub>r_in_hhom [intro]: "\\\<^sub>r : src u \ trg s\" using \\<^sub>r_in_hom src_cod [of \\<^sub>r] trg_cod [of \\<^sub>r] by (metis \src w\<^sub>r = src u\ \.leg1_simps(4) arr_dom in_hhomI in_homE r\<^sub>0s\<^sub>1.cospan src_hcomp trg_hcomp vconn_implies_hpar(1) vconn_implies_hpar(2)) have [simp]: "src \\<^sub>r = src u" using \\<^sub>r_in_hhom by auto have [simp]: "trg \\<^sub>r = trg s" using \\<^sub>r_in_hhom by auto have [simp]: "dom \\<^sub>r = r\<^sub>0 \ w\<^sub>r" using \\<^sub>r_in_hom by blast have [simp]: "cod \\<^sub>r = s \ u" using \\<^sub>r_in_hom by blast have \\<^sub>r_in_hom [intro]: "\\\<^sub>r : ?v \ r\<^sub>1 \ w\<^sub>r\" using w\<^sub>r\\<^sub>r\\<^sub>r by simp have \\<^sub>r_in_hhom [intro]: "\\\<^sub>r : src u \ trg r\" using \\<^sub>r_in_hom src_dom [of \\<^sub>r] trg_dom [of \\<^sub>r] by (metis \src w\<^sub>r = 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 \\<^sub>r = src u" using \\<^sub>r_in_hhom by auto have [simp]: "trg \\<^sub>r = trg r" using \\<^sub>r_in_hhom by auto have [simp]: "dom \\<^sub>r = ?v" using \\<^sub>r_in_hom by auto have [simp]: "cod \\<^sub>r = r\<^sub>1 \ w\<^sub>r" using \\<^sub>r_in_hom by auto have iso_\\<^sub>r: "iso \\<^sub>r" using w\<^sub>r\\<^sub>r\\<^sub>r by simp obtain w\<^sub>s \\<^sub>s \\<^sub>s where w\<^sub>s\\<^sub>s\\<^sub>s: "ide w\<^sub>s \ \\\<^sub>s : s\<^sub>0 \ w\<^sub>s \ u\ \ \\\<^sub>s : r\<^sub>0 \ w\<^sub>r \ s\<^sub>1 \ w\<^sub>s\ \ iso \\<^sub>s \ \.composite_cell w\<^sub>s \\<^sub>s \ \\<^sub>s = \\<^sub>r" using u w\<^sub>r\\<^sub>r\\<^sub>r \.T1 [of u \\<^sub>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 w\<^sub>s: "ide w\<^sub>s" using w\<^sub>s\\<^sub>s\\<^sub>s by simp have [simp]: "src w\<^sub>s = src u" using w\<^sub>s\\<^sub>s\\<^sub>s src_cod by (metis \.leg0_simps(2) \.tab_simps(2) \\<^sub>r_in_hom arrI hseqI' ideD(1) seqE seq_if_composable src_hcomp vconn_implies_hpar(3)) have [simp]: "trg w\<^sub>s = src \" using w\<^sub>s\\<^sub>s\\<^sub>s by (metis \.tab_simps(2) arr_dom in_homE not_arr_null seq_if_composable) have \\<^sub>s_in_hom [intro]: "\\\<^sub>s : s\<^sub>0 \ w\<^sub>s \ u\" using w\<^sub>s\\<^sub>s\\<^sub>s by simp have \\<^sub>s_in_hhom [intro]: "\\\<^sub>s : src u \ src s\" using \\<^sub>s_in_hom src_cod trg_cod by (metis \\<^sub>r_in_hom arrI hseqE in_hhom_def seqE vconn_implies_hpar(1) vconn_implies_hpar(3) w\<^sub>s\\<^sub>s\\<^sub>s) have [simp]: "src \\<^sub>s = src u" using \\<^sub>s_in_hhom by auto have [simp]: "trg \\<^sub>s = src s" using \\<^sub>s_in_hhom by auto have [simp]: "dom \\<^sub>s = s\<^sub>0 \ w\<^sub>s" using \\<^sub>s_in_hom by blast have [simp]: "cod \\<^sub>s = u" using \\<^sub>s_in_hom by blast have \\<^sub>s_in_hom [intro]: "\\\<^sub>s : r\<^sub>0 \ w\<^sub>r \ s\<^sub>1 \ w\<^sub>s\" using w\<^sub>s\\<^sub>s\\<^sub>s by simp have \\<^sub>s_in_hhom [intro]: "\\\<^sub>s : src u \ trg s\" using \\<^sub>s_in_hom src_dom trg_cod by (metis \src \\<^sub>r = src u\ \trg \\<^sub>r = trg s\ \\<^sub>r_in_hom in_hhomI in_homE src_dom trg_dom) have [simp]: "src \\<^sub>s = src u" using \\<^sub>s_in_hhom by auto have [simp]: "trg \\<^sub>s = trg s" using \\<^sub>s_in_hhom by auto have [simp]: "dom \\<^sub>s = r\<^sub>0 \ w\<^sub>r" using \\<^sub>s_in_hom by auto have [simp]: "cod \\<^sub>s = s\<^sub>1 \ w\<^sub>s" using \\<^sub>s_in_hom by auto have iso_\\<^sub>s: "iso \\<^sub>s" using w\<^sub>s\\<^sub>s\\<^sub>s by simp obtain w \\<^sub>t \\<^sub>t where w\\<^sub>t\\<^sub>t: "ide w \ \\\<^sub>t : p\<^sub>0 \ w \ w\<^sub>s\ \ \\\<^sub>t : w\<^sub>r \ p\<^sub>1 \ w\ \ iso \\<^sub>t \ (s\<^sub>1 \ \\<^sub>t) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \ (r\<^sub>0 \ \\<^sub>t) = \\<^sub>s" using w\<^sub>r\\<^sub>r\\<^sub>r w\<^sub>s\\<^sub>s\\<^sub>s iso_\\<^sub>s r\<^sub>0s\<^sub>1.\_biuniversal_prop(1) [of w\<^sub>s w\<^sub>r \\<^sub>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 w\\<^sub>t\\<^sub>t by simp have [simp]: "src w = src u" using w\\<^sub>t\\<^sub>t src_cod by (metis \\<^sub>s_in_hom \src \\<^sub>s = src u\ arrI seqE src_hcomp src_vcomp vseq_implies_hpar(1)) have [simp]: "trg w = src p\<^sub>0" using w\\<^sub>t\\<^sub>t by (metis \\<^sub>s_in_hom arrI not_arr_null r\<^sub>0s\<^sub>1.\_simps(2) seqE seq_if_composable) have \\<^sub>t_in_hom [intro]: "\\\<^sub>t : p\<^sub>0 \ w \ w\<^sub>s\" using w\\<^sub>t\\<^sub>t by simp have \\<^sub>t_in_hhom [intro]: "\\\<^sub>t : src u \ src \\" using \\<^sub>t_in_hom src_cod trg_cod \src w\<^sub>s = src u\ \trg w\<^sub>s = src \\ by fastforce have [simp]: "src \\<^sub>t = src u" using \\<^sub>t_in_hhom by auto have [simp]: "trg \\<^sub>t = src \" using \\<^sub>t_in_hhom by auto have [simp]: "dom \\<^sub>t = p\<^sub>0 \ w" using \\<^sub>t_in_hom by blast have (* [simp]: *) "cod \\<^sub>t = w\<^sub>s" using \\<^sub>t_in_hom by blast have \\<^sub>t_in_hom [intro]: "\\\<^sub>t : w\<^sub>r \ p\<^sub>1 \ w\" using w\\<^sub>t\\<^sub>t by simp have \\<^sub>t_in_hhom [intro]: "\\\<^sub>t : src u \ src \\" using \\<^sub>t_in_hom src_dom trg_dom \src w\<^sub>r = src u\ \trg w\<^sub>r = src \\ by fastforce have [simp]: "src \\<^sub>t = src u" using \\<^sub>t_in_hhom by auto have [simp]: "trg \\<^sub>t = src \" using \\<^sub>t_in_hhom by auto have (* [simp]: *) "dom \\<^sub>t = w\<^sub>r" using \\<^sub>t_in_hom by auto have [simp]: "cod \\<^sub>t = p\<^sub>1 \ w" using \\<^sub>t_in_hom by auto have iso_\\<^sub>t: "iso \\<^sub>t" using w\\<^sub>t\\<^sub>t by simp define \ where "\ = \\<^sub>s \ (s\<^sub>0 \ \\<^sub>t) \ \[s\<^sub>0, p\<^sub>0, w]" have \: "\\ : (s\<^sub>0 \ p\<^sub>0) \ w \ u\" proof (unfold \_def, intro comp_in_homI) show "\\[s\<^sub>0, p\<^sub>0, w] : (s\<^sub>0 \ p\<^sub>0) \ w \ s\<^sub>0 \ p\<^sub>0 \ w\" using w\\<^sub>t\\<^sub>t by auto show "\s\<^sub>0 \ \\<^sub>t : s\<^sub>0 \ p\<^sub>0 \ w \ s\<^sub>0 \ w\<^sub>s\" using w\\<^sub>t\\<^sub>t by auto show "\\\<^sub>s : s\<^sub>0 \ w\<^sub>s \ u\" using w\<^sub>s\\<^sub>s\\<^sub>s by simp qed define \ where "\ = \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w] \ (r\<^sub>1 \ \\<^sub>t) \ \\<^sub>r" have \: "\\ : ?v \ (r\<^sub>1 \ p\<^sub>1) \ w\" proof (unfold \_def, intro comp_in_homI) show "\\\<^sub>r : ?v \ r\<^sub>1 \ w\<^sub>r\" using w\<^sub>r\\<^sub>r\\<^sub>r by blast show "\r\<^sub>1 \ \\<^sub>t : r\<^sub>1 \ w\<^sub>r \ r\<^sub>1 \ p\<^sub>1 \ w\" using w\\<^sub>t\\<^sub>t by (intro hcomp_in_vhom, auto) show "\\\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w] : r\<^sub>1 \ p\<^sub>1 \ w \ (r\<^sub>1 \ p\<^sub>1) \ w\" using w\\<^sub>t\\<^sub>t assoc_in_hom by (simp add: \.T0.antipar(1)) qed have iso_\: "iso \" using \ w\\<^sub>t\\<^sub>t w\<^sub>r\\<^sub>r\\<^sub>r \.T0.antipar(1) iso_inv_iso apply (unfold \_def, intro isos_compose) by auto have *: "arr ((s \ \\<^sub>s) \ \[s, s\<^sub>0, w\<^sub>s] \ (\ \ w\<^sub>s) \ (s\<^sub>1 \ \\<^sub>t) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \ (r\<^sub>0 \ \\<^sub>t))" using w\<^sub>s\\<^sub>s\\<^sub>s w\\<^sub>t\\<^sub>t \\<^sub>r_in_hom comp_assoc by auto have "((r \ s) \ \) \ \[r \ s, s\<^sub>0 \ p\<^sub>0, w] \ (tab \ w) \ \ = \" proof - have "seq (r \ \\<^sub>r) (\[r, r\<^sub>0, w\<^sub>r] \ (\ \ w\<^sub>r) \ \\<^sub>r)" using w\<^sub>r\\<^sub>r\\<^sub>r \.base_simps(2) composable by fastforce hence "\ = \\<^sup>-\<^sup>1[r, s, u] \ \.composite_cell w\<^sub>r \\<^sub>r \ \\<^sub>r" using w\<^sub>r\\<^sub>r\\<^sub>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) w\<^sub>s\\<^sub>s\\<^sub>s) also have "... = \\<^sup>-\<^sup>1[r, s, u] \ \.composite_cell w\<^sub>r (\.composite_cell w\<^sub>s \\<^sub>s \ \\<^sub>s) \ \\<^sub>r" using w\<^sub>s\\<^sub>s\\<^sub>s by simp also have "... = \\<^sup>-\<^sup>1[r, s, u] \ (r \ (s \ \\<^sub>s) \ \[s, s\<^sub>0, w\<^sub>s] \ (\ \ w\<^sub>s) \ (s\<^sub>1 \ \\<^sub>t) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \ (r\<^sub>0 \ \\<^sub>t)) \ \[r, r\<^sub>0, w\<^sub>r] \ (\ \ w\<^sub>r) \ \\<^sub>r" using w\\<^sub>t\\<^sub>t comp_assoc by simp text \Rearrange to create \\\ and \\\, leaving \tab\ in the middle.\ also have "... = \\<^sup>-\<^sup>1[r, s, u] \ (r \ (s \ \\<^sub>s) \ \[s, s\<^sub>0, w\<^sub>s] \ ((\ \ w\<^sub>s) \ (s\<^sub>1 \ \\<^sub>t)) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \ (r\<^sub>0 \ \\<^sub>t)) \ \[r, r\<^sub>0, w\<^sub>r] \ (\ \ w\<^sub>r) \ \\<^sub>r" using comp_assoc by presburger also have "... = \\<^sup>-\<^sup>1[r, s, u] \ (r \ (s \ \\<^sub>s) \ (\[s, s\<^sub>0, w\<^sub>s] \ ((s \ s\<^sub>0) \ \\<^sub>t)) \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \ (r\<^sub>0 \ \\<^sub>t)) \ \[r, r\<^sub>0, w\<^sub>r] \ (\ \ w\<^sub>r) \ \\<^sub>r" proof - have "(\ \ w\<^sub>s) \ (s\<^sub>1 \ \\<^sub>t) = \ \ \\<^sub>t" using comp_arr_dom comp_cod_arr interchange by (metis \cod \\<^sub>t = w\<^sub>s\ \.tab_simps(1) \.tab_simps(4) arrI w\\<^sub>t\\<^sub>t) also have "... = ((s \ s\<^sub>0) \ \\<^sub>t) \ (\ \ p\<^sub>0 \ w)" using comp_arr_dom comp_cod_arr interchange w\<^sub>s\\<^sub>s\\<^sub>s w\\<^sub>t\\<^sub>t \.tab_in_hom by (metis \dom \\<^sub>t = p\<^sub>0 \ w\ \.tab_simps(5) arrI) finally have "(\ \ w\<^sub>s) \ (s\<^sub>1 \ \\<^sub>t) = ((s \ s\<^sub>0) \ \\<^sub>t) \ (\ \ p\<^sub>0 \ w)" by simp thus ?thesis using comp_assoc by presburger qed also have "... = \\<^sup>-\<^sup>1[r, s, u] \ (r \ (s \ \\<^sub>s) \ (s \ s\<^sub>0 \ \\<^sub>t) \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \ (r\<^sub>0 \ \\<^sub>t)) \ \[r, r\<^sub>0, w\<^sub>r] \ (\ \ w\<^sub>r) \ \\<^sub>r" using assoc_naturality [of s s\<^sub>0 \\<^sub>t] w\\<^sub>t\\<^sub>t comp_assoc \cod \\<^sub>t = w\<^sub>s\ arrI by force also have "... = \\<^sup>-\<^sup>1[r, s, u] \ (r \ (s \ \\<^sub>s) \ (s \ s\<^sub>0 \ \\<^sub>t)) \ (r \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \ (r\<^sub>0 \ \\<^sub>t)) \ \[r, r\<^sub>0, w\<^sub>r] \ (\ \ w\<^sub>r) \ \\<^sub>r" proof - have "r \ (s \ \\<^sub>s) \ (s \ s\<^sub>0 \ \\<^sub>t) \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \ (r\<^sub>0 \ \\<^sub>t) = (r \ (s \ \\<^sub>s) \ (s \ s\<^sub>0 \ \\<^sub>t)) \ (r \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \ (r\<^sub>0 \ \\<^sub>t))" proof - have "seq ((s \ \\<^sub>s) \ (s \ s\<^sub>0 \ \\<^sub>t)) (\[s, s\<^sub>0, p\<^sub>0 \ w] \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \ (r\<^sub>0 \ \\<^sub>t))" proof - have "seq (s \ \\<^sub>s) ((s \ s\<^sub>0 \ \\<^sub>t) \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \ (r\<^sub>0 \ \\<^sub>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 \ \\<^sub>s) \ (s \ s\<^sub>0 \ \\<^sub>t)" "\[s, s\<^sub>0, p\<^sub>0 \ w] \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \ (r\<^sub>0 \ \\<^sub>t)"] w\<^sub>s\\<^sub>s\\<^sub>s w\\<^sub>t\\<^sub>t comp_assoc by simp qed thus ?thesis using comp_assoc by presburger qed also have "... = \\<^sup>-\<^sup>1[r, s, u] \ (r \ (s \ \\<^sub>s) \ (s \ s\<^sub>0 \ \\<^sub>t)) \ (r \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \ ((r \ r\<^sub>0 \ \\<^sub>t) \ \[r, r\<^sub>0, w\<^sub>r]) \ (\ \ w\<^sub>r) \ \\<^sub>r" proof - have "seq (\[s, s\<^sub>0, p\<^sub>0 \ w] \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) (r\<^sub>0 \ \\<^sub>t)" using 1 r\<^sub>0s\<^sub>1.p\<^sub>1_simps w\\<^sub>t\\<^sub>t apply (intro seqI' comp_in_homI) by auto hence "r \ (\[s, s\<^sub>0, p\<^sub>0 \ w] \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \ (r\<^sub>0 \ \\<^sub>t) = (r \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \ (r \ r\<^sub>0 \ \\<^sub>t)" using whisker_left by simp thus ?thesis using comp_assoc by simp qed also have "... = \\<^sup>-\<^sup>1[r, s, u] \ (r \ (s \ \\<^sub>s) \ (s \ s\<^sub>0 \ \\<^sub>t)) \ (r \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \ \[r, r\<^sub>0, p\<^sub>1 \ w] \ (((r \ r\<^sub>0) \ \\<^sub>t) \ (\ \ w\<^sub>r)) \ \\<^sub>r" proof - have "(r \ r\<^sub>0 \ \\<^sub>t) \ \[r, r\<^sub>0, w\<^sub>r] = \[r, r\<^sub>0, p\<^sub>1 \ w] \ ((r \ r\<^sub>0) \ \\<^sub>t)" using assoc_naturality [of r r\<^sub>0 \\<^sub>t] \\<^sub>t_in_hom by auto thus ?thesis using comp_assoc by presburger qed also have "... = (\\<^sup>-\<^sup>1[r, s, u] \ (r \ (s \ \\<^sub>s) \ (s \ s\<^sub>0 \ \\<^sub>t))) \ (r \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \ \[r, r\<^sub>0, p\<^sub>1 \ w] \ (\ \ p\<^sub>1 \ w) \ (r\<^sub>1 \ \\<^sub>t) \ \\<^sub>r" proof - have "((r \ r\<^sub>0) \ \\<^sub>t) \ (\ \ w\<^sub>r) = \ \ \\<^sub>t" using comp_arr_dom comp_cod_arr interchange by (metis \dom \\<^sub>t = w\<^sub>r\ \.tab_simps(1) \.tab_simps(5) arrI w\\<^sub>t\\<^sub>t) also have "... = (\ \ p\<^sub>1 \ w) \ (r\<^sub>1 \ \\<^sub>t)" using comp_arr_dom comp_cod_arr interchange by (metis \cod \\<^sub>t = p\<^sub>1 \ w\ \trg \\<^sub>t = src \\ \.T0.antipar(1) \.tab_simps(1) \.tab_simps(2) \.tab_simps(4) r\<^sub>0s\<^sub>1.base_simps(2) trg.preserves_reflects_arr trg_hcomp) finally have "((r \ r\<^sub>0) \ \\<^sub>t) \ (\ \ w\<^sub>r) = (\ \ p\<^sub>1 \ w) \ (r\<^sub>1 \ \\<^sub>t)" by simp thus ?thesis using comp_assoc by presburger qed also have "... = ((r \ s) \ \\<^sub>s \ (s\<^sub>0 \ \\<^sub>t)) \ \\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0 \ w] \ (r \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ ((\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w]) \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \ \[r, r\<^sub>0, p\<^sub>1 \ w] \ (\ \ p\<^sub>1 \ w) \ (r\<^sub>1 \ \\<^sub>t) \ \\<^sub>r" proof - have "\\<^sup>-\<^sup>1[r, s, u] \ (r \ (s \ \\<^sub>s) \ (s \ s\<^sub>0 \ \\<^sub>t)) = ((r \ s) \ \\<^sub>s \ (s\<^sub>0 \ \\<^sub>t)) \ \\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0 \ w]" proof - have "seq (s \ \\<^sub>s) (s \ s\<^sub>0 \ \\<^sub>t)" using \\<^sub>s_in_hom \\<^sub>s_in_hhom \\<^sub>t_in_hom \\<^sub>t_in_hhom 1 calculation by blast moreover have "src r = trg (s \ \\<^sub>s)" using composable hseqI by force ultimately have "\\<^sup>-\<^sup>1[r, s, u] \ (r \ (s \ \\<^sub>s) \ (s \ s\<^sub>0 \ \\<^sub>t)) = (\\<^sup>-\<^sup>1[r, s, u] \ (r \ s \ \\<^sub>s)) \ (r \ s \ s\<^sub>0 \ \\<^sub>t)" using whisker_left comp_assoc by simp also have "... = ((r \ s) \ \\<^sub>s) \ \\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ w\<^sub>s] \ (r \ s \ s\<^sub>0 \ \\<^sub>t)" using assoc_naturality comp_assoc by (metis \cod \\<^sub>s = u\ \dom \\<^sub>s = s\<^sub>0 \ w\<^sub>s\ \trg \\<^sub>s = src s\ \.base_simps(2-4) \.base_simps(2-4) arrI assoc'_naturality composable w\<^sub>s\\<^sub>s\\<^sub>s) also have "... = (((r \ s) \ \\<^sub>s) \ ((r \ s) \ s\<^sub>0 \ \\<^sub>t)) \ \\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0 \ w]" proof - have "\\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ w\<^sub>s] \ (r \ s \ s\<^sub>0 \ \\<^sub>t) = ((r \ s) \ s\<^sub>0 \ \\<^sub>t) \ \\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0 \ w]" using arrI hseq_char assoc'_naturality [of r s "s\<^sub>0 \ \\<^sub>t"] \cod \\<^sub>t = w\<^sub>s\ composable by auto thus ?thesis using comp_assoc by auto qed also have "... = ((r \ s) \ \\<^sub>s \ (s\<^sub>0 \ \\<^sub>t)) \ \\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0 \ 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) \ \\<^sub>s \ (s\<^sub>0 \ \\<^sub>t)) \ \\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0 \ w] \ ((r \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ \[s \ s\<^sub>0, p\<^sub>0, w] \ ((\ \ p\<^sub>0) \ w) \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w])) \ \[r, r\<^sub>0, p\<^sub>1 \ w] \ (\ \ p\<^sub>1 \ w) \ (r\<^sub>1 \ \\<^sub>t) \ \\<^sub>r" proof - have "(\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] = \[s \ s\<^sub>0, p\<^sub>0, w] \ ((\ \ p\<^sub>0) \ w)" using assoc_naturality [of \ p\<^sub>0 w] by (simp add: w\\<^sub>t\\<^sub>t) thus ?thesis using comp_assoc by presburger qed also have "... = ((r \ s) \ \\<^sub>s \ (s\<^sub>0 \ \\<^sub>t)) \ \\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0 \ w] \ (r \ \[s, s\<^sub>0, p\<^sub>0 \ w]) \ (r \ \[s \ s\<^sub>0, p\<^sub>0, w]) \ (r \ (\ \ p\<^sub>0) \ w) \ (r \ r\<^sub>0s\<^sub>1.\ \ w) \ ((r \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \ \[r, r\<^sub>0, p\<^sub>1 \ w]) \ (\ \ p\<^sub>1 \ w) \ (r\<^sub>1 \ \\<^sub>t) \ \\<^sub>r" using r\<^sub>0s\<^sub>1.p\<^sub>1_simps w\\<^sub>t\\<^sub>t whisker_left comp_assoc by force also have "... = ((r \ s) \ \\<^sub>s \ (s\<^sub>0 \ \\<^sub>t)) \ \\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0 \ w] \ (r \ \[s, s\<^sub>0, p\<^sub>0 \ w]) \ (r \ \[s \ s\<^sub>0, p\<^sub>0, w]) \ (r \ (\ \ p\<^sub>0) \ w) \ (r \ r\<^sub>0s\<^sub>1.\ \ w) \ (\[r, r\<^sub>0 \ p\<^sub>1, w] \ (\[r, r\<^sub>0, p\<^sub>1] \ w) \ (\\<^sup>-\<^sup>1[r \ r\<^sub>0, p\<^sub>1, w]) \ (\ \ p\<^sub>1 \ w)) \ (r\<^sub>1 \ \\<^sub>t) \ \\<^sub>r" proof - have "(r \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \ \[r, r\<^sub>0, p\<^sub>1 \ w] = \[r, r\<^sub>0 \ p\<^sub>1, w] \ (\[r, r\<^sub>0, p\<^sub>1] \ w) \ \\<^sup>-\<^sup>1[r \ r\<^sub>0, p\<^sub>1, w]" proof - have 1: "(r \ \[r\<^sub>0, p\<^sub>1, w]) \ \[r, r\<^sub>0 \ p\<^sub>1, w] \ (\[r, r\<^sub>0, p\<^sub>1] \ w) = \[r, r\<^sub>0, p\<^sub>1 \ w] \ \[r \ r\<^sub>0, p\<^sub>1, w]" using pentagon by (simp add: \.T0.antipar(1) w) moreover have 2: "seq \[r, r\<^sub>0, p\<^sub>1 \ w] \[r \ r\<^sub>0, p\<^sub>1, w]" using \.T0.antipar(1) w by simp moreover have "inv (r \ \[r\<^sub>0, p\<^sub>1, w]) = r \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]" using \.T0.antipar(1) w by simp ultimately have "\[r, r\<^sub>0 \ p\<^sub>1, w] \ (\[r, r\<^sub>0, p\<^sub>1] \ w) = ((r \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \ \[r, r\<^sub>0, p\<^sub>1 \ w]) \ \[r \ r\<^sub>0, p\<^sub>1, w]" using \.T0.antipar(1) w comp_assoc invert_side_of_triangle(1) [of "\[r, r\<^sub>0, p\<^sub>1 \ w] \ \[r \ r\<^sub>0, p\<^sub>1, w]" "r \ \[r\<^sub>0, p\<^sub>1, w]" "\[r, r\<^sub>0 \ p\<^sub>1, w] \ (\[r, r\<^sub>0, p\<^sub>1] \ w)"] by simp hence "(r \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \ \[r, r\<^sub>0, p\<^sub>1 \ w] = (\[r, r\<^sub>0 \ p\<^sub>1, w] \ (\[r, r\<^sub>0, p\<^sub>1] \ w)) \ \\<^sup>-\<^sup>1[r \ r\<^sub>0, p\<^sub>1, w]" using \.T0.antipar(1) w invert_side_of_triangle(2) [of "\[r, r\<^sub>0 \ p\<^sub>1, w] \ (\[r, r\<^sub>0, p\<^sub>1] \ w)" "(r \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \ \[r, r\<^sub>0, p\<^sub>1 \ w]" "\[r \ r\<^sub>0, p\<^sub>1, w]"] using \trg w = src p\<^sub>0\ by simp thus ?thesis using comp_assoc by presburger qed thus ?thesis using comp_assoc by presburger qed also have "... = ((r \ s) \ \\<^sub>s \ (s\<^sub>0 \ \\<^sub>t)) \ \\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0 \ w] \ (r \ \[s, s\<^sub>0, p\<^sub>0 \ w]) \ (r \ \[s \ s\<^sub>0, p\<^sub>0, w]) \ (r \ (\ \ p\<^sub>0) \ w) \ ((r \ r\<^sub>0s\<^sub>1.\ \ w) \ \[r, r\<^sub>0 \ p\<^sub>1, w]) \ (\[r, r\<^sub>0, p\<^sub>1] \ w) \ ((\ \ p\<^sub>1) \ w) \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w] \ (r\<^sub>1 \ \\<^sub>t) \ \\<^sub>r" proof - have "\\<^sup>-\<^sup>1[r \ r\<^sub>0, p\<^sub>1, w] \ (\ \ p\<^sub>1 \ w) = ((\ \ p\<^sub>1) \ w) \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]" using assoc'_naturality [of \ p\<^sub>1 w] by (simp add: \.T0.antipar(1) w\\<^sub>t\\<^sub>t) thus ?thesis using comp_assoc by presburger qed also have "... = ((r \ s) \ \\<^sub>s \ (s\<^sub>0 \ \\<^sub>t)) \ \\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0 \ w] \ (r \ \[s, s\<^sub>0, p\<^sub>0 \ w]) \ (r \ \[s \ s\<^sub>0, p\<^sub>0, w]) \ ((r \ (\ \ p\<^sub>0) \ w) \ \[r, s\<^sub>1 \ p\<^sub>0, w]) \ ((r \ r\<^sub>0s\<^sub>1.\) \ w) \ (\[r, r\<^sub>0, p\<^sub>1] \ w) \ ((\ \ p\<^sub>1) \ w) \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w] \ (r\<^sub>1 \ \\<^sub>t) \ \\<^sub>r" proof - have "(r \ r\<^sub>0s\<^sub>1.\ \ w) \ \[r, r\<^sub>0 \ p\<^sub>1, w] = \[r, s\<^sub>1 \ p\<^sub>0, w] \ ((r \ r\<^sub>0s\<^sub>1.\) \ w)" using assoc_naturality [of r r\<^sub>0s\<^sub>1.\ w] r\<^sub>0s\<^sub>1.cospan w\\<^sub>t\\<^sub>t by auto thus ?thesis using comp_assoc by presburger qed also have "... = ((r \ s) \ \\<^sub>s \ (s\<^sub>0 \ \\<^sub>t)) \ \\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0 \ w] \ (r \ \[s, s\<^sub>0, p\<^sub>0 \ w]) \ (r \ \[s \ s\<^sub>0, p\<^sub>0, w]) \ \[r, (s \ s\<^sub>0) \ p\<^sub>0, w] \ (((r \ \ \ p\<^sub>0) \ w) \ ((r \ r\<^sub>0s\<^sub>1.\) \ w) \ (\[r, r\<^sub>0, p\<^sub>1] \ w) \ ((\ \ p\<^sub>1) \ w)) \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w] \ (r\<^sub>1 \ \\<^sub>t) \ \\<^sub>r" proof - have "(r \ (\ \ p\<^sub>0) \ w) \ \[r, s\<^sub>1 \ p\<^sub>0, w] = \[r, (s \ s\<^sub>0) \ p\<^sub>0, w] \ ((r \ \ \ p\<^sub>0) \ 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 "\ \ p\<^sub>0" w] composable by auto qed thus ?thesis using comp_assoc by presburger qed also have "... = ((r \ s) \ \\<^sub>s \ (s\<^sub>0 \ \\<^sub>t)) \ (\\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0 \ w] \ (r \ \[s, s\<^sub>0, p\<^sub>0 \ w]) \ (r \ \[s \ s\<^sub>0, p\<^sub>0, w]) \ \[r, (s \ s\<^sub>0) \ p\<^sub>0, w] \ ((r \ \\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \ w) \ (\[r, s, s\<^sub>0 \ p\<^sub>0] \ w)) \ (tab \ w) \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w] \ (r\<^sub>1 \ \\<^sub>t) \ \\<^sub>r" proof - have "((r \ \ \ p\<^sub>0) \ w) \ ((r \ r\<^sub>0s\<^sub>1.\) \ w) \ (\[r, r\<^sub>0, p\<^sub>1] \ w) \ ((\ \ p\<^sub>1) \ w) = (r \ \ \ p\<^sub>0) \ (r \ r\<^sub>0s\<^sub>1.\) \ \[r, r\<^sub>0, p\<^sub>1] \ (\ \ p\<^sub>1) \ w" using w \.T0.antipar(1) composable whisker_right by auto also have "... = (((r \ \\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \ (\[r, s, s\<^sub>0 \ p\<^sub>0] \ \\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0]) \ (r \ \[s, s\<^sub>0, p\<^sub>0])) \ (r \ \ \ p\<^sub>0)) \ (r \ r\<^sub>0s\<^sub>1.\) \ \[r, r\<^sub>0, p\<^sub>1] \ (\ \ p\<^sub>1) \ w" proof - have "((r \ \\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \ (\[r, s, s\<^sub>0 \ p\<^sub>0] \ \\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0]) \ (r \ \[s, s\<^sub>0, p\<^sub>0])) \ (r \ \ \ p\<^sub>0) = r \ \ \ p\<^sub>0" proof - have "((r \ \\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \ (\[r, s, s\<^sub>0 \ p\<^sub>0] \ \\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0]) \ (r \ \[s, s\<^sub>0, p\<^sub>0])) \ (r \ \ \ p\<^sub>0) = ((r \ \\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \ ((r \ s \ s\<^sub>0 \ p\<^sub>0) \ (r \ \[s, s\<^sub>0, p\<^sub>0]))) \ (r \ \ \ p\<^sub>0)" using comp_assoc_assoc' by (simp add: composable) also have "... = ((r \ \\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \ (r \ \[s, s\<^sub>0, p\<^sub>0])) \ (r \ \ \ p\<^sub>0)" using comp_cod_arr by (simp add: composable) also have "... = ((r \ (s \ s\<^sub>0) \ p\<^sub>0)) \ (r \ \ \ p\<^sub>0)" 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) r\<^sub>0s\<^sub>1.ide_leg0 r\<^sub>0s\<^sub>1.leg0_simps(2) r\<^sub>0s\<^sub>1.p\<^sub>0_simps hcomp_simps(1)) also have "... = r \ \ \ p\<^sub>0" using comp_cod_arr by (simp add: composable) finally show ?thesis by blast qed thus ?thesis by simp qed also have "... = (r \ \\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \ \[r, s, s\<^sub>0 \ p\<^sub>0] \ (\\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0]) \ (r \ \[s, s\<^sub>0, p\<^sub>0]) \ (r \ \ \ p\<^sub>0) \ (r \ r\<^sub>0s\<^sub>1.\) \ \[r, r\<^sub>0, p\<^sub>1] \ (\ \ p\<^sub>1) \ w" using comp_assoc by presburger also have "... = (r \ \\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \ \[r, s, s\<^sub>0 \ p\<^sub>0] \ tab \ w" using tab_def by simp also have "... = ((r \ \\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \ w) \ (\[r, s, s\<^sub>0 \ p\<^sub>0] \ w) \ (tab \ w)" using w \.T0.antipar(1) composable comp_assoc whisker_right by auto finally have "((r \ \ \ p\<^sub>0) \ w) \ ((r \ r\<^sub>0s\<^sub>1.\) \ w) \ (\[r, r\<^sub>0, p\<^sub>1] \ w) \ ((\ \ p\<^sub>1) \ w) = ((r \ \\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \ w) \ (\[r, s, s\<^sub>0 \ p\<^sub>0] \ w) \ (tab \ w)" by simp thus ?thesis using comp_assoc by presburger qed also have "... = (((r \ s) \ \\<^sub>s \ (s\<^sub>0 \ \\<^sub>t)) \ ((r \ s) \ \[s\<^sub>0, p\<^sub>0, w])) \ \[r \ s, s\<^sub>0 \ p\<^sub>0, w] \ (tab \ w) \ \" proof - have "\\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0 \ w] \ (r \ \[s, s\<^sub>0, p\<^sub>0 \ w]) \ (r \ \[s \ s\<^sub>0, p\<^sub>0, w]) \ \[r, (s \ s\<^sub>0) \ p\<^sub>0, w] \ ((r \ \\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \ w) \ (\[r, s, s\<^sub>0 \ p\<^sub>0] \ w) = ((r \ s) \ \[s\<^sub>0, p\<^sub>0, w]) \ \[r \ s, s\<^sub>0 \ p\<^sub>0, w]" proof - have "\\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0 \ w] \ (r \ \[s, s\<^sub>0, p\<^sub>0 \ w]) \ (r \ \[s \ s\<^sub>0, p\<^sub>0, w]) \ \[r, (s \ s\<^sub>0) \ p\<^sub>0, w] \ ((r \ \\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \ w) \ (\[r, s, s\<^sub>0 \ p\<^sub>0] \ w) = \\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\r\<^bold>\, \<^bold>\s\<^bold>\, \<^bold>\s\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\w\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\r\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\s\<^bold>\, \<^bold>\s\<^sub>0\<^bold>\, \<^bold>\p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\w\<^bold>\\<^bold>]) \<^bold>\ (\<^bold>\r\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\s\<^bold>\ \<^bold>\ \<^bold>\s\<^sub>0\<^bold>\, \<^bold>\p\<^sub>0\<^bold>\, \<^bold>\w\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\r\<^bold>\, (\<^bold>\s\<^bold>\ \<^bold>\ \<^bold>\s\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\p\<^sub>0\<^bold>\, \<^bold>\w\<^bold>\\<^bold>] \<^bold>\ ((\<^bold>\r\<^bold>\ \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\s\<^bold>\, \<^bold>\s\<^sub>0\<^bold>\, \<^bold>\p\<^sub>0\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\w\<^bold>\) \<^bold>\ (\<^bold>\\<^bold>[\<^bold>\r\<^bold>\, \<^bold>\s\<^bold>\, \<^bold>\s\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\p\<^sub>0\<^bold>\\<^bold>] \<^bold>\ \<^bold>\w\<^bold>\)\" using w comp_assoc \'_def \_def composable by simp also have "... = \((\<^bold>\r\<^bold>\ \<^bold>\ \<^bold>\s\<^bold>\) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\s\<^sub>0\<^bold>\, \<^bold>\p\<^sub>0\<^bold>\, \<^bold>\w\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\r\<^bold>\ \<^bold>\ \<^bold>\s\<^bold>\, \<^bold>\s\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\p\<^sub>0\<^bold>\, \<^bold>\w\<^bold>\\<^bold>]\" using w composable apply (intro E.eval_eqI) by simp_all also have "... = ((r \ s) \ \[s\<^sub>0, p\<^sub>0, w]) \ \[r \ s, s\<^sub>0 \ p\<^sub>0, 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, s\<^sub>0 \ p\<^sub>0, w] \ (tab \ w) \ \" proof - have "((r \ s) \ \\<^sub>s \ (s\<^sub>0 \ \\<^sub>t)) \ ((r \ s) \ \[s\<^sub>0, p\<^sub>0, 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, s\<^sub>0 \ p\<^sub>0, w] \ (tab \ w) \ \ = \" by simp qed thus "\w \ \. ide w \ \\ : (s\<^sub>0 \ p\<^sub>0) \ w \ u\ \ \\ : dom \ \ (r\<^sub>1 \ p\<^sub>1) \ w\ \ iso \ \ composite_cell w \ \ \ = \" using w\\<^sub>t\\<^sub>t \ \ iso_\ comp_assoc by metis qed show "\u w w' \ \' \. \ ide w; ide w'; \\ : (s\<^sub>0 \ p\<^sub>0) \ w \ u\; \\' : (s\<^sub>0 \ p\<^sub>0) \ w' \ u\; \\ : (r\<^sub>1 \ p\<^sub>1) \ w \ (r\<^sub>1 \ p\<^sub>1) \ w'\; composite_cell w \ = composite_cell w' \' \ \ \ \ \!\. \\ : w \ w'\ \ \ = (r\<^sub>1 \ p\<^sub>1) \ \ \ \ = \' \ ((s\<^sub>0 \ p\<^sub>0) \ \)" proof - fix u w w' \ \' \ assume w: "ide w" assume w': "ide w'" assume \: "\\ : (s\<^sub>0 \ p\<^sub>0) \ w \ u\" assume \': "\\' : (s\<^sub>0 \ p\<^sub>0) \ w' \ u\" assume \: "\\ : (r\<^sub>1 \ p\<^sub>1) \ w \ (r\<^sub>1 \ p\<^sub>1) \ w'\" assume eq: "composite_cell w \ = composite_cell w' \' \ \" interpret uw\w'\'\: uw\w'\'\ V H \ \ src trg \r \ s\ tab \s\<^sub>0 \ p\<^sub>0\ \r\<^sub>1 \ p\<^sub>1\ 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 \\\\<^sub>r : r\<^sub>1 \ p\<^sub>1 \ w \ r\<^sub>1 \ p\<^sub>1 \ w'\\ (obtained by composing \\\ with associativities) and ``everything to the right'' as \\\<^sub>r\ and \\\<^sub>r'\. This yields \\\\<^sub>r : p\<^sub>1 \ w \ p\<^sub>1 \ w'\\. Next, apply property \\.T2\ to obtain \\\\<^sub>s : p\<^sub>0 \ w \ p\<^sub>0 \ w'\\. For this use \\\\<^sub>s : s\<^sub>0 \ p\<^sub>0 \ w \ u\\ and \\\\<^sub>s' : s\<^sub>0 \ p\<^sub>0 \ w'\\ obtained by composing \\\ and \\'\ with associativities. We also need \\\\<^sub>s : s\<^sub>1 \ p\<^sub>0 \ w \ s\<^sub>1 \ p\<^sub>0 \ w'\\. To get this, transport \r\<^sub>0 \ \\<^sub>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 ?w\<^sub>r = "p\<^sub>1 \ w" have w\<^sub>r: "ide ?w\<^sub>r" by simp let ?w\<^sub>r' = "p\<^sub>1 \ w'" have w\<^sub>r': "ide ?w\<^sub>r'" by simp define \\<^sub>r where "\\<^sub>r = (s \ \) \ (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]) \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]" have \\<^sub>r: "\\\<^sub>r : r\<^sub>0 \ ?w\<^sub>r \ s \ u\" unfolding \\<^sub>r_def using \.T0.antipar(1) by fastforce define \\<^sub>r' where "\\<^sub>r' = (s \ \') \ (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \ \[s, s\<^sub>0, p\<^sub>0 \ w'] \ (\ \ p\<^sub>0 \ w') \ \[s\<^sub>1, p\<^sub>0, w'] \ (r\<^sub>0s\<^sub>1.\ \ w') \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w']" have \\<^sub>r': "\\\<^sub>r' : r\<^sub>0 \ ?w\<^sub>r' \ s \ u\" unfolding \\<^sub>r'_def using \.T0.antipar(1) by fastforce define \\<^sub>r where "\\<^sub>r = \[r\<^sub>1, p\<^sub>1, w'] \ \ \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]" have \\<^sub>r: "\\\<^sub>r : r\<^sub>1 \ ?w\<^sub>r \ r\<^sub>1 \ ?w\<^sub>r'\" unfolding \\<^sub>r_def using \.T0.antipar(1) by force have eq\<^sub>r: "\.composite_cell ?w\<^sub>r \\<^sub>r = \.composite_cell ?w\<^sub>r' \\<^sub>r' \ \\<^sub>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 ?w\<^sub>r \\<^sub>r = \[r, s, u] \ composite_cell w \ \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]" using \\<^sub>r_def technical uw\w'\'\.uw\.uw\ by blast also have "... = \[r, s, u] \ (((r \ s) \ \') \ \[r \ s, s\<^sub>0 \ p\<^sub>0, w'] \ (tab \ w') \ \) \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]" using eq comp_assoc by simp also have "... = (r \ \\<^sub>r') \ \[r, r\<^sub>0, ?w\<^sub>r'] \ (\ \ ?w\<^sub>r') \ \\<^sub>r" proof - have "\[r, s, u] \ (composite_cell w' \' \ \) \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w] = \[r, s, u] \ composite_cell w' \' \ \ \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]" using comp_assoc by presburger also have "... = (r \ (s \ \') \ (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \ \[s, s\<^sub>0, p\<^sub>0 \ w'] \ (\ \ p\<^sub>0 \ w') \ \[s\<^sub>1, p\<^sub>0, w'] \ (r\<^sub>0s\<^sub>1.\ \ w') \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w']) \ \[r, r\<^sub>0, p\<^sub>1 \ w'] \ (\ \ p\<^sub>1 \ w') \ \[r\<^sub>1, p\<^sub>1, w'] \ \ \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]" proof - have "\[r, s, u] \ composite_cell w' \' \ \ \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w] = \[r, s, u] \ composite_cell w' \' \ ((\\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w'] \ \[r\<^sub>1, p\<^sub>1, w']) \ \) \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]" proof - have "(\\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w'] \ \[r\<^sub>1, p\<^sub>1, 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, s\<^sub>0 \ p\<^sub>0, w'] \ (tab \ w') \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w']) \ \[r\<^sub>1, p\<^sub>1, w'] \ \ \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]" using comp_assoc by presburger also have "... = ((r \ (s \ \') \ (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \ \[s, s\<^sub>0, p\<^sub>0 \ w'] \ (\ \ p\<^sub>0 \ w') \ \[s\<^sub>1, p\<^sub>0, w'] \ (r\<^sub>0s\<^sub>1.\ \ w') \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w']) \ \[r, r\<^sub>0, p\<^sub>1 \ w'] \ (\ \ p\<^sub>1 \ w')) \ \[r\<^sub>1, p\<^sub>1, w'] \ \ \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]" using \\<^sub>r'_def technical [of w' \' u ?w\<^sub>r' \\<^sub>r'] comp_assoc by fastforce finally show ?thesis using comp_assoc by simp qed finally show ?thesis using \\<^sub>r'_def \\<^sub>r_def comp_assoc by auto qed finally show ?thesis using comp_assoc by presburger qed have 1: "\!\. \\ : ?w\<^sub>r \ ?w\<^sub>r'\ \ \\<^sub>r = \\<^sub>r' \ (r\<^sub>0 \ \) \ \\<^sub>r = r\<^sub>1 \ \" using eq\<^sub>r \.T2 [of ?w\<^sub>r ?w\<^sub>r' \\<^sub>r "s \ u" \\<^sub>r' \\<^sub>r] w\<^sub>r w\<^sub>r' \\<^sub>r \\<^sub>r' \\<^sub>r by blast obtain \\<^sub>r where \\<^sub>r: "\\\<^sub>r : ?w\<^sub>r \ ?w\<^sub>r'\ \ \\<^sub>r = \\<^sub>r' \ (r\<^sub>0 \ \\<^sub>r) \ \\<^sub>r = r\<^sub>1 \ \\<^sub>r" using 1 by blast let ?w\<^sub>s = "p\<^sub>0 \ w" have w\<^sub>s: "ide ?w\<^sub>s" by simp let ?w\<^sub>s' = "p\<^sub>0 \ w'" have w\<^sub>s': "ide ?w\<^sub>s'" by simp define \\<^sub>s where "\\<^sub>s = \ \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]" have \\<^sub>s: "\\\<^sub>s : s\<^sub>0 \ p\<^sub>0 \ w \ u\" using \\<^sub>s_def by auto define \\<^sub>s' where "\\<^sub>s' = \' \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']" have \\<^sub>s': "\\\<^sub>s' : s\<^sub>0 \ p\<^sub>0 \ w' \ u\" using \\<^sub>s'_def by auto define \\<^sub>s where "\\<^sub>s = \[s\<^sub>1, p\<^sub>0, w'] \ (r\<^sub>0s\<^sub>1.\ \ w') \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w'] \ (r\<^sub>0 \ \\<^sub>r) \ \[r\<^sub>0, p\<^sub>1, w] \ (inv r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w]" have \\<^sub>s: "\\\<^sub>s : s\<^sub>1 \ ?w\<^sub>s \ s\<^sub>1 \ ?w\<^sub>s'\" unfolding \\<^sub>s_def using \\<^sub>r r\<^sub>0s\<^sub>1.\_in_hom(2) r\<^sub>0s\<^sub>1.\_uniqueness(2) \.T0.antipar(1) apply (intro comp_in_homI) apply auto by auto have eq\<^sub>s: "\.composite_cell (p\<^sub>0 \ w) (\ \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]) = \.composite_cell (p\<^sub>0 \ w') (\' \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \ \\<^sub>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 (p\<^sub>0 \ w') (\' \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \ \\<^sub>s = (\\<^sub>r' \ (r\<^sub>0 \ \\<^sub>r)) \ \[r\<^sub>0, p\<^sub>1, w] \ (inv r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w]" using \\<^sub>s_def \\<^sub>r'_def whisker_left comp_assoc by simp also have "... = \\<^sub>r \ \[r\<^sub>0, p\<^sub>1, w] \ (inv r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w]" using \\<^sub>r by simp also have "... = ((s \ \) \ (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w])) \ \[s, s\<^sub>0, ?w\<^sub>s] \ (\ \ ?w\<^sub>s) \ \[s\<^sub>1, p\<^sub>0, w] \ ((r\<^sub>0s\<^sub>1.\ \ w) \ (\\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \ \[r\<^sub>0, p\<^sub>1, w]) \ (inv r\<^sub>0s\<^sub>1.\ \ w)) \ \\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w]" using \\<^sub>r_def comp_assoc by simp also have "... = (s \ \) \ \.composite_cell (p\<^sub>0 \ w) \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]" proof - have "(\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ ((r\<^sub>0s\<^sub>1.\ \ w) \ (\\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \ \[r\<^sub>0, p\<^sub>1, w]) \ (inv r\<^sub>0s\<^sub>1.\ \ w)) \ \\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w] = \ \ p\<^sub>0 \ w" proof - have "\\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \ \[r\<^sub>0, p\<^sub>1, w] = cod (inv r\<^sub>0s\<^sub>1.\ \ w)" using r\<^sub>0s\<^sub>1.\_uniqueness(2) \.T0.antipar(1) comp_assoc_assoc' by simp text \Here the fact that \\\ is a retraction is used.\ moreover have "(r\<^sub>0s\<^sub>1.\ \ w) \ (inv r\<^sub>0s\<^sub>1.\ \ w) = cod \\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w]" using r\<^sub>0s\<^sub>1.\_uniqueness(2) comp_arr_inv' whisker_right [of w r\<^sub>0s\<^sub>1.\ "inv r\<^sub>0s\<^sub>1.\"] by simp moreover have "\[s\<^sub>1, p\<^sub>0, w] \ \\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w] = dom (\ \ p\<^sub>0 \ w)" using r\<^sub>0s\<^sub>1.base_simps(2) hseq_char comp_assoc_assoc' by auto moreover have "hseq (inv r\<^sub>0s\<^sub>1.\) w" using r\<^sub>0s\<^sub>1.\_uniqueness(2) by (intro hseqI, auto) moreover have "hseq \ (p\<^sub>0 \ 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 (p\<^sub>0 \ w) (\ \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w])" using \\<^sub>s_def whisker_left by (metis \.ide_base \\<^sub>s arrI comp_assoc) finally show ?thesis by simp qed hence 2: "\!\. \\ : ?w\<^sub>s \ ?w\<^sub>s'\ \ \\<^sub>s = \\<^sub>s' \ (s\<^sub>0 \ \) \ \\<^sub>s = s\<^sub>1 \ \" using \.T2 [of ?w\<^sub>s ?w\<^sub>s' \\<^sub>s u \\<^sub>s' \\<^sub>s] w\<^sub>s w\<^sub>s' \\<^sub>s \\<^sub>s' \\<^sub>s by (metis \\<^sub>s'_def \\<^sub>s_def) obtain \\<^sub>s where \\<^sub>s: "\\\<^sub>s : ?w\<^sub>s \ ?w\<^sub>s'\ \ \\<^sub>s = \\<^sub>s' \ (s\<^sub>0 \ \\<^sub>s) \ \\<^sub>s = s\<^sub>1 \ \\<^sub>s" using 2 by blast have eq\<^sub>t: "(s\<^sub>1 \ \\<^sub>s) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] = (s\<^sub>1 \ ?w\<^sub>s') \ \[s\<^sub>1, p\<^sub>0, w'] \ (r\<^sub>0s\<^sub>1.\ \ w') \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w'] \ (r\<^sub>0 \ \\<^sub>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 "(s\<^sub>1 \ ?w\<^sub>s') \ \[s\<^sub>1, p\<^sub>0, w'] \ (r\<^sub>0s\<^sub>1.\ \ w') \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w'] \ (r\<^sub>0 \ \\<^sub>r) = \\<^sub>s \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]" proof - have "\\<^sub>s \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] = (\[s\<^sub>1, p\<^sub>0, w'] \ (r\<^sub>0s\<^sub>1.\ \ w') \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w']) \ (r\<^sub>0 \ \\<^sub>r) \ \[r\<^sub>0, p\<^sub>1, w] \ ((inv r\<^sub>0s\<^sub>1.\ \ w) \ (\\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w] \ \[s\<^sub>1, p\<^sub>0, w]) \ (r\<^sub>0s\<^sub>1.\ \ w)) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]" using \\<^sub>s_def comp_assoc by metis also have "... = (\[s\<^sub>1, p\<^sub>0, w'] \ (r\<^sub>0s\<^sub>1.\ \ w') \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w']) \ (r\<^sub>0 \ \\<^sub>r)" proof - have "(r\<^sub>0 \ \\<^sub>r) \ \[r\<^sub>0, p\<^sub>1, w] \ ((inv r\<^sub>0s\<^sub>1.\ \ w) \ (\\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w] \ \[s\<^sub>1, p\<^sub>0, w]) \ (r\<^sub>0s\<^sub>1.\ \ w)) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] = r\<^sub>0 \ \\<^sub>r" proof - have "(r\<^sub>0 \ \\<^sub>r) \ \[r\<^sub>0, p\<^sub>1, w] \ ((inv r\<^sub>0s\<^sub>1.\ \ w) \ (\\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w] \ \[s\<^sub>1, p\<^sub>0, w]) \ (r\<^sub>0s\<^sub>1.\ \ w)) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] = (r\<^sub>0 \ \\<^sub>r) \ \[r\<^sub>0, p\<^sub>1, w] \ ((inv r\<^sub>0s\<^sub>1.\ \ w) \ (r\<^sub>0s\<^sub>1.\ \ w)) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]" using r\<^sub>0s\<^sub>1.\_uniqueness(2) comp_assoc_assoc' comp_cod_arr by simp (* Used here that \ is a section. *) also have "... = (r\<^sub>0 \ \\<^sub>r) \ \[r\<^sub>0, p\<^sub>1, w] \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]" using r\<^sub>0s\<^sub>1.\_uniqueness(2) comp_inv_arr' \.T0.antipar(1) whisker_right [of w "inv r\<^sub>0s\<^sub>1.\" r\<^sub>0s\<^sub>1.\] comp_cod_arr by simp also have "... = r\<^sub>0 \ \\<^sub>r" proof - have "hseq r\<^sub>0 \\<^sub>r" using \\<^sub>s \\<^sub>s_def by blast thus ?thesis using comp_assoc_assoc' comp_arr_dom by (metis (no_types) \\<^sub>r \.ide_leg0 comp_assoc_assoc'(1) hcomp_simps(3) hseq_char ide_char in_homE r\<^sub>0s\<^sub>1.ide_leg1 r\<^sub>0s\<^sub>1.p\<^sub>1_simps w w\<^sub>r) qed finally show ?thesis by blast qed thus ?thesis by simp qed also have "... = \[s\<^sub>1, p\<^sub>0, w'] \ (r\<^sub>0s\<^sub>1.\ \ w') \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w'] \ (r\<^sub>0 \ \\<^sub>r)" using comp_assoc by presburger also have "... = (s\<^sub>1 \ ?w\<^sub>s') \ \[s\<^sub>1, p\<^sub>0, w'] \ (r\<^sub>0s\<^sub>1.\ \ w') \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w'] \ (r\<^sub>0 \ \\<^sub>r)" proof - have "(s\<^sub>1 \ ?w\<^sub>s') \ \[s\<^sub>1, p\<^sub>0, w'] = \[s\<^sub>1, p\<^sub>0, w']" using comp_cod_arr by simp thus ?thesis using comp_assoc by metis qed finally show ?thesis by simp qed also have "... = (s\<^sub>1 \ \\<^sub>s) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]" using \\<^sub>s by simp finally show ?thesis by simp qed have 3: "\!\. \\ : w \ w'\ \ \\<^sub>s = (p\<^sub>0 \ w') \ (p\<^sub>0 \ \) \ p\<^sub>1 \ \ = \\<^sub>r" using w w' w\<^sub>s' w\<^sub>r \\<^sub>r \\<^sub>s eq\<^sub>t r\<^sub>0s\<^sub>1.\_biuniversal_prop(2) [of ?w\<^sub>s' ?w\<^sub>r w w' \\<^sub>s "p\<^sub>0 \ w'" \\<^sub>r] by blast obtain \ where \: "\\ : w \ w'\ \ \\<^sub>s = (p\<^sub>0 \ w') \ (p\<^sub>0 \ \) \ p\<^sub>1 \ \ = \\<^sub>r" using 3 by blast show "\!\. \\ : w \ w'\ \ \ = (r\<^sub>1 \ p\<^sub>1) \ \ \ \ = \' \ ((s\<^sub>0 \ p\<^sub>0) \ \)" proof - have "\\. \\ : w \ w'\ \ \ = (r\<^sub>1 \ p\<^sub>1) \ \ \ \ = \' \ ((s\<^sub>0 \ p\<^sub>0) \ \)" proof - have "\ = \' \ ((s\<^sub>0 \ p\<^sub>0) \ \)" proof - have "\' \ ((s\<^sub>0 \ p\<^sub>0) \ \) = (\\<^sub>s' \ \[s\<^sub>0, p\<^sub>0, w']) \ ((s\<^sub>0 \ p\<^sub>0) \ \)" using \\<^sub>s'_def comp_arr_dom comp_assoc comp_assoc_assoc'(2) by auto also have "... = (\\<^sub>s' \ (s\<^sub>0 \ p\<^sub>0 \ \)) \ \[s\<^sub>0, p\<^sub>0, w]" using assoc_naturality [of s\<^sub>0 p\<^sub>0 \] comp_assoc by (metis \ \\<^sub>r \.leg0_simps(4-5) r\<^sub>0s\<^sub>1.leg0_simps(4-5) r\<^sub>0s\<^sub>1.leg1_simps(3) hseqE in_homE leg0_simps(2)) also have "... = \\<^sub>s \ \[s\<^sub>0, p\<^sub>0, w]" by (metis \ \\<^sub>s arrI comp_ide_arr w\<^sub>s') also have "... = \" using \\<^sub>s_def comp_assoc comp_arr_dom comp_assoc_assoc' by simp finally show ?thesis by simp qed moreover have "\ = (r\<^sub>1 \ p\<^sub>1) \ \" proof - have "\ = \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w'] \ \\<^sub>r \ \[r\<^sub>1, p\<^sub>1, w]" proof - have "\\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w'] \ \\<^sub>r \ \[r\<^sub>1, p\<^sub>1, w] = (\\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w'] \ \[r\<^sub>1, p\<^sub>1, w']) \ \ \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w] \ \[r\<^sub>1, p\<^sub>1, w]" using \\<^sub>r_def comp_assoc by simp also have "... = \" using comp_arr_dom comp_cod_arr by (metis \.ide_leg1 r\<^sub>0s\<^sub>1.ide_leg1 comp_assoc_assoc'(2) hseqE ideD(1) uw\w'\'\.\_simps(1) uw\w'\'\.\_simps(4-5) leg1_simps(2) w w' w\<^sub>r w\<^sub>r') finally show ?thesis by simp qed also have "... = \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w'] \ (r\<^sub>1 \ \\<^sub>r) \ \[r\<^sub>1, p\<^sub>1, w]" using \\<^sub>r by simp also have "... = \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w'] \ \[r\<^sub>1, p\<^sub>1, w'] \ ((r\<^sub>1 \ p\<^sub>1) \ \)" using assoc_naturality [of r\<^sub>1 p\<^sub>1 \] by (metis \ \\<^sub>r \.ide_leg1 r\<^sub>0s\<^sub>1.leg1_simps(5-6) hseqE ide_char in_homE leg1_simps(2)) also have "... = (\\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w'] \ \[r\<^sub>1, p\<^sub>1, w']) \ ((r\<^sub>1 \ p\<^sub>1) \ \)" using comp_assoc by presburger also have "... = (r\<^sub>1 \ p\<^sub>1) \ \" using comp_cod_arr by (metis \.ide_leg1 r\<^sub>0s\<^sub>1.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' w\<^sub>r') finally show ?thesis by simp qed ultimately show "\\. \\ : w \ w'\ \ \ = (r\<^sub>1 \ p\<^sub>1) \ \ \ \ = \' \ ((s\<^sub>0 \ p\<^sub>0) \ \)" using \ by blast qed moreover have "\\'. \\' : w \ w'\ \ \ = (r\<^sub>1 \ p\<^sub>1) \ \' \ \ = \' \ ((s\<^sub>0 \ p\<^sub>0) \ \') \ \' = \" proof - fix \' assume \': "\\' : w \ w'\ \ \ = (r\<^sub>1 \ p\<^sub>1) \ \' \ \ = \' \ ((s\<^sub>0 \ p\<^sub>0) \ \')" show "\' = \" proof - let ?P\<^sub>r = "\\. \\ : ?w\<^sub>r \ ?w\<^sub>r'\ \ \\<^sub>r = \\<^sub>r' \ (r\<^sub>0 \ \) \ \\<^sub>r = r\<^sub>1 \ \" let ?P\<^sub>s = "\\. \\ : ?w\<^sub>s \ ?w\<^sub>s'\ \ \\<^sub>s = \\<^sub>s' \ (s\<^sub>0 \ \) \ \\<^sub>s = s\<^sub>1 \ \" let ?\\<^sub>r' = "p\<^sub>1 \ \'" let ?\\<^sub>s' = "p\<^sub>0 \ \'" let ?P\<^sub>t = "\\. \\ : w \ w'\ \ \\<^sub>s = (p\<^sub>0 \ w') \ (p\<^sub>0 \ \) \ p\<^sub>1 \ \ = \\<^sub>r" have "hseq p\<^sub>0 \'" proof (intro hseqI) show "\\' : src \ \ src p\<^sub>0\" 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 "\p\<^sub>0 : src p\<^sub>0 \ src s\<^sub>0\" by simp qed hence "hseq p\<^sub>1 \'" using hseq_char by simp have "\?\\<^sub>r' : ?w\<^sub>r \ ?w\<^sub>r'\" using \' by auto moreover have "\\<^sub>r = \\<^sub>r' \ (r\<^sub>0 \ ?\\<^sub>r')" proof - text \ Note that @{term \\<^sub>r} is the composite of ``everything to the right'' of @{term "\ \ ?w\<^sub>r"}, and similarly for @{term \\<^sub>r'}. We can factor @{term \\<^sub>r} as @{term "(s \ \) \ X w"}, where @{term "X w"} is a composite of @{term \} and @{term \}. We can similarly factor @{term \\<^sub>r'} as @{term "(s \ \') \ X w'"}. Then @{term "\\<^sub>r' \ (r\<^sub>0 \ ?\\<^sub>r') = (s \ \') \ X w' \ (r\<^sub>0 \ ?\\<^sub>r')"}, which equals @{term "(s \ \') \ (s \ (s\<^sub>0 \ p\<^sub>0) \ ?\\<^sub>r') \ X w = \\<^sub>r"}. \ let ?X = "\w. (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]) \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]" have "\\<^sub>r' \ (r\<^sub>0 \ ?\\<^sub>r') = (s \ \') \ ?X w' \ (r\<^sub>0 \ ?\\<^sub>r')" using \\<^sub>r'_def comp_assoc by simp also have "... = (s \ \') \ (s \ (s\<^sub>0 \ p\<^sub>0) \ \') \ ?X w" proof - have "(s \ \') \ ((s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \ \[s, s\<^sub>0, p\<^sub>0 \ w'] \ (\ \ p\<^sub>0 \ w') \ \[s\<^sub>1, p\<^sub>0, w'] \ (r\<^sub>0s\<^sub>1.\ \ w') \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w']) \ (r\<^sub>0 \ p\<^sub>1 \ \') = (s \ \') \ (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \ \[s, s\<^sub>0, p\<^sub>0 \ w'] \ (\ \ p\<^sub>0 \ w') \ \[s\<^sub>1, p\<^sub>0, w'] \ (r\<^sub>0s\<^sub>1.\ \ w') \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w'] \ (r\<^sub>0 \ p\<^sub>1 \ \')" using comp_assoc by presburger also have "... = (s \ \') \ (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \ \[s, s\<^sub>0, p\<^sub>0 \ w'] \ (\ \ p\<^sub>0 \ w') \ \[s\<^sub>1, p\<^sub>0, w'] \ ((r\<^sub>0s\<^sub>1.\ \ w') \ ((r\<^sub>0 \ p\<^sub>1) \ \')) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]" using assoc'_naturality [of r\<^sub>0 p\<^sub>1 \'] comp_assoc by (metis \' \\p\<^sub>1 \ \' : p\<^sub>1 \ w \ p\<^sub>1 \ w'\\ \.T0.antipar(1) \.leg0_in_hom(2) r\<^sub>0s\<^sub>1.leg1_simps(4-6) r\<^sub>0s\<^sub>1.base_simps(2) hcomp_in_vhomE in_homE trg_hcomp) also have "... = (s \ \') \ (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \ \[s, s\<^sub>0, p\<^sub>0 \ w'] \ (\ \ p\<^sub>0 \ w') \ (\[s\<^sub>1, p\<^sub>0, w'] \ ((s\<^sub>1 \ p\<^sub>0) \ \')) \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]" proof - have "(r\<^sub>0s\<^sub>1.\ \ w') \ ((r\<^sub>0 \ p\<^sub>1) \ \') = r\<^sub>0s\<^sub>1.\ \ \'" using \' interchange [of r\<^sub>0s\<^sub>1.\ "r\<^sub>0 \ p\<^sub>1" w' \'] comp_arr_dom comp_cod_arr by auto also have "... = ((s\<^sub>1 \ p\<^sub>0) \ \') \ (r\<^sub>0s\<^sub>1.\ \ w)" using \' interchange \hseq p\<^sub>0 \'\ comp_arr_dom comp_cod_arr by (metis comp_arr_ide r\<^sub>0s\<^sub>1.\_simps(1,5) seqI' uw\w'\'\.uw\.w_in_hom(2) w) finally have "(r\<^sub>0s\<^sub>1.\ \ w') \ ((r\<^sub>0 \ p\<^sub>1) \ \') = ((s\<^sub>1 \ p\<^sub>0) \ \') \ (r\<^sub>0s\<^sub>1.\ \ w)" by simp thus ?thesis using comp_assoc by presburger qed also have "... = (s \ \') \ (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \ \[s, s\<^sub>0, p\<^sub>0 \ w'] \ ((\ \ p\<^sub>0 \ w') \ (s\<^sub>1 \ p\<^sub>0 \ \')) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]" using \' assoc_naturality [of s\<^sub>1 p\<^sub>0 \'] comp_assoc by (metis \.leg1_simps(2) \.leg1_simps(3,5-6) r\<^sub>0s\<^sub>1.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 \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \ (\[s, s\<^sub>0, p\<^sub>0 \ w'] \ ((s \ s\<^sub>0) \ p\<^sub>0 \ \')) \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]" proof - have "(\ \ p\<^sub>0 \ w') \ (s\<^sub>1 \ p\<^sub>0 \ \') = \ \ p\<^sub>0 \ \'" using \' interchange [of \ s\<^sub>1 "p\<^sub>0 \ w'" "p\<^sub>0 \ \'"] whisker_left \hseq p\<^sub>0 \'\comp_arr_dom comp_cod_arr by (metis \.tab_simps(1) \.tab_simps(4) hcomp_simps(4) in_homE r\<^sub>0s\<^sub>1.leg0_simps(5)) also have "... = ((s \ s\<^sub>0) \ p\<^sub>0 \ \') \ (\ \ p\<^sub>0 \ w)" using \' interchange [of "s \ s\<^sub>0" \ "p\<^sub>0 \ \'" "p\<^sub>0 \ w"] whisker_left comp_arr_dom comp_cod_arr \hseq p\<^sub>0 \'\ by (metis \.tab_simps(1) \.tab_simps(5) hcomp_simps(3) in_homE r\<^sub>0s\<^sub>1.leg0_simps(4)) finally have "(\ \ p\<^sub>0 \ w') \ (s\<^sub>1 \ p\<^sub>0 \ \') = ((s \ s\<^sub>0) \ p\<^sub>0 \ \') \ (\ \ p\<^sub>0 \ w)" by simp thus ?thesis using comp_assoc by presburger qed also have "... = (s \ \') \ (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \ ((s \ s\<^sub>0 \ p\<^sub>0 \ \') \ \[s, s\<^sub>0, p\<^sub>0 \ w]) \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]" using \' assoc_naturality [of s s\<^sub>0 "p\<^sub>0 \ \'"] \hseq p\<^sub>0 \'\ by force also have "... = (s \ \') \ ((s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \ (s \ s\<^sub>0 \ p\<^sub>0 \ \')) \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]" using comp_assoc by presburger also have "... = (s \ \') \ ((s \ (s\<^sub>0 \ p\<^sub>0) \ \') \ (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w])) \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]" proof - have "(s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \ (s \ s\<^sub>0 \ p\<^sub>0 \ \') = (s \ (s\<^sub>0 \ p\<^sub>0) \ \') \ (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w])" proof - have "(s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \ (s \ s\<^sub>0 \ p\<^sub>0 \ \') = s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w'] \ (s\<^sub>0 \ p\<^sub>0 \ \')" proof - have "seq \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w'] (s\<^sub>0 \ p\<^sub>0 \ \')" proof (* It seems to be too time-consuming for auto to solve these. *) show "\s\<^sub>0 \ p\<^sub>0 \ \' : s\<^sub>0 \ p\<^sub>0 \ w \ s\<^sub>0 \ p\<^sub>0 \ w'\" using \' by (intro hcomp_in_vhom, auto) show "\\\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w'] : s\<^sub>0 \ p\<^sub>0 \ w' \ (s\<^sub>0 \ p\<^sub>0) \ w'\" by auto qed thus ?thesis using w w' \' whisker_left by simp qed also have "... = s \ ((s\<^sub>0 \ p\<^sub>0) \ \') \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]" using \' \hseq p\<^sub>0 \'\ assoc'_naturality [of s\<^sub>0 p\<^sub>0 \'] by fastforce also have "... = (s \ (s\<^sub>0 \ p\<^sub>0) \ \') \ (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w])" proof - have "seq ((s\<^sub>0 \ p\<^sub>0) \ \') \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]" proof (* Same here. *) show "\\\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w] : s\<^sub>0 \ p\<^sub>0 \ w \ (s\<^sub>0 \ p\<^sub>0) \ w\" by auto show "\(s\<^sub>0 \ p\<^sub>0) \ \' : (s\<^sub>0 \ p\<^sub>0) \ w \ (s\<^sub>0 \ p\<^sub>0) \ 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 \ (s\<^sub>0 \ p\<^sub>0) \ \') \ ?X w" using comp_assoc by presburger finally show ?thesis by simp qed also have "... = \\<^sub>r" using \\<^sub>r_def \' uw\w'\'\.uw\.\_simps(1) whisker_left \.ide_base comp_assoc by simp finally show ?thesis by simp qed moreover have "\\<^sub>r = r\<^sub>1 \ ?\\<^sub>r'" proof - have "\\<^sub>r = \[r\<^sub>1, p\<^sub>1, w'] \ ((r\<^sub>1 \ p\<^sub>1) \ \') \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]" using \\<^sub>r_def \' by simp also have "... = \[r\<^sub>1, p\<^sub>1, w'] \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w'] \ (r\<^sub>1 \ p\<^sub>1 \ \')" using \' assoc'_naturality by (metis \.leg1_simps(5-6) r\<^sub>0s\<^sub>1.leg1_simps(5-6) hcomp_in_vhomE hseqE in_homE uw\w'\'\.\_simps(1) leg1_in_hom(2)) also have "... = (\[r\<^sub>1, p\<^sub>1, w'] \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w']) \ (r\<^sub>1 \ p\<^sub>1 \ \')" using comp_assoc by presburger also have "... = r\<^sub>1 \ p\<^sub>1 \ \'" using comp_cod_arr by (metis (no_types, lifting) \\<^sub>r \.ide_leg1 r\<^sub>0s\<^sub>1.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' w\<^sub>r') finally show ?thesis by simp qed ultimately have P\<^sub>r': "?P\<^sub>r ?\\<^sub>r'" by simp have eq\<^sub>r: "\\<^sub>r = ?\\<^sub>r'" using 1 \\<^sub>r P\<^sub>r' by blast have "\?\\<^sub>s' : ?w\<^sub>s \ ?w\<^sub>s'\" using \' by auto moreover have "\\<^sub>s = \\<^sub>s' \ (s\<^sub>0 \ ?\\<^sub>s')" using \' \hseq p\<^sub>0 \'\ \.leg0_simps(2,4-5) \.leg1_simps(3) \\<^sub>s'_def \\<^sub>s_def assoc'_naturality hseqE in_homE comp_assoc r\<^sub>0s\<^sub>1.leg0_simps(4-5) r\<^sub>0s\<^sub>1.p\<^sub>0_simps by metis moreover have "\\<^sub>s = s\<^sub>1 \ ?\\<^sub>s'" proof - have "\[s\<^sub>1, p\<^sub>0, w'] \ (r\<^sub>0s\<^sub>1.\ \ w') \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w'] \ (r\<^sub>0 \ \\<^sub>r) \ \[r\<^sub>0, p\<^sub>1, w] \ (inv r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w] = \[s\<^sub>1, p\<^sub>0, w'] \ (r\<^sub>0s\<^sub>1.\ \ w') \ (\\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w'] \ (r\<^sub>0 \ p\<^sub>1 \ \')) \ \[r\<^sub>0, p\<^sub>1, w] \ (inv r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w]" using eq\<^sub>r comp_assoc by simp also have "... = \[s\<^sub>1, p\<^sub>0, w'] \ ((r\<^sub>0s\<^sub>1.\ \ w') \ ((r\<^sub>0 \ p\<^sub>1) \ \')) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \ \[r\<^sub>0, p\<^sub>1, w] \ (inv r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w]" proof - have "\\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w'] \ (r\<^sub>0 \ p\<^sub>1 \ \') = ((r\<^sub>0 \ p\<^sub>1) \ \') \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]" using \' assoc'_naturality \hseq p\<^sub>1 \'\ by (metis \.leg0_simps(2,4-5) \.leg1_simps(3) r\<^sub>0s\<^sub>1.leg1_simps(5-6) hseqE in_homE leg1_simps(2)) thus ?thesis using comp_assoc by presburger qed also have "... = (\[s\<^sub>1, p\<^sub>0, w'] \ ((s\<^sub>1 \ p\<^sub>0) \ \')) \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \ \[r\<^sub>0, p\<^sub>1, w] \ (inv r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w]" proof - have "(r\<^sub>0s\<^sub>1.\ \ w') \ ((r\<^sub>0 \ p\<^sub>1) \ \') = r\<^sub>0s\<^sub>1.\ \ \'" using \' interchange [of r\<^sub>0s\<^sub>1.\ "r\<^sub>0 \ p\<^sub>1" w' \'] comp_arr_dom comp_cod_arr by auto also have "... = ((s\<^sub>1 \ p\<^sub>0) \ \') \ (r\<^sub>0s\<^sub>1.\ \ w)" using \' interchange \hseq p\<^sub>0 \'\ comp_arr_dom comp_cod_arr by (metis in_homE r\<^sub>0s\<^sub>1.\_simps(1,5)) finally have "(r\<^sub>0s\<^sub>1.\ \ w') \ ((r\<^sub>0 \ p\<^sub>1) \ \') = ((s\<^sub>1 \ p\<^sub>0) \ \') \ (r\<^sub>0s\<^sub>1.\ \ w)" by simp thus ?thesis using comp_assoc by presburger qed also have "... = (s\<^sub>1 \ ?\\<^sub>s') \ \[s\<^sub>1, p\<^sub>0, w] \ ((r\<^sub>0s\<^sub>1.\ \ w) \ (\\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \ \[r\<^sub>0, p\<^sub>1, w]) \ (inv r\<^sub>0s\<^sub>1.\ \ w)) \ \\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w]" proof - have "\[s\<^sub>1, p\<^sub>0, w'] \ ((s\<^sub>1 \ p\<^sub>0) \ \') = (s\<^sub>1 \ ?\\<^sub>s') \ \[s\<^sub>1, p\<^sub>0, w]" using \' assoc_naturality [of s\<^sub>1 p\<^sub>0 \'] \hseq p\<^sub>0 \'\ by auto thus ?thesis using comp_assoc by presburger qed also have "... = s\<^sub>1 \ ?\\<^sub>s'" proof - have "\\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \ \[r\<^sub>0, p\<^sub>1, w] = cod (inv r\<^sub>0s\<^sub>1.\ \ w)" using r\<^sub>0s\<^sub>1.\_uniqueness(2) \.T0.antipar(1) comp_assoc_assoc' by simp text \Here the fact that \\\ is a retraction is used.\ moreover have "(r\<^sub>0s\<^sub>1.\ \ w) \ (inv r\<^sub>0s\<^sub>1.\ \ w) = cod \\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w]" using r\<^sub>0s\<^sub>1.\_uniqueness(2) comp_arr_inv' whisker_right [of w r\<^sub>0s\<^sub>1.\ "inv r\<^sub>0s\<^sub>1.\"] by simp moreover have "cod (inv r\<^sub>0s\<^sub>1.\ \ w) \ (inv r\<^sub>0s\<^sub>1.\ \ w) = inv r\<^sub>0s\<^sub>1.\ \ w" using \\<^sub>s_def \\<^sub>s by (meson arrI comp_cod_arr seqE) ultimately show ?thesis using \' \hseq p\<^sub>0 \'\ comp_arr_dom comp_cod_arr comp_assoc_assoc' whisker_left [of s\<^sub>1 "p\<^sub>0 \ \'" "p\<^sub>0 \ w"] whisker_left [of p\<^sub>0] by (metis \.ide_leg1 assoc'_simps(1) hseqE ideD(1) in_homE r\<^sub>0s\<^sub>1.ide_leg0 r\<^sub>0s\<^sub>1.p\<^sub>0_simps w w\<^sub>s) qed finally show ?thesis using \\<^sub>s_def by simp qed ultimately have P\<^sub>s': "?P\<^sub>s ?\\<^sub>s'" by simp have eq\<^sub>s: "\\<^sub>s = ?\\<^sub>s'" using 2 \\<^sub>s P\<^sub>s' by blast have "?P\<^sub>t \'" using \' comp_cod_arr \\p\<^sub>0 \ \' : p\<^sub>0 \ w \ p\<^sub>0 \ w'\\ eq\<^sub>r eq\<^sub>s 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 \s\<^sub>0 \ p\<^sub>0\ \r\<^sub>1 \ p\<^sub>1\ using composite_is_tabulation by simp sublocale composite_tabulation_in_maps \ tabulation_in_maps V H \ \ src trg \r \ s\ tab \s\<^sub>0 \ p\<^sub>0\ \r\<^sub>1 \ p\<^sub>1\ 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 ("\_ : _ \\<^sub>B _\") notation B.in_hhom ("\_ : _ \\<^sub>B _\") notation B.in_hom ("\_ : _ \\<^sub>B _\") notation B.isomorphic (infix "\\<^sub>B" 50) notation B.iso_class ("\_\\<^sub>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 \\<^sub>B b\ \ B.is_left_adjoint f \ F = \f\\<^sub>B}" 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 \\<^sub>B b\" and "f \ F \ B.is_left_adjoint f" and "f \ F \ F = \f\\<^sub>B" 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 \\<^sub>B b\ \ B.is_left_adjoint f' \ F = \f'\\<^sub>B" using assms by auto show "f \ F \ \f : a \\<^sub>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 = \f\\<^sub>B" 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 \\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>B h" using h Comp_def by auto have "Comp G F = \g \ f\\<^sub>B \ 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 \ f\\<^sub>B" proof show "\g \ f\\<^sub>B \ Comp G F" unfolding Comp_def B.iso_class_def using fg by auto show "Comp G F \ \g \ f\\<^sub>B" proof fix h' assume h': "h' \ Comp G F" obtain f' g' where f'g': "f' \ F \ g' \ G \ g' \ f' \\<^sub>B h'" using h' Comp_def by auto have 1: "f' \\<^sub>B f \ g' \\<^sub>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' \\<^sub>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' \\<^sub>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 \\<^sub>B h'" shows "Comp G F = Comp G' F'" proof - obtain f g where fg: "f \ F \ g \ G \ g \ f \\<^sub>B h" using assms comp_def by auto obtain f' g' where f'g': "f' \ F' \ g' \ G' \ g' \ f' \\<^sub>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 = \h\\<^sub>B" proof show "Comp G F \ \h\\<^sub>B" proof fix h' assume h': "h' \ Comp G F" obtain f g where fg: "f \ F \ g \ G \ g \ f \\<^sub>B h" using assms by auto obtain f' g' where f'g': "f' \ F \ g' \ G \ g' \ f' \\<^sub>B h'" using h' by auto have "f \\<^sub>B f' \ g \\<^sub>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 \\<^sub>B g' \ f'" using fg f'g' B.hcomp_ide_isomorphic B.hcomp_isomorphic_ide B.isomorphic_transitive by metis thus "h' \ \h\\<^sub>B" using fg f'g' B.isomorphic_symmetric B.isomorphic_transitive B.iso_class_def [of h] by blast qed show "\h\\<^sub>B \ Comp G F" proof (unfold B.iso_class_def Comp_def) obtain f g where 1: "f \ F \ g \ G \ g \ f \\<^sub>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 \\<^sub>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 \ \a\\<^sub>B \ 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 \\<^sub>B b\ \ B.is_left_adjoint f \ F = \f\\<^sub>B" using F by blast obtain g where g: "\g : b \\<^sub>B c\ \ B.is_left_adjoint g \ G = \g\\<^sub>B" using G by blast have "{h. B.is_iso_class F \ B.is_iso_class G \ (\f g. f \ F \ \f : a \\<^sub>B b\ \ g \ G \ \g : b \\<^sub>B c\ \ g \ f \\<^sub>B h)} = \g \ f\\<^sub>B" proof show "{h. B.is_iso_class F \ B.is_iso_class G \ (\f g. f \ F \ \f : a \\<^sub>B b\ \ g \ G \ \g : b \\<^sub>B c\ \ g \ f \\<^sub>B h)} \ \g \ f\\<^sub>B" proof fix h assume "h \ {h. B.is_iso_class F \ B.is_iso_class G \ (\f g. f \ F \ \f : a \\<^sub>B b\ \ g \ G \ \g : b \\<^sub>B c\ \ g \ f \\<^sub>B h)}" hence h: "B.is_iso_class F \ B.is_iso_class G \ (\f g. f \ F \ \f : a \\<^sub>B b\ \ g \ G \ \g : b \\<^sub>B c\ \ g \ f \\<^sub>B h)" by simp show "h \ \g \ f\\<^sub>B" proof - obtain f' g' where f'g': "g' \ G \ f' \ F \ g' \ f' \\<^sub>B h" using h by auto obtain \ where \: "\\ : f \\<^sub>B f'\ \ B.iso \" using f f'g' F B.iso_class_def by auto obtain \ where \: "\\ : g \\<^sub>B g'\ \ B.iso \" using g f'g' G B.iso_class_def by auto have 1: "\\ \ \ : g \ f \\<^sub>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 \ f\\<^sub>B \ {h. B.is_iso_class F \ B.is_iso_class G \ (\f g. f \ F \ \f : a \\<^sub>B b\ \ g \ G \ \g : b \\<^sub>B c\ \ g \ f \\<^sub>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 \\<^sub>B b\ \ g \ G \ \g : b \\<^sub>B c\ \ g \ f \\<^sub>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 \ f\\<^sub>B = 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 \a\\<^sub>B = F" proof - fix a b F assume a: "a \ Collect B.obj" assume F: "F \ Hom a b" obtain f where f: "\f : a \\<^sub>B b\ \ B.is_left_adjoint f \ F = \f\\<^sub>B" using F by auto have *: "\f'. f' \ F \ \f' : a \\<^sub>B b\ \ B.ide f' \ f \\<^sub>B f'" using f B.iso_class_def by force show "Comp F \a\\<^sub>B = F" proof show "Comp F \a\\<^sub>B \ F" proof fix h assume "h \ Comp F \a\\<^sub>B" hence h: "\f' a'. f' \ F \ a' \ \a\\<^sub>B \ f' \ a' \\<^sub>B h" unfolding Comp_def by auto obtain f' a' where f'a': "f' \ F \ a' \ \a\\<^sub>B \ f' \ a' \\<^sub>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 \\<^sub>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 \\<^sub>B f' \ a'" using f'a' * [of f'] B.iso_class_def B.hcomp_ide_isomorphic by auto also have "f' \ a' \\<^sub>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 \a\\<^sub>B" 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 \\<^sub>B h" proof - have "f \ a \\<^sub>B f" using f B.iso_runit [of f] B.isomorphic_def B.left_adjoint_is_ide by blast also have "f \\<^sub>B h" using h * by simp finally show ?thesis by blast qed ultimately show "h \ Comp F \a\\<^sub>B" 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 \b\\<^sub>B 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 \\<^sub>B b\ \ B.is_left_adjoint f \ F = \f\\<^sub>B" using F by auto have *: "\f'. f' \ F \ \f' : a \\<^sub>B b\ \ B.ide f' \ f \\<^sub>B f'" using f B.iso_class_def by force show "Comp \b\\<^sub>B F = F" proof show "Comp \b\\<^sub>B F \ F" proof fix h assume "h \ Comp \b\\<^sub>B F" hence h: "\b' f'. b' \ \b\\<^sub>B \ f' \ F \ b' \ f' \\<^sub>B h" unfolding Comp_def by auto obtain b' f' where b'f': "b' \ \b\\<^sub>B \ f' \ F \ b' \ f' \\<^sub>B h" using h by auto have "f \\<^sub>B h" proof - have "f \\<^sub>B b \ f" using f B.iso_lunit' [of f] B.isomorphic_def B.left_adjoint_is_ide by blast also have "... \\<^sub>B b \ f'" using f b'f' B.iso_class_def B.hcomp_ide_isomorphic by (elim conjE B.in_hhomE, auto) also have "... \\<^sub>B b' \ f'" using b'f' * [of f'] B.iso_class_def B.hcomp_isomorphic_ide by auto also have "... \\<^sub>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 \b\\<^sub>B 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 \\<^sub>B h" proof - have "b \ f \\<^sub>B f" using f B.iso_lunit [of f] B.isomorphic_def B.left_adjoint_is_ide by blast also have "f \\<^sub>B h" using h * by simp finally show ?thesis by blast qed ultimately show "h \ Comp \b\\<^sub>B 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 \\<^sub>B gf \ h \ gf \\<^sub>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 \\<^sub>B x" proof - have "h \ g \ f \\<^sub>B h \ gf" using 1 hgf B.hcomp_ide_isomorphic by (metis (full_types) B.isomorphic_implies_hpar(1) B.isomorphic_reflexive B.isomorphic_symmetric B.seq_if_composable) also have "h \ gf \\<^sub>B x" using 1 by simp finally show ?thesis by blast qed moreover have "(h \ g) \ f \\<^sub>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 \\<^sub>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" using 1 F G H hgf B.is_iso_class_def is_iso_class_Comp [of H G] B.isomorphic_reflexive [of "h \ g"] apply (intro in_CompI) apply auto[7] apply blast apply simp by (meson B.isomorphic_symmetric B.isomorphic_transitive) 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 \\<^sub>B hg \ hg \ f \\<^sub>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 \\<^sub>B x" proof - have "(h \ g) \ f \\<^sub>B hg \ f" using 1 F G H hgf by (simp add: B.hcomp_isomorphic_ide) also have "hg \ f \\<^sub>B x" using 1 by simp finally show ?thesis by blast qed moreover have "(h \ g) \ f \\<^sub>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 ("\_ : _ \\<^sub>B _\") lemma Map_memb_in_hhom: assumes "\F : A \ B\" and "f \ Map F" shows "\f : Dom A \\<^sub>B Dom B\" proof - have "\f : Dom F \\<^sub>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 \\<^sub>B b\" shows "\MkArr a b \f\\<^sub>B : 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 \\<^sub>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 \\<^sub>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) \\<^sub>B g \ f" proof - have "g \ f \ Map (G \ F)" proof - have 1: "f \ Map F \ g \ Map G \ g \ f \\<^sub>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 iso_inv_iso 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 F\\<^sub>B" 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 \\<^sub>B trg f" proof - have "f \ g \ Map (F \ G)" proof - have 1: "f \ Map F \ g \ Map G \ f \ g \\<^sub>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 iso_inv_iso 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 F\\<^sub>B" 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 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 \\<^sub>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 \\<^sub>B Dom F\" using \.ide_right \.antipar f_in_hhom by auto have fg: "B.equivalence_pair f g" using B.equivalence_pair_def \.equivalence_in_bicategory_axioms by auto have g: "\g : Cod F \\<^sub>B Dom F\ \ B.is_left_adjoint g \ \g\\<^sub>B = \g\\<^sub>B" 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) \f\\<^sub>B" 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) \g\\<^sub>B" 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 \g\\<^sub>B \f\\<^sub>B" using 1 comp_char [of ?G F] `Map F = B.iso_class f` by simp also have "... = \g \ f\\<^sub>B" proof - have "g \ f \ Comp \g\\<^sub>B \f\\<^sub>B" 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 f\\<^sub>B" 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 F\\<^sub>B" 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 \f\\<^sub>B \g\\<^sub>B" using 1 comp_char [of F ?G] `Map F = \f\\<^sub>B` by simp also have "... = \f \ g\\<^sub>B" proof - have "f \ g \ Comp \f\\<^sub>B \g\\<^sub>B" 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 = \f\\<^sub>B` by auto qed also have "... = \trg f\\<^sub>B" proof - have "trg f \ \f \ g\\<^sub>B" 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 F\\<^sub>B" 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) \f\\<^sub>B" lemma ide_CLS_obj: assumes "B.obj a" shows "ide \\a\\" - by (simp add: assms) + 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 \\<^sub>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 \f\\<^sub>B \g\\<^sub>B)" proof - have "\\f \ g\\ = MkArr (src g) (trg f) \f \ g\\<^sub>B" using assms B.left_adjoint_is_ide by simp moreover have "\f \ g\\<^sub>B = Comp \f\\<^sub>B \g\\<^sub>B" proof show "\f \ g\\<^sub>B \ Comp \f\\<^sub>B \g\\<^sub>B" 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 \f\\<^sub>B \g\\<^sub>B \ \f \ g\\<^sub>B" by (metis Comp_eq_iso_class_memb \\f \ g\\<^sub>B \ Comp \f\\<^sub>B \g\\<^sub>B\ 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 \\<^sub>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 \f\\<^sub>B \g\\<^sub>B)" 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) \\<^sub>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 \\<^sub>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) \\<^sub>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 F\\<^sub>B" 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 \\<^sub>B src (REP A)" using assms by (metis (no_types, lifting) ideD(1) ide_char 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 \\<^sub>B trg (REP A)" using assms ide_char 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\\ \\<^sub>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 \\<^sub>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 F\\<^sub>B \REP G\\<^sub>B" using assms seq_char null_char by (metis (no_types, lifting) CLS_REP Map.simps(1) Map_comp) have "Comp \REP F\\<^sub>B \REP G\\<^sub>B = \REP F \ REP G\\<^sub>B" proof - have "REP F \ REP G \ Comp \REP F\\<^sub>B \REP G\\<^sub>B" 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 \\<^sub>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 F\\<^sub>B \ Hom (Cod G) (Cod F)" using assms 1 2 arr_char [of F] by auto moreover have "\REP G\\<^sub>B \ 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 F\\<^sub>B \REP G\\<^sub>B" 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 \PRJ\<^sub>0\ and \PRJ\<^sub>1\ 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 \PRJ\<^sub>0 = CLS (tab\<^sub>0 (k\<^sup>* \ h))\ and \PRJ\<^sub>1 = CLS (tab\<^sub>1 (k\<^sup>* \ h))\. That these constitute a pullback in \Maps(B)\ follows from the fact that \tab\<^sub>0 (k\<^sup>* \ h)\ and \tab\<^sub>1 (k\<^sup>* \ 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\<^sup>* \ 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 PRJ\<^sub>0 where "PRJ\<^sub>0 \ \K H. if cospan K H then \\B.tab\<^sub>0 ((REP K)\<^sup>* \ (REP H))\\ else null" definition PRJ\<^sub>1 where "PRJ\<^sub>1 \ \K H. if cospan K H then \\B.tab\<^sub>1 ((REP K)\<^sup>* \ (REP H))\\ else null" interpretation elementary_category_with_pullbacks comp PRJ\<^sub>0 PRJ\<^sub>1 proof show "\H K. \ cospan K H \ PRJ\<^sub>0 K H = null" unfolding PRJ\<^sub>0_def by auto show "\H K. \ cospan K H \ PRJ\<^sub>1 K H = null" unfolding PRJ\<^sub>1_def by auto show "\H K. cospan K H \ commutative_square K H (PRJ\<^sub>1 K H) (PRJ\<^sub>0 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.tab\<^sub>0 (k\<^sup>* \ h)" let ?g = "B.tab\<^sub>1 (k\<^sup>* \ 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.p\<^sub>0_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.p\<^sub>1_simps k_def span) show "commutative_square K H (PRJ\<^sub>1 K H) (PRJ\<^sub>0 K H)" proof show "cospan K H" using HK by simp show "dom K = cod (PRJ\<^sub>1 K H)" using seq' PRJ\<^sub>1_def HK h_def k_def by auto show "span (PRJ\<^sub>1 K H) (PRJ\<^sub>0 K H)" unfolding PRJ\<^sub>0_def PRJ\<^sub>1_def using HK span h_def k_def by simp show "K \ PRJ\<^sub>1 K H = H \ PRJ\<^sub>0 K H" proof - have iso: "h \ ?f \\<^sub>B k \ ?g" using hk.\_uniqueness B.isomorphic_symmetric B.isomorphic_def by blast have *: "Comp (Map H) \?f\\<^sub>B = Comp (Map K) \?g\\<^sub>B" proof (intro Comp_eqI) show "h \ ?f \ Comp (Map H) \B.tab\<^sub>0 (k\<^sup>* \ h)\\<^sub>B" proof (unfold Comp_def) have "B.is_iso_class \?f\\<^sub>B" 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.tab\<^sub>0 (k\<^sup>* \ h)\\<^sub>B" by (simp add: B.ide_in_iso_class(1)) moreover have "\?f : src (B.tab\<^sub>0 (k\<^sup>* \ h)) \\<^sub>B Dom H\" using seq seq_char by simp moreover have "h \ Map H" by fact moreover have "\h : Dom H \\<^sub>B Cod H\" by (simp add: HK h_def) moreover have "h \ ?f \\<^sub>B h \ ?f" using B.isomorphic_reflexive by auto ultimately show "h \ B.tab\<^sub>0 (k\<^sup>* \ h) \ {h'. B.is_iso_class \B.tab\<^sub>0 (k\<^sup>* \ h)\\<^sub>B \ B.is_iso_class (Map H) \ (\f g. f \ \B.tab\<^sub>0 (k\<^sup>* \ h)\\<^sub>B \ g \ Map H \ g \ f \\<^sub>B h')}" by auto qed show "k \ ?g \ Comp (Map K) \B.tab\<^sub>1 (k\<^sup>* \ h)\\<^sub>B" proof (unfold Comp_def) have "B.is_iso_class \?g\\<^sub>B" 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.tab\<^sub>1 (k\<^sup>* \ h)\\<^sub>B" by (simp add: B.ide_in_iso_class(1)) moreover have "\?g : src (B.tab\<^sub>1 (k\<^sup>* \ h)) \\<^sub>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 \\<^sub>B Cod K\" by (simp add: HK k_def) moreover have "k \ ?g \\<^sub>B k \ ?g" using B.isomorphic_reflexive iso B.isomorphic_implies_hpar(2) by auto ultimately show "k \ B.tab\<^sub>1 (k\<^sup>* \ h) \ {h'. B.is_iso_class \B.tab\<^sub>1 (k\<^sup>* \ h)\\<^sub>B \ B.is_iso_class (Map K) \ (\f g. f \ \B.tab\<^sub>1 (k\<^sup>* \ h)\\<^sub>B \ g \ Map K \ g \ f \\<^sub>B h')}" by auto qed show "h \ ?f \\<^sub>B k \ ?g" using iso by simp qed have "K \ PRJ\<^sub>1 K H = K \ \\?g\\" unfolding PRJ\<^sub>1_def using HK h_def k_def by simp also have "... = MkArr (src ?g) (Cod K) (Comp (Map K) \?g\\<^sub>B)" using seq' comp_char [of K "\\?g\\"] by simp also have "... = MkArr (src ?f) (Cod H) (Comp (Map H) \?f\\<^sub>B)" using * HK cod_char by auto also have "... = comp H \\?f\\" using seq comp_char [of H "\\?f\\"] by simp also have "... = comp H (PRJ\<^sub>0 K H)" unfolding PRJ\<^sub>0_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 (PRJ\<^sub>1 K H) E = V \ comp (PRJ\<^sub>0 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.tab\<^sub>0 (k\<^sup>* \ h)" let ?g = "B.tab\<^sub>1 (k\<^sup>* \ 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 \\<^sub>B k \ ?v" proof - have "h \ ?u \\<^sub>B REP (H \ U)" proof - have "h \ ?u \\<^sub>B REP H \ ?u" proof - have "h \\<^sub>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 "... \\<^sub>B REP (H \ U)" using seq_HU isomorphic_hcomp_REP isomorphic_def by blast finally show ?thesis by blast qed also have "... \\<^sub>B REP (K \ V)" using seq_HU cs B.isomorphic_reflexive by auto also have "... \\<^sub>B (k \ ?v)" proof - have "... \\<^sub>B REP K \ ?v" using seq_KV isomorphic_hcomp_REP B.isomorphic_def B.isomorphic_symmetric by blast also have "... \\<^sub>B k \ ?v" proof - have "k \\<^sub>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 \\<^sub>B ?u \ ?v \\<^sub>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 \\<^sub>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. PRJ\<^sub>1 K H \ E = V \ PRJ\<^sub>0 K H \ E = U" proof - have "PRJ\<^sub>1 K H \ ?W = V \ PRJ\<^sub>0 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 PRJ\<^sub>0_def PRJ\<^sub>1_def by auto qed moreover have "\W'. PRJ\<^sub>1 K H \ W' = V \ PRJ\<^sub>0 K H \ W' = U \ W' = ?W" proof - fix W' assume "PRJ\<^sub>1 K H \ W' = V \ PRJ\<^sub>0 K H \ W' = U" hence W': "\W' : dom U \ dom \\?f\\\ \ \\?f\\ \ W' = U \ \\?g\\ \ W' = V" using PRJ\<^sub>0_def PRJ\<^sub>1_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' \\<^sub>B ?u" proof - have "?f \ ?w' \\<^sub>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' \\<^sub>B ?u" using W' isomorphic_hcomp_REP cs by blast finally show ?thesis by blast qed have gw'_iso_v: "?g \ ?w' \\<^sub>B ?v" proof - have "?g \ ?w' \\<^sub>B REP \\?g\\ \ ?w'" proof - have "?g \\<^sub>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 "... \\<^sub>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' \\<^sub>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 PRJ\<^sub>0 PRJ\<^sub>1" .. lemma is_category_with_pullbacks: shows "category_with_pullbacks comp" .. sublocale elementary_category_with_pullbacks comp PRJ\<^sub>0 PRJ\<^sub>1 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 r\<^sub>0s\<^sub>1: cospan_of_maps_in_bicategory_of_spans V H \ \ src trg s\<^sub>1 r\<^sub>0 using \.leg0_is_map \.leg1_is_map composable by unfold_locales auto lemma prj_char: shows "Maps.PRJ\<^sub>0 \\r\<^sub>0\\ \\s\<^sub>1\\ = \\prj\<^sub>0 s\<^sub>1 r\<^sub>0\\" and "Maps.PRJ\<^sub>1 \\r\<^sub>0\\ \\s\<^sub>1\\ = \\prj\<^sub>1 s\<^sub>1 r\<^sub>0\\" proof - have "Maps.arr (Maps.MkArr (src s\<^sub>0) (trg s) \s\<^sub>1\)" using \.leg1_in_hom Maps.CLS_in_hom \.leg1_is_map Maps.arr_char by auto moreover have "Maps.arr (Maps.MkArr (src r\<^sub>0) (trg s) \r\<^sub>0\)" using Maps.CLS_in_hom composable r\<^sub>0s\<^sub>1.k_is_map by fastforce moreover have "Maps.cod (Maps.MkArr (src r\<^sub>0) (trg s) \r\<^sub>0\) = Maps.cod (Maps.MkArr (src s\<^sub>0) (trg s) \s\<^sub>1\)" unfolding Maps.arr_char using \.leg1_in_hom \.leg0_in_hom by (simp add: Maps.cod_char calculation(1) calculation(2)) ultimately have "Maps.PRJ\<^sub>0 \\r\<^sub>0\\ \\s\<^sub>1\\ = \\tab\<^sub>0 ((Maps.REP (Maps.MkArr (src r\<^sub>0) (trg s) \r\<^sub>0\))\<^sup>* \ Maps.REP (Maps.MkArr (src s\<^sub>0) (trg s) \s\<^sub>1\))\\ \ Maps.PRJ\<^sub>1 (Maps.CLS r\<^sub>0) (Maps.CLS s\<^sub>1) = \\tab\<^sub>1 ((Maps.REP (Maps.MkArr (src r\<^sub>0) (trg s) \r\<^sub>0\))\<^sup>* \ Maps.REP (Maps.MkArr (src s\<^sub>0) (trg s) \s\<^sub>1\))\\" unfolding Maps.PRJ\<^sub>0_def Maps.PRJ\<^sub>1_def using Maps.CLS_in_hom \.leg1_is_map \.leg0_is_map composable by simp moreover have "r\<^sub>0\<^sup>* \ s\<^sub>1 \ (Maps.REP (Maps.MkArr (src r\<^sub>0) (trg s) \r\<^sub>0\))\<^sup>* \ Maps.REP (Maps.MkArr (src s\<^sub>0) (trg s) \s\<^sub>1\)" proof - have "r\<^sub>0 \ Maps.REP (Maps.MkArr (src r\<^sub>0) (trg s) \r\<^sub>0\)" using Maps.REP_CLS composable isomorphic_symmetric r\<^sub>0s\<^sub>1.k_is_map by fastforce hence 3: "isomorphic r\<^sub>0\<^sup>* (Maps.REP (Maps.MkArr (src r\<^sub>0) (trg s) \r\<^sub>0\))\<^sup>*" using \.leg0_is_map by (simp add: isomorphic_to_left_adjoint_implies_isomorphic_right_adjoint) moreover have 4: "s\<^sub>1 \ Maps.REP (Maps.MkArr (src s\<^sub>0) (trg s) \s\<^sub>1\)" using Maps.REP_CLS isomorphic_symmetric r\<^sub>0s\<^sub>1.h_is_map by fastforce ultimately show ?thesis proof - have 1: "src r\<^sub>0\<^sup>* = trg s\<^sub>1" using \.T0.antipar(2) r\<^sub>0s\<^sub>1.cospan by argo have 2: "ide s\<^sub>1" by simp have "src (Maps.REP (Maps.MkArr (src r\<^sub>0) (trg s) \r\<^sub>0\))\<^sup>* = trg s\<^sub>1" by (metis 3 \.T0.antipar(2) isomorphic_implies_hpar(3) r\<^sub>0s\<^sub>1.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.PRJ\<^sub>0 \\r\<^sub>0\\ \\s\<^sub>1\\ = \\prj\<^sub>0 s\<^sub>1 r\<^sub>0\\ \ Maps.PRJ\<^sub>1 \\r\<^sub>0\\ \\s\<^sub>1\\ = \\prj\<^sub>1 s\<^sub>1 r\<^sub>0\\" using r\<^sub>0s\<^sub>1.isomorphic_implies_same_tab by simp show "Maps.PRJ\<^sub>0 \\r\<^sub>0\\ \\s\<^sub>1\\ = \\prj\<^sub>0 s\<^sub>1 r\<^sub>0\\" using 1 by simp show "Maps.PRJ\<^sub>1 \\r\<^sub>0\\ \\s\<^sub>1\\ = \\prj\<^sub>1 s\<^sub>1 r\<^sub>0\\" 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.PRJ\<^sub>0 Maps.PRJ\<^sub>1 .. notation isomorphic (infix "\" 50) text \ A 1-cell \r\ in a bicategory of spans is a map if and only if the ``input leg'' \tab\<^sub>0 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_tab\<^sub>0_is_equivalence: shows "is_left_adjoint r \ equivalence_map (tab\<^sub>0 r)" proof assume 1: "equivalence_map (tab\<^sub>0 r)" have 2: "equivalence_pair (tab\<^sub>0 r) (tab\<^sub>0 r)\<^sup>*" proof - obtain g' \' \' where \'\': "equivalence_in_bicategory V H \ \ src trg (tab\<^sub>0 r) g' \' \'" using 1 equivalence_map_def by auto have "adjoint_pair (tab\<^sub>0 r) g'" using \'\' equivalence_pair_def equivalence_pair_is_adjoint_pair by blast moreover have "adjoint_pair (tab\<^sub>0 r) (tab\<^sub>0 r)\<^sup>*" using T0.adjunction_in_bicategory_axioms adjoint_pair_def by auto ultimately have "g' \ (tab\<^sub>0 r)\<^sup>*" using left_adjoint_determines_right_up_to_iso by simp thus ?thesis using \'\' equivalence_pair_def equivalence_pair_isomorphic_right by blast qed obtain \' \' where \'\': "equivalence_in_bicategory V H \ \ src trg (tab\<^sub>0 r) (tab\<^sub>0 r)\<^sup>* \' \'" using 2 equivalence_pair_def by auto interpret \'\': equivalence_in_bicategory V H \ \ src trg \tab\<^sub>0 r\ \(tab\<^sub>0 r)\<^sup>*\ \' \' using \'\' by auto have "is_left_adjoint (tab\<^sub>0 r)\<^sup>*" using 2 equivalence_pair_is_adjoint_pair equivalence_pair_symmetric by blast hence "is_left_adjoint (tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>*)" 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 (tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>*)" using 1 yields_isomorphic_representation left_adjoint_preserved_by_iso' isomorphic_symmetric isomorphic_def by meson hence "is_left_adjoint (tab\<^sub>0 r)\<^sup>*" using is_ide BS4 [of "tab\<^sub>1 r" "(tab\<^sub>0 r)\<^sup>*"] by auto hence "is_left_adjoint ((tab\<^sub>0 r)\<^sup>* \ tab\<^sub>0 r) \ is_left_adjoint (tab\<^sub>0 r \ (tab\<^sub>0 r)\<^sup>*)" using left_adjoints_compose T0.antipar by simp hence 3: "iso \ \ iso \" using BS3 [of "src (tab\<^sub>0 r)" "(tab\<^sub>0 r)\<^sup>* \ tab\<^sub>0 r" \ \] BS3 [of "tab\<^sub>0 r \ (tab\<^sub>0 r)\<^sup>*" "trg (tab\<^sub>0 r)" \ \] T0.unit_in_hom T0.counit_in_hom obj_is_self_adjoint by auto hence "equivalence_in_bicategory V H \ \ src trg (tab\<^sub>0 r) (tab\<^sub>0 r)\<^sup>* \ \" apply unfold_locales by auto thus "equivalence_map (tab\<^sub>0 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 "tab\<^sub>0 r \ tab\<^sub>1 r" proof - have "tab\<^sub>0 r \ r \ tab\<^sub>0 r" proof - have "trg (tab\<^sub>0 r) = r" using assms by auto moreover have "\\\<^sup>-\<^sup>1[tab\<^sub>0 r] : tab\<^sub>0 r \ trg (tab\<^sub>0 r) \ tab\<^sub>0 r\ \ iso \\<^sup>-\<^sup>1[tab\<^sub>0 r]" using assms by simp ultimately show ?thesis unfolding isomorphic_def by metis qed also have "... \ tab\<^sub>1 r" proof - have "\tab : tab\<^sub>1 r \ r \ tab\<^sub>0 r\" using tab_in_hom by simp moreover have "is_left_adjoint (r \ tab\<^sub>0 r)" using assms left_adjoints_compose obj_is_self_adjoint by simp ultimately show ?thesis using BS3 [of "tab\<^sub>1 r" "r \ tab\<^sub>0 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 = \\tab\<^sub>0 r\\, Leg1 = \\tab\<^sub>1 r\\\" using assms Maps.CLS_in_hom [of "tab\<^sub>0 r"] Maps.CLS_in_hom [of "tab\<^sub>1 r"] tab\<^sub>0_in_hom tab\<^sub>1_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 \\ = (\ \ r\<^sub>0) \ \\ ``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 \ r\<^sub>0 r\<^sub>1 + \: tabulation_in_maps V H \ \ src trg s \ s\<^sub>0 s\<^sub>1 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 r\<^sub>0 :: 'a and r\<^sub>1 :: 'a and s :: 'a and \ :: 'a and s\<^sub>0 :: 'a and s\<^sub>1 :: 'a and \ :: 'a + assumes in_hom: "\\ : r \ s\" begin abbreviation (input) \ where "\ \ (\ \ r\<^sub>0) \ \" lemma \_in_hom [intro]: shows "\\ : src \ \ trg \\" and "\\ : r\<^sub>1 \ s \ r\<^sub>0\" proof - show "\\ : r\<^sub>1 \ s \ r\<^sub>0\" 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 \ = r\<^sub>1" and "cod \ = s \ r\<^sub>0" using \_in_hom by auto abbreviation is_induced_map where "is_induced_map w \ \.is_induced_by_cell w r\<^sub>0 \" 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 \ (\\ \. \\ : s\<^sub>0 \ w \ r\<^sub>0\ \ \\ : r\<^sub>1 \ s\<^sub>1 \ w\ \ iso \ \ \ = (s \ \) \ \[s, s\<^sub>0, 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 \ = r\<^sub>1" by auto interpret w: arrow_of_spans_of_maps_to_tabulation V H \ \ src trg r\<^sub>0 \dom \\ s \ s\<^sub>0 s\<^sub>1 w proof - have "arrow_of_spans_of_maps_to_tabulation V H \ \ src trg r\<^sub>0 r\<^sub>1 s \ s\<^sub>0 s\<^sub>1 w" using w apply unfold_locales by auto thus "arrow_of_spans_of_maps_to_tabulation V H \ \ src trg r\<^sub>0 (dom \) s \ s\<^sub>0 s\<^sub>1 w" using 1 by simp qed show "arrow_of_spans_of_maps V H \ \ src trg r\<^sub>0 (dom \) s\<^sub>0 s\<^sub>1 w" using w.arrow_of_spans_of_maps_axioms by auto show "\.composite_cell w w.the_\ \ w.the_\ = \" proof - obtain \ \ where \\: "\\ : s\<^sub>0 \ w \ r\<^sub>0\ \ \\ : r\<^sub>1 \ s\<^sub>1 \ w\ \ iso \ \ \ = (s \ \) \ \[s, s\<^sub>0, w] \ (\ \ w) \ \" using w w.the_\_props(1) by auto have "(s \ \) \ \[s, s\<^sub>0, 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 r\<^sub>0 r\<^sub>1 s\<^sub>0 s\<^sub>1 w using w in_hom by auto interpret w: arrow_of_spans_of_maps_to_tabulation V H \ \ src trg r\<^sub>0 r\<^sub>1 s \ s\<^sub>0 s\<^sub>1 w .. have "dom \ = r\<^sub>1" 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 \ \\ : s\<^sub>0 \ w \ r\<^sub>0\ \ \\ : r\<^sub>1 \ s\<^sub>1 \ w\ \ iso \ \ \ = (s \ \) \ \[s, s\<^sub>0, 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 r\<^sub>0 \ src s\<^sub>0\" and "\chine: chine \ chine\" proof - show "\chine : src r\<^sub>0 \ src s\<^sub>0\" 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 r\<^sub>0" and "trg chine = src s\<^sub>0" 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 r\<^sub>0 r\<^sub>1 s\<^sub>0 s\<^sub>1 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 r\<^sub>0 r\<^sub>1 s \ s\<^sub>0 s\<^sub>1 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 "(\ \ r\<^sub>0) \ \ = (s \ the_\) \ \[s, s\<^sub>0, 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 r\<^sub>0 r\<^sub>1 s\<^sub>0 s\<^sub>1 w using assms in_hom by auto interpret w: arrow_of_spans_of_maps_to_tabulation V H \ \ src trg r\<^sub>0 r\<^sub>1 s \ s\<^sub>0 s\<^sub>1 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 r\<^sub>0 \dom \\ s\<^sub>0 s\<^sub>1 w' proof show "is_left_adjoint r\<^sub>0" by (simp add: \.satisfies_T0) show "is_left_adjoint (dom \)" by (simp add: \.leg1_is_map) show "ide w'" using assms by force show "\\. \\ : s\<^sub>0 \ w' \ r\<^sub>0\" 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 "\(s\<^sub>1 \ \) \ w.the_\ : r\<^sub>1 \ s\<^sub>1 \ w'\ \ iso ((s\<^sub>1 \ \) \ w.the_\)" proof (intro conjI) show "\(s\<^sub>1 \ \) \ w.the_\ : r\<^sub>1 \ s\<^sub>1 \ w'\" using \ w.the_\_props by (intro comp_in_homI, auto) thus "iso ((s\<^sub>1 \ \) \ w.the_\)" using \ w.the_\_props by (meson \.ide_leg1 arrI iso_hcomp hseqE ide_is_iso isos_compose seqE) qed hence "\\. \\ : r\<^sub>1 \ s\<^sub>1 \ w'\ \ iso \" by auto thus "\\. \\ : dom \ \ s\<^sub>1 \ w'\ \ iso \" by auto qed interpret w': arrow_of_spans_of_maps_to_tabulation V H \ \ src trg r\<^sub>0 \dom \\ s \ s\<^sub>0 s\<^sub>1 w' .. show "arrow_of_spans_of_maps V H \ \ src trg r\<^sub>0 (dom \) s\<^sub>0 s\<^sub>1 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_\ \ (s\<^sub>0 \ inv \)" proof - have "\w.the_\ \ (s\<^sub>0 \ inv \) : s\<^sub>0 \ w' \ r\<^sub>0\" using w.the_\_props \ by (intro comp_in_homI, auto) moreover have "\w'.the_\ : s\<^sub>0 \ w' \ r\<^sub>0\" using w'.the_\_props by simp ultimately show ?thesis using w'.leg0_uniquely_isomorphic(2) by blast qed moreover have "w'.the_\ = (s\<^sub>1 \ \) \ w.the_\" proof - have "\(s\<^sub>1 \ \) \ w.the_\ : dom \ \ s\<^sub>1 \ w'\" using w.the_\_props \ by (intro comp_in_homI, auto) moreover have "iso ((s\<^sub>1 \ \) \ 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_\ \ (s\<^sub>0 \ inv \)) \ (s\<^sub>1 \ \) \ w.the_\" by simp also have "... = (s \ w.the_\ \ (s\<^sub>0 \ inv \)) \ \[s, s\<^sub>0, w'] \ (\ \ w') \ (s\<^sub>1 \ \) \ w.the_\" using comp_assoc by presburger also have "... = (s \ w.the_\) \ ((s \ s\<^sub>0 \ inv \) \ \[s, s\<^sub>0, w'] \ (\ \ w') \ (s\<^sub>1 \ \)) \ w.the_\" using 1 comp_assoc w'.the_\_simps(1) whisker_left by auto also have "... = (s \ w.the_\) \ (\[s, s\<^sub>0, w] \ (\ \ w)) \ w.the_\" proof - have "(s \ s\<^sub>0 \ inv \) \ \[s, s\<^sub>0, w'] \ (\ \ w') \ (s\<^sub>1 \ \) = \[s, s\<^sub>0, w] \ (\ \ w)" proof - have "(s \ s\<^sub>0 \ inv \) \ \[s, s\<^sub>0, w'] \ (\ \ w') \ (s\<^sub>1 \ \) = \[s, s\<^sub>0, w] \ ((s \ s\<^sub>0) \ inv \) \ (\ \ w') \ (s\<^sub>1 \ \)" proof - have "(s \ s\<^sub>0 \ inv \) \ \[s, s\<^sub>0, w'] = \[s, s\<^sub>0, w] \ ((s \ s\<^sub>0) \ inv \)" using assms \ assoc_naturality [of s s\<^sub>0 "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, s\<^sub>0, w] \ (\ \ w) \ (s\<^sub>1 \ inv \) \ (s\<^sub>1 \ \)" proof - have "((s \ s\<^sub>0) \ inv \) \ (\ \ w') = (\ \ w) \ (s\<^sub>1 \ inv \)" using \ comp_arr_dom comp_cod_arr in_hhom_def interchange [of "s \ s\<^sub>0" \ "inv \" w'] interchange [of \ s\<^sub>1 w "inv \"] by auto thus ?thesis using comp_assoc by metis qed also have "... = \[s, s\<^sub>0, w] \ (\ \ w)" proof - have "(\ \ w) \ (s\<^sub>1 \ inv \) \ (s\<^sub>1 \ \) = \ \ w" proof - have "(\ \ w) \ (s\<^sub>1 \ inv \) \ (s\<^sub>1 \ \) = (\ \ w) \ (s\<^sub>1 \ inv \ \ \)" using \ whisker_left in_hhom_def by auto also have "... = (\ \ w) \ (s\<^sub>1 \ w)" using \ comp_inv_arr' by auto also have "... = \ \ w" using whisker_right [of w \ s\<^sub>1] 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 s\<^sub>0 \ src r\<^sub>0\ \ \w' : src r\<^sub>0 \ src s\<^sub>0\ \ \\ : r\<^sub>0 \ w \ s\<^sub>0\ \ \\ : s\<^sub>1 \ r\<^sub>1 \ w\ \ iso \ \ \ = (s \ \) \ \[s, r\<^sub>0, w] \ (\ \ w) \ \ \ \\' : s\<^sub>0 \ w' \ r\<^sub>0\ \ \\' : r\<^sub>1 \ s\<^sub>1 \ w'\ \ iso \' \ \ = (s \ \') \ \[s, s\<^sub>0, w'] \ (\ \ w') \ \'" using r_eq_s \.apex_unique_up_to_equivalence [of \ r\<^sub>0 r\<^sub>1] \.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) r\<^sub>0 where "r\<^sub>0 \ tab\<^sub>0 r" abbreviation (input) r\<^sub>1 where "r\<^sub>1 \ tab\<^sub>1 r" abbreviation (input) s\<^sub>0 where "s\<^sub>0 \ tab\<^sub>0 s" abbreviation (input) s\<^sub>1 where "s\<^sub>1 \ tab\<^sub>1 s" lemma is_arrow_of_tabulations_in_maps: shows "arrow_of_tabulations_in_maps V H \ \ src trg r r.tab r\<^sub>0 r\<^sub>1 s s.tab s\<^sub>0 s\<^sub>1 \" 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 r\<^sub>0 r\<^sub>1 s s.tab s\<^sub>0 s\<^sub>1 \ using is_arrow_of_tabulations_in_maps by simp interpretation r: arrow_of_tabulations_in_maps V H \ \ src trg r r.tab r\<^sub>0 r\<^sub>1 r r.tab r\<^sub>0 r\<^sub>1 r using r.is_arrow_of_tabulations_in_maps by simp lemma \_in_terms_of_\: shows "\ = r.T0.trnr\<^sub>\ (cod \) \ \ inv (r.T0.trnr\<^sub>\ r r.tab)" proof - have \: "arr \" using in_hom by auto have "\ \ r.T0.trnr\<^sub>\ r r.tab = r.T0.trnr\<^sub>\ s \" proof - have "\ \ r.T0.trnr\<^sub>\ r r.tab = (\ \ \[r]) \ (r \ r.\) \ \[r, tab\<^sub>0 r, (tab\<^sub>0 r)\<^sup>*] \ (r.tab \ (tab\<^sub>0 r)\<^sup>*)" unfolding r.T0.trnr\<^sub>\_def using comp_assoc by presburger also have "... = \[s] \ ((\ \ src \) \ (r \ r.\)) \ \[r, tab\<^sub>0 r, (tab\<^sub>0 r)\<^sup>*] \ (r.tab \ (tab\<^sub>0 r)\<^sup>*)" using \ runit_naturality comp_assoc by (metis in_hom in_homE) also have "... = \[s] \ (s \ r.\) \ ((\ \ tab\<^sub>0 r \ (tab\<^sub>0 r)\<^sup>*) \ \[r, tab\<^sub>0 r, (tab\<^sub>0 r)\<^sup>*]) \ (r.tab \ (tab\<^sub>0 r)\<^sup>*)" 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.\) \ (\ \ tab\<^sub>0 r \ (tab\<^sub>0 r)\<^sup>*)" using in_hom interchange [of s \ r.\ "tab\<^sub>0 r \ (tab\<^sub>0 r)\<^sup>*"] 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.\) \ (\ \ tab\<^sub>0 r \ (tab\<^sub>0 r)\<^sup>*)" by blast thus ?thesis using comp_assoc by presburger qed also have "... = \[s] \ (s \ r.\) \ \[s, tab\<^sub>0 r, (tab\<^sub>0 r)\<^sup>*] \ ((\ \ tab\<^sub>0 r) \ (tab\<^sub>0 r)\<^sup>*) \ (r.tab \ (tab\<^sub>0 r)\<^sup>*)" proof - have "(\ \ tab\<^sub>0 r \ (tab\<^sub>0 r)\<^sup>*) \ \[r, tab\<^sub>0 r, (tab\<^sub>0 r)\<^sup>*] = \[s, tab\<^sub>0 r, (tab\<^sub>0 r)\<^sup>*] \ ((\ \ tab\<^sub>0 r) \ (tab\<^sub>0 r)\<^sup>*)" using \ assoc_naturality [of \ "tab\<^sub>0 r" "(tab\<^sub>0 r)\<^sup>*"] 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, tab\<^sub>0 r, (tab\<^sub>0 r)\<^sup>*] \ ((\ \ tab\<^sub>0 r) \ r.tab \ (tab\<^sub>0 r)\<^sup>*)" using \ whisker_right \_simps(1) by auto also have "... = r.T0.trnr\<^sub>\ s \" unfolding r.T0.trnr\<^sub>\_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 \ r\<^sub>0 r\<^sub>1 + \: tabulation_in_maps V H \ \ src trg s \ s\<^sub>0 s\<^sub>1 + \: tabulation_in_maps V H \ \ src trg t \ t\<^sub>0 t\<^sub>1 + \: arrow_of_tabulations_in_maps V H \ \ src trg r \ r\<^sub>0 r\<^sub>1 s \ s\<^sub>0 s\<^sub>1 \ + \: arrow_of_tabulations_in_maps V H \ \ src trg s \ s\<^sub>0 s\<^sub>1 t \ t\<^sub>0 t\<^sub>1 \ 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 r\<^sub>0 :: 'a and r\<^sub>1 :: 'a and s :: 'a and \ :: 'a and s\<^sub>0 :: 'a and s\<^sub>1 :: 'a and t :: 'a and \ :: 'a and t\<^sub>0 :: 'a and t\<^sub>1 :: '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 \ r\<^sub>0 r\<^sub>1 t \ t\<^sub>0 t\<^sub>1 \\ \ \\ 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 \ r\<^sub>0 r\<^sub>1 t \ t\<^sub>0 t\<^sub>1 (\ \ \)" .. 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) \ \\<^sup>-\<^sup>1[t\<^sub>0, ?g, ?f]" let ?\ = "\[t\<^sub>1, ?g, ?f] \ (\.the_\ \ ?f) \ \.the_\" have \: "\?\ : t\<^sub>0 \ ?g \ ?f \ r\<^sub>0\" using f g \.the_\_props \.the_\_props by (intro comp_in_homI hcomp_in_vhom, auto+) have \: "\?\ : r\<^sub>1 \ t\<^sub>1 \ ?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 r\<^sub>0 r\<^sub>1 t\<^sub>0 t\<^sub>1 \?g \ ?f\ proof show "ide (?g \ ?f)" by simp show "\\. \\ : t\<^sub>0 \ ?g \ ?f \ r\<^sub>0\" using \ by blast show "\\. \\ : r\<^sub>1 \ t\<^sub>1 \ ?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 r\<^sub>0 (dom \) t\<^sub>0 t\<^sub>1 (?g \ ?f)" using gf.arrow_of_spans_of_maps_axioms by simp have "((t \ gf.the_\) \ \[t, t\<^sub>0, ?g \ ?f] \ (\ \ ?g \ ?f)) \ gf.the_\ = \" proof - have "\ = (\ \ r\<^sub>0) \ (\ \ r\<^sub>0) \ \" using whisker_right comp_assoc by (metis \_simps(1) hseqE ide_u seqE) also have "... = ((\ \ r\<^sub>0) \ (s \ \.the_\)) \ \[s, s\<^sub>0, ?f] \ (\ \ ?f) \ \.the_\" using \.\_naturality comp_assoc by simp also have "... = (t \ \.the_\) \ ((\ \ s\<^sub>0 \ ?f) \ \[s, s\<^sub>0, ?f]) \ (\ \ ?f) \ \.the_\" proof - have "(\ \ r\<^sub>0) \ (s \ \.the_\) = \ \ \.the_\" using f comp_arr_dom comp_cod_arr \.the_\_props \.in_hom interchange [of \ s r\<^sub>0 \.the_\] by (metis in_homE) also have "... = (t \ \.the_\) \ (\ \ s\<^sub>0 \ ?f)" using f comp_arr_dom comp_cod_arr \.the_\_props \.in_hom interchange [of t \ \.the_\ "s\<^sub>0 \ ?f"] by (metis in_homE) finally have "(\ \ r\<^sub>0) \ (s \ \.the_\) = (t \ \.the_\) \ (\ \ s\<^sub>0 \ ?f)" by simp thus ?thesis using comp_assoc by presburger qed also have "... = (t \ \.the_\) \ \[t, s\<^sub>0, ?f] \ (((\ \ s\<^sub>0) \ ?f) \ (\ \ ?f)) \ \.the_\" proof - have "(\ \ s\<^sub>0 \ ?f) \ \[s, s\<^sub>0, ?f] = \[t, s\<^sub>0, ?f] \ ((\ \ s\<^sub>0) \ ?f)" using f assoc_naturality [of \ s\<^sub>0 ?f] \.in_hom by auto thus ?thesis using comp_assoc by presburger qed also have "... = (t \ \.the_\) \ \[t, s\<^sub>0, ?f] \ (\.\ \ ?f) \ \.the_\" using whisker_right comp_assoc by simp also have "... = (t \ \.the_\) \ \[t, s\<^sub>0, ?f] \ ((t \ \.the_\) \ \[t, t\<^sub>0, ?g] \ (\ \ ?g) \ \.the_\ \ ?f) \ \.the_\" using \.\_naturality by simp also have "... = (t \ \.the_\) \ \[t, s\<^sub>0, ?f] \ (((t \ \.the_\) \ ?f) \ (\[t, t\<^sub>0, ?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, s\<^sub>0, ?f] \ ((t \ \.the_\) \ ?f)) \ (\[t, t\<^sub>0, ?g] \ ?f) \ ((\ \ ?g) \ ?f) \ (\.the_\ \ ?f) \ \.the_\" using comp_assoc by presburger also have "... = (t \ \.the_\) \ (t \ \.the_\ \ ?f) \ (\[t, t\<^sub>0 \ ?g, ?f] \ (\[t, t\<^sub>0, ?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 \ \\<^sup>-\<^sup>1[t\<^sub>0, ?g, ?f]) \ \[t, t\<^sub>0, ?g \ ?f] \ (\[t \ t\<^sub>0, ?g, ?f] \ ((\ \ ?g) \ ?f)) \ (\.the_\ \ ?f) \ \.the_\" proof - have "seq \[t, t\<^sub>0, ?g \ ?f] \[t \ t\<^sub>0, ?g, ?f]" using f g by fastforce moreover have "inv (t \ \[t\<^sub>0, ?g, ?f]) = t \ \\<^sup>-\<^sup>1[t\<^sub>0, ?g, ?f]" using f g by simp moreover have "iso (t \ \[t\<^sub>0, ?g, ?f])" using f g by simp have "\[t, t\<^sub>0 \ ?g, ?f] \ (\[t, t\<^sub>0, ?g] \ ?f) = (t \ \\<^sup>-\<^sup>1[t\<^sub>0, ?g, ?f]) \ \[t, t\<^sub>0, ?g \ ?f] \ \[t \ t\<^sub>0, ?g, ?f]" proof - have "seq \[t, t\<^sub>0, ?g \ ?f] \[t \ t\<^sub>0, ?g, ?f]" using f g by fastforce moreover have "inv (t \ \[t\<^sub>0, ?g, ?f]) = t \ \\<^sup>-\<^sup>1[t\<^sub>0, ?g, ?f]" using f g by simp moreover have "iso (t \ \[t\<^sub>0, ?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 \ \\<^sup>-\<^sup>1[t\<^sub>0, ?g, ?f]) \ \[t, t\<^sub>0, ?g \ ?f] \ (\ \ ?g \ ?f) \ \[t\<^sub>1, ?g, ?f] \ (\.the_\ \ ?f) \ \.the_\" using f g assoc_naturality [of \ ?g ?f] comp_assoc by simp also have "... = (t \ \.the_\ \ (\.the_\ \ ?f) \ \\<^sup>-\<^sup>1[t\<^sub>0, ?g, ?f]) \ \[t, t\<^sub>0, ?g \ ?f] \ (\ \ ?g \ ?f) \ \[t\<^sub>1, ?g, ?f] \ (\.the_\ \ ?f) \ \.the_\" proof - have 1: "seq \.the_\ ((\.the_\ \ ?f) \ \\<^sup>-\<^sup>1[t\<^sub>0, ?g, ?f])" using \_eq by auto hence "t \ (\.the_\ \ ?f) \ \\<^sup>-\<^sup>1[t\<^sub>0, ?g, ?f] = (t \ \.the_\ \ ?f) \ (t \ \\<^sup>-\<^sup>1[t\<^sub>0, ?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, t\<^sub>0, ?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, t\<^sub>0, ?g \ ?f] \ (\ \ ?g \ ?f)) \ arrow_of_spans_of_maps.the_\ (\) (\) (dom ((\ \ \ \ r\<^sub>0) \ \)) t\<^sub>1 (?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 \ r\<^sub>0 r\<^sub>1 t \ t\<^sub>0 t\<^sub>1 "\ \ \" 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 \ r\<^sub>0 r\<^sub>1 + \: tabulation_in_maps V H \ \ src trg s \ s\<^sub>0 s\<^sub>1 + \: tabulation_in_maps V H \ \ src trg t \ t\<^sub>0 t\<^sub>1 + \: tabulation_in_maps V H \ \ src trg u \ u\<^sub>0 u\<^sub>1 + \\: composite_tabulation_in_maps V H \ \ src trg r \ r\<^sub>0 r\<^sub>1 s \ s\<^sub>0 s\<^sub>1 + \\: composite_tabulation_in_maps V H \ \ src trg t \ t\<^sub>0 t\<^sub>1 u \ u\<^sub>0 u\<^sub>1 + \: arrow_of_tabulations_in_maps V H \ \ src trg r \ r\<^sub>0 r\<^sub>1 t \ t\<^sub>0 t\<^sub>1 \ + \: arrow_of_tabulations_in_maps V H \ \ src trg s \ s\<^sub>0 s\<^sub>1 u \ u\<^sub>0 u\<^sub>1 \ 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 r\<^sub>0 :: 'a and r\<^sub>1 :: 'a and s :: 'a and \ :: 'a and s\<^sub>0 :: 'a and s\<^sub>1 :: 'a and t :: 'a and \ :: 'a and t\<^sub>0 :: 'a and t\<^sub>1 :: 'a and u :: 'a and \ :: 'a and u\<^sub>0 :: 'a and u\<^sub>1 :: '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 \s\<^sub>0 \ \\.p\<^sub>0\ \r\<^sub>1 \ \\.p\<^sub>1\ \t \ u\ \\.tab \u\<^sub>0 \ \\.p\<^sub>0\ \t\<^sub>1 \ \\.p\<^sub>1\ \\ \ \\ 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 (s\<^sub>0 \ \\.p\<^sub>0) (r\<^sub>1 \ \\.p\<^sub>1) (t \ u) \\.tab (u\<^sub>0 \ \\.p\<^sub>0) (t\<^sub>1 \ \\.p\<^sub>1) (\ \ \)" .. sublocale arrow_of_tabulations_in_maps V H \ \ src trg \r \ s\ \\.tab \s\<^sub>0 \ \\.p\<^sub>0\ \r\<^sub>1 \ \\.p\<^sub>1\ \t \ u\ \\.tab \u\<^sub>0 \ \\.p\<^sub>0\ \t\<^sub>1 \ \\.p\<^sub>1\ \\ \ \\ using is_arrow_of_tabulations_in_maps by simp interpretation Maps: maps_category V H \ \ src trg .. notation Maps.comp (infixr "\" 55) interpretation r\<^sub>0s\<^sub>1: cospan_of_maps_in_bicategory_of_spans \(\)\ \(\)\ \ \ src trg s\<^sub>1 r\<^sub>0 using \.leg0_is_map \.leg1_is_map \\.composable apply unfold_locales by auto interpretation r\<^sub>0s\<^sub>1: arrow_of_tabulations_in_maps \(\)\ \(\)\ \ \ src trg \r\<^sub>0\<^sup>* \ s\<^sub>1\ r\<^sub>0s\<^sub>1.tab r\<^sub>0s\<^sub>1.p\<^sub>0 r\<^sub>0s\<^sub>1.p\<^sub>1 \r\<^sub>0\<^sup>* \ s\<^sub>1\ r\<^sub>0s\<^sub>1.tab r\<^sub>0s\<^sub>1.p\<^sub>0 r\<^sub>0s\<^sub>1.p\<^sub>1 \r\<^sub>0\<^sup>* \ s\<^sub>1\ using r\<^sub>0s\<^sub>1.is_arrow_of_tabulations_in_maps by simp interpretation t\<^sub>0u\<^sub>1: cospan_of_maps_in_bicategory_of_spans \(\)\ \(\)\ \ \ src trg u\<^sub>1 t\<^sub>0 using \.leg0_is_map \.leg1_is_map \\.composable apply unfold_locales by auto interpretation t\<^sub>0u\<^sub>1: arrow_of_tabulations_in_maps \(\)\ \(\)\ \ \ src trg \t\<^sub>0\<^sup>* \ u\<^sub>1\ t\<^sub>0u\<^sub>1.tab t\<^sub>0u\<^sub>1.p\<^sub>0 t\<^sub>0u\<^sub>1.p\<^sub>1 \t\<^sub>0\<^sup>* \ u\<^sub>1\ t\<^sub>0u\<^sub>1.tab t\<^sub>0u\<^sub>1.p\<^sub>0 t\<^sub>0u\<^sub>1.p\<^sub>1 \t\<^sub>0\<^sup>* \ u\<^sub>1\ using t\<^sub>0u\<^sub>1.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 "\\.p\<^sub>0 \ chine \ \.chine \ \\.p\<^sub>0" and "\\.p\<^sub>1 \ chine \ \.chine \ \\.p\<^sub>1" proof - have 1: "arrow_of_spans_of_maps V H \ \ src trg (s\<^sub>0 \ \\.p\<^sub>0) (r\<^sub>1 \ \\.p\<^sub>1) (u\<^sub>0 \ \\.p\<^sub>0) (t\<^sub>1 \ \\.p\<^sub>1) chine \ (((t \ u) \ the_\) \ \[t \ u, u\<^sub>0 \ \\.p\<^sub>0, chine] \ (\\.tab \ chine)) \ the_\ = ((\ \ \) \ s\<^sub>0 \ \\.p\<^sub>0) \ \\.tab" using chine_is_induced_map by simp let ?u\<^sub>\ = "u \ s\<^sub>0 \ \\.p\<^sub>0" let ?w\<^sub>\ = "\.chine \ \\.p\<^sub>1" let ?w\<^sub>\' = "\\.p\<^sub>1 \ chine" have u\<^sub>\: "ide ?u\<^sub>\" using \.u_simps(3) by auto have w\<^sub>\: "ide ?w\<^sub>\ \ is_left_adjoint ?w\<^sub>\" by (simp add: \.is_map \.T0.antipar(1) left_adjoints_compose) have w\<^sub>\': "ide ?w\<^sub>\' \ is_left_adjoint ?w\<^sub>\'" by (simp add: is_map left_adjoints_compose) let ?\\<^sub>\ = "\[u, s\<^sub>0, \\.p\<^sub>0] \ ((\ \ s\<^sub>0) \ \ \ \\.p\<^sub>0) \ r\<^sub>0s\<^sub>1.\ \ (\.the_\ \ \\.p\<^sub>1) \ \\<^sup>-\<^sup>1[t\<^sub>0, \.chine, \\.p\<^sub>1]" let ?\\<^sub>\' = "(u \ the_\) \ \[u, u\<^sub>0 \ \\.p\<^sub>0, chine] \ (\[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ ((\ \ \\.p\<^sub>0) \ chine) \ (t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[t\<^sub>0, \\.p\<^sub>1, chine]" let ?\\<^sub>\ = "\[t\<^sub>1, \\.p\<^sub>1, chine] \ the_\ \ (inv \.the_\ \ \\.p\<^sub>1) \ \\<^sup>-\<^sup>1[t\<^sub>1, \.chine, \\.p\<^sub>1]" have \\<^sub>\: "\?\\<^sub>\ : t\<^sub>0 \ ?w\<^sub>\ \ ?u\<^sub>\\" using \.T0.antipar(1) \.the_\_in_hom \.u_simps(3) by (intro comp_in_homI, auto) have \\<^sub>\': "\?\\<^sub>\' : t\<^sub>0 \ ?w\<^sub>\' \ ?u\<^sub>\\" proof (intro comp_in_homI) show "\\\<^sup>-\<^sup>1[t\<^sub>0, \\.p\<^sub>1, chine] : t\<^sub>0 \ \\.p\<^sub>1 \ chine \ (t\<^sub>0 \ \\.p\<^sub>1) \ chine\" using t\<^sub>0u\<^sub>1.p\<^sub>1_simps assoc'_in_hom by simp show "\t\<^sub>0u\<^sub>1.\ \ chine : (t\<^sub>0 \ \\.p\<^sub>1) \ chine \ (u\<^sub>1 \ \\.p\<^sub>0) \ chine\" using \.T0.antipar(1) by (intro hcomp_in_vhom, auto) show "\(\ \ \\.p\<^sub>0) \ chine : (u\<^sub>1 \ \\.p\<^sub>0) \ chine \ ((u \ u\<^sub>0) \ \\.p\<^sub>0) \ chine\" by (intro hcomp_in_vhom, auto) show "\\[u, u\<^sub>0, \\.p\<^sub>0] \ chine : ((u \ u\<^sub>0) \ \\.p\<^sub>0) \ chine \ (u \ u\<^sub>0 \ \\.p\<^sub>0) \ chine\" using assoc_in_hom by auto show "\\[u, u\<^sub>0 \ \\.p\<^sub>0, chine] : (u \ u\<^sub>0 \ \\.p\<^sub>0) \ chine \ u \ (u\<^sub>0 \ \\.p\<^sub>0) \ chine\" by auto show "\u \ the_\ : u \ (u\<^sub>0 \ \\.p\<^sub>0) \ chine \ ?u\<^sub>\\" by (intro hcomp_in_vhom, auto) qed have \\<^sub>\: "\?\\<^sub>\ : t\<^sub>1 \ ?w\<^sub>\ \ t\<^sub>1 \ ?w\<^sub>\'\" proof (intro comp_in_homI) show "\\\<^sup>-\<^sup>1[t\<^sub>1, \.chine, \\.p\<^sub>1] : t\<^sub>1 \ ?w\<^sub>\ \ (t\<^sub>1 \ \.chine) \ \\.p\<^sub>1\" using \.T0.antipar(1) by auto show "\inv \.the_\ \ \\.p\<^sub>1 : (t\<^sub>1 \ \.chine) \ \\.p\<^sub>1 \ r\<^sub>1 \ \\.p\<^sub>1\" using \.the_\_props \.T0.antipar(1) by (intro hcomp_in_vhom, auto) show "\the_\ : r\<^sub>1 \ \\.p\<^sub>1 \ (t\<^sub>1 \ \\.p\<^sub>1) \ chine\" using the_\_in_hom(2) by simp show "\\[t\<^sub>1, \\.p\<^sub>1, chine] : (t\<^sub>1 \ \\.p\<^sub>1) \ chine \ t\<^sub>1 \ ?w\<^sub>\'\" using t\<^sub>0u\<^sub>1.p\<^sub>1_simps assoc_in_hom by simp qed define LHS where "LHS = (t \ ?\\<^sub>\) \ \[t, t\<^sub>0, ?w\<^sub>\] \ (\ \ ?w\<^sub>\)" have LHS: "\LHS : t\<^sub>1 \ ?w\<^sub>\ \ t \ ?u\<^sub>\\" proof (unfold LHS_def, intro comp_in_homI) show "\\ \ ?w\<^sub>\ : t\<^sub>1 \ \.chine \ \\.p\<^sub>1 \ (t \ t\<^sub>0) \ ?w\<^sub>\\" using \.T0.antipar(1) by (intro hcomp_in_vhom, auto) show "\\[t, t\<^sub>0, ?w\<^sub>\] : (t \ t\<^sub>0) \ ?w\<^sub>\ \ t \ t\<^sub>0 \ ?w\<^sub>\\" using \.T0.antipar(1) by auto show "\t \ ?\\<^sub>\ : t \ t\<^sub>0 \ ?w\<^sub>\ \ t \ ?u\<^sub>\\" proof - have "src t = trg (t\<^sub>0 \ \.chine \ r\<^sub>0s\<^sub>1.p\<^sub>1)" by (metis \.u_simps(3) \.ide_base \.ide_leg0 \.leg1_simps(3) \\.composable \\<^sub>\ arrI assoc_simps(3) r\<^sub>0s\<^sub>1.ide_u r\<^sub>0s\<^sub>1.p\<^sub>0_simps trg_vcomp vconn_implies_hpar(2)) thus ?thesis using \\<^sub>\ by blast qed qed define RHS where "RHS = ((t \ ?\\<^sub>\') \ \[t, t\<^sub>0, ?w\<^sub>\'] \ (\ \ ?w\<^sub>\')) \ ?\\<^sub>\" have RHS: "\RHS : t\<^sub>1 \ ?w\<^sub>\ \ t \ ?u\<^sub>\\" unfolding RHS_def proof show "\?\\<^sub>\ : t\<^sub>1 \ ?w\<^sub>\ \ t\<^sub>1 \ ?w\<^sub>\'\" using \\<^sub>\ by simp show "\(t \ ?\\<^sub>\') \ \[t, t\<^sub>0, ?w\<^sub>\'] \ (\ \ ?w\<^sub>\') : t\<^sub>1 \ ?w\<^sub>\' \ t \ ?u\<^sub>\\" proof show "\\[t, t\<^sub>0, ?w\<^sub>\'] \ (\ \ ?w\<^sub>\') : t\<^sub>1 \ ?w\<^sub>\' \ t \ t\<^sub>0 \ ?w\<^sub>\'\" using \.T0.antipar(1) by fastforce show "\t \ ?\\<^sub>\' : t \ t\<^sub>0 \ ?w\<^sub>\' \ t \ ?u\<^sub>\\" using w\<^sub>\' \\<^sub>\' \.leg0_simps(2) \.leg0_simps(3) hseqI' ideD(1) t\<^sub>0u\<^sub>1.p\<^sub>1_simps trg_hcomp \.base_in_hom(2) hcomp_in_vhom by presburger qed qed have eq: "LHS = RHS" proof - have "\\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ LHS \ \[t\<^sub>1, \.chine, \\.p\<^sub>1] \ (\.the_\ \ \\.p\<^sub>1) = \" 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 \\.p\<^sub>1} on the right. \ have "\\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ LHS \ \[t\<^sub>1, \.chine, \\.p\<^sub>1] \ (\.the_\ \ \\.p\<^sub>1) = \\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ (t \ \[u, s\<^sub>0, \\.p\<^sub>0] \ ((\ \ s\<^sub>0) \ \ \ \\.p\<^sub>0) \ r\<^sub>0s\<^sub>1.\ \ (\.the_\ \ \\.p\<^sub>1) \ \\<^sup>-\<^sup>1[t\<^sub>0, \.chine, \\.p\<^sub>1]) \ \[t, t\<^sub>0, \.chine \ \\.p\<^sub>1] \ (\ \ \.chine \ \\.p\<^sub>1) \ \[t\<^sub>1, \.chine, \\.p\<^sub>1] \ (\.the_\ \ \\.p\<^sub>1)" unfolding LHS_def using comp_assoc by presburger also have "... = \\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ (t \ \[u, s\<^sub>0, \\.p\<^sub>0]) \ (t \ (\ \ s\<^sub>0) \ \ \ \\.p\<^sub>0) \ (t \ r\<^sub>0s\<^sub>1.\) \ (t \ \.the_\ \ \\.p\<^sub>1) \ (t \ \\<^sup>-\<^sup>1[t\<^sub>0, \.chine, \\.p\<^sub>1]) \ \[t, t\<^sub>0, \.chine \ \\.p\<^sub>1] \ ((\ \ \.chine \ \\.p\<^sub>1) \ \[t\<^sub>1, \.chine, \\.p\<^sub>1]) \ (\.the_\ \ \\.p\<^sub>1)" proof - have "t \ \[u, s\<^sub>0, \\.p\<^sub>0] \ ((\ \ s\<^sub>0) \ \ \ \\.p\<^sub>0) \ r\<^sub>0s\<^sub>1.\ \ (\.the_\ \ \\.p\<^sub>1) \ \\<^sup>-\<^sup>1[t\<^sub>0, \.chine, \\.p\<^sub>1] = (t \ \[u, s\<^sub>0, \\.p\<^sub>0]) \ (t \ (\ \ s\<^sub>0) \ \ \ \\.p\<^sub>0) \ (t \ r\<^sub>0s\<^sub>1.\) \ (t \ \.the_\ \ \\.p\<^sub>1) \ (t \ \\<^sup>-\<^sup>1[t\<^sub>0, \.chine, \\.p\<^sub>1])" using whisker_left \.ide_base \\<^sub>\ arrI seqE by (metis (full_types)) thus ?thesis using comp_assoc by presburger qed also have "... = \\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ (t \ \[u, s\<^sub>0, \\.p\<^sub>0]) \ (t \ (\ \ s\<^sub>0) \ \ \ \\.p\<^sub>0) \ (t \ r\<^sub>0s\<^sub>1.\) \ (t \ \.the_\ \ \\.p\<^sub>1) \ ((t \ \\<^sup>-\<^sup>1[t\<^sub>0, \.chine, \\.p\<^sub>1]) \ \[t, t\<^sub>0, \.chine \ \\.p\<^sub>1] \ \[t \ t\<^sub>0, \.chine, \\.p\<^sub>1]) \ ((\ \ \.chine) \ \\.p\<^sub>1) \ (\.the_\ \ \\.p\<^sub>1)" proof - have "(\ \ \.chine \ \\.p\<^sub>1) \ \[t\<^sub>1, \.chine, \\.p\<^sub>1] = \[t \ t\<^sub>0, \.chine, \\.p\<^sub>1] \ ((\ \ \.chine) \ \\.p\<^sub>1)" using assoc_naturality by (metis \.w_simps(2-6) \.leg1_simps(3) \\.leg1_simps(2) \.tab_simps(1) \.tab_simps(2,4-5) hseqE r\<^sub>0s\<^sub>1.leg1_simps(5) r\<^sub>0s\<^sub>1.leg1_simps(6)) thus ?thesis using comp_assoc by presburger qed also have "... = \\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ (t \ \[u, s\<^sub>0, \\.p\<^sub>0]) \ (t \ (\ \ s\<^sub>0) \ \ \ \\.p\<^sub>0) \ (t \ r\<^sub>0s\<^sub>1.\) \ ((t \ \.the_\ \ \\.p\<^sub>1) \ \[t, t\<^sub>0 \ \.chine, \\.p\<^sub>1]) \ (\[t, t\<^sub>0, \.chine] \ \\.p\<^sub>1) \ ((\ \ \.chine) \ \\.p\<^sub>1) \ (\.the_\ \ \\.p\<^sub>1)" proof - have "(t \ \\<^sup>-\<^sup>1[t\<^sub>0, \.chine, \\.p\<^sub>1]) \ \[t, t\<^sub>0, \.chine \ \\.p\<^sub>1] \ \[t \ t\<^sub>0, \.chine, \\.p\<^sub>1] = \[t, t\<^sub>0 \ \.chine, \\.p\<^sub>1] \ (\[t, t\<^sub>0, \.chine] \ \\.p\<^sub>1)" proof - have "seq \[t, t\<^sub>0, \.chine \ \\.p\<^sub>1] \[t \ t\<^sub>0, \.chine, \\.p\<^sub>1]" by (simp add: \.T0.antipar(1)) moreover have "inv (t \ \[t\<^sub>0, \.chine, \\.p\<^sub>1]) = t \ \\<^sup>-\<^sup>1[t\<^sub>0, \.chine, \\.p\<^sub>1]" using \.T0.antipar(1) inv_hcomp [of t "\[t\<^sub>0, \.chine, \\.p\<^sub>1]"] by simp ultimately show ?thesis using pentagon \.T0.antipar(1) iso_hcomp invert_side_of_triangle(1) [of "\[t, t\<^sub>0, \.chine \ \\.p\<^sub>1] \ \[t \ t\<^sub>0, \.chine, \\.p\<^sub>1]" "t \ \[t\<^sub>0, \.chine, \\.p\<^sub>1]" "\[t, t\<^sub>0 \ \.chine, \\.p\<^sub>1] \ (\[t, t\<^sub>0, \.chine] \ \\.p\<^sub>1)"] by simp qed thus ?thesis using comp_assoc by presburger qed also have "... = \\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ (t \ \[u, s\<^sub>0, \\.p\<^sub>0]) \ (t \ (\ \ s\<^sub>0) \ \ \ \\.p\<^sub>0) \ (t \ r\<^sub>0s\<^sub>1.\) \ \[t, r\<^sub>0, \\.p\<^sub>1] \ (((t \ \.the_\) \ \\.p\<^sub>1) \ (\[t, t\<^sub>0, \.chine] \ \\.p\<^sub>1) \ ((\ \ \.chine) \ \\.p\<^sub>1)) \ (\.the_\ \ \\.p\<^sub>1)" proof - have "(t \ \.the_\ \ \\.p\<^sub>1) \ \[t, t\<^sub>0 \ \.chine, \\.p\<^sub>1] = \[t, r\<^sub>0, \\.p\<^sub>1] \ ((t \ \.the_\) \ \\.p\<^sub>1)" using assoc_naturality [of t \.the_\ \\.p\<^sub>1] \.\_simps(3) \\.leg1_simps(2) hseq_char by auto thus ?thesis using comp_assoc by presburger qed also have "... = \\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ (t \ \[u, s\<^sub>0, \\.p\<^sub>0]) \ (t \ (\ \ s\<^sub>0) \ \ \ \\.p\<^sub>0) \ (t \ r\<^sub>0s\<^sub>1.\) \ \[t, r\<^sub>0, \\.p\<^sub>1] \ ((\ \ r\<^sub>0) \ \ \ \\.p\<^sub>1)" using whisker_right \.T0.antipar(1) \.\_simps(1) \.\_naturality comp_assoc by fastforce also have "... = \\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ ((t \ \[u, s\<^sub>0, \\.p\<^sub>0]) \ (t \ (\ \ s\<^sub>0) \ \\.p\<^sub>0)) \ (t \ \ \ \\.p\<^sub>0) \ (t \ r\<^sub>0s\<^sub>1.\) \ \[t, r\<^sub>0, \\.p\<^sub>1] \ ((\ \ r\<^sub>0) \ \ \ \\.p\<^sub>1)" proof - have "t \ (\ \ s\<^sub>0) \ \ \ \\.p\<^sub>0 = (t \ (\ \ s\<^sub>0) \ \\.p\<^sub>0) \ (t \ \ \ \\.p\<^sub>0)" using whisker_left whisker_right \.T0.antipar(1) by (metis (full_types) \.\_simps(1) \.ide_base \\<^sub>\ arrI r\<^sub>0s\<^sub>1.ide_u seqE) thus ?thesis using comp_assoc by presburger qed also have "... = \\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ (t \ \ \ s\<^sub>0 \ \\.p\<^sub>0) \ (t \ \[s, s\<^sub>0, \\.p\<^sub>0]) \ (t \ \ \ \\.p\<^sub>0) \ (t \ r\<^sub>0s\<^sub>1.\) \ \[t, r\<^sub>0, \\.p\<^sub>1] \ ((\ \ r\<^sub>0) \ \ \ \\.p\<^sub>1)" proof - have "(t \ \[u, s\<^sub>0, \\.p\<^sub>0]) \ (t \ (\ \ s\<^sub>0) \ \\.p\<^sub>0) = t \ \[u, s\<^sub>0, \\.p\<^sub>0] \ ((\ \ s\<^sub>0) \ \\.p\<^sub>0)" using \.in_hom whisker_left by auto also have "... = t \ (\ \ s\<^sub>0 \ \\.p\<^sub>0) \ \[s, s\<^sub>0, \\.p\<^sub>0]" using assoc_naturality [of \ s\<^sub>0 \\.p\<^sub>0] \.in_hom by auto also have "... = (t \ \ \ s\<^sub>0 \ \\.p\<^sub>0) \ (t \ \[s, s\<^sub>0, \\.p\<^sub>0])" proof - have "seq (\ \ s\<^sub>0 \ \\.p\<^sub>0) \[s, s\<^sub>0, \\.p\<^sub>0]" 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 (\ \ s\<^sub>0 \ \\.p\<^sub>0) = s \ s\<^sub>0 \ \\.p\<^sub>0" 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 "... = \\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ (t \ \ \ s\<^sub>0 \ \\.p\<^sub>0) \ (t \ \[s, s\<^sub>0, \\.p\<^sub>0]) \ (t \ \ \ \\.p\<^sub>0) \ (t \ r\<^sub>0s\<^sub>1.\) \ (\[t, r\<^sub>0, \\.p\<^sub>1] \ ((\ \ r\<^sub>0) \ \\.p\<^sub>1)) \ (\ \ \\.p\<^sub>1)" using whisker_right comp_assoc by simp also have "... = \\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ (t \ \ \ s\<^sub>0 \ \\.p\<^sub>0) \ (t \ \[s, s\<^sub>0, \\.p\<^sub>0]) \ (t \ \ \ \\.p\<^sub>0) \ ((t \ r\<^sub>0s\<^sub>1.\) \ (\ \ r\<^sub>0 \ \\.p\<^sub>1)) \ \[r, r\<^sub>0, \\.p\<^sub>1] \ (\ \ \\.p\<^sub>1)" using assoc_naturality [of \ r\<^sub>0 \\.p\<^sub>1] \.in_hom \.T0.antipar(1) comp_assoc by fastforce also have "... = \\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ ((t \ \ \ s\<^sub>0 \ \\.p\<^sub>0) \ (t \ \[s, s\<^sub>0, \\.p\<^sub>0]) \ (t \ \ \ \\.p\<^sub>0)) \ (\ \ s\<^sub>1 \ \\.p\<^sub>0) \ (r \ r\<^sub>0s\<^sub>1.\) \ \[r, r\<^sub>0, \\.p\<^sub>1] \ (\ \ \\.p\<^sub>1)" proof - have "(t \ r\<^sub>0s\<^sub>1.\) \ (\ \ r\<^sub>0 \ \\.p\<^sub>1) = \ \ r\<^sub>0s\<^sub>1.\" using comp_cod_arr comp_arr_dom \.in_hom interchange comp_ide_arr by (metis \.base_in_hom(2) \.ide_base r\<^sub>0s\<^sub>1.\_simps(1) r\<^sub>0s\<^sub>1.\_simps(4) seqI') also have "... = (\ \ s\<^sub>1 \ \\.p\<^sub>0) \ (r \ r\<^sub>0s\<^sub>1.\)" using r\<^sub>0s\<^sub>1.\_in_hom comp_cod_arr comp_arr_dom \.in_hom interchange by (metis in_homE) finally have "(t \ r\<^sub>0s\<^sub>1.\) \ (\ \ r\<^sub>0 \ \\.p\<^sub>1) = (\ \ s\<^sub>1 \ \\.p\<^sub>0) \ (r \ r\<^sub>0s\<^sub>1.\)" by simp thus ?thesis using comp_assoc by presburger qed also have "... = \\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ ((t \ (\ \ s\<^sub>0 \ \\.p\<^sub>0) \ \[s, s\<^sub>0, \\.p\<^sub>0] \ (\ \ \\.p\<^sub>0)) \ (\ \ s\<^sub>1 \ \\.p\<^sub>0)) \ (r \ r\<^sub>0s\<^sub>1.\) \ \[r, r\<^sub>0, \\.p\<^sub>1] \ (\ \ \\.p\<^sub>1)" using whisker_left \.T0.antipar(1) \\.composable \.in_hom comp_assoc by auto also have "... = \\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ (\ \ (\ \ s\<^sub>0 \ \\.p\<^sub>0) \ \[s, s\<^sub>0, \\.p\<^sub>0] \ (\ \ \\.p\<^sub>0)) \ (r \ r\<^sub>0s\<^sub>1.\) \ \[r, r\<^sub>0, \\.p\<^sub>1] \ (\ \ \\.p\<^sub>1)" proof - have "(t \ (\ \ s\<^sub>0 \ \\.p\<^sub>0) \ \[s, s\<^sub>0, \\.p\<^sub>0] \ (\ \ \\.p\<^sub>0)) \ (\ \ s\<^sub>1 \ \\.p\<^sub>0) = \ \ (\ \ s\<^sub>0 \ \\.p\<^sub>0) \ \[s, s\<^sub>0, \\.p\<^sub>0] \ (\ \ \\.p\<^sub>0)" proof - have "\(\ \ s\<^sub>0 \ \\.p\<^sub>0) \ \[s, s\<^sub>0, \\.p\<^sub>0] \ (\ \ \\.p\<^sub>0) : s\<^sub>1 \ \\.p\<^sub>0 \ u \ s\<^sub>0 \ \\.p\<^sub>0\" 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 "... = (\\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ (\ \ \ \ s\<^sub>0 \ \\.p\<^sub>0)) \ (r \ \[s, s\<^sub>0, \\.p\<^sub>0] \ (\ \ \\.p\<^sub>0)) \ (r \ r\<^sub>0s\<^sub>1.\) \ \[r, r\<^sub>0, \\.p\<^sub>1] \ (\ \ \\.p\<^sub>1)" proof - have "\ \ (\ \ s\<^sub>0 \ \\.p\<^sub>0) \ \[s, s\<^sub>0, \\.p\<^sub>0] \ (\ \ \\.p\<^sub>0) = (\ \ \ \ s\<^sub>0 \ \\.p\<^sub>0) \ (r \ \[s, s\<^sub>0, \\.p\<^sub>0] \ (\ \ \\.p\<^sub>0))" proof - have "seq (\ \ s\<^sub>0 \ \\.p\<^sub>0) (\[s, s\<^sub>0, \\.p\<^sub>0] \ (\ \ \\.p\<^sub>0))" 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 "... = ((\ \ \) \ s\<^sub>0 \ \\.p\<^sub>0) \ \\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ \\.p\<^sub>0] \ (r \ \[s, s\<^sub>0, \\.p\<^sub>0] \ (\ \ \\.p\<^sub>0)) \ (r \ r\<^sub>0s\<^sub>1.\) \ \[r, r\<^sub>0, \\.p\<^sub>1] \ (\ \ \\.p\<^sub>1)" proof - have "\\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ (\ \ \ \ s\<^sub>0 \ \\.p\<^sub>0) = ((\ \ \) \ s\<^sub>0 \ \\.p\<^sub>0) \ \\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ \\.p\<^sub>0]" 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 "... = \\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ RHS \ \[t\<^sub>1, \.chine, \\.p\<^sub>1] \ (\.the_\ \ \\.p\<^sub>1)" proof - text \Now cancel @{term \.the_\} and its inverse.\ have "\\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ RHS \ \[t\<^sub>1, \.chine, \\.p\<^sub>1] \ (\.the_\ \ \\.p\<^sub>1) = \\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ (t \ (u \ the_\) \ \[u, u\<^sub>0 \ \\.p\<^sub>0, chine] \ (\[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ ((\ \ \\.p\<^sub>0) \ chine) \ (t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[t\<^sub>0, \\.p\<^sub>1, chine]) \ \[t, t\<^sub>0, \\.p\<^sub>1 \ chine] \ (\ \ \\.p\<^sub>1 \ chine) \ \[t\<^sub>1, \\.p\<^sub>1, chine] \ the_\ \ (inv \.the_\ \ \\.p\<^sub>1) \ ((\\<^sup>-\<^sup>1[t\<^sub>1, \.chine, \\.p\<^sub>1]) \ \[t\<^sub>1, \.chine, \\.p\<^sub>1]) \ (\.the_\ \ \\.p\<^sub>1)" unfolding RHS_def using comp_assoc by presburger also have "... = \\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ (t \ (u \ the_\) \ \[u, u\<^sub>0 \ \\.p\<^sub>0, chine] \ (\[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ ((\ \ \\.p\<^sub>0) \ chine) \ (t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[t\<^sub>0, \\.p\<^sub>1, chine]) \ \[t, t\<^sub>0, \\.p\<^sub>1 \ chine] \ (\ \ \\.p\<^sub>1 \ chine) \ \[t\<^sub>1, \\.p\<^sub>1, chine] \ the_\" proof - have "the_\ \ (inv \.the_\ \ \\.p\<^sub>1) \ ((\\<^sup>-\<^sup>1[t\<^sub>1, \.chine, \\.p\<^sub>1]) \ \[t\<^sub>1, \.chine, \\.p\<^sub>1]) \ (\.the_\ \ \\.p\<^sub>1) = the_\ \ (inv \.the_\ \ \\.p\<^sub>1) \ ((t\<^sub>1 \ \.chine) \ \\.p\<^sub>1) \ (\.the_\ \ \\.p\<^sub>1)" using comp_inv_arr \.T0.antipar(1) comp_assoc_assoc' by simp also have "... = the_\ \ (inv \.the_\ \ \\.p\<^sub>1) \ (\.the_\ \ \\.p\<^sub>1)" using comp_cod_arr \.T0.antipar(1) by simp also have "... = the_\ \ (r\<^sub>1 \ \\.p\<^sub>1)" using whisker_right [of \\.p\<^sub>1] r\<^sub>0s\<^sub>1.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 "... = (\\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ (t \ u \ the_\)) \ (t \ \[u, u\<^sub>0 \ \\.p\<^sub>0, chine]) \ (t \ \[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ (t \ (\ \ \\.p\<^sub>0) \ chine) \ (t \ t\<^sub>0u\<^sub>1.\ \ chine) \ (t \ \\<^sup>-\<^sup>1[t\<^sub>0, \\.p\<^sub>1, chine]) \ \[t, t\<^sub>0, \\.p\<^sub>1 \ chine] \ (\ \ \\.p\<^sub>1 \ chine) \ \[t\<^sub>1, \\.p\<^sub>1, chine] \ the_\" proof - have "arr ((u \ the_\) \ \[u, u\<^sub>0 \ \\.p\<^sub>0, chine] \ (\[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ ((\ \ \\.p\<^sub>0) \ chine) \ (t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[t\<^sub>0, \\.p\<^sub>1, chine])" using \\<^sub>\' by blast moreover have "arr (\[u, u\<^sub>0 \ \\.p\<^sub>0, chine] \ (\[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ ((\ \ \\.p\<^sub>0) \ chine) \ (t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[t\<^sub>0, \\.p\<^sub>1, chine])" using calculation by blast moreover have "arr ((\[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ ((\ \ \\.p\<^sub>0) \ chine) \ (t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[t\<^sub>0, \\.p\<^sub>1, chine])" using calculation by blast moreover have "arr (((\ \ \\.p\<^sub>0) \ chine) \ (t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[t\<^sub>0, \\.p\<^sub>1, chine])" using calculation by blast moreover have "arr ((t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[t\<^sub>0, \\.p\<^sub>1, chine])" using calculation by blast ultimately have "t \ (u \ the_\) \ \[u, u\<^sub>0 \ \\.p\<^sub>0, chine] \ (\[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ ((\ \ \\.p\<^sub>0) \ chine) \ (t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[t\<^sub>0, \\.p\<^sub>1, chine] = (t \ u \ the_\) \ (t \ \[u, u\<^sub>0 \ \\.p\<^sub>0, chine]) \ (t \ \[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ (t \ (\ \ \\.p\<^sub>0) \ chine) \ (t \ t\<^sub>0u\<^sub>1.\ \ chine) \ (t \ \\<^sup>-\<^sup>1[t\<^sub>0, \\.p\<^sub>1, chine])" using whisker_left \.T0.antipar(1) \.ide_base by presburger thus ?thesis using comp_assoc by presburger qed also have "... = ((t \ u) \ the_\) \ \\<^sup>-\<^sup>1[t, u, (u\<^sub>0 \ \\.p\<^sub>0) \ chine] \ (t \ \[u, u\<^sub>0 \ \\.p\<^sub>0, chine]) \ (t \ \[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ (t \ (\ \ \\.p\<^sub>0) \ chine) \ (t \ t\<^sub>0u\<^sub>1.\ \ chine) \ (t \ \\<^sup>-\<^sup>1[t\<^sub>0, \\.p\<^sub>1, chine]) \ \[t, t\<^sub>0, \\.p\<^sub>1 \ chine] \ ((\ \ \\.p\<^sub>1 \ chine) \ \[t\<^sub>1, \\.p\<^sub>1, chine]) \ the_\" using assoc'_naturality [of t u the_\] \\.composable \_simps(3) comp_assoc by auto also have "... = ((t \ u) \ the_\) \ \\<^sup>-\<^sup>1[t, u, (u\<^sub>0 \ \\.p\<^sub>0) \ chine] \ (t \ \[u, u\<^sub>0 \ \\.p\<^sub>0, chine]) \ (t \ \[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ (t \ (\ \ \\.p\<^sub>0) \ chine) \ (t \ t\<^sub>0u\<^sub>1.\ \ chine) \ ((t \ \\<^sup>-\<^sup>1[t\<^sub>0, \\.p\<^sub>1, chine]) \ \[t, t\<^sub>0, \\.p\<^sub>1 \ chine] \ \[t \ t\<^sub>0, \\.p\<^sub>1, chine]) \ ((\ \ \\.p\<^sub>1) \ chine) \ the_\" proof - have "(\ \ \\.p\<^sub>1 \ chine) \ \[t\<^sub>1, \\.p\<^sub>1, chine] = \[t \ t\<^sub>0, \\.p\<^sub>1, chine] \ ((\ \ \\.p\<^sub>1) \ 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 t\<^sub>0u\<^sub>1.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_\) \ \\<^sup>-\<^sup>1[t, u, (u\<^sub>0 \ \\.p\<^sub>0) \ chine] \ (t \ \[u, u\<^sub>0 \ \\.p\<^sub>0, chine]) \ (t \ \[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ (t \ (\ \ \\.p\<^sub>0) \ chine) \ ((t \ t\<^sub>0u\<^sub>1.\ \ chine) \ \[t, t\<^sub>0 \ \\.p\<^sub>1, chine]) \ (\[t, t\<^sub>0, \\.p\<^sub>1] \ chine) \ ((\ \ \\.p\<^sub>1) \ chine) \ the_\" proof - have "(t \ \\<^sup>-\<^sup>1[t\<^sub>0, \\.p\<^sub>1, chine]) \ \[t, t\<^sub>0, \\.p\<^sub>1 \ chine] \ \[t \ t\<^sub>0, \\.p\<^sub>1, chine] = \[t, t\<^sub>0 \ \\.p\<^sub>1, chine] \ (\[t, t\<^sub>0, \\.p\<^sub>1] \ chine)" using pentagon t\<^sub>0u\<^sub>1.p\<^sub>1_simps uw\ \.T0.antipar(1) iso_hcomp comp_assoc_assoc' invert_side_of_triangle(1) [of "\[t, t\<^sub>0, \\.p\<^sub>1 \ chine] \ \[t \ t\<^sub>0, \\.p\<^sub>1, chine]" "t \ \[t\<^sub>0, \\.p\<^sub>1, chine]" "\[t, t\<^sub>0 \ \\.p\<^sub>1, chine] \ (\[t, t\<^sub>0, \\.p\<^sub>1] \ chine)"] by auto thus ?thesis using comp_assoc by presburger qed also have "... = ((t \ u) \ the_\) \ \\<^sup>-\<^sup>1[t, u, (u\<^sub>0 \ \\.p\<^sub>0) \ chine] \ (t \ \[u, u\<^sub>0 \ \\.p\<^sub>0, chine]) \ (t \ \[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ ((t \ (\ \ \\.p\<^sub>0) \ chine) \ \[t, u\<^sub>1 \ \\.p\<^sub>0, chine]) \ ((t \ t\<^sub>0u\<^sub>1.\) \ chine) \ (\[t, t\<^sub>0, \\.p\<^sub>1] \ chine) \ ((\ \ \\.p\<^sub>1) \ chine) \ the_\" proof - have "(t \ t\<^sub>0u\<^sub>1.\ \ chine) \ \[t, t\<^sub>0 \ \\.p\<^sub>1, chine] = \[t, u\<^sub>1 \ \\.p\<^sub>0, chine] \ ((t \ t\<^sub>0u\<^sub>1.\) \ chine)" using assoc_naturality [of t t\<^sub>0u\<^sub>1.\ chine] t\<^sub>0u\<^sub>1.cospan by auto thus ?thesis using comp_assoc by presburger qed also have "... = ((t \ u) \ the_\) \ \\<^sup>-\<^sup>1[t, u, (u\<^sub>0 \ \\.p\<^sub>0) \ chine] \ (t \ \[u, u\<^sub>0 \ \\.p\<^sub>0, chine]) \ (t \ \[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ \[t, (u \ u\<^sub>0) \ \\.p\<^sub>0, chine] \ ((t \ \ \ \\.p\<^sub>0) \ chine) \ ((t \ t\<^sub>0u\<^sub>1.\) \ chine) \ (\[t, t\<^sub>0, \\.p\<^sub>1] \ chine) \ ((\ \ \\.p\<^sub>1) \ chine) \ the_\" proof - have "(t \ (\ \ \\.p\<^sub>0) \ chine) \ \[t, u\<^sub>1 \ \\.p\<^sub>0, chine] = \[t, (u \ u\<^sub>0) \ \\.p\<^sub>0, chine] \ ((t \ (\ \ \\.p\<^sub>0)) \ chine)" using assoc_naturality [of t "\ \ \\.p\<^sub>0" chine] by (simp add: \\.composable) thus ?thesis using comp_assoc by presburger qed also have "... = ((t \ u) \ the_\) \ \\<^sup>-\<^sup>1[t, u, (u\<^sub>0 \ \\.p\<^sub>0) \ chine] \ (t \ \[u, u\<^sub>0 \ \\.p\<^sub>0, chine]) \ (t \ \[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ \[t, (u \ u\<^sub>0) \ \\.p\<^sub>0, chine] \ ((t \ \\<^sup>-\<^sup>1[u, u\<^sub>0, \\.p\<^sub>0]) \ chine) \ ((t \ \[u, u\<^sub>0, \\.p\<^sub>0]) \ chine) \ ((t \ \ \ \\.p\<^sub>0) \ chine) \ ((t \ t\<^sub>0u\<^sub>1.\) \ chine) \ (\[t, t\<^sub>0, \\.p\<^sub>1] \ chine) \ ((\ \ \\.p\<^sub>1) \ chine) \ the_\" proof - have "(((t \ \\<^sup>-\<^sup>1[u, u\<^sub>0, \\.p\<^sub>0]) \ chine) \ ((t \ \[u, u\<^sub>0, \\.p\<^sub>0]) \ chine)) \ ((t \ \ \ \\.p\<^sub>0) \ chine) = ((t \ ((u \ u\<^sub>0) \ \\.p\<^sub>0)) \ chine) \ ((t \ \ \ \\.p\<^sub>0) \ chine)" using whisker_right whisker_left [of t "\\<^sup>-\<^sup>1[u, u\<^sub>0, \\.p\<^sub>0]" "\[u, u\<^sub>0, \\.p\<^sub>0]"] \\.composable comp_assoc_assoc' by simp also have "... = (t \ \ \ \\.p\<^sub>0) \ chine" using comp_cod_arr \\.composable by simp finally have "(((t \ \\<^sup>-\<^sup>1[u, u\<^sub>0, \\.p\<^sub>0]) \ chine) \ ((t \ \[u, u\<^sub>0, \\.p\<^sub>0]) \ chine)) \ ((t \ \ \ \\.p\<^sub>0) \ chine) = (t \ \ \ \\.p\<^sub>0) \ chine" by simp thus ?thesis using comp_assoc by metis qed also have "... = ((t \ u) \ the_\) \ \\<^sup>-\<^sup>1[t, u, (u\<^sub>0 \ \\.p\<^sub>0) \ chine] \ (t \ \[u, u\<^sub>0 \ \\.p\<^sub>0, chine]) \ (t \ \[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ \[t, (u \ u\<^sub>0) \ \\.p\<^sub>0, chine] \ ((t \ \\<^sup>-\<^sup>1[u, u\<^sub>0, \\.p\<^sub>0]) \ chine) \ (((\[t, u, u\<^sub>0 \ \\.p\<^sub>0] \ chine) \ (\\<^sup>-\<^sup>1[t, u, u\<^sub>0 \ \\.p\<^sub>0] \ chine)) \ ((t \ \[u, u\<^sub>0, \\.p\<^sub>0]) \ chine)) \ ((t \ \ \ \\.p\<^sub>0) \ chine) \ ((t \ t\<^sub>0u\<^sub>1.\) \ chine) \ (\[t, t\<^sub>0, \\.p\<^sub>1] \ chine) \ ((\ \ \\.p\<^sub>1) \ chine) \ the_\" proof - have "((\[t, u, u\<^sub>0 \ \\.p\<^sub>0] \ chine) \ (\\<^sup>-\<^sup>1[t, u, u\<^sub>0 \ \\.p\<^sub>0] \ chine)) \ ((t \ \[u, u\<^sub>0, \\.p\<^sub>0]) \ chine) = ((t \ \[u, u\<^sub>0, \\.p\<^sub>0]) \ chine)" using comp_inv_arr' comp_cod_arr \\.composable comp_assoc_assoc' whisker_right [of chine "\[t, u, u\<^sub>0 \ \\.p\<^sub>0]" "\\<^sup>-\<^sup>1[t, u, u\<^sub>0 \ \\.p\<^sub>0]"] by simp thus ?thesis using comp_assoc by simp qed also have "... = ((t \ u) \ the_\) \ \\<^sup>-\<^sup>1[t, u, (u\<^sub>0 \ \\.p\<^sub>0) \ chine] \ (t \ \[u, u\<^sub>0 \ \\.p\<^sub>0, chine]) \ (t \ \[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ \[t, (u \ u\<^sub>0) \ \\.p\<^sub>0, chine] \ ((t \ \\<^sup>-\<^sup>1[u, u\<^sub>0, \\.p\<^sub>0]) \ chine) \ (\[t, u, u\<^sub>0 \ \\.p\<^sub>0] \ chine) \ ((\\<^sup>-\<^sup>1[t, u, u\<^sub>0 \ \\.p\<^sub>0] \ chine) \ ((t \ \[u, u\<^sub>0, \\.p\<^sub>0]) \ chine) \ ((t \ \ \ \\.p\<^sub>0) \ chine) \ ((t \ t\<^sub>0u\<^sub>1.\) \ chine) \ (\[t, t\<^sub>0, \\.p\<^sub>1] \ chine) \ ((\ \ \\.p\<^sub>1) \ chine)) \ the_\" using comp_assoc by presburger also have "... = ((t \ u) \ the_\) \ (\\<^sup>-\<^sup>1[t, u, (u\<^sub>0 \ \\.p\<^sub>0) \ chine] \ (t \ \[u, u\<^sub>0 \ \\.p\<^sub>0, chine]) \ (t \ \[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ \[t, (u \ u\<^sub>0) \ \\.p\<^sub>0, chine] \ ((t \ \\<^sup>-\<^sup>1[u, u\<^sub>0, \\.p\<^sub>0]) \ chine) \ (\[t, u, u\<^sub>0 \ \\.p\<^sub>0] \ chine)) \ (\\.tab \ chine) \ the_\" proof - have "(\\<^sup>-\<^sup>1[t, u, u\<^sub>0 \ \\.p\<^sub>0] \ chine) \ ((t \ \[u, u\<^sub>0, \\.p\<^sub>0]) \ chine) \ ((t \ \ \ \\.p\<^sub>0) \ chine) \ ((t \ t\<^sub>0u\<^sub>1.\) \ chine) \ (\[t, t\<^sub>0, \\.p\<^sub>1] \ chine) \ ((\ \ \\.p\<^sub>1) \ 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, u\<^sub>0 \ \\.p\<^sub>0, chine] \ (\\.tab \ chine) \ the_\" proof - have "\\<^sup>-\<^sup>1[t, u, (u\<^sub>0 \ \\.p\<^sub>0) \ chine] \ (t \ \[u, u\<^sub>0 \ \\.p\<^sub>0, chine]) \ (t \ \[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ \[t, (u \ u\<^sub>0) \ \\.p\<^sub>0, chine] \ ((t \ \\<^sup>-\<^sup>1[u, u\<^sub>0, \\.p\<^sub>0]) \ chine) \ (\[t, u, u\<^sub>0 \ \\.p\<^sub>0] \ chine) = \\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\t\<^bold>\, \<^bold>\u\<^bold>\, (\<^bold>\u\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\\\.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\chine\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\t\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\u\<^bold>\, \<^bold>\u\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\\\.p\<^sub>0\<^bold>\, \<^bold>\chine\<^bold>\\<^bold>]) \<^bold>\ (\<^bold>\t\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\u\<^bold>\, \<^bold>\u\<^sub>0\<^bold>\, \<^bold>\\\.p\<^sub>0\<^bold>\\<^bold>] \<^bold>\ \<^bold>\chine\<^bold>\) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\t\<^bold>\, (\<^bold>\u\<^bold>\ \<^bold>\ \<^bold>\u\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\\\.p\<^sub>0\<^bold>\, \<^bold>\chine\<^bold>\\<^bold>] \<^bold>\ ((\<^bold>\t\<^bold>\ \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\u\<^bold>\, \<^bold>\u\<^sub>0\<^bold>\, \<^bold>\\\.p\<^sub>0\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\chine\<^bold>\) \<^bold>\ (\<^bold>\\<^bold>[\<^bold>\t\<^bold>\, \<^bold>\u\<^bold>\, \<^bold>\u\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\\\.p\<^sub>0\<^bold>\\<^bold>] \<^bold>\ \<^bold>\chine\<^bold>\)\" using \'_def \_def \\.composable by simp also have "... = \\<^bold>\\<^bold>[\<^bold>\t\<^bold>\ \<^bold>\ \<^bold>\u\<^bold>\, \<^bold>\u\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\\\.p\<^sub>0\<^bold>\, \<^bold>\chine\<^bold>\\<^bold>]\" using \\.composable apply (intro E.eval_eqI) by simp_all also have "... = \[t \ u, u\<^sub>0 \ \\.p\<^sub>0, 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 "\\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ LHS \ \[t\<^sub>1, \.chine, \\.p\<^sub>1] \ (\.the_\ \ \\.p\<^sub>1) = \\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ RHS \ \[t\<^sub>1, \.chine, \\.p\<^sub>1] \ (\.the_\ \ \\.p\<^sub>1)" by blast (* * TODO: This is common enough that there should be "cancel_iso_left" and * "cancel_iso_right" rules for doing it. *) hence "(LHS \ \[t\<^sub>1, \.chine, \\.p\<^sub>1]) \ (\.the_\ \ \\.p\<^sub>1) = (RHS \ \[t\<^sub>1, \.chine, \\.p\<^sub>1]) \ (\.the_\ \ \\.p\<^sub>1)" using u\<^sub>\ r\<^sub>0s\<^sub>1.ide_u LHS RHS iso_is_section [of "\\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0]"] section_is_mono monoE \\.composable comp_assoc by (metis (no_types, lifting) \_simps(1) \.ide_base \\\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ r\<^sub>0s\<^sub>1.p\<^sub>0] \ LHS \ \[t\<^sub>1, \.chine, r\<^sub>0s\<^sub>1.p\<^sub>1] \ (\.the_\ \ r\<^sub>0s\<^sub>1.p\<^sub>1) = ((\ \ \) \ s\<^sub>0 \ r\<^sub>0s\<^sub>1.p\<^sub>0) \ \\.tab\ \.ide_base hseq_char ideD(1) ide_u iso_assoc') hence 1: "LHS \ \[t\<^sub>1, \.chine, \\.p\<^sub>1] = RHS \ \[t\<^sub>1, \.chine, \\.p\<^sub>1]" using epiE LHS RHS iso_is_retraction retraction_is_epi \\.composable \.the_\_props iso_hcomp by (metis \_simps(1) \.the_\_simps(2) \((\ \ \) \ s\<^sub>0 \ r\<^sub>0s\<^sub>1.p\<^sub>0) \ \\.tab = \\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ r\<^sub>0s\<^sub>1.p\<^sub>0] \ RHS \ \[t\<^sub>1, \.chine, r\<^sub>0s\<^sub>1.p\<^sub>1] \ (\.the_\ \ r\<^sub>0s\<^sub>1.p\<^sub>1)\ \.leg1_simps(3) ide_is_iso local.comp_assoc r\<^sub>0s\<^sub>1.ide_leg1 r\<^sub>0s\<^sub>1.p\<^sub>1_simps seqE) show "LHS = RHS" proof - have "epi \[t\<^sub>1, \.chine, \\.p\<^sub>1]" using iso_is_retraction retraction_is_epi \.T0.antipar(1) by simp moreover have "seq LHS \[t\<^sub>1, \.chine, \\.p\<^sub>1]" using LHS \.T0.antipar(1) by auto moreover have "seq RHS \[t\<^sub>1, \.chine, \\.p\<^sub>1]" using RHS \.T0.antipar(1) by auto ultimately show ?thesis using epiE 1 by blast qed qed have 1: "\!\. \\ : ?w\<^sub>\ \ ?w\<^sub>\'\ \ ?\\<^sub>\ = t\<^sub>1 \ \ \ ?\\<^sub>\ = ?\\<^sub>\' \ (t\<^sub>0 \ \)" using LHS_def RHS_def u\<^sub>\ w\<^sub>\ w\<^sub>\' \\<^sub>\ \\<^sub>\ \\<^sub>\' eq \.T2 [of ?w\<^sub>\ ?w\<^sub>\' ?\\<^sub>\ ?u\<^sub>\ ?\\<^sub>\' ?\\<^sub>\] by fastforce obtain \\<^sub>\ where \\<^sub>\: "\\\<^sub>\ : ?w\<^sub>\ \ ?w\<^sub>\'\ \ ?\\<^sub>\ = t\<^sub>1 \ \\<^sub>\ \ ?\\<^sub>\ = ?\\<^sub>\' \ (t\<^sub>0 \ \\<^sub>\)" using 1 by auto text \ At this point we could show that @{term \\<^sub>\} is invertible using \BS3\, but we want to avoid using \BS3\ if possible and we also want to establish a characterization of @{term "inv \\<^sub>\"}. So we show the invertibility of @{term \\<^sub>\} directly, using a few more applications of \T2\. \ have iso_\\<^sub>\: "iso ?\\<^sub>\" using uw\ \\<^sub>\ the_\_props \.the_\_props iso_inv_iso hseqI' iso_assoc' \.hseq_leg\<^sub>0 iso_inv_iso iso_hcomp apply (intro isos_compose) apply (metis \.is_ide \\.leg1_simps(2) \.ide_leg1 \.leg1_simps(2) \.leg1_simps(3) hseqE r\<^sub>0s\<^sub>1.ide_leg1 hcomp_simps(1) vconn_implies_hpar(3)) apply (metis \\.leg1_simps(2) hseqE ide_is_iso r\<^sub>0s\<^sub>1.ide_leg1 src_inv vconn_implies_hpar(1)) apply blast apply blast apply blast apply (metis \.ide_leg1 \.leg1_simps(3) hseqE ide_char iso_assoc t\<^sub>0u\<^sub>1.ide_leg1 t\<^sub>0u\<^sub>1.p\<^sub>1_simps w\<^sub>\') by blast hence eq': "((t \ ?\\<^sub>\') \ \[t, t\<^sub>0, ?w\<^sub>\'] \ (\ \ ?w\<^sub>\')) = ((t \ ?\\<^sub>\) \ \[t, t\<^sub>0, ?w\<^sub>\] \ (\ \ ?w\<^sub>\)) \ inv ?\\<^sub>\" proof - have "seq ((t \ ?\\<^sub>\') \ \[t, t\<^sub>0, ?w\<^sub>\'] \ (\ \ ?w\<^sub>\')) ?\\<^sub>\" using LHS RHS_def eq by blast hence "(t \ ?\\<^sub>\') \ \[t, t\<^sub>0, ?w\<^sub>\'] \ (\ \ ?w\<^sub>\') = (((t \ ?\\<^sub>\') \ \[t, t\<^sub>0, ?w\<^sub>\'] \ (\ \ ?w\<^sub>\')) \ ?\\<^sub>\) \ inv ?\\<^sub>\" by (meson invert_side_of_triangle(2) iso_\\<^sub>\) thus ?thesis using LHS_def RHS_def eq by argo qed have 2: "\!\. \\ : ?w\<^sub>\' \ ?w\<^sub>\\ \ inv ?\\<^sub>\ = t\<^sub>1 \ \ \ ?\\<^sub>\' = ?\\<^sub>\ \ (t\<^sub>0 \ \)" using u\<^sub>\ w\<^sub>\ w\<^sub>\' \\<^sub>\ \\<^sub>\ \\<^sub>\' eq' \.T2 [of ?w\<^sub>\' ?w\<^sub>\ ?\\<^sub>\'?u\<^sub>\ ?\\<^sub>\ "inv ?\\<^sub>\"] iso_\\<^sub>\ comp_assoc by blast obtain \\<^sub>\' where \\<^sub>\': "\\\<^sub>\' : ?w\<^sub>\' \ ?w\<^sub>\\ \ inv ?\\<^sub>\ = t\<^sub>1 \ \\<^sub>\' \ ?\\<^sub>\' = ?\\<^sub>\ \ (t\<^sub>0 \ \\<^sub>\')" using 2 by auto have "inverse_arrows \\<^sub>\ \\<^sub>\'" proof have "\\\<^sub>\' \ \\<^sub>\ : ?w\<^sub>\ \ ?w\<^sub>\\" using \\<^sub>\ \\<^sub>\' by auto moreover have "t\<^sub>1 \ \\<^sub>\' \ \\<^sub>\ = t\<^sub>1 \ ?w\<^sub>\" using \\<^sub>\ \\<^sub>\' whisker_left \\<^sub>\ iso_\\<^sub>\ comp_inv_arr' by (metis (no_types, lifting) \.ide_leg1 calculation in_homE) moreover have "?\\<^sub>\ = ?\\<^sub>\ \ (t\<^sub>0 \ \\<^sub>\' \ \\<^sub>\)" proof - have "?\\<^sub>\ = ?\\<^sub>\' \ (t\<^sub>0 \ \\<^sub>\)" using \\<^sub>\ by simp also have "... = ?\\<^sub>\ \ (t\<^sub>0 \ \\<^sub>\') \ (t\<^sub>0 \ \\<^sub>\)" using \\<^sub>\' comp_assoc by simp also have "... = ?\\<^sub>\ \ (t\<^sub>0 \ \\<^sub>\' \ \\<^sub>\)" using \\<^sub>\ \\<^sub>\' whisker_left by (metis (full_types) \.ide_leg0 seqI') finally show ?thesis by simp qed moreover have "\\. \\ : ?w\<^sub>\ \ ?w\<^sub>\\ \ t\<^sub>1 \ \ = t\<^sub>1 \ ?w\<^sub>\ \ ?\\<^sub>\ = ?\\<^sub>\ \ (t\<^sub>0 \ \) \ \ = ?w\<^sub>\" proof - have "\!\. \\ : ?w\<^sub>\ \ ?w\<^sub>\\ \ t\<^sub>1 \ ?w\<^sub>\ = t\<^sub>1 \ \ \ ?\\<^sub>\ = ?\\<^sub>\ \ (t\<^sub>0 \ \)" proof - have "(t \ ?\\<^sub>\) \ \[t, t\<^sub>0, ?w\<^sub>\] \ (\ \ ?w\<^sub>\) = ((t \ ?\\<^sub>\) \ \[t, t\<^sub>0, ?w\<^sub>\] \ (\ \ ?w\<^sub>\)) \ (t\<^sub>1 \ ?w\<^sub>\)" by (metis LHS LHS_def comp_arr_dom in_homE) thus ?thesis using w\<^sub>\ \\<^sub>\ \.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\<^sub>\ \ ?w\<^sub>\\ \ t\<^sub>1 \ \ = t\<^sub>1 \ ?w\<^sub>\ \ ?\\<^sub>\ = ?\\<^sub>\ \ (t\<^sub>0 \ \) \ \ = ?w\<^sub>\" by (metis \\<^sub>\ comp_arr_dom ide_in_hom(2) in_homE w\<^sub>\) qed ultimately have "\\<^sub>\' \ \\<^sub>\ = ?w\<^sub>\" by simp thus "ide (\\<^sub>\' \ \\<^sub>\)" using w\<^sub>\ by simp have "\\\<^sub>\ \ \\<^sub>\' : ?w\<^sub>\' \ ?w\<^sub>\'\" using \\<^sub>\ \\<^sub>\' by auto moreover have "t\<^sub>1 \ \\<^sub>\ \ \\<^sub>\' = t\<^sub>1 \ ?w\<^sub>\'" by (metis \\<^sub>\ \\<^sub>\ \\<^sub>\' \.ide_leg1 calculation comp_arr_inv' in_homE iso_\\<^sub>\ whisker_left) moreover have "?\\<^sub>\' = ?\\<^sub>\' \ (t\<^sub>0 \ \\<^sub>\ \ \\<^sub>\')" proof - have "?\\<^sub>\' = ?\\<^sub>\ \ (t\<^sub>0 \ \\<^sub>\')" using \\<^sub>\' by simp also have "... = ?\\<^sub>\' \ (t\<^sub>0 \ \\<^sub>\) \ (t\<^sub>0 \ \\<^sub>\')" using \\<^sub>\ comp_assoc by simp also have "... = ?\\<^sub>\' \ (t\<^sub>0 \ \\<^sub>\ \ \\<^sub>\')" using \\<^sub>\ \\<^sub>\' whisker_left by (metis (full_types) \.ide_leg0 seqI') finally show ?thesis by simp qed moreover have "\\. \\ : ?w\<^sub>\' \ ?w\<^sub>\'\ \ t\<^sub>1 \ \ = t\<^sub>1 \ ?w\<^sub>\' \ ?\\<^sub>\' = ?\\<^sub>\' \ (t\<^sub>0 \ \) \ \ = ?w\<^sub>\'" proof - have "\!\. \\ : ?w\<^sub>\' \ ?w\<^sub>\'\ \ t\<^sub>1 \ ?w\<^sub>\' = t\<^sub>1 \ \ \ ?\\<^sub>\' = ?\\<^sub>\' \ (t\<^sub>0 \ \)" proof - have "(t \ ?\\<^sub>\') \ \[t, t\<^sub>0, ?w\<^sub>\'] \ (\ \ ?w\<^sub>\') = ((t \ ?\\<^sub>\') \ \[t, t\<^sub>0, ?w\<^sub>\'] \ (\ \ ?w\<^sub>\')) \ (t\<^sub>1 \ ?w\<^sub>\')" proof - have 1: "t\<^sub>1 \ \\<^sub>\ \ \\<^sub>\' = (t\<^sub>1 \ \\<^sub>\) \ (t\<^sub>1 \ \\<^sub>\')" by (meson \\<^sub>\ \\<^sub>\' \.ide_leg1 seqI' whisker_left) have "((LHS \ inv ?\\<^sub>\) \ (t\<^sub>1 \ \\<^sub>\)) \ (t\<^sub>1 \ \\<^sub>\') = LHS \ inv ?\\<^sub>\" using LHS_def RHS_def \\<^sub>\ \\<^sub>\' eq eq' by argo thus ?thesis unfolding LHS_def using 1 by (simp add: calculation(2) eq' comp_assoc) qed thus ?thesis using w\<^sub>\' \\<^sub>\' \.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) t\<^sub>0u\<^sub>1.base_simps(2) t\<^sub>0u\<^sub>1.leg1_simps(4) by presburger qed thus "\\. \\ : ?w\<^sub>\' \ ?w\<^sub>\'\ \ t\<^sub>1 \ \ = t\<^sub>1 \ ?w\<^sub>\' \ ?\\<^sub>\' = ?\\<^sub>\' \ (t\<^sub>0 \ \) \ \ = ?w\<^sub>\'" by (metis \\<^sub>\' comp_arr_dom ide_in_hom(2) in_homE w\<^sub>\') qed ultimately have "\\<^sub>\ \ \\<^sub>\' = ?w\<^sub>\'" by simp thus "ide (\\<^sub>\ \ \\<^sub>\')" using w\<^sub>\' by simp qed thus "\\.p\<^sub>1 \ chine \ \.chine \ \\.p\<^sub>1" using w\<^sub>\ w\<^sub>\' \\<^sub>\ isomorphic_symmetric isomorphic_def by blast have iso_\\<^sub>\: "iso \\<^sub>\" using \inverse_arrows \\<^sub>\ \\<^sub>\'\ by auto have \\<^sub>\'_eq: "\\<^sub>\' = inv \\<^sub>\" using \inverse_arrows \\<^sub>\ \\<^sub>\'\ inverse_unique by blast let ?w\<^sub>\ = "\\.p\<^sub>0 \ chine" let ?w\<^sub>\' = "\.chine \ \\.p\<^sub>0" let ?u\<^sub>\ = "s\<^sub>0 \ \\.p\<^sub>0" let ?\\<^sub>\ = "the_\ \ \\<^sup>-\<^sup>1[u\<^sub>0, \\.p\<^sub>0, chine]" let ?\\<^sub>\' = "(\.the_\ \ \\.p\<^sub>0) \ \\<^sup>-\<^sup>1[u\<^sub>0, \.chine, \\.p\<^sub>0]" let ?\\<^sub>\ = "\[u\<^sub>1, \.chine, \\.p\<^sub>0] \ (\.the_\ \ \\.p\<^sub>0) \ r\<^sub>0s\<^sub>1.\ \ (\.the_\ \ \\.p\<^sub>1) \ \\<^sup>-\<^sup>1[t\<^sub>0, \.chine, \\.p\<^sub>1] \ (t\<^sub>0 \ inv \\<^sub>\) \ \[t\<^sub>0, \\.p\<^sub>1, chine] \ (inv t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine]" have w\<^sub>\: "ide ?w\<^sub>\ \ is_left_adjoint ?w\<^sub>\" using is_map left_adjoints_compose by simp have w\<^sub>\': "ide ?w\<^sub>\' \ is_left_adjoint ?w\<^sub>\'" using \.is_map left_adjoints_compose by (simp add: is_map left_adjoints_compose) have 1: "\!\. \\ : ?w\<^sub>\ \ ?w\<^sub>\'\ \ ?\\<^sub>\ = u\<^sub>1 \ \ \ ?\\<^sub>\ = ?\\<^sub>\' \ (u\<^sub>0 \ \)" proof - have \\<^sub>\: "\?\\<^sub>\ : u\<^sub>0 \ ?w\<^sub>\ \ ?u\<^sub>\\" by auto have \\<^sub>\': "\?\\<^sub>\' : u\<^sub>0 \ ?w\<^sub>\' \ ?u\<^sub>\\" by fastforce have \\<^sub>\: "\?\\<^sub>\ : u\<^sub>1 \ ?w\<^sub>\ \ u\<^sub>1 \ ?w\<^sub>\'\" proof (intro comp_in_homI) show "\\\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine] : u\<^sub>1 \ \\.p\<^sub>0 \ chine \ (u\<^sub>1 \ \\.p\<^sub>0) \ chine\" by auto show "\inv t\<^sub>0u\<^sub>1.\ \ chine : (u\<^sub>1 \ \\.p\<^sub>0) \ chine \ (t\<^sub>0 \ \\.p\<^sub>1) \ chine\" using t\<^sub>0u\<^sub>1.\_in_hom(2) t\<^sub>0u\<^sub>1.\_uniqueness(2) by auto show "\\[t\<^sub>0, \\.p\<^sub>1, chine] : (t\<^sub>0 \ \\.p\<^sub>1) \ chine \ t\<^sub>0 \ \\.p\<^sub>1 \ chine\" using \.T0.antipar(1) by auto show "\t\<^sub>0 \ inv \\<^sub>\ : t\<^sub>0 \ \\.p\<^sub>1 \ chine \ t\<^sub>0 \ \.chine \ \\.p\<^sub>1\" using \\<^sub>\ iso_\\<^sub>\ using \.T0.antipar(1) by auto show "\\\<^sup>-\<^sup>1[t\<^sub>0, \.chine, \\.p\<^sub>1] : t\<^sub>0 \ \.chine \ \\.p\<^sub>1 \ (t\<^sub>0 \ \.chine) \ \\.p\<^sub>1\" using \.T0.antipar(1) by auto show "\\.the_\ \ \\.p\<^sub>1 : (t\<^sub>0 \ \.chine) \ \\.p\<^sub>1 \ r\<^sub>0 \ \\.p\<^sub>1\" using \.T0.antipar(1) by auto show "\r\<^sub>0s\<^sub>1.\ : r\<^sub>0 \ \\.p\<^sub>1 \ s\<^sub>1 \ \\.p\<^sub>0\" by auto show "\\.the_\ \ \\.p\<^sub>0 : s\<^sub>1 \ \\.p\<^sub>0 \ (u\<^sub>1 \ \.chine) \ \\.p\<^sub>0\" by auto show "\\[u\<^sub>1, \.chine, \\.p\<^sub>0] : (u\<^sub>1 \ \.chine) \ \\.p\<^sub>0 \ u\<^sub>1 \ \.chine \ \\.p\<^sub>0\" by auto qed text \ The proof of the equation below needs to make use of the equation \\\<^sub>\' = \\<^sub>\ \ (t\<^sub>0 \ \\<^sub>\')\ from the previous section. So the overall strategy is to work toward an expression of the form \\\<^sub>\ \ (t\<^sub>0 \ \\<^sub>\')\ and perform the replacement to eliminate \\\<^sub>\'\. \ have eq\<^sub>\: "(u \ ?\\<^sub>\) \ \[u, u\<^sub>0, ?w\<^sub>\] \ (\ \ ?w\<^sub>\) = ((u \ ?\\<^sub>\') \ \[u, u\<^sub>0, ?w\<^sub>\'] \ (\ \ ?w\<^sub>\')) \ ?\\<^sub>\" proof - let ?LHS = "(u \ ?\\<^sub>\) \ \[u, u\<^sub>0, ?w\<^sub>\] \ (\ \ ?w\<^sub>\)" let ?RHS = "((u \ ?\\<^sub>\') \ \[u, u\<^sub>0, ?w\<^sub>\'] \ (\ \ ?w\<^sub>\')) \ ?\\<^sub>\" have "?RHS = (u \ (\.the_\ \ \\.p\<^sub>0) \ \\<^sup>-\<^sup>1[u\<^sub>0, \.chine, \\.p\<^sub>0]) \ \[u, u\<^sub>0, \.chine \ \\.p\<^sub>0] \ (\ \ \.chine \ \\.p\<^sub>0) \ \[u\<^sub>1, \.chine, \\.p\<^sub>0] \ (\.the_\ \ \\.p\<^sub>0) \ r\<^sub>0s\<^sub>1.\ \ (\.the_\ \ \\.p\<^sub>1) \ \\<^sup>-\<^sup>1[t\<^sub>0, \.chine, \\.p\<^sub>1] \ (t\<^sub>0 \ inv \\<^sub>\) \ \[t\<^sub>0, \\.p\<^sub>1, chine] \ (inv t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine]" using comp_assoc by presburger also have "... = (u \ \.the_\ \ \\.p\<^sub>0) \ ((u \ \\<^sup>-\<^sup>1[u\<^sub>0, \.chine, \\.p\<^sub>0]) \ \[u, u\<^sub>0, \.chine \ \\.p\<^sub>0]) \ (\ \ \.chine \ \\.p\<^sub>0) \ \[u\<^sub>1, \.chine, \\.p\<^sub>0] \ (\.the_\ \ \\.p\<^sub>0) \ r\<^sub>0s\<^sub>1.\ \ (\.the_\ \ \\.p\<^sub>1) \ \\<^sup>-\<^sup>1[t\<^sub>0, \.chine, \\.p\<^sub>1] \ (t\<^sub>0 \ inv \\<^sub>\) \ \[t\<^sub>0, \\.p\<^sub>1, chine] \ (inv t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine]" proof - have "u \ (\.the_\ \ \\.p\<^sub>0) \ \\<^sup>-\<^sup>1[u\<^sub>0, \.chine, \\.p\<^sub>0] = (u \ \.the_\ \ \\.p\<^sub>0) \ (u \ \\<^sup>-\<^sup>1[u\<^sub>0, \.chine, \\.p\<^sub>0])" using whisker_left \.ide_base \\<^sub>\' by blast thus ?thesis using comp_assoc by presburger qed also have "... = ((u \ \.the_\ \ \\.p\<^sub>0) \ \[u, u\<^sub>0 \ \.chine, \\.p\<^sub>0]) \ (\[u, u\<^sub>0, \.chine] \ \\.p\<^sub>0) \ \\<^sup>-\<^sup>1[u \ u\<^sub>0, \.chine, \\.p\<^sub>0] \ (\ \ \.chine \ \\.p\<^sub>0) \ \[u\<^sub>1, \.chine, \\.p\<^sub>0] \ (\.the_\ \ \\.p\<^sub>0) \ r\<^sub>0s\<^sub>1.\ \ (\.the_\ \ \\.p\<^sub>1) \ \\<^sup>-\<^sup>1[t\<^sub>0, \.chine, \\.p\<^sub>1] \ (t\<^sub>0 \ inv \\<^sub>\) \ \[t\<^sub>0, \\.p\<^sub>1, chine] \ (inv t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine]" proof - have "seq (u \ \[u\<^sub>0, \.chine, \\.p\<^sub>0]) (\[u, u\<^sub>0 \ \.chine, \\.p\<^sub>0] \ (\[u, u\<^sub>0, \.chine] \ \\.p\<^sub>0))" by auto moreover have "src u = trg \[u\<^sub>0, \.chine, \\.p\<^sub>0]" by simp moreover have "inv (u \ \[u\<^sub>0, \.chine, \\.p\<^sub>0]) = u \ \\<^sup>-\<^sup>1[u\<^sub>0, \.chine, \\.p\<^sub>0]" by simp moreover have "iso (u \ \[u\<^sub>0, \.chine, \\.p\<^sub>0])" by simp moreover have "iso \[u \ u\<^sub>0, \.chine, \\.p\<^sub>0]" by simp ultimately have "(u \ \\<^sup>-\<^sup>1[u\<^sub>0, \.chine, \\.p\<^sub>0]) \ \[u, u\<^sub>0, \.chine \ \\.p\<^sub>0] = \[u, u\<^sub>0 \ \.chine, \\.p\<^sub>0] \ (\[u, u\<^sub>0, \.chine] \ \\.p\<^sub>0) \ \\<^sup>-\<^sup>1[u \ u\<^sub>0, \.chine, \\.p\<^sub>0]" using pentagon hseqI' comp_assoc invert_opposite_sides_of_square [of "u \ \[u\<^sub>0, \.chine, \\.p\<^sub>0]" "\[u, u\<^sub>0 \ \.chine, \\.p\<^sub>0] \ (\[u, u\<^sub>0, \.chine] \ \\.p\<^sub>0)" "\[u, u\<^sub>0, \.chine \ \\.p\<^sub>0]" "\[u \ u\<^sub>0, \.chine, \\.p\<^sub>0]"] 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 r\<^sub>0s\<^sub>1.ide_u r\<^sub>0s\<^sub>1.p\<^sub>0_simps hcomp_simps(1) by presburger thus ?thesis using comp_assoc by presburger qed also have "... = \[u, s\<^sub>0, \\.p\<^sub>0] \ ((u \ \.the_\) \ \\.p\<^sub>0) \ (\[u, u\<^sub>0, \.chine] \ \\.p\<^sub>0) \ (\\<^sup>-\<^sup>1[u \ u\<^sub>0, \.chine, \\.p\<^sub>0] \ (\ \ \.chine \ \\.p\<^sub>0)) \ \[u\<^sub>1, \.chine, \\.p\<^sub>0] \ (\.the_\ \ \\.p\<^sub>0) \ r\<^sub>0s\<^sub>1.\ \ (\.the_\ \ \\.p\<^sub>1) \ \\<^sup>-\<^sup>1[t\<^sub>0, \.chine, \\.p\<^sub>1] \ (t\<^sub>0 \ inv \\<^sub>\) \ \[t\<^sub>0, \\.p\<^sub>1, chine] \ (inv t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine]" proof - have "(u \ \.the_\ \ \\.p\<^sub>0) \ \[u, u\<^sub>0 \ \.chine, \\.p\<^sub>0] = \[u, s\<^sub>0, \\.p\<^sub>0] \ ((u \ \.the_\) \ \\.p\<^sub>0)" using assoc_naturality [of u \.the_\ \\.p\<^sub>0] \.\_simps(3) by auto thus ?thesis using comp_assoc by presburger qed also have "... = \[u, s\<^sub>0, \\.p\<^sub>0] \ ((u \ \.the_\) \ \\.p\<^sub>0) \ (\[u, u\<^sub>0, \.chine] \ \\.p\<^sub>0) \ ((\ \ \.chine) \ \\.p\<^sub>0) \ \\<^sup>-\<^sup>1[u\<^sub>1, \.chine, \\.p\<^sub>0] \ \[u\<^sub>1, \.chine, \\.p\<^sub>0] \ (\.the_\ \ \\.p\<^sub>0) \ r\<^sub>0s\<^sub>1.\ \ (\.the_\ \ \\.p\<^sub>1) \ \\<^sup>-\<^sup>1[t\<^sub>0, \.chine, \\.p\<^sub>1] \ (t\<^sub>0 \ inv \\<^sub>\) \ \[t\<^sub>0, \\.p\<^sub>1, chine] \ (inv t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine]" proof - have "\\<^sup>-\<^sup>1[u \ u\<^sub>0, \.chine, \\.p\<^sub>0] \ (\ \ \.chine \ \\.p\<^sub>0) = ((\ \ \.chine) \ \\.p\<^sub>0) \ \\<^sup>-\<^sup>1[u\<^sub>1, \.chine, \\.p\<^sub>0]" using assoc'_naturality [of \ \.chine \\.p\<^sub>0] by simp thus ?thesis using comp_assoc by presburger qed also have "... = \[u, s\<^sub>0, \\.p\<^sub>0] \ ((u \ \.the_\) \ \\.p\<^sub>0) \ (\[u, u\<^sub>0, \.chine] \ \\.p\<^sub>0) \ ((\ \ \.chine) \ \\.p\<^sub>0) \ ((\\<^sup>-\<^sup>1[u\<^sub>1, \.chine, \\.p\<^sub>0] \ \[u\<^sub>1, \.chine, \\.p\<^sub>0]) \ (\.the_\ \ \\.p\<^sub>0)) \ r\<^sub>0s\<^sub>1.\ \ (\.the_\ \ \\.p\<^sub>1) \ \\<^sup>-\<^sup>1[t\<^sub>0, \.chine, \\.p\<^sub>1] \ (t\<^sub>0 \ inv \\<^sub>\) \ \[t\<^sub>0, \\.p\<^sub>1, chine] \ (inv t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine]" using comp_assoc by metis also have "... = \[u, s\<^sub>0, \\.p\<^sub>0] \ (((u \ \.the_\) \ \\.p\<^sub>0) \ (\[u, u\<^sub>0, \.chine] \ \\.p\<^sub>0) \ ((\ \ \.chine) \ \\.p\<^sub>0) \ (\.the_\ \ \\.p\<^sub>0)) \ r\<^sub>0s\<^sub>1.\ \ (\.the_\ \ \\.p\<^sub>1) \ \\<^sup>-\<^sup>1[t\<^sub>0, \.chine, \\.p\<^sub>1] \ (t\<^sub>0 \ inv \\<^sub>\) \ \[t\<^sub>0, \\.p\<^sub>1, chine] \ (inv t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine]" proof - have "(\\<^sup>-\<^sup>1[u\<^sub>1, \.chine, \\.p\<^sub>0] \ \[u\<^sub>1, \.chine, \\.p\<^sub>0]) \ (\.the_\ \ \\.p\<^sub>0) = \.the_\ \ \\.p\<^sub>0" using comp_inv_arr' comp_cod_arr by auto thus ?thesis using comp_assoc by simp qed also have "... = (\[u, s\<^sub>0, \\.p\<^sub>0] \ ((\ \ s\<^sub>0) \ \ \ \\.p\<^sub>0) \ r\<^sub>0s\<^sub>1.\ \ (\.the_\ \ \\.p\<^sub>1) \ \\<^sup>-\<^sup>1[t\<^sub>0, \.chine, \\.p\<^sub>1]) \ (t\<^sub>0 \ inv \\<^sub>\) \ \[t\<^sub>0, \\.p\<^sub>1, chine] \ (inv t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine]" proof - have "arr ((u \ \.the_\) \ \[u, u\<^sub>0, \.chine] \ (\ \ \.chine) \ \.the_\)" using \.\_simps(3) by simp hence "((u \ \.the_\) \ \\.p\<^sub>0) \ (\[u, u\<^sub>0, \.chine] \ \\.p\<^sub>0) \ ((\ \ \.chine) \ \\.p\<^sub>0) \ (\.the_\ \ \\.p\<^sub>0) = (u \ \.the_\) \ \[u, u\<^sub>0, \.chine] \ (\ \ \.chine) \ \.the_\ \ \\.p\<^sub>0" using whisker_right by simp also have "... = (\ \ s\<^sub>0) \ \ \ \\.p\<^sub>0" using \.\_naturality by simp finally have "((u \ \.the_\) \ \\.p\<^sub>0) \ (\[u, u\<^sub>0, \.chine] \ \\.p\<^sub>0) \ ((\ \ \.chine) \ \\.p\<^sub>0) \ (\.the_\ \ \\.p\<^sub>0) = (\ \ s\<^sub>0) \ \ \ \\.p\<^sub>0" by simp thus ?thesis using comp_assoc by presburger qed also have "... = (?\\<^sub>\ \ (t\<^sub>0 \ inv \\<^sub>\)) \ \[t\<^sub>0, \\.p\<^sub>1, chine] \ (inv t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine]" using comp_assoc by presburger also have "... = ?\\<^sub>\' \ \[t\<^sub>0, \\.p\<^sub>1, chine] \ (inv t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine]" using \\<^sub>\' \\<^sub>\'_eq by simp also have "... = (u \ the_\) \ \[u, u\<^sub>0 \ \\.p\<^sub>0, chine] \ (\[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ ((\ \ \\.p\<^sub>0) \ chine) \ ((t\<^sub>0u\<^sub>1.\ \ chine) \ ((\\<^sup>-\<^sup>1[t\<^sub>0, \\.p\<^sub>1, chine]) \ \[t\<^sub>0, \\.p\<^sub>1, chine]) \ (inv t\<^sub>0u\<^sub>1.\ \ chine)) \ \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine]" using comp_assoc by presburger also have "... = (u \ the_\) \ \[u, u\<^sub>0 \ \\.p\<^sub>0, chine] \ (\[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ ((\ \ \\.p\<^sub>0) \ chine) \ \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine]" proof - have "((t\<^sub>0u\<^sub>1.\ \ chine) \ ((\\<^sup>-\<^sup>1[t\<^sub>0, \\.p\<^sub>1, chine]) \ \[t\<^sub>0, \\.p\<^sub>1, chine]) \ (inv t\<^sub>0u\<^sub>1.\ \ chine)) \ \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine] = \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine]" proof - have "((t\<^sub>0u\<^sub>1.\ \ chine) \ ((\\<^sup>-\<^sup>1[t\<^sub>0, \\.p\<^sub>1, chine]) \ \[t\<^sub>0, \\.p\<^sub>1, chine]) \ (inv t\<^sub>0u\<^sub>1.\ \ chine)) \ \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine] = ((t\<^sub>0u\<^sub>1.\ \ chine) \ ((t\<^sub>0 \ \\.p\<^sub>1) \ chine) \ (inv t\<^sub>0u\<^sub>1.\ \ chine)) \ \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine]" using comp_inv_arr' \.T0.antipar(1) by auto also have "... = ((t\<^sub>0u\<^sub>1.\ \ chine) \ (inv t\<^sub>0u\<^sub>1.\ \ chine)) \ \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine]" using comp_cod_arr t\<^sub>0u\<^sub>1.\_uniqueness by simp also have "... = (t\<^sub>0u\<^sub>1.\ \ inv t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine]" using whisker_right t\<^sub>0u\<^sub>1.\_uniqueness by simp also have "... = ((u\<^sub>1 \ \\.p\<^sub>0) \ chine) \ \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine]" using comp_arr_inv' \.T0.antipar(1) t\<^sub>0u\<^sub>1.\_uniqueness by simp also have "... = \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, 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, u\<^sub>0 \ \\.p\<^sub>0, chine] \ (\[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ \\<^sup>-\<^sup>1[u \ u\<^sub>0, \\.p\<^sub>0, chine]) \ (\ \ ?w\<^sub>\)" using assoc'_naturality [of \ \\.p\<^sub>0 chine] comp_assoc by auto also have "... = ((u \ the_\) \ (u \ \\<^sup>-\<^sup>1[u\<^sub>0, \\.p\<^sub>0, chine])) \ \[u, u\<^sub>0, ?w\<^sub>\] \ (\ \ ?w\<^sub>\)" using uw\ pentagon comp_assoc inv_hcomp invert_opposite_sides_of_square [of "u \ \[u\<^sub>0, \\.p\<^sub>0, chine]" "\[u, u\<^sub>0 \ \\.p\<^sub>0, chine] \ (\[u, u\<^sub>0, \\.p\<^sub>0] \ chine)" "\[u, u\<^sub>0, ?w\<^sub>\]" "\[u \ u\<^sub>0, \\.p\<^sub>0, chine]"] \.base_simps(2) \.ide_base \.ide_leg0 \.leg0_simps(2) assoc'_eq_inv_assoc ide_hcomp hcomp_simps(1) t\<^sub>0u\<^sub>1.ide_u by force also have "... = (u \ the_\ \ \\<^sup>-\<^sup>1[u\<^sub>0, \\.p\<^sub>0, chine]) \ \[u, u\<^sub>0, ?w\<^sub>\] \ (\ \ ?w\<^sub>\)" using whisker_left comp_assoc by simp finally show ?thesis by simp qed show "\!\. \\ : ?w\<^sub>\ \ ?w\<^sub>\'\ \ ?\\<^sub>\ = u\<^sub>1 \ \ \ ?\\<^sub>\ = ?\\<^sub>\' \ (u\<^sub>0 \ \)" using w\<^sub>\ w\<^sub>\' \\<^sub>\ \\<^sub>\' \\<^sub>\ eq\<^sub>\ \.T2 [of ?w\<^sub>\ ?w\<^sub>\' ?\\<^sub>\ ?u\<^sub>\ ?\\<^sub>\' ?\\<^sub>\] by fast qed obtain \\<^sub>\ where \\<^sub>\: "\\\<^sub>\ : ?w\<^sub>\ \ ?w\<^sub>\'\ \ ?\\<^sub>\ = u\<^sub>1 \ \\<^sub>\ \ ?\\<^sub>\ = ?\\<^sub>\' \ (u\<^sub>0 \ \\<^sub>\)" using 1 by auto show "?w\<^sub>\ \ ?w\<^sub>\'" using w\<^sub>\ w\<^sub>\' \\<^sub>\ BS3 [of ?w\<^sub>\ ?w\<^sub>\' \\<^sub>\ \\<^sub>\] isomorphic_def by auto qed lemma comp_L: shows "Maps.seq \\t\<^sub>0\\ \\\.chine \ \\.p\<^sub>1\\" and "\\t\<^sub>0\\ \ \\\.chine \ \\.p\<^sub>1\\ = Maps.MkArr (src (\.chine \ \\.p\<^sub>1)) (src t) (Maps.Comp \t\<^sub>0\ \\.chine \ \\.p\<^sub>1\)" proof - show "Maps.seq \\t\<^sub>0\\ \\\.chine \ \\.p\<^sub>1\\" proof - have "is_left_adjoint (\.chine \ \\.p\<^sub>1)" using \.is_map r\<^sub>0s\<^sub>1.leg1_is_map left_adjoints_compose r\<^sub>0s\<^sub>1.p\<^sub>1_simps by auto thus ?thesis using Maps.CLS_in_hom r\<^sub>0s\<^sub>1.leg1_is_map apply (intro Maps.seqI') apply blast using Maps.CLS_in_hom [of t\<^sub>0] \.leg0_is_map \\.leg1_in_hom by auto qed thus "\\t\<^sub>0\\ \ \\\.chine \ \\.p\<^sub>1\\ = Maps.MkArr (src (\.chine \ \\.p\<^sub>1)) (src t) (Maps.Comp \t\<^sub>0\ \\.chine \ \\.p\<^sub>1\)" using Maps.comp_char by auto qed lemma comp_R: shows "Maps.seq \\u\<^sub>1\\ \\\.chine \ \\.p\<^sub>0\\" and "\\u\<^sub>1\\ \ \\\.chine \ \\.p\<^sub>0\\ = Maps.MkArr (src r\<^sub>0s\<^sub>1.p\<^sub>0) (trg u) (Maps.Comp \u\<^sub>1\ \\.chine \ r\<^sub>0s\<^sub>1.p\<^sub>0\)" proof - show "Maps.seq \\u\<^sub>1\\ \\\.chine \ \\.p\<^sub>0\\" proof - have "is_left_adjoint (\.chine \ \\.p\<^sub>0)" using \.is_map r\<^sub>0s\<^sub>1.leg0_is_map left_adjoints_compose [of \.chine \\.p\<^sub>0] by simp thus ?thesis using Maps.CLS_in_hom \.leg1_is_map apply (intro Maps.seqI') apply blast using Maps.CLS_in_hom [of u\<^sub>1] \.leg1_is_map by simp qed thus "\\u\<^sub>1\\ \ \\\.chine \ \\.p\<^sub>0\\ = Maps.MkArr (src r\<^sub>0s\<^sub>1.p\<^sub>0) (trg u) (Maps.Comp \u\<^sub>1\ \\.chine \ r\<^sub>0s\<^sub>1.p\<^sub>0\)" using Maps.comp_char by auto qed lemma comp_L_eq_comp_R: shows "\\t\<^sub>0\\ \ \\\.chine \ \\.p\<^sub>1\\ = \\u\<^sub>1\\ \ \\\.chine \ \\.p\<^sub>0\\" proof (intro Maps.arr_eqI) show "Maps.seq \\t\<^sub>0\\ \\\.chine \ \\.p\<^sub>1\\" using comp_L(1) by simp show "Maps.seq \\u\<^sub>1\\ \\\.chine \ \\.p\<^sub>0\\" using comp_R(1) by simp show "Maps.Dom (\\t\<^sub>0\\ \ \\\.chine \ \\.p\<^sub>1\\) = Maps.Dom (\\u\<^sub>1\\ \ \\\.chine \ \\.p\<^sub>0\\)" 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' r\<^sub>0s\<^sub>1.leg1_simps(3) hcomp_simps(1)) show "Maps.Cod (\\t\<^sub>0\\ \ \\\.chine \ \\.p\<^sub>1\\) = Maps.Cod (\\u\<^sub>1\\ \ \\\.chine \ \\.p\<^sub>0\\)" by (metis Maps.Cod.simps(1) \\.composable comp_L(2) comp_R(2)) have A: "Maps.Map (\\t\<^sub>0\\ \ \\\.chine \ \\.p\<^sub>1\\) = Maps.Comp \t\<^sub>0\ \\.chine \ \\.p\<^sub>1\" using comp_L(1) Maps.comp_char by auto have B: "Maps.Map (\\u\<^sub>1\\ \ \\\.chine \ \\.p\<^sub>0\\) = Maps.Comp \u\<^sub>1\ \\.chine \ \\.p\<^sub>0\" using comp_R(1) Maps.comp_char by auto have C: "Maps.Comp \t\<^sub>0\ \\.chine \ \\.p\<^sub>1\ = Maps.Comp \u\<^sub>1\ \\.chine \ \\.p\<^sub>0\" proof (intro Maps.Comp_eqI) show "t\<^sub>0 \ \.chine \ \\.p\<^sub>1 \ Maps.Comp \t\<^sub>0\ \\.chine \ \\.p\<^sub>1\" proof (intro Maps.in_CompI) show "is_iso_class \\.chine \ \\.p\<^sub>1\" using prj_chine(2) is_iso_classI isomorphic_implies_hpar(2) by blast show "is_iso_class \t\<^sub>0\" using is_iso_classI by auto show "\.chine \ \\.p\<^sub>1 \ \\.chine \ \\.p\<^sub>1\" using ide_in_iso_class prj_chine(2) isomorphic_implies_hpar(2) by blast show "t\<^sub>0 \ \t\<^sub>0\" using ide_in_iso_class by simp show "t\<^sub>0 \ \.chine \ \\.p\<^sub>1 \ t\<^sub>0 \ \.chine \ \\.p\<^sub>1" using isomorphic_reflexive prj_chine(2) isomorphic_implies_hpar(2) by auto qed show "u\<^sub>1 \ \.chine \ \\.p\<^sub>0 \ Maps.Comp \u\<^sub>1\ \\.chine \ \\.p\<^sub>0\" proof (intro Maps.in_CompI) show "is_iso_class \\.chine \ \\.p\<^sub>0\" using is_iso_classI by simp show "is_iso_class \u\<^sub>1\" using is_iso_classI by simp show "\.chine \ \\.p\<^sub>0 \ \\.chine \ \\.p\<^sub>0\" using ide_in_iso_class by simp show "u\<^sub>1 \ iso_class u\<^sub>1" using ide_in_iso_class by simp show "u\<^sub>1 \ \.chine \ \\.p\<^sub>0 \ u\<^sub>1 \ \.chine \ \\.p\<^sub>0" using isomorphic_reflexive isomorphic_implies_hpar(2) by auto qed show "t\<^sub>0 \ \.chine \ \\.p\<^sub>1 \ u\<^sub>1 \ \.chine \ \\.p\<^sub>0" proof - have "t\<^sub>0 \ \.chine \ \\.p\<^sub>1 \ (t\<^sub>0 \ \.chine) \ \\.p\<^sub>1" using assoc'_in_hom [of t\<^sub>0 \.chine \\.p\<^sub>1] iso_assoc' isomorphic_def r\<^sub>0s\<^sub>1.p\<^sub>1_simps by auto also have "... \ r\<^sub>0 \ \\.p\<^sub>1" using \.leg0_uniquely_isomorphic hcomp_isomorphic_ide by (simp add: \.T0.antipar(1)) also have "... \ s\<^sub>1 \ \\.p\<^sub>0" using isomorphic_def r\<^sub>0s\<^sub>1.\_uniqueness(2) by blast also have "... \ (u\<^sub>1 \ \.chine) \ \\.p\<^sub>0" using \.leg1_uniquely_isomorphic hcomp_isomorphic_ide by auto also have "... \ u\<^sub>1 \ \.chine \ \\.p\<^sub>0" using assoc_in_hom [of u\<^sub>1 \.chine \\.p\<^sub>0] iso_assoc isomorphic_def by auto finally show ?thesis by simp qed qed show "Maps.Map (\\t\<^sub>0\\ \ \\\.chine \ \\.p\<^sub>1\\) = Maps.Map (\\u\<^sub>1\\ \ \\\.chine \ \\.p\<^sub>0\\)" using A B C by simp qed lemma csq: shows "Maps.commutative_square \\t\<^sub>0\\ \\u\<^sub>1\\ \\\.chine \ \\.p\<^sub>1\\ \\\.chine \ \\.p\<^sub>0\\" proof show "Maps.cospan \\t\<^sub>0\\ \\u\<^sub>1\\" 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 \ \\.p\<^sub>1\\ \\\.chine \ \\.p\<^sub>0\\" 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 \\t\<^sub>0\\ = Maps.cod \\\.chine \ \\.p\<^sub>1\\" using comp_L(1) by auto show "\\t\<^sub>0\\ \ \\\.chine \ \\.p\<^sub>1\\ = \\u\<^sub>1\\ \ \\\.chine \ \\.p\<^sub>0\\" using comp_L_eq_comp_R by simp qed lemma CLS_chine: shows "\\chine\\ = Maps.tuple \\\.chine \ \\.p\<^sub>1\\ \\t\<^sub>0\\ \\u\<^sub>1\\ \\\.chine \ \\.p\<^sub>0\\" proof - let ?T = "Maps.tuple \\\.chine \ \\.p\<^sub>1\\ \\t\<^sub>0\\ \\u\<^sub>1\\ \\\.chine \ \\.p\<^sub>0\\" have "\!l. \\t\<^sub>0u\<^sub>1.p\<^sub>1\\ \ l = \\\.chine \ \\.p\<^sub>1\\ \ \\t\<^sub>0u\<^sub>1.p\<^sub>0\\ \ l = \\\.chine \ \\.p\<^sub>0\\" using csq \\.prj_char Maps.universal [of "\\t\<^sub>0\\" "\\u\<^sub>1\\" "\\\.chine \ \\.p\<^sub>1\\" "\\\.chine \ \\.p\<^sub>0\\"] by simp moreover have "\\\\.p\<^sub>1\\ \ ?T = \\\.chine \ \\.p\<^sub>1\\ \ \\\\.p\<^sub>0\\ \ ?T = \\\.chine \ \\.p\<^sub>0\\" using csq \\.prj_char Maps.prj_tuple [of "\\t\<^sub>0\\" "\\u\<^sub>1\\" "\\\.chine \ \\.p\<^sub>1\\" "\\\.chine \ \\.p\<^sub>0\\"] by simp moreover have "\\t\<^sub>0u\<^sub>1.p\<^sub>1\\ \ \\chine\\ = \\\.chine \ \\.p\<^sub>1\\ \ \\t\<^sub>0u\<^sub>1.p\<^sub>0\\ \ \\chine\\ = \\\.chine \ \\.p\<^sub>0\\" using prj_chine \\.leg0_is_map \\.leg1_is_map is_map t\<^sub>0u\<^sub>1.leg1_is_map t\<^sub>0u\<^sub>1.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.tab\<^sub>1 \ r.tab\<^sub>0\<^sup>* \ r\ and \s.tab\<^sub>1 \ s.tab\<^sub>0\<^sup>* \ s\ to transform \\\ to \\\' : r.tab\<^sub>1 \ r.tab\<^sub>0\<^sup>* \ s.tab\<^sub>1 \ s.tab\<^sub>0\<^sup>*\\. We then take the adjoint transpose of \\'\ to obtain \\\ : r.tab\<^sub>1 \ (s.tab\<^sub>1 \ s.tab\<^sub>0\<^sup>*) \ r.tab\<^sub>0\\. The 2-cell \\\ induces a map \w\ which is an arrow of spans from \(r.tab\<^sub>0, r.tab\<^sub>1)\ to \(s.tab\<^sub>0, s.tab\<^sub>1)\. 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\<^sup>*\. \ context bicategory_of_spans begin interpretation Maps: maps_category V H \ \ src trg .. interpretation Span: span_bicategory Maps.comp Maps.PRJ\<^sub>0 Maps.PRJ\<^sub>1 .. 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 \)) (tab\<^sub>0 (dom \)) (cod \) (tab_of_ide (cod \)) (tab\<^sub>0 (cod \)) (tab\<^sub>1 (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 \)) (tab\<^sub>0 (dom \)) (cod \) (tab_of_ide (cod \)) (tab\<^sub>0 (cod \)) (tab\<^sub>1 (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 \tab\<^sub>0 (dom \)\ \tab\<^sub>1 (dom \)\ \cod \\ \.s.tab \tab\<^sub>0 (cod \)\ \tab\<^sub>1 (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 (tab\<^sub>0 (dom \)) \ src (tab\<^sub>0 (cod \))\" and "is_left_adjoint (spn \)" and "tab\<^sub>0 (cod \) \ spn \ \ tab\<^sub>0 (dom \)" and "tab\<^sub>1 (cod \) \ spn \ \ tab\<^sub>1 (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 \tab\<^sub>0 (dom \)\ \tab\<^sub>1 (dom \)\ \cod \\ \.s.tab \tab\<^sub>0 (cod \)\ \tab\<^sub>1 (cod \)\ \ using \.is_arrow_of_tabulations_in_maps by simp show "\spn \ : src (tab\<^sub>0 (dom \)) \ src (tab\<^sub>0 (cod \))\" using spn_def by simp show "is_left_adjoint (spn \)" using spn_def by (simp add: \.is_map) show "tab\<^sub>0 (cod \) \ spn \ \ tab\<^sub>0 (dom \)" using spn_def isomorphic_def \.leg0_uniquely_isomorphic(1) by auto show "tab\<^sub>1 (cod \) \ spn \ \ tab\<^sub>1 (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 (tab\<^sub>0 (dom \)) \ src (tab\<^sub>0 (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 (tab\<^sub>0 (dom \))" and "trg (spn \) = src (tab\<^sub>0 (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 (tab\<^sub>0 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 \tab\<^sub>0 r\ \tab\<^sub>1 r\ r r.tab \tab\<^sub>0 r\ \tab\<^sub>1 r\ r using r.is_arrow_of_tabulations_in_maps by simp interpret tab: tabulation V H \ \ src trg r \r.tab\ \tab\<^sub>0 r\ \dom r.tab\ using assms r.tab_is_tabulation by simp interpret tab: tabulation_in_maps V H \ \ src trg r \r.tab\ \tab\<^sub>0 r\ \dom r.tab\ by (unfold_locales, simp_all) have "tab.is_induced_by_cell (spn r) (tab\<^sub>0 r) r.tab" using spn_def comp_ide_arr r.chine_is_induced_map by auto thus ?thesis using tab.induced_map_unique [of "tab\<^sub>0 r" "r.tab" "spn r" "src r.s\<^sub>0"] 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 \tab\<^sub>0 (dom \)\ \tab\<^sub>1 (dom \)\ \cod \\ \.s.tab \tab\<^sub>0 (cod \)\ \tab\<^sub>1 (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 \tab\<^sub>0 (dom \)\ \tab\<^sub>1 (dom \)\ \dom \\ \.r.tab \tab\<^sub>0 (dom \)\ \tab\<^sub>1 (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 \tab\<^sub>0 (dom \)\ \tab\<^sub>1 (dom \)\ \dom \\ \.r.tab \tab\<^sub>0 (dom \)\ \tab\<^sub>1 (dom \)\ \cod \\ \.s.tab \tab\<^sub>0 (cod \)\ \tab\<^sub>1 (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) SPN\<^sub>0 where "SPN\<^sub>0 r \ Span.mkIde \\tab\<^sub>0 r\\ \\tab\<^sub>1 r\\" definition SPN where "SPN \ \ if arr \ then \Chn = \\spn \\\, Dom = \Leg0 = \\tab\<^sub>0 (dom \)\\, Leg1 = \\tab\<^sub>1 (dom \)\\\, Cod = \Leg0 = \\tab\<^sub>0 (cod \)\\, Leg1 = \\tab\<^sub>1 (cod \)\\\\ else Span.null" lemma Dom_SPN [simp]: assumes "arr \" shows "Dom (SPN \) = \Leg0 = \\tab\<^sub>0 (dom \)\\, Leg1 = \\tab\<^sub>1 (dom \)\\\" using assms SPN_def by simp lemma Cod_SPN [simp]: assumes "arr \" shows "Cod (SPN \) = \Leg0 = \\tab\<^sub>0 (cod \)\\, Leg1 = \\tab\<^sub>1 (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 \) (SPN\<^sub>0 (dom \)) (SPN\<^sub>0 (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 = \\tab\<^sub>0 (dom \)\\, Leg1 = \\tab\<^sub>1 (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 = \\tab\<^sub>0 (cod \)\\, Leg1 = \\tab\<^sub>1 (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 "tab\<^sub>0 (cod \)" "spn \"] by simp show "Cod.leg1 \ Chn (SPN \) = Dom.leg1" unfolding SPN_def using assms spn_props [of \] Maps.comp_CLS [of "tab\<^sub>1 (cod \)" "spn \"] by simp qed qed show "Span.dom (SPN \) = SPN\<^sub>0 (dom \)" using assms 1 Span.dom_char Dom'.apex_def SPN_def by simp show "Span.cod (SPN \) = SPN\<^sub>0 (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 (tab\<^sub>0 (dom \))) (src \) \tab\<^sub>0 (dom \)\)" proof - have "src (tab\<^sub>0 (dom \)) \ Collect obj" using \ by simp moreover have "src \ \ Collect obj" using \ by simp moreover have "\tab\<^sub>0 (dom \)\ \ Maps.Hom (src (tab\<^sub>0 (local.dom \))) (src \)" proof - have "\tab\<^sub>0 (dom \) : src (tab\<^sub>0 (dom \)) \ src \\" using \ by simp moreover have "is_left_adjoint (tab\<^sub>0 (dom \))" using \ tab\<^sub>0_simps [of "dom \"] by auto ultimately show ?thesis by auto qed ultimately show ?thesis by simp qed have "\spn (dom \)\ = \src (tab\<^sub>0 (dom \))\" using \ spn_ide iso_class_eqI by auto hence "SPN (dom \) = SPN\<^sub>0 (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 (tab\<^sub>0 (cod \))) (src \) \tab\<^sub>0 (cod \)\)" proof - have "src (tab\<^sub>0 (cod \)) \ Collect obj" using \ by simp moreover have "src \ \ Collect obj" using \ by simp moreover have "\tab\<^sub>0 (cod \)\ \ Maps.Hom (src (tab\<^sub>0 (cod \))) (src \)" proof - have "\tab\<^sub>0 (cod \) : src (tab\<^sub>0 (cod \)) \ src \\" using \ by simp moreover have "is_left_adjoint (tab\<^sub>0 (cod \))" using \ by simp ultimately show ?thesis by auto qed ultimately show ?thesis by simp qed have "\spn (cod \)\ = \src (tab\<^sub>0 (cod \))\" using \ spn_ide iso_class_eqI by auto hence "SPN (cod \) = SPN\<^sub>0 (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 (tab\<^sub>0 (cod \))) (src (tab\<^sub>0 (cod \))) \spn \\ \ Maps.MkArr (src (tab\<^sub>0 (dom \))) (src (tab\<^sub>0 (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 (tab\<^sub>0 (dom \))) (src (tab\<^sub>0 (cod \))) (Maps.Comp \spn \\ \spn \\)" proof - have "Maps.seq (Maps.MkArr (src (tab\<^sub>0 (cod \))) (src (tab\<^sub>0 (cod \))) \spn \\) (Maps.MkArr (src (tab\<^sub>0 (dom \))) (src (tab\<^sub>0 (cod \))) \spn \\)" proof show "Maps.in_hom (Maps.MkArr (src (tab\<^sub>0 (local.dom \))) (src (tab\<^sub>0 (cod \))) \spn \\) (Maps.MkIde (src (tab\<^sub>0 (dom \)))) (Maps.MkIde (src (tab\<^sub>0 (cod \))))" proof - have "src (tab\<^sub>0 (dom \)) \ Collect obj" using in_hhom_def seq by auto moreover have "src (tab\<^sub>0 (cod \)) \ Collect obj" using seq by auto moreover have "\spn \\ \ Maps.Hom (src (tab\<^sub>0 (dom \))) (src (tab\<^sub>0 (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 (tab\<^sub>0 (cod \))) (src (tab\<^sub>0 (cod \))) \spn \\) (Maps.MkIde (src (tab\<^sub>0 (cod \)))) (Maps.MkIde (src (tab\<^sub>0 (cod \))))" proof - have "src (tab\<^sub>0 (cod \)) \ Collect obj" using in_hhom_def seq by auto moreover have "src (tab\<^sub>0 (cod \)) \ Collect obj" using seq by auto moreover have "\spn \\ \ Maps.Hom (src (tab\<^sub>0 (cod \))) (src (tab\<^sub>0 (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 (tab\<^sub>0 (cod \))) (src (tab\<^sub>0 (cod \))) \spn \\" "Maps.MkArr (src (tab\<^sub>0 (dom \))) (src (tab\<^sub>0 (cod \))) \spn \\"] by simp qed also have "... = Maps.MkArr (src (tab\<^sub>0 (dom \))) (src (tab\<^sub>0 (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 \tab\<^sub>0 (src \)\ \tab\<^sub>1 (src \)\ \src \\ src.tab \tab\<^sub>0 (src \)\ \tab\<^sub>1 (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 ?tab\<^sub>0 = "Maps.MkArr (src (tab\<^sub>0 (src \))) (src \) \tab\<^sub>0 (src \)\" have tab\<^sub>0_src: "\tab\<^sub>0 (src \) : src (tab\<^sub>0 (src \)) \ src \\ \ is_left_adjoint (tab\<^sub>0 (src \)) \ \tab\<^sub>0 (src \)\ = \tab\<^sub>0 (src \)\" using \ by simp have tab\<^sub>0: "Maps.arr ?tab\<^sub>0" using \ Maps.arr_MkArr tab\<^sub>0_src by blast let ?tab\<^sub>1 = "Maps.MkArr (src (tab\<^sub>0 (src \))) (src \) \tab\<^sub>1 (src \)\" have tab\<^sub>1_src: "\tab\<^sub>1 (src \) : src (tab\<^sub>0 (src \)) \ src \\ \ is_left_adjoint (tab\<^sub>1 (src \)) \ \tab\<^sub>1 (src \)\ = \tab\<^sub>1 (src \)\" using \ by simp have tab\<^sub>1: "Maps.arr ?tab\<^sub>1" using \ Maps.arr_MkArr tab\<^sub>1_src by blast interpret tab: span_in_category Maps.comp \\Leg0 = ?tab\<^sub>0, Leg1 = ?tab\<^sub>1\\ using tab\<^sub>0 tab\<^sub>1 Maps.dom_char Maps.cod_char by (unfold_locales, simp) have "src \ \ tab\<^sub>0 (src \) \ tab\<^sub>0 (src \)" using \ iso_lunit isomorphic_def by (metis lunit_in_hom(2) src.ide_u src.u_simps(3) src_src) hence "src \ \ tab\<^sub>0 (src \) \ tab\<^sub>1 (src \)" using \ src.obj_has_symmetric_tab isomorphic_transitive by blast have "\tab\<^sub>0 (src \)\ \ Maps.Hom (src (tab\<^sub>0 (src \))) (src \)" using \ tab\<^sub>0_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 (tab\<^sub>0 (src \))) (src (tab\<^sub>0 (src \))) \spn (src \)\, Dom = \Leg0 = ?tab\<^sub>0, Leg1 = ?tab\<^sub>1\, Cod = \Leg0 = ?tab\<^sub>0, Leg1 = ?tab\<^sub>1\\" 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 ?tab\<^sub>0_dom = "Maps.MkArr (src (tab\<^sub>0 (dom \))) (src \) \tab\<^sub>0 (dom \)\" have "Maps.arr ?tab\<^sub>0_dom" proof - have "\tab\<^sub>0 (dom \) : src (tab\<^sub>0 (dom \)) \ src \\ \ is_left_adjoint (tab\<^sub>0 (dom \)) \ \tab\<^sub>0 (dom \)\ = \tab\<^sub>0 (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 (tab\<^sub>0 (src \))"} is invertible in \Maps(B)\ and determines an invertible arrow of spans from @{term "SPN (src \)"} to @{term "Span.src (SPN \)"}. \ let ?\ = "\Chn = ?tab\<^sub>0, 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 ?tab\<^sub>0 SPN_src.dom.apex src_SPN.cod.apex" using \ tab\<^sub>0 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 \ ?tab\<^sub>0 = SPN_src.dom.leg0" proof - have "Maps.seq src_SPN.cod.leg0 ?tab\<^sub>0" using \ src_SPN tab\<^sub>0 Maps.dom_char Maps.cod_char by (intro Maps.seqI, auto) moreover have "Maps.Comp \src \\ \tab\<^sub>0 (src \)\ = \tab\<^sub>0 (src \)\" proof - have "tab\<^sub>0 (src \) \ Maps.Comp \src \\ \tab\<^sub>0 (src \)\" using \ is_iso_classI ide_in_iso_class [of "src \"] ide_in_iso_class [of "tab\<^sub>0 (src \)"] \src \ \ tab\<^sub>0 (src \) \ tab\<^sub>0 (src \)\ by auto thus ?thesis using Maps.Comp_eq_iso_class_memb \\tab\<^sub>0 (src \)\ \ Maps.Hom (src (tab\<^sub>0 (src \))) (src \)\ \\src \\ \ Maps.Hom (src \) (src \)\ by meson qed ultimately show ?thesis using \ Maps.comp_char [of src_SPN.cod.leg0 ?tab\<^sub>0] src_SPN by simp qed show "src_SPN.cod.leg1 \ ?tab\<^sub>0 = SPN_src.dom.leg1" proof - have "Maps.seq src_SPN.cod.leg1 ?tab\<^sub>0" using \ src_SPN tab\<^sub>0 Maps.dom_char Maps.cod_char by (intro Maps.seqI, auto) moreover have "Maps.Comp \src \\ \tab\<^sub>0 (src \)\ = \tab\<^sub>1 (src \)\" proof - have "tab\<^sub>1 (src \) \ Maps.Comp \src \\ \tab\<^sub>0 (src \)\" using \ is_iso_classI ide_in_iso_class [of "src \"] ide_in_iso_class [of "tab\<^sub>0 (src \)"] \isomorphic (src \ \ tab\<^sub>0 (src \)) (tab\<^sub>1 (src \))\ by auto thus ?thesis using Maps.Comp_eq_iso_class_memb \\tab\<^sub>0 (src \)\ \ Maps.Hom (src (tab\<^sub>0 (src \))) (src \)\ \\src \\ \ Maps.Hom (src \) (src \)\ by meson qed ultimately show ?thesis using \ Maps.comp_char [of src_SPN.cod.leg1 ?tab\<^sub>0] src_SPN by simp qed qed have "Span.in_hom ?\ (SPN (src \)) (Span.src (SPN \))" using \ tab\<^sub>0 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 ?tab\<^sub>0" using \ tab\<^sub>0 ide_in_iso_class src.is_map_iff_tab\<^sub>0_is_equivalence obj_is_self_adjoint Maps.iso_char' [of ?tab\<^sub>0] 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 ?tab\<^sub>0 = "Maps.MkArr (src (tab\<^sub>0 (trg \))) (trg \) \tab\<^sub>0 (trg \)\" have tab\<^sub>0_trg: "\tab\<^sub>0 (trg \) : src (tab\<^sub>0 (trg \)) \ trg \\ \ is_left_adjoint (tab\<^sub>0 (trg \)) \ \tab\<^sub>0 (trg \)\ = \tab\<^sub>0 (trg \)\" using \ by simp have tab\<^sub>0: "Maps.arr ?tab\<^sub>0" using \ Maps.arr_MkArr tab\<^sub>0_trg by blast let ?tab\<^sub>1 = "Maps.MkArr (src (tab\<^sub>0 (trg \))) (trg \) \tab\<^sub>1 (trg \)\" have tab\<^sub>1_trg: "\tab\<^sub>1 (trg \) : src (tab\<^sub>0 (trg \)) \ trg \\ \ is_left_adjoint (tab\<^sub>1 (trg \)) \ \tab\<^sub>1 (trg \)\ = \tab\<^sub>1 (trg \)\" using \ by simp have tab\<^sub>1: "Maps.arr ?tab\<^sub>1" using \ Maps.arr_MkArr tab\<^sub>1_trg by blast interpret tab: span_in_category Maps.comp \\Leg0 = ?tab\<^sub>0, Leg1 = ?tab\<^sub>1\\ using tab\<^sub>0 tab\<^sub>1 Maps.dom_char Maps.cod_char by (unfold_locales, simp) have "trg \ \ tab\<^sub>1 (trg \) \ tab\<^sub>0 (trg \)" proof - have "\\[tab\<^sub>1 (trg \)] : trg \ \ tab\<^sub>1 (trg \) \ tab\<^sub>1 (trg \)\" using \ by simp moreover have "iso \[tab\<^sub>1 (trg \)]" using \ iso_lunit by simp ultimately have "trg \ \ tab\<^sub>1 (trg \) \ tab\<^sub>1 (trg \)" using isomorphic_def by auto also have "tab\<^sub>1 (trg \) \ tab\<^sub>0 (trg \)" using \ trg.obj_has_symmetric_tab isomorphic_symmetric by auto finally show ?thesis by blast qed hence "trg \ \ tab\<^sub>1 (trg \) \ tab\<^sub>1 (trg \)" using \ trg.obj_has_symmetric_tab isomorphic_transitive by blast have "\tab\<^sub>1 (trg \)\ \ Maps.Hom (src (tab\<^sub>0 (trg \))) (trg \)" proof - have "\tab\<^sub>1 (trg \) : src (tab\<^sub>0 (trg \)) \ trg \\ \ is_left_adjoint (tab\<^sub>0 (trg \)) \ \tab\<^sub>0 (trg \)\ = \tab\<^sub>0 (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 (tab\<^sub>1 (trg \))) (src (tab\<^sub>1 (trg \))) \spn (trg \)\, Dom = \Leg0 = ?tab\<^sub>0, Leg1 = ?tab\<^sub>1\, Cod = \Leg0 = ?tab\<^sub>0, Leg1 = ?tab\<^sub>1\\" 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 ?tab\<^sub>1_dom = "Maps.MkArr (src (tab\<^sub>1 (dom \))) (trg \) \tab\<^sub>1 (dom \)\" have "Maps.arr ?tab\<^sub>1_dom" proof - have "\tab\<^sub>1 (dom \) : src (tab\<^sub>1 (dom \)) \ trg \\ \ is_left_adjoint (tab\<^sub>1 (dom \)) \ \tab\<^sub>1 (dom \)\ = \tab\<^sub>1 (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 = ?tab\<^sub>1, 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 ?tab\<^sub>1 SPN_trg.dom.apex trg_SPN.cod.apex" using \ tab\<^sub>0 tab\<^sub>1 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 ?tab\<^sub>1 = SPN_trg.dom.leg0" proof - have "Maps.seq trg_SPN.cod.leg0 ?tab\<^sub>1" using \ trg_SPN tab\<^sub>1 Maps.dom_char Maps.cod_char by (intro Maps.seqI, auto) moreover have "Maps.Comp \trg \\ \tab\<^sub>1 (trg \)\ = \tab\<^sub>1 (trg \)\" proof - have "tab\<^sub>1 (trg \) \ Maps.Comp \trg \\ \tab\<^sub>1 (trg \)\" using \ is_iso_classI ide_in_iso_class [of "trg \"] ide_in_iso_class [of "tab\<^sub>1 (trg \)"] \trg \ \ tab\<^sub>1 (trg \) \ tab\<^sub>1 (trg \)\ by auto thus ?thesis using Maps.Comp_eq_iso_class_memb \iso_class (tab\<^sub>1 (trg \)) \ Maps.Hom (src (tab\<^sub>0 (trg \))) (trg \)\ \iso_class (trg \) \ Maps.Hom (trg \) (trg \)\ by meson qed moreover have "\tab\<^sub>1 (trg \)\ = \tab\<^sub>0 (trg \)\" using \ iso_class_eqI trg.obj_has_symmetric_tab by auto ultimately show ?thesis using \ Maps.comp_char [of trg_SPN.cod.leg0 ?tab\<^sub>1] trg_SPN by simp qed show "trg_SPN.cod.leg1 \ ?tab\<^sub>1 = SPN_trg.dom.leg1" proof - have "Maps.seq trg_SPN.cod.leg1 ?tab\<^sub>1" using \ trg_SPN tab\<^sub>1 Maps.dom_char Maps.cod_char by (intro Maps.seqI, auto) moreover have "Maps.Comp \trg \\ \tab\<^sub>1 (trg \)\ = \tab\<^sub>1 (trg \)\" proof - have "tab\<^sub>1 (trg \) \ Maps.Comp \trg \\ \tab\<^sub>1 (trg \)\" using \ is_iso_classI ide_in_iso_class [of "trg \"] ide_in_iso_class [of "tab\<^sub>1 (trg \)"] \trg \ \ tab\<^sub>1 (trg \) \ tab\<^sub>1 (trg \)\ by auto thus ?thesis using Maps.Comp_eq_iso_class_memb \\tab\<^sub>1 (trg \)\ \ Maps.Hom (src (tab\<^sub>0 (trg \))) (trg \)\ \\trg \\ \ Maps.Hom (trg \) (trg \)\ by meson qed ultimately show ?thesis using \ Maps.comp_char [of trg_SPN.cod.leg1 ?tab\<^sub>1] trg_SPN by simp qed qed have \: "Span.in_hom ?\ (SPN (trg \)) (Span.trg (SPN \))" using \ tab\<^sub>0 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 ?tab\<^sub>1" proof - have "Maps.iso ?tab\<^sub>0" using \ tab\<^sub>0 ide_in_iso_class trg.is_map_iff_tab\<^sub>0_is_equivalence obj_is_self_adjoint Maps.iso_char' [of ?tab\<^sub>0] by auto moreover have "?tab\<^sub>0 = ?tab\<^sub>1" proof - have "\tab\<^sub>0 (trg \)\ = \tab\<^sub>1 (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 \tab\<^sub>0 r\ \tab\<^sub>1 r\ r r.tab \tab\<^sub>0 r\ \tab\<^sub>1 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 \tab\<^sub>0 s\ \tab\<^sub>1 s\ s s.tab \tab\<^sub>0 s\ \tab\<^sub>1 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 \tab\<^sub>0 r\ \tab\<^sub>1 r\ s s.tab \tab\<^sub>0 s\ \tab\<^sub>1 s\ r r.tab \tab\<^sub>0 r\ \tab\<^sub>1 r\ s s.tab \tab\<^sub>0 s\ \tab\<^sub>1 s\ r s using composable by unfold_locales auto abbreviation p\<^sub>0 where "p\<^sub>0 \ \\.p\<^sub>0" abbreviation p\<^sub>1 where "p\<^sub>1 \ \\.p\<^sub>1" 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 \tab\<^sub>0 s \ \\.p\<^sub>0\ \tab\<^sub>1 r \ \\.p\<^sub>1\ \r \ s\ tab \tab\<^sub>0 (r \ s)\ \tab\<^sub>1 (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 (tab\<^sub>0 s \ \\.p\<^sub>0) (tab\<^sub>1 r \ \\.p\<^sub>1) (r \ s) tab (tab\<^sub>0 (r \ s)) (tab\<^sub>1 (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 "tab\<^sub>0 (r \ s) \ cmp \ tab\<^sub>0 s \ \\.p\<^sub>0" and "tab\<^sub>1 (r \ s) \ cmp \ tab\<^sub>1 r \ \\.p\<^sub>1" 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 isomoprhisms 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.PRJ\<^sub>0 Maps.PRJ\<^sub>1 .. 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 R\<^sub>0 where "R\<^sub>0 = \\tab\<^sub>0 r\\" definition R\<^sub>1 where "R\<^sub>1 = \\tab\<^sub>1 r\\" definition S\<^sub>0 where "S\<^sub>0 = \\tab\<^sub>0 s\\" definition S\<^sub>1 where "S\<^sub>1 = \\tab\<^sub>1 s\\" lemma span_legs_eq: shows "Leg0 (Dom (SPN r)) = R\<^sub>0" and "Leg1 (Dom (SPN r)) = R\<^sub>1" and "Leg0 (Dom (SPN s)) = S\<^sub>0" and "Leg1 (Dom (SPN s)) = S\<^sub>1" using SPN_def R\<^sub>0_def R\<^sub>1_def S\<^sub>0_def S\<^sub>1_def composable by auto lemma R\<^sub>0_in_hom [intro]: shows "Maps.in_hom R\<^sub>0 (Maps.MkIde (src r.s\<^sub>0)) (Maps.MkIde (src r))" by (simp add: Maps.MkArr_in_hom' R\<^sub>0_def) lemma R\<^sub>1_in_hom [intro]: shows "Maps.in_hom R\<^sub>1 (Maps.MkIde (src r.s\<^sub>0)) (Maps.MkIde (trg r))" by (simp add: Maps.MkArr_in_hom' R\<^sub>1_def) lemma S\<^sub>0_in_hom [intro]: shows "Maps.in_hom S\<^sub>0 (Maps.MkIde (src s.s\<^sub>0)) (Maps.MkIde (src s))" by (simp add: Maps.MkArr_in_hom' S\<^sub>0_def) lemma S\<^sub>1_in_hom [intro]: shows "Maps.in_hom S\<^sub>1 (Maps.MkIde (src s.s\<^sub>0)) (Maps.MkIde (trg s))" by (simp add: Maps.MkArr_in_hom' S\<^sub>1_def) lemma RS_simps [simp]: shows "Maps.arr R\<^sub>0" and "Maps.dom R\<^sub>0 = Maps.MkIde (src r.s\<^sub>0)" and "Maps.cod R\<^sub>0 = Maps.MkIde (src r)" and "Maps.Dom R\<^sub>0 = src r.s\<^sub>0" and "Maps.Cod R\<^sub>0 = src r" and "Maps.arr R\<^sub>1" and "Maps.dom R\<^sub>1 = Maps.MkIde (src r.s\<^sub>0)" and "Maps.cod R\<^sub>1 = Maps.MkIde (trg r)" and "Maps.Dom R\<^sub>1 = src r.s\<^sub>0" and "Maps.Cod R\<^sub>1 = trg r" and "Maps.arr S\<^sub>0" and "Maps.dom S\<^sub>0 = Maps.MkIde (src s.s\<^sub>0)" and "Maps.cod S\<^sub>0 = Maps.MkIde (src s)" and "Maps.Dom S\<^sub>0 = src s.s\<^sub>0" and "Maps.Cod S\<^sub>0 = src s" and "Maps.arr S\<^sub>1" and "Maps.dom S\<^sub>1 = Maps.MkIde (src s.s\<^sub>0)" and "Maps.cod S\<^sub>1 = Maps.MkIde (trg s)" and "Maps.Dom S\<^sub>1 = src s.s\<^sub>0" and "Maps.Cod S\<^sub>1 = trg s" using R\<^sub>0_in_hom R\<^sub>1_in_hom S\<^sub>0_in_hom S\<^sub>1_in_hom composable by (auto simp add: R\<^sub>0_def R\<^sub>1_def S\<^sub>0_def S\<^sub>1_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 \(tab\<^sub>0 r)\<^sup>* \ tab\<^sub>1 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_S\<^sub>1: map_in_bicategory V H \ \ src trg \Maps.REP S\<^sub>1\ using Maps.REP_in_Map Maps.arr_char Maps.in_HomD S\<^sub>1_def apply unfold_locales by (meson Maps.REP_in_hhom(2) S\<^sub>1_in_hom) interpret REP_R\<^sub>0: map_in_bicategory V H \ \ src trg \Maps.REP R\<^sub>0\ using Maps.REP_in_Map Maps.arr_char Maps.in_HomD R\<^sub>0_def apply unfold_locales by (meson Maps.REP_in_hhom(2) R\<^sub>0_in_hom) have "Maps.Dom SPN_r_SPN_s.leg0 = Maps.Dom (S\<^sub>0 \ Maps.PRJ\<^sub>0 R\<^sub>0 S\<^sub>1)" using composable Span.hcomp_def S\<^sub>0_def R\<^sub>0_def S\<^sub>1_def by simp also have "... = Maps.Dom \\tab\<^sub>0 ((Maps.REP R\<^sub>0)\<^sup>* \ Maps.REP S\<^sub>1)\\" proof - have "is_left_adjoint (tab\<^sub>0 ((Maps.REP R\<^sub>0)\<^sup>* \ Maps.REP S\<^sub>1))" proof - have "ide ((Maps.REP R\<^sub>0)\<^sup>* \ Maps.REP S\<^sub>1)" proof - have "src (Maps.REP R\<^sub>0)\<^sup>* = trg (Maps.REP S\<^sub>1)" using REP_R\<^sub>0.is_map REP_S\<^sub>1.is_map left_adjoint_is_ide R\<^sub>0_def S\<^sub>1_def by (metis (no_types, lifting) Maps.REP_CLS REP_R\<^sub>0.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 (S\<^sub>0 \ \\tab\<^sub>0 ((Maps.REP R\<^sub>0)\<^sup>* \ Maps.REP S\<^sub>1)\\) = src (tab\<^sub>0 ((Maps.REP R\<^sub>0)\<^sup>* \ Maps.REP S\<^sub>1))" proof - have "Maps.arr (\\prj\<^sub>0 (Maps.REP S\<^sub>1) (Maps.REP R\<^sub>0)\\)" using Maps.CLS_in_hom Maps.prj0_simps(1) Maps.PRJ\<^sub>0_def composable by fastforce moreover have "Maps.Dom S\<^sub>0 = Maps.Cod \\prj\<^sub>0 (Maps.REP S\<^sub>1) (Maps.REP R\<^sub>0)\\" proof - have "Maps.Cod \\prj\<^sub>0 (Maps.REP S\<^sub>1) (Maps.REP R\<^sub>0)\\ = trg (prj\<^sub>0 (Maps.REP S\<^sub>1) (Maps.REP R\<^sub>0))" by simp also have "... = src (Maps.REP S\<^sub>1)" proof - have "ide (Maps.REP S\<^sub>1)" by simp moreover have "is_left_adjoint (Maps.REP R\<^sub>0)" by auto moreover have "trg (Maps.REP S\<^sub>1) = trg (Maps.REP R\<^sub>0)" by (simp add: composable) ultimately show ?thesis using S\<^sub>1_def Maps.REP_CLS r.leg0_is_map s.leg1_is_map by simp qed also have "... = src (tab\<^sub>0 s)" using tab\<^sub>0_in_hom(1) by simp also have "... = Maps.Dom S\<^sub>0" using S\<^sub>0_def by simp finally show ?thesis by simp qed ultimately have "Maps.Dom (S\<^sub>0 \ \\tab\<^sub>0 ((Maps.REP R\<^sub>0)\<^sup>* \ Maps.REP S\<^sub>1)\\) = Maps.Dom \\tab\<^sub>0 ((Maps.REP R\<^sub>0)\<^sup>* \ Maps.REP S\<^sub>1)\\" using Maps.CLS_in_hom by simp thus ?thesis by simp qed ultimately show ?thesis using Maps.PRJ\<^sub>0_def composable Maps.Dom.simps(1) RS_simps(1) RS_simps(16) RS_simps(18) RS_simps(3) by presburger qed also have "... = src (tab\<^sub>0 ((Maps.REP R\<^sub>0)\<^sup>* \ Maps.REP S\<^sub>1))" by simp finally have "Maps.Dom SPN_r_SPN_s.leg0 = src (tab\<^sub>0 ((Maps.REP R\<^sub>0)\<^sup>* \ Maps.REP S\<^sub>1))" by simp also have "... = src (tab\<^sub>0 (r.s\<^sub>0\<^sup>* \ s.s\<^sub>1))" proof - interpret r\<^sub>0's\<^sub>1: identity_in_bicategory_of_spans V H \ \ src trg \r.s\<^sub>0\<^sup>* \ s.s\<^sub>1\ using composable by (unfold_locales, simp) have "(Maps.REP R\<^sub>0)\<^sup>* \ Maps.REP S\<^sub>1 \ r.s\<^sub>0\<^sup>* \ s.s\<^sub>1" proof - have "(Maps.REP R\<^sub>0)\<^sup>* \ r.s\<^sub>0\<^sup>*" proof - have 1: "adjoint_pair (Maps.REP R\<^sub>0) (Maps.REP R\<^sub>0)\<^sup>*" using REP_R\<^sub>0.is_map left_adjoint_extends_to_adjoint_pair by blast moreover have "adjoint_pair r.s\<^sub>0 (Maps.REP R\<^sub>0)\<^sup>*" proof - have "Maps.REP R\<^sub>0 \ r.s\<^sub>0" unfolding R\<^sub>0_def using Maps.REP_CLS r.leg0_is_map composable by force thus ?thesis using 1 adjoint_pair_preserved_by_iso isomorphic_def REP_R\<^sub>0.triangle_in_hom(4) REP_R\<^sub>0.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 S\<^sub>1 \ s.s\<^sub>1" unfolding S\<^sub>1_def using Maps.REP_CLS s.leg1_is_map composable by force moreover have "\a. a \ (tab\<^sub>0 r)\<^sup>* \ tab\<^sub>1 s \ (Maps.REP R\<^sub>0)\<^sup>* \ Maps.REP S\<^sub>1 \ a" using calculation composable isomorphic_implies_hpar(3) hcomp_ide_isomorphic hcomp_isomorphic_ide [of "(Maps.REP R\<^sub>0)\<^sup>*" "r.s\<^sub>0\<^sup>*" s.s\<^sub>1] by auto ultimately show ?thesis using isomorphic_transitive by blast qed thus ?thesis using r\<^sub>0's\<^sub>1.isomorphic_implies_same_tab isomorphic_symmetric by metis qed also have "... = src \\.tab" using VV.ide_char VV.arr_char 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 r\<^sub>0's\<^sub>1: two_composable_identities_in_bicategory_of_spans V H \ \ src trg \(Maps.REP R\<^sub>0)\<^sup>*\ \Maps.REP S\<^sub>1\ proof show "ide (Maps.REP S\<^sub>1)" using Maps.REP_in_Map Maps.arr_char left_adjoint_is_ide by (meson Maps.REP_in_hhom(2) S\<^sub>1_in_hom) show "ide (Maps.REP R\<^sub>0)\<^sup>*" using Maps.REP_in_Map Maps.arr_char left_adjoint_is_ide Maps.REP_in_hhom(2) R\<^sub>0_in_hom by auto show "src (Maps.REP R\<^sub>0)\<^sup>* = trg (Maps.REP S\<^sub>1)" using Maps.REP_in_hhom(2) R\<^sub>0_in_hom composable by auto qed interpretation R\<^sub>0'S\<^sub>1: identity_in_bicategory_of_spans V H \ \ src trg \(tab\<^sub>0 r)\<^sup>* \ tab\<^sub>1 s\ by (unfold_locales, simp add: composable) lemma prj_tab_agreement: shows "(tab\<^sub>0 r)\<^sup>* \ tab\<^sub>1 s \ (Maps.REP R\<^sub>0)\<^sup>* \ Maps.REP S\<^sub>1" and "\\.p\<^sub>0 \ prj\<^sub>0 (Maps.REP S\<^sub>1) (Maps.REP R\<^sub>0)" and "\\.p\<^sub>1 \ prj\<^sub>1 (Maps.REP S\<^sub>1) (Maps.REP R\<^sub>0)" proof - show 1: "(tab\<^sub>0 r)\<^sup>* \ tab\<^sub>1 s \ (Maps.REP R\<^sub>0)\<^sup>* \ Maps.REP S\<^sub>1" proof - have "(tab\<^sub>0 r)\<^sup>* \ (Maps.REP R\<^sub>0)\<^sup>*" using Maps.REP_CLS isomorphic_symmetric R\<^sub>0_def composable r.satisfies_T0 isomorphic_to_left_adjoint_implies_isomorphic_right_adjoint by fastforce moreover have "tab\<^sub>1 s \ Maps.REP S\<^sub>1" by (metis Maps.REP_CLS isomorphic_symmetric S\<^sub>1_def s.leg1_is_map s.leg1_simps(3-4)) moreover have "src (Maps.REP R\<^sub>0)\<^sup>* = trg (tab\<^sub>1 s)" using composable r.T0.antipar right_adjoint_simps(2) by fastforce ultimately show ?thesis using hcomp_isomorphic_ide [of "(tab\<^sub>0 r)\<^sup>*" "(Maps.REP R\<^sub>0)\<^sup>*" "tab\<^sub>1 s"] hcomp_ide_isomorphic isomorphic_transitive composable by auto qed show "\\.p\<^sub>0 \ tab\<^sub>0 ((Maps.REP R\<^sub>0)\<^sup>* \ Maps.REP S\<^sub>1)" using 1 R\<^sub>0'S\<^sub>1.isomorphic_implies_same_tab isomorphic_reflexive by auto show "\\.p\<^sub>1 \ tab\<^sub>1 ((Maps.REP R\<^sub>0)\<^sup>* \ Maps.REP S\<^sub>1)" using 1 R\<^sub>0'S\<^sub>1.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 \\.p\<^sub>0)" proof - have "Span.chine_hcomp (SPN r) (SPN s) = Maps.MkIde (src (tab\<^sub>0 ((Maps.REP R\<^sub>0)\<^sup>* \ Maps.REP S\<^sub>1)))" using Span.chine_hcomp_ide_ide [of "SPN r" "SPN s"] composable Maps.pbdom_def Maps.PRJ\<^sub>0_def Maps.CLS_in_hom Maps.dom_char R\<^sub>0_def S\<^sub>1_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 \\.p\<^sub>0)" 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.PRJ\<^sub>0 Maps.PRJ\<^sub>1 .. 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 \ 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.map\<^sub>0 (src s)) (SPN.map\<^sub>0 (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_char VV.arr_char by simp interpret rs: two_composable_identities_in_bicategory_of_spans V H \ \ src trg r s using rs VV.ide_char VV.arr_char apply unfold_locales by auto interpret cmp: arrow_of_tabulations_in_maps V H \ \ src trg \r \ s\ rs.\\.tab \tab\<^sub>0 s \ rs.\\.p\<^sub>0\ \tab\<^sub>1 r \ rs.\\.p\<^sub>1\ \r \ s\ rs.tab \tab\<^sub>0 (r \ s)\ \tab\<^sub>1 (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 "\tab\<^sub>0 (r \ s) : src (tab\<^sub>0 (r \ s)) \ src s\ \ is_left_adjoint (tab\<^sub>0 (r \ s)) \ \tab\<^sub>0 (r \ s)\ = \tab\<^sub>0 (r \ s)\" by simp hence "\f. \f : src (tab\<^sub>0 (r \ s)) \ src s\ \ is_left_adjoint f \ \tab\<^sub>0 (r \ s)\ = \f\" by blast moreover have "\f. \f : src (tab\<^sub>0 (r \ s)) \ trg r\ \ is_left_adjoint f \ \tab\<^sub>1 (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 r\<^sub>0's\<^sub>1: two_composable_identities_in_bicategory_of_spans V H \ \ src trg \(Maps.REP rs.R\<^sub>0)\<^sup>*\ \Maps.REP rs.S\<^sub>1\ proof show "ide (Maps.REP rs.S\<^sub>1)" using Maps.REP_in_Map Maps.arr_char left_adjoint_is_ide by (meson Maps.REP_in_hhom(2) rs.S\<^sub>1_in_hom) show "ide (Maps.REP rs.R\<^sub>0)\<^sup>*" using Maps.REP_in_Map Maps.arr_char left_adjoint_is_ide Maps.REP_in_hhom(2) rs.R\<^sub>0_in_hom by auto show "src (Maps.REP rs.R\<^sub>0)\<^sup>* = trg (Maps.REP rs.S\<^sub>1)" using Maps.REP_in_hhom(2) rs.R\<^sub>0_in_hom rs.composable by auto qed interpret R\<^sub>0'S\<^sub>1: identity_in_bicategory_of_spans V H \ \ src trg \(tab\<^sub>0 r)\<^sup>* \ tab\<^sub>1 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.S\<^sub>0 (Maps.PRJ\<^sub>0 rs.R\<^sub>0 rs.S\<^sub>1)" using rs.composable Span.hcomp_def rs.S\<^sub>0_def rs.R\<^sub>0_def rs.S\<^sub>1_def by simp have SPN_r_SPN_s_leg1_eq: "SPN_r_SPN_s.leg1 = Maps.comp rs.R\<^sub>1 (Maps.PRJ\<^sub>1 rs.R\<^sub>0 rs.S\<^sub>1)" using rs.composable Span.hcomp_def rs.R\<^sub>1_def rs.R\<^sub>0_def rs.S\<^sub>1_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 = \\tab\<^sub>0 (r \ s)\\" unfolding SPN_def using rs by simp have SPN_rs_leg1_eq: "SPN_rs.leg1 = \\tab\<^sub>1 (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.S\<^sub>0 "Maps.PRJ\<^sub>0 rs.R\<^sub>0 rs.S\<^sub>1"] 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 \tab\<^sub>0 (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 \tab\<^sub>0 s\ \rs.\\.p\<^sub>0\" proof - have "Maps.Comp \tab\<^sub>0 (r \ s)\ \rs.cmp\ = \tab\<^sub>0 (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 \tab\<^sub>0 s\ \rs.\\.p\<^sub>0\" 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.S\<^sub>0 (Maps.PRJ\<^sub>0 rs.R\<^sub>0 rs.S\<^sub>1)" by (intro Maps.seqI, simp_all add: rs.composable) moreover have "\prj\<^sub>0 (Maps.REP rs.S\<^sub>1) (Maps.REP rs.R\<^sub>0)\ = \rs.\\.p\<^sub>0\" using "rs.prj_tab_agreement" iso_class_eqI by auto moreover have "Maps.Dom (Maps.PRJ\<^sub>0 rs.R\<^sub>0 rs.S\<^sub>1) = src rs.\\.p\<^sub>0" using rs.prj_tab_agreement Maps.PRJ\<^sub>0_def rs.composable isomorphic_implies_hpar(3) by auto ultimately show ?thesis using SPN_r_SPN_s_leg0_eq Maps.comp_char [of rs.S\<^sub>0 "Maps.PRJ\<^sub>0 rs.R\<^sub>0 rs.S\<^sub>1"] rs.S\<^sub>0_def Maps.PRJ\<^sub>0_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.R\<^sub>1 "Maps.PRJ\<^sub>1 rs.R\<^sub>0 rs.S\<^sub>1"] 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 \tab\<^sub>1 (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 \tab\<^sub>1 r\ \rs.\\.p\<^sub>1\" proof - have "Maps.Comp \tab\<^sub>1 (r \ s)\ \rs.cmp\ = \tab\<^sub>1 (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 \tab\<^sub>1 r\ \rs.\\.p\<^sub>1\" using Maps.Comp_eq_iso_class_memb Maps.CLS_hcomp [of "tab\<^sub>1 r" rs.\\.p\<^sub>1] 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.R\<^sub>1 (Maps.PRJ\<^sub>1 rs.R\<^sub>0 rs.S\<^sub>1)" by (intro Maps.seqI, simp_all add: rs.composable) moreover have "\prj\<^sub>1 (Maps.REP rs.S\<^sub>1) (Maps.REP rs.R\<^sub>0)\ = \rs.\\.p\<^sub>1\" using rs.prj_tab_agreement iso_class_eqI by auto moreover have "Maps.Dom (Maps.PRJ\<^sub>1 rs.R\<^sub>0 rs.S\<^sub>1) = src rs.\\.p\<^sub>1" using rs.prj_tab_agreement Maps.PRJ\<^sub>1_def rs.composable isomorphic_implies_hpar(3) by auto ultimately show ?thesis using SPN_r_SPN_s_leg1_eq Maps.comp_char [of rs.R\<^sub>1 "Maps.PRJ\<^sub>1 rs.R\<^sub>0 rs.S\<^sub>1"] rs.R\<^sub>1_def Maps.PRJ\<^sub>1_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_char rs.composable by auto thus "Span.in_hhom (CMP r s) (SPN.map\<^sub>0 (src s)) (SPN.map\<^sub>0 (trg r))" using assms VV.ide_char VV.arr_char VV.in_hom_char SPN.FF_def Span.src_dom [of "CMP r s"] Span.trg_dom [of "CMP r s"] by fastforce 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.map\<^sub>0 (src s)" and "Span.trg (CMP r s) = SPN.map\<^sub>0 (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_char VV.arr_char 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 \\ by (auto simp add: VV.arr_char) qed have RHS: "Span.in_hom ?RHS (HoSPN_SPN.map (VV.dom \\)) (SPNoH.map (VV.cod \\))" using \\ by (auto simp add: VV.arr_char) 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_char VV.arr_char by unfold_locales auto interpret cod_\_\: two_composable_identities_in_bicategory_of_spans V H \ \ src trg \cod ?\\ \cod ?\\ using \\ VV.ide_char VV.arr_char by unfold_locales auto interpret \_\: horizontal_composite_of_arrows_of_tabulations_in_maps V H \ \ src trg \dom ?\\ \tab_of_ide (dom ?\)\ \tab\<^sub>0 (dom ?\)\ \tab\<^sub>1 (dom ?\)\ \dom ?\\ \tab_of_ide (dom ?\)\ \tab\<^sub>0 (dom ?\)\ \tab\<^sub>1 (dom ?\)\ \cod ?\\ \tab_of_ide (cod ?\)\ \tab\<^sub>0 (cod ?\)\ \tab\<^sub>1 (cod ?\)\ \cod ?\\ \tab_of_ide (cod ?\)\ \tab\<^sub>0 (cod ?\)\ \tab\<^sub>1 (cod ?\)\ ?\ ?\ using \\ VV.arr_char 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 ?\\)\ \tab\<^sub>0 (dom ?\\)\ \tab\<^sub>1 (dom ?\\)\ \cod ?\\\ \tab_of_ide (cod ?\\)\ \tab\<^sub>0 (cod ?\\)\ \tab\<^sub>1 (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_char VV.cod_char CMP_def by simp also have "... = \\cod_\_\.cmp\\ \ Span.chine_hcomp (SPN (fst \\)) (SPN (snd \\))" using \\ VV.arr_char SPN.FF_def Span.hcomp_def by simp finally show ?thesis by blast qed have Chn_RHS_eq: "Chn ?RHS = Maps.MkArr (src (tab\<^sub>0 (dom ?\ \ dom ?\))) (src (tab\<^sub>0 (cod ?\ \ cod ?\))) \\\.chine\ \ Maps.MkArr (src dom_\_\.\\.p\<^sub>0) (src (tab_of_ide (dom ?\ \ dom ?\))) \dom_\_\.cmp\" proof - have "Chn ?RHS = Chn (SPN (?\ \ ?\)) \ Maps.MkArr (src dom_\_\.\\.p\<^sub>0) (src (tab_of_ide (dom ?\ \ dom ?\))) \dom_\_\.cmp\" using \\ RHS Span.vcomp_def VV.arr_char CMP_def Span.arr_char Span.not_arr_Null by auto moreover have "Chn (SPN (?\ \ ?\)) = Maps.MkArr (src (tab\<^sub>0 (dom ?\ \ dom ?\))) (src (tab\<^sub>0 (cod ?\ \ cod ?\))) \\\.chine\" proof - have "Chn (SPN (?\ \ ?\)) = Maps.MkArr (src (tab\<^sub>0 (dom ?\ \ dom ?\))) (src (tab\<^sub>0 (cod ?\ \ cod ?\))) \spn ?\\\" using \\ SPN_def by simp also have "... = Maps.MkArr (src (tab\<^sub>0 (dom ?\ \ dom ?\))) (src (tab\<^sub>0 (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_\_\.\\.p\<^sub>0) (src (tab_of_ide (cod ?\ \ cod ?\))) \cod_\_\.cmp\ \ Span.chine_hcomp (SPN ?\) (SPN ?\)" let ?Chn_RHS = "Maps.MkArr (src (tab\<^sub>0 (dom ?\ \ dom ?\))) (src (tab\<^sub>0 (cod ?\ \ cod ?\))) \\\.chine\ \ Maps.MkArr (src dom_\_\.\\.p\<^sub>0) (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 \r\<^sub>0\<^sup>* \ s\<^sub>1\. 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_char 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_char 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_\_\.S\<^sub>1_def dom_\_\.RS_simps(1) dom_\_\.R\<^sub>0_def Maps.pbdom_in_hom by simp thus ?thesis using \\ VV.arr_char 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_char by simp also have "... = Maps.Dom (Maps.dom (Span.chine_hcomp (SPN (dom ?\)) (SPN (dom ?\))))" using \\ VV.arr_char 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_\_\.\\.p\<^sub>0" 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_char Maps.seq_char by auto also have "... = src (tab\<^sub>0 (cod ?\ \ cod ?\))" using \\ VV.arr_char 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_char Maps.Map_comp Maps.comp_char [of "Maps.MkArr (src (tab\<^sub>0 (dom ?\ \ dom ?\))) (src (tab\<^sub>0 (cod ?\ \ cod ?\))) \\\.chine\" "Maps.MkArr (src dom_\_\.\\.p\<^sub>0) (src (tab_of_ide (dom ?\ \ dom ?\))) \dom_\_\.cmp\"] by simp also have "... = \\\.chine \ dom_\_\.cmp\" proof - have "\dom_\_\.cmp\ \ Maps.Hom (src dom_\_\.\\.p\<^sub>0) (src (tab\<^sub>0 (dom ?\ \ dom ?\)))" proof - have "\dom_\_\.cmp\ \ Maps.Hom (src dom_\_\.\\.tab) (src (tab_of_ide (dom ?\ \ dom ?\)))" using \\ VV.arr_char dom_\_\.cmp_props(1-3) by (metis (mono_tags, lifting) equivalence_is_left_adjoint mem_Collect_eq) thus ?thesis using \\ VV.arr_char dom_\\.tab_simps(2) by simp qed moreover have "\\\.chine\ \ Maps.Hom (src (tab\<^sub>0 (dom ?\ \ dom ?\))) (src (tab\<^sub>0 (cod ?\ \ cod ?\)))" using \\ VV.arr_char \\.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_char \\.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_\_\.\\.p\<^sub>1)) (Maps.CLS (tab\<^sub>0 (cod ?\))) (Maps.CLS (tab\<^sub>1 (cod ?\))) (Maps.CLS (spn ?\ \ dom_\_\.\\.p\<^sub>0))))" proof - have "Maps.Map ?Chn_LHS = Maps.Comp \cod_\_\.cmp\ (Maps.Map (Span.chine_hcomp (SPN ?\) (SPN ?\)))" using \\ 1 VV.arr_char Maps.Map_comp cod_\\.tab_simps(2) Maps.comp_char [of "Maps.MkArr (src cod_\_\.\\.p\<^sub>0) (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_\_\.\\.p\<^sub>1)) (Maps.CLS (tab\<^sub>0 (cod ?\))) (Maps.CLS (tab\<^sub>1 (cod ?\))) (Maps.CLS (spn ?\ \ dom_\_\.\\.p\<^sub>0))" proof - have "Maps.PRJ\<^sub>0 (Maps.MkArr (src (tab\<^sub>0 (dom ?\))) (trg ?\) \tab\<^sub>0 (dom ?\)\) (Maps.MkArr (src (tab\<^sub>0 (dom ?\))) (trg ?\) \tab\<^sub>1 (dom ?\)\) = \\dom_\_\.\\.p\<^sub>0\\ \ Maps.PRJ\<^sub>1 (Maps.MkArr (src (tab\<^sub>0 (dom ?\))) (trg ?\) \tab\<^sub>0 (dom ?\)\) (Maps.MkArr (src (tab\<^sub>0 (dom ?\))) (trg ?\) \tab\<^sub>1 (dom ?\)\) = \\dom_\_\.\\.p\<^sub>1\\" proof - interpret X: identity_in_bicategory_of_spans V H \ \ src trg \(tab\<^sub>0 (dom ?\))\<^sup>* \ tab\<^sub>1 (dom ?\)\ using \\ VV.arr_char by (unfold_locales, simp) have "Maps.PRJ\<^sub>0 (Maps.MkArr (src (tab\<^sub>0 (dom ?\))) (trg ?\) \tab\<^sub>0 (dom ?\)\) (Maps.MkArr (src (tab\<^sub>0 (dom ?\))) (trg ?\) \tab\<^sub>1 (dom ?\)\) = \\tab\<^sub>0 ((Maps.REP (Maps.MkArr (src (tab\<^sub>0 (dom ?\))) (trg (snd \\)) \tab\<^sub>0 (dom ?\)\))\<^sup>* \ Maps.REP (Maps.MkArr (src (tab\<^sub>0 (dom ?\))) (trg ?\) \tab\<^sub>1 (dom ?\)\))\\" unfolding Maps.PRJ\<^sub>0_def using \\ VV.arr_char dom_\_\.RS_simps(1) dom_\_\.RS_simps(16) dom_\_\.RS_simps(18) dom_\_\.RS_simps(3) dom_\_\.R\<^sub>0_def dom_\_\.S\<^sub>1_def by auto moreover have "Maps.PRJ\<^sub>1 (Maps.MkArr (src (tab\<^sub>0 (dom ?\))) (trg ?\) \tab\<^sub>0 (dom ?\)\) (Maps.MkArr (src (tab\<^sub>0 (dom ?\))) (trg ?\) \tab\<^sub>1 (dom ?\)\) = \\tab\<^sub>1 ((Maps.REP (Maps.MkArr (src (tab\<^sub>0 (dom ?\))) (trg (snd \\)) \tab\<^sub>0 (dom ?\)\))\<^sup>* \ Maps.REP (Maps.MkArr (src (tab\<^sub>0 (dom ?\))) (trg ?\) \tab\<^sub>1 (dom ?\)\))\\" unfolding Maps.PRJ\<^sub>1_def using \\ VV.arr_char dom_\_\.RS_simps(1) dom_\_\.RS_simps(16) dom_\_\.RS_simps(18) dom_\_\.RS_simps(3) dom_\_\.R\<^sub>0_def dom_\_\.S\<^sub>1_def by auto moreover have "(Maps.REP (Maps.MkArr (src (tab\<^sub>0 (dom ?\))) (trg (snd \\)) \tab\<^sub>0 (dom ?\)\))\<^sup>* \ Maps.REP (Maps.MkArr (src (tab\<^sub>0 (dom ?\))) (trg ?\) \tab\<^sub>1 (dom ?\)\) \ (tab\<^sub>0 (dom ?\))\<^sup>* \ tab\<^sub>1 (dom ?\)" using VV.arr_char \\ dom_\_\.S\<^sub>1_def dom_\_\.s.leg1_simps(3) dom_\_\.s.leg1_simps(4) trg_dom dom_\_\.R\<^sub>0_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_char SPN_def isomorphic_reflexive Maps.comp_CLS [of "spn ?\" dom_\_\.\\.p\<^sub>1 "spn ?\ \ dom_\_\.\\.p\<^sub>1"] Maps.comp_CLS [of "spn ?\" dom_\_\.\\.p\<^sub>0 "spn ?\ \ dom_\_\.\\.p\<^sub>0"] by simp qed moreover have "Maps.Dom (Span.chine_hcomp (SPN ?\) (SPN ?\)) = src dom_\_\.\\.p\<^sub>0" 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_\_\.\\.p\<^sub>1\\ \\tab\<^sub>0 (cod ?\)\\ \\tab\<^sub>1 (cod ?\)\\ \\spn (snd \\) \ dom_\_\.\\.p\<^sub>0\\" 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_\_\.\\.p\<^sub>0) (src cod_\_\.\\.p\<^sub>0)" proof - have "\\_\.chine : src dom_\_\.\\.p\<^sub>0 \ src cod_\_\.\\.p\<^sub>0\" using \\ VV.arr_char 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_\_\.\\.p\<^sub>0) (src (tab\<^sub>0 (cod ?\ \ cod ?\)))" proof - have "\cod_\_\.cmp : src cod_\_\.\\.p\<^sub>0 \ src (tab\<^sub>0 (cod ?\ \ cod ?\))\" using \\ VV.arr_char 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 blast 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 \tab\<^sub>0 (dom ?\) \ dom_\_\.\\.p\<^sub>0\ \tab\<^sub>1 (dom ?\) \ dom_\_\.\\.p\<^sub>1\ \dom ?\\\ \tab_of_ide (dom ?\ \ dom ?\)\ \tab\<^sub>0 (dom ?\ \ dom ?\)\ \tab\<^sub>1 (dom ?\ \ dom ?\)\ \dom ?\\\ using \\ VV.arr_char dom_\_\.cmp_interpretation by simp interpret cod_cmp: identity_arrow_of_tabulations_in_maps V H \ \ src trg \cod ?\\\ cod_\_\.\\.tab \tab\<^sub>0 (cod ?\) \ cod_\_\.\\.p\<^sub>0\ \tab\<^sub>1 (cod ?\) \ cod_\_\.\\.p\<^sub>1\ \cod ?\\\ \tab_of_ide (cod ?\ \ cod ?\)\ \tab\<^sub>0 (cod ?\ \ cod ?\)\ \tab\<^sub>1 (cod ?\ \ cod ?\)\ \cod ?\\\ using \\ VV.arr_char cod_\_\.cmp_interpretation by simp interpret L: vertical_composite_of_arrows_of_tabulations_in_maps V H \ \ src trg \dom ?\\\ \dom_\_\.\\.tab\ \tab\<^sub>0 (dom ?\) \ dom_\_\.\\.p\<^sub>0\ \tab\<^sub>1 (dom ?\) \ dom_\_\.\\.p\<^sub>1\ \dom ?\\\ \tab_of_ide (dom ?\\)\ \tab\<^sub>0 (dom ?\\)\ \tab\<^sub>1 (dom ?\\)\ \cod ?\\\ cod_\\.tab \tab\<^sub>0 (cod ?\\)\ \tab\<^sub>1 (cod ?\\)\ \dom ?\\\ \?\ \ ?\\ using \\ VV.arr_char 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\ \tab\<^sub>0 (dom ?\) \ dom_\_\.\\.p\<^sub>0\ \tab\<^sub>1 (dom ?\) \ dom_\_\.\\.p\<^sub>1\ \cod ?\\\ \cod_\_\.\\.tab\ \tab\<^sub>0 (cod ?\) \ cod_\_\.\\.p\<^sub>0\ \tab\<^sub>1 (cod ?\) \ cod_\_\.\\.p\<^sub>1\ \cod ?\\\ cod_\\.tab \tab\<^sub>0 (cod ?\\)\ \tab\<^sub>1 (cod ?\\)\ \?\ \ ?\\ \cod ?\\\ using \\ VV.arr_char cod_\_\.cmp_in_hom by unfold_locales auto have "\\.chine \ dom_\_\.cmp \ L.chine" using \\ VV.arr_char 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_char 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_char VV.arr_char \.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 \tab\<^sub>0 f\ \tab\<^sub>1 f\ f f.tab \tab\<^sub>0 f\ \tab\<^sub>1 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 \tab\<^sub>0 h\ \tab\<^sub>1 h\ h h.tab \tab\<^sub>0 h\ \tab\<^sub>1 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.PRJ\<^sub>0 Maps.PRJ\<^sub>1 .. 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 \tab\<^sub>0 h \ THfgh.\\.p\<^sub>0\ \tab\<^sub>1 (f \ g) \ THfgh.\\.p\<^sub>1\ 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 \tab\<^sub>0 (g \ h) \ TfHgh.\\.p\<^sub>0\ \tab\<^sub>1 f \ TfHgh.\\.p\<^sub>1\ using TfHgh.\\.composite_is_tabulation by simp interpretation Tfg_Hfg: arrow_of_tabulations_in_maps V H \ \ src trg \f \ g\ Tfg.\\.tab \tab\<^sub>0 g \ Tfg.\\.p\<^sub>0\ \tab\<^sub>1 f \ Tfg.\\.p\<^sub>1\ \f \ g\ \tab_of_ide (f \ g)\ \tab\<^sub>0 (f \ g)\ \tab\<^sub>1 (f \ g)\ \f \ g\ by unfold_locales auto interpretation Tgh_Hgh: arrow_of_tabulations_in_maps V H \ \ src trg \g \ h\ Tgh.\\.tab \tab\<^sub>0 h \ Tgh.\\.p\<^sub>0\ \tab\<^sub>1 g \ Tgh.\\.p\<^sub>1\ \g \ h\ \tab_of_ide (g \ h)\ \tab\<^sub>0 (g \ h)\ \tab\<^sub>1 (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 \tab\<^sub>0 h \ THfgh.\\.p\<^sub>0\ \tab\<^sub>1 (f \ g) \ THfgh.\\.p\<^sub>1\ \(f \ g) \ h\ \tab_of_ide ((f \ g) \ h)\ \tab\<^sub>0 ((f \ g) \ h)\ \tab\<^sub>1 ((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 \tab\<^sub>0 (g \ h) \ TfHgh.\\.p\<^sub>0\ \tab\<^sub>1 f \ TfHgh.\\.p\<^sub>1\ \f \ g \ h\ \tab_of_ide (f \ g \ h)\ \tab\<^sub>0 (f \ g \ h)\ \tab\<^sub>1 (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 \tab\<^sub>0 g \ Tfg.\\.p\<^sub>0\ \tab\<^sub>1 f \ Tfg.\\.p\<^sub>1\ h \tab_of_ide h\ \tab\<^sub>0 h\ \tab\<^sub>1 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 \tab\<^sub>0 g \ Tfg.\\.p\<^sub>0\ \tab\<^sub>1 f \ Tfg.\\.p\<^sub>1\ h \tab_of_ide h\ \tab\<^sub>0 h\ \tab\<^sub>1 h\ \f \ g\ \tab_of_ide (f \ g)\ \tab\<^sub>0 (f \ g)\ \tab\<^sub>1 (f \ g)\ h \tab_of_ide h\ \tab\<^sub>0 h\ \tab\<^sub>1 h\ \f \ g\ h .. interpretation TfTgh: composite_tabulation_in_maps V H \ \ src trg f \tab_of_ide f\ \tab\<^sub>0 f\ \tab\<^sub>1 f\ \g \ h\ Tgh.\\.tab \tab\<^sub>0 h \ Tgh.\\.p\<^sub>0\ \tab\<^sub>1 g \ Tgh.\\.p\<^sub>1\ 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\ \tab\<^sub>0 f\ \tab\<^sub>1 f\ \g \ h\ Tgh.\\.tab \tab\<^sub>0 h \ Tgh.\\.p\<^sub>0\ \tab\<^sub>1 g \ Tgh.\\.p\<^sub>1\ f \tab_of_ide f\ \tab\<^sub>0 f\ \tab\<^sub>1 f\ \g \ h\ \tab_of_ide (g \ h)\ \tab\<^sub>0 (g \ h)\ \tab\<^sub>1 (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\ \tab\<^sub>0 f\ \tab\<^sub>1 f\ \g \ h\ Tgh.\\.tab \tab\<^sub>0 h \ Tgh.\\.p\<^sub>0\ \tab\<^sub>1 g \ Tgh.\\.p\<^sub>1\ f \tab_of_ide f\ \tab\<^sub>0 f\ \tab\<^sub>1 f\ \g \ h\ Tgh.\\.tab \tab\<^sub>0 h \ Tgh.\\.p\<^sub>0\ \tab\<^sub>1 g \ Tgh.\\.p\<^sub>1\ 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.\\.p\<^sub>1. *) interpretation TTfgh_TfTgh: arrow_of_tabulations_in_maps V H \ \ src trg \(f \ g) \ h\ TTfgh.tab \tab\<^sub>0 h \ TTfgh.p\<^sub>0\ \(tab\<^sub>1 f \ Tfg.\\.p\<^sub>1) \ TTfgh.p\<^sub>1\ \f \ g \ h\ TfTgh.tab \(tab\<^sub>0 h \ Tgh.\\.p\<^sub>0) \ TfTgh.p\<^sub>0\ \tab\<^sub>1 f \ TfTgh.p\<^sub>1\ \\[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)))\ \tab\<^sub>0 (dom (\ (f, g, h)))\ \tab\<^sub>1 (dom (\ (f, g, h)))\ \cod (\ (f, g, h))\ \tab_of_ide (cod (\ (f, g, h)))\ \tab\<^sub>0 (cod (\ (f, g, h)))\ \tab\<^sub>1 (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)) (tab\<^sub>0 ((f \ g) \ h)) (tab\<^sub>1 ((f \ g) \ h)) (f \ g \ h) (tab_of_ide (f \ g \ h)) (tab\<^sub>0 (f \ g \ h)) (tab\<^sub>1 (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)))) (tab\<^sub>0 (dom (\ (f, g, h)))) (tab\<^sub>1 (dom (\ (f, g, h)))) (cod (\ (f, g, h))) (tab_of_ide (cod (\ (f, g, h)))) (tab\<^sub>0 (cod (\ (f, g, h)))) (tab\<^sub>1 (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.PRJ\<^sub>0 Maps.PRJ\<^sub>1 \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.Prj\<^sub>1\<^sub>1 = \\Tfg.\\.p\<^sub>1 \ TTfgh.p\<^sub>1\\" and "SPN_fgh.Prj\<^sub>0\<^sub>1 = \\Tfg.\\.p\<^sub>0 \ TTfgh.p\<^sub>1\\" and "SPN_fgh.Prj\<^sub>0 = \\TTfgh.p\<^sub>0\\" and "SPN_fgh.Prj\<^sub>1 = \\TfTgh.p\<^sub>1\\" and "SPN_fgh.Prj\<^sub>1\<^sub>0 = \\Tgh.\\.p\<^sub>1 \ TfTgh.p\<^sub>0\\" and "SPN_fgh.Prj\<^sub>0\<^sub>0 = \\Tgh.\\.p\<^sub>0 \ TfTgh.p\<^sub>0\\" proof - show "SPN_fgh.Prj\<^sub>1\<^sub>1 = \\Tfg.\\.p\<^sub>1 \ TTfgh.p\<^sub>1\\" proof - have "ide (Tfg.\\.p\<^sub>1 \ TTfgh.p\<^sub>1)" 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) tab\<^sub>1_simps(1)) thus ?thesis using fg gh Tfg.\\.prj_char TTfgh.prj_char isomorphic_reflexive Maps.comp_CLS [of "tab\<^sub>0 g" Tfg.\\.p\<^sub>0 "tab\<^sub>0 g \ Tfg.\\.p\<^sub>0"] Maps.comp_CLS [of Tfg.\\.p\<^sub>1 TTfgh.p\<^sub>1 "Tfg.\\.p\<^sub>1 \ TTfgh.p\<^sub>1"] by (simp add: TTfgh.composable Tfg.\\.T0.antipar(2)) qed show "SPN_fgh.Prj\<^sub>0\<^sub>1 = \\Tfg.\\.p\<^sub>0 \ TTfgh.p\<^sub>1\\" proof - have "ide (Tfg.\\.p\<^sub>0 \ TTfgh.p\<^sub>1)" by (metis TTfgh.leg1_simps(2) bicategory_of_spans.tab\<^sub>0_simps(1) bicategory_of_spans.tab\<^sub>1_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) tab\<^sub>1_simps(3)) thus ?thesis using fg gh Tfg.\\.prj_char TTfgh.prj_char isomorphic_reflexive Maps.comp_CLS [of "tab\<^sub>0 g" Tfg.\\.p\<^sub>0 "tab\<^sub>0 g \ Tfg.\\.p\<^sub>0"] Maps.comp_CLS [of Tfg.\\.p\<^sub>0 TTfgh.p\<^sub>1 "Tfg.\\.p\<^sub>0 \ TTfgh.p\<^sub>1"] by (simp add: Tfg.\\.T0.antipar(2) THfgh.composable) qed show "SPN_fgh.Prj\<^sub>0 = \\TTfgh.p\<^sub>0\\" using isomorphic_reflexive TTfgh.prj_char Tfg.\\.prj_char Maps.comp_CLS [of "tab\<^sub>0 g" Tfg.\\.p\<^sub>0 "tab\<^sub>0 g \ Tfg.\\.p\<^sub>0"] by (simp add: Tfg.composable) show "SPN_fgh.Prj\<^sub>1 = \\TfTgh.p\<^sub>1\\" using Tgh.\\.prj_char isomorphic_reflexive Tgh.composable Maps.comp_CLS [of "tab\<^sub>1 g" Tgh.\\.p\<^sub>1 "tab\<^sub>1 g \ Tgh.\\.p\<^sub>1"] TfTgh.prj_char by simp show "SPN_fgh.Prj\<^sub>1\<^sub>0 = \\Tgh.\\.p\<^sub>1 \ TfTgh.p\<^sub>0\\" using fg gh Tgh.\\.prj_char TfTgh.prj_char isomorphic_reflexive Maps.comp_CLS [of "tab\<^sub>1 g" "prj\<^sub>1 (tab\<^sub>1 h) (tab\<^sub>0 g)" "tab\<^sub>1 g \ Tgh.\\.p\<^sub>1"] Maps.comp_CLS [of Tgh.\\.p\<^sub>1 TfTgh.p\<^sub>0 "Tgh.\\.p\<^sub>1 \ TfTgh.p\<^sub>0"] by simp show "SPN_fgh.Prj\<^sub>0\<^sub>0 = \\Tgh.\\.p\<^sub>0 \ TfTgh.p\<^sub>0\\" using fg gh Tgh.\\.prj_char TfTgh.prj_char isomorphic_reflexive Maps.comp_CLS [of "tab\<^sub>1 g" "Tgh.\\.p\<^sub>1" "tab\<^sub>1 g \ Tgh.\\.p\<^sub>1"] Maps.comp_CLS [of Tgh.\\.p\<^sub>0 TfTgh.p\<^sub>0 "Tgh.\\.p\<^sub>0 \ TfTgh.p\<^sub>0"] 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_char VV.arr_char 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_char VVV.ide_char VV.arr_char VV.ide_char 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.map\<^sub>0 (src h)) (SPN.map\<^sub>0 (trg h))" using SPN.preserves_src SPN.preserves_trg by simp show 2: "Span.in_hhom (CMP f g) (SPN.map\<^sub>0 (trg h)) (SPN.map\<^sub>0 (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_char VV.arr_char \.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_char VV.arr_char SPN.FF_def by simp also have "... = Span.dom (Span.assoc (SPN f) (SPN g) (SPN h))" using fg gh VVV.arr_char VVV.ide_char VV.arr_char VV.ide_char 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_char VVV.arr_char VV.ide_char VV.arr_char HoVH_def by simp also have "... = Span.cod RHS" using arr_RHS fg gh \.map_simp_ide VV.ide_char VV.arr_char 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_char VV.arr_char Span.Chn_vcomp by auto moreover have "Chn (SPN \[f, g, h]) = Maps.CLS HHfgh_HfHgh.chine" using fg gh SPN_def VVV.arr_char VV.arr_char 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.PRJ\<^sub>1 \\tab\<^sub>0 g \ Tfg.\\.p\<^sub>0\\ \\tab\<^sub>1 h\\) \\tab\<^sub>0 (f \ g)\\ \\tab\<^sub>1 h\\ (\\spn h\\ \ Maps.PRJ\<^sub>0 \\tab\<^sub>0 g \ Tfg.\\.p\<^sub>0\\ \\tab\<^sub>1 h\\)" proof - have "\\tab\<^sub>0 g\\ \ \\Tfg.\\.p\<^sub>0\\ = \\tab\<^sub>0 g \ Tfg.\\.p\<^sub>0\\" using fg isomorphic_reflexive Maps.comp_CLS [of "tab\<^sub>0 g" "Tfg.\\.p\<^sub>0" "tab\<^sub>0 g \ Tfg.\\.p\<^sub>0"] by simp moreover have "span_in_category.apex Maps.comp \Leg0 = \\tab\<^sub>0 h\\, Leg1 = \\tab\<^sub>1 h\\\ = \\spn h\\" proof - interpret h: span_in_category Maps.comp \\Leg0 = \\tab\<^sub>0 h\\, Leg1 = \\tab\<^sub>1 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)\ \tab\<^sub>0 (dom h)\ \tab\<^sub>1 (dom h)\ \cod h\ \tab_of_ide (cod h)\ \tab\<^sub>0 (cod h)\ \tab\<^sub>1 (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 (tab\<^sub>0 h))" using h.is_induced_map_iff dom_h.\_eq_\ h.apex_is_induced_by_cell by force hence "src (tab\<^sub>0 h) \ h.chine" using h.chine_is_induced_map h.induced_map_unique by simp thus "\src (tab\<^sub>0 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.p\<^sub>1\\ \\tab\<^sub>0 (f \ g)\\ \\tab\<^sub>1 h\\ \\h.chine \ TTfgh.p\<^sub>0\\" using TTfgh_THfgh.CLS_chine by simp also have "... = Maps.tuple (\\Tfg_Hfg.chine\\ \ \\TTfgh.p\<^sub>1\\) \\tab\<^sub>0 (f \ g)\\ \\tab\<^sub>1 h\\ (\\h.chine\\ \ \\TTfgh.p\<^sub>0\\)" proof - have "\\Tfg_Hfg.chine \ TTfgh.p\<^sub>1\\ = \\Tfg_Hfg.chine\\ \ \\TTfgh.p\<^sub>1\\" proof - have "is_left_adjoint TTfgh.p\<^sub>1" using Tfg.\\.T0.antipar(2) THfgh.composable by simp moreover have "Tfg_Hfg.chine \ TTfgh.p\<^sub>1 \ Tfg_Hfg.chine \ TTfgh.p\<^sub>1" 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.p\<^sub>1 "Tfg_Hfg.chine \ TTfgh.p\<^sub>1"] by simp qed moreover have "\\h.chine \ TTfgh.p\<^sub>0\\ = \\h.chine\\ \ \\TTfgh.p\<^sub>0\\" proof - have "is_left_adjoint TTfgh.p\<^sub>0" by (simp add: Tfg.\\.T0.antipar(2) THfgh.composable) moreover have "h.chine \ TTfgh.p\<^sub>0 \ h.chine \ TTfgh.p\<^sub>0" 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.p\<^sub>0 "h.chine \ TTfgh.p\<^sub>0"] by simp qed ultimately show ?thesis by argo qed also have "... = Maps.tuple (\\Tfg.cmp\\ \ Maps.PRJ\<^sub>1 \\tab\<^sub>0 g \ Tfg.\\.p\<^sub>0\\ \\tab\<^sub>1 h\\) \\tab\<^sub>0 (f \ g)\\ \\tab\<^sub>1 h\\ (\\spn h\\ \ Maps.PRJ\<^sub>0 \\tab\<^sub>0 g \ Tfg.\\.p\<^sub>0\\ \\tab\<^sub>1 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.Prj\<^sub>0\<^sub>1 SPN_fgh.\.leg0 SPN_fgh.\.leg1 SPN_fgh.Prj\<^sub>0" abbreviation tuple_ABC where "tuple_ABC \ Maps.tuple SPN_fgh.Prj\<^sub>1\<^sub>1 SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1) tuple_BC" abbreviation tuple_BC' where "tuple_BC' \ Maps.tuple \\Tfg.\\.p\<^sub>0 \ TTfgh.p\<^sub>1\\ \\tab\<^sub>0 g\\ \\tab\<^sub>1 h\\ \\TTfgh.p\<^sub>0\\" abbreviation tuple_ABC' where "tuple_ABC' \ Maps.tuple \\Tfg.\\.p\<^sub>1 \ TTfgh.p\<^sub>1\\ \\tab\<^sub>0 f\\ \\tab\<^sub>1 g \ Tgh.\\.p\<^sub>1\\ tuple_BC'" lemma csq: shows "Maps.commutative_square SPN_fgh.\.leg0 SPN_fgh.\.leg1 SPN_fgh.Prj\<^sub>0\<^sub>1 SPN_fgh.Prj\<^sub>0" and "Maps.commutative_square SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1) SPN_fgh.Prj\<^sub>1\<^sub>1 tuple_BC" proof - show 1: "Maps.commutative_square SPN_fgh.\.leg0 SPN_fgh.\.leg1 SPN_fgh.Prj\<^sub>0\<^sub>1 SPN_fgh.Prj\<^sub>0" proof show "Maps.cospan SPN_fgh.\.leg0 SPN_fgh.\.leg1" using SPN_fgh.\\.legs_form_cospan(1) by simp show "Maps.span SPN_fgh.Prj\<^sub>0\<^sub>1 SPN_fgh.Prj\<^sub>0" using SPN_fgh.prj_simps(2-3,5-6) by presburger show "Maps.dom SPN_fgh.\.leg0 = Maps.cod SPN_fgh.Prj\<^sub>0\<^sub>1" 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.Prj\<^sub>0\<^sub>1 = SPN_fgh.\.leg1 \ SPN_fgh.Prj\<^sub>0" 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.\\.prj\<^sub>1) SPN_fgh.Prj\<^sub>1\<^sub>1 tuple_BC" proof show "Maps.cospan SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1)" using fg gh SPN_fgh.prj_simps(10) by blast show "Maps.span SPN_fgh.Prj\<^sub>1\<^sub>1 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.Prj\<^sub>1\<^sub>1" using fg gh SPN_f.dom.leg_simps(2) SPN_fgh.prj_simps(7) by auto show "SPN_fgh.\.leg0 \ SPN_fgh.Prj\<^sub>1\<^sub>1 = (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1) \ 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.Prj\<^sub>1\<^sub>1 = \\Tfg.\\.p\<^sub>1 \ TTfgh.p\<^sub>1\\" using prj_char by simp moreover have "SPN_fgh.\.leg0 = \\tab\<^sub>0 f\\" by simp moreover have "SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1 = \\tab\<^sub>1 g \ Tgh.\\.p\<^sub>1\\" using Tgh.\\.prj_char isomorphic_reflexive Tgh.composable Maps.comp_CLS [of "tab\<^sub>1 g" Tgh.\\.p\<^sub>1 "tab\<^sub>1 g \ Tgh.\\.p\<^sub>1"] 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.p\<^sub>0)) (Maps.MkIde (src Tgh.\\.p\<^sub>0))" proof show 1: "Maps.arr tuple_BC" using csq(1) by simp have 2: "Maps.commutative_square \\tab\<^sub>0 g\\ \\tab\<^sub>1 h\\ \\Tfg.\\.p\<^sub>0 \ TTfgh.p\<^sub>1\\ \\TTfgh.p\<^sub>0\\" by (metis Tfg.S\<^sub>0_def Tfg.span_legs_eq(3) Tgh.S\<^sub>1_def Tgh.span_legs_eq(4) csq(1) prj_char(2) prj_char(3)) show "Maps.dom tuple_BC = Maps.MkIde (src TTfgh.p\<^sub>0)" proof - have "Maps.dom tuple_BC' = Maps.dom \\Tfg.\\.p\<^sub>0 \ TTfgh.p\<^sub>1\\" 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.p\<^sub>0)" 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.p\<^sub>0)" by blast thus ?thesis using tuple_ABC_eq_ABC' by simp qed show "Maps.cod tuple_BC = Maps.MkIde (src Tgh.\\.p\<^sub>0)" proof - have "Maps.cod tuple_BC' = Maps.pbdom \\tab\<^sub>0 g\\ \\tab\<^sub>1 h\\" using 1 2 fg gh Maps.tuple_in_hom by blast also have "... = Maps.MkIde (src Tgh.\\.p\<^sub>0)" 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.S\<^sub>0_def Tfg.span_legs_eq(3) Tgh.S\<^sub>1_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.p\<^sub>0)) (Maps.MkIde (src TfTgh.p\<^sub>0))" 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.p\<^sub>0)" 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.p\<^sub>0)" 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.p\<^sub>0)" 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.p\<^sub>0)" 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_char VV.arr_char 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_char VV.arr_char 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_char VVV.arr_char VV.ide_char VV.arr_char 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.p\<^sub>1\\) \\tab\<^sub>0 f\\ \\tab\<^sub>1 (g \ h)\\ (\\Tgh_Hgh.chine\\ \ \\TfTgh.p\<^sub>0\\)" proof - interpret f: span_in_category Maps.comp \\Leg0 = Maps.MkArr (src (tab\<^sub>0 f)) (trg g) \tab\<^sub>0 f\, Leg1 = Maps.MkArr (src (tab\<^sub>0 f)) (trg f) \tab\<^sub>1 f\\\ using f.determines_span by (simp add: Tfg.composable) interpret f: arrow_of_tabulations_in_maps V H \ \ src trg f f.tab \tab\<^sub>0 f\ \tab\<^sub>1 f\ f f.tab \tab\<^sub>0 f\ \tab\<^sub>1 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.R\<^sub>0_def Tfg.composable by auto show "Maps.Cod f.apex = Maps.Cod \\f.chine\\" using f.apex_def Tfg.RS_simps(2) Tfg.R\<^sub>0_def Tfg.composable by auto show "Maps.Map f.apex = Maps.Map \\f.chine\\" proof - have "Maps.Map f.apex = \src (tab\<^sub>0 f)\" using f.apex_def Maps.dom_char Tfg.RS_simps(2) Tfg.R\<^sub>0_def Tfg.composable by auto also have "... = \f.chine\" proof (intro iso_class_eqI) have "f.is_induced_map (src (tab\<^sub>0 f))" using f.apex_is_induced_by_cell comp_cod_arr by auto thus "src (tab\<^sub>0 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 "tab\<^sub>1 g" Tgh.\\.p\<^sub>1 "tab\<^sub>1 g \ Tgh.\\.p\<^sub>1"] Tgh.cmp_def TfTgh.prj_char by simp qed also have "... = Maps.tuple \\f.chine \ TfTgh.p\<^sub>1\\ \\tab\<^sub>0 f\\ \\tab\<^sub>1 (g \ h)\\ \\Tgh_Hgh.chine \ TfTgh.p\<^sub>0\\" using isomorphic_reflexive TfHgh.composable f.is_map TfHgh.composable Tgh_Hgh.is_map Maps.comp_CLS [of f.chine TfTgh.p\<^sub>1 "f.chine \ TfTgh.p\<^sub>1"] Maps.comp_CLS [of Tgh_Hgh.chine TfTgh.p\<^sub>0 "Tgh_Hgh.chine \ TfTgh.p\<^sub>0"] 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 g\<^sub>0h\<^sub>1: cospan_of_maps_in_bicategory_of_spans V H \ \ src trg \tab\<^sub>1 h\ \tab\<^sub>0 g\ using gh by unfold_locales auto interpretation f\<^sub>0g\<^sub>1: cospan_of_maps_in_bicategory_of_spans V H \ \ src trg \tab\<^sub>1 g\ \tab\<^sub>0 f\ using fg by unfold_locales auto interpretation f\<^sub>0gh\<^sub>1: cospan_of_maps_in_bicategory_of_spans V H \ \ src trg \tab\<^sub>1 g \ Tgh.\\.p\<^sub>1\ \tab\<^sub>0 f\ using fg gh Tgh.\\.leg1_is_map by unfold_locales auto interpretation fg\<^sub>0h\<^sub>1: cospan_of_maps_in_bicategory_of_spans V H \ \ src trg \tab\<^sub>1 h\ \tab\<^sub>0 g \ Tfg.p\<^sub>0\ using TTfgh.r\<^sub>0s\<^sub>1_is_cospan by simp lemma src_tab_eq: shows "(\\<^sup>-\<^sup>1[f, g, h] \ tab\<^sub>0 h \ TTfgh.p\<^sub>0) \ 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] \ tab\<^sub>0 h \ TTfgh.p\<^sub>0) \ 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] \ tab\<^sub>0 h \ TTfgh.p\<^sub>0)" by (simp add: fg gh) moreover have "inv (\[f, g, h] \ tab\<^sub>0 h \ TTfgh.p\<^sub>0) = \\<^sup>-\<^sup>1[f, g, h] \ tab\<^sub>0 h \ TTfgh.p\<^sub>0" 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] \ tab\<^sub>0 h \ TTfgh.p\<^sub>0" 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.p\<^sub>1 \ TTfgh_TfTgh.chine\\ = \\Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1\\" and "\\Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine\\ = \\Tfg.p\<^sub>0 \ TTfgh.p\<^sub>1\\" and "\\Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine\\ = \\TTfgh.p\<^sub>0\\" proof - have 1: "ide TfTgh.p\<^sub>1" by (simp add: TfTgh.composable) have 2: "ide TTfgh_TfTgh.chine" by simp have 3: "src TfTgh.p\<^sub>1 = trg TTfgh_TfTgh.chine" using TTfgh_TfTgh.chine_in_hom(1) by simp have 4: "src (tab\<^sub>1 f) = trg TfTgh.p\<^sub>1" 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 u\<^sub>f where "u\<^sub>f = g \ h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0" define w\<^sub>f where "w\<^sub>f = Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1" define w\<^sub>f' where "w\<^sub>f' = TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine" define \\<^sub>f where "\\<^sub>f = (g \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ (g \ h.tab \ TTfgh.p\<^sub>0) \ (g \ fg\<^sub>0h\<^sub>1.\) \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ (\[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ ((g.tab \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1) \ (f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" define \\<^sub>f' where "\\<^sub>f' = (g \ h \ TTfgh_TfTgh.the_\) \ can (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ (((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine]" define \\<^sub>f where "\\<^sub>f = \[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ TTfgh_TfTgh.the_\ \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" have w\<^sub>f: "ide w\<^sub>f" using w\<^sub>f_def fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by auto have w\<^sub>f_is_map: "is_left_adjoint w\<^sub>f" using w\<^sub>f_def fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by (simp add: left_adjoints_compose) have w\<^sub>f': "ide w\<^sub>f'" unfolding w\<^sub>f'_def by simp have w\<^sub>f'_is_map: "is_left_adjoint w\<^sub>f'" unfolding w\<^sub>f'_def using 3 TTfgh_TfTgh.is_map f\<^sub>0gh\<^sub>1.leg1_is_map by (simp add: left_adjoints_compose) have \\<^sub>f: "\\\<^sub>f : tab\<^sub>0 f \ w\<^sub>f \ u\<^sub>f\" proof (unfold \\<^sub>f_def w\<^sub>f_def u\<^sub>f_def, intro comp_in_homI) show "\\\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] : tab\<^sub>0 f \ Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1 \ (tab\<^sub>0 f \ Tfg.p\<^sub>1) \ TTfgh.p\<^sub>1\" using f\<^sub>0g\<^sub>1.leg1_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.cospan g\<^sub>0h\<^sub>1.cospan by auto show "\f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1 : (tab\<^sub>0 f \ Tfg.p\<^sub>1) \ TTfgh.p\<^sub>1 \ (tab\<^sub>1 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1\" using f\<^sub>0g\<^sub>1.\_in_hom(2) Tfg.\\.T0.antipar(1) by (intro hcomp_in_vhom, auto) show "\(g.tab \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1 : (tab\<^sub>1 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1 \ ((g \ tab\<^sub>0 g) \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1\" using Tfg.\\.T0.antipar(1) by (intro hcomp_in_vhom, auto) show "\\[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1 : ((g \ tab\<^sub>0 g) \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1 \ (g \ tab\<^sub>0 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1\" using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by (intro hcomp_in_vhom, auto) show "\\[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] : (g \ tab\<^sub>0 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1 \ g \ (tab\<^sub>0 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1\" using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by auto show "\g \ fg\<^sub>0h\<^sub>1.\ : g \ (tab\<^sub>0 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1 \ g \ tab\<^sub>1 h \ TTfgh.p\<^sub>0\" using fg\<^sub>0h\<^sub>1.\_in_hom fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by (intro hcomp_in_vhom, auto) show "\g \ h.tab \ TTfgh.p\<^sub>0 : g \ tab\<^sub>1 h \ TTfgh.p\<^sub>0 \ g \ (h \ tab\<^sub>0 h) \ TTfgh.p\<^sub>0\" using gh fg\<^sub>0h\<^sub>1.\_in_hom fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by (intro hcomp_in_vhom, auto) show "\g \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0] : g \ (h \ tab\<^sub>0 h) \ TTfgh.p\<^sub>0 \ g \ h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0\" using gh fg\<^sub>0h\<^sub>1.\_in_hom fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by (intro hcomp_in_vhom, auto) qed have \\<^sub>f': "\\\<^sub>f' : tab\<^sub>0 f \ w\<^sub>f' \ u\<^sub>f\" proof (unfold \\<^sub>f'_def w\<^sub>f'_def u\<^sub>f_def, intro comp_in_homI) show "\\\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] : tab\<^sub>0 f \ TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine \ (tab\<^sub>0 f \ TfTgh.p\<^sub>1) \ TTfgh_TfTgh.chine\" using "1" "2" "3" "4" assoc'_in_hom(2) f.ide_u f.leg1_simps(3) by auto show "\f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine : (tab\<^sub>0 f \ TfTgh.p\<^sub>1) \ TTfgh_TfTgh.chine \ ((tab\<^sub>1 g \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine\" using f\<^sub>0gh\<^sub>1.\_in_hom(2) by (intro hcomp_in_vhom, auto) show "\((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine : ((tab\<^sub>1 g \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine \ (((g \ tab\<^sub>0 g) \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine\" using f\<^sub>0gh\<^sub>1.cospan g\<^sub>0h\<^sub>1.cospan by (intro hcomp_in_vhom, auto) show "\(\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine : (((g \ tab\<^sub>0 g) \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine \ ((g \ tab\<^sub>0 g \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine\" using f\<^sub>0gh\<^sub>1.cospan g\<^sub>0h\<^sub>1.cospan by (intro hcomp_in_vhom, auto) show "\((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine : ((g \ tab\<^sub>0 g \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine \ ((g \ tab\<^sub>1 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine\" using f\<^sub>0gh\<^sub>1.cospan g\<^sub>0h\<^sub>1.cospan g\<^sub>0h\<^sub>1.\_in_hom(2) by (intro hcomp_in_vhom, auto) show "\((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine : ((g \ tab\<^sub>1 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine \ ((g \ (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine\" using f\<^sub>0gh\<^sub>1.cospan g\<^sub>0h\<^sub>1.cospan by (intro hcomp_in_vhom, auto) show "\can (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) : ((g \ (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine \ g \ h \ ((tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine\" using f\<^sub>0gh\<^sub>1.cospan g\<^sub>0h\<^sub>1.cospan by auto show "\g \ h \ TTfgh_TfTgh.the_\ : g \ h \ ((tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine \ g \ h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0\" using f\<^sub>0gh\<^sub>1.cospan g\<^sub>0h\<^sub>1.cospan TTfgh_TfTgh.the_\_in_hom by (intro hcomp_in_vhom, auto) qed have \\<^sub>f: "\\\<^sub>f : tab\<^sub>1 f \ w\<^sub>f \ tab\<^sub>1 f \ w\<^sub>f'\" proof (unfold \\<^sub>f_def w\<^sub>f_def w\<^sub>f'_def, intro comp_in_homI) show "\\\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] : tab\<^sub>1 f \ Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1 \ (tab\<^sub>1 f \ Tfg.p\<^sub>1) \ TTfgh.p\<^sub>1\" using TTfgh.leg1_in_hom(2) assoc'_in_hom by auto show "\TTfgh_TfTgh.the_\ : (tab\<^sub>1 f \ Tfg.p\<^sub>1) \ TTfgh.p\<^sub>1 \ (tab\<^sub>1 f \ TfTgh.p\<^sub>1) \ TTfgh_TfTgh.chine\" using TTfgh_TfTgh.the_\_in_hom TTfgh_TfTgh.the_\_props by simp show "\\[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] : (tab\<^sub>1 f \ TfTgh.p\<^sub>1) \ TTfgh_TfTgh.chine \ tab\<^sub>1 f \ TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine\" using 1 2 3 4 by auto qed have iso_\\<^sub>f: "iso \\<^sub>f" unfolding \\<^sub>f_def using 1 2 3 4 \\<^sub>f \\<^sub>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 f\<^sub>0g\<^sub>1.cospan g.ide_leg1 h.ide_leg1 h.leg1_simps(4) hcomp_in_vhomE ide_hcomp iso_assoc' tab\<^sub>1_simps(1)) using TTfgh_TfTgh.the_\_props(2) f.ide_leg1 iso_assoc by blast+ have u\<^sub>f: "ide u\<^sub>f" using \\<^sub>f ide_cod by blast have w\<^sub>f_in_hhom: "in_hhom w\<^sub>f (src u\<^sub>f) (src (tab\<^sub>0 f))" using u\<^sub>f w\<^sub>f u\<^sub>f_def w\<^sub>f_def by simp have w\<^sub>f'_in_hhom: "in_hhom w\<^sub>f' (src u\<^sub>f) (src (tab\<^sub>0 f))" using u\<^sub>f w\<^sub>f' w\<^sub>f'_def u\<^sub>f_def by simp have 5: "\!\. \\ : w\<^sub>f \ w\<^sub>f'\ \ \\<^sub>f = tab\<^sub>1 f \ \ \ \\<^sub>f = \\<^sub>f' \ (tab\<^sub>0 f \ \)" proof - have eq\<^sub>f: "f.composite_cell w\<^sub>f \\<^sub>f = f.composite_cell w\<^sub>f' \\<^sub>f' \ \\<^sub>f" proof - text \ I don't see any alternative here to just grinding out the calculation. The idea is to bring \f.composite_cell w\<^sub>f \\<^sub>f\ into a form in which \src_tab_eq\ can be applied to eliminate \\\<^sub>f\ in favor of \\\<^sub>f'\. \ have "f.composite_cell w\<^sub>f \\<^sub>f = (f \ g \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ (f \ g \ h.tab \ TTfgh.p\<^sub>0) \ (f \ g \ fg\<^sub>0h\<^sub>1.\) \ (f \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ (f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ (f \ (g.tab \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1) \ (f \ f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ (f \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]) \ \[f, tab\<^sub>0 f, Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1] \ (f.tab \ Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1)" unfolding w\<^sub>f_def \\<^sub>f_def using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps Tgh.composable whisker_left comp_assoc by simp (* 20 sec *) also have "... = (\[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ \\<^sup>-\<^sup>1[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ (f \ g \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0])) \ (f \ g \ h.tab \ TTfgh.p\<^sub>0) \ (f \ g \ fg\<^sub>0h\<^sub>1.\) \ (f \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ (f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ (f \ (g.tab \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1) \ (f \ f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ (f \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]) \ \[f, tab\<^sub>0 f, Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1] \ (f.tab \ Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1)" proof - have "(\[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ \\<^sup>-\<^sup>1[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0]) \ (f \ g \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) = f \ g \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>1_simps comp_cod_arr comp_assoc_assoc' by simp thus ?thesis using comp_assoc by simp qed also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ (\\<^sup>-\<^sup>1[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ (f \ g \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0])) \ (f \ g \ h.tab \ TTfgh.p\<^sub>0) \ (f \ g \ fg\<^sub>0h\<^sub>1.\) \ (f \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ (f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ (f \ (g.tab \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1) \ (f \ f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ (f \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]) \ \[f, tab\<^sub>0 f, Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1] \ (f.tab \ Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1)" using comp_assoc by presburger also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ ((f \ g) \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ (\\<^sup>-\<^sup>1[f, g, (h \ tab\<^sub>0 h) \ TTfgh.p\<^sub>0] \ (f \ g \ h.tab \ TTfgh.p\<^sub>0)) \ (f \ g \ fg\<^sub>0h\<^sub>1.\) \ (f \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ (f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ (f \ (g.tab \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1) \ (f \ f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ (f \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]) \ \[f, tab\<^sub>0 f, Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1] \ (f.tab \ Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1)" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps comp_assoc assoc'_naturality [of f g "\[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]"] by simp also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ ((f \ g) \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ ((f \ g) \ h.tab \ TTfgh.p\<^sub>0) \ (\\<^sup>-\<^sup>1[f, g, tab\<^sub>1 h \ TTfgh.p\<^sub>0] \ (f \ g \ fg\<^sub>0h\<^sub>1.\)) \ (f \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ (f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ (f \ (g.tab \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1) \ (f \ f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ (f \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]) \ \[f, tab\<^sub>0 f, Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1] \ (f.tab \ Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1)" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps comp_assoc assoc'_naturality [of f g "h.tab \ TTfgh.p\<^sub>0"] by simp also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ ((f \ g) \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ ((f \ g) \ h.tab \ TTfgh.p\<^sub>0) \ ((f \ g) \ fg\<^sub>0h\<^sub>1.\) \ \\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1] \ (f \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ (f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ (f \ (g.tab \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1) \ (f \ f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ ((f \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]) \ \[f, tab\<^sub>0 f, Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1]) \ (f.tab \ Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1)" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps comp_assoc assoc'_naturality [of f g fg\<^sub>0h\<^sub>1.\] by simp also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ ((f \ g) \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ ((f \ g) \ h.tab \ TTfgh.p\<^sub>0) \ ((f \ g) \ fg\<^sub>0h\<^sub>1.\) \ \\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1] \ (f \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ (f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ (f \ (g.tab \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1) \ ((f \ f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ \[f, tab\<^sub>0 f \ Tfg.p\<^sub>1, TTfgh.p\<^sub>1]) \ (\[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[f \ tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \ (f.tab \ Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1)" proof - have "(f \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]) \ \[f, tab\<^sub>0 f, Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1] = \[f, tab\<^sub>0 f \ Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \ (\[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[f \ tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" proof - have "(f \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]) \ \[f, tab\<^sub>0 f, Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1] = \(\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\tab\<^sub>0 f\<^bold>\, \<^bold>\Tfg.p\<^sub>1\<^bold>\, \<^bold>\TTfgh.p\<^sub>1\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\f\<^bold>\, \<^bold>\tab\<^sub>0 f\<^bold>\, \<^bold>\Tfg.p\<^sub>1\<^bold>\ \<^bold>\ \<^bold>\TTfgh.p\<^sub>1\<^bold>\\<^bold>]\" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps \'_def \_def by simp also have "... = \\<^bold>\\<^bold>[\<^bold>\f\<^bold>\, \<^bold>\tab\<^sub>0 f\<^bold>\ \<^bold>\ \<^bold>\Tfg.p\<^sub>1\<^bold>\, \<^bold>\TTfgh.p\<^sub>1\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\\<^bold>[\<^bold>\f\<^bold>\, \<^bold>\tab\<^sub>0 f\<^bold>\, \<^bold>\Tfg.p\<^sub>1\<^bold>\\<^bold>] \<^bold>\ \<^bold>\TTfgh.p\<^sub>1\<^bold>\) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 f\<^bold>\, \<^bold>\Tfg.p\<^sub>1\<^bold>\, \<^bold>\TTfgh.p\<^sub>1\<^bold>\\<^bold>]\" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps by (intro E.eval_eqI, simp_all) also have "... = \[f, tab\<^sub>0 f \ Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \ (\[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[f \ tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps \'_def \_def by simp finally show ?thesis by blast qed thus ?thesis using comp_assoc by presburger qed also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ ((f \ g) \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ ((f \ g) \ h.tab \ TTfgh.p\<^sub>0) \ ((f \ g) \ fg\<^sub>0h\<^sub>1.\) \ \\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1] \ (f \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ (f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ ((f \ (g.tab \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1) \ \[f, tab\<^sub>1 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ ((f \ f\<^sub>0g\<^sub>1.\) \ TTfgh.p\<^sub>1) \ (\[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[f \ tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \ (f.tab \ Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1)" proof - have "(f \ f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ \[f, tab\<^sub>0 f \ Tfg.p\<^sub>1, TTfgh.p\<^sub>1] = \[f, tab\<^sub>1 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ ((f \ f\<^sub>0g\<^sub>1.\) \ TTfgh.p\<^sub>1)" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.\_in_hom assoc_naturality [of f f\<^sub>0g\<^sub>1.\ TTfgh.p\<^sub>1] by simp thus ?thesis using comp_assoc by presburger qed also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ ((f \ g) \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ ((f \ g) \ h.tab \ TTfgh.p\<^sub>0) \ ((f \ g) \ fg\<^sub>0h\<^sub>1.\) \ \\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1] \ (f \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ (f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ \[f, (g \ tab\<^sub>0 g) \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ ((f \ (g.tab \ Tfg.p\<^sub>0)) \ TTfgh.p\<^sub>1) \ ((f \ f\<^sub>0g\<^sub>1.\) \ TTfgh.p\<^sub>1) \ (\[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[f \ tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \ (f.tab \ Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1)" proof - have "(f \ (g.tab \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1) \ \[f, tab\<^sub>1 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] = \[f, (g \ tab\<^sub>0 g) \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ ((f \ (g.tab \ Tfg.p\<^sub>0)) \ TTfgh.p\<^sub>1)" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.\_in_hom assoc_naturality [of f "g.tab \ Tfg.p\<^sub>0" TTfgh.p\<^sub>1] by simp thus ?thesis using comp_assoc by presburger qed also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ ((f \ g) \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ ((f \ g) \ h.tab \ TTfgh.p\<^sub>0) \ ((f \ g) \ fg\<^sub>0h\<^sub>1.\) \ \\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1] \ (f \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ ((f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ \[f, (g \ tab\<^sub>0 g) \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ ((f \ (g.tab \ Tfg.p\<^sub>0)) \ TTfgh.p\<^sub>1) \ ((f \ f\<^sub>0g\<^sub>1.\) \ TTfgh.p\<^sub>1) \ (\[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \ TTfgh.p\<^sub>1) \ ((f.tab \ Tfg.p\<^sub>1) \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps assoc'_naturality [of f.tab Tfg.p\<^sub>1 TTfgh.p\<^sub>1] comp_assoc by simp also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ ((f \ g) \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ ((f \ g) \ h.tab \ TTfgh.p\<^sub>0) \ ((f \ g) \ fg\<^sub>0h\<^sub>1.\) \ \\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1] \ (f \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ (f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ \[f, (g \ tab\<^sub>0 g) \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ ((((f \ \\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \ TTfgh.p\<^sub>1) \ ((f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \ TTfgh.p\<^sub>1)) \ ((f \ (g.tab \ Tfg.p\<^sub>0)) \ TTfgh.p\<^sub>1)) \ ((f \ f\<^sub>0g\<^sub>1.\) \ TTfgh.p\<^sub>1) \ (\[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \ TTfgh.p\<^sub>1) \ ((f.tab \ Tfg.p\<^sub>1) \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" proof - have "(((f \ \\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \ TTfgh.p\<^sub>1) \ ((f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \ TTfgh.p\<^sub>1)) \ ((f \ (g.tab \ Tfg.p\<^sub>0)) \ TTfgh.p\<^sub>1) = (f \ (g.tab \ Tfg.p\<^sub>0)) \ TTfgh.p\<^sub>1" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>1_simps comp_cod_arr whisker_right comp_assoc_assoc' whisker_left [of f "\\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tfg.p\<^sub>0]" "\[g, tab\<^sub>0 g, Tfg.p\<^sub>0]"] by simp thus ?thesis using comp_assoc by simp qed also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ ((f \ g) \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ ((f \ g) \ h.tab \ TTfgh.p\<^sub>0) \ ((f \ g) \ fg\<^sub>0h\<^sub>1.\) \ \\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1] \ (f \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ (f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ \[f, (g \ tab\<^sub>0 g) \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ ((f \ \\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \ TTfgh.p\<^sub>1) \ ((f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \ TTfgh.p\<^sub>1) \ ((f \ (g.tab \ Tfg.p\<^sub>0)) \ TTfgh.p\<^sub>1) \ ((f \ f\<^sub>0g\<^sub>1.\) \ TTfgh.p\<^sub>1) \ (\[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \ TTfgh.p\<^sub>1) \ ((f.tab \ Tfg.p\<^sub>1) \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" using comp_assoc by presburger also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ ((f \ g) \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ ((f \ g) \ h.tab \ TTfgh.p\<^sub>0) \ ((f \ g) \ fg\<^sub>0h\<^sub>1.\) \ \\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1] \ (f \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ (f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ \[f, (g \ tab\<^sub>0 g) \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ ((f \ \\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \ TTfgh.p\<^sub>1) \ (((\[f, g, tab\<^sub>0 g \ Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ (\\<^sup>-\<^sup>1[f, g, tab\<^sub>0 g \ Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1)) \ ((f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \ TTfgh.p\<^sub>1)) \ ((f \ (g.tab \ Tfg.p\<^sub>0)) \ TTfgh.p\<^sub>1) \ ((f \ f\<^sub>0g\<^sub>1.\) \ TTfgh.p\<^sub>1) \ (\[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \ TTfgh.p\<^sub>1) \ ((f.tab \ Tfg.p\<^sub>1) \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" proof - have "((\[f, g, tab\<^sub>0 g \ Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ (\\<^sup>-\<^sup>1[f, g, tab\<^sub>0 g \ Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1)) \ ((f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \ TTfgh.p\<^sub>1) = (f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \ TTfgh.p\<^sub>1" using fg fg\<^sub>0h\<^sub>1.p\<^sub>1_simps comp_cod_arr comp_assoc_assoc' whisker_right [of TTfgh.p\<^sub>1 "\[f, g, tab\<^sub>0 g \ Tfg.p\<^sub>0]" "\\<^sup>-\<^sup>1[f, g, tab\<^sub>0 g \ Tfg.p\<^sub>0]"] by simp thus ?thesis using comp_assoc by simp qed also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ ((f \ g) \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ ((f \ g) \ h.tab \ TTfgh.p\<^sub>0) \ ((f \ g) \ fg\<^sub>0h\<^sub>1.\) \ \\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1] \ (f \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ (f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ \[f, (g \ tab\<^sub>0 g) \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ ((f \ \\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \ TTfgh.p\<^sub>1) \ (\[f, g, tab\<^sub>0 g \ Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ ((\\<^sup>-\<^sup>1[f, g, tab\<^sub>0 g \ Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ ((f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \ TTfgh.p\<^sub>1) \ ((f \ (g.tab \ Tfg.p\<^sub>0)) \ TTfgh.p\<^sub>1) \ ((f \ f\<^sub>0g\<^sub>1.\) \ TTfgh.p\<^sub>1) \ (\[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \ TTfgh.p\<^sub>1) \ ((f.tab \ Tfg.p\<^sub>1) \ TTfgh.p\<^sub>1)) \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" using comp_assoc by presburger also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ ((f \ g) \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ ((f \ g) \ h.tab \ TTfgh.p\<^sub>0) \ ((f \ g) \ fg\<^sub>0h\<^sub>1.\) \ (\\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1] \ (f \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ (f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ \[f, (g \ tab\<^sub>0 g) \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ ((f \ \\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \ TTfgh.p\<^sub>1) \ (\[f, g, tab\<^sub>0 g \ Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1)) \ (\\<^sup>-\<^sup>1[f, g, tab\<^sub>0 g \ Tfg.p\<^sub>0] \ (f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \ (f \ (g.tab \ Tfg.p\<^sub>0)) \ (f \ f\<^sub>0g\<^sub>1.\) \ \[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \ (f.tab \ Tfg.p\<^sub>1) \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps whisker_right comp_assoc by simp also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ ((f \ g) \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ ((f \ g) \ h.tab \ TTfgh.p\<^sub>0) \ ((f \ g) \ fg\<^sub>0h\<^sub>1.\) \ \[f \ g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ (\\<^sup>-\<^sup>1[f, g, tab\<^sub>0 g \ Tfg.p\<^sub>0] \ (f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \ (f \ (g.tab \ Tfg.p\<^sub>0)) \ (f \ f\<^sub>0g\<^sub>1.\) \ \[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \ (f.tab \ Tfg.p\<^sub>1) \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" proof - have "\\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1] \ (f \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ (f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ \[f, (g \ tab\<^sub>0 g) \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ ((f \ \\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \ TTfgh.p\<^sub>1) \ (\[f, g, tab\<^sub>0 g \ Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) = \\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\f\<^bold>\, \<^bold>\g\<^bold>\, (\<^bold>\tab\<^sub>0 g\<^bold>\ \<^bold>\ \<^bold>\Tfg.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh.p\<^sub>1\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\g\<^bold>\, \<^bold>\tab\<^sub>0 g\<^bold>\ \<^bold>\ \<^bold>\Tfg.p\<^sub>0\<^bold>\, \<^bold>\TTfgh.p\<^sub>1\<^bold>\\<^bold>]) \<^bold>\ (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\g\<^bold>\, \<^bold>\tab\<^sub>0 g\<^bold>\, \<^bold>\Tfg.p\<^sub>0\<^bold>\\<^bold>] \<^bold>\ \<^bold>\TTfgh.p\<^sub>1\<^bold>\) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\f\<^bold>\, (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 g\<^bold>\) \<^bold>\ \<^bold>\Tfg.p\<^sub>0\<^bold>\, \<^bold>\TTfgh.p\<^sub>1\<^bold>\\<^bold>] \<^bold>\ ((\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\g\<^bold>\, \<^bold>\tab\<^sub>0 g\<^bold>\, \<^bold>\Tfg.p\<^sub>0\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\TTfgh.p\<^sub>1\<^bold>\) \<^bold>\ (\<^bold>\\<^bold>[\<^bold>\f\<^bold>\, \<^bold>\g\<^bold>\, \<^bold>\tab\<^sub>0 g\<^bold>\ \<^bold>\ \<^bold>\Tfg.p\<^sub>0\<^bold>\\<^bold>] \<^bold>\ \<^bold>\TTfgh.p\<^sub>1\<^bold>\)\" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps \'_def \_def by simp also have "... = \\<^bold>\\<^bold>[\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\, \<^bold>\tab\<^sub>0 g\<^bold>\ \<^bold>\ \<^bold>\Tfg.p\<^sub>0\<^bold>\, \<^bold>\TTfgh.p\<^sub>1\<^bold>\\<^bold>]\" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps by (intro E.eval_eqI, auto) also have "... = \[f \ g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps \'_def \_def by simp finally show ?thesis using comp_assoc by presburger qed also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ ((\[f \ g, h, tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ \\<^sup>-\<^sup>1[f \ g, h, tab\<^sub>0 h \ TTfgh.p\<^sub>0]) \ ((f \ g) \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0])) \ ((f \ g) \ h.tab \ TTfgh.p\<^sub>0) \ ((f \ g) \ fg\<^sub>0h\<^sub>1.\) \ \[f \ g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ (\\<^sup>-\<^sup>1[f, g, tab\<^sub>0 g \ Tfg.p\<^sub>0] \ (f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \ (f \ (g.tab \ Tfg.p\<^sub>0)) \ (f \ f\<^sub>0g\<^sub>1.\) \ \[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \ (f.tab \ Tfg.p\<^sub>1) \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" proof - have "(\[f \ g, h, tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ \\<^sup>-\<^sup>1[f \ g, h, tab\<^sub>0 h \ TTfgh.p\<^sub>0]) \ ((f \ g) \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) = ((f \ g) \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0])" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps comp_cod_arr comp_assoc_assoc' by simp thus ?thesis by simp qed also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ \[f \ g, h, tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ (\\<^sup>-\<^sup>1[f \ g, h, tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ ((f \ g) \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ ((f \ g) \ h.tab \ TTfgh.p\<^sub>0) \ ((f \ g) \ fg\<^sub>0h\<^sub>1.\) \ \[f \ g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ (\\<^sup>-\<^sup>1[f, g, tab\<^sub>0 g \ Tfg.p\<^sub>0] \ (f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \ (f \ (g.tab \ Tfg.p\<^sub>0)) \ (f \ f\<^sub>0g\<^sub>1.\) \ \[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \ (f.tab \ Tfg.p\<^sub>1) \ TTfgh.p\<^sub>1)) \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" using comp_assoc by presburger also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ \[f \ g, h, tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ TTfgh.tab \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" using TTfgh.tab_def Tfg.\\.tab_def by simp also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ \[f \ g, h, tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ ((\\<^sup>-\<^sup>1[f, g, h] \ tab\<^sub>0 h \ TTfgh.p\<^sub>0) \ ((f \ g \ h) \ TTfgh_TfTgh.the_\) \ \[f \ g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[f, g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0] \ (f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ (f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ (g \ h.tab \ Tgh.p\<^sub>0) \ (g \ g\<^sub>0h\<^sub>1.\) \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ (f \ f\<^sub>0gh\<^sub>1.\) \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ (f.tab \ TfTgh.p\<^sub>1) \ TTfgh_TfTgh.chine) \ TTfgh_TfTgh.the_\) \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" 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 w\<^sub>f' \\<^sub>f' \ \\<^sub>f\.\ also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ \[f \ g, h, tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ ((\\<^sup>-\<^sup>1[f, g, h] \ tab\<^sub>0 h \ TTfgh.p\<^sub>0) \ ((f \ g \ h) \ TTfgh_TfTgh.the_\) \ \[f \ g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[f, g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0] \ (f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ ((f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ (f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ (f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ (f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ (f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ (f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ (f \ f\<^sub>0gh\<^sub>1.\)) \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ (f.tab \ TfTgh.p\<^sub>1) \ TTfgh_TfTgh.chine) \ TTfgh_TfTgh.the_\) \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" proof - have "f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ (g \ h.tab \ Tgh.p\<^sub>0) \ (g \ g\<^sub>0h\<^sub>1.\) \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0 = (f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ (f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ (f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ (f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ (f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ (f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0)" using fg gh whisker_right whisker_left by simp thus ?thesis using comp_assoc by presburger qed also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ \[f \ g, h, tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ (\\<^sup>-\<^sup>1[f, g, h] \ tab\<^sub>0 h \ TTfgh.p\<^sub>0) \ ((f \ g \ h) \ TTfgh_TfTgh.the_\) \ \[f \ g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[f, g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ ((f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ TTfgh_TfTgh.chine) \ ((f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ f\<^sub>0gh\<^sub>1.\) \ TTfgh_TfTgh.chine) \ (\[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ TTfgh_TfTgh.chine) \ ((f.tab \ TfTgh.p\<^sub>1) \ TTfgh_TfTgh.chine) \ TTfgh_TfTgh.the_\ \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" proof - have "\\<^sup>-\<^sup>1[f, g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0] \ (f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ (f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ (f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ (f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ (f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ (f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ (f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ (f \ f\<^sub>0gh\<^sub>1.\) \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ (f.tab \ TfTgh.p\<^sub>1) \ TTfgh_TfTgh.chine = (\\<^sup>-\<^sup>1[f, g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ ((f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ TTfgh_TfTgh.chine) \ ((f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ f\<^sub>0gh\<^sub>1.\) \ TTfgh_TfTgh.chine) \ (\[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ TTfgh_TfTgh.chine) \ ((f.tab \ TfTgh.p\<^sub>1) \ TTfgh_TfTgh.chine)" proof - have "arr (\\<^sup>-\<^sup>1[f, g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0] \ (f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ (f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ (f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ (f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ (f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ (f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ (f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ (f \ f\<^sub>0gh\<^sub>1.\) \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ (f.tab \ TfTgh.p\<^sub>1))" using fg gh by (intro seqI' comp_in_homI) auto (* * TODO: Find a way to generate the following consequences automatically * without having to list them. *) moreover have "arr ((f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ (f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ (f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ (f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ (f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ (f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ (f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ (f \ f\<^sub>0gh\<^sub>1.\) \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ (f.tab \ TfTgh.p\<^sub>1))" using calculation by blast moreover have "arr ((f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ (f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ (f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ (f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ (f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ (f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ (f \ f\<^sub>0gh\<^sub>1.\) \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ (f.tab \ TfTgh.p\<^sub>1))" using calculation by blast moreover have "arr ((f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ (f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ (f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ (f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ (f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ (f \ f\<^sub>0gh\<^sub>1.\) \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ (f.tab \ TfTgh.p\<^sub>1))" using calculation by blast moreover have "arr ((f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ (f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ (f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ (f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ (f \ f\<^sub>0gh\<^sub>1.\) \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ (f.tab \ TfTgh.p\<^sub>1))" using calculation by blast moreover have "arr ((f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ (f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ (f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ (f \ f\<^sub>0gh\<^sub>1.\) \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ (f.tab \ TfTgh.p\<^sub>1))" using calculation by blast moreover have "arr ((f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ (f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ (f \ f\<^sub>0gh\<^sub>1.\) \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ (f.tab \ TfTgh.p\<^sub>1))" using calculation by blast moreover have "arr ((f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ (f \ f\<^sub>0gh\<^sub>1.\) \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ (f.tab \ TfTgh.p\<^sub>1))" using calculation by blast moreover have "arr ((f \ f\<^sub>0gh\<^sub>1.\) \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ (f.tab \ TfTgh.p\<^sub>1))" using calculation by blast moreover have "arr (\[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ (f.tab \ TfTgh.p\<^sub>1))" using calculation by blast ultimately show ?thesis using whisker_right by auto qed thus ?thesis using comp_assoc by presburger qed also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ \[f \ g, h, tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ (\\<^sup>-\<^sup>1[f, g, h] \ tab\<^sub>0 h \ TTfgh.p\<^sub>0) \ ((f \ g \ h) \ TTfgh_TfTgh.the_\) \ \[f \ g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[f, g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ ((f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ TTfgh_TfTgh.chine) \ ((f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ f\<^sub>0gh\<^sub>1.\) \ TTfgh_TfTgh.chine) \ (\[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ TTfgh_TfTgh.chine) \ (((f.tab \ TfTgh.p\<^sub>1) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ \[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine]) \ TTfgh_TfTgh.the_\ \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" proof - have "((f.tab \ TfTgh.p\<^sub>1) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ \[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] = (f.tab \ TfTgh.p\<^sub>1) \ TTfgh_TfTgh.chine" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps comp_arr_dom comp_assoc_assoc' by simp thus ?thesis by simp qed also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ \[f \ g, h, tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ (\\<^sup>-\<^sup>1[f, g, h] \ tab\<^sub>0 h \ TTfgh.p\<^sub>0) \ ((f \ g \ h) \ TTfgh_TfTgh.the_\) \ \[f \ g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[f, g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ ((f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ TTfgh_TfTgh.chine) \ ((f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ f\<^sub>0gh\<^sub>1.\) \ TTfgh_TfTgh.chine) \ (\[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ TTfgh_TfTgh.chine) \ (((f.tab \ TfTgh.p\<^sub>1) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine]) \ \[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ TTfgh_TfTgh.the_\ \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" using comp_assoc by presburger also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ \[f \ g, h, tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ (\\<^sup>-\<^sup>1[f, g, h] \ tab\<^sub>0 h \ TTfgh.p\<^sub>0) \ ((f \ g \ h) \ TTfgh_TfTgh.the_\) \ \[f \ g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[f, g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ ((f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ TTfgh_TfTgh.chine) \ ((f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ f\<^sub>0gh\<^sub>1.\) \ TTfgh_TfTgh.chine) \ (\[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f \ tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ (f.tab \ TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine) \ \[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ TTfgh_TfTgh.the_\ \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" proof - have "((f.tab \ TfTgh.p\<^sub>1) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] = \\<^sup>-\<^sup>1[f \ tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ (f.tab \ TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine)" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps assoc'_naturality [of f.tab TfTgh.p\<^sub>1 TTfgh_TfTgh.chine] by simp thus ?thesis using comp_assoc by presburger qed also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ \[f \ g, h, tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ (\\<^sup>-\<^sup>1[f, g, h] \ tab\<^sub>0 h \ TTfgh.p\<^sub>0) \ ((f \ g \ h) \ TTfgh_TfTgh.the_\) \ \[f \ g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[f, g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ ((f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ TTfgh_TfTgh.chine) \ ((f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ f\<^sub>0gh\<^sub>1.\) \ TTfgh_TfTgh.chine) \ (\[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f \ tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ ((\\<^sup>-\<^sup>1[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine] \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine]) \ (f.tab \ TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine)) \ \[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ TTfgh_TfTgh.the_\ \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" proof - have "(\\<^sup>-\<^sup>1[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine] \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine]) \ (f.tab \ TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine) = f.tab \ TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps comp_cod_arr comp_assoc_assoc' by simp thus ?thesis using comp_assoc by simp qed also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ \[f \ g, h, tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ (\\<^sup>-\<^sup>1[f, g, h] \ tab\<^sub>0 h \ TTfgh.p\<^sub>0) \ ((f \ g \ h) \ TTfgh_TfTgh.the_\) \ \[f \ g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[f, g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ ((f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ TTfgh_TfTgh.chine) \ ((f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ f\<^sub>0gh\<^sub>1.\) \ TTfgh_TfTgh.chine) \ ((\[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f \ tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine]) \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine] \ (f.tab \ TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine) \ \[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ TTfgh_TfTgh.the_\ \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" using comp_assoc by presburger also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ \[f \ g, h, tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ (\\<^sup>-\<^sup>1[f, g, h] \ tab\<^sub>0 h \ TTfgh.p\<^sub>0) \ ((f \ g \ h) \ TTfgh_TfTgh.the_\) \ \[f \ g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[f, g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ (((f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ TTfgh_TfTgh.chine) \ ((f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ f\<^sub>0gh\<^sub>1.\) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f, tab\<^sub>0 f \ TfTgh.p\<^sub>1, TTfgh_TfTgh.chine]) \ (f \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine]) \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine] \ (f.tab \ TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine) \ \[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ TTfgh_TfTgh.the_\ \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" proof - have "(\[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f \ tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine] = \\<^sup>-\<^sup>1[f, tab\<^sub>0 f \ TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ (f \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine])" proof - have "(\[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f \ tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine] = \(\<^bold>\\<^bold>[\<^bold>\f\<^bold>\, \<^bold>\tab\<^sub>0 f\<^bold>\, \<^bold>\TfTgh.p\<^sub>1\<^bold>\\<^bold>] \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 f\<^bold>\, \<^bold>\TfTgh.p\<^sub>1\<^bold>\, \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\f\<^bold>\, \<^bold>\tab\<^sub>0 f\<^bold>\, \<^bold>\TfTgh.p\<^sub>1\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>]\" using \'_def \_def by simp also have "... = \\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\f\<^bold>\, \<^bold>\tab\<^sub>0 f\<^bold>\ \<^bold>\ \<^bold>\TfTgh.p\<^sub>1\<^bold>\, \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\tab\<^sub>0 f\<^bold>\, \<^bold>\TfTgh.p\<^sub>1\<^bold>\, \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>])\" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps by (intro E.eval_eqI, auto) also have "... = \\<^sup>-\<^sup>1[f, tab\<^sub>0 f \ TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ (f \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, 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 \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ \[f \ g, h, tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ (\\<^sup>-\<^sup>1[f, g, h] \ tab\<^sub>0 h \ TTfgh.p\<^sub>0) \ ((f \ g \ h) \ TTfgh_TfTgh.the_\) \ \[f \ g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[f, g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f, (g \ h) \ (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ (f \ (\\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ ((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ (\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine) \ (f \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine])) \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine] \ (f.tab \ TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine) \ \[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ TTfgh_TfTgh.the_\ \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" proof - have "((f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ TTfgh_TfTgh.chine) \ ((f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ f\<^sub>0gh\<^sub>1.\) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f, tab\<^sub>0 f \ TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] = \\<^sup>-\<^sup>1[f, (g \ h) \ (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ (f \ (\\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ (\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine)" proof - (* * This one can't be shortcut with a straight coherence-based proof, * due to the presence of f\<^sub>0gh\<^sub>1.\, g\<^sub>0h\<^sub>1.\, 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, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ TTfgh_TfTgh.chine) \ ((f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ f\<^sub>0gh\<^sub>1.\) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f, tab\<^sub>0 f \ TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] = ((f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ TTfgh_TfTgh.chine) \ ((f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f, (tab\<^sub>1 g \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \ (f \ f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine)" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps assoc'_naturality [of f f\<^sub>0gh\<^sub>1.\ TTfgh_TfTgh.chine] comp_assoc by simp also have "... = ((f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ TTfgh_TfTgh.chine) \ ((f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f, ((g \ tab\<^sub>0 g) \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \ (f \ ((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine)" proof - have "((f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f, (tab\<^sub>1 g \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] = \\<^sup>-\<^sup>1[f, ((g \ tab\<^sub>0 g) \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (f \ ((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine)" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps assoc'_naturality [of f "(g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0" TTfgh_TfTgh.chine] by simp thus ?thesis using comp_assoc by presburger qed also have "... = ((f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ TTfgh_TfTgh.chine) \ ((f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f, (g \ tab\<^sub>0 g \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \ (f \ (\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine)" proof - have "((f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f, ((g \ tab\<^sub>0 g) \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] = \\<^sup>-\<^sup>1[f, (g \ tab\<^sub>0 g \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (f \ (\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine)" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps assoc'_naturality [of f "\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0" TTfgh_TfTgh.chine] by simp thus ?thesis using comp_assoc by presburger qed also have "... = ((f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ TTfgh_TfTgh.chine) \ ((f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f, (g \ tab\<^sub>1 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \ (f \ ((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ (\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine)" proof - have "((f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f, (g \ tab\<^sub>0 g \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] = \\<^sup>-\<^sup>1[f, (g \ tab\<^sub>1 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (f \ ((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine)" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps assoc'_naturality [of f "(g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0" TTfgh_TfTgh.chine] by simp thus ?thesis using comp_assoc by presburger qed also have "... = (((f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ TTfgh_TfTgh.chine) \ ((f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f, (g \ (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \ (f \ ((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ (\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine)" proof - have "((f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f, (g \ tab\<^sub>1 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] = \\<^sup>-\<^sup>1[f, (g \ (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (f \ ((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine)" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps assoc'_naturality [of f "(g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0" TTfgh_TfTgh.chine] by simp thus ?thesis using comp_assoc by presburger qed also have "... = \\<^sup>-\<^sup>1[f, (g \ h) \ (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ (f \ (\\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ (\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine)" proof - (* OK, we can perhaps shortcut the last few steps... *) have "((f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ TTfgh_TfTgh.chine) \ ((f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f, (g \ (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] = \((\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\, \<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\, \<^bold>\TfTgh.p\<^sub>0\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \<^bold>\ ((\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\g\<^bold>\, \<^bold>\h\<^bold>\, \<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\\<^bold>] \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \<^bold>\ ((\<^bold>\f\<^bold>\ \<^bold>\ (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\h\<^bold>\, \<^bold>\tab\<^sub>0 h\<^bold>\, \<^bold>\Tgh.p\<^sub>0\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\f\<^bold>\, (\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\, \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>]\" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps \'_def \_def by simp also have "... = \\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\f\<^bold>\, (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) \<^bold>\ (\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\, \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\f\<^bold>\ \<^bold>\ (\<^bold>\\<^bold>[\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\, \<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\, \<^bold>\TfTgh.p\<^sub>0\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \<^bold>\ (\<^bold>\f\<^bold>\ \<^bold>\ (\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\g\<^bold>\, \<^bold>\h\<^bold>\, \<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\\<^bold>] \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \<^bold>\ (\<^bold>\f\<^bold>\ \<^bold>\ ((\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\h\<^bold>\, \<^bold>\tab\<^sub>0 h\<^bold>\, \<^bold>\Tgh.p\<^sub>0\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)\" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps apply (intro E.eval_eqI) by simp_all also have "... = \\<^sup>-\<^sup>1[f, (g \ h) \ (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ (f \ (\\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine)" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_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 \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ \[f \ g, h, tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ (\\<^sup>-\<^sup>1[f, g, h] \ tab\<^sub>0 h \ TTfgh.p\<^sub>0) \ ((f \ g \ h) \ TTfgh_TfTgh.the_\)) \ \[f \ g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[f, g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f, (g \ h) \ (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ (f \ (\\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ (((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine]) \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine] \ (f.tab \ TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine) \ \[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ TTfgh_TfTgh.the_\ \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" proof - have "(f \ ((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ (\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine) \ (f \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine]) = (f \ (((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine])" proof - have "arr ((((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine])" using fg gh apply (intro seqI) by auto moreover have "arr ((((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine])" using calculation by blast moreover have "arr (((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine])" using calculation by blast moreover have "arr ((((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine])" using calculation by blast moreover have "arr ((f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, 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 \ g \ h \ TTfgh_TfTgh.the_\) \ (\[f, g, h \ ((tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine] \ \[f \ g, h, ((tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[f, g, h] \ ((tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \[f \ g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[f, g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f, (g \ h) \ (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ (f \ (\\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine)) \ (f \ (((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine]) \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine] \ (f.tab \ TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine) \ \[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ TTfgh_TfTgh.the_\ \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" proof - have "\[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ \[f \ g, h, tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ (\\<^sup>-\<^sup>1[f, g, h] \ tab\<^sub>0 h \ TTfgh.p\<^sub>0) \ ((f \ g \ h) \ TTfgh_TfTgh.the_\) = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ (\[f \ g, h, tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ (((f \ g) \ h) \ TTfgh_TfTgh.the_\)) \ (\\<^sup>-\<^sup>1[f, g, h] \ ((tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine)" proof - have "(\\<^sup>-\<^sup>1[f, g, h] \ tab\<^sub>0 h \ TTfgh.p\<^sub>0) \ ((f \ g \ h) \ TTfgh_TfTgh.the_\) = \\<^sup>-\<^sup>1[f, g, h] \ TTfgh_TfTgh.the_\" using fg gh comp_arr_dom comp_cod_arr interchange [of "\\<^sup>-\<^sup>1[f, g, h]" "f \ g \ h" "tab\<^sub>0 h \ TTfgh.p\<^sub>0" TTfgh_TfTgh.the_\] by simp also have "... = (((f \ g) \ h) \ TTfgh_TfTgh.the_\) \ (\\<^sup>-\<^sup>1[f, g, h] \ ((tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine)" using fg gh comp_arr_dom comp_cod_arr interchange [of "(f \ g) \ h" "\\<^sup>-\<^sup>1[f, g, h]" TTfgh_TfTgh.the_\ "((tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine"] by simp finally show ?thesis using comp_assoc by presburger qed also have "... = (\[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ ((f \ g) \ h \ TTfgh_TfTgh.the_\)) \ \[f \ g, h, ((tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[f, g, h] \ ((tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ 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 \ ((tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine] \ \[f \ g, h, ((tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[f, g, h] \ ((tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ 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 (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)) \ (f \ (((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine])) \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine] \ (f.tab \ TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine) \ \[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ TTfgh_TfTgh.the_\ \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" proof - have "\[f, g, h \ ((tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine] \ \[f \ g, h, ((tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[f, g, h] \ ((tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \[f \ g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[f, g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f, (g \ h) \ (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ (f \ (\\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) = f \ can (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)" proof - have "\[f, g, h \ ((tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine] \ \[f \ g, h, ((tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[f, g, h] \ ((tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \[f \ g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[f, g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f, (g \ h) \ (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ (f \ (\\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) = \\<^bold>\\<^bold>[\<^bold>\f\<^bold>\, \<^bold>\g\<^bold>\, \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\<^bold>[\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\, \<^bold>\h\<^bold>\, ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\f\<^bold>\, \<^bold>\g\<^bold>\, \<^bold>\h\<^bold>\\<^bold>] \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\, (\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\, \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\f\<^bold>\, \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\, (\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\\<^bold>] \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\f\<^bold>\, (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) \<^bold>\ (\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\, \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\, \<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\, \<^bold>\TfTgh.p\<^sub>0\<^bold>\\<^bold>] \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \<^bold>\ (\<^bold>\f\<^bold>\ \<^bold>\ (\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\g\<^bold>\, \<^bold>\h\<^bold>\, \<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\\<^bold>] \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \<^bold>\ (\<^bold>\f\<^bold>\ \<^bold>\ ((\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\h\<^bold>\, \<^bold>\tab\<^sub>0 h\<^bold>\, \<^bold>\Tgh.p\<^sub>0\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)\" using \'_def \_def by simp also have "... = can (\<^bold>\f\<^bold>\ \<^bold>\ (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)) (\<^bold>\f\<^bold>\ \<^bold>\ (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\))" using fg gh apply (unfold can_def) apply (intro E.eval_eqI) by simp_all also have "... = f \ can (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)" 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 (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ (((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine]) \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine] \ (f.tab \ TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine) \ \[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ TTfgh_TfTgh.the_\ \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" proof - have "((f \ g \ h \ TTfgh_TfTgh.the_\) \ (f \ can (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)) \ (f \ (((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine])) = f \ (g \ h \ TTfgh_TfTgh.the_\) \ can (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ (((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine]" proof - have 1: "arr ((g \ h \ TTfgh_TfTgh.the_\) \ can (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ (((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine])" using fg gh apply (intro seqI' comp_in_homI) by auto moreover have 2: "arr (can (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ (((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, 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 w\<^sub>f' \\<^sub>f' \ \\<^sub>f" unfolding w\<^sub>f'_def \\<^sub>f'_def \\<^sub>f_def using comp_assoc by presburger finally show ?thesis by blast qed show ?thesis using w\<^sub>f w\<^sub>f' \\<^sub>f \\<^sub>f' \\<^sub>f f.T2 [of w\<^sub>f w\<^sub>f' \\<^sub>f u\<^sub>f \\<^sub>f' \\<^sub>f] eq\<^sub>f by fast qed obtain \\<^sub>f where \\<^sub>f: "\\\<^sub>f : w\<^sub>f \ w\<^sub>f'\ \ \\<^sub>f = tab\<^sub>1 f \ \\<^sub>f \ \\<^sub>f = \\<^sub>f' \ (tab\<^sub>0 f \ \\<^sub>f)" using 5 by auto show "\\TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine\\ = \\Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1\\" proof - have "iso \\<^sub>f" using \\<^sub>f BS3 w\<^sub>f_is_map w\<^sub>f'_is_map by blast hence "isomorphic w\<^sub>f w\<^sub>f'" using \\<^sub>f isomorphic_def isomorphic_symmetric by auto thus ?thesis using w\<^sub>f w\<^sub>f_def w\<^sub>f'_def Maps.CLS_eqI isomorphic_symmetric by auto qed text \ On to the next equation: \[ \\\Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine\\ = \\Tfg.p\<^sub>0 \ TTfgh.p\<^sub>1\\\. \] We have to make use of the equation \\\<^sub>f = \\<^sub>f' \ (tab\<^sub>0 f \ \\<^sub>f)\ in this part, similarly to how the equation \src_tab_eq\ was used to replace \TTfgh.tab\ in the first part. \ define u\<^sub>g where "u\<^sub>g = h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0" define w\<^sub>g where "w\<^sub>g = Tfg.p\<^sub>0 \ TTfgh.p\<^sub>1" define w\<^sub>g' where "w\<^sub>g' = Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine" define \\<^sub>g where "\\<^sub>g = \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0] \ (h.tab \ TTfgh.p\<^sub>0) \ fg\<^sub>0h\<^sub>1.\ \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]" define \\<^sub>g' where "\\<^sub>g' = (h \ TTfgh_TfTgh.the_\) \ can (\<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ ((h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ (g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]" define \\<^sub>g where "\\<^sub>g = \[tab\<^sub>1 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \[tab\<^sub>1 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ (tab\<^sub>0 f \ \\<^sub>f) \ \[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \ (inv f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]" have u\<^sub>g: "ide u\<^sub>g" unfolding u\<^sub>g_def by simp have w\<^sub>g: "ide w\<^sub>g" unfolding w\<^sub>g_def using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by auto have w\<^sub>g_is_map: "is_left_adjoint w\<^sub>g" unfolding w\<^sub>g_def using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps left_adjoints_compose by simp have w\<^sub>g': "ide w\<^sub>g'" unfolding w\<^sub>g'_def by simp have w\<^sub>g'_is_map: "is_left_adjoint w\<^sub>g'" unfolding w\<^sub>g'_def using TTfgh_TfTgh.is_map left_adjoints_compose by simp have \\<^sub>g: "\\\<^sub>g : tab\<^sub>0 g \ w\<^sub>g \ u\<^sub>g\" unfolding w\<^sub>g_def u\<^sub>g_def \\<^sub>g_def using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps fg\<^sub>0h\<^sub>1.\_in_hom apply (intro comp_in_homI) by auto have \\<^sub>g': "\\\<^sub>g' : tab\<^sub>0 g \ w\<^sub>g' \ u\<^sub>g\" unfolding w\<^sub>g'_def u\<^sub>g_def \\<^sub>g'_def using fg\<^sub>0h\<^sub>1.p\<^sub>0_simps apply (intro comp_in_homI) by auto have w\<^sub>g_in_hhom: "in_hhom w\<^sub>g (src u\<^sub>g) (src (tab\<^sub>0 g))" unfolding w\<^sub>g_def u\<^sub>g_def by auto have w\<^sub>g'_in_hhom: "in_hhom w\<^sub>g' (src u\<^sub>g) (src (tab\<^sub>0 g))" unfolding w\<^sub>g'_def u\<^sub>g_def by auto have \\<^sub>g: "\\\<^sub>g : tab\<^sub>1 g \ w\<^sub>g \ tab\<^sub>1 g \ w\<^sub>g'\" proof (unfold \\<^sub>g_def w\<^sub>g_def, intro comp_in_homI) (* auto can solve this, but it's too slow *) show "\\\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] : tab\<^sub>1 g \ Tfg.p\<^sub>0 \ TTfgh.p\<^sub>1 \ (tab\<^sub>1 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1\" using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by auto show "\inv f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1 : (tab\<^sub>1 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1 \ (tab\<^sub>0 f \ Tfg.p\<^sub>1) \ TTfgh.p\<^sub>1\" using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.\_in_hom f\<^sub>0g\<^sub>1.\_uniqueness(2) by (intro hcomp_in_vhom) auto show "\\[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] : (tab\<^sub>0 f \ Tfg.p\<^sub>1) \ TTfgh.p\<^sub>1 \ tab\<^sub>0 f \ Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1\" using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps \\<^sub>f w\<^sub>f_def w\<^sub>f'_def by auto show "\tab\<^sub>0 f \ \\<^sub>f : tab\<^sub>0 f \ Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1 \ tab\<^sub>0 f \ TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine\" using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps \\<^sub>f w\<^sub>f_def w\<^sub>f'_def by auto show "\\\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] : tab\<^sub>0 f \ TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine \ (tab\<^sub>0 f \ TfTgh.p\<^sub>1) \ TTfgh_TfTgh.chine\" by auto show "\f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine : (tab\<^sub>0 f \ TfTgh.p\<^sub>1) \ TTfgh_TfTgh.chine \ ((tab\<^sub>1 g \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine\" using f\<^sub>0gh\<^sub>1.\_in_hom by (intro hcomp_in_vhom) auto show "\\[tab\<^sub>1 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] : ((tab\<^sub>1 g \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine \ (tab\<^sub>1 g \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine\" by auto show "\\[tab\<^sub>1 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] : (tab\<^sub>1 g \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine \ tab\<^sub>1 g \ w\<^sub>g'\" using w\<^sub>g'_def by auto qed have eq\<^sub>g: "g.composite_cell w\<^sub>g \\<^sub>g = g.composite_cell w\<^sub>g' \\<^sub>g' \ \\<^sub>g" proof - have "g.composite_cell w\<^sub>g \\<^sub>g = (g \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0] \ (h.tab \ TTfgh.p\<^sub>0) \ fg\<^sub>0h\<^sub>1.\ \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0 \ TTfgh.p\<^sub>1] \ (g.tab \ Tfg.p\<^sub>0 \ TTfgh.p\<^sub>1)" unfolding w\<^sub>g_def \\<^sub>g_def by simp also have "... = (g \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ (g \ h.tab \ TTfgh.p\<^sub>0) \ (g \ fg\<^sub>0h\<^sub>1.\) \ ((g \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0 \ TTfgh.p\<^sub>1]) \ (g.tab \ Tfg.p\<^sub>0 \ TTfgh.p\<^sub>1)" using fg gh f\<^sub>0g\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps whisker_left comp_assoc by simp also have "... = (g \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ (g \ h.tab \ TTfgh.p\<^sub>0) \ (g \ fg\<^sub>0h\<^sub>1.\) \ (\[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ (\\<^sup>-\<^sup>1[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ (g \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]))) \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0 \ TTfgh.p\<^sub>1] \ (g.tab \ Tfg.p\<^sub>0 \ TTfgh.p\<^sub>1)" proof - have "(\[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ \\<^sup>-\<^sup>1[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ (g \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) = g \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]" using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps comp_cod_arr comp_assoc_assoc' by simp thus ?thesis by (simp add: comp_assoc) qed also have "... = (g \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ (g \ h.tab \ TTfgh.p\<^sub>0) \ (g \ fg\<^sub>0h\<^sub>1.\) \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ (\\<^sup>-\<^sup>1[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ (g \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0 \ TTfgh.p\<^sub>1]) \ (g.tab \ Tfg.p\<^sub>0 \ TTfgh.p\<^sub>1)" using comp_assoc by presburger also have "... = (g \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ (g \ h.tab \ TTfgh.p\<^sub>0) \ (g \ fg\<^sub>0h\<^sub>1.\) \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ (\[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ (\\<^sup>-\<^sup>1[g \ tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ (g.tab \ Tfg.p\<^sub>0 \ TTfgh.p\<^sub>1))" using fg gh f\<^sub>0g\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps comp_assoc fg\<^sub>0h\<^sub>1.p\<^sub>1_simps pentagon' iso_inv_iso invert_opposite_sides_of_square [of "\\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1" "(\\<^sup>-\<^sup>1[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ (g \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1])" "\\<^sup>-\<^sup>1[g \ tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]" "\\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tfg.p\<^sub>0 \ TTfgh.p\<^sub>1]"] by simp also have "... = (g \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ (g \ h.tab \ TTfgh.p\<^sub>0) \ (g \ fg\<^sub>0h\<^sub>1.\) \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ (\[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ ((g.tab \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]" using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps assoc'_naturality [of g.tab Tfg.p\<^sub>0 TTfgh.p\<^sub>1] by simp also have "... = (g \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ (g \ h.tab \ TTfgh.p\<^sub>0) \ (g \ fg\<^sub>0h\<^sub>1.\) \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ (\[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ ((g.tab \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1) \ (f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \ \[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \ (inv f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]" proof - have "(f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \ \[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \ (inv f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] = ((f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ (\\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \ \[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]) \ (inv f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1)) \ \\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]" using comp_assoc by presburger also have "... = ((f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ ((tab\<^sub>0 f \ Tfg.p\<^sub>1) \ TTfgh.p\<^sub>1) \ (inv f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1)) \ \\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]" using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps whisker_right comp_assoc_assoc' by simp also have "... = ((f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ (inv f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1)) \ \\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]" using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.\_uniqueness comp_cod_arr by simp also have "... = ((tab\<^sub>1 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]" proof - have "(f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ (inv f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) = f\<^sub>0g\<^sub>1.\ \ inv f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1" using f\<^sub>0g\<^sub>1.\_uniqueness whisker_right by simp also have "... = (tab\<^sub>1 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1" using f\<^sub>0g\<^sub>1.\_uniqueness comp_arr_inv' by simp finally show ?thesis by simp qed also have "... = \\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]" using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps comp_cod_arr by simp finally have "(f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \ \[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \ (inv f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] = \\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]" by simp thus ?thesis using comp_assoc by presburger qed also have "... = \\<^sub>f \ \[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \ (inv f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]" unfolding \\<^sub>f_def using comp_assoc by presburger also have "... = \\<^sub>f' \ (tab\<^sub>0 f \ \\<^sub>f) \ \[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \ (inv f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]" using \\<^sub>f comp_assoc by simp also have "... = (g \ h \ TTfgh_TfTgh.the_\) \ can (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ (((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ (tab\<^sub>0 f \ \\<^sub>f) \ \[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \ (inv f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]" unfolding \\<^sub>f'_def using comp_assoc by presburger also have "... = (g \ h \ TTfgh_TfTgh.the_\) \ can (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ (((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (\\<^sup>-\<^sup>1[tab\<^sub>1 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[tab\<^sub>1 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \[tab\<^sub>1 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]) \ \[tab\<^sub>1 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine)) \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ (tab\<^sub>0 f \ \\<^sub>f) \ \[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \ (inv f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]" proof - have "(\\<^sup>-\<^sup>1[tab\<^sub>1 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[tab\<^sub>1 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \[tab\<^sub>1 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]) \ \[tab\<^sub>1 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \ (f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine) = f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine" using f\<^sub>0gh\<^sub>1.p\<^sub>0_simps comp_cod_arr comp_arr_dom comp_assoc_assoc' by simp thus ?thesis using comp_assoc by simp qed also have "... = (g \ h \ TTfgh_TfTgh.the_\) \ can (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ (((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>1 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \ \\<^sup>-\<^sup>1[tab\<^sub>1 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \\<^sub>g" unfolding \\<^sub>g_def using comp_assoc by presburger also have "... = (g \ h \ TTfgh_TfTgh.the_\) \ can (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ (((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[(g \ tab\<^sub>0 g) \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>1 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]) \ \\<^sub>g" proof - have "(((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>1 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] = \\<^sup>-\<^sup>1[(g \ tab\<^sub>0 g) \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ ((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine)" using f\<^sub>0gh\<^sub>1.p\<^sub>0_simps assoc'_naturality [of "(g.tab \ Tgh.p\<^sub>1)" TfTgh.p\<^sub>0 TTfgh_TfTgh.chine] by simp thus ?thesis using comp_assoc by presburger qed also have "... = (g \ h \ TTfgh_TfTgh.the_\) \ can (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ (((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[(g \ tab\<^sub>0 g) \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[g \ tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (g.tab \ Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sub>g" proof - have "((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>1 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] = \\<^sup>-\<^sup>1[g \ tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (g.tab \ Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine)" using g\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0gh\<^sub>1.p\<^sub>0_simps assoc'_naturality [of g.tab Tgh.p\<^sub>1 "TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine"] by simp thus ?thesis using comp_assoc by presburger qed also have "... = (g \ h \ TTfgh_TfTgh.the_\) \ can (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ (((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[(g \ tab\<^sub>0 g) \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[g \ tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ ((\\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]) \ (g.tab \ Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine)) \ \\<^sub>g" proof - have "(\\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]) \ (g.tab \ Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) = g.tab \ Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine" using comp_cod_arr comp_assoc_assoc' by simp thus ?thesis using comp_assoc g\<^sub>0h\<^sub>1.\_in_hom by simp qed also have "... = (g \ h \ TTfgh_TfTgh.the_\) \ can (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ (((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((\\<^sup>-\<^sup>1[g \ tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[g, tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \[g, tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]) \ \[g \ tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine)) \ \\<^sup>-\<^sup>1[(g \ tab\<^sub>0 g) \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[g \ tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ ((\\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]) \ (g.tab \ Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine)) \ \\<^sub>g" proof - have "(\\<^sup>-\<^sup>1[g \ tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[g, tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \[g, tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]) \ \[g \ tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) = (\\<^sup>-\<^sup>1[g \ tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ ((g \ (tab\<^sub>0 g \ Tgh.p\<^sub>1)) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \[g \ tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine)" using g\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0gh\<^sub>1.p\<^sub>0_simps comp_assoc comp_assoc_assoc' by simp also have "... = (\\<^sup>-\<^sup>1[g \ tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ \[g \ tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine)" using g\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0gh\<^sub>1.p\<^sub>0_simps comp_cod_arr comp_assoc_assoc' by simp also have "... = (((g \ (tab\<^sub>0 g \ Tgh.p\<^sub>1)) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine)" using g\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0gh\<^sub>1.p\<^sub>0_simps whisker_right comp_assoc_assoc' by simp also have "... = (\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine" using g\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0gh\<^sub>1.p\<^sub>0_simps comp_cod_arr by simp finally show ?thesis by presburger qed also have "... = (g \ h \ TTfgh_TfTgh.the_\) \ can (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ ((((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[g \ tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[g, tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]) \ \[g, tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \[g \ tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[(g \ tab\<^sub>0 g) \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[g \ tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (g.tab \ Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine)) \ \\<^sub>g" using comp_assoc by presburger also have "... = (g \ h \ TTfgh_TfTgh.the_\) \ (can (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ \\<^sup>-\<^sup>1[g \ (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[g, (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]) \ (g \ (h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ (g \ g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ (\[g, tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \[g \ tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[(g \ tab\<^sub>0 g) \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[g \ tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]) \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (g.tab \ Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sub>g" proof - have "(((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[g \ tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[g, tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] = \\<^sup>-\<^sup>1[g \ (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[g, (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (g \ (h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ (g \ g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine)" proof - have "(((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[g \ tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[g, tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] = (((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[g \ tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \ \\<^sup>-\<^sup>1[g, tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]" using comp_assoc by presburger also have "... = ((((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[g \ tab\<^sub>1 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \ ((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[g, tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]" proof - have "(((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[g \ tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] = \\<^sup>-\<^sup>1[g \ tab\<^sub>1 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ ((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine)" using gh f\<^sub>0gh\<^sub>1.p\<^sub>0_simps assoc'_naturality [of "g \ g\<^sub>0h\<^sub>1.\" TfTgh.p\<^sub>0 TTfgh_TfTgh.chine] by simp thus ?thesis using comp_assoc by presburger qed also have "... = \\<^sup>-\<^sup>1[g \ (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ ((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ ((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[g, tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]" proof - have "(((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[g \ tab\<^sub>1 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] = \\<^sup>-\<^sup>1[g \ (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ ((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine)" using gh f\<^sub>0gh\<^sub>1.p\<^sub>0_simps assoc'_naturality [of "g \ h.tab \ Tgh.p\<^sub>0" TfTgh.p\<^sub>0 TTfgh_TfTgh.chine] by simp thus ?thesis using comp_assoc by presburger qed also have "... = \\<^sup>-\<^sup>1[g \ (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[g, tab\<^sub>1 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]) \ (g \ g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine)" proof - have "((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[g, tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] = \\<^sup>-\<^sup>1[g, tab\<^sub>1 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (g \ g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine)" using gh f\<^sub>0gh\<^sub>1.p\<^sub>0_simps assoc'_naturality [of g g\<^sub>0h\<^sub>1.\ "TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine"] by simp thus ?thesis using comp_assoc by presburger qed also have "... = \\<^sup>-\<^sup>1[g \ (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[g, (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (g \ (h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ (g \ g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine)" proof - have "((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[g, tab\<^sub>1 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] = \\<^sup>-\<^sup>1[g, (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (g \ (h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine)" using gh f\<^sub>0gh\<^sub>1.p\<^sub>0_simps assoc'_naturality [of g "h.tab \ Tgh.p\<^sub>0" "TfTgh.p\<^sub>0 \ 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 (\<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)) \ (g \ (h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ (g \ g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ (g \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine])) \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (g.tab \ Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sub>g" proof - have "can (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ \\<^sup>-\<^sup>1[g \ (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[g, (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] = g \ can (\<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)" proof - have "\\<^sup>-\<^sup>1[g \ (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] = can (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) ((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)" proof - have "\\<^sup>-\<^sup>1[g \ (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] = \\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\, \<^bold>\TfTgh.p\<^sub>0\<^bold>\, \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>]\" using gh f\<^sub>0gh\<^sub>1.p\<^sub>0_simps canI_associator_0 \'_def \_def by simp also have "... = can (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) ((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)" unfolding can_def using gh apply (intro E.eval_eqI) by simp_all finally show ?thesis by blast qed moreover have "\\<^sup>-\<^sup>1[g, (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] = can ((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (\<^bold>\g\<^bold>\ \<^bold>\ ((\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)" proof - have "\\<^sup>-\<^sup>1[g, (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] = \\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\g\<^bold>\, (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\, \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>]\" using gh f\<^sub>0gh\<^sub>1.p\<^sub>0_simps canI_associator_0 \'_def \_def by simp also have "... = can ((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (\<^bold>\g\<^bold>\ \<^bold>\ ((\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)" 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, tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \[g \ tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[(g \ tab\<^sub>0 g) \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[g \ tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] = g \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]" proof - have "\[g, tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \[g \ tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[(g \ tab\<^sub>0 g) \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[g \ tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] = \\<^bold>\\<^bold>[\<^bold>\g\<^bold>\, \<^bold>\tab\<^sub>0 g\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>1\<^bold>\, \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\<^bold>[\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 g\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>1\<^bold>\, \<^bold>\TfTgh.p\<^sub>0\<^bold>\, \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>] \<^bold>\ ((\<^bold>\\<^bold>[\<^bold>\g\<^bold>\, \<^bold>\tab\<^sub>0 g\<^bold>\, \<^bold>\Tgh.p\<^sub>1\<^bold>\\<^bold>] \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[(\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 g\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>1\<^bold>\, \<^bold>\TfTgh.p\<^sub>0\<^bold>\, \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 g\<^bold>\, \<^bold>\Tgh.p\<^sub>1\<^bold>\, \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\g\<^bold>\, \<^bold>\tab\<^sub>0 g\<^bold>\, \<^bold>\Tgh.p\<^sub>1\<^bold>\ \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>]\" using gh g\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0gh\<^sub>1.p\<^sub>0_simps \'_def \_def by simp also have "... = \\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\tab\<^sub>0 g\<^bold>\, \<^bold>\Tgh.p\<^sub>1\<^bold>\, \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>]\" apply (intro E.eval_eqI) by simp_all also have "... = g \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]" using gh g\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0gh\<^sub>1.p\<^sub>0_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 (\<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)) \ ((h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ (g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]) \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (g.tab \ Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sub>g" proof - have "arr ((h \ TTfgh_TfTgh.the_\) \ (can (\<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)) \ ((h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ (g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine])" using gh f\<^sub>0gh\<^sub>1.p\<^sub>0_simps apply (intro seqI) by auto moreover have "arr ((can (\<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)) \ ((h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ (g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine])" using calculation by blast moreover have "arr (((h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ (g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine])" using calculation by blast moreover have "arr ((g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine])" using calculation by blast ultimately have "(g \ h \ TTfgh_TfTgh.the_\) \ (g \ can (\<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)) \ (g \ (h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ (g \ g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ (g \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]) = g \ (h \ TTfgh_TfTgh.the_\) \ (can (\<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)) \ ((h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ (g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]" using whisker_left by simp (* 20 sec *) thus ?thesis by simp qed also have "... = g.composite_cell w\<^sub>g' \\<^sub>g' \ \\<^sub>g" unfolding w\<^sub>g'_def \\<^sub>g'_def using comp_assoc by presburger finally show ?thesis by simp qed have 6: "\!\. \\ : w\<^sub>g \ w\<^sub>g'\ \ \\<^sub>g = tab\<^sub>1 g \ \ \ \\<^sub>g = \\<^sub>g' \ (tab\<^sub>0 g \ \)" using w\<^sub>g w\<^sub>g' \\<^sub>g \\<^sub>g' \\<^sub>g eq\<^sub>g g.T2 [of w\<^sub>g w\<^sub>g' \\<^sub>g u\<^sub>g \\<^sub>g' \\<^sub>g] by blast obtain \\<^sub>g where \\<^sub>g: "\\\<^sub>g : w\<^sub>g \ w\<^sub>g'\ \ \\<^sub>g = tab\<^sub>1 g \ \\<^sub>g \ \\<^sub>g = \\<^sub>g' \ (tab\<^sub>0 g \ \\<^sub>g)" using 6 by auto show "\\Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine\\ = \\Tfg.p\<^sub>0 \ TTfgh.p\<^sub>1\\" proof - have "iso \\<^sub>g" using \\<^sub>g BS3 w\<^sub>g_is_map w\<^sub>g'_is_map by blast hence "isomorphic w\<^sub>g w\<^sub>g'" using \\<^sub>g isomorphic_def isomorphic_symmetric by auto thus ?thesis using w\<^sub>g w\<^sub>g' w\<^sub>g_def w\<^sub>g'_def Maps.CLS_eqI by auto qed text \Now the last equation: similar, but somewhat simpler.\ define u\<^sub>h where "u\<^sub>h = tab\<^sub>0 h \ TTfgh.p\<^sub>0" define w\<^sub>h where "w\<^sub>h = TTfgh.p\<^sub>0" define w\<^sub>h' where "w\<^sub>h' = Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine" define \\<^sub>h where "\\<^sub>h = tab\<^sub>0 h \ TTfgh.p\<^sub>0" define \\<^sub>h' where "\\<^sub>h' = TTfgh_TfTgh.the_\ \ \\<^sup>-\<^sup>1[tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]" define \\<^sub>h where "\\<^sub>h = \[tab\<^sub>1 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (tab\<^sub>0 g \ \\<^sub>g) \ \[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ inv fg\<^sub>0h\<^sub>1.\" have u\<^sub>h: "ide u\<^sub>h" unfolding u\<^sub>h_def by simp have w\<^sub>h: "ide w\<^sub>h" unfolding w\<^sub>h_def by simp have w\<^sub>h_is_map: "is_left_adjoint w\<^sub>h" unfolding w\<^sub>h_def by simp have w\<^sub>h': "ide w\<^sub>h'" unfolding w\<^sub>h'_def by simp have w\<^sub>h'_is_map: "is_left_adjoint w\<^sub>h'" unfolding w\<^sub>h'_def using g\<^sub>0h\<^sub>1.p\<^sub>0_simps f\<^sub>0gh\<^sub>1.p\<^sub>0_simps TTfgh_TfTgh.is_map left_adjoints_compose by simp have \\<^sub>h: "\\\<^sub>h : tab\<^sub>0 h \ w\<^sub>h \ u\<^sub>h\" unfolding \\<^sub>h_def w\<^sub>h_def u\<^sub>h_def by auto have \\<^sub>h': "\\\<^sub>h' : tab\<^sub>0 h \ w\<^sub>h' \ u\<^sub>h\" unfolding \\<^sub>h'_def w\<^sub>h'_def u\<^sub>h_def using g\<^sub>0h\<^sub>1.p\<^sub>0_simps f\<^sub>0gh\<^sub>1.p\<^sub>0_simps by (intro comp_in_homI) auto have \\<^sub>h: "\\\<^sub>h : tab\<^sub>1 h \ w\<^sub>h \ tab\<^sub>1 h \ w\<^sub>h'\" proof (unfold \\<^sub>h_def w\<^sub>h_def w\<^sub>h'_def, intro comp_in_homI) (* auto can solve this, but it's too slow *) show "\inv fg\<^sub>0h\<^sub>1.\ : tab\<^sub>1 h \ TTfgh.p\<^sub>0 \ (tab\<^sub>0 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1\" using fg\<^sub>0h\<^sub>1.\_uniqueness by blast show "\\[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] : (tab\<^sub>0 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1 \ tab\<^sub>0 g \ Tfg.p\<^sub>0 \ TTfgh.p\<^sub>1\" using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by auto show "\tab\<^sub>0 g \ \\<^sub>g : tab\<^sub>0 g \ Tfg.p\<^sub>0 \ TTfgh.p\<^sub>1 \ tab\<^sub>0 g \ Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine\" using \\<^sub>g w\<^sub>g_def w\<^sub>g'_def fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by auto show "\\\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] : tab\<^sub>0 g \ Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine \ (tab\<^sub>0 g \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine\" using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by auto show "\g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine : (tab\<^sub>0 g \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine \ (tab\<^sub>1 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine\" using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by force show "\\[tab\<^sub>1 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] : (tab\<^sub>1 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine \ tab\<^sub>1 h \ Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine\" using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by auto qed have eq\<^sub>h: "h.composite_cell w\<^sub>h \\<^sub>h = h.composite_cell w\<^sub>h' \\<^sub>h' \ \\<^sub>h" proof - text \ Once again, the strategy is to form the subexpression \[ \\[h, tab\<^sub>0 h, TTfgh.p\<^sub>0] \ (h.tab \ TTfgh.p\<^sub>0) \ fg\<^sub>0h\<^sub>1.\ \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]\ \] which is equal to \\\<^sub>g\, so that we can make use of the equation \\\<^sub>g = \\<^sub>g' \ (tab\<^sub>0 g \ \\<^sub>g)\. \ have "h.composite_cell w\<^sub>h \\<^sub>h = (h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0) \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0] \ (h.tab \ TTfgh.p\<^sub>0)" unfolding w\<^sub>h_def \\<^sub>h_def by simp also have "... = \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0] \ (h.tab \ TTfgh.p\<^sub>0)" proof - have "(h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0) \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0] = \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]" using comp_cod_arr by simp thus ?thesis using comp_assoc by metis qed also have "... = (\[h, tab\<^sub>0 h, TTfgh.p\<^sub>0] \ (h.tab \ TTfgh.p\<^sub>0) \ fg\<^sub>0h\<^sub>1.\ \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ \[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ inv fg\<^sub>0h\<^sub>1.\" proof - have "(h.tab \ TTfgh.p\<^sub>0) \ fg\<^sub>0h\<^sub>1.\ \ (\\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ \[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ inv fg\<^sub>0h\<^sub>1.\ = (h.tab \ TTfgh.p\<^sub>0) \ fg\<^sub>0h\<^sub>1.\ \ ((tab\<^sub>0 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1) \ inv fg\<^sub>0h\<^sub>1.\" using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps comp_assoc_assoc' by simp also have "... = (h.tab \ TTfgh.p\<^sub>0) \ fg\<^sub>0h\<^sub>1.\ \ inv fg\<^sub>0h\<^sub>1.\" using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps fg\<^sub>0h\<^sub>1.\_uniqueness comp_cod_arr by simp also have "... = (h.tab \ TTfgh.p\<^sub>0) \ (tab\<^sub>1 h \ TTfgh.p\<^sub>0)" using comp_arr_inv' fg\<^sub>0h\<^sub>1.\_uniqueness by simp also have "... = h.tab \ TTfgh.p\<^sub>0" using comp_arr_dom fg\<^sub>0h\<^sub>1.p\<^sub>0_simps by simp finally have "(h.tab \ TTfgh.p\<^sub>0) \ fg\<^sub>0h\<^sub>1.\ \ (\\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ \[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ inv fg\<^sub>0h\<^sub>1.\ = h.tab \ TTfgh.p\<^sub>0" by simp thus ?thesis using comp_assoc by simp qed also have "... = \\<^sub>g \ \[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ inv fg\<^sub>0h\<^sub>1.\" unfolding \\<^sub>g_def by simp also have "... = (\\<^sub>g' \ (tab\<^sub>0 g \ \\<^sub>g)) \ \[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ inv fg\<^sub>0h\<^sub>1.\" using \\<^sub>g by simp also have "... = (h \ TTfgh_TfTgh.the_\) \ can (\<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ ((h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ (g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (tab\<^sub>0 g \ \\<^sub>g) \ \[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ inv fg\<^sub>0h\<^sub>1.\" unfolding \\<^sub>g'_def using comp_assoc by presburger also have "... = (h \ TTfgh_TfTgh.the_\) \ can (\<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ ((\\<^sup>-\<^sup>1[h \ tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \[h \ tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]) \ ((h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine)) \ (g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (tab\<^sub>0 g \ \\<^sub>g) \ \[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ inv fg\<^sub>0h\<^sub>1.\" proof - have "(\\<^sup>-\<^sup>1[h \ tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \[h \ tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]) \ ((h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) = (h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0 \ 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 (\<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ \\<^sup>-\<^sup>1[h \ tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (\[h \ tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ ((h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine)) \ (g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (tab\<^sub>0 g \ \\<^sub>g) \ \[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ inv fg\<^sub>0h\<^sub>1.\" using comp_assoc by presburger also have "... = (h \ TTfgh_TfTgh.the_\) \ can (\<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ \\<^sup>-\<^sup>1[h \ tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (h.tab \ Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \[tab\<^sub>1 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (tab\<^sub>0 g \ \\<^sub>g) \ \[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ inv fg\<^sub>0h\<^sub>1.\" using assoc_naturality [of h.tab Tgh.p\<^sub>0 "TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine"] comp_assoc by simp also have "... = (h \ TTfgh_TfTgh.the_\) \ can (\<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ \\<^sup>-\<^sup>1[h \ tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ ((\\<^sup>-\<^sup>1[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]) \ (h.tab \ Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine)) \ \[tab\<^sub>1 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (tab\<^sub>0 g \ \\<^sub>g) \ \[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ inv fg\<^sub>0h\<^sub>1.\" proof - have "(\\<^sup>-\<^sup>1[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]) \ (h.tab \ Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) = h.tab \ Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ 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 (\<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ (\\<^sup>-\<^sup>1[h \ tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine])) \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (h.tab \ Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \[tab\<^sub>1 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (tab\<^sub>0 g \ \\<^sub>g) \ \[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ inv fg\<^sub>0h\<^sub>1.\" using comp_assoc by presburger also have "... = ((h \ TTfgh_TfTgh.the_\) \ (h \ can (((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\))) \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (h.tab \ Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \[tab\<^sub>1 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (tab\<^sub>0 g \ \\<^sub>g) \ \[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ inv fg\<^sub>0h\<^sub>1.\" proof - have "can (\<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ (\\<^sup>-\<^sup>1[h \ tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]) = can (\<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ \\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\, \<^bold>\Tgh.p\<^sub>0\<^bold>\, \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\h\<^bold>\, \<^bold>\tab\<^sub>0 h\<^bold>\, \<^bold>\Tgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>]\" using \'_def \_def by simp also have "... = can (\<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ can (((\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)" proof - have "\\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\, \<^bold>\Tgh.p\<^sub>0\<^bold>\, \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\h\<^bold>\, \<^bold>\tab\<^sub>0 h\<^bold>\, \<^bold>\Tgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>]\ = can (((\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)" unfolding can_def apply (intro E.eval_eqI) by simp_all thus ?thesis by simp qed also have "... = can (\<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)" by simp also have "... = h \ can (((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)" using whisker_can_left_0 by simp finally have "can (\<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ (\\<^sup>-\<^sup>1[h \ tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]) = h \ can (((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)" by simp thus ?thesis using comp_assoc by presburger qed also have "... = (h \ TTfgh_TfTgh.the_\ \ can (((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)) \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (h.tab \ Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \[tab\<^sub>1 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (tab\<^sub>0 g \ \\<^sub>g) \ \[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ inv fg\<^sub>0h\<^sub>1.\" using whisker_left [of h] comp_assoc by simp also have "... = (h \ TTfgh_TfTgh.the_\ \ \\<^sup>-\<^sup>1[tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]) \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (h.tab \ Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \[tab\<^sub>1 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (tab\<^sub>0 g \ \\<^sub>g) \ \[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ inv fg\<^sub>0h\<^sub>1.\" proof - have "can (((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) = \\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\, \<^bold>\TfTgh.p\<^sub>0\<^bold>\, \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\tab\<^sub>0 h\<^bold>\, \<^bold>\Tgh.p\<^sub>0\<^bold>\, \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>]\" unfolding can_def apply (intro E.eval_eqI) by auto also have "... = \\<^sup>-\<^sup>1[tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]" using \'_def \_def by simp finally have "can (((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) = \\<^sup>-\<^sup>1[tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]" by simp thus ?thesis by simp qed also have "... = h.composite_cell w\<^sub>h' \\<^sub>h' \ \\<^sub>h" unfolding w\<^sub>h'_def \\<^sub>h'_def \\<^sub>h_def using comp_assoc by presburger finally show ?thesis by simp qed have 7: "\!\. \\ : w\<^sub>h \ w\<^sub>h'\ \ \\<^sub>h = tab\<^sub>1 h \ \ \ \\<^sub>h = \\<^sub>h' \ (tab\<^sub>0 h \ \)" using w\<^sub>h w\<^sub>h' \\<^sub>h \\<^sub>h' \\<^sub>h eq\<^sub>h h.T2 [of w\<^sub>h w\<^sub>h' \\<^sub>h u\<^sub>h \\<^sub>h' \\<^sub>h] by blast obtain \\<^sub>h where \\<^sub>h: "\\\<^sub>h : w\<^sub>h \ w\<^sub>h'\ \ \\<^sub>h = tab\<^sub>1 h \ \\<^sub>h \ \\<^sub>h = \\<^sub>h' \ (tab\<^sub>0 h \ \\<^sub>h)" using 7 by auto show "\\Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine\\ = \\TTfgh.p\<^sub>0\\" proof - have "iso \\<^sub>h" using \\<^sub>h BS3 w\<^sub>h_is_map w\<^sub>h'_is_map by blast hence "isomorphic w\<^sub>h w\<^sub>h'" using \\<^sub>h isomorphic_def isomorphic_symmetric by auto thus ?thesis using w\<^sub>h w\<^sub>h' w\<^sub>h_def w\<^sub>h'_def Maps.CLS_eqI [of w\<^sub>h w\<^sub>h'] 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 (prj\<^sub>0 (tab\<^sub>1 g \ prj\<^sub>1 (tab\<^sub>1 h) (tab\<^sub>0 g)) (tab\<^sub>0 f))" by (metis (lifting) Maps.Dom.simps(1) Maps.seq_char SPN_fgh.prj_chine_assoc(1) SPN_fgh.prj_simps(1) calculation f\<^sub>0gh\<^sub>1.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.p\<^sub>0)" 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.\\.prj\<^sub>1" tuple_ABC "\\TTfgh_TfTgh.chine\\"]) show "Maps.cospan SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1)" using SPN_fgh.\\.dom.is_span SPN_fgh.\\.leg1_composite SPN_fgh.cospan_\\ by auto show "Maps.seq SPN_fgh.Prj\<^sub>1 tuple_ABC" using 0 tuple_ABC_in_hom SPN_fgh.prj_in_hom(4) by auto show "Maps.seq SPN_fgh.Prj\<^sub>1 \\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.Prj\<^sub>1 \\trg TTfgh_TfTgh.chine\\ SPN_fgh.\.apex" proof show "Maps.cospan SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1)" using SPN_fgh.prj_in_hom(4) by blast show "\\trg TTfgh_TfTgh.chine\\ = Maps.pbdom SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1)" proof - have "\\trg TTfgh_TfTgh.chine\\ = Maps.MkIde (src TfTgh.p\<^sub>0)" by simp also have "... = Maps.pbdom SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1)" 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.Prj\<^sub>0\<^sub>1 SPN_fgh.Prj\<^sub>0" proof show "Maps.cospan SPN_fgh.\.leg0 SPN_fgh.\.leg1" using SPN_fgh.\\.legs_form_cospan(1) by simp show "Maps.span SPN_fgh.Prj\<^sub>0\<^sub>1 SPN_fgh.Prj\<^sub>0" using SPN_fgh.prj_simps(2-3,5-6) by presburger show "Maps.dom SPN_fgh.\.leg0 = Maps.cod SPN_fgh.Prj\<^sub>0\<^sub>1" 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.Prj\<^sub>0\<^sub>1 = SPN_fgh.\.leg1 \ SPN_fgh.Prj\<^sub>0" 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.\\.prj\<^sub>1) SPN_fgh.Prj\<^sub>1\<^sub>1 tuple_BC" proof show "Maps.cospan SPN_fgh.\.leg0 (Maps.comp SPN_fgh.\.leg1 SPN_fgh.\\.prj\<^sub>1)" using fg gh SPN_fgh.prj_simps(10) by blast show "Maps.span SPN_fgh.Prj\<^sub>1\<^sub>1 tuple_BC" using fg gh csq(2) by blast show "Maps.dom SPN_fgh.\.leg0 = Maps.cod SPN_fgh.Prj\<^sub>1\<^sub>1" using fg gh SPN_f.dom.leg_simps(2) SPN_fgh.prj_simps(7) by auto show "SPN_fgh.\.leg0 \ SPN_fgh.Prj\<^sub>1\<^sub>1 = (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1) \ tuple_BC" using 2 fg gh Maps.comp_assoc csq(2) Maps.prj_tuple [of SPN_fgh.\.leg0 SPN_fgh.\.leg1 SPN_fgh.Prj\<^sub>0\<^sub>1 SPN_fgh.Prj\<^sub>0] by blast qed show "SPN_fgh.Prj\<^sub>1 \ tuple_ABC = SPN_fgh.Prj\<^sub>1 \ Maps.CLS TTfgh_TfTgh.chine" proof - have "SPN_fgh.Prj\<^sub>1 \ tuple_ABC = SPN_fgh.Prj\<^sub>1\<^sub>1" using csq(2) by simp also have "... = \\Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1\\" using prj_char by simp also have "... = \\TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine\\" using prj_chine(1) by simp also have "... = \\TfTgh.p\<^sub>1\\ \ \\TTfgh_TfTgh.chine\\" proof - have "is_left_adjoint TfTgh.p\<^sub>1" by (simp add: fg) moreover have "is_left_adjoint TTfgh_TfTgh.chine" using TTfgh_TfTgh.is_map by simp moreover have "TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine \ TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine" using fg gh isomorphic_reflexive by simp ultimately show ?thesis using Maps.comp_CLS [of TfTgh.p\<^sub>1 TTfgh_TfTgh.chine "TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine"] by simp qed finally show ?thesis using prj_char by simp qed show "Maps.PRJ\<^sub>0 SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1) \ tuple_ABC = Maps.PRJ\<^sub>0 SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1) \ \\TTfgh_TfTgh.chine\\" proof - have "Maps.PRJ\<^sub>0 SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1) \ tuple_ABC = tuple_BC" using csq(2) Maps.prj_tuple [of SPN_fgh.\.leg0 "SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1" SPN_fgh.Prj\<^sub>1\<^sub>1 tuple_BC] by simp also have "... = Maps.comp (Maps.PRJ\<^sub>0 SPN_fgh.\.leg0 (Maps.comp SPN_fgh.\.leg1 SPN_fgh.\\.prj\<^sub>1)) \\TTfgh_TfTgh.chine\\" proof (intro Maps.prj_joint_monic [of SPN_fgh.\.leg0 SPN_fgh.\.leg1 tuple_BC "Maps.PRJ\<^sub>0 SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1) \ \\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.\\.prj\<^sub>1 tuple_BC" proof show "Maps.in_hom tuple_BC (Maps.MkIde (src TTfgh.p\<^sub>0)) (Maps.MkIde (src Tgh.p\<^sub>0))" using tuple_BC_in_hom by simp show "Maps.in_hom SPN_fgh.\\.prj\<^sub>1 (Maps.MkIde (src Tgh.p\<^sub>0)) SPN_fgh.\.apex" proof - have "Maps.pbdom SPN_fgh.\.leg0 SPN_fgh.\.leg1 = Maps.MkIde (src Tgh.p\<^sub>0)" 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.\\.prj\<^sub>1 (Maps.PRJ\<^sub>0 SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1) \ \\TTfgh_TfTgh.chine\\)" proof show "Maps.in_hom SPN_fgh.\\.prj\<^sub>1 (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.PRJ\<^sub>0 SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1) \ \\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.PRJ\<^sub>0 SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1)) \\trg TTfgh_TfTgh.chine\\ (Maps.pbdom SPN_fgh.\.leg0 SPN_fgh.\.leg1)" proof show "Maps.cospan SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1)" using SPN_fgh.prj_in_hom(4) by blast show "\\trg TTfgh_TfTgh.chine\\ = Maps.pbdom SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1)" proof - have "\\trg TTfgh_TfTgh.chine\\ = Maps.MkIde (src TfTgh.p\<^sub>0)" by simp also have "... = Maps.pbdom SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1)" 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.\\.prj\<^sub>1)" 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.\\.prj\<^sub>0 \ tuple_BC = SPN_fgh.\\.prj\<^sub>0 \ Maps.PRJ\<^sub>0 SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1) \ \\TTfgh_TfTgh.chine\\" proof - have "SPN_fgh.\\.prj\<^sub>0 \ tuple_BC = SPN_fgh.Prj\<^sub>0" using csq(1) by simp also have "... = SPN_fgh.\\.prj\<^sub>0 \ Maps.PRJ\<^sub>0 SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1) \ \\TTfgh_TfTgh.chine\\" proof - have "SPN_fgh.\\.prj\<^sub>0 \ Maps.PRJ\<^sub>0 SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1) \ \\TTfgh_TfTgh.chine\\ = \\Tgh.p\<^sub>0\\ \ \\TfTgh.p\<^sub>0\\ \ \\TTfgh_TfTgh.chine\\" using fg gh Tgh.\\.prj_char TfTgh.prj_char(1) isomorphic_reflexive Maps.comp_CLS [of "tab\<^sub>1 g" "prj\<^sub>1 (tab\<^sub>1 h) (tab\<^sub>0 g)" "tab\<^sub>1 g \ Tgh.p\<^sub>1"] by simp also have "... = \\Tgh.p\<^sub>0\\ \ \\TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine\\" using fg gh TTfgh_TfTgh.is_map isomorphic_reflexive Maps.comp_CLS [of TfTgh.p\<^sub>0 TTfgh_TfTgh.chine "TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine"] by simp also have "... = \\Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine\\" using fg gh TTfgh_TfTgh.is_map left_adjoints_compose isomorphic_reflexive Maps.comp_CLS [of Tgh.p\<^sub>0 "TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine" "Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine"] by simp also have "... = \\TTfgh.p\<^sub>0\\" using prj_chine(3) by simp also have "... = SPN_fgh.Prj\<^sub>0" using prj_char by simp finally show ?thesis by argo qed finally show ?thesis by blast qed show "SPN_fgh.\\.prj\<^sub>1 \ tuple_BC = SPN_fgh.\\.prj\<^sub>1 \ Maps.PRJ\<^sub>0 SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1) \ \\TTfgh_TfTgh.chine\\" proof - have "SPN_fgh.\\.prj\<^sub>1 \ tuple_BC = SPN_fgh.Prj\<^sub>0\<^sub>1" using csq(1) by simp also have "... = SPN_fgh.\\.prj\<^sub>1 \ Maps.PRJ\<^sub>0 SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1) \ \\TTfgh_TfTgh.chine\\" proof - have "SPN_fgh.\\.prj\<^sub>1 \ Maps.PRJ\<^sub>0 SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1) \ \\TTfgh_TfTgh.chine\\ = \\Tgh.p\<^sub>1\\ \ \\TfTgh.p\<^sub>0\\ \ \\TTfgh_TfTgh.chine\\" using fg gh Tgh.\\.prj_char TfTgh.prj_char(1) isomorphic_reflexive Maps.comp_CLS [of "tab\<^sub>1 g" "prj\<^sub>1 (tab\<^sub>1 h) (tab\<^sub>0 g)" "tab\<^sub>1 g \ Tgh.p\<^sub>1"] by simp also have "... = \\Tgh.p\<^sub>1\\ \ \\TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine\\" using fg gh TTfgh_TfTgh.is_map isomorphic_reflexive Maps.comp_CLS [of TfTgh.p\<^sub>0 TTfgh_TfTgh.chine "TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine"] by simp also have "... = \\Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine\\" using fg gh TTfgh_TfTgh.is_map left_adjoints_compose isomorphic_reflexive Maps.comp_CLS [of Tgh.p\<^sub>1 "TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine" "Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine"] by simp also have "... = \\Tfg.p\<^sub>0 \ TTfgh.p\<^sub>1\\" using prj_chine(2) by simp also have "... = SPN_fgh.Prj\<^sub>0\<^sub>1" 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 \tab\<^sub>0 h \ TTfgh.p\<^sub>0\ \(tab\<^sub>1 f \ Tfg.p\<^sub>1) \ TTfgh.p\<^sub>1\ \f \ g \ h\ TfTgh.tab \(tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0\ \tab\<^sub>1 f \ TfTgh.p\<^sub>1\ \f \ g \ h\ TfHgh.\\.tab \tab\<^sub>0 (g \ h) \ TfHgh.\\.p\<^sub>0\ \tab\<^sub>1 f \ TfHgh.\\.p\<^sub>1\ \\[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 \tab\<^sub>0 h \ TTfgh.p\<^sub>0\ \(tab\<^sub>1 f \ Tfg.p\<^sub>1) \ TTfgh.p\<^sub>1\ \f \ g \ h\ TfHgh.\\.tab \tab\<^sub>0 (g \ h) \ TfHgh.\\.p\<^sub>0\ \tab\<^sub>1 f \ TfHgh.\\.p\<^sub>1\ \f \ g \ h\ HfHgh.tab \tab\<^sub>0 (f \ g \ h)\ \tab\<^sub>1 (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 \tab\<^sub>0 h \ THfgh.\\.p\<^sub>0\ \tab\<^sub>1 (f \ g) \ THfgh.\\.p\<^sub>1\ \(f \ g) \ h\ HHfgh.tab \tab\<^sub>0 ((f \ g) \ h)\ \tab\<^sub>1 ((f \ g) \ h)\ \f \ g \ h\ HfHgh.tab \tab\<^sub>0 (f \ g \ h)\ \tab\<^sub>1 (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 \tab\<^sub>0 h \ TTfgh.p\<^sub>0\ \(tab\<^sub>1 f \ Tfg.p\<^sub>1) \ TTfgh.p\<^sub>1\ \(f \ g) \ h\ THfgh.\\.tab \tab\<^sub>0 h \ THfgh.\\.p\<^sub>0\ \tab\<^sub>1 (f \ g) \ THfgh.\\.p\<^sub>1\ \f \ g \ h\ HfHgh.tab \tab\<^sub>0 (f \ g \ h)\ \tab\<^sub>1 (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.PRJ\<^sub>0 Maps.PRJ\<^sub>1 .. 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 \\\\. 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 \)\ \tab\<^sub>0 (dom \)\ \tab\<^sub>1 (dom \)\ \cod \\ \tab_of_ide (cod \)\ \tab\<^sub>0 (cod \)\ \tab\<^sub>1 (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 \)\ \tab\<^sub>0 (dom \)\ \tab\<^sub>1 (dom \)\ \cod \\ \tab_of_ide (cod \)\ \tab\<^sub>0 (cod \)\ \tab\<^sub>1 (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.map\<^sub>0 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_char by blast moreover have "Span.equivalent_objects (SPN.map\<^sub>0 ?a) a'" proof - have "SPN.map\<^sub>0 ?a = a'" proof (intro Span.arr_eqI) have "Chn (SPN.map\<^sub>0 ?a) = Chn (Span.src (SPN ?a))" using a a' by auto also have "... = Maps.MkIde (Maps.Dom (Chn a'))" proof - have "Maps.arr \\tab\<^sub>0 (dom (Maps.Dom (Chn a')))\\" proof - have "is_left_adjoint (tab\<^sub>0 (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 by simp + using a a' SPN_def Span.src_def Maps.cod_char obj_simps by simp qed also have "... = Chn a'" using a' Maps.MkIde_Dom Span.obj_char Span.ide_char by simp finally show "Chn (SPN.map\<^sub>0 ?a) = Chn a'" by simp show "Span.par (SPN.map\<^sub>0 (Maps.Dom (Chn a'))) a'" using a a' Span.obj_char apply (intro conjI) using SPN.map\<^sub>0_simps(1) Span.obj_def apply blast apply simp apply (metis (no_types, lifting) SPN.map\<^sub>0_def SPN.preserves_arr Span.obj_src \Chn (SPN.map\<^sub>0 (Maps.Dom (Chn a'))) = Chn a'\ obj_def) by (metis (no_types, lifting) SPN.map\<^sub>0_def SPN.preserves_arr Span.obj_src \Chn (SPN.map\<^sub>0 (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.map\<^sub>0 a) a'" by auto qed show "\a b g. \obj a; obj b; Span.in_hhom g (SPN.map\<^sub>0 a) (SPN.map\<^sub>0 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.map\<^sub>0 a) (SPN.map\<^sub>0 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 force 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 force let ?f = "Maps.REP g.leg1 \ (Maps.REP g.leg0)\<^sup>*" 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.map\<^sub>0 b))))" using g_in_hhom Span.trg_def [of g] by auto also have "... = src (Maps.REP (Maps.cod \\tab\<^sub>0 b\\))" using b arr_b SPN.map\<^sub>0_def Span.src_def SPN_in_hom by auto also have "... = src (Maps.REP \\trg (tab\<^sub>0 b)\\)" using b Maps.CLS_in_hom [of "tab\<^sub>0 b"] by force also have "... = src (Maps.REP \\b\\)" using b by fastforce also have "... = b" - using b by auto + using b obj_simps by auto finally show ?thesis by simp qed ultimately show ?thesis by argo qed show "\(Maps.REP g.leg0)\<^sup>* : 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.map\<^sub>0 a))))" using g_in_hhom Span.src_def [of g] by auto also have "... = src (Maps.REP (Maps.cod \\tab\<^sub>0 a\\))" using a arr_a SPN.map\<^sub>0_def Span.src_def SPN_in_hom by auto also have "... = src (Maps.REP \\trg (tab\<^sub>0 a)\\)" using a Maps.CLS_in_hom [of "tab\<^sub>0 a"] by force also have "... = src (Maps.REP \\a\\)" using a by fastforce also have "... = a" - using a by auto + 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 = \\tab\<^sub>0 ?f\\, Leg1 = \\tab\<^sub>1 ?f\\\, Cod = \Leg0 = \\tab\<^sub>0 ?f\\, Leg1 = \\tab\<^sub>1 ?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 \(tab\<^sub>0 f, \', tab\<^sub>1 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 \(tab\<^sub>0 f, tab\<^sub>1 f)\ to \(REP g.leg0, \, REP g.leg1)\. Its isomorphism class is an invertible arrow of spans in maps from \(CLS (tab\<^sub>0 f), CLS (tab\<^sub>1 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 \tab\<^sub>0 ?f\ \tab\<^sub>1 ?f\ ?f f.tab \tab\<^sub>0 ?f\ \tab\<^sub>1 ?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\<^sub>\ (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 (tab\<^sub>0 ?f) \ src (Maps.REP g.leg0)\ \ \w' : src (Maps.REP g.leg0) \ src (tab\<^sub>0 ?f)\ \ \\ : Maps.REP g.leg0 \ w \ tab\<^sub>0 ?f\ \ \\ : tab\<^sub>1 ?f \ Maps.REP g.leg1 \ w\ \ iso \ \ f.tab = (?f \ \) \ \[?f, Maps.REP g.leg0, w] \ (REP_leg0.trnr\<^sub>\ (Maps.REP g.leg1) ?f \ w) \ \ \ \\' : tab\<^sub>0 ?f \ w' \ Maps.REP g.leg0\ \ \\' : Maps.REP g.leg1 \ tab\<^sub>1 ?f \ w'\ \ iso \' \ REP_leg0.trnr\<^sub>\ (Maps.REP g.leg1) ?f = (?f \ \') \ \[?f, tab\<^sub>0 ?f, w'] \ (f.tab \ w') \ \'" using f.apex_unique_up_to_equivalence [of "REP_leg0.trnr\<^sub>\ (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 (tab\<^sub>0 ?f) \ src (Maps.REP g.leg0)\ \ \w' : src (Maps.REP g.leg0) \ src (tab\<^sub>0 ?f)\ \ \\ : Maps.REP g.leg0 \ w \ tab\<^sub>0 ?f\ \ \\ : tab\<^sub>1 ?f \ Maps.REP g.leg1 \ w\ \ iso \ \ f.tab = (?f \ \) \ \[?f, Maps.REP g.leg0, w] \ (REP_leg0.trnr\<^sub>\ (Maps.REP g.leg1) ?f \ w) \ \ \ \\' : tab\<^sub>0 ?f \ w' \ Maps.REP g.leg0\ \ \\' : Maps.REP g.leg1 \ tab\<^sub>1 ?f \ w'\ \ iso \' \ REP_leg0.trnr\<^sub>\ (Maps.REP g.leg1) ?f = (?f \ \') \ \[?f, tab\<^sub>0 ?f, w'] \ (f.tab \ w') \ \'" using 3 by meson hence w\\: "equivalence_map w \ \w : src (tab\<^sub>0 ?f) \ src (Maps.REP g.leg0)\ \ \\ : Maps.REP g.leg0 \ w \ tab\<^sub>0 ?f\ \ \\ : tab\<^sub>1 ?f \ Maps.REP g.leg1 \ w\ \ iso \" using equivalence_map_def equivalence_pair_def equivalence_pair_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.arrI) 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.arrI) apply auto proof - let ?f = "tab\<^sub>1 (Maps.REP g.leg1 \ (Maps.REP g.leg0)\<^sup>*)" assume 1: "\Maps.REP g.leg1 : Maps.Dom g.apex \ Maps.Cod g.leg1\" have "\?f : src (tab\<^sub>0 (Maps.REP g.leg1 \ (Maps.REP g.leg0)\<^sup>*)) \ Maps.Cod g.leg1\ \ is_left_adjoint ?f \ \tab\<^sub>1 (Maps.REP g.leg1 \ (Maps.REP g.leg0)\<^sup>*)\ = \?f\" using 1 by simp thus "\f. \f : src (tab\<^sub>0 (Maps.REP g.leg1 \ (Maps.REP g.leg0)\<^sup>*)) \ Maps.Cod g.leg1\ \ is_left_adjoint f \ \tab\<^sub>1 (Maps.REP g.leg1 \ (Maps.REP g.leg0)\<^sup>*)\ = \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 (tab\<^sub>0 (Maps.REP g.leg1 \ (Maps.REP g.leg0)\<^sup>*)))" 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 (tab\<^sub>0 ?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 "... = \\tab\<^sub>0 ?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 (tab\<^sub>0 ?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 "... = \\tab\<^sub>1 ?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 \ tab\<^sub>1 (Maps.REP g.leg1 \ (Maps.REP g.leg0)\<^sup>*)" 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 (tab\<^sub>0 (Maps.REP g.leg1 \ (Maps.REP g.leg0)\<^sup>*))) (src (Maps.REP g.leg0)\<^sup>*) (iso_class (tab\<^sub>0 (Maps.REP g.leg1 \ (Maps.REP g.leg0)\<^sup>*))), Leg1 = Maps.MkArr (src (tab\<^sub>0 (Maps.REP g.leg1 \ (Maps.REP g.leg0)\<^sup>*))) (Maps.Cod g.leg1) (iso_class (tab\<^sub>1 (Maps.REP g.leg1 \ (Maps.REP g.leg0)\<^sup>*)))\\ 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 = \\tab\<^sub>0 (Maps.REP g.leg1 \ (Maps.REP g.leg0)\<^sup>*)\\, Leg1 = \\tab\<^sub>1 (Maps.REP g.leg1 \ (Maps.REP g.leg0)\<^sup>*)\\\, Cod = \Leg0 = \\tab\<^sub>0 (Maps.REP g.leg1 \ (Maps.REP g.leg0)\<^sup>*)\\, Leg1 = \\tab\<^sub>1 (Maps.REP g.leg1 \ (Maps.REP g.leg0)\<^sup>*)\\\\" 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 \tab\<^sub>0 s\ \tab\<^sub>1 s\ s s.tab \tab\<^sub>0 s\ \tab\<^sub>1 s\ s using s.is_arrow_of_tabulations_in_maps by simp have \_dom_leg0_eq: "\.dom.leg0 = \\tab\<^sub>0 r\\" using \ Span.dom_char SPN_def [of r] by auto have \_dom_leg1_eq: "\.dom.leg1 = \\tab\<^sub>1 r\\" using \ Span.dom_char SPN_def [of r] by auto have \_cod_leg0_eq: "\.cod.leg0 = \\tab\<^sub>0 s\\" using \ Span.cod_char SPN_def [of s] by auto have \_cod_leg1_eq: "\.cod.leg1 = \\tab\<^sub>1 s\\" using \ Span.cod_char SPN_def [of s] by auto have 1: "tab\<^sub>0 s \ Maps.REP \.chine \ tab\<^sub>0 r" proof - have "tab\<^sub>0 s \ Maps.REP \.chine \ Maps.REP \.cod.leg0 \ Maps.REP \.chine" proof - have "Maps.REP \.cod.leg0 \ tab\<^sub>0 s" using \_cod_leg0_eq Maps.CLS_REP Maps.CLS_eqI Maps.REP_CLS s.satisfies_T0 by presburger thus ?thesis using hcomp_isomorphic_ide [of "Maps.REP \.cod.leg0" "tab\<^sub>0 s" "Maps.REP \.chine"] isomorphic_symmetric Maps.seq_char 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 "... \ tab\<^sub>0 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 \: "\\ : tab\<^sub>0 s \ Maps.REP \.chine \ tab\<^sub>0 r\ \ iso \" using 1 by blast have 2: "tab\<^sub>1 s \ Maps.REP \.chine \ tab\<^sub>1 r" proof - have "tab\<^sub>1 s \ Maps.REP \.chine \ Maps.REP \.cod.leg1 \ Maps.REP \.chine" proof - have "Maps.REP \.cod.leg1 \ tab\<^sub>1 s" using \_cod_leg1_eq Maps.CLS_REP Maps.CLS_eqI Maps.REP_CLS s.leg1_is_map by presburger thus ?thesis using hcomp_isomorphic_ide [of "Maps.REP \.cod.leg1" "tab\<^sub>1 s" "Maps.REP \.chine"] isomorphic_symmetric Maps.seq_char 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 "... \ tab\<^sub>1 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 \: "\\ : tab\<^sub>1 r \ tab\<^sub>1 s \ Maps.REP \.chine\ \ iso \" using 2 isomorphic_symmetric by blast define \ where "\ \ (s \ \) \ \[s, tab\<^sub>0 s, Maps.REP \.chine] \ (s.tab \ Maps.REP \.chine) \ \" have \: "\\ : tab\<^sub>1 r \ s \ tab\<^sub>0 r\" proof (unfold \_def, intro comp_in_homI) show "\\ : tab\<^sub>1 r \ tab\<^sub>1 s \ Maps.REP \.chine\" using \ by simp show 3: "\s.tab \ Maps.REP \.chine : tab\<^sub>1 s \ Maps.REP \.chine \ (s \ tab\<^sub>0 s) \ Maps.REP \.chine\" apply (intro hcomp_in_vhom) apply auto using "1" by fastforce show "\\[s, tab\<^sub>0 s, Maps.REP \.chine] : (s \ tab\<^sub>0 s) \ Maps.REP \.chine \ s \ tab\<^sub>0 s \ Maps.REP \.chine\" using s assoc_in_hom [of s "tab\<^sub>0 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 \ tab\<^sub>0 s \ Maps.REP \.chine \ s \ tab\<^sub>0 r\" using 1 s \ isomorphic_implies_hpar(4) src_eq by auto qed define \ where "\ \ r.T0.trnr\<^sub>\ s \ \ inv (r.T0.trnr\<^sub>\ r r.tab)" have \: "\\ : r \ s\" proof (unfold \_def, intro comp_in_homI) show "\inv (r.T0.trnr\<^sub>\ r r.tab) : r \ tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>*\" using r.yields_isomorphic_representation by fastforce show "\r.T0.trnr\<^sub>\ s \ : tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>* \ s\" using s \ src_eq r.T0.adjoint_transpose_right(2) [of s "tab\<^sub>1 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 \tab\<^sub>0 r\ \tab\<^sub>1 r\ s s.tab \tab\<^sub>0 s\ \tab\<^sub>1 s\ \ using \.is_arrow_of_tabulations_in_maps by simp have \_eq: "\ = \.\" proof - have "r.T0.trnr\<^sub>\ s \ \ inv (r.T0.trnr\<^sub>\ r r.tab) = r.T0.trnr\<^sub>\ s \.\ \ inv (r.T0.trnr\<^sub>\ r r.tab)" using \ \_def \.\_in_terms_of_\ by auto hence "r.T0.trnr\<^sub>\ s \ = r.T0.trnr\<^sub>\ 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\<^sub>\ s" "hom (tab\<^sub>1 r) (s \ tab\<^sub>0 r)" "hom (tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>*) s"] inj_on_def [of "r.T0.trnr\<^sub>\ s" "hom (tab\<^sub>1 r) (s \ tab\<^sub>0 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 diff --git a/thys/Bicategory/InternalAdjunction.thy b/thys/Bicategory/InternalAdjunction.thy --- a/thys/Bicategory/InternalAdjunction.thy +++ b/thys/Bicategory/InternalAdjunction.thy @@ -1,3435 +1,3435 @@ (* Title: InternalAdjunction Author: Eugene W. Stark , 2019 Maintainer: Eugene W. Stark *) section "Adjunctions in a Bicategory" theory InternalAdjunction imports CanonicalIsos Strictness begin text \ An \emph{internal adjunction} in a bicategory is a four-tuple \(f, g, \, \)\, where \f\ and \g\ are antiparallel 1-cells and \\\ : src f \ g \ f\\ and \\\ : f \ g \ src g\\ are 2-cells, such that the familiar ``triangle'' (or ``zig-zag'') identities are satisfied. We state the triangle identities in two equivalent forms, each of which is convenient in certain situations. \ locale adjunction_in_bicategory = adjunction_data_in_bicategory + assumes triangle_left: "(\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \) = \\<^sup>-\<^sup>1[f] \ \[f]" and triangle_right: "(g \ \) \ \[g, f, g] \ (\ \ g) = \\<^sup>-\<^sup>1[g] \ \[g]" begin lemma triangle_left': shows "\[f] \ (\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \) \ \\<^sup>-\<^sup>1[f] = f" using triangle_left triangle_equiv_form by simp lemma triangle_right': shows "\[g] \ (g \ \) \ \[g, f, g] \ (\ \ g) \ \\<^sup>-\<^sup>1[g] = g" using triangle_right triangle_equiv_form by simp end text \ Internal adjunctions have a number of properties, which we now develop, that generalize those of ordinary adjunctions involving functors and natural transformations. \ context bicategory begin lemma adjunction_unit_determines_counit: assumes "adjunction_in_bicategory (\) (\) \ \ src trg f g \ \" and "adjunction_in_bicategory (\) (\) \ \ src trg f g \ \'" shows "\ = \'" proof - interpret E: adjunction_in_bicategory V H \ \ src trg f g \ \ using assms(1) by auto interpret E': adjunction_in_bicategory V H \ \ src trg f g \ \' using assms(2) by auto text \ Note that since we want to prove the the result for an arbitrary bicategory, not just in for a strict bicategory, the calculation is a little more involved than one might expect from a treatment that suppresses canonical isomorphisms. \ have "\ \ \[f \ g] = \[trg f] \ (\ \ trg f)" using runit_naturality [of \] by simp have 1: "\[f \ g] = (f \ \[g]) \ \[f, g, src g]" using E.antipar runit_hcomp by simp have "\ = \ \ (f \ \[g] \ (g \ \') \ \[g, f, g] \ (\ \ g) \ \\<^sup>-\<^sup>1[g])" using E'.triangle_right' comp_arr_dom by simp also have "... = \ \ (f \ \[g]) \ (f \ g \ \') \ (f \ \[g, f, g]) \ (f \ \ \ g) \ (f \ \\<^sup>-\<^sup>1[g])" using E.antipar whisker_left by simp also have "... = \ \ ((f \ \[g]) \ (f \ g \ \')) \ (f \ \[g, f, g]) \ (f \ \ \ g) \ (f \ \\<^sup>-\<^sup>1[g])" using comp_assoc by simp also have "... = \ \ \[f \ g] \ (\\<^sup>-\<^sup>1[f, g, src g] \ (f \ g \ \')) \ (f \ \[g, f, g]) \ (f \ \ \ g) \ (f \ \\<^sup>-\<^sup>1[g])" proof - have "f \ \[g] = \[f \ g] \ \\<^sup>-\<^sup>1[f, g, src g]" using E.antipar(1) runit_hcomp(3) by auto thus ?thesis using comp_assoc by simp qed also have "... = (\ \ \[f \ g]) \ ((f \ g) \ \') \ \\<^sup>-\<^sup>1[f, g, f \ g] \ (f \ \[g, f, g]) \ (f \ \ \ g) \ (f \ \\<^sup>-\<^sup>1[g])" using E.antipar E'.counit_in_hom assoc'_naturality [of f g \'] comp_assoc by simp also have "... = \[trg f] \ ((\ \ trg f) \ ((f \ g) \ \')) \ \\<^sup>-\<^sup>1[f, g, f \ g] \ (f \ \[g, f, g]) \ (f \ \ \ g) \ (f \ \\<^sup>-\<^sup>1[g])" using E.antipar E.counit_in_hom runit_naturality [of \] comp_assoc by simp also have "... = (\[src g] \ (src g \ \')) \ (\ \ f \ g) \ \\<^sup>-\<^sup>1[f, g, f \ g] \ (f \ \[g, f, g]) \ (f \ \ \ g) \ (f \ \\<^sup>-\<^sup>1[g])" proof - have "(\ \ trg f) \ ((f \ g) \ \') = (src g \ \') \ (\ \ f \ g)" using E.antipar interchange E.counit_in_hom comp_arr_dom comp_cod_arr by (metis E'.counit_simps(1-3) E.counit_simps(1-3)) thus ?thesis using E.antipar comp_assoc unitor_coincidence by simp qed also have "... = \' \ \[f \ g] \ (\ \ f \ g) \ \\<^sup>-\<^sup>1[f, g, f \ g] \ (f \ \[g, f, g]) \ (f \ \ \ g) \ (f \ \\<^sup>-\<^sup>1[g])" proof - have "\[src g] \ (src g \ \') = \' \ \[f \ g]" using E.antipar lunit_naturality [of \'] by simp thus ?thesis using comp_assoc by simp qed also have "... = \' \ (\[f] \ g) \ (\\<^sup>-\<^sup>1[trg f, f, g] \ (\ \ f \ g)) \ \\<^sup>-\<^sup>1[f, g, f \ g] \ (f \ \[g, f, g]) \ (f \ \ \ g) \ (f \ \\<^sup>-\<^sup>1[g])" using E.antipar lunit_hcomp comp_assoc by simp also have "... = \' \ (\[f] \ g) \ ((\ \ f) \ g) \ (\\<^sup>-\<^sup>1[f \ g, f, g] \ \\<^sup>-\<^sup>1[f, g, f \ g] \ (f \ \[g, f, g])) \ (f \ \ \ g) \ (f \ \\<^sup>-\<^sup>1[g])" using E.antipar assoc'_naturality [of \ f g] comp_assoc by simp also have "... = \' \ (\[f] \ g) \ ((\ \ f) \ g) \ (\\<^sup>-\<^sup>1[f, g, f] \ g) \ (\\<^sup>-\<^sup>1[f, g \ f, g] \ (f \ \ \ g)) \ (f \ \\<^sup>-\<^sup>1[g])" proof - have "\\<^sup>-\<^sup>1[f \ g, f, g] \ \\<^sup>-\<^sup>1[f, g, f \ g] \ (f \ \[g, f, g]) = (\\<^sup>-\<^sup>1[f, g, f] \ g) \ \\<^sup>-\<^sup>1[f, g \ f, g]" using 1 E.antipar iso_assoc' iso_inv_iso pentagon' comp_assoc invert_side_of_triangle(2) [of "\\<^sup>-\<^sup>1[f \ g, f, g] \ \\<^sup>-\<^sup>1[f, g, f \ g]" "(\\<^sup>-\<^sup>1[f, g, f] \ g) \ \\<^sup>-\<^sup>1[f, g \ f, g]" "f \ \\<^sup>-\<^sup>1[g, f, g]"] by simp thus ?thesis using comp_assoc by simp qed also have "... = \' \ (\[f] \ g) \ ((\ \ f) \ g) \ (\\<^sup>-\<^sup>1[f, g, f] \ g) \ ((f \ \) \ g) \ \\<^sup>-\<^sup>1[f, trg g, g] \ (f \ \\<^sup>-\<^sup>1[g])" using E.antipar assoc'_naturality [of f \ g] comp_assoc by simp also have "... = \' \ (\[f] \ g) \ ((\ \ f) \ g) \ (\\<^sup>-\<^sup>1[f, g, f] \ g) \ ((f \ \) \ g) \ (\\<^sup>-\<^sup>1[f] \ g)" proof - have "\\<^sup>-\<^sup>1[f, trg g, g] \ (f \ \\<^sup>-\<^sup>1[g]) = \\<^sup>-\<^sup>1[f] \ g" proof - have "\\<^sup>-\<^sup>1[f] \ g = inv (\[f] \ g)" using E.antipar by simp also have "... = inv ((f \ \[g]) \ \[f, trg g, g])" using E.antipar by (simp add: triangle) also have "... = \\<^sup>-\<^sup>1[f, trg g, g] \ (f \ \\<^sup>-\<^sup>1[g])" using E.antipar inv_comp by simp finally show ?thesis by simp qed thus ?thesis by simp qed also have "... = \' \ (\[f] \ (\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \) \ \\<^sup>-\<^sup>1[f] \ g)" using E.antipar whisker_right by simp also have "... = \'" using E.triangle_left' comp_arr_dom by simp finally show ?thesis by simp qed end subsection "Adjoint Transpose" context adjunction_in_bicategory begin interpretation E: self_evaluation_map V H \ \ src trg .. notation E.eval ("\_\") text \ Just as for an ordinary adjunction between categories, an adjunction in a bicategory determines bijections between hom-sets. There are two versions of this relationship: depending on whether the transposition is occurring on the left (\emph{i.e.}~``output'') side or the right (\emph{i.e.}~``input'') side. \ definition trnl\<^sub>\ where "trnl\<^sub>\ v \ \ (g \ \) \ \[g, f, v] \ (\ \ v) \ \\<^sup>-\<^sup>1[v]" definition trnl\<^sub>\ where "trnl\<^sub>\ u \ \ \[u] \ (\ \ u) \ \\<^sup>-\<^sup>1[f, g, u] \ (f \ \)" lemma adjoint_transpose_left: assumes "ide u" and "ide v" and "src f = trg v" and "src g = trg u" shows "trnl\<^sub>\ v \ hom (f \ v) u \ hom v (g \ u)" and "trnl\<^sub>\ u \ hom v (g \ u) \ hom (f \ v) u" and "\\ : f \ v \ u\ \ trnl\<^sub>\ u (trnl\<^sub>\ v \) = \" and "\\ : v \ g \ u\ \ trnl\<^sub>\ v (trnl\<^sub>\ u \) = \" and "bij_betw (trnl\<^sub>\ v) (hom (f \ v) u) (hom v (g \ u))" and "bij_betw (trnl\<^sub>\ u) (hom v (g \ u)) (hom (f \ v) u)" proof - show A: "trnl\<^sub>\ v \ hom (f \ v) u \ hom v (g \ u)" using assms antipar trnl\<^sub>\_def by fastforce show B: "trnl\<^sub>\ u \ hom v (g \ u) \ hom (f \ v) u" using assms antipar trnl\<^sub>\_def by fastforce show C: "\\. \\ : f \ v \ u\ \ trnl\<^sub>\ u (trnl\<^sub>\ v \) = \" proof - fix \ assume \: "\\ : f \ v \ u\" have "trnl\<^sub>\ u (trnl\<^sub>\ v \) = \[u] \ (\ \ u) \ \\<^sup>-\<^sup>1[f, g, u] \ (f \ (g \ \) \ \[g, f, v] \ (\ \ v) \ \\<^sup>-\<^sup>1[v])" using trnl\<^sub>\_def trnl\<^sub>\_def by simp also have "... = \[u] \ (\ \ u) \ (\\<^sup>-\<^sup>1[f, g, u] \ (f \ g \ \)) \ (f \ \[g, f, v]) \ (f \ \ \ v) \ (f \ \\<^sup>-\<^sup>1[v])" using assms \ antipar whisker_left comp_assoc by auto also have "... = \[u] \ ((\ \ u) \ ((f \ g) \ \)) \ \\<^sup>-\<^sup>1[f, g, f \ v] \ (f \ \[g, f, v]) \ (f \ \ \ v) \ (f \ \\<^sup>-\<^sup>1[v])" using assms \ antipar assoc'_naturality [of f g \] comp_assoc by fastforce also have "... = \[u] \ (trg u \ \) \ (\ \ f \ v) \ \\<^sup>-\<^sup>1[f, g, f \ v] \ (f \ \[g, f, v]) \ (f \ \ \ v) \ (f \ \\<^sup>-\<^sup>1[v])" proof - have "(\ \ u) \ ((f \ g) \ \) = (trg u \ \) \ (\ \ f \ v)" using assms \ antipar comp_cod_arr comp_arr_dom interchange [of "trg u" \ \ "f \ v"] interchange [of \ "f \ g" u \] by auto thus ?thesis using comp_assoc by simp qed also have "... = \[u] \ (trg u \ \) \ \[trg f, f, v] \ ((\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \) \ v) \ \\<^sup>-\<^sup>1[f, trg v, v] \ (f \ \\<^sup>-\<^sup>1[v])" proof - have 1: "(\ \ f \ v) \ \\<^sup>-\<^sup>1[f, g, f \ v] \ (f \ \[g, f, v]) \ (f \ \ \ v) = \[trg f, f, v] \ ((\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \) \ v) \ \\<^sup>-\<^sup>1[f, trg v, v]" proof - have "(\ \ f \ v) \ \\<^sup>-\<^sup>1[f, g, f \ v] \ (f \ \[g, f, v]) \ (f \ \ \ v) = (\ \ f \ v) \ \[f \ g, f, v] \ (\\<^sup>-\<^sup>1[f, g, f] \ v) \ \\<^sup>-\<^sup>1[f, g \ f, v] \ (f \ \ \ v)" proof - have "(\ \ f \ v) \ \\<^sup>-\<^sup>1[f, g, f \ v] \ (f \ \[g, f, v]) \ (f \ \ \ v) = (\ \ f \ v) \ (\\<^sup>-\<^sup>1[f, g, f \ v] \ (f \ \[g, f, v])) \ (f \ \ \ v)" using comp_assoc by simp also have "... = (\ \ f \ v) \ \[f \ g, f, v] \ (\\<^sup>-\<^sup>1[f, g, f] \ v) \ \\<^sup>-\<^sup>1[f, g \ f, v] \ (f \ \ \ v)" proof - have "\\<^sup>-\<^sup>1[f, g, f \ v] \ (f \ \[g, f, v]) = \[f \ g, f, v] \ (\\<^sup>-\<^sup>1[f, g, f] \ v) \ \\<^sup>-\<^sup>1[f, g \ f, v]" using assms antipar canI_associator_0 whisker_can_left_0 whisker_can_right_0 canI_associator_hcomp(1-3) by simp thus ?thesis using comp_assoc by simp qed finally show ?thesis by blast qed also have "... = ((\ \ f \ v) \ \[f \ g, f, v]) \ (\\<^sup>-\<^sup>1[f, g, f] \ v) \ ((f \ \) \ v) \ \\<^sup>-\<^sup>1[f, trg v, v]" using assms \ antipar assoc'_naturality [of f \ v] comp_assoc by simp also have "... = (\[trg f, f, v] \ ((\ \ f) \ v)) \ (\\<^sup>-\<^sup>1[f, g, f] \ v) \ ((f \ \) \ v) \ \\<^sup>-\<^sup>1[f, trg v, v]" using assms \ antipar assoc_naturality [of \ f v] by simp also have "... = \[trg f, f, v] \ (((\ \ f) \ v) \ (\\<^sup>-\<^sup>1[f, g, f] \ v) \ ((f \ \) \ v)) \ \\<^sup>-\<^sup>1[f, trg v, v]" using comp_assoc by simp also have "... = \[trg f, f, v] \ ((\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \) \ v) \ \\<^sup>-\<^sup>1[f, trg v, v]" using assms \ antipar whisker_right by simp finally show ?thesis by simp qed show ?thesis using 1 comp_assoc by metis qed also have "... = \[u] \ (trg u \ \) \ \[trg f, f, v] \ (\\<^sup>-\<^sup>1[f] \ \[f] \ v) \ \\<^sup>-\<^sup>1[f, trg v, v] \ (f \ \\<^sup>-\<^sup>1[v])" using assms \ antipar triangle_left by simp also have "... = \[u] \ (trg u \ \) \ can (\<^bold>\trg u\<^bold>\\<^sub>0 \<^bold>\ \<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\v\<^bold>\) (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\v\<^bold>\)" using assms \ antipar canI_unitor_0 canI_associator_1 canI_associator_1(1-2) [of f v] whisker_can_right_0 whisker_can_left_0 by simp also have "... = \[u] \ (trg u \ \) \ \\<^sup>-\<^sup>1[f \ v]" unfolding can_def using assms antipar comp_arr_dom comp_cod_arr \_ide_simp by simp also have "... = (\[u] \ \\<^sup>-\<^sup>1[u]) \ \" using assms \ lunit'_naturality [of \] comp_assoc by auto also have "... = \" using assms \ comp_cod_arr iso_lunit comp_arr_inv inv_is_inverse by auto finally show "trnl\<^sub>\ u (trnl\<^sub>\ v \) = \" by simp qed show D: "\\. \\ : v \ g \ u\ \ trnl\<^sub>\ v (trnl\<^sub>\ u \) = \" proof - fix \ assume \: "\\ : v \ g \ u\" have "trnl\<^sub>\ v (trnl\<^sub>\ u \) = (g \ \[u] \ (\ \ u) \ \\<^sup>-\<^sup>1[f, g, u] \ (f \ \)) \ \[g, f, v] \ (\ \ v) \ \\<^sup>-\<^sup>1[v]" using trnl\<^sub>\_def trnl\<^sub>\_def by simp also have "... = (g \ \[u]) \ (g \ \ \ u) \ (g \ \\<^sup>-\<^sup>1[f, g, u]) \ (g \ f \ \) \ \[g, f, v] \ (\ \ v) \ \\<^sup>-\<^sup>1[v]" using assms \ antipar interchange [of g "g \ g \ g"] comp_assoc by auto also have "... = ((g \ \[u]) \ (g \ \ \ u) \ (g \ \\<^sup>-\<^sup>1[f, g, u]) \ \[g, f, g \ u] \ (\ \ g \ u)) \ (trg v \ \) \ \\<^sup>-\<^sup>1[v]" proof - have "(g \ f \ \) \ \[g, f, v] \ (\ \ v) \ \\<^sup>-\<^sup>1[v] = \[g, f, g \ u] \ (\ \ g \ u) \ (trg v \ \) \ \\<^sup>-\<^sup>1[v]" proof - have "(g \ f \ \) \ \[g, f, v] \ (\ \ v) \ \\<^sup>-\<^sup>1[v] = \[g, f, g \ u] \ ((g \ f) \ \) \ (\ \ v) \ \\<^sup>-\<^sup>1[v]" proof - have "(g \ f \ \) \ \[g, f, v] = \[g, f, g \ u] \ ((g \ f) \ \)" using assms \ antipar assoc_naturality [of g f \] by auto thus ?thesis using assms comp_assoc by metis qed also have "... = \[g, f, g \ u] \ (\ \ g \ u) \ (trg v \ \) \ \\<^sup>-\<^sup>1[v]" proof - have "((g \ f) \ \) \ (\ \ v) = (\ \ g \ u) \ (trg v \ \)" using assms \ antipar comp_arr_dom comp_cod_arr interchange [of "g \ f" \ \ v] interchange [of \ "trg v" "g \ u" \] by auto thus ?thesis using comp_assoc by metis qed finally show ?thesis by blast qed thus ?thesis using comp_assoc by simp qed also have "... = \[g \ u] \ (trg v \ \) \ \\<^sup>-\<^sup>1[v]" proof - have "(g \ \[u]) \ (g \ \ \ u) \ (g \ \\<^sup>-\<^sup>1[f, g, u]) \ \[g, f, g \ u] \ (\ \ g \ u) = \[g \ u]" proof - have "(g \ \[u]) \ (g \ \ \ u) \ (g \ \\<^sup>-\<^sup>1[f, g, u]) \ \[g, f, g \ u] \ (\ \ g \ u) = (g \ \[u]) \ \[g, trg u, u] \ ((g \ \) \ \[g, f, g] \ (\ \ g) \ u) \ \\<^sup>-\<^sup>1[trg v, g, u]" proof - have "(g \ \[u]) \ (g \ \ \ u) \ (g \ \\<^sup>-\<^sup>1[f, g, u]) \ \[g, f, g \ u] \ (\ \ g \ u) = (g \ \[u]) \ (g \ \ \ u) \ (g \ \\<^sup>-\<^sup>1[f, g, u]) \ \[g, f, g \ u] \ ((\ \ g \ u) \ \[trg v, g, u]) \ \\<^sup>-\<^sup>1[trg v, g, u]" using assms antipar comp_arr_dom comp_assoc comp_assoc_assoc'(1) by simp also have "... = (g \ \[u]) \ (g \ \ \ u) \ (g \ \\<^sup>-\<^sup>1[f, g, u]) \ \[g, f, g \ u] \ (\[g \ f, g, u] \ ((\ \ g) \ u)) \ \\<^sup>-\<^sup>1[trg v, g, u]" using assms antipar assoc_naturality [of \ g u] by simp also have "... = (g \ \[u]) \ (g \ \ \ u) \ ((g \ \\<^sup>-\<^sup>1[f, g, u]) \ \[g, f, g \ u] \ \[g \ f, g, u]) \ ((\ \ g) \ u) \ \\<^sup>-\<^sup>1[trg v, g, u]" using comp_assoc by simp also have "... = (g \ \[u]) \ ((\[g, trg u, u] \ \\<^sup>-\<^sup>1[g, trg u, u]) \ (g \ \ \ u)) \ ((g \ \\<^sup>-\<^sup>1[f, g, u]) \ \[g, f, g \ u] \ \[g \ f, g, u]) \ ((\ \ g) \ u) \ \\<^sup>-\<^sup>1[trg v, g, u]" proof - have "(\[g, trg u, u] \ \\<^sup>-\<^sup>1[g, trg u, u]) \ (g \ \ \ u) = g \ \ \ u" using assms antipar comp_cod_arr comp_assoc_assoc'(1) by simp thus ?thesis using comp_assoc by simp qed also have "... = (g \ \[u]) \ \[g, trg u, u] \ (\\<^sup>-\<^sup>1[g, trg u, u] \ (g \ \ \ u)) \ (g \ \\<^sup>-\<^sup>1[f, g, u]) \ \[g, f, g \ u] \ \[g \ f, g, u] \ ((\ \ g) \ u) \ \\<^sup>-\<^sup>1[trg v, g, u]" using comp_assoc by simp also have "... = (g \ \[u]) \ \[g, trg u, u] \ (((g \ \) \ u) \ (\\<^sup>-\<^sup>1[g, f \ g, u] \ (g \ \\<^sup>-\<^sup>1[f, g, u]) \ \[g, f, g \ u] \ \[g \ f, g, u]) \ ((\ \ g) \ u)) \ \\<^sup>-\<^sup>1[trg v, g, u]" using assms antipar assoc'_naturality [of g \ u] comp_assoc by simp also have "... = (g \ \[u]) \ \[g, trg u, u] \ ((g \ \) \ \[g, f, g] \ (\ \ g) \ u) \ \\<^sup>-\<^sup>1[trg v, g, u]" proof - have "\\<^sup>-\<^sup>1[g, f \ g, u] \ (g \ \\<^sup>-\<^sup>1[f, g, u]) \ \[g, f, g \ u] \ \[g \ f, g, u] = \[g, f, g] \ u" using assms antipar canI_associator_0 whisker_can_left_0 whisker_can_right_0 canI_associator_hcomp by simp hence "((g \ \) \ u) \ (\\<^sup>-\<^sup>1[g, f \ g, u] \ (g \ \\<^sup>-\<^sup>1[f, g, u]) \ \[g, f, g \ u] \ \[g \ f, g, u]) \ ((\ \ g) \ u) = (g \ \) \ \[g, f, g] \ (\ \ g) \ u" using assms antipar whisker_right by simp thus ?thesis by simp qed finally show ?thesis by blast qed also have "... = (g \ \[u]) \ \[g, trg u, u] \ (\\<^sup>-\<^sup>1[g] \ \[g] \ u) \ \\<^sup>-\<^sup>1[trg g, g, u]" using assms antipar triangle_right by simp also have "... = can (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\u\<^bold>\) (\<^bold>\trg g\<^bold>\\<^sub>0 \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\u\<^bold>\)" proof - have "(g \ \[u]) \ \[g, trg u, u] \ (\\<^sup>-\<^sup>1[g] \ \[g] \ u) \ \\<^sup>-\<^sup>1[trg g, g, u] = ((g \ \[u]) \ \[g, trg u, u] \ (\\<^sup>-\<^sup>1[g] \ \[g] \ u) \ \\<^sup>-\<^sup>1[trg g, g, u])" using comp_assoc by simp moreover have "... = can (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\u\<^bold>\) (\<^bold>\trg g\<^bold>\\<^sub>0 \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\u\<^bold>\)" using assms antipar canI_unitor_0 canI_associator_1 [of g u] inv_can whisker_can_left_0 whisker_can_right_0 by simp ultimately show ?thesis by simp qed also have "... = \[g \ u]" unfolding can_def using assms comp_arr_dom comp_cod_arr \_ide_simp by simp finally show ?thesis by simp qed thus ?thesis by simp qed also have "... = (\[g \ u] \ \\<^sup>-\<^sup>1[g \ u]) \ \" using assms \ lunit'_naturality comp_assoc by auto also have "... = \" using assms \ comp_cod_arr iso_lunit comp_arr_inv inv_is_inverse by auto finally show "trnl\<^sub>\ v (trnl\<^sub>\ u \) = \" by simp qed show "bij_betw (trnl\<^sub>\ v) (hom (f \ v) u) (hom v (g \ u))" using A B C D by (intro bij_betwI, auto) show "bij_betw (trnl\<^sub>\ u) (hom v (g \ u)) (hom (f \ v) u)" using A B C D by (intro bij_betwI, auto) qed lemma trnl\<^sub>\_comp: assumes "ide u" and "seq \ \" and "src f = trg \" shows "trnl\<^sub>\ u (\ \ \) = trnl\<^sub>\ u \ \ (f \ \)" using assms trnl\<^sub>\_def whisker_left [of f \ \] comp_assoc by auto definition trnr\<^sub>\ where "trnr\<^sub>\ v \ \ (\ \ f) \ \\<^sup>-\<^sup>1[v, g, f] \ (v \ \) \ \\<^sup>-\<^sup>1[v]" definition trnr\<^sub>\ where "trnr\<^sub>\ u \ \ \[u] \ (u \ \) \ \[u, f, g] \ (\ \ g)" lemma adjoint_transpose_right: assumes "ide u" and "ide v" and "src v = trg g" and "src u = trg f" shows "trnr\<^sub>\ v \ hom (v \ g) u \ hom v (u \ f)" and "trnr\<^sub>\ u \ hom v (u \ f) \ hom (v \ g) u" and "\\ : v \ g \ u\ \ trnr\<^sub>\ u (trnr\<^sub>\ v \) = \" and "\\ : v \ u \ f\ \ trnr\<^sub>\ v (trnr\<^sub>\ u \) = \" and "bij_betw (trnr\<^sub>\ v) (hom (v \ g) u) (hom v (u \ f))" and "bij_betw (trnr\<^sub>\ u) (hom v (u \ f)) (hom (v \ g) u)" proof - show A: "trnr\<^sub>\ v \ hom (v \ g) u \ hom v (u \ f)" using assms antipar trnr\<^sub>\_def by fastforce show B: "trnr\<^sub>\ u \ hom v (u \ f) \ hom (v \ g) u" using assms antipar trnr\<^sub>\_def by fastforce show C: "\\. \\ : v \ g \ u\ \ trnr\<^sub>\ u (trnr\<^sub>\ v \) = \" proof - fix \ assume \: "\\ : v \ g \ u\" have "trnr\<^sub>\ u (trnr\<^sub>\ v \) = \[u] \ (u \ \) \ \[u, f, g] \ ((\ \ f) \ \\<^sup>-\<^sup>1[v, g, f] \ (v \ \) \ \\<^sup>-\<^sup>1[v] \ g)" unfolding trnr\<^sub>\_def trnr\<^sub>\_def by simp also have "... = \[u] \ (u \ \) \ (\[u, f, g] \ ((\ \ f) \ g)) \ (\\<^sup>-\<^sup>1[v, g, f] \ g) \ ((v \ \) \ g) \ (\\<^sup>-\<^sup>1[v] \ g)" using assms \ antipar whisker_right comp_assoc by auto also have "... = \[u] \ (u \ \) \ ((\ \ f \ g) \ \[v \ g, f, g]) \ (\\<^sup>-\<^sup>1[v, g, f] \ g) \ ((v \ \) \ g) \ (\\<^sup>-\<^sup>1[v] \ g)" using assms \ antipar assoc_naturality [of \ f g] by auto also have "... = \[u] \ ((u \ \) \ (\ \ f \ g)) \ \[v \ g, f, g] \ (\\<^sup>-\<^sup>1[v, g, f] \ g) \ ((v \ \) \ g) \ (\\<^sup>-\<^sup>1[v] \ g)" using comp_assoc by auto also have "... = \[u] \ (\ \ src u) \ ((v \ g) \ \) \ \[v \ g, f, g] \ (\\<^sup>-\<^sup>1[v, g, f] \ g) \ ((v \ \) \ g) \ (\\<^sup>-\<^sup>1[v] \ g)" proof - have "(u \ \) \ (\ \ f \ g) = (\ \ src u) \ ((v \ g) \ \)" using assms \ antipar comp_arr_dom comp_cod_arr interchange [of \ "v \ g" "src u" \] interchange [of u \ \ "f \ g"] by auto thus ?thesis using comp_assoc by simp qed also have "... = \[u] \ (\ \ src u) \ (((v \ g) \ \) \ \[v \ g, f, g] \ (\\<^sup>-\<^sup>1[v, g, f] \ g) \ ((v \ \) \ g)) \ (\\<^sup>-\<^sup>1[v] \ g)" using comp_assoc by simp also have "... = \[u] \ (\ \ src u) \ (\\<^sup>-\<^sup>1[v, g, src u] \ (v \ (g \ \) \ \[g, f, g] \ (\ \ g)) \ \[v, src v, g]) \ (\\<^sup>-\<^sup>1[v] \ g)" proof - have "((v \ g) \ \) \ \[v \ g, f, g] \ (\\<^sup>-\<^sup>1[v, g, f] \ g) \ ((v \ \) \ g) = \\<^sup>-\<^sup>1[v, g, src u] \ (v \ (g \ \) \ \[g, f, g] \ (\ \ g)) \ \[v, src v, g]" proof - have "((v \ g) \ \) \ \[v \ g, f, g] \ (\\<^sup>-\<^sup>1[v, g, f] \ g) \ ((v \ \) \ g) = ((\\<^sup>-\<^sup>1[v, g, src u] \ \[v, g, src u]) \ ((v \ g) \ \)) \ \[v \ g, f, g] \ (\\<^sup>-\<^sup>1[v, g, f] \ g) \ ((v \ \) \ g)" proof - have "arr v \ dom v = v \ cod v = v" using assms(2) ide_char by blast moreover have "arr g \ dom g = g \ cod g = g" using ide_right ide_char by blast ultimately show ?thesis by (metis (no_types) antipar(2) assms(3-4) assoc_naturality counit_simps(1,3,5) hcomp_reassoc(1) comp_assoc) qed also have "... = \\<^sup>-\<^sup>1[v, g, src u] \ (\[v, g, src u] \ ((v \ g) \ \)) \ \[v \ g, f, g] \ (\\<^sup>-\<^sup>1[v, g, f] \ g) \ ((v \ \) \ g)" using comp_assoc by simp also have "... = \\<^sup>-\<^sup>1[v, g, src u] \ ((v \ g \ \) \ \[v, g, f \ g]) \ \[v \ g, f, g] \ (\\<^sup>-\<^sup>1[v, g, f] \ g) \ (\\<^sup>-\<^sup>1[v, g \ f, g] \ \[v, g \ f, g]) \ ((v \ \) \ g)" proof - have "\[v, g, src u] \ ((v \ g) \ \) = (v \ g \ \) \ \[v, g, f \ g]" using assms antipar assoc_naturality [of v g \] by simp moreover have "(\\<^sup>-\<^sup>1[v, g \ f, g] \ \[v, g \ f, g]) \ ((v \ \) \ g) = (v \ \) \ g" using assms antipar comp_cod_arr comp_assoc_assoc'(2) by simp ultimately show ?thesis by simp qed also have "... = \\<^sup>-\<^sup>1[v, g, src u] \ (v \ g \ \) \ \[v, g, f \ g] \ \[v \ g, f, g] \ (\\<^sup>-\<^sup>1[v, g, f] \ g) \ \\<^sup>-\<^sup>1[v, g \ f, g] \ \[v, g \ f, g] \ ((v \ \) \ g)" using comp_assoc by simp also have "... = \\<^sup>-\<^sup>1[v, g, src u] \ ((v \ g \ \) \ (\[v, g, f \ g] \ \[v \ g, f, g] \ (\\<^sup>-\<^sup>1[v, g, f] \ g) \ \\<^sup>-\<^sup>1[v, g \ f, g]) \ (v \ \ \ g)) \ \[v, src v, g]" using assms antipar assoc_naturality [of v \ g] comp_assoc by simp also have "... = \\<^sup>-\<^sup>1[v, g, src u] \ ((v \ g \ \) \ (v \ \[g, f, g]) \ (v \ \ \ g)) \ \[v, src v, g]" proof - have "\[v, g, f \ g] \ \[v \ g, f, g] \ (\\<^sup>-\<^sup>1[v, g, f] \ g) \ \\<^sup>-\<^sup>1[v, g \ f, g] = v \ \[g, f, g]" using assms antipar canI_associator_0 canI_associator_hcomp whisker_can_left_0 whisker_can_right_0 by simp thus ?thesis using assms antipar whisker_left by simp qed also have "... = \\<^sup>-\<^sup>1[v, g, src u] \ (v \ (g \ \) \ \[g, f, g] \ (\ \ g)) \ \[v, src v, g]" using assms antipar whisker_left by simp finally show ?thesis by simp qed thus ?thesis by auto qed also have "... = \[u] \ (\ \ src u) \ \\<^sup>-\<^sup>1[v, g, src u] \ (v \ \\<^sup>-\<^sup>1[g] \ \[g]) \ \[v, src v, g] \ (\\<^sup>-\<^sup>1[v] \ g)" using triangle_right comp_assoc by simp also have "... = \[u] \ (\ \ src u) \ \\<^sup>-\<^sup>1[v \ g]" proof - have "\\<^sup>-\<^sup>1[v, g, src u] \ (v \ \\<^sup>-\<^sup>1[g] \ \[g]) \ \[v, src v, g] \ (\\<^sup>-\<^sup>1[v] \ g) = \\<^sup>-\<^sup>1[v \ g]" proof - have "\\<^sup>-\<^sup>1[v, g, src u] \ (v \ \\<^sup>-\<^sup>1[g] \ \[g]) \ \[v, src v, g] \ (\\<^sup>-\<^sup>1[v] \ g) = \\<^sup>-\<^sup>1[v, g, trg f] \ can (\<^bold>\v\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\src g\<^bold>\\<^sub>0) (\<^bold>\v\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\)" using assms canI_unitor_0 canI_associator_1(2-3) whisker_can_left_0(1) whisker_can_right_0 by simp also have "... = \\<^sup>-\<^sup>1[v, g, src g] \ can (\<^bold>\v\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\src g\<^bold>\\<^sub>0) (\<^bold>\v\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\)" using antipar by simp (* TODO: There should be an alternate version of whisker_can_left for this. *) also have "... = \\<^sup>-\<^sup>1[v, g, src g] \ (v \ can (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\src g\<^bold>\\<^sub>0) \<^bold>\g\<^bold>\)" using assms canI_unitor_0(2) whisker_can_left_0 by simp also have "... = \\<^sup>-\<^sup>1[v, g, src g] \ (v \ \\<^sup>-\<^sup>1[g])" using assms canI_unitor_0(2) by simp also have "... = \\<^sup>-\<^sup>1[v \ g]" using assms runit_hcomp(2) by simp finally have "\\<^sup>-\<^sup>1[v, g, src u] \ (v \ \\<^sup>-\<^sup>1[g] \ \[g]) \ \[v, src v, g] \ (\\<^sup>-\<^sup>1[v] \ g) = \\<^sup>-\<^sup>1[v \ g]" by simp thus ?thesis by simp qed thus ?thesis by simp qed also have "... = (\[u] \ \\<^sup>-\<^sup>1[u]) \ \" using assms \ runit'_naturality [of \] comp_assoc by auto also have "... = \" using \ comp_cod_arr iso_runit inv_is_inverse comp_arr_inv by auto finally show "trnr\<^sub>\ u (trnr\<^sub>\ v \) = \" by simp qed show D: "\\. \\ : v \ u \ f\ \ trnr\<^sub>\ v (trnr\<^sub>\ u \) = \" proof - fix \ assume \: "\\ : v \ u \ f\" have "trnr\<^sub>\ v (trnr\<^sub>\ u \) = (\[u] \ (u \ \) \ \[u, f, g] \ (\ \ g) \ f) \ \\<^sup>-\<^sup>1[v, g, f] \ (v \ \) \ \\<^sup>-\<^sup>1[v]" unfolding trnr\<^sub>\_def trnr\<^sub>\_def by simp also have "... = (\[u] \ f) \ ((u \ \) \ f) \ (\[u, f, g] \ f) \ (((\ \ g) \ f) \ \\<^sup>-\<^sup>1[v, g, f]) \ (v \ \) \ \\<^sup>-\<^sup>1[v]" using assms \ antipar whisker_right comp_assoc by auto also have "... = (\[u] \ f) \ ((u \ \) \ f) \ (\[u, f, g] \ f) \ (\\<^sup>-\<^sup>1[u \ f, g, f] \ (\ \ g \ f)) \ (v \ \) \ \\<^sup>-\<^sup>1[v]" using assms \ antipar assoc'_naturality [of \ g f] by auto also have "... = (\[u] \ f) \ ((u \ \) \ f) \ (\[u, f, g] \ f) \ \\<^sup>-\<^sup>1[u \ f, g, f] \ ((\ \ g \ f) \ (v \ \)) \ \\<^sup>-\<^sup>1[v]" using comp_assoc by simp also have "... = (\[u] \ f) \ ((u \ \) \ f) \ (\[u, f, g] \ f) \ \\<^sup>-\<^sup>1[u \ f, g, f] \ (((u \ f) \ \) \ (\ \ src v)) \ \\<^sup>-\<^sup>1[v]" proof - have "(\ \ g \ f) \ (v \ \) = ((u \ f) \ \) \ (\ \ src v)" using assms \ antipar interchange [of "u \ f" \ \ "src v"] interchange [of \ v "g \ f" \] comp_arr_dom comp_cod_arr by auto thus ?thesis by simp qed also have "... = ((\[u] \ f) \ ((u \ \) \ f) \ ((\[u, f, g] \ f) \ \\<^sup>-\<^sup>1[u \ f, g, f]) \ ((u \ f) \ \)) \ (\ \ src v) \ \\<^sup>-\<^sup>1[v]" using comp_assoc by simp also have "... = ((\[u] \ f) \ ((u \ \) \ f) \ (\\<^sup>-\<^sup>1[u, f \ g, f] \ (u \ \\<^sup>-\<^sup>1[f, g, f]) \ \[u, f, g \ f]) \ ((u \ f) \ \)) \ (\ \ src v) \ \\<^sup>-\<^sup>1[v]" using assms antipar canI_associator_hcomp canI_associator_0 whisker_can_left_0 whisker_can_right_0 by simp also have "... = ((\[u] \ f) \ (((u \ \) \ f) \ \\<^sup>-\<^sup>1[u, f \ g, f]) \ (u \ \\<^sup>-\<^sup>1[f, g, f]) \ (\[u, f, g \ f]) \ ((u \ f) \ \)) \ (\ \ src v) \ \\<^sup>-\<^sup>1[v]" using comp_assoc by simp also have "... = ((\[u] \ f) \ (\\<^sup>-\<^sup>1[u, src u, f] \ (u \ \ \ f)) \ (u \ \\<^sup>-\<^sup>1[f, g, f]) \ ((u \ f \ \) \ \[u, f, src f])) \ (\ \ src v) \ \\<^sup>-\<^sup>1[v]" using assms antipar assoc'_naturality [of u \ f] assoc_naturality [of u f \] by auto also have "... = (\[u] \ f) \ \\<^sup>-\<^sup>1[u, src u, f] \ ((u \ \ \ f) \ (u \ \\<^sup>-\<^sup>1[f, g, f]) \ (u \ f \ \)) \ \[u, f, src f] \ (\ \ src v) \ \\<^sup>-\<^sup>1[v]" using comp_assoc by simp also have "... = (\[u] \ f) \ \\<^sup>-\<^sup>1[u, src u, f] \ (u \ (\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \)) \ \[u, f, src f] \ (\ \ src v) \ \\<^sup>-\<^sup>1[v]" using assms antipar whisker_left by auto also have "... = ((\[u] \ f) \ (\\<^sup>-\<^sup>1[u, src u, f] \ (u \ \\<^sup>-\<^sup>1[f] \ \[f]) \ \[u, f, src f])) \ (\ \ src v) \ \\<^sup>-\<^sup>1[v]" using assms antipar triangle_left comp_assoc by simp also have "... = \[u \ f] \ (\ \ src v) \ \\<^sup>-\<^sup>1[v]" proof - have "(\[u] \ f) \ \\<^sup>-\<^sup>1[u, src u, f] \ (u \ \\<^sup>-\<^sup>1[f] \ \[f]) \ \[u, f, src f] = ((u \ \[f]) \ (\[u, src u, f] \ \\<^sup>-\<^sup>1[u, src u, f])) \ (u \ \\<^sup>-\<^sup>1[f] \ \[f]) \ \[u, f, src f]" using assms ide_left ide_right antipar triangle comp_assoc by metis also have "... = (u \ \[f]) \ \[u, f, src f]" using assms antipar canI_associator_1 canI_unitor_0 whisker_can_left_0 whisker_can_right_0 canI_associator_1 by simp also have "... = \[u \ f]" using assms antipar runit_hcomp by simp finally show ?thesis by simp qed also have "... = (\[u \ f] \ \\<^sup>-\<^sup>1[u \ f]) \ \" using assms \ runit'_naturality [of \] comp_assoc by auto also have "... = \" using assms \ comp_cod_arr comp_arr_inv inv_is_inverse iso_runit by auto finally show "trnr\<^sub>\ v (trnr\<^sub>\ u \) = \" by auto qed show "bij_betw (trnr\<^sub>\ v) (hom (v \ g) u) (hom v (u \ f))" using A B C D by (intro bij_betwI, auto) show "bij_betw (trnr\<^sub>\ u) (hom v (u \ f)) (hom (v \ g) u)" using A B C D by (intro bij_betwI, auto) qed lemma trnr\<^sub>\_comp: assumes "ide v" and "seq \ \" and "src \ = trg f" shows "trnr\<^sub>\ v (\ \ \) = (\ \ f) \ trnr\<^sub>\ v \" using assms trnr\<^sub>\_def whisker_right comp_assoc by auto end text \ It is useful to have at hand the simpler versions of the preceding results that hold in a normal bicategory and in a strict bicategory. \ locale adjunction_in_normal_bicategory = normal_bicategory + adjunction_in_bicategory begin lemma triangle_left: shows "(\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \) = f" using triangle_left strict_lunit strict_runit by simp lemma triangle_right: shows "(g \ \) \ \[g, f, g] \ (\ \ g) = g" using triangle_right strict_lunit strict_runit by simp lemma trnr\<^sub>\_eq: assumes "ide u" and "ide v" and "src v = trg g" and "src u = trg f" and "\ \ hom (v \ g) u" shows "trnr\<^sub>\ v \ = (\ \ f) \ \\<^sup>-\<^sup>1[v, g, f] \ (v \ \)" unfolding trnr\<^sub>\_def using assms antipar strict_runit' comp_arr_ide [of "\\<^sup>-\<^sup>1[v]" "v \ \"] hcomp_arr_obj by auto lemma trnr\<^sub>\_eq: assumes "ide u" and "ide v" and "src v = trg g" and "src u = trg f" and "\ \ hom v (u \ f)" shows "trnr\<^sub>\ u \ = (u \ \) \ \[u, f, g] \ (\ \ g)" unfolding trnr\<^sub>\_def using assms antipar strict_runit comp_ide_arr hcomp_arr_obj by auto lemma trnl\<^sub>\_eq: assumes "ide u" and "ide v" and "src f = trg v" and "src g = trg u" and "\ \ hom (f \ v) u" shows "trnl\<^sub>\ v \ = (g \ \) \ \[g, f, v] \ (\ \ v)" using assms trnl\<^sub>\_def antipar strict_lunit comp_arr_dom hcomp_obj_arr by auto lemma trnl\<^sub>\_eq: assumes "ide u" and "ide v" and "src f = trg v" and "src g = trg u" and "\ \ hom v (g \ u)" shows "trnl\<^sub>\ u \ = (\ \ u) \ \\<^sup>-\<^sup>1[f, g, u] \ (f \ \)" using assms trnl\<^sub>\_def antipar strict_lunit comp_cod_arr hcomp_obj_arr by auto end locale adjunction_in_strict_bicategory = strict_bicategory + adjunction_in_normal_bicategory begin lemma triangle_left: shows "(\ \ f) \ (f \ \) = f" using ide_left ide_right antipar triangle_left strict_assoc' comp_cod_arr by (metis dom_eqI ideD(1) seqE) lemma triangle_right: shows "(g \ \) \ (\ \ g) = g" using ide_left ide_right antipar triangle_right strict_assoc comp_cod_arr by (metis ideD(1) ideD(2) seqE) lemma trnr\<^sub>\_eq: assumes "ide u" and "ide v" and "src v = trg g" and "src u = trg f" and "\ \ hom (v \ g) u" shows "trnr\<^sub>\ v \ = (\ \ f) \ (v \ \)" proof - have "trnr\<^sub>\ v \ = (\ \ f) \ \\<^sup>-\<^sup>1[v, g, f] \ (v \ \)" using assms trnr\<^sub>\_eq [of u v \] by simp also have "... = (\ \ f) \ (v \ \)" proof - have "\\<^sup>-\<^sup>1[v, g, f] \ (v \ \) = (v \ \)" proof - have "ide \\<^sup>-\<^sup>1[v, g, f]" using assms antipar strict_assoc' by simp moreover have "seq \\<^sup>-\<^sup>1[v, g, f] (v \ \)" using assms antipar by simp ultimately show ?thesis using comp_ide_arr by simp qed thus ?thesis by simp qed finally show ?thesis by simp qed lemma trnr\<^sub>\_eq: assumes "ide u" and "ide v" and "src v = trg g" and "src u = trg f" and "\ \ hom v (u \ f)" shows "trnr\<^sub>\ u \ = (u \ \) \ (\ \ g)" proof - have "trnr\<^sub>\ u \ = (u \ \) \ \[u, f, g] \ (\ \ g)" using assms trnr\<^sub>\_eq [of u v \] by simp also have "... = (u \ \) \ (\ \ g)" proof - have "\[u, f, g] \ (\ \ g) = (\ \ g)" proof - have "ide \[u, f, g]" using assms antipar strict_assoc by simp moreover have "seq \[u, f, g] (\ \ g)" using assms antipar by force ultimately show ?thesis using comp_ide_arr by simp qed thus ?thesis by simp qed finally show ?thesis by simp qed lemma trnl\<^sub>\_eq: assumes "ide u" and "ide v" and "src f = trg v" and "src g = trg u" and "\ \ hom (f \ v) u" shows "trnl\<^sub>\ v \ = (g \ \) \ (\ \ v)" proof - have "trnl\<^sub>\ v \ = (g \ \) \ \[g, f, v] \ (\ \ v)" using assms trnl\<^sub>\_eq [of u v \] by simp also have "... = (g \ \) \ (\ \ v)" proof - have "seq \[g, f, v] (\ \ v)" using assms antipar unit_in_hom by simp thus ?thesis using assms antipar trnl\<^sub>\_eq strict_assoc comp_ide_arr [of "\[g, f, v]" "\ \ v"] by simp qed finally show ?thesis by blast qed lemma trnl\<^sub>\_eq: assumes "ide u" and "ide v" and "src f = trg v" and "src g = trg u" and "\ \ hom v (g \ u)" shows "trnl\<^sub>\ u \ = (\ \ u) \ (f \ \)" proof - have "trnl\<^sub>\ u \ = (\ \ u) \ \\<^sup>-\<^sup>1[f, g, u] \ (f \ \)" using assms trnl\<^sub>\_eq [of u v \] by simp also have "... = ((\ \ u) \ \\<^sup>-\<^sup>1[f, g, u]) \ (f \ \)" using comp_assoc by simp also have "... = (\ \ u) \ (f \ \)" proof - have "seq (\ \ u) \\<^sup>-\<^sup>1[f, g, u]" using assms antipar unit_in_hom by simp thus ?thesis using assms antipar trnl\<^sub>\_eq strict_assoc' comp_arr_ide ide_left ide_right by metis qed finally show ?thesis by simp qed end subsection "Preservation Properties for Adjunctions" text \ Here we show that adjunctions are preserved under isomorphisms of the left and right adjoints. \ context bicategory begin interpretation E: self_evaluation_map V H \ \ src trg .. notation E.eval ("\_\") definition adjoint_pair where "adjoint_pair f g \ \\ \. adjunction_in_bicategory V H \ \ src trg f g \ \" (* These would normally be called "maps", but that name is too heavily used already. *) abbreviation is_left_adjoint where "is_left_adjoint f \ \g. adjoint_pair f g" abbreviation is_right_adjoint where "is_right_adjoint g \ \f. adjoint_pair f g" lemma adjoint_pair_antipar: assumes "adjoint_pair f g" shows "ide f" and "ide g" and "src f = trg g" and "src g = trg f" proof - obtain \ \ where A: "adjunction_in_bicategory V H \ \ src trg f g \ \" using assms adjoint_pair_def by auto interpret A: adjunction_in_bicategory V H \ \ src trg f g \ \ using A by auto show "ide f" by simp show "ide g" by simp show "src f = trg g" using A.antipar by simp show "src g = trg f" using A.antipar by simp qed lemma left_adjoint_is_ide: assumes "is_left_adjoint f" shows "ide f" using assms adjoint_pair_antipar by auto lemma right_adjoint_is_ide: assumes "is_right_adjoint f" shows "ide f" using assms adjoint_pair_antipar by auto lemma adjunction_preserved_by_iso_right: assumes "adjunction_in_bicategory V H \ \ src trg f g \ \" and "\\ : g \ g'\" and "iso \" shows "adjunction_in_bicategory V H \ \ src trg f g' ((\ \ f) \ \) (\ \ (f \ inv \))" proof interpret A: adjunction_in_bicategory V H \ \ src trg f g \ \ using assms by auto show "ide f" by simp show "ide g'" using assms(2) isomorphic_def by auto show "\(\ \ f) \ \ : src f \ g' \ f\" using assms A.antipar by fastforce show "\\ \ (f \ inv \) : f \ g' \ src g'\" proof show "\f \ inv \ : f \ g' \ f \ g\" using assms A.antipar by (intro hcomp_in_vhom, auto) show "\\ : f \ g \ src g'\" using assms A.counit_in_hom A.antipar by auto qed show "(\ \ (f \ inv \) \ f) \ \\<^sup>-\<^sup>1[f, g', f] \ (f \ (\ \ f) \ \) = \\<^sup>-\<^sup>1[f] \ \[f]" proof - have "(\ \ (f \ inv \) \ f) \ \\<^sup>-\<^sup>1[f, g', f] \ (f \ (\ \ f) \ \) = (\ \ f) \ (((f \ inv \) \ f) \ \\<^sup>-\<^sup>1[f, g', f]) \ (f \ \ \ f) \ (f \ \)" using assms A.antipar whisker_right whisker_left comp_assoc by auto also have "... = (\ \ f) \ (\\<^sup>-\<^sup>1[f, g, f] \ (f \ inv \ \ f)) \ (f \ \ \ f) \ (f \ \)" using assms A.antipar assoc'_naturality [of f "inv \" f] by auto also have "... = (\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ ((f \ inv \ \ f) \ (f \ \ \ f)) \ (f \ \)" using comp_assoc by simp also have "... = (\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ g \ f) \ (f \ \)" using assms A.antipar comp_inv_arr inv_is_inverse whisker_left whisker_right [of f "inv \" \] by auto also have "... = (\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \)" using assms A.antipar comp_cod_arr by simp also have "... = \\<^sup>-\<^sup>1[f] \ \[f]" using A.triangle_left by simp finally show ?thesis by simp qed show "(g' \ \ \ (f \ inv \)) \ \[g', f, g'] \ ((\ \ f) \ \ \ g') = \\<^sup>-\<^sup>1[g'] \ \[g']" proof - have "(g' \ \ \ (f \ inv \)) \ \[g', f, g'] \ ((\ \ f) \ \ \ g') = (g' \ \) \ ((g' \ f \ inv \) \ \[g', f, g']) \ ((\ \ f) \ g') \ (\ \ g')" using assms A.antipar whisker_left whisker_right comp_assoc by auto also have "... = (g' \ \) \ (\[g', f, g] \ ((g' \ f) \ inv \)) \ ((\ \ f) \ g') \ (\ \ g')" using assms A.antipar assoc_naturality [of g' f "inv \"] by auto also have "... = (g' \ \) \ \[g', f, g] \ (((g' \ f) \ inv \) \ ((\ \ f) \ g')) \ (\ \ g')" using comp_assoc by simp also have "... = (g' \ \) \ (\[g', f, g] \ ((\ \ f) \ g)) \ ((g \ f) \ inv \) \ (\ \ g')" proof - have "((g' \ f) \ inv \) \ ((\ \ f) \ g') = (\ \ f) \ inv \" using assms A.antipar comp_arr_dom comp_cod_arr interchange [of "g' \ f" "\ \ f" "inv \" g'] by auto also have "... = ((\ \ f) \ g) \ ((g \ f) \ inv \)" using assms A.antipar comp_arr_dom comp_cod_arr interchange [of "\ \ f" "g \ f" g "inv \"] by auto finally show ?thesis using comp_assoc by simp qed also have "... = ((g' \ \) \ (\ \ f \ g)) \ \[g, f, g] \ (\ \ g) \ (trg g \ inv \)" proof - have "\[g', f, g] \ ((\ \ f) \ g) = (\ \ f \ g) \ \[g, f, g]" using assms A.antipar assoc_naturality [of \ f g] by auto moreover have "((g \ f) \ inv \) \ (\ \ g') = (\ \ g) \ (trg g \ inv \)" using assms A.antipar comp_arr_dom comp_cod_arr interchange [of "g \ f" \ "inv \" g'] interchange [of \ "trg g" g "inv \"] by auto ultimately show ?thesis using comp_assoc by simp qed also have "... = ((\ \ src g) \ (g \ \)) \ \[g, f, g] \ (\ \ g) \ (trg g \ inv \)" using assms A.antipar comp_arr_dom comp_cod_arr interchange [of g' \ \ "f \ g"] interchange [of \ g "src g" \] by auto also have "... = (\ \ src g) \ ((g \ \) \ \[g, f, g] \ (\ \ g)) \ (trg g \ inv \)" using comp_assoc by simp also have "... = ((\ \ src g) \ \\<^sup>-\<^sup>1[g]) \ \[g] \ (trg g \ inv \)" using assms A.antipar A.triangle_right comp_cod_arr comp_assoc by simp also have "... = (\\<^sup>-\<^sup>1[g'] \ \) \ inv \ \ \[g']" using assms A.antipar runit'_naturality [of \] lunit_naturality [of "inv \"] by auto also have "... = \\<^sup>-\<^sup>1[g'] \ (\ \ inv \) \ \[g']" using comp_assoc by simp also have "... = \\<^sup>-\<^sup>1[g'] \ \[g']" using assms comp_cod_arr comp_arr_inv' by auto finally show ?thesis by simp qed qed lemma adjunction_preserved_by_iso_left: assumes "adjunction_in_bicategory V H \ \ src trg f g \ \" and "\\ : f \ f'\" and "iso \" shows "adjunction_in_bicategory V H \ \ src trg f' g ((g \ \) \ \) (\ \ (inv \ \ g))" proof interpret A: adjunction_in_bicategory V H \ \ src trg f g \ \ using assms by auto show "ide g" by simp show "ide f'" using assms(2) isomorphic_def by auto show "\(g \ \) \ \ : src f' \ g \ f'\" proof show "\\ : src f' \ g \ f\" using assms A.unit_in_hom by auto show "\g \ \ : g \ f \ g \ f'\" using assms A.antipar by fastforce qed show "\\ \ (inv \ \ g) : f' \ g \ src g\" proof show "\inv \ \ g : f' \ g \ f \ g\" using assms A.antipar by (intro hcomp_in_vhom, auto) show "\\ : f \ g \ src g\" using assms A.antipar by auto qed show "(g \ \ \ (inv \ \ g)) \ \[g, f', g] \ ((g \ \) \ \ \ g) = \\<^sup>-\<^sup>1[g] \ \[g]" proof - have "(g \ \ \ (inv \ \ g)) \ \[g, f', g] \ ((g \ \) \ \ \ g) = (g \ \) \ ((g \ inv \ \ g) \ \[g, f', g]) \ ((g \ \) \ g) \ (\ \ g)" using assms A.antipar whisker_left whisker_right comp_assoc by auto also have "... = (g \ \) \ (\[g, f, g] \ ((g \ inv \) \ g)) \ ((g \ \) \ g) \ (\ \ g)" using assms A.antipar assoc_naturality [of g "inv \" g] by auto also have "... = (g \ \) \ \[g, f, g] \ (((g \ inv \) \ g) \ ((g \ \) \ g)) \ (\ \ g)" using comp_assoc by simp also have "... = (g \ \) \ \[g, f, g] \ ((g \ f) \ g) \ (\ \ g)" using assms A.antipar comp_inv_arr inv_is_inverse whisker_right whisker_left [of g "inv \" \] by auto also have "... = (g \ \) \ \[g, f, g] \ (\ \ g)" using assms A.antipar comp_cod_arr by simp also have "... = \\<^sup>-\<^sup>1[g] \ \[g]" using A.triangle_right by simp finally show ?thesis by simp qed show "(\ \ (inv \ \ g) \ f') \ \\<^sup>-\<^sup>1[f', g, f'] \ (f' \ (g \ \) \ \) = \\<^sup>-\<^sup>1[f'] \ \[f']" proof - have "(\ \ (inv \ \ g) \ f') \ \\<^sup>-\<^sup>1[f', g, f'] \ (f' \ (g \ \) \ \) = (\ \ f') \ (((inv \ \ g) \ f') \ \\<^sup>-\<^sup>1[f', g, f']) \ (f' \ g \ \) \ (f' \ \)" using assms A.antipar whisker_right whisker_left comp_assoc by auto also have "... = (\ \ f') \ (\\<^sup>-\<^sup>1[f, g, f'] \ (inv \ \ g \ f')) \ (f' \ g \ \) \ (f' \ \)" using assms A.antipar assoc'_naturality [of "inv \" g f'] by auto also have "... = (\ \ f') \ \\<^sup>-\<^sup>1[f, g, f'] \ ((inv \ \ g \ f') \ (f' \ g \ \)) \ (f' \ \)" using comp_assoc by simp also have "... = (\ \ f') \ (\\<^sup>-\<^sup>1[f, g, f'] \ (f \ g \ \)) \ (inv \ \ g \ f) \ (f' \ \)" proof - have "(inv \ \ g \ f') \ (f' \ g \ \) = (f \ g \ \) \ (inv \ \ g \ f)" using assms(2-3) A.antipar comp_arr_dom comp_cod_arr interchange [of "inv \" f' "g \ f'" "g \ \"] interchange [of f "inv \" "g \ \" "g \ f"] by auto thus ?thesis using comp_assoc by simp qed also have "... = ((\ \ f') \ ((f \ g) \ \)) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \) \ (inv \ \ src f)" proof - have "\\<^sup>-\<^sup>1[f, g, f'] \ (f \ g \ \) = ((f \ g) \ \) \ \\<^sup>-\<^sup>1[f, g, f]" using assms A.antipar assoc'_naturality [of f g \] by auto moreover have "(inv \ \ g \ f) \ (f' \ \) = (f \ \) \ (inv \ \ src f)" using assms A.antipar comp_arr_dom comp_cod_arr interchange [of "inv \" f' "g \ f" \] interchange [of f "inv \" \ "src f"] by auto ultimately show ?thesis using comp_assoc by simp qed also have "... = ((trg f \ \) \ (\ \ f)) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \) \ (inv \ \ src f)" using assms A.antipar comp_arr_dom comp_cod_arr interchange [of \ "f \ g" f' \] interchange [of "trg f" \ \ f] by (metis A.counit_simps(1) A.counit_simps(2) A.counit_simps(3) in_homE) also have "... = (trg f \ \) \ ((\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \)) \ (inv \ \ src f)" using comp_assoc by simp also have "... = ((trg f \ \) \ \\<^sup>-\<^sup>1[f]) \ \[f] \ (inv \ \ src f)" using assms A.antipar A.triangle_left comp_cod_arr comp_assoc by simp also have "... = (\\<^sup>-\<^sup>1[f'] \ \) \ inv \ \ \[f']" using assms A.antipar lunit'_naturality runit_naturality [of "inv \"] by auto also have "... = \\<^sup>-\<^sup>1[f'] \ (\ \ inv \) \ \[f']" using comp_assoc by simp also have "... = \\<^sup>-\<^sup>1[f'] \ \[f']" using assms comp_cod_arr comp_arr_inv inv_is_inverse by auto finally show ?thesis by simp qed qed lemma adjoint_pair_preserved_by_iso: assumes "adjoint_pair f g" and "\\ : f \ f'\" and "iso \" and "\\ : g \ g'\" and "iso \" shows "adjoint_pair f' g'" proof - obtain \ \ where A: "adjunction_in_bicategory V H \ \ src trg f g \ \" using assms adjoint_pair_def by auto have "adjunction_in_bicategory V H \ \ src trg f g' ((\ \ f) \ \) (\ \ (f \ inv \))" using assms A adjunction_preserved_by_iso_right by blast hence "adjunction_in_bicategory V H \ \ src trg f' g' ((g' \ \) \ (\ \ f) \ \) ((\ \ (f \ inv \)) \ (inv \ \ g'))" using assms adjunction_preserved_by_iso_left by blast thus ?thesis using adjoint_pair_def by auto qed lemma left_adjoint_preserved_by_iso: assumes "is_left_adjoint f" and "\\ : f \ f'\" and "iso \" shows "is_left_adjoint f'" proof - obtain g where g: "adjoint_pair f g" using assms by auto have "adjoint_pair f' g" using assms g adjoint_pair_preserved_by_iso [of f g \ f' g g] adjoint_pair_antipar [of f g] by auto thus ?thesis by auto qed lemma right_adjoint_preserved_by_iso: assumes "is_right_adjoint g" and "\\ : g \ g'\" and "iso \" shows "is_right_adjoint g'" proof - obtain f where f: "adjoint_pair f g" using assms by auto have "adjoint_pair f g'" using assms f adjoint_pair_preserved_by_iso [of f g f f \ g'] adjoint_pair_antipar [of f g] by auto thus ?thesis by auto qed lemma left_adjoint_preserved_by_iso': assumes "is_left_adjoint f" and "f \ f'" shows "is_left_adjoint f'" using assms isomorphic_def left_adjoint_preserved_by_iso by blast lemma right_adjoint_preserved_by_iso': assumes "is_right_adjoint g" and "g \ g'" shows "is_right_adjoint g'" using assms isomorphic_def right_adjoint_preserved_by_iso by blast lemma obj_self_adjunction: assumes "obj a" shows "adjunction_in_bicategory V H \ \ src trg a a \\<^sup>-\<^sup>1[a] \[a]" proof show 1: "ide a" using assms by auto show "\\\<^sup>-\<^sup>1[a] : src a \ a \ a\" using assms 1 by auto show "\\[a] : a \ a \ src a\" using assms 1 by fastforce show "(\[a] \ a) \ \\<^sup>-\<^sup>1[a, a, a] \ (a \ \\<^sup>-\<^sup>1[a]) = \\<^sup>-\<^sup>1[a] \ \[a]" using assms 1 canI_unitor_1 canI_associator_1(2) canI_associator_3 - whisker_can_right_1 whisker_can_left_1 can_Ide_self + whisker_can_right_1 whisker_can_left_1 can_Ide_self obj_simps by simp show "(a \ \[a]) \ \[a, a, a] \ (\\<^sup>-\<^sup>1[a] \ a) = \\<^sup>-\<^sup>1[a] \ \[a]" using assms 1 canI_unitor_1 canI_associator_1(2) canI_associator_3 whisker_can_right_1 whisker_can_left_1 can_Ide_self by simp qed lemma obj_is_self_adjoint: assumes "obj a" shows "adjoint_pair a a" and "is_left_adjoint a" and "is_right_adjoint a" using assms obj_self_adjunction adjoint_pair_def by auto end subsection "Pseudofunctors and Adjunctions" context pseudofunctor begin lemma preserves_adjunction: assumes "adjunction_in_bicategory V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C f g \ \" shows "adjunction_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D (F f) (F g) (D.inv (\ (g, f)) \\<^sub>D F \ \\<^sub>D \ (src\<^sub>C f)) (D.inv (\ (trg\<^sub>C f)) \\<^sub>D F \ \\<^sub>D \ (f, g))" proof - interpret adjunction_in_bicategory V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C f g \ \ using assms by auto interpret A: adjunction_data_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D \F f\ \F g\ \D.inv (\ (g, f)) \\<^sub>D F \ \\<^sub>D \ (src\<^sub>C f)\ \D.inv (\ (trg\<^sub>C f)) \\<^sub>D F \ \\<^sub>D \ (f, g)\ using adjunction_data_in_bicategory_axioms preserves_adjunction_data by auto show "adjunction_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D (F f) (F g) (D.inv (\ (g, f)) \\<^sub>D F \ \\<^sub>D \ (src\<^sub>C f)) (D.inv (\ (trg\<^sub>C f)) \\<^sub>D F \ \\<^sub>D \ (f, g))" proof show "(D.inv (\ (trg\<^sub>C f)) \\<^sub>D F \ \\<^sub>D \ (f, g) \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f, F g, F f] \\<^sub>D (F f \\<^sub>D D.inv (\ (g, f)) \\<^sub>D F \ \\<^sub>D \ (src\<^sub>C f)) = D.lunit' (F f) \\<^sub>D \\<^sub>D[F f]" proof - have 1: "D.iso (\ (f, g \\<^sub>C f) \\<^sub>D (F f \\<^sub>D \ (g, f)))" using antipar C.VV.ide_char C.VV.arr_char D.iso_is_arr FF_def by (intro D.isos_compose D.seqI, simp_all) have "(D.inv (\ (trg\<^sub>C f)) \\<^sub>D F \ \\<^sub>D \ (f, g) \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f, F g, F f] \\<^sub>D (F f \\<^sub>D D.inv (\ (g, f)) \\<^sub>D F \ \\<^sub>D \ (src\<^sub>C f)) = (D.inv (\ (trg\<^sub>C f)) \\<^sub>D F \ \\<^sub>D \ (f, g) \\<^sub>D F f) \\<^sub>D (D.inv (\ (f, g)) \\<^sub>D F f) \\<^sub>D D.inv (\ (f \\<^sub>C g, f)) \\<^sub>D F \\<^sub>C\<^sup>-\<^sup>1[f, g, f] \\<^sub>D \ (f, g \\<^sub>C f) \\<^sub>D (F f \\<^sub>D \ (g, f)) \\<^sub>D (F f \\<^sub>D D.inv (\ (g, f)) \\<^sub>D F \ \\<^sub>D \ (src\<^sub>C f))" proof - have "\\<^sub>D\<^sup>-\<^sup>1[F f, F g, F f] = (D.inv (\ (f, g)) \\<^sub>D F f) \\<^sub>D D.inv (\ (f \\<^sub>C g, f)) \\<^sub>D F \\<^sub>C\<^sup>-\<^sup>1[f, g, f] \\<^sub>D \ (f, g \\<^sub>C f) \\<^sub>D (F f \\<^sub>D \ (g, f))" proof - have "\\<^sub>D\<^sup>-\<^sup>1[F f, F g, F f] \\<^sub>D D.inv (\ (f, g \\<^sub>C f) \\<^sub>D (F f \\<^sub>D \ (g, f))) = D.inv (F \\<^sub>C[f, g, f] \\<^sub>D \ (f \\<^sub>C g, f) \\<^sub>D (\ (f, g) \\<^sub>D F f))" proof - have "D.inv (\ (f, g \\<^sub>C f) \\<^sub>D (F f \\<^sub>D \ (g, f)) \\<^sub>D \\<^sub>D[F f, F g, F f]) = D.inv (F \\<^sub>C[f, g, f] \\<^sub>D \ (f \\<^sub>C g, f) \\<^sub>D (\ (f, g) \\<^sub>D F f))" using antipar assoc_coherence by simp moreover have "D.inv (\ (f, g \\<^sub>C f) \\<^sub>D (F f \\<^sub>D \ (g, f)) \\<^sub>D \\<^sub>D[F f, F g, F f]) = \\<^sub>D\<^sup>-\<^sup>1[F f, F g, F f] \\<^sub>D D.inv (\ (f, g \\<^sub>C f) \\<^sub>D (F f \\<^sub>D \ (g, f)))" proof - have "D.seq (\ (f, g \\<^sub>C f) \\<^sub>D (F f \\<^sub>D \ (g, f))) \\<^sub>D[F f, F g, F f]" using antipar by fastforce thus ?thesis using 1 antipar D.comp_assoc D.inv_comp [of "\\<^sub>D[F f, F g, F f]" "\ (f, g \\<^sub>C f) \\<^sub>D (F f \\<^sub>D \ (g, f))"] by auto qed ultimately show ?thesis by simp qed moreover have 2: "D.iso (F \\<^sub>C[f, g, f] \\<^sub>D \ (f \\<^sub>C g, f) \\<^sub>D (\ (f, g) \\<^sub>D F f))" using antipar D.isos_compose C.VV.ide_char C.VV.arr_char \_simps(4) by simp ultimately have "\\<^sub>D\<^sup>-\<^sup>1[F f, F g, F f] = D.inv (F \\<^sub>C[f, g, f] \\<^sub>D \ (f \\<^sub>C g, f) \\<^sub>D (\ (f, g) \\<^sub>D F f)) \\<^sub>D (\ (f, g \\<^sub>C f) \\<^sub>D (F f \\<^sub>D \ (g, f)))" using 1 2 antipar D.invert_side_of_triangle(2) D.inv_inv D.iso_inv_iso D.arr_inv by metis moreover have "D.inv (F \\<^sub>C[f, g, f] \\<^sub>D \ (f \\<^sub>C g, f) \\<^sub>D (\ (f, g) \\<^sub>D F f)) = (D.inv (\ (f, g)) \\<^sub>D F f) \\<^sub>D D.inv (\ (f \\<^sub>C g, f)) \\<^sub>D F \\<^sub>C\<^sup>-\<^sup>1[f, g, f]" proof - have "D.inv (F \\<^sub>C[f, g, f] \\<^sub>D \ (f \\<^sub>C g, f) \\<^sub>D (\ (f, g) \\<^sub>D F f)) = D.inv (\ (f \\<^sub>C g, f) \\<^sub>D (\ (f, g) \\<^sub>D F f)) \\<^sub>D F \\<^sub>C\<^sup>-\<^sup>1[f, g, f]" using antipar D.isos_compose C.VV.arr_char \_simps(4) preserves_inv D.inv_comp by simp also have "... = (D.inv (\ (f, g) \\<^sub>D F f) \\<^sub>D D.inv (\ (f \\<^sub>C g, f))) \\<^sub>D F \\<^sub>C\<^sup>-\<^sup>1[f, g, f]" using antipar D.inv_comp C.VV.ide_char C.VV.arr_char \_simps(4) by simp also have "... = ((D.inv (\ (f, g)) \\<^sub>D F f) \\<^sub>D D.inv (\ (f \\<^sub>C g, f))) \\<^sub>D F \\<^sub>C\<^sup>-\<^sup>1[f, g, f]" using antipar C.VV.ide_char C.VV.arr_char by simp also have "... = (D.inv (\ (f, g)) \\<^sub>D F f) \\<^sub>D D.inv (\ (f \\<^sub>C g, f)) \\<^sub>D F \\<^sub>C\<^sup>-\<^sup>1[f, g, f]" using D.comp_assoc by simp finally show ?thesis by simp qed ultimately show ?thesis using D.comp_assoc by simp qed thus ?thesis using D.comp_assoc by simp qed also have "... = (D.inv (\ (trg\<^sub>C f)) \\<^sub>D F \ \\<^sub>D F f) \\<^sub>D D.inv (\ (f \\<^sub>C g, f)) \\<^sub>D F \\<^sub>C\<^sup>-\<^sup>1[f, g, f] \\<^sub>D \ (f, g \\<^sub>C f) \\<^sub>D (F f \\<^sub>D F \ \\<^sub>D \ (src\<^sub>C f))" proof - have "... = ((D.inv (\ (trg\<^sub>C f)) \\<^sub>D F \ \\<^sub>D F f) \\<^sub>D (\ (f, g) \\<^sub>D F f)) \\<^sub>D (D.inv (\ (f, g)) \\<^sub>D F f) \\<^sub>D D.inv (\ (f \\<^sub>C g, f)) \\<^sub>D F \\<^sub>C\<^sup>-\<^sup>1[f, g, f] \\<^sub>D \ (f, g \\<^sub>C f) \\<^sub>D (F f \\<^sub>D \ (g, f)) \\<^sub>D ((F f \\<^sub>D D.inv (\ (g, f))) \\<^sub>D (F f \\<^sub>D F \ \\<^sub>D \ (src\<^sub>C f)))" proof - have "D.inv (\ (trg\<^sub>C f)) \\<^sub>D F \ \\<^sub>D \ (f, g) \\<^sub>D F f = (D.inv (\ (trg\<^sub>C f)) \\<^sub>D F \ \\<^sub>D F f) \\<^sub>D (\ (f, g) \\<^sub>D F f)" using ide_left ide_right antipar D.whisker_right \_char(2) by (metis A.counit_simps(1) A.ide_left D.comp_assoc) moreover have "F f \\<^sub>D D.inv (\ (g, f)) \\<^sub>D F \ \\<^sub>D \ (src\<^sub>C f) = (F f \\<^sub>D D.inv (\ (g, f))) \\<^sub>D (F f \\<^sub>D F \ \\<^sub>D \ (src\<^sub>C f))" using antipar \_char(2) D.whisker_left by simp ultimately show ?thesis by simp qed also have "... = (D.inv (\ (trg\<^sub>C f)) \\<^sub>D F \ \\<^sub>D F f) \\<^sub>D (((\ (f, g) \\<^sub>D F f) \\<^sub>D (D.inv (\ (f, g)) \\<^sub>D F f)) \\<^sub>D D.inv (\ (f \\<^sub>C g, f))) \\<^sub>D F \\<^sub>C\<^sup>-\<^sup>1[f, g, f] \\<^sub>D \ (f, g \\<^sub>C f) \\<^sub>D (((F f \\<^sub>D \ (g, f)) \\<^sub>D (F f \\<^sub>D D.inv (\ (g, f)))) \\<^sub>D (F f \\<^sub>D F \ \\<^sub>D \ (src\<^sub>C f)))" using D.comp_assoc by simp also have "... = (D.inv (\ (trg\<^sub>C f)) \\<^sub>D F \ \\<^sub>D F f) \\<^sub>D D.inv (\ (f \\<^sub>C g, f)) \\<^sub>D F \\<^sub>C\<^sup>-\<^sup>1[f, g, f] \\<^sub>D \ (f, g \\<^sub>C f) \\<^sub>D (F f \\<^sub>D F \ \\<^sub>D \ (src\<^sub>C f))" proof - have "((F f \\<^sub>D \ (g, f)) \\<^sub>D (F f \\<^sub>D D.inv (\ (g, f)))) \\<^sub>D (F f \\<^sub>D F \ \\<^sub>D \ (src\<^sub>C f)) = F f \\<^sub>D F \ \\<^sub>D \ (src\<^sub>C f)" proof - have "(F f \\<^sub>D \ (g, f)) \\<^sub>D (F f \\<^sub>D D.inv (\ (g, f))) = F f \\<^sub>D F (g \\<^sub>C f)" using antipar \_char(2) D.comp_arr_inv D.inv_is_inverse D.whisker_left [of "F f" "\ (g, f)" "D.inv (\ (g, f))"] by simp moreover have "D.seq (F f \\<^sub>D F (g \\<^sub>C f)) (F f \\<^sub>D F \ \\<^sub>D \ (src\<^sub>C f))" using antipar by fastforce ultimately show ?thesis using D.comp_cod_arr by auto qed moreover have "((\ (f, g) \\<^sub>D F f) \\<^sub>D (D.inv (\ (f, g)) \\<^sub>D F f)) \\<^sub>D D.inv (\ (f \\<^sub>C g, f)) = D.inv (\ (f \\<^sub>C g, f))" using antipar D.comp_arr_inv D.inv_is_inverse D.comp_cod_arr D.whisker_right [of "F f" "\ (f, g)" "D.inv (\ (f, g))"] by simp ultimately show ?thesis by simp qed finally show ?thesis by simp qed also have "... = (D.inv (\ (trg\<^sub>C f)) \\<^sub>D F f) \\<^sub>D D.inv (\ (trg\<^sub>C f, f)) \\<^sub>D F (\ \\<^sub>C f) \\<^sub>D ((\ (f \\<^sub>C g, f) \\<^sub>D D.inv (\ (f \\<^sub>C g, f))) \\<^sub>D F \\<^sub>C\<^sup>-\<^sup>1[f, g, f]) \\<^sub>D ((\ (f, g \\<^sub>C f) \\<^sub>D D.inv (\ (f, g \\<^sub>C f))) \\<^sub>D F (f \\<^sub>C \)) \\<^sub>D \ (f, src\<^sub>C f) \\<^sub>D (F f \\<^sub>D \ (src\<^sub>C f))" proof - have "(D.inv (\ (trg\<^sub>C f)) \\<^sub>D F \ \\<^sub>D F f) \\<^sub>D D.inv (\ (f \\<^sub>C g, f)) \\<^sub>D F \\<^sub>C\<^sup>-\<^sup>1[f, g, f] \\<^sub>D \ (f, g \\<^sub>C f) \\<^sub>D (F f \\<^sub>D F \ \\<^sub>D \ (src\<^sub>C f)) = ((D.inv (\ (trg\<^sub>C f)) \\<^sub>D F f) \\<^sub>D (F \ \\<^sub>D F f)) \\<^sub>D D.inv (\ (f \\<^sub>C g, f)) \\<^sub>D F \\<^sub>C\<^sup>-\<^sup>1[f, g, f] \\<^sub>D \ (f, g \\<^sub>C f) \\<^sub>D ((F f \\<^sub>D F \) \\<^sub>D (F f \\<^sub>D \ (src\<^sub>C f)))" using antipar D.comp_assoc D.whisker_left D.whisker_right \_char(2) by simp moreover have "F \ \\<^sub>D F f = D.inv (\ (trg\<^sub>C f, f)) \\<^sub>D F (\ \\<^sub>C f) \\<^sub>D \ (f \\<^sub>C g, f)" using antipar \.naturality [of "(\, f)"] C.VV.arr_char FF_def D.invert_side_of_triangle(1) by simp moreover have "F f \\<^sub>D F \ = D.inv (\ (f, g \\<^sub>C f)) \\<^sub>D F (f \\<^sub>C \) \\<^sub>D \ (f, src\<^sub>C f)" using antipar \.naturality [of "(f, \)"] C.VV.arr_char FF_def D.invert_side_of_triangle(1) by simp ultimately show ?thesis using D.comp_assoc by simp qed also have "... = ((D.inv (\ (trg\<^sub>C f)) \\<^sub>D F f) \\<^sub>D D.inv (\ (trg\<^sub>C f, f))) \\<^sub>D (F (\ \\<^sub>C f) \\<^sub>D F \\<^sub>C\<^sup>-\<^sup>1[f, g, f] \\<^sub>D F (f \\<^sub>C \)) \\<^sub>D \ (f, src\<^sub>C f) \\<^sub>D (F f \\<^sub>D \ (src\<^sub>C f))" using antipar D.comp_arr_inv' D.comp_cod_arr D.comp_assoc by simp also have "... = D.inv (\ (trg\<^sub>C f, f) \\<^sub>D (\ (trg\<^sub>C f) \\<^sub>D F f)) \\<^sub>D F ((\ \\<^sub>C f) \\<^sub>C \\<^sub>C\<^sup>-\<^sup>1[f, g, f] \\<^sub>C (f \\<^sub>C \)) \\<^sub>D \ (f, src\<^sub>C f) \\<^sub>D (F f \\<^sub>D \ (src\<^sub>C f))" proof - have "(D.inv (\ (trg\<^sub>C f)) \\<^sub>D F f) \\<^sub>D D.inv (\ (trg\<^sub>C f, f)) = D.inv (\ (trg\<^sub>C f, f) \\<^sub>D (\ (trg\<^sub>C f) \\<^sub>D F f))" proof - have "D.iso (\ (trg\<^sub>C f, f))" using antipar by simp moreover have "D.iso (\ (trg\<^sub>C f) \\<^sub>D F f)" using antipar \_char(2) by simp moreover have "D.seq (\ (trg\<^sub>C f, f)) (\ (trg\<^sub>C f) \\<^sub>D F f)" using antipar D.iso_is_arr calculation(2) apply (intro D.seqI D.hseqI) by auto ultimately show ?thesis using antipar D.inv_comp \_char(2) by simp qed moreover have "F (\ \\<^sub>C f) \\<^sub>D F \\<^sub>C\<^sup>-\<^sup>1[f, g, f] \\<^sub>D F (f \\<^sub>C \) = F ((\ \\<^sub>C f) \\<^sub>C \\<^sub>C\<^sup>-\<^sup>1[f, g, f] \\<^sub>C (f \\<^sub>C \))" using antipar by simp ultimately show ?thesis by simp qed also have "... = (D.lunit' (F f) \\<^sub>D F \\<^sub>C[f]) \\<^sub>D F (C.lunit' f \\<^sub>C \\<^sub>C[f]) \\<^sub>D (D.inv (F \\<^sub>C[f]) \\<^sub>D \\<^sub>D[F f])" proof - have "F ((\ \\<^sub>C f) \\<^sub>C \\<^sub>C\<^sup>-\<^sup>1[f, g, f] \\<^sub>C (f \\<^sub>C \)) = F (C.lunit' f \\<^sub>C \\<^sub>C[f])" using triangle_left by simp moreover have "D.inv (\ (trg\<^sub>C f, f) \\<^sub>D (\ (trg\<^sub>C f) \\<^sub>D F f)) = D.lunit' (F f) \\<^sub>D F \\<^sub>C[f]" proof - have 0: "D.iso (\ (trg\<^sub>C f, f) \\<^sub>D (\ (trg\<^sub>C f) \\<^sub>D F f))" using \_char(2) apply (intro D.isos_compose D.seqI) by auto show ?thesis proof - have 1: "D.iso (F \\<^sub>C[f])" using C.iso_lunit preserves_iso by auto moreover have "D.iso (F \\<^sub>C[f] \\<^sub>D \ (trg\<^sub>C f, f) \\<^sub>D (\ (trg\<^sub>C f) \\<^sub>D F f))" by (metis (no_types) A.ide_left D.iso_lunit ide_left lunit_coherence) moreover have "D.inv (D.inv (F \\<^sub>C[f])) = F \\<^sub>C[f]" using 1 D.inv_inv by blast ultimately show ?thesis by (metis 0 D.inv_comp D.invert_side_of_triangle(2) D.iso_inv_iso D.iso_is_arr ide_left lunit_coherence) qed qed moreover have "\ (f, src\<^sub>C f) \\<^sub>D (F f \\<^sub>D \ (src\<^sub>C f)) = D.inv (F \\<^sub>C[f]) \\<^sub>D \\<^sub>D[F f]" using ide_left runit_coherence preserves_iso C.iso_runit D.invert_side_of_triangle(1) by (metis A.ide_left D.runit_simps(1)) ultimately show ?thesis by simp qed also have "... = D.lunit' (F f) \\<^sub>D ((F \\<^sub>C[f] \\<^sub>D F (C.lunit' f)) \\<^sub>D (F \\<^sub>C[f] \\<^sub>D D.inv (F \\<^sub>C[f]))) \\<^sub>D \\<^sub>D[F f]" using D.comp_assoc by simp also have "... = D.lunit' (F f) \\<^sub>D \\<^sub>D[F f]" using D.comp_cod_arr C.iso_runit C.iso_lunit preserves_iso D.comp_arr_inv' preserves_inv by force finally show ?thesis by blast qed show "(F g \\<^sub>D D.inv (\ (trg\<^sub>C f)) \\<^sub>D F \ \\<^sub>D \ (f, g)) \\<^sub>D \\<^sub>D[F g, F f, F g] \\<^sub>D (D.inv (\ (g, f)) \\<^sub>D F \ \\<^sub>D \ (src\<^sub>C f) \\<^sub>D F g) = D.runit' (F g) \\<^sub>D \\<^sub>D[F g]" proof - have "\\<^sub>D[F g, F f, F g] = D.inv (\ (g, f \\<^sub>C g) \\<^sub>D (F g \\<^sub>D \ (f, g))) \\<^sub>D F \\<^sub>C[g, f, g] \\<^sub>D \ (g \\<^sub>C f, g) \\<^sub>D (\ (g, f) \\<^sub>D F g)" proof - have "D.iso (\ (g, f \\<^sub>C g) \\<^sub>D (F g \\<^sub>D \ (f, g)))" using antipar D.iso_is_arr apply (intro D.isos_compose, auto) by (metis C.iso_assoc D.comp_assoc D.seqE ide_left ide_right preserves_assoc(1) preserves_iso) have "F \\<^sub>C[g, f, g] \\<^sub>D \ (g \\<^sub>C f, g) \\<^sub>D (\ (g, f) \\<^sub>D F g) = \ (g, f \\<^sub>C g) \\<^sub>D (F g \\<^sub>D \ (f, g)) \\<^sub>D \\<^sub>D[F g, F f, F g]" using antipar assoc_coherence by simp moreover have "D.seq (F \\<^sub>C[g, f, g]) (\ (g \\<^sub>C f, g) \\<^sub>D (\ (g, f) \\<^sub>D F g))" proof (intro D.seqI) show 1: "D.hseq (\ (g, f)) (F g)" using antipar C.VV.arr_char by simp show "D.arr (\ (g \\<^sub>C f, g))" using antipar C.VV.arr_char by simp show "D.dom (\ (g \\<^sub>C f, g)) = D.cod (\ (g, f) \\<^sub>D F g)" proof - have "D.iso (\ (g, f) \\<^sub>D F g)" using antipar by simp moreover have "D.iso (\ (g \\<^sub>C f, g))" using antipar by simp ultimately show ?thesis by (metis C.iso_assoc D.comp_assoc D.iso_is_arr D.seqE \F \\<^sub>C[g, f, g] \\<^sub>D \ (g \\<^sub>C f, g) \\<^sub>D (\ (g, f) \\<^sub>D F g) = \ (g, f \\<^sub>C g) \\<^sub>D (F g \\<^sub>D \ (f, g)) \\<^sub>D \\<^sub>D[F g, F f, F g]\ antipar(1) antipar(2) ide_left ide_right preserves_assoc(1) preserves_iso) qed show "D.arr (F \\<^sub>C[g, f, g])" using antipar by simp show "D.dom (F \\<^sub>C[g, f, g]) = D.cod (\ (g \\<^sub>C f, g) \\<^sub>D (\ (g, f) \\<^sub>D F g))" proof - have "D.iso (\ (g, f) \\<^sub>D F g)" using antipar by simp moreover have "D.seq (\ (g \\<^sub>C f, g)) (\ (g, f) \\<^sub>D F g)" using antipar D.iso_is_arr apply (intro D.seqI) by auto ultimately show ?thesis using antipar by simp qed qed ultimately show ?thesis using \D.iso (\ (g, f \\<^sub>C g) \\<^sub>D (F g \\<^sub>D \ (f, g)))\ D.invert_side_of_triangle(1) D.comp_assoc by auto qed hence "(F g \\<^sub>D D.inv (\ (trg\<^sub>C f)) \\<^sub>D F \ \\<^sub>D \ (f, g)) \\<^sub>D \\<^sub>D[F g, F f, F g] \\<^sub>D (D.inv (\ (g, f)) \\<^sub>D F \ \\<^sub>D \ (src\<^sub>C f) \\<^sub>D F g) = (F g \\<^sub>D (D.inv (\ (trg\<^sub>C f)) \\<^sub>D F \) \\<^sub>D \ (f, g)) \\<^sub>D D.inv (\ (g, f \\<^sub>C g) \\<^sub>D (F g \\<^sub>D \ (f, g))) \\<^sub>D F \\<^sub>C[g, f, g] \\<^sub>D \ (g \\<^sub>C f, g) \\<^sub>D (\ (g, f) \\<^sub>D F g) \\<^sub>D (D.inv (\ (g, f)) \\<^sub>D F \ \\<^sub>D \ (src\<^sub>C f) \\<^sub>D F g)" using D.comp_assoc by simp also have "... = ((F g \\<^sub>D D.inv (\ (trg\<^sub>C f)) \\<^sub>D F \) \\<^sub>D (F g \\<^sub>D \ (f, g))) \\<^sub>D D.inv (\ (g, f \\<^sub>C g) \\<^sub>D (F g \\<^sub>D \ (f, g))) \\<^sub>D F \\<^sub>C[g, f, g] \\<^sub>D \ (g \\<^sub>C f, g) \\<^sub>D (\ (g, f) \\<^sub>D F g) \\<^sub>D ((D.inv (\ (g, f)) \\<^sub>D F g) \\<^sub>D (F \ \\<^sub>D \ (src\<^sub>C f) \\<^sub>D F g))" proof - have "F g \\<^sub>D (D.inv (\ (trg\<^sub>C f)) \\<^sub>D F \) \\<^sub>D \ (f, g) = (F g \\<^sub>D D.inv (\ (trg\<^sub>C f)) \\<^sub>D F \) \\<^sub>D (F g \\<^sub>D \ (f, g))" proof - have "D.seq (D.inv (\ (trg\<^sub>C f)) \\<^sub>D F \) (\ (f, g))" using antipar D.comp_assoc by simp thus ?thesis using antipar D.whisker_left by simp qed moreover have "D.inv (\ (g, f)) \\<^sub>D F \ \\<^sub>D \ (src\<^sub>C f) \\<^sub>D F g = (D.inv (\ (g, f)) \\<^sub>D F g) \\<^sub>D (F \ \\<^sub>D \ (src\<^sub>C f) \\<^sub>D F g)" using antipar D.whisker_right by simp ultimately show ?thesis using D.comp_assoc by simp qed also have "... = (F g \\<^sub>D D.inv (\ (trg\<^sub>C f)) \\<^sub>D F \) \\<^sub>D (((F g \\<^sub>D \ (f, g)) \\<^sub>D D.inv (F g \\<^sub>D \ (f, g))) \\<^sub>D D.inv (\ (g, f \\<^sub>C g))) \\<^sub>D F \\<^sub>C[g, f, g] \\<^sub>D \ (g \\<^sub>C f, g) \\<^sub>D ((\ (g, f) \\<^sub>D F g) \\<^sub>D (D.inv (\ (g, f)) \\<^sub>D F g)) \\<^sub>D (F \ \\<^sub>D \ (src\<^sub>C f) \\<^sub>D F g)" proof - have "D.inv (\ (g, f \\<^sub>C g) \\<^sub>D (F g \\<^sub>D \ (f, g))) = D.inv (F g \\<^sub>D \ (f, g)) \\<^sub>D D.inv (\ (g, f \\<^sub>C g))" proof - have "D.iso (\ (g, f \\<^sub>C g))" using antipar by simp moreover have "D.iso (F g \\<^sub>D \ (f, g))" using antipar by simp moreover have "D.seq (\ (g, f \\<^sub>C g)) (F g \\<^sub>D \ (f, g))" using antipar \_in_hom A.ide_right D.iso_is_arr apply (intro D.seqI D.hseqI) by auto ultimately show ?thesis using antipar D.inv_comp by simp qed thus ?thesis using D.comp_assoc by simp qed also have "... = (F g \\<^sub>D D.inv (\ (trg\<^sub>C f)) \\<^sub>D F \) \\<^sub>D D.inv (\ (g, f \\<^sub>C g)) \\<^sub>D F \\<^sub>C[g, f, g] \\<^sub>D \ (g \\<^sub>C f, g) \\<^sub>D (F \ \\<^sub>D \ (src\<^sub>C f) \\<^sub>D F g)" proof - have "((\ (g, f) \\<^sub>D F g) \\<^sub>D (D.inv (\ (g, f)) \\<^sub>D F g)) \\<^sub>D (F \ \\<^sub>D \ (src\<^sub>C f) \\<^sub>D F g) = (F \ \\<^sub>D \ (src\<^sub>C f) \\<^sub>D F g)" proof - have "(\ (g, f) \\<^sub>D F g) \\<^sub>D (D.inv (\ (g, f)) \\<^sub>D F g) = F (g \\<^sub>C f) \\<^sub>D F g" using antipar D.comp_arr_inv' D.whisker_right [of "F g" "\ (g, f)" "D.inv (\ (g, f))"] by simp thus ?thesis using antipar D.comp_cod_arr D.whisker_right by simp qed moreover have "((F g \\<^sub>D \ (f, g)) \\<^sub>D D.inv (F g \\<^sub>D \ (f, g))) \\<^sub>D D.inv (\ (g, f \\<^sub>C g)) = D.inv (\ (g, f \\<^sub>C g))" using antipar D.comp_arr_inv' D.comp_cod_arr D.whisker_left [of "F g" "\ (f, g)" "D.inv (\ (f, g))"] by simp ultimately show ?thesis by simp qed also have "... = (F g \\<^sub>D D.inv (\ (trg\<^sub>C f))) \\<^sub>D ((F g \\<^sub>D F \) \\<^sub>D D.inv (\ (g, f \\<^sub>C g))) \\<^sub>D F \\<^sub>C[g, f, g] \\<^sub>D (\ (g \\<^sub>C f, g) \\<^sub>D (F \ \\<^sub>D F g)) \\<^sub>D (\ (src\<^sub>C f) \\<^sub>D F g)" using antipar D.whisker_left D.whisker_right \_char(2) D.comp_assoc by simp also have "... = (F g \\<^sub>D D.inv (\ (trg\<^sub>C f))) \\<^sub>D D.inv (\ (g, src\<^sub>C g)) \\<^sub>D (F (g \\<^sub>C \) \\<^sub>D F \\<^sub>C[g, f, g] \\<^sub>D F (\ \\<^sub>C g)) \\<^sub>D \ (trg\<^sub>C g, g) \\<^sub>D (\ (src\<^sub>C f) \\<^sub>D F g)" proof - have "(F g \\<^sub>D F \) \\<^sub>D D.inv (\ (g, f \\<^sub>C g)) = D.inv (\ (g, src\<^sub>C g)) \\<^sub>D F (g \\<^sub>C \)" using antipar C.VV.arr_char \.naturality [of "(g, \)"] FF_def D.invert_opposite_sides_of_square by simp moreover have "\ (g \\<^sub>C f, g) \\<^sub>D (F \ \\<^sub>D F g) = F (\ \\<^sub>C g) \\<^sub>D \ (trg\<^sub>C g, g)" using antipar C.VV.arr_char \.naturality [of "(\, g)"] FF_def by simp ultimately show ?thesis using D.comp_assoc by simp qed also have "... = ((F g \\<^sub>D D.inv (\ (trg\<^sub>C f))) \\<^sub>D D.inv (\ (g, src\<^sub>C g)) \\<^sub>D F (C.runit' g)) \\<^sub>D (F \\<^sub>C[g] \\<^sub>D \ (trg\<^sub>C g, g) \\<^sub>D (\ (src\<^sub>C f) \\<^sub>D F g))" proof - have "F (g \\<^sub>C \) \\<^sub>D F \\<^sub>C[g, f, g] \\<^sub>D F (\ \\<^sub>C g) = F (C.runit' g) \\<^sub>D F \\<^sub>C[g]" using ide_left ide_right antipar triangle_right by (metis C.comp_in_homE C.seqI' preserves_comp triangle_in_hom(2)) thus ?thesis using D.comp_assoc by simp qed also have "... = D.runit' (F g) \\<^sub>D \\<^sub>D[F g]" proof - have "D.inv \\<^sub>D[F g] = (F g \\<^sub>D D.inv (\ (trg\<^sub>C f))) \\<^sub>D D.inv (\ (g, src\<^sub>C g)) \\<^sub>D F (C.runit' g)" proof - have "D.runit' (F g) = D.inv (F \\<^sub>C[g] \\<^sub>D \ (g, src\<^sub>C g) \\<^sub>D (F g \\<^sub>D \ (src\<^sub>C g)))" using runit_coherence by simp also have "... = (F g \\<^sub>D D.inv (\ (trg\<^sub>C f))) \\<^sub>D D.inv (\ (g, src\<^sub>C g)) \\<^sub>D F (C.runit' g)" proof - have "D.inv (F \\<^sub>C[g] \\<^sub>D \ (g, src\<^sub>C g) \\<^sub>D (F g \\<^sub>D \ (src\<^sub>C g))) = D.inv (F g \\<^sub>D \ (src\<^sub>C g)) \\<^sub>D D.inv (\ (g, src\<^sub>C g)) \\<^sub>D F (C.runit' g)" proof - have "D.iso (F \\<^sub>C[g])" using preserves_iso by simp moreover have 1: "D.iso (\ (g, src\<^sub>C g) \\<^sub>D (F g \\<^sub>D \ (src\<^sub>C g)))" using preserves_iso \_char(2) D.arrI D.seqE ide_right runit_coherence by (intro D.isos_compose D.seqI, auto) moreover have "D.seq (F \\<^sub>C[g]) (\ (g, src\<^sub>C g) \\<^sub>D (F g \\<^sub>D \ (src\<^sub>C g)))" using ide_right A.ide_right D.runit_simps(1) runit_coherence by metis ultimately have "D.inv (F \\<^sub>C[g] \\<^sub>D \ (g, src\<^sub>C g) \\<^sub>D (F g \\<^sub>D \ (src\<^sub>C g))) = D.inv (\ (g, src\<^sub>C g) \\<^sub>D (F g \\<^sub>D \ (src\<^sub>C g))) \\<^sub>D F (C.runit' g)" using C.iso_runit preserves_inv D.inv_comp by simp moreover have "D.inv (\ (g, src\<^sub>C g) \\<^sub>D (F g \\<^sub>D \ (src\<^sub>C g))) = D.inv (F g \\<^sub>D \ (src\<^sub>C g)) \\<^sub>D D.inv (\ (g, src\<^sub>C g))" proof - have "D.seq (\ (g, src\<^sub>C g)) (F g \\<^sub>D \ (src\<^sub>C g))" using 1 antipar preserves_iso \_char(2) by fast (* * TODO: The fact that auto cannot do this step is probably what is blocking * the whole thing from being done by auto. *) thus ?thesis using 1 antipar preserves_iso \_char(2) D.inv_comp by auto qed ultimately show ?thesis using D.comp_assoc by simp qed thus ?thesis using antipar \_char(2) preserves_iso by simp qed finally show ?thesis by simp qed thus ?thesis using antipar lunit_coherence by simp qed finally show ?thesis by simp qed qed qed lemma preserves_adjoint_pair: assumes "C.adjoint_pair f g" shows "D.adjoint_pair (F f) (F g)" using assms C.adjoint_pair_def D.adjoint_pair_def preserves_adjunction by blast lemma preserves_left_adjoint: assumes "C.is_left_adjoint f" shows "D.is_left_adjoint (F f)" using assms preserves_adjoint_pair by auto lemma preserves_right_adjoint: assumes "C.is_right_adjoint g" shows "D.is_right_adjoint (F g)" using assms preserves_adjoint_pair by auto end context equivalence_pseudofunctor begin lemma reflects_adjunction: assumes "C.ide f" and "C.ide g" and "\\ : src\<^sub>C f \\<^sub>C g \\<^sub>C f\" and "\\ : f \\<^sub>C g \\<^sub>C src\<^sub>C g\" and "adjunction_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D (F f) (F g) (D.inv (\ (g, f)) \\<^sub>D F \ \\<^sub>D \ (src\<^sub>C f)) (D.inv (\ (trg\<^sub>C f)) \\<^sub>D F \ \\<^sub>D \ (f, g))" shows "adjunction_in_bicategory V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C f g \ \" proof - let ?\' = "D.inv (\ (g, f)) \\<^sub>D F \ \\<^sub>D \ (src\<^sub>C f)" let ?\' = "D.inv (\ (trg\<^sub>C f)) \\<^sub>D F \ \\<^sub>D \ (f, g)" interpret A': adjunction_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D \F f\ \F g\ ?\' ?\' using assms(5) by auto interpret A: adjunction_data_in_bicategory V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C f g \ \ using assms(1-4) by (unfold_locales, auto) show ?thesis proof show "(\ \\<^sub>C f) \\<^sub>C \\<^sub>C\<^sup>-\<^sup>1[f, g, f] \\<^sub>C (f \\<^sub>C \) = \\<^sub>C\<^sup>-\<^sup>1[f] \\<^sub>C \\<^sub>C[f]" proof - have 1: "C.par ((\ \\<^sub>C f) \\<^sub>C \\<^sub>C\<^sup>-\<^sup>1[f, g, f] \\<^sub>C (f \\<^sub>C \)) (\\<^sub>C\<^sup>-\<^sup>1[f] \\<^sub>C \\<^sub>C[f])" using assms A.antipar by simp moreover have "F ((\ \\<^sub>C f) \\<^sub>C \\<^sub>C\<^sup>-\<^sup>1[f, g, f] \\<^sub>C (f \\<^sub>C \)) = F (\\<^sub>C\<^sup>-\<^sup>1[f] \\<^sub>C \\<^sub>C[f])" proof - have "F ((\ \\<^sub>C f) \\<^sub>C \\<^sub>C\<^sup>-\<^sup>1[f, g, f] \\<^sub>C (f \\<^sub>C \)) = F (\ \\<^sub>C f) \\<^sub>D F \\<^sub>C\<^sup>-\<^sup>1[f, g, f] \\<^sub>D F (f \\<^sub>C \)" using 1 by auto also have "... = (F (\ \\<^sub>C f) \\<^sub>D \ (f \\<^sub>C g, f)) \\<^sub>D (\ (f, g) \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f, F g, F f] \\<^sub>D (F f \\<^sub>D D.inv (\ (g, f))) \\<^sub>D (D.inv (\ (f, g \\<^sub>C f)) \\<^sub>D F (f \\<^sub>C \))" using assms A.antipar preserves_assoc(2) D.comp_assoc by auto also have "... = \ (trg\<^sub>C f, f) \\<^sub>D ((F \ \\<^sub>D F f) \\<^sub>D (\ (f, g) \\<^sub>D F f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f, F g, F f] \\<^sub>D ((F f \\<^sub>D D.inv (\ (g, f))) \\<^sub>D (F f \\<^sub>D F \)) \\<^sub>D D.inv (\ (f, src\<^sub>C f))" proof - have "F (\ \\<^sub>C f) \\<^sub>D \ (f \\<^sub>C g, f) = \ (trg\<^sub>C f, f) \\<^sub>D (F \ \\<^sub>D F f)" using assms \.naturality [of "(\, f)"] FF_def C.VV.arr_char by simp moreover have "D.inv (\ (f, g \\<^sub>C f)) \\<^sub>D F (f \\<^sub>C \) = (F f \\<^sub>D F \) \\<^sub>D D.inv (\ (f, src\<^sub>C f))" proof - have "F (f \\<^sub>C \) \\<^sub>D \ (f, src\<^sub>C f) = \ (f, g \\<^sub>C f) \\<^sub>D (F f \\<^sub>D F \)" using assms \.naturality [of "(f, \)"] FF_def C.VV.arr_char A.antipar by simp thus ?thesis using assms A.antipar \_components_are_iso C.VV.arr_char \_in_hom FF_def D.invert_opposite_sides_of_square [of "\ (f, g \\<^sub>C f)" "F f \\<^sub>D F \" "F (f \\<^sub>C \)" "\ (f, src\<^sub>C f)"] by fastforce qed ultimately show ?thesis using D.comp_assoc by simp qed also have "... = \ (trg\<^sub>C f, f) \\<^sub>D (F \ \\<^sub>D \ (f, g) \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f, F g, F f] \\<^sub>D (F f \\<^sub>D D.inv (\ (g, f)) \\<^sub>D F \) \\<^sub>D D.inv (\ (f, src\<^sub>C f))" using assms A.antipar \_in_hom A.ide_left A.ide_right A'.ide_left A'.ide_right D.whisker_left [of "F f" "D.inv (\ (g, f))" "F \"] D.whisker_right [of "F f" "F \" "\ (f, g)"] by (metis A'.counit_in_vhom A'.unit_simps(1)D.arrI D.comp_assoc D.src.preserves_reflects_arr D.src_vcomp D.vseq_implies_hpar(1) \_simps(2)) also have "... = \ (trg\<^sub>C f, f) \\<^sub>D (\ (trg\<^sub>C f) \\<^sub>D ?\' \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f, F g, F f] \\<^sub>D (F f \\<^sub>D ?\' \\<^sub>D D.inv (\ (src\<^sub>C f))) \\<^sub>D D.inv (\ (f, src\<^sub>C f))" proof - have "F \ \\<^sub>D \ (f, g) = \ (trg\<^sub>C f) \\<^sub>D ?\'" proof - have "D.iso (\ (trg\<^sub>C f))" using A.ide_left C.ideD(1) \_char(2) by blast thus ?thesis by (metis A'.counit_simps(1) D.comp_assoc D.comp_cod_arr D.inv_is_inverse D.seqE D.comp_arr_inv) qed moreover have "D.inv (\ (g, f)) \\<^sub>D F \ = ?\' \\<^sub>D D.inv (\ (src\<^sub>C f))" using assms(2) \_char D.comp_arr_inv D.inv_is_inverse D.comp_assoc D.comp_cod_arr by (metis A'.unit_simps(1) A.antipar(1) C.ideD(1) C.obj_trg D.invert_side_of_triangle(2)) ultimately show ?thesis by simp qed also have "... = \ (trg\<^sub>C f, f) \\<^sub>D ((\ (trg\<^sub>C f) \\<^sub>D F f) \\<^sub>D (?\' \\<^sub>D F f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f, F g, F f] \\<^sub>D ((F f \\<^sub>D ?\') \\<^sub>D (F f \\<^sub>D D.inv (\ (src\<^sub>C f)))) \\<^sub>D D.inv (\ (f, src\<^sub>C f))" using assms A.antipar A'.antipar \_char D.whisker_left D.whisker_right by simp also have "... = \ (trg\<^sub>C f, f) \\<^sub>D (\ (trg\<^sub>C f) \\<^sub>D F f) \\<^sub>D ((?\' \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f, F g, F f] \\<^sub>D (F f \\<^sub>D ?\')) \\<^sub>D (F f \\<^sub>D D.inv (\ (src\<^sub>C f))) \\<^sub>D D.inv (\ (f, src\<^sub>C f))" using D.comp_assoc by simp also have "... = (\ (trg\<^sub>C f, f) \\<^sub>D (\ (trg\<^sub>C f) \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f]) \\<^sub>D \\<^sub>D[F f] \\<^sub>D (F f \\<^sub>D D.inv (\ (src\<^sub>C f))) \\<^sub>D D.inv (\ (f, src\<^sub>C f))" using A'.triangle_left D.comp_assoc by simp also have "... = F \\<^sub>C\<^sup>-\<^sup>1[f] \\<^sub>D F \\<^sub>C[f]" using assms A.antipar preserves_lunit(2) preserves_runit(1) by simp also have "... = F (\\<^sub>C\<^sup>-\<^sup>1[f] \\<^sub>C \\<^sub>C[f])" using assms by simp finally show ?thesis by simp qed ultimately show ?thesis using is_faithful by blast qed show "(g \\<^sub>C \) \\<^sub>C \\<^sub>C[g, f, g] \\<^sub>C (\ \\<^sub>C g) = \\<^sub>C\<^sup>-\<^sup>1[g] \\<^sub>C \\<^sub>C[g]" proof - have 1: "C.par ((g \\<^sub>C \) \\<^sub>C \\<^sub>C g f g \\<^sub>C (\ \\<^sub>C g)) (\\<^sub>C\<^sup>-\<^sup>1[g] \\<^sub>C \\<^sub>C[g])" using assms A.antipar by auto moreover have "F ((g \\<^sub>C \) \\<^sub>C \\<^sub>C[g, f, g] \\<^sub>C (\ \\<^sub>C g)) = F (\\<^sub>C\<^sup>-\<^sup>1[g] \\<^sub>C \\<^sub>C[g])" proof - have "F ((g \\<^sub>C \) \\<^sub>C \\<^sub>C g f g \\<^sub>C (\ \\<^sub>C g)) = F (g \\<^sub>C \) \\<^sub>D F \\<^sub>C[g, f, g] \\<^sub>D F (\ \\<^sub>C g)" using 1 by auto also have "... = (F (g \\<^sub>C \) \\<^sub>D \ (g, f \\<^sub>C g)) \\<^sub>D (F g \\<^sub>D \ (f, g)) \\<^sub>D \\<^sub>D[F g, F f, F g] \\<^sub>D (D.inv (\ (g, f)) \\<^sub>D F g) \\<^sub>D (D.inv (\ (g \\<^sub>C f, g)) \\<^sub>D F (\ \\<^sub>C g))" using assms A.antipar preserves_assoc(1) [of g f g] D.comp_assoc by auto also have "... = \ (g, src\<^sub>C g) \\<^sub>D ((F g \\<^sub>D F \) \\<^sub>D (F g \\<^sub>D \ (f, g))) \\<^sub>D \\<^sub>D[F g, F f, F g] \\<^sub>D ((D.inv (\ (g, f)) \\<^sub>D F g) \\<^sub>D (F \ \\<^sub>D F g)) \\<^sub>D D.inv (\ (trg\<^sub>C g, g))" proof - have "F (g \\<^sub>C \) \\<^sub>D \ (g, f \\<^sub>C g) = \ (g, src\<^sub>C g) \\<^sub>D (F g \\<^sub>D F \)" using assms \.naturality [of "(g, \)"] FF_def C.VV.arr_char by auto moreover have "D.inv (\ (g \\<^sub>C f, g)) \\<^sub>D F (\ \\<^sub>C g) = (F \ \\<^sub>D F g) \\<^sub>D D.inv (\ (trg\<^sub>C g, g))" proof - have "F (\ \\<^sub>C g) \\<^sub>D \ (trg\<^sub>C g, g) = \ (g \\<^sub>C f, g) \\<^sub>D (F \ \\<^sub>D F g)" using assms \.naturality [of "(\, g)"] FF_def C.VV.arr_char A.antipar by auto thus ?thesis using assms A.antipar \_components_are_iso C.VV.arr_char FF_def D.invert_opposite_sides_of_square [of "\ (g \\<^sub>C f, g)" "F \ \\<^sub>D F g" "F (\ \\<^sub>C g)" "\ (trg\<^sub>C g, g)"] by fastforce qed ultimately show ?thesis using D.comp_assoc by simp qed also have " ... = \ (g, src\<^sub>C g) \\<^sub>D (F g \\<^sub>D F \ \\<^sub>D \ (f, g)) \\<^sub>D \\<^sub>D[F g, F f, F g] \\<^sub>D (D.inv (\ (g, f)) \\<^sub>D F \ \\<^sub>D F g) \\<^sub>D D.inv (\ (trg\<^sub>C g, g))" proof - have "(F g \\<^sub>D F \) \\<^sub>D (F g \\<^sub>D \ (f, g)) = F g \\<^sub>D F \ \\<^sub>D \ (f, g)" using assms A.antipar D.whisker_left by (metis A'.counit_simps(1) A'.ide_right D.seqE) moreover have "(D.inv (\ (g, f)) \\<^sub>D F g) \\<^sub>D (F \ \\<^sub>D F g) = D.inv (\ (g, f)) \\<^sub>D F \ \\<^sub>D F g" using assms A.antipar D.whisker_right by simp ultimately show ?thesis by simp qed also have "... = \ (g, src\<^sub>C g) \\<^sub>D (F g \\<^sub>D \ (trg\<^sub>C f) \\<^sub>D ?\') \\<^sub>D \\<^sub>D[F g, F f, F g] \\<^sub>D (?\' \\<^sub>D D.inv (\ (src\<^sub>C f)) \\<^sub>D F g) \\<^sub>D D.inv (\ (trg\<^sub>C g, g))" proof - have "F \ \\<^sub>D \ (f, g) = \ (trg\<^sub>C f) \\<^sub>D ?\'" using \_char D.comp_arr_inv D.inv_is_inverse D.comp_assoc D.comp_cod_arr by (metis A'.counit_simps(1) C.ideD(1) C.obj_trg D.seqE assms(1)) moreover have "D.inv (\ (g, f)) \\<^sub>D F \ = ?\' \\<^sub>D D.inv (\ (src\<^sub>C f))" using \_char D.comp_arr_inv D.inv_is_inverse D.comp_assoc D.comp_cod_arr by (metis A'.unit_simps(1) A.unit_simps(1) A.unit_simps(5) C.obj_trg D.invert_side_of_triangle(2)) ultimately show ?thesis by simp qed also have "... = \ (g, src\<^sub>C g) \\<^sub>D (F g \\<^sub>D \ (trg\<^sub>C f)) \\<^sub>D ((F g \\<^sub>D ?\') \\<^sub>D \\<^sub>D[F g, F f, F g] \\<^sub>D (?\' \\<^sub>D F g)) \\<^sub>D (D.inv (\ (src\<^sub>C f)) \\<^sub>D F g) \\<^sub>D D.inv (\ (trg\<^sub>C g, g))" using assms A.antipar \_char D.whisker_left D.whisker_right D.comp_assoc by simp also have "... = \ (g, src\<^sub>C g) \\<^sub>D (F g \\<^sub>D \ (trg\<^sub>C f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D[F g] \\<^sub>D (D.inv (\ (src\<^sub>C f)) \\<^sub>D F g) \\<^sub>D D.inv (\ (trg\<^sub>C g, g))" using A'.triangle_right D.comp_assoc by simp also have "... = F \\<^sub>C\<^sup>-\<^sup>1[g] \\<^sub>D F \\<^sub>C[g]" using assms A.antipar preserves_lunit(1) preserves_runit(2) D.comp_assoc by simp also have "... = F (\\<^sub>C\<^sup>-\<^sup>1[g] \\<^sub>C \\<^sub>C[g])" using assms by simp finally show ?thesis by simp qed ultimately show ?thesis using is_faithful by blast qed qed qed lemma reflects_adjoint_pair: assumes "C.ide f" and "C.ide g" and "src\<^sub>C f = trg\<^sub>C g" and "src\<^sub>C g = trg\<^sub>C f" and "D.adjoint_pair (F f) (F g)" shows "C.adjoint_pair f g" proof - obtain \' \' where A': "adjunction_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D (F f) (F g) \' \'" using assms D.adjoint_pair_def by auto interpret A': adjunction_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D \F f\ \F g\ \' \' using A' by auto have 1: "\\ (g, f) \\<^sub>D \' \\<^sub>D D.inv (\ (src\<^sub>C f)) : F (src\<^sub>C f) \\<^sub>D F (g \\<^sub>C f)\" using assms \_char [of "src\<^sub>C f"] A'.unit_in_hom by (intro D.comp_in_homI, auto) have 2: "\\ (trg\<^sub>C f) \\<^sub>D \' \\<^sub>D D.inv (\ (f, g)): F (f \\<^sub>C g) \\<^sub>D F (trg\<^sub>C f)\" using assms \_in_hom [of f g] \_char [of "trg\<^sub>C f"] A'.counit_in_hom by (intro D.comp_in_homI, auto) obtain \ where \: "\\ : src\<^sub>C f \\<^sub>C g \\<^sub>C f\ \ F \ = \ (g, f) \\<^sub>D \' \\<^sub>D D.inv (\ (src\<^sub>C f))" using assms 1 A'.unit_in_hom \_in_hom locally_full by fastforce have \': "\' = D.inv (\ (g, f)) \\<^sub>D F \ \\<^sub>D \ (src\<^sub>C f)" using assms 1 \ \_in_hom \.components_are_iso C.VV.ide_char C.VV.arr_char D.iso_inv_iso \_components_are_iso \_char(2) D.invert_side_of_triangle(1) [of "F \" "\ (g, f)" "\' \\<^sub>D D.inv (\ (src\<^sub>C f))"] D.invert_side_of_triangle(2) [of "D.inv (\ (g, f)) \\<^sub>D F \" \' "D.inv (\ (src\<^sub>C f))"] by (metis (no_types, lifting) C.ideD(1) C.obj_trg D.arrI D.comp_assoc D.inv_inv) obtain \ where \: "\\ : f \\<^sub>C g \\<^sub>C trg\<^sub>C f\ \ F \ = \ (trg\<^sub>C f) \\<^sub>D \' \\<^sub>D D.inv (\ (f, g))" using assms 2 A'.counit_in_hom \_in_hom locally_full by fastforce have \': "\' = D.inv (\ (trg\<^sub>C f)) \\<^sub>D F \ \\<^sub>D \ (f, g)" using assms 2 \ \_in_hom \.components_are_iso C.VV.ide_char C.VV.arr_char D.iso_inv_iso \_char(2) D.comp_assoc D.invert_side_of_triangle(1) [of "F \" "\ (trg\<^sub>C f)" "\' \\<^sub>D D.inv (\ (f, g))"] D.invert_side_of_triangle(2) [of "D.inv (\ (trg\<^sub>C f)) \\<^sub>D F \" \' "D.inv (\ (f, g))"] by (metis (no_types, lifting) C.arrI C.ideD(1) C.obj_trg D.inv_inv \_components_are_iso preserves_arr) have "adjunction_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D (F f) (F g) (D.inv (\ (g, f)) \\<^sub>D F \ \\<^sub>D \ (src\<^sub>C f)) (D.inv (\ (trg\<^sub>C f)) \\<^sub>D F \ \\<^sub>D \ (f, g))" using A'.adjunction_in_bicategory_axioms \' \' by simp hence "adjunction_in_bicategory V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C f g \ \" using assms \ \ reflects_adjunction by simp thus ?thesis using C.adjoint_pair_def by auto qed lemma reflects_left_adjoint: assumes "C.ide f" and "D.is_left_adjoint (F f)" shows "C.is_left_adjoint f" proof - obtain g' where g': "D.adjoint_pair (F f) g'" using assms D.adjoint_pair_def by auto obtain g where g: "\g : trg\<^sub>C f \\<^sub>C src\<^sub>C f\ \ C.ide g \ D.isomorphic (F g) g'" using assms g' locally_essentially_surjective [of "trg\<^sub>C f" "src\<^sub>C f" g'] D.adjoint_pair_antipar [of "F f" g'] by auto obtain \ where \: "\\ : g' \\<^sub>D F g\ \ D.iso \" using g D.isomorphic_def D.isomorphic_symmetric by metis have "D.adjoint_pair (F f) (F g)" using assms g g' \ D.adjoint_pair_preserved_by_iso [of "F f" g' "F f" "F f" \ "F g"] by auto thus ?thesis using assms g reflects_adjoint_pair [of f g] D.adjoint_pair_antipar C.in_hhom_def by auto qed lemma reflects_right_adjoint: assumes "C.ide g" and "D.is_right_adjoint (F g)" shows "C.is_right_adjoint g" proof - obtain f' where f': "D.adjoint_pair f' (F g)" using assms D.adjoint_pair_def by auto obtain f where f: "\f : trg\<^sub>C g \\<^sub>C src\<^sub>C g\ \ C.ide f \ D.isomorphic (F f) f'" using assms f' locally_essentially_surjective [of "trg\<^sub>C g" "src\<^sub>C g" f'] D.adjoint_pair_antipar [of f' "F g"] by auto obtain \ where \: "\\ : f' \\<^sub>D F f\ \ D.iso \" using f D.isomorphic_def D.isomorphic_symmetric by metis have "D.adjoint_pair (F f) (F g)" using assms f f' \ D.adjoint_pair_preserved_by_iso [of f' "F g" \ "F f" "F g" "F g"] by auto thus ?thesis using assms f reflects_adjoint_pair [of f g] D.adjoint_pair_antipar C.in_hhom_def by auto qed end subsection "Composition of Adjunctions" text \ We first consider the strict case, then extend to all bicategories using strictification. \ locale composite_adjunction_in_strict_bicategory = strict_bicategory V H \ \ src trg + fg: adjunction_in_strict_bicategory V H \ \ src trg f g \ \ + hk: adjunction_in_strict_bicategory V H \ \ src trg h k \ \ for V :: "'a \ 'a \ 'a" (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 \ :: "'a" and \ :: "'a" and h :: "'a" and k :: "'a" and \ :: "'a" and \ :: "'a" + assumes composable: "src h = trg f" begin abbreviation \ where "\ \ (g \ \ \ f) \ \" abbreviation \ where "\ \ \ \ (h \ \ \ k)" interpretation adjunction_data_in_bicategory V H \ \ src trg \h \ f\ \g \ k\ \ \ proof show "ide (h \ f)" using composable by simp show "ide (g \ k)" using fg.antipar hk.antipar composable by simp show "\\ : src (h \ f) \ (g \ k) \ h \ f\" proof show "\\ : src (h \ f) \ g \ f\" using fg.antipar hk.antipar composable \ide (h \ f)\ by auto show "\g \ \ \ f : g \ f \ (g \ k) \ h \ f\" proof - have "\g \ \ \ f : g \ trg f \ f \ g \ (k \ h) \ f\" using fg.antipar hk.antipar composable hk.unit_in_hom apply (intro hcomp_in_vhom) by auto thus ?thesis using hcomp_obj_arr hcomp_assoc by fastforce qed qed show "\\ : (h \ f) \ g \ k \ src (g \ k)\" proof show "\h \ \ \ k : (h \ f) \ g \ k \ h \ k\" proof - have "\h \ \ \ k : h \ (f \ g) \ k \ h \ trg f \ k\" using composable fg.antipar(1-2) hk.antipar(1) by fastforce thus ?thesis using fg.antipar hk.antipar composable hk.unit_in_hom hcomp_obj_arr hcomp_assoc by simp qed show "\\ : h \ k \ src (g \ k)\" using fg.antipar hk.antipar composable hk.unit_in_hom by auto qed qed sublocale adjunction_in_strict_bicategory V H \ \ src trg \h \ f\ \g \ k\ \ \ proof show "(\ \ h \ f) \ \\<^sup>-\<^sup>1[h \ f, g \ k, h \ f] \ ((h \ f) \ \) = \\<^sup>-\<^sup>1[h \ f] \ \[h \ f]" proof - have "(\ \ h \ f) \ \\<^sup>-\<^sup>1[h \ f, g \ k, h \ f] \ ((h \ f) \ \) = (\ \ (h \ \ \ k) \ h \ f) \ ((h \ f) \ (g \ \ \ f) \ \)" using fg.antipar hk.antipar composable strict_assoc comp_ide_arr ide_left ide_right antipar(1) antipar(2) by (metis arrI seqE strict_assoc' triangle_in_hom(1)) also have "... = (\ \ h \ f) \ ((h \ \ \ (k \ h) \ f) \ (h \ (f \ g) \ \ \ f)) \ (h \ f \ \)" using fg.antipar hk.antipar composable whisker_left [of "h \ f"] whisker_right comp_assoc hcomp_assoc by simp also have "... = (\ \ h \ f) \ (h \ (\ \ (k \ h)) \ ((f \ g) \ \) \ f) \ (h \ f \ \)" using fg.antipar hk.antipar composable whisker_left whisker_right hcomp_assoc by simp also have "... = (\ \ h \ f) \ (h \ (trg f \ \) \ (\ \ trg f) \ f) \ (h \ f \ \)" using fg.antipar hk.antipar composable comp_arr_dom comp_cod_arr interchange [of \ "f \ g" "k \ h" \] interchange [of "trg f" \ \ "trg f"] by (metis fg.counit_simps(1) fg.counit_simps(2) fg.counit_simps(3) hk.unit_simps(1) hk.unit_simps(2) hk.unit_simps(3)) also have "... = (\ \ h \ f) \ (h \ \ \ \ \ f) \ (h \ f \ \)" using fg.antipar hk.antipar composable hcomp_obj_arr hcomp_arr_obj by (metis fg.counit_simps(1) fg.counit_simps(4) hk.unit_simps(1) hk.unit_simps(5) obj_src) also have "... = ((\ \ h \ f) \ (h \ \ \ f)) \ ((h \ \ \ f) \ (h \ f \ \))" using fg.antipar hk.antipar composable whisker_left whisker_right comp_assoc by simp also have "... = ((\ \ h) \ (h \ \) \ f) \ (h \ (\ \ f) \ (f \ \))" using fg.antipar hk.antipar composable whisker_left whisker_right hcomp_assoc by simp also have "... = h \ f" using fg.antipar hk.antipar composable fg.triangle_left hk.triangle_left by simp also have "... = \\<^sup>-\<^sup>1[h \ f] \ \[h \ f]" using fg.antipar hk.antipar composable strict_lunit' strict_runit by simp finally show ?thesis by simp qed show "((g \ k) \ \) \ \[g \ k, h \ f, g \ k] \ (\ \ g \ k) = \\<^sup>-\<^sup>1[g \ k] \ \[g \ k]" proof - have "((g \ k) \ \) \ \[g \ k, h \ f, g \ k] \ (\ \ g \ k) = ((g \ k) \ \ \ (h \ \ \ k)) \ ((g \ \ \ f) \ \ \ g \ k)" using fg.antipar hk.antipar composable strict_assoc comp_ide_arr ide_left ide_right by (metis antipar(1) antipar(2) arrI seqE triangle_in_hom(2)) also have "... = (g \ k \ \) \ ((g \ (k \ h) \ \ \ k) \ (g \ \ \ (f \ g) \ k)) \ (\ \ g \ k)" using fg.antipar hk.antipar composable whisker_left [of "g \ k"] whisker_right comp_assoc hcomp_assoc by simp also have "... = (g \ k \ \) \ (g \ ((k \ h) \ \) \ (\ \ f \ g) \ k) \ (\ \ g \ k)" using fg.antipar hk.antipar composable whisker_left whisker_right hcomp_assoc by simp also have "... = (g \ k \ \) \ (g \ (\ \ src g) \ (src g \ \) \ k) \ (\ \ g \ k)" using fg.antipar hk.antipar composable interchange [of "k \ h" \ \ "f \ g"] interchange [of \ "src g" "src g" \] comp_arr_dom comp_cod_arr by (metis fg.counit_simps(1) fg.counit_simps(2) fg.counit_simps(3) hk.unit_simps(1) hk.unit_simps(2) hk.unit_simps(3)) also have "... = (g \ k \ \) \ (g \ \ \ \ \ k) \ (\ \ g \ k)" using fg.antipar hk.antipar composable hcomp_obj_arr [of "src g" \] hcomp_arr_obj [of \ "src g"] by simp also have "... = ((g \ k \ \) \ (g \ \ \ k)) \ (g \ \ \ k) \ (\ \ g \ k)" using fg.antipar hk.antipar composable whisker_left whisker_right comp_assoc by simp also have "... = (g \ (k \ \) \ (\ \ k)) \ ((g \ \) \ (\ \ g) \ k)" using fg.antipar hk.antipar composable whisker_left whisker_right hcomp_assoc by simp also have "... = g \ k" using fg.antipar hk.antipar composable fg.triangle_right hk.triangle_right by simp also have "... = \\<^sup>-\<^sup>1[g \ k] \ \[g \ k]" using fg.antipar hk.antipar composable strict_lunit strict_runit' by simp finally show ?thesis by simp qed qed lemma is_adjunction_in_strict_bicategory: shows "adjunction_in_strict_bicategory V H \ \ src trg (h \ f) (g \ k) \ \" .. end context strict_bicategory begin lemma left_adjoints_compose: assumes "is_left_adjoint f" and "is_left_adjoint f'" and "src f' = trg f" shows "is_left_adjoint (f' \ f)" proof - obtain g \ \ where fg: "adjunction_in_bicategory V H \ \ src trg f g \ \" using assms adjoint_pair_def by auto interpret fg: adjunction_in_bicategory V H \ \ src trg f g \ \ using fg by auto obtain g' \' \' where f'g': "adjunction_in_bicategory V H \ \ src trg f' g' \' \'" using assms adjoint_pair_def by auto interpret f'g': adjunction_in_bicategory V H \ \ src trg f' g' \' \' using f'g' by auto interpret f'fgg': composite_adjunction_in_strict_bicategory V H \ \ src trg f g \ \ f' g' \' \' using assms apply unfold_locales by simp have "adjoint_pair (f' \ f) (g \ g')" using adjoint_pair_def f'fgg'.adjunction_in_bicategory_axioms by auto thus ?thesis by auto qed lemma right_adjoints_compose: assumes "is_right_adjoint g" and "is_right_adjoint g'" and "src g = trg g'" shows "is_right_adjoint (g \ g')" proof - obtain f \ \ where fg: "adjunction_in_bicategory V H \ \ src trg f g \ \" using assms adjoint_pair_def by auto interpret fg: adjunction_in_bicategory V H \ \ src trg f g \ \ using fg by auto obtain f' \' \' where f'g': "adjunction_in_bicategory V H \ \ src trg f' g' \' \'" using assms adjoint_pair_def by auto interpret f'g': adjunction_in_bicategory V H \ \ src trg f' g' \' \' using f'g' by auto interpret f'fgg': composite_adjunction_in_strict_bicategory V H \ \ src trg f g \ \ f' g' \' \' using assms fg.antipar f'g'.antipar apply unfold_locales by simp have "adjoint_pair (f' \ f) (g \ g')" using adjoint_pair_def f'fgg'.adjunction_in_bicategory_axioms by auto thus ?thesis by auto qed end text \ We now use strictification to extend the preceding results to an arbitrary bicategory. We only prove that ``left adjoints compose'' and ``right adjoints compose''; I did not work out formulas for the unit and counit of the composite adjunction in the non-strict case. \ context bicategory begin interpretation S: strictified_bicategory V H \ \ src trg .. notation S.vcomp (infixr "\\<^sub>S" 55) notation S.hcomp (infixr "\\<^sub>S" 53) notation S.in_hom ("\_ : _ \\<^sub>S _\") notation S.in_hhom ("\_ : _ \\<^sub>S _\") interpretation UP: fully_faithful_functor V S.vcomp S.UP using S.UP_is_fully_faithful_functor by auto interpretation UP: equivalence_pseudofunctor V H \ \ src trg S.vcomp S.hcomp S.\ S.\ S.src S.trg S.UP S.\ using S.UP_is_equivalence_pseudofunctor by auto lemma left_adjoints_compose: assumes "is_left_adjoint f" and "is_left_adjoint f'" and "src f = trg f'" shows "is_left_adjoint (f \ f')" proof - have "S.is_left_adjoint (S.UP f) \ S.is_left_adjoint (S.UP f')" using assms UP.preserves_left_adjoint by simp moreover have "S.src (S.UP f) = S.trg (S.UP f')" using assms left_adjoint_is_ide by simp ultimately have "S.is_left_adjoint (S.hcomp (S.UP f) (S.UP f'))" using S.left_adjoints_compose by simp moreover have "S.isomorphic (S.hcomp (S.UP f) (S.UP f')) (S.UP (f \ f'))" proof - have "\S.\ (f, f') : S.hcomp (S.UP f) (S.UP f') \\<^sub>S S.UP (f \ f')\" using assms left_adjoint_is_ide UP.\_in_hom by simp moreover have "S.iso (S.\ (f, f'))" using assms left_adjoint_is_ide by simp ultimately show ?thesis using S.isomorphic_def by blast qed ultimately have "S.is_left_adjoint (S.UP (f \ f'))" using S.left_adjoint_preserved_by_iso S.isomorphic_def by blast thus "is_left_adjoint (f \ f')" using assms left_adjoint_is_ide UP.reflects_left_adjoint by simp qed lemma right_adjoints_compose: assumes "is_right_adjoint g" and "is_right_adjoint g'" and "src g' = trg g" shows "is_right_adjoint (g' \ g)" proof - have "S.is_right_adjoint (S.UP g) \ S.is_right_adjoint (S.UP g')" using assms UP.preserves_right_adjoint by simp moreover have "S.src (S.UP g') = S.trg (S.UP g)" using assms right_adjoint_is_ide by simp ultimately have "S.is_right_adjoint (S.hcomp (S.UP g') (S.UP g))" using S.right_adjoints_compose by simp moreover have "S.isomorphic (S.hcomp (S.UP g') (S.UP g)) (S.UP (g' \ g))" proof - have "\S.\ (g', g) : S.hcomp (S.UP g') (S.UP g) \\<^sub>S S.UP (g' \ g)\" using assms right_adjoint_is_ide UP.\_in_hom by simp moreover have "S.iso (S.\ (g', g))" using assms right_adjoint_is_ide by simp ultimately show ?thesis using S.isomorphic_def by blast qed ultimately have "S.is_right_adjoint (S.UP (g' \ g))" using S.right_adjoint_preserved_by_iso S.isomorphic_def by blast thus "is_right_adjoint (g' \ g)" using assms right_adjoint_is_ide UP.reflects_right_adjoint by simp qed end subsection "Choosing Right Adjoints" text \ It will be useful in various situations to suppose that we have made a choice of right adjoint for each left adjoint ({\it i.e.} each ``map'') in a bicategory. \ locale chosen_right_adjoints = bicategory begin (* Global notation is evil! *) no_notation Transitive_Closure.rtrancl ("(_\<^sup>*)" [1000] 999) definition some_right_adjoint ("_\<^sup>*" [1000] 1000) where "f\<^sup>* \ SOME g. adjoint_pair f g" definition some_unit where "some_unit f \ SOME \. \\. adjunction_in_bicategory V H \ \ src trg f f\<^sup>* \ \" definition some_counit where "some_counit f \ SOME \. adjunction_in_bicategory V H \ \ src trg f f\<^sup>* (some_unit f) \" lemma left_adjoint_extends_to_adjunction: assumes "is_left_adjoint f" shows "adjunction_in_bicategory V H \ \ src trg f f\<^sup>* (some_unit f) (some_counit f)" using assms some_right_adjoint_def adjoint_pair_def some_unit_def some_counit_def someI_ex [of "\g. adjoint_pair f g"] someI_ex [of "\\. \\. adjunction_in_bicategory V H \ \ src trg f f\<^sup>* \ \"] someI_ex [of "\\. adjunction_in_bicategory V H \ \ src trg f f\<^sup>* (some_unit f) \"] by auto lemma left_adjoint_extends_to_adjoint_pair: assumes "is_left_adjoint f" shows "adjoint_pair f f\<^sup>*" using assms adjoint_pair_def left_adjoint_extends_to_adjunction by blast lemma right_adjoint_in_hom [intro]: assumes "is_left_adjoint f" shows "\f\<^sup>* : trg f \ src f\" and "\f\<^sup>* : f\<^sup>* \ f\<^sup>*\" using assms left_adjoint_extends_to_adjoint_pair adjoint_pair_antipar [of f "f\<^sup>*"] by auto lemma right_adjoint_simps [simp]: assumes "is_left_adjoint f" shows "ide f\<^sup>*" and "src f\<^sup>* = trg f" and "trg f\<^sup>* = src f" and "dom f\<^sup>* = f\<^sup>*" and "cod f\<^sup>* = f\<^sup>*" using assms right_adjoint_in_hom left_adjoint_extends_to_adjoint_pair apply auto using assms right_adjoint_is_ide [of "f\<^sup>*"] by blast end locale map_in_bicategory = bicategory + chosen_right_adjoints + fixes f :: 'a assumes is_map: "is_left_adjoint f" begin abbreviation \ where "\ \ some_unit f" abbreviation \ where "\ \ some_counit f" sublocale adjunction_in_bicategory V H \ \ src trg f \f\<^sup>*\ \ \ using is_map left_adjoint_extends_to_adjunction by simp end subsection "Equivalences Refine to Adjoint Equivalences" text \ In this section, we show that, just as an equivalence between categories can always be refined to an adjoint equivalence, an internal equivalence in a bicategory can also always be so refined. The proof, which follows that of Theorem 3.3 from \cite{nlab-adjoint-equivalence}, makes use of the fact that if an internal equivalence satisfies one of the triangle identities, then it also satisfies the other. \ locale adjoint_equivalence_in_bicategory = equivalence_in_bicategory + adjunction_in_bicategory begin lemma dual_adjoint_equivalence: shows "adjoint_equivalence_in_bicategory V H \ \ src trg g f (inv \) (inv \)" proof - interpret gf: equivalence_in_bicategory V H \ \ src trg g f \inv \\ \inv \\ using dual_equivalence by simp show ?thesis proof show "(inv \ \ g) \ \\<^sup>-\<^sup>1[g, f, g] \ (g \ inv \) = \\<^sup>-\<^sup>1[g] \ \[g]" proof - have "(inv \ \ g) \ \\<^sup>-\<^sup>1[g, f, g] \ (g \ inv \) = inv ((g \ \) \ \[g, f, g] \ (\ \ g))" using antipar inv_comp iso_inv_iso isos_compose comp_assoc by simp also have "... = inv (\\<^sup>-\<^sup>1[g] \ \[g])" using triangle_right by simp also have "... = \\<^sup>-\<^sup>1[g] \ \[g]" using iso_inv_iso inv_comp by simp finally show ?thesis by blast qed show "(f \ inv \) \ \[f, g, f] \ (inv \ \ f) = \\<^sup>-\<^sup>1[f] \ \[f]" proof - have "(f \ inv \) \ \[f, g, f] \ (inv \ \ f) = inv ((\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \))" using antipar inv_comp iso_inv_iso isos_compose comp_assoc by simp also have "... = inv (\\<^sup>-\<^sup>1[f] \ \[f])" using triangle_left by simp also have "... = \\<^sup>-\<^sup>1[f] \ \[f]" using iso_inv_iso inv_comp by simp finally show ?thesis by blast qed qed qed end context strict_bicategory begin lemma equivalence_refines_to_adjoint_equivalence: assumes "equivalence_map f" and "\g : trg f \ src f\" and "ide g" and "\\ : src f \ g \ f\" and "iso \" shows "\!\. adjoint_equivalence_in_bicategory V H \ \ src trg f g \ \" proof - obtain g' \' \' where E': "equivalence_in_bicategory V H \ \ src trg f g' \' \'" using assms equivalence_map_def by auto interpret E': equivalence_in_bicategory V H \ \ src trg f g' \' \' using E' by auto let ?a = "src f" and ?b = "trg f" (* TODO: in_homE cannot be applied automatically to a conjunction. Must break down! *) have f_in_hhom: "\f : ?a \ ?b\" and ide_f: "ide f" using assms equivalence_map_def by auto have g_in_hhom: "\g : ?b \ ?a\" and ide_g: "ide g" using assms by auto have g'_in_hhom: "\g' : ?b \ ?a\" and ide_g': "ide g'" using assms f_in_hhom E'.antipar by auto have \_in_hom: "\\ : ?a \ g \ f\" and iso_\: "iso \" using assms by auto have a: "obj ?a" and b: "obj ?b" using f_in_hhom by auto have \_in_hhom: "\\ : ?a \ ?a\" using a src_dom trg_dom \_in_hom by fastforce text \ The following is quoted from \cite{nlab-adjoint-equivalence}: \begin{quotation} ``Since \g \ gfg' \ g'\, the isomorphism \fg' \ 1\ also induces an isomorphism \fg \ 1\, which we denote \\\. Now \\\ and \\\ may not satisfy the zigzag identities, but if we define \\\ by \\ \ (f \ \\<^sup>-\<^sup>1) \ (f \ g \ \\<^sup>-\<^sup>1) : f \ g \ 1\, then we can verify, using string diagram notation as above, that \\\ satisfies one zigzag identity, and hence (by the previous lemma) also the other. Finally, if \\': fg \ 1\ is any other isomorphism satisfying the zigzag identities with \\\, then we have: \[ \\' = \' \ (\ f g) \ (f \ g) = \ \ (f g \') \ (f \ g) = \\ \] using the interchange law and two zigzag identities. This shows uniqueness.'' \end{quotation} \ have 1: "isomorphic g g'" proof - have "isomorphic g (g \ ?b)" using assms hcomp_arr_obj isomorphic_reflexive by auto also have "isomorphic ... (g \ f \ g')" using assms f_in_hhom g_in_hhom g'_in_hhom E'.counit_in_vhom E'.counit_is_iso isomorphic_def hcomp_ide_isomorphic isomorphic_symmetric by (metis E'.counit_simps(5) in_hhomE trg_trg) also have "isomorphic ... (?a \ g')" using assms f_in_hhom g_in_hhom g'_in_hhom ide_g' E'.unit_in_vhom E'.unit_is_iso isomorphic_def hcomp_isomorphic_ide isomorphic_symmetric by (metis hcomp_assoc hcomp_isomorphic_ide in_hhomE src_src) also have "isomorphic ... g'" using assms by (simp add: E'.antipar(1) hcomp_obj_arr isomorphic_reflexive) finally show ?thesis by blast qed moreover have "isomorphic (f \ g') ?b" using E'.counit_is_iso isomorphicI [of \'] by auto hence 2: "isomorphic (f \ g) ?b" using assms 1 ide_f hcomp_ide_isomorphic [of f g g'] isomorphic_transitive isomorphic_symmetric by (metis in_hhomE) obtain \ where \: "\\ : f \ g \ ?b\ \ iso \" using 2 by auto have \_in_hom: "\\ : f \ g \ ?b\" and iso_\: "iso \" using \ by auto have \_in_hhom: "\\ : ?b \ ?b\" using b src_cod trg_cod \_in_hom by fastforce text \ At the time of this writing, the definition of \\\ given on nLab \cite{nlab-adjoint-equivalence} had an apparent typo: the expression \f \ g \ \\<^sup>-\<^sup>1\ should read \\\<^sup>-\<^sup>1 \ f \ g\, as we have used here. \ let ?\ = "\ \ (f \ inv \ \ g) \ (inv \ \ f \ g)" have \_in_hom: "\?\ : f \ g \ ?b\" proof - have "\f \ inv \ \ g : f \ g \ f \ g \ f \ g\" proof - have "\inv \ : g \ f \ ?a\" using \_in_hom iso_\ by auto hence "\f \ inv \ \ g : f \ (g \ f) \ g \ f \ ?a \ g\" using assms by (intro hcomp_in_vhom, auto) hence "\f \ inv \ \ g : f \ (g \ f) \ g \ f \ g\" using assms f_in_hhom hcomp_obj_arr by (metis in_hhomE) moreover have "f \ (g \ f) \ g = f \ g \ f \ g" using hcomp_assoc by simp ultimately show ?thesis by simp qed moreover have "\inv \ \ f \ g : f \ g \ f \ g \ f \ g\" proof - have "\inv \ \ f \ g : ?b \ f \ g \ (f \ g) \ f \ g\" using assms \_in_hom iso_\ by (intro hcomp_in_vhom, auto) moreover have "(f \ g) \ f \ g = f \ g \ f \ g" using hcomp_assoc by simp moreover have "?b \ f \ g = f \ g" using f_in_hhom g_in_hhom b hcomp_obj_arr [of ?b "f \ g"] by fastforce ultimately show ?thesis by simp qed ultimately show "\?\ : f \ g \ ?b\" using \_in_hom by blast qed have "iso ?\" using f_in_hhom g_in_hhom \_in_hhom ide_f ide_g \_in_hom iso_\ \_in_hhom \_in_hom iso_\ iso_inv_iso isos_compose by (metis \_in_hom arrI hseqE ide_is_iso iso_hcomp seqE) have 4: "\inv \ \ f : ?b \ f \ f \ g \ f\" proof - have "\inv \ \ f : ?b \ f \ (f \ g) \ f\" using \_in_hom iso_\ f_in_hhom by (intro hcomp_in_vhom, auto) thus ?thesis using hcomp_assoc by simp qed text \ First show \?\\ and \\\ satisfy the ``left'' triangle identity. \ have triangle_left: "(?\ \ f) \ (f \ \) = f" proof - have "(?\ \ f) \ (f \ \) = (\ \ f) \ (f \ inv \ \ g \ f) \ (inv \ \ f \ g \ f) \ (?b \ f \ \)" proof - have "f \ \ = ?b \ f \ \" using b \_in_hhom hcomp_obj_arr [of ?b "f \ \"] by fastforce moreover have "\ \ (f \ inv \ \ g) \ (inv \ \ f \ g) \ f = (\ \ f) \ ((f \ inv \ \ g) \ f) \ ((inv \ \ f \ g) \ f)" using ide_f ide_g \_in_hhom \_in_hom iso_\ \_in_hhom \_in_hom iso_\ whisker_right by (metis \_in_hom arrI in_hhomE seqE) moreover have "... = (\ \ f) \ (f \ inv \ \ g \ f) \ (inv \ \ f \ g \ f)" using hcomp_assoc by simp ultimately show ?thesis using comp_assoc by simp qed also have "... = (\ \ f) \ ((f \ inv \ \ g \ f) \ (f \ g \ f \ \)) \ (inv \ \ f)" proof - have "((inv \ \ f) \ (g \ f)) \ ((?b \ f) \ \) = (inv \ \ f) \ (?b \ f) \ (g \ f) \ \" proof - have "seq (inv \ \ f) (?b \ f)" using a b 4 ide_f ide_g \_in_hhom \_in_hom iso_\ \_in_hhom \_in_hom iso_\ by blast moreover have "seq (g \ f) \" using f_in_hhom g_in_hhom ide_g ide_f \_in_hom by fast ultimately show ?thesis using interchange [of "inv \ \ f" "?b \ f" "g \ f" \] by simp qed also have "... = inv \ \ f \ \" proof - have "(inv \ \ f) \ (?b \ f) = inv \ \ f" using 4 comp_arr_dom by auto moreover have "(g \ f) \ \ = \" using \_in_hom comp_cod_arr by auto ultimately show ?thesis using hcomp_assoc by simp qed also have "... = (f \ g) \ inv \ \ (f \ \) \ (f \ ?a)" proof - have "(f \ g) \ inv \ = inv \" using \_in_hom iso_\ comp_cod_arr by auto moreover have "(f \ \) \ (f \ ?a) = f \ \" proof - have "\f \ \ : f \ ?a \ f \ g \ f\" using \_in_hom by fastforce thus ?thesis using comp_arr_dom by blast qed ultimately show ?thesis by argo qed also have "... = ((f \ g) \ (f \ \)) \ (inv \ \ (f \ ?a))" proof - have "seq (f \ g) (inv \)" using \_in_hom iso_\ comp_cod_arr by auto moreover have "seq (f \ \) (f \ ?a)" using f_in_hhom \_in_hom by force ultimately show ?thesis using interchange by simp qed also have "... = (f \ g \ f \ \) \ (inv \ \ f)" using hcomp_arr_obj hcomp_assoc by auto finally have "((inv \ \ f) \ (g \ f)) \ ((?b \ f) \ \) = (f \ g \ f \ \) \ (inv \ \ f)" by simp thus ?thesis using comp_assoc hcomp_assoc by simp qed also have "... = (\ \ f) \ ((f \ ?a \ \) \ (f \ inv \ \ ?a)) \ (inv \ \ f)" proof - have "(f \ inv \ \ g \ f) \ (f \ (g \ f) \ \) = f \ (inv \ \ g \ f) \ ((g \ f) \ \)" proof - have "seq ((inv \ \ g) \ f) ((g \ f) \ \)" proof - have "seq (inv \ \ g \ f) ((g \ f) \ \)" using f_in_hhom ide_f g_in_hhom ide_g \_in_hhom \_in_hom iso_\ apply (intro seqI) apply blast apply blast by fastforce thus ?thesis using hcomp_assoc by simp qed hence "(f \ (inv \ \ g) \ f) \ (f \ (g \ f) \ \) = f \ ((inv \ \ g) \ f) \ ((g \ f) \ \)" using whisker_left by simp thus ?thesis using hcomp_assoc by simp qed also have "... = f \ (?a \ \) \ (inv \ \ ?a)" proof - have "(inv \ \ g \ f) \ ((g \ f) \ \) = (?a \ \) \ (inv \ \ ?a)" proof - have "(inv \ \ g \ f) \ ((g \ f) \ \) = inv \ \ (g \ f) \ (g \ f) \ \" proof - have "seq (inv \) (g \ f)" using g_in_hhom ide_g \_in_hom iso_\ by force moreover have "seq (g \ f) \" using g_in_hhom ide_g \_in_hom by fastforce ultimately show ?thesis using interchange by fastforce qed also have "... = inv \ \ \" using \_in_hom iso_\ comp_arr_dom comp_cod_arr by auto also have "... = ?a \ inv \ \ \ \ ?a" using \_in_hom iso_\ comp_arr_dom comp_cod_arr by auto also have "... = (?a \ \) \ (inv \ \ ?a)" proof - have "seq ?a (inv \)" using a \_in_hom iso_\ ideD [of ?a] by (elim objE, auto) moreover have "seq \ ?a" using a \_in_hom by fastforce ultimately show ?thesis using interchange by blast qed finally show ?thesis by simp qed thus ?thesis by argo qed also have "... = (f \ ?a \ \) \ (f \ inv \ \ ?a)" proof - have "seq (?a \ \) (inv \ \ ?a)" proof (intro seqI') show "\inv \ \ ?a : (g \ f) \ ?a \ ?a \ ?a\" using a g_in_hhom \_in_hom iso_\ hseqI ide_f ide_g apply (elim in_homE in_hhomE, intro hcomp_in_vhom) by auto show "\?a \ \ : ?a \ ?a \ ?a \ g \ f\" using a \_in_hom hseqI by (intro hcomp_in_vhom, auto) qed thus ?thesis using whisker_left by simp qed finally show ?thesis using hcomp_assoc by simp qed also have "... = (\ \ f) \ ((f \ \) \ (f \ inv \)) \ (inv \ \ f)" using a \_in_hhom iso_\ hcomp_obj_arr [of ?a \] hcomp_arr_obj [of "inv \" ?a] by auto also have "... = (\ \ f) \ (inv \ \ f)" proof - have "(f \ \) \ (f \ inv \) = f \ \ \ inv \" using \_in_hhom iso_\ whisker_left inv_in_hom by auto moreover have "f \ \ \ inv \ = f \ g \ f" using \_in_hom iso_\ comp_arr_inv inv_is_inverse by auto moreover have "(f \ g \ f) \ (inv \ \ f) = inv \ \ f" proof - have "\inv \ \ f : ?b \ f \ f \ g \ f\" proof - have "\inv \ \ f : ?b \ f \ (f \ g) \ f\" using \_in_hom iso_\ by (intro hcomp_in_vhom, auto) thus ?thesis using hcomp_assoc by simp qed moreover have "f \ g \ f = cod (inv \ \ f)" using \_in_hhom iso_\ hcomp_assoc calculation by auto ultimately show ?thesis using comp_cod_arr by auto qed ultimately show ?thesis by simp qed also have "... = ?b \ f" proof - have "(\ \ f) \ (inv \ \ f) = \ \ inv \ \ f" using \_in_hhom iso_\ whisker_right by auto moreover have "\ \ inv \ = ?b" using \_in_hom iso_\ comp_arr_inv inv_is_inverse by auto ultimately show ?thesis by simp qed also have "... = f" using hcomp_obj_arr by auto finally show ?thesis by blast qed (* TODO: Putting this earlier breaks some steps in the proof. *) interpret E: equivalence_in_strict_bicategory V H \ \ src trg f g \ ?\ using ide_g \_in_hom \_in_hom g_in_hhom `iso \` `iso ?\` by (unfold_locales, auto) text \ Apply ``triangle left if and only iff right'' to show the ``right'' triangle identity. \ have triangle_right: "((g \ \ \ (f \ inv \ \ g) \ (inv \ \ f \ g)) \ (\ \ g) = g)" using triangle_left E.triangle_left_iff_right by simp text \ Use the two triangle identities to establish an adjoint equivalence and show that there is only one choice for the counit. \ show "\!\. adjoint_equivalence_in_bicategory V H \ \ src trg f g \ \" proof - have "adjoint_equivalence_in_bicategory V H \ \ src trg f g \ ?\" proof show "(?\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \) = \\<^sup>-\<^sup>1[f] \ \[f]" proof - have "(?\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \) = (?\ \ f) \ (f \ \)" proof - have "seq \\<^sup>-\<^sup>1[f, g, f] (f \ \)" using E.antipar by (intro seqI, auto) hence "\\<^sup>-\<^sup>1[f, g, f] \ (f \ \) = f \ \" using ide_f ide_g E.antipar triangle_right strict_assoc' comp_ide_arr by presburger thus ?thesis by simp qed also have "... = f" using triangle_left by simp also have "... = \\<^sup>-\<^sup>1[f] \ \[f]" using strict_lunit strict_runit by simp finally show ?thesis by simp qed show "(g \ ?\) \ \[g, f, g] \ (\ \ g) = \\<^sup>-\<^sup>1[g] \ \[g]" proof - have "(g \ ?\) \ \[g, f, g] \ (\ \ g) = (g \ ?\) \ (\ \ g)" proof - have "seq \[g, f, g] (\ \ g)" using E.antipar by (intro seqI, auto) hence "\[g, f, g] \ (\ \ g) = \ \ g" using ide_f ide_g E.antipar triangle_right strict_assoc comp_ide_arr by presburger thus ?thesis by simp qed also have "... = g" using triangle_right by simp also have "... = \\<^sup>-\<^sup>1[g] \ \[g]" using strict_lunit strict_runit by simp finally show ?thesis by blast qed qed moreover have "\\ \'. \ adjoint_equivalence_in_bicategory (\) (\) \ \ src trg f g \ \; adjoint_equivalence_in_bicategory (\) (\) \ \ src trg f g \ \' \ \ \ = \'" using adjunction_unit_determines_counit by (meson adjoint_equivalence_in_bicategory.axioms(2)) ultimately show ?thesis by auto qed qed end text \ We now apply strictification to generalize the preceding result to an arbitrary bicategory. \ context bicategory begin interpretation S: strictified_bicategory V H \ \ src trg .. notation S.vcomp (infixr "\\<^sub>S" 55) notation S.hcomp (infixr "\\<^sub>S" 53) notation S.in_hom ("\_ : _ \\<^sub>S _\") notation S.in_hhom ("\_ : _ \\<^sub>S _\") interpretation UP: fully_faithful_functor V S.vcomp S.UP using S.UP_is_fully_faithful_functor by auto interpretation UP: equivalence_pseudofunctor V H \ \ src trg S.vcomp S.hcomp S.\ S.\ S.src S.trg S.UP S.\ using S.UP_is_equivalence_pseudofunctor by auto interpretation UP: pseudofunctor_into_strict_bicategory V H \ \ src trg S.vcomp S.hcomp S.\ S.\ S.src S.trg S.UP S.\ .. lemma equivalence_refines_to_adjoint_equivalence: assumes "equivalence_map f" and "\g : trg f \ src f\" and "ide g" and "\\ : src f \ g \ f\" and "iso \" shows "\!\. adjoint_equivalence_in_bicategory V H \ \ src trg f g \ \" proof - text \ To unpack the consequences of the assumptions, we need to obtain an interpretation of @{locale equivalence_in_bicategory}, even though we don't need the associated data other than \f\, \a\, and \b\. \ obtain g' \ \ where E: "equivalence_in_bicategory V H \ \ src trg f g' \ \" using assms equivalence_map_def by auto interpret E: equivalence_in_bicategory V H \ \ src trg f g' \ \ using E by auto let ?a = "src f" and ?b = "trg f" have ide_f: "ide f" by simp have f_in_hhom: "\f : ?a \ ?b\" by simp have a: "obj ?a" and b: "obj ?b" by auto have 1: "S.equivalence_map (S.UP f)" using assms UP.preserves_equivalence_maps by simp let ?\' = "S.inv (S.\ (g, f)) \\<^sub>S S.UP \ \\<^sub>S UP.\ ?a" have 2: "\S.UP \ : S.UP ?a \\<^sub>S S.UP (g \ f)\" using assms UP.preserves_hom [of \ "src f" "g \ f"] by auto have 4: "\?\' : UP.map\<^sub>0 ?a \\<^sub>S S.UP g \\<^sub>S S.UP f\ \ S.iso ?\'" proof (intro S.comp_in_homI conjI) have 3: "S.iso (S.\ (g, f))" using assms UP.\_components_are_iso by auto show "\S.inv (S.\ (g, f)) : S.UP (g \ f) \\<^sub>S S.UP g \\<^sub>S S.UP f\" using assms 3 UP.\_in_hom(2) [of g f] UP.FF_def by auto moreover show "\UP.\ ?a : UP.map\<^sub>0 ?a \\<^sub>S S.UP ?a\" by auto moreover show "\S.UP \ : S.UP ?a \\<^sub>S S.UP (g \ f)\" using 2 by simp ultimately show "S.iso (S.inv (S.\ (g, f)) \\<^sub>S S.UP \ \\<^sub>S UP.\ ?a)" using assms 3 a UP.\_char(2) S.iso_inv_iso apply (intro S.isos_compose) by auto qed have ex_un_\': "\!\'. adjoint_equivalence_in_bicategory S.vcomp S.hcomp S.\ S.\ S.src S.trg (S.UP f) (S.UP g) ?\' \'" proof - have "\S.UP g : S.trg (S.UP f) \\<^sub>S S.src (S.UP f)\" using assms(2) by auto moreover have "S.ide (S.UP g)" by (simp add: assms(3)) ultimately show ?thesis using 1 4 S.equivalence_refines_to_adjoint_equivalence S.UP_map\<^sub>0_obj by simp qed obtain \' where \': "adjoint_equivalence_in_bicategory S.vcomp S.hcomp S.\ S.\ S.src S.trg (S.UP f) (S.UP g) ?\' \'" using ex_un_\' by auto interpret E': adjoint_equivalence_in_bicategory S.vcomp S.hcomp S.\ S.\ S.src S.trg \S.UP f\ \S.UP g\ ?\' \' using \' by auto let ?\' = "UP.\ ?b \\<^sub>S \' \\<^sub>S S.inv (S.\ (f, g))" have \': "\?\' : S.UP (f \ g) \\<^sub>S S.UP ?b\" using assms(2-3) S.UP_map\<^sub>0_obj apply (intro S.in_homI) by auto have ex_un_\: "\!\. \\ : f \ g \ ?b\ \ S.UP \ = ?\'" proof - have "\\. \\ : f \ g \ ?b\ \ S.UP \ = ?\'" proof - have "src (f \ g) = src ?b \ trg (f \ g) = trg ?b" proof - have "arr (f \ g)" using assms(2) f_in_hhom by blast thus ?thesis using assms(2) f_in_hhom by (elim hseqE, auto) qed moreover have "ide (f \ g)" using assms(2-3) by auto ultimately show ?thesis using \' UP.locally_full by auto qed moreover have "\\ \. \ \\ : f \ g \ ?b\; S.UP \ = ?\'; \\ : f \ g \ ?b\; S.UP \ = ?\' \ \ \ = \" proof - fix \ \ assume \: "\\ : f \ g \ ?b\" and \: "\\ : f \ g \ ?b\" and 1: "S.UP \ = ?\'" and 2: "S.UP \ = ?\'" have "par \ \" using \ \ by fastforce thus "\ = \" using 1 2 UP.is_faithful [of \ \] by simp qed ultimately show ?thesis by auto qed have iso_\': "S.iso ?\'" proof (intro S.isos_compose) show "S.iso (S.inv (S.\ (f, g)))" using assms UP.\_components_are_iso S.iso_inv_iso by auto show "S.iso \'" using E'.counit_is_iso by blast show "S.iso (UP.\ ?b)" using b UP.\_char(2) by simp show "S.seq (UP.\ ?b) (\' \\<^sub>S S.inv (S.\ (f, g)))" proof (intro S.seqI') show "\UP.\ ?b : UP.map\<^sub>0 ?b \\<^sub>S S.UP ?b\" using b UP.\_char by simp show "\\' \\<^sub>S S.inv (S.\ (f, g)) : S.UP (f \ g) \\<^sub>S UP.map\<^sub>0 ?b\" using assms UP.\_components_are_iso VV.arr_char S.\_in_hom [of "(f, g)"] E'.counit_in_hom S.UP_map\<^sub>0_obj apply (intro S.comp_in_homI) by auto qed thus "S.seq \' (S.inv (S.\ (f, g)))" by auto qed obtain \ where \: "\\ : f \ g \ ?b\ \ S.UP \ = ?\'" using ex_un_\ by auto interpret E'': equivalence_in_bicategory V H \ \ src trg f g \ \ using assms(1,3-5) apply unfold_locales apply simp_all using assms(2) \ apply auto[1] using \ iso_\' UP.reflects_iso [of \] by auto interpret E'': adjoint_equivalence_in_bicategory V H \ \ src trg f g \ \ proof show "(\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \) = \\<^sup>-\<^sup>1[f] \ \[f]" proof - have "S.UP ((\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \)) = S.\ (trg f, f) \\<^sub>S (S.UP \ \\<^sub>S S.\ (f, g) \\<^sub>S S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.\ (g, f)) \\<^sub>S S.UP \) \\<^sub>S S.inv (S.\ (f, src f))" using E''.UP_triangle(3) by simp also have "... = S.\ (trg f, f) \\<^sub>S (UP.\ ?b \\<^sub>S \' \\<^sub>S S.inv (S.\ (f, g)) \\<^sub>S S.\ (f, g) \\<^sub>S S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.\ (g, f)) \\<^sub>S S.UP \) \\<^sub>S S.inv (S.\ (f, src f))" using \ S.comp_assoc by simp also have "... = S.\ (trg f, f) \\<^sub>S (UP.\ ?b \\<^sub>S \' \\<^sub>S S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.\ (g, f)) \\<^sub>S S.UP \) \\<^sub>S S.inv (S.\ (f, src f))" proof - have "\' \\<^sub>S S.inv (S.\ (f, g)) \\<^sub>S S.\ (f, g) = \'" proof - have "S.iso (S.\ (f, g))" using assms by auto moreover have "S.dom (S.\ (f, g)) = S.UP f \\<^sub>S S.UP g" using assms by auto ultimately have "S.inv (S.\ (f, g)) \\<^sub>S S.\ (f, g) = S.UP f \\<^sub>S S.UP g" using S.comp_inv_arr' by simp thus ?thesis using S.comp_arr_dom E'.counit_in_hom(2) by simp qed thus ?thesis by argo qed also have "... = S.\ (trg f, f) \\<^sub>S (UP.\ ?b \\<^sub>S S.UP f) \\<^sub>S ((\' \\<^sub>S S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.\ (g, f))) \\<^sub>S (S.UP f \\<^sub>S S.UP \)) \\<^sub>S S.inv (S.\ (f, src f))" proof - have "UP.\ ?b \\<^sub>S \' \\<^sub>S S.UP f = (UP.\ ?b \\<^sub>S S.UP f) \\<^sub>S (\' \\<^sub>S S.UP f)" using assms b UP.\_char S.whisker_right S.UP_map\<^sub>0_obj by auto moreover have "S.UP f \\<^sub>S S.inv (S.\ (g, f)) \\<^sub>S S.UP \ = (S.UP f \\<^sub>S S.inv (S.\ (g, f))) \\<^sub>S (S.UP f \\<^sub>S S.UP \)" using assms S.whisker_left S.comp_assoc by auto ultimately show ?thesis using S.comp_assoc by simp qed also have "... = (S.\ (trg f, f) \\<^sub>S (UP.\ ?b \\<^sub>S S.UP f)) \\<^sub>S (S.UP f \\<^sub>S S.inv (UP.\ (src f))) \\<^sub>S S.inv (S.\ (f, src f))" proof - have "(\' \\<^sub>S S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.\ (g, f))) \\<^sub>S (S.UP f \\<^sub>S S.UP \) = (S.UP f \\<^sub>S S.inv (UP.\ (src f)))" proof - have "(S.UP f \\<^sub>S S.inv (S.\ (g, f))) \\<^sub>S (S.UP f \\<^sub>S S.UP \) = S.UP f \\<^sub>S S.inv (S.\ (g, f)) \\<^sub>S S.UP \" using assms S.whisker_left by auto hence "((\' \\<^sub>S S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.\ (g, f))) \\<^sub>S (S.UP f \\<^sub>S S.UP \)) = ((\' \\<^sub>S S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.\ (g, f)) \\<^sub>S S.UP \))" by simp also have "... = ((\' \\<^sub>S S.UP f) \\<^sub>S S.\' (S.UP f) (S.UP g) (S.UP f)) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.\ (g, f)) \\<^sub>S S.UP \)" proof - have "(\' \\<^sub>S S.UP f) \\<^sub>S S.\' (S.UP f) (S.UP g) (S.UP f) = \' \\<^sub>S S.UP f" proof - have "\\' \\<^sub>S S.UP f : (S.UP f \\<^sub>S S.UP g) \\<^sub>S S.UP f \\<^sub>S S.trg (S.UP f) \\<^sub>S S.UP f\" using assms by (intro S.hcomp_in_vhom, auto) moreover have "\S.\' (S.UP f) (S.UP g) (S.UP f) : S.UP f \\<^sub>S S.UP g \\<^sub>S S.UP f \\<^sub>S (S.UP f \\<^sub>S S.UP g) \\<^sub>S S.UP f\" using assms S.assoc'_in_hom by auto ultimately show ?thesis using assms S.strict_assoc' S.iso_assoc S.hcomp_assoc E'.antipar S.comp_arr_ide S.seqI' by (metis (no_types, lifting) E'.ide_left E'.ide_right) qed thus ?thesis using S.comp_assoc by simp qed also have "... = ((\' \\<^sub>S S.UP f) \\<^sub>S S.\' (S.UP f) (S.UP g) (S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.\ (g, f)) \\<^sub>S S.UP \))" using S.comp_assoc by simp also have "... = (S.UP f \\<^sub>S S.inv (UP.\ (src f)))" proof - have "(\' \\<^sub>S S.UP f) \\<^sub>S S.\' (S.UP f) (S.UP g) (S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.\ (g, f)) \\<^sub>S S.UP \) = (S.UP f \\<^sub>S S.inv (UP.\ (src f)))" proof - have "(\' \\<^sub>S S.UP f) \\<^sub>S S.\' (S.UP f) (S.UP g) (S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.\ (g, f)) \\<^sub>S S.UP \) \\<^sub>S (S.UP f \\<^sub>S UP.\ ?a) = S.lunit' (S.UP f) \\<^sub>S S.runit (S.UP f)" proof - have "(\' \\<^sub>S S.UP f) \\<^sub>S S.\' (S.UP f) (S.UP g) (S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.\ (g, f)) \\<^sub>S S.UP \) \\<^sub>S (S.UP f \\<^sub>S UP.\ ?a) = (\' \\<^sub>S S.UP f) \\<^sub>S S.\' (S.UP f) (S.UP g) (S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.\ (g, f)) \\<^sub>S S.UP \ \\<^sub>S UP.\ ?a)" proof - have "S.seq (S.inv (S.\ (g, f)) \\<^sub>S S.UP \) (UP.\ ?a)" using assms UP.\_char UP.\_components_are_iso E'.unit_simps(1) S.comp_assoc by presburger hence "(S.UP f \\<^sub>S S.inv (S.\ (g, f)) \\<^sub>S S.UP \) \\<^sub>S (S.UP f \\<^sub>S UP.\ ?a) = S.UP f \\<^sub>S S.inv (S.\ (g, f)) \\<^sub>S S.UP \ \\<^sub>S UP.\ ?a" using assms UP.\_char UP.\_components_are_iso S.comp_assoc S.whisker_left [of "S.UP f" "S.inv (S.\ (g, f)) \\<^sub>S S.UP \" "UP.\ ?a"] by simp thus ?thesis by simp qed thus ?thesis using assms E'.triangle_left UP.\_components_are_iso UP.\_char by argo qed also have "... = S.UP f" using S.strict_lunit' S.strict_runit by simp finally have 1: "((\' \\<^sub>S S.UP f) \\<^sub>S S.\' (S.UP f) (S.UP g) (S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.\ (g, f)) \\<^sub>S S.UP \)) \\<^sub>S (S.UP f \\<^sub>S UP.\ ?a) = S.UP f" using S.comp_assoc by simp have "(\' \\<^sub>S S.UP f) \\<^sub>S S.\' (S.UP f) (S.UP g) (S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.\ (g, f)) \\<^sub>S S.UP \) = S.UP f \\<^sub>S (S.UP f \\<^sub>S S.inv (UP.\ ?a))" proof - have "S.arr (S.UP f)" using assms by simp moreover have "S.iso (S.UP f \\<^sub>S UP.\ ?a)" using assms UP.\_char S.UP_map\<^sub>0_obj by auto moreover have "S.inv (S.UP f \\<^sub>S UP.\ ?a) = S.UP f \\<^sub>S S.inv (UP.\ ?a)" using assms a UP.\_char S.UP_map\<^sub>0_obj by auto ultimately show ?thesis using assms 1 UP.\_char UP.\_components_are_iso S.invert_side_of_triangle(2) [of "S.UP f" "(\' \\<^sub>S S.UP f) \\<^sub>S S.\' (S.UP f) (S.UP g) (S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.\ (g, f)) \\<^sub>S S.UP \)" "S.UP f \\<^sub>S UP.\ ?a"] by presburger qed also have "... = S.UP f \\<^sub>S S.inv (UP.\ ?a)" proof - have "\S.UP f \\<^sub>S S.inv (UP.\ ?a) : S.UP f \\<^sub>S S.UP ?a \\<^sub>S S.UP f \\<^sub>S UP.map\<^sub>0 ?a\" using assms ide_f f_in_hhom UP.\_char [of ?a] S.inv_in_hom apply (intro S.hcomp_in_vhom) apply auto[1] apply blast by auto moreover have "S.UP f \\<^sub>S UP.map\<^sub>0 ?a = S.UP f" using a S.hcomp_arr_obj S.UP_map\<^sub>0_obj by auto finally show ?thesis using S.comp_cod_arr by blast qed finally show ?thesis by auto qed thus ?thesis using S.comp_assoc by simp qed finally show ?thesis by simp qed thus ?thesis using S.comp_assoc by simp qed also have "... = S.UP \\<^sup>-\<^sup>1[f] \\<^sub>S S.UP \[f]" proof - have "S.\ (trg f, f) \\<^sub>S (UP.\ ?b \\<^sub>S S.UP f) = S.UP \\<^sup>-\<^sup>1[f]" proof - have "S.UP f = S.UP \[f] \\<^sub>S S.\ (trg f, f) \\<^sub>S (UP.\ (trg f) \\<^sub>S S.UP f)" using UP.lunit_coherence iso_lunit S.strict_lunit by simp thus ?thesis using UP.image_of_unitor(3) ide_f by presburger qed moreover have "(S.UP f \\<^sub>S S.inv (UP.\ (src f))) \\<^sub>S S.inv (S.\ (f, src f)) = S.UP \[f]" proof - have "S.UP \[f] \\<^sub>S S.\ (f, src f) \\<^sub>S (S.UP f \\<^sub>S UP.\ (src f)) = S.UP f" using UP.runit_coherence [of f] S.strict_runit by simp moreover have "S.iso (S.\ (f, src f) \\<^sub>S (S.UP f \\<^sub>S UP.\ (src f)))" using UP.\_char UP.\_components_are_iso VV.arr_char S.UP_map\<^sub>0_obj apply (intro S.isos_compose S.seqI) by simp_all ultimately have "S.UP \[f] = S.UP f \\<^sub>S S.inv (S.\ (f, src f) \\<^sub>S (S.UP f \\<^sub>S UP.\ (src f)))" using S.invert_side_of_triangle(2) [of "S.UP f" "S.UP \[f]" "S.\ (f, src f) \\<^sub>S (S.UP f \\<^sub>S UP.\ (src f))"] ideD(1) ide_f by blast thus ?thesis using ide_f UP.image_of_unitor(2) [of f] by argo qed ultimately show ?thesis using S.comp_assoc by simp qed also have "... = S.UP (\\<^sup>-\<^sup>1[f] \ \[f])" by simp finally have "S.UP ((\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \)) = S.UP (\\<^sup>-\<^sup>1[f] \ \[f])" by simp moreover have "par ((\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \)) (\\<^sup>-\<^sup>1[f] \ \[f])" proof - have "\(\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \) : f \ src f \ trg f \ f\" using E''.triangle_in_hom(1) by simp moreover have "\\\<^sup>-\<^sup>1[f] \ \[f] : f \ src f \ trg f \ f\" by auto ultimately show ?thesis by (metis in_homE) qed ultimately show ?thesis using UP.is_faithful by blast qed thus "(g \ \) \ \[g, f, g] \ (\ \ g) = \\<^sup>-\<^sup>1[g] \ \[g]" using E''.triangle_left_implies_right by simp qed show ?thesis using E''.adjoint_equivalence_in_bicategory_axioms E''.adjunction_in_bicategory_axioms adjunction_unit_determines_counit adjoint_equivalence_in_bicategory_def by metis qed lemma equivalence_map_extends_to_adjoint_equivalence: assumes "equivalence_map f" shows "\g \ \. adjoint_equivalence_in_bicategory V H \ \ src trg f g \ \" proof - obtain g \ \' where E: "equivalence_in_bicategory V H \ \ src trg f g \ \'" using assms equivalence_map_def by auto interpret E: equivalence_in_bicategory V H \ \ src trg f g \ \' using E by auto obtain \ where A: "adjoint_equivalence_in_bicategory V H \ \ src trg f g \ \" using assms equivalence_refines_to_adjoint_equivalence [of f g \] E.antipar E.unit_is_iso E.unit_in_hom by auto show ?thesis using E A by blast qed end subsection "Uniqueness of Adjoints" text \ Left and right adjoints determine each other up to isomorphism. \ context strict_bicategory begin lemma left_adjoint_determines_right_up_to_iso: assumes "adjoint_pair f g" and "adjoint_pair f g'" shows "g \ g'" proof - obtain \ \ where A: "adjunction_in_bicategory V H \ \ src trg f g \ \" using assms adjoint_pair_def by auto interpret A: adjunction_in_bicategory V H \ \ src trg f g \ \ using A by auto interpret A: adjunction_in_strict_bicategory V H \ \ src trg f g \ \ .. obtain \' \' where A': "adjunction_in_bicategory V H \ \ src trg f g' \' \'" using assms adjoint_pair_def by auto interpret A': adjunction_in_bicategory V H \ \ src trg f g' \' \' using A' by auto interpret A': adjunction_in_strict_bicategory V H \ \ src trg f g' \' \' .. let ?\ = "A'.trnl\<^sub>\ g \" have "\?\: g \ g'\" using A'.trnl\<^sub>\_eq A'.adjoint_transpose_left(1) [of "trg f" g] A.antipar A'.antipar hcomp_arr_obj by auto moreover have "iso ?\" proof (intro isoI) let ?\ = "A.trnl\<^sub>\ g' \'" show "inverse_arrows ?\ ?\" proof show "ide (?\ \ ?\)" proof - have 1: "ide (trg f) \ trg (trg f) = trg f" by simp have "?\ \ ?\ = (g' \ \) \ ((\' \ g) \ (g \ \')) \ (\ \ g')" using 1 A.antipar A'.antipar A.trnl\<^sub>\_eq [of "trg f" g' \'] A'.trnl\<^sub>\_eq [of "trg f" g \] comp_assoc A.counit_in_hom A'.counit_in_hom by simp also have "... = ((g' \ \) \ (g' \ f \ g \ \')) \ ((\' \ g \ f \ g') \ (\ \ g'))" proof - have "(\' \ g) \ (g \ \') = (\' \ g \ trg f) \ (src f \ g \ \')" using A.antipar A'.antipar hcomp_arr_obj hcomp_obj_arr [of "src f" "g \ \'"] hseqI' by (metis A'.counit_simps(1) A'.counit_simps(5) A.ide_right ideD(1) obj_trg trg_hcomp) also have "... = \' \ g \ \'" using A.antipar A'.antipar interchange [of \' "src f" "g \ trg f" "g \ \'"] whisker_left comp_arr_dom comp_cod_arr by simp also have "... = ((g' \ f) \ g \ \') \ (\' \ g \ (f \ g'))" using A.ide_left A.ide_right A'.ide_right A.antipar A'.antipar A'.unit_in_hom A'.counit_in_hom interchange whisker_left comp_arr_dom comp_cod_arr by (metis A'.counit_simps(1-2,5) A'.unit_simps(1,3) hseqI' ide_char) also have "... = (g' \ f \ g \ \') \ (\' \ g \ f \ g')" using hcomp_assoc by simp finally show ?thesis using comp_assoc by simp qed also have "... = (g' \ \') \ ((g' \ (\ \ f) \ g') \ (g' \ (f \ \) \ g')) \ (\' \ g')" proof - have "(g' \ \) \ (g' \ f \ g \ \') = (g' \ \') \ (g' \ \ \ f \ g')" proof - have "(g' \ \) \ (g' \ f \ g \ \') = g' \ \ \ \'" proof - have "\ \ (f \ g \ \') = \ \ \'" using A.ide_left A.ide_right A.antipar A'.antipar hcomp_arr_obj comp_arr_dom comp_cod_arr interchange obj_src trg_src by (metis A'.counit_simps(1,3) A.counit_simps(1-2,4) hcomp_assoc) thus ?thesis using A.antipar A'.antipar whisker_left [of g' \ "f \ g \ \'"] by (simp add: hcomp_assoc) qed also have "... = (g' \ \') \ (g' \ \ \ f \ g')" proof - have "\ \ \' = \' \ (\ \ f \ g')" using A.ide_left A.ide_right A'.ide_right A.antipar A'.antipar hcomp_obj_arr hcomp_arr_obj comp_arr_dom comp_cod_arr interchange obj_src trg_src by (metis A'.counit_simps(1-2,5) A.counit_simps(1,3-4) arr_cod not_arr_null seq_if_composable) thus ?thesis using A.ide_left A.ide_right A'.ide_right A.antipar A'.antipar whisker_left by (metis A'.counit_simps(1,5) A.counit_simps(1,4) hseqI') qed finally show ?thesis by simp qed moreover have "(\' \ g \ f \ g') \ (\ \ g') = (g' \ f \ \ \ g') \ (\' \ g')" proof - have "(\' \ g \ f \ g') \ (\ \ g') = \' \ \ \ g'" proof - have "(\' \ g \ f) \ \ = \' \ \" using A.ide_left A.ide_right A.antipar A'.antipar A'.unit_in_hom hcomp_arr_obj interchange comp_arr_dom comp_cod_arr by (metis A'.unit_simps(1-2,4) A.unit_simps(1,3,5) hcomp_obj_arr obj_trg) thus ?thesis using A.antipar A'.antipar whisker_right [of g' "\' \ g \ f" \] by (simp add: hcomp_assoc) qed also have "... = (g' \ f \ \ \ g') \ (\' \ g')" proof - have "\' \ \ = (g' \ f \ \) \ \'" using A.ide_left A.ide_right A.antipar A'.antipar A'.unit_in_hom hcomp_arr_obj comp_arr_dom comp_cod_arr hcomp_assoc interchange by (metis A'.unit_simps(1,3-4) A.unit_simps(1-2) obj_src) thus ?thesis using A.ide_left A.ide_right A.antipar A'.antipar A'.unit_in_hom hcomp_arr_obj whisker_right [of g' "g' \ f \ \" \'] by (metis A'.ide_right A'.unit_simps(1,4) A.unit_simps(1,5) hseqI' hcomp_assoc) qed finally show ?thesis by simp qed ultimately show ?thesis using comp_assoc hcomp_assoc by simp qed also have "... = (g' \ \') \ ((g' \ f) \ g') \ (\' \ g')" proof - have "(g' \ (\ \ f) \ g') \ (g' \ (f \ \) \ g') = g' \ f \ g'" proof - have "(g' \ (\ \ f) \ g') \ (g' \ (f \ \) \ g') = g' \ ((\ \ f) \ g') \ ((f \ \) \ g')" using A.ide_left A.ide_right A.antipar A'.antipar A'.unit_in_hom A'.counit_in_hom whisker_left [of g' "(\ \ f) \ g'" "(f \ \) \ g'"] by (metis A'.ide_right A.triangle_left hseqI' ideD(1) whisker_right) also have "... = g' \ (\ \ f) \ (f \ \) \ g'" using A.antipar A'.antipar whisker_right [of g' "\ \ f" "f \ \"] by (simp add: A.triangle_left) also have "... = g' \ f \ g'" using A.triangle_left by simp finally show ?thesis by simp qed thus ?thesis using hcomp_assoc by simp qed also have "... = (g' \ \') \ (\' \ g')" using A.antipar A'.antipar A'.unit_in_hom A'.counit_in_hom comp_cod_arr by (metis A'.ide_right A'.triangle_in_hom(2) A.ide_left arrI assoc_is_natural_2 ide_char seqE strict_assoc) also have "... = g'" using A'.triangle_right by simp finally have "?\ \ ?\ = g'" by simp thus ?thesis by simp qed show "ide (?\ \ ?\)" proof - have 1: "ide (trg f) \ trg (trg f) = trg f" by simp have "?\ \ ?\ = (g \ \') \ ((\ \ g') \ (g' \ \)) \ (\' \ g)" using A.antipar A'.antipar A'.trnl\<^sub>\_eq [of "trg f" g \] A.trnl\<^sub>\_eq [of "trg f" g' \'] comp_assoc A.counit_in_hom A'.counit_in_hom by simp also have "... = ((g \ \') \ (g \ f \ g' \ \)) \ ((\ \ g' \ f \ g) \ (\' \ g))" proof - have "(\ \ g') \ (g' \ \) = (\ \ g' \ trg f) \ (src f \ g' \ \)" using A.antipar A'.antipar hcomp_arr_obj hcomp_obj_arr hseqI' by (metis A'.ide_right A.unit_simps(1,4) hcomp_assoc hcomp_obj_arr ideD(1) obj_src) also have "... = \ \ g' \ \" using A.ide_left A.ide_right A'.ide_right A.antipar A'.antipar A.unit_in_hom A.counit_in_hom interchange by (metis "1" A.counit_simps(5) A.unit_simps(4) hseqI' ide_def ide_in_hom(2) not_arr_null seqI' src.preserves_ide) also have "... = ((g \ f) \ g' \ \) \ (\ \ g' \ (f \ g))" using A'.ide_right A'.antipar interchange ide_char comp_arr_dom comp_cod_arr hseqI' by (metis A.counit_simps(1-2,5) A.unit_simps(1,3)) also have "... = (g \ f \ g' \ \) \ (\ \ g' \ f \ g)" using hcomp_assoc by simp finally show ?thesis using comp_assoc by simp qed also have "... = (g \ \) \ ((g \ (\' \ f) \ g) \ (g \ (f \ \') \ g)) \ (\ \ g)" proof - have "(g \ \') \ (g \ f \ g' \ \) = (g \ \) \ (g \ \' \ f \ g)" proof - have "(g \ \') \ (g \ f \ g' \ \) = g \ \' \ \" proof - have "\' \ (f \ g' \ \) = \' \ \" using A.ide_left A.ide_right A'.ide_right A.antipar A'.antipar hcomp_arr_obj comp_arr_dom comp_cod_arr interchange obj_src trg_src hcomp_assoc by (metis A.counit_simps(1,3) A'.counit_simps(1-2,4)) thus ?thesis using A.antipar A'.antipar whisker_left [of g \' "f \ g' \ \"] by (simp add: hcomp_assoc) qed also have "... = (g \ \) \ (g \ \' \ f \ g)" proof - have "\' \ \ = \ \ (\' \ f \ g)" using A.ide_left A.ide_right A'.ide_right A.antipar A'.antipar hcomp_obj_arr hcomp_arr_obj comp_arr_dom comp_cod_arr interchange obj_src trg_src by (metis A.counit_simps(1-2,5) A'.counit_simps(1,3-4) arr_cod not_arr_null seq_if_composable) thus ?thesis using A.ide_left A.ide_right A'.ide_right A.antipar A'.antipar whisker_left by (metis A.counit_simps(1,5) A'.counit_simps(1,4) hseqI') qed finally show ?thesis by simp qed moreover have "(\ \ g' \ f \ g) \ (\' \ g) = (g \ f \ \' \ g) \ (\ \ g)" proof - have "(\ \ g' \ f \ g) \ (\' \ g) = \ \ \' \ g" proof - have "(\ \ g' \ f) \ \' = \ \ \'" using A.antipar A'.antipar A.unit_in_hom hcomp_arr_obj comp_arr_dom comp_cod_arr hcomp_obj_arr interchange by (metis A'.unit_simps(1,3,5) A.unit_simps(1-2,4) obj_trg) thus ?thesis using A.antipar A'.antipar whisker_right [of g "\ \ g' \ f" \'] by (simp add: hcomp_assoc) qed also have "... = ((g \ f) \ \' \ g) \ (\ \ src f \ g)" using A.ide_left A.ide_right A'.ide_right A.antipar A'.antipar A.unit_in_hom A'.unit_in_hom comp_arr_dom comp_cod_arr interchange by (metis A'.unit_simps(1-2,4) A.unit_simps(1,3) hseqI' ide_char) also have "... = (g \ f \ \' \ g) \ (\ \ g)" using A.antipar A'.antipar hcomp_assoc by (simp add: hcomp_obj_arr) finally show ?thesis by simp qed ultimately show ?thesis using comp_assoc hcomp_assoc by simp qed also have "... = (g \ \) \ ((g \ f) \ g) \ (\ \ g)" proof - have "(g \ (\' \ f) \ g) \ (g \ (f \ \') \ g) = g \ f \ g" proof - have "(g \ (\' \ f) \ g) \ (g \ (f \ \') \ g) = g \ ((\' \ f) \ g) \ ((f \ \') \ g)" using A.ide_left A.ide_right A'.ide_right A.antipar A'.antipar A.unit_in_hom A.counit_in_hom whisker_left by (metis A'.triangle_left hseqI' ideD(1) whisker_right) also have "... = g \ (\' \ f) \ (f \ \') \ g" using A.antipar A'.antipar whisker_right [of g "\' \ f" "f \ \'"] by (simp add: A'.triangle_left) also have "... = g \ f \ g" using A'.triangle_left by simp finally show ?thesis by simp qed thus ?thesis using hcomp_assoc by simp qed also have "... = (g \ \) \ (\ \ g)" using A.antipar A'.antipar A.unit_in_hom A.counit_in_hom comp_cod_arr by (metis A.ide_left A.ide_right A.triangle_in_hom(2) arrI assoc_is_natural_2 ide_char seqE strict_assoc) also have "... = g" using A.triangle_right by simp finally have "?\ \ ?\ = g" by simp moreover have "ide g" by simp ultimately show ?thesis by simp qed qed qed ultimately show ?thesis using isomorphic_def by auto qed end text \ We now use strictification to extend to arbitrary bicategories. \ context bicategory begin interpretation S: strictified_bicategory V H \ \ src trg .. notation S.vcomp (infixr "\\<^sub>S" 55) notation S.hcomp (infixr "\\<^sub>S" 53) notation S.in_hom ("\_ : _ \\<^sub>S _\") notation S.in_hhom ("\_ : _ \\<^sub>S _\") interpretation UP: equivalence_pseudofunctor V H \ \ src trg S.vcomp S.hcomp S.\ S.\ S.src S.trg S.UP S.\ using S.UP_is_equivalence_pseudofunctor by auto interpretation UP: pseudofunctor_into_strict_bicategory V H \ \ src trg S.vcomp S.hcomp S.\ S.\ S.src S.trg S.UP S.\ .. interpretation UP: fully_faithful_functor V S.vcomp S.UP using S.UP_is_fully_faithful_functor by auto lemma left_adjoint_determines_right_up_to_iso: assumes "adjoint_pair f g" and "adjoint_pair f g'" shows "g \ g'" proof - have 0: "ide g \ ide g'" using assms adjoint_pair_def adjunction_in_bicategory_def adjunction_data_in_bicategory_def adjunction_data_in_bicategory_axioms_def by metis have 1: "S.adjoint_pair (S.UP f) (S.UP g) \ S.adjoint_pair (S.UP f) (S.UP g')" using assms UP.preserves_adjoint_pair by simp obtain \ where \: "\\ : S.UP g \\<^sub>S S.UP g'\ \ S.iso \" using 1 S.left_adjoint_determines_right_up_to_iso S.isomorphic_def by blast obtain \ where \: "\\ : g \ g'\ \ S.UP \ = \" using 0 \ UP.is_full [of g' g \] by auto have "\\ : g \ g'\ \ iso \" using \ \ UP.reflects_iso by auto thus ?thesis using isomorphic_def by auto qed lemma right_adjoint_determines_left_up_to_iso: assumes "adjoint_pair f g" and "adjoint_pair f' g" shows "f \ f'" proof - obtain \ \ where A: "adjunction_in_bicategory V H \ \ src trg f g \ \" using assms adjoint_pair_def by auto interpret A: adjunction_in_bicategory V H \ \ src trg f g \ \ using A by auto obtain \' \' where A': "adjunction_in_bicategory V H \ \ src trg f' g \' \'" using assms adjoint_pair_def by auto interpret A': adjunction_in_bicategory V H \ \ src trg f' g \' \' using A' by auto interpret Cop: op_bicategory V H \ \ src trg .. interpret Aop: adjunction_in_bicategory V Cop.H Cop.\ \ Cop.src Cop.trg g f \ \ using A.antipar A.triangle_left A.triangle_right Cop.assoc_ide_simp Cop.lunit_ide_simp Cop.runit_ide_simp by (unfold_locales, auto) interpret Aop': adjunction_in_bicategory V Cop.H Cop.\ \ Cop.src Cop.trg g f' \' \' using A'.antipar A'.triangle_left A'.triangle_right Cop.assoc_ide_simp Cop.lunit_ide_simp Cop.runit_ide_simp by (unfold_locales, auto) show ?thesis using Aop.adjunction_in_bicategory_axioms Aop'.adjunction_in_bicategory_axioms Cop.left_adjoint_determines_right_up_to_iso Cop.adjoint_pair_def by blast qed end context chosen_right_adjoints begin lemma isomorphic_to_left_adjoint_implies_isomorphic_right_adjoint: assumes "is_left_adjoint f" and "f \ h" shows "f\<^sup>* \ h\<^sup>*" proof - have 1: "adjoint_pair f f\<^sup>*" using assms left_adjoint_extends_to_adjoint_pair by blast moreover have "adjoint_pair h f\<^sup>*" using assms 1 adjoint_pair_preserved_by_iso isomorphic_symmetric isomorphic_reflexive by (meson isomorphic_def right_adjoint_simps(1)) thus ?thesis using left_adjoint_determines_right_up_to_iso left_adjoint_extends_to_adjoint_pair by blast qed end context bicategory begin lemma equivalence_is_adjoint: assumes "equivalence_map f" shows equivalence_is_left_adjoint: "is_left_adjoint f" and equivalence_is_right_adjoint: "is_right_adjoint f" proof - obtain g \ \ where fg: "adjoint_equivalence_in_bicategory V H \ \ src trg f g \ \" using assms equivalence_map_extends_to_adjoint_equivalence by blast interpret fg: adjoint_equivalence_in_bicategory V H \ \ src trg f g \ \ using fg by simp interpret gf: adjoint_equivalence_in_bicategory V H \ \ src trg g f \inv \\ \inv \\ using fg.dual_adjoint_equivalence by simp show "is_left_adjoint f" using fg.adjunction_in_bicategory_axioms adjoint_pair_def by auto show "is_right_adjoint f" using gf.adjunction_in_bicategory_axioms adjoint_pair_def by auto qed lemma right_adjoint_to_equivalence_is_equivalence: assumes "equivalence_map f" and "adjoint_pair f g" shows "equivalence_map g" proof - obtain \ \ where fg: "adjunction_in_bicategory (\) (\) \ \ src trg f g \ \" using assms adjoint_pair_def by auto interpret fg: adjunction_in_bicategory \(\)\ \(\)\ \ \ src trg f g \ \ using fg by simp obtain g' \ \ where fg': "equivalence_in_bicategory (\) (\) \ \ src trg f g' \ \" using assms equivalence_map_def by auto interpret fg': equivalence_in_bicategory \(\)\ \(\)\ \ \ src trg f g' \ \ using fg' by auto obtain \' where \': "adjoint_equivalence_in_bicategory (\) (\) \ \ src trg f g' \ \'" using assms equivalence_refines_to_adjoint_equivalence [of f g' \] fg'.antipar fg'.unit_in_hom fg'.unit_is_iso by auto interpret \': adjoint_equivalence_in_bicategory \(\)\ \(\)\ \ \ src trg f g' \ \' using \' by simp have 1: "g \ g'" using fg.adjunction_in_bicategory_axioms \'.adjunction_in_bicategory_axioms left_adjoint_determines_right_up_to_iso adjoint_pair_def by blast obtain \ where \: "\\ : g' \ g\ \ iso \" using 1 isomorphic_def isomorphic_symmetric by metis have "equivalence_in_bicategory (\) (\) \ \ src trg f g ((\ \ f) \ \) (\' \ (f \ inv \))" using \ equivalence_preserved_by_iso_right \'.equivalence_in_bicategory_axioms by simp hence "equivalence_pair f g" using equivalence_pair_def by auto thus ?thesis using equivalence_pair_symmetric equivalence_pair_def equivalence_map_def by blast qed lemma left_adjoint_to_equivalence_is_equivalence: assumes "equivalence_map f" and "adjoint_pair g f" shows "equivalence_map g" proof - obtain \ \ where gf: "adjunction_in_bicategory (\) (\) \ \ src trg g f \ \" using assms adjoint_pair_def by auto interpret gf: adjunction_in_bicategory \(\)\ \(\)\ \ \ src trg g f \ \ using gf by simp obtain g' where 1: "equivalence_pair g' f" using assms equivalence_map_def equivalence_pair_def equivalence_pair_symmetric by blast obtain \ \ where g'f: "equivalence_in_bicategory (\) (\) \ \ src trg g' f \ \" using assms 1 equivalence_pair_def by auto interpret g'f: equivalence_in_bicategory \(\)\ \(\)\ \ \ src trg g' f \ \ using g'f by auto obtain \' where \': "adjoint_equivalence_in_bicategory (\) (\) \ \ src trg g' f \ \'" using assms 1 equivalence_refines_to_adjoint_equivalence [of g' f \] g'f.antipar g'f.unit_in_hom g'f.unit_is_iso equivalence_pair_def equivalence_map_def by auto interpret \': adjoint_equivalence_in_bicategory \(\)\ \(\)\ \ \ src trg g' f \ \' using \' by simp have 1: "g \ g'" using gf.adjunction_in_bicategory_axioms \'.adjunction_in_bicategory_axioms right_adjoint_determines_left_up_to_iso adjoint_pair_def by blast obtain \ where \: "\\ : g' \ g\ \ iso \" using 1 isomorphic_def isomorphic_symmetric by metis have "equivalence_in_bicategory (\) (\) \ \ src trg g f ((f \ \) \ \) (\' \ (inv \ \ f))" using \ equivalence_preserved_by_iso_left \'.equivalence_in_bicategory_axioms by simp hence "equivalence_pair g f" using equivalence_pair_def by auto thus ?thesis using equivalence_pair_symmetric equivalence_pair_def equivalence_map_def by blast qed lemma equivalence_pair_is_adjoint_pair: assumes "equivalence_pair f g" shows "adjoint_pair f g" proof - obtain \ \ where \\: "equivalence_in_bicategory V H \ \ src trg f g \ \" using assms equivalence_pair_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 \\ equivalence_map_def \\.antipar \\.unit_in_hom \\.unit_is_iso \\.ide_right equivalence_refines_to_adjoint_equivalence [of f g \] by force interpret \\': adjoint_equivalence_in_bicategory V H \ \ src trg f g \ \' using \\' by auto show ?thesis using \\' adjoint_pair_def \\'.adjunction_in_bicategory_axioms by auto qed lemma equivalence_pair_isomorphic_right: assumes "equivalence_pair f g" shows "equivalence_pair f g' \ g \ g'" proof show "g \ g' \ equivalence_pair f g'" using assms equivalence_pair_def isomorphic_def equivalence_preserved_by_iso_right by metis assume g': "equivalence_pair f g'" show "g \ g'" using assms g' equivalence_pair_is_adjoint_pair left_adjoint_determines_right_up_to_iso by blast qed lemma equivalence_pair_isomorphic_left: assumes "equivalence_pair f g" shows "equivalence_pair f' g \ f \ f'" proof show "f \ f' \ equivalence_pair f' g" using assms equivalence_pair_def isomorphic_def equivalence_preserved_by_iso_left by metis assume f': "equivalence_pair f' g" show "f \ f'" using assms f' equivalence_pair_is_adjoint_pair right_adjoint_determines_left_up_to_iso by blast qed end end diff --git a/thys/Bicategory/Prebicategory.thy b/thys/Bicategory/Prebicategory.thy --- a/thys/Bicategory/Prebicategory.thy +++ b/thys/Bicategory/Prebicategory.thy @@ -1,3330 +1,3335 @@ (* Title: PreBicategory Author: Eugene W. Stark , 2019 Maintainer: Eugene W. Stark *) text \ The objective of this section is to construct a formalization of bicategories that is compatible with our previous formulation of categories \cite{Category3-AFP} and that permits us to carry over unchanged as much of the work done on categories as possible. For these reasons, we conceive of a bicategory in what might be regarded as a somewhat unusual fashion. Rather than a traditional development, which would typically define a bicategory to be essentially ``a `category' whose homs themselves have the structure of categories,'' here we regard a bicategory as ``a (vertical) category that has been equipped with a suitable (horizontal) weak composition.'' Stated another way, we think of a bicategory as a generalization of a monoidal category in which the tensor product is a partial operation, rather than a total one. Our definition of bicategory can thus be summarized as follows: a bicategory is a (vertical) category that has been equipped with idempotent endofunctors \src\ and \trg\ that assign to each arrow its ``source'' and ``target'' subject to certain commutativity constraints, a partial binary operation \\\ of horizontal composition that is suitably functorial on the ``hom-categories'' determined by the assignment of sources and targets, ``associativity'' isomorphisms \\\[f, g, h] : (f \ g) \ h \ f \ (g \ h)\\ for each horizontally composable triple of vertical identities \f\, \g\, \h\, subject to the usual naturality and coherence conditions, and for each ``object'' \a\ (defined to be an arrow that is its own source and target) a ``unit isomorphism'' \\\[a] : a \ a \ a\\. As is the case for monoidal categories, the unit isomorphisms and associator isomorphisms together enable a canonical definition of left and right ``unit'' isomorphisms \\\[f] : a \ f \ f\\ and \\\[f] : f \ a \ f\\ when \f\ is a vertical identity horizontally composable on the left or right by \a\, and it can be shown that these are the components of natural transformations. The definition of bicategory just sketched shares with a more traditional version the requirement that assignments of source and target are given as basic data, and these assignments determine horizontal composability in the sense that arrows \\\ and \\\ are composable if the chosen source of \\\ coincides with the chosen target of \\\. Thus it appears, at least on its face, that composability of arrows depends on an assignment of sources and targets. We are interested in establishing whether this is essential or whether bicategories can be formalized in a completely ``object-free'' fashion. It turns out that we can obtain such an object-free formalization through a rather direct generalization of the approach we used in the formalization of categories. Specifically, we define a \emph{weak composition} to be a partial binary operation \\\ on the arrow type of a ``vertical'' category \V\, such that the domain of definition of this operation is itself a category (of ``horizontally composable pairs of arrows''), the operation is functorial, and it is subject to certain matching conditions which include those satisfied by a category. From the axioms for a weak composition we can prove the existence of ``hom-categories'', which are subcategories of \V\ consisting of arrows horizontally composable on the left or right by a specified vertical identity. A \emph{weak unit} is defined to be a vertical identity \a\ such that \a \ a \ a\ and is such that the mappings \a \ \\ and \\ \ a\ are fully faithful endofunctors of the subcategories of \V\ consisting of the arrows for which they are defined. We define the \emph{sources} of an arrow \\\ to be the weak units that are horizontally composable with \\\ on the right, and the \emph{targets} of \\\ to be the weak units that are horizontally composable with \\\ on the left. An \emph{associative weak composition} is defined to be a weak composition that is equipped with ``associator'' isomorphisms \\\[f, g, h] : (f \ g) \ h \ f \ (g \ h)\\ for horizontally composable vertical identities \f\, \g\, \h\, subject to the usual naturality and coherence conditions. A \emph{prebicategory} is defined to be an associative weak composition for which every arrow has a source and a target. We show that the sets of sources and targets of each arrow in a prebicategory is an isomorphism class of weak units, and that horizontal composability of arrows \\\ and \\\ is characterized by the set of sources of \\\ being equal to the set of targets of \\\. We show that prebicategories are essentially ``bicategories without objects''. Given a prebicategory, we may choose an arbitrary representative of each isomorphism class of weak units and declare these to be ``objects'' (this is analogous to choosing a particular unit object in a monoidal category). For each object we may choose a particular \emph{unit isomorphism} \\\[a] : a \ a \ a\\. This choice, together with the associator isomorphisms, enables a canonical definition of left and right unit isomorphisms \\\[f] : a \ f \ f\\ and \\\[f] : f \ a \ f\\ when \f\ is a vertical identity horizontally composable on the left or right by \a\, and it can be shown that these are the components of natural isomorphisms. We may then define ``the source'' of an arrow to be the chosen representative of the set of its sources and ``the target'' to be the chosen representative of the set of its targets. We show that the resulting structure is a bicategory, in which horizontal composability as given by the weak composition coincides with the ``syntactic'' version determined by the chosen sources and targets. Conversely, a bicategory determines a prebicategory, essentially by forgetting the sources, targets and unit isomorphisms. These results make it clear that the assignment of sources and targets to arrows in a bicategory is basically a convenience and that horizontal composability of arrows is not dependent on a particular choice. \ theory Prebicategory imports Category3.EquivalenceOfCategories Category3.Subcategory IsomorphismClass begin section "Weak Composition" text \ In this section we define a locale \weak_composition\, which formalizes a functorial operation of ``horizontal'' composition defined on an underlying ``vertical'' category. The definition is expressed without the presumption of the existence of any sort of ``objects'' that determine horizontal composability. Rather, just as we did in showing that the @{locale partial_magma} locale supported the definition of ``identity arrow'' as a kind of unit for vertical composition which ultimately served as a basis for the definition of ``domain'' and ``codomain'' of an arrow, here we show that the \weak_composition\ locale supports a definition of \emph{weak unit} for horizontal composition which can ultimately be used to define the \emph{sources} and \emph{targets} of an arrow with respect to horizontal composition. In particular, the definition of weak composition involves axioms that relate horizontal and vertical composability. As a consequence of these axioms, for any fixed arrow \\\, the sets of arrows horizontally composable on the left and on the right with \\\ form subcategories with respect to vertical composition. We define the sources of \\\ to be the weak units that are composable with \\\ on the right, and the targets of \\\ to be the weak units that are composable with \\\ on the left. Weak units are then characterized as arrows that are members of the set of their own sources (or, equivalently, of their own targets). \ subsection "Definition" locale weak_composition = category V + VxV: product_category V V + VoV: subcategory VxV.comp \\\\. fst \\ \ snd \\ \ null\ + "functor" VoV.comp V \\\\. fst \\ \ snd \\\ for V :: "'a comp" (infixr "\" 55) and H :: "'a comp" (infixr "\" 53) + assumes left_connected: "seq \ \' \ \ \ \ \ null \ \' \ \ \ null" and right_connected: "seq \ \' \ \ \ \ \ null \ \ \ \' \ null" and match_1: "\ \ \ \ \ null; (\ \ \) \ \ \ null \ \ \ \ \ \ null" and match_2: "\ \ \ (\ \ \) \ null; \ \ \ \ null \ \ \ \ \ \ null" and match_3: "\ \ \ \ \ null; \ \ \ \ null \ \ (\ \ \) \ \ \ null" and match_4: "\ \ \ \ \ null; \ \ \ \ null \ \ \ \ (\ \ \) \ null" begin text \ We think of the arrows of the vertical category as ``2-cells'' and the vertical identities as ``1-cells''. In the formal development, the predicate @{term arr} (``arrow'') will have its normal meaning with respect to the vertical composition, hence @{term "arr \"} will mean, essentially, ``\\\ is a 2-cell''. This is somewhat unfortunate, as it is traditional when discussing bicategories to use the term ``arrow'' to refer to the 1-cells. However, we are trying to carry over all the formalism that we have already developed for categories and apply it to bicategories with as little change and redundancy as possible. It becomes too confusing to try to repurpose the name @{term arr} to mean @{term ide} and to introduce a replacement for the name @{term arr}, so we will simply tolerate the situation. In informal text, we will prefer the terms ``2-cell'' and ``1-cell'' over (vertical) ``arrow'' and ``identity'' when there is a chance for confusion. We do, however, make the following adjustments in notation for @{term in_hom} so that it is distinguished from the notion @{term in_hhom} (``in horizontal hom'') to be introduced subsequently. \ no_notation in_hom ("\_ : _ \ _\") notation in_hom ("\_ : _ \ _\") lemma is_partial_magma: shows "partial_magma H" proof show "\!n. \f. n \ f = n \ f \ n = n" proof show 1: "\f. null \ f = null \ f \ null = null" using is_extensional VoV.inclusion VoV.arr_char by force show "\n. \f. n \ f = n \ f \ n = n \ n = null" using 1 VoV.arr_char is_extensional not_arr_null by metis qed qed interpretation H: partial_magma H using is_partial_magma by auto text \ Either \match_1\ or \match_2\ seems essential for the next result, which states that the nulls for the horizontal and vertical compositions coincide. \ lemma null_agreement [simp]: shows "H.null = null" by (metis VoV.inclusion VxV.not_arr_null match_1 H.comp_null(1)) lemma composable_implies_arr: assumes "\ \ \ \ null" shows "arr \" and "arr \" using assms is_extensional VoV.arr_char VoV.inclusion by auto lemma hcomp_null [simp]: shows "null \ \ = null" and "\ \ null = null" using H.comp_null by auto lemma hcomp_simps\<^sub>W\<^sub>C [simp]: assumes "\ \ \ \ null" shows "arr (\ \ \)" and "dom (\ \ \) = dom \ \ dom \" and "cod (\ \ \) = cod \ \ cod \" using assms preserves_arr preserves_dom preserves_cod VoV.arr_char VoV.inclusion by force+ lemma ide_hcomp\<^sub>W\<^sub>C [simp]: assumes "ide f" and "ide g" and "g \ f \ null" shows "ide (g \ f)" using assms preserves_ide VoV.ide_char by force lemma hcomp_in_hom\<^sub>W\<^sub>C [intro]: assumes "\ \ \ \ null" shows "\\ \ \ : dom \ \ dom \ \ cod \ \ cod \\" using assms by auto text \ Horizontal composability of arrows is determined by horizontal composability of their domains and codomains (defined with respect to vertical composition). \ lemma hom_connected: shows "\ \ \ \ null \ dom \ \ \ \ null" and "\ \ \ \ null \ \ \ dom \ \ null" and "\ \ \ \ null \ cod \ \ \ \ null" and "\ \ \ \ null \ \ \ cod \ \ null" proof - show "\ \ \ \ null \ dom \ \ \ \ null" using left_connected [of \ "dom \" \] composable_implies_arr arr_dom_iff_arr by force show "\ \ \ \ null \ cod \ \ \ \ null" using left_connected [of "cod \" \ \] composable_implies_arr arr_cod_iff_arr by force show "\ \ \ \ null \ \ \ dom \ \ null" using right_connected [of \ "dom \" \] composable_implies_arr arr_dom_iff_arr by force show "\ \ \ \ null \ \ \ cod \ \ null" using right_connected [of "cod \" \ \] composable_implies_arr arr_cod_iff_arr by force qed lemma isomorphic_implies_equicomposable: assumes "f \ g" shows "\ \ f \ null \ \ \ g \ null" and "f \ \ \ null \ g \ \ \ null" using assms isomorphic_def hom_connected by auto lemma interchange: assumes "seq \ \" and "seq \ \" shows "(\ \ \) \ (\ \ \) = (\ \ \) \ (\ \ \)" proof - have "\ \ \ = null \ ?thesis" by (metis assms comp_null(2) dom_comp hom_connected(1-2)) moreover have "\ \ \ \ null \ ?thesis" proof - assume \\: "\ \ \ \ null" have 1: "VoV.arr (\, \)" using \\ VoV.arr_char by auto have \\: "(\, \) \ VoV.hom (VoV.cod (\, \)) (VoV.cod (\, \))" proof - have "VoV.arr (\, \)" using assms 1 hom_connected VoV.arr_char by (elim seqE, auto, metis) thus ?thesis using assms \\ VoV.dom_char VoV.cod_char by fastforce qed show ?thesis proof - have "VoV.seq (\, \) (\, \)" using assms 1 \\ \\ VoV.seqI by blast thus ?thesis using assms 1 \\ \\ VoV.comp_char preserves_comp [of "(\, \)" "(\, \)"] VoV.seqI by fastforce qed qed ultimately show ?thesis by blast qed lemma paste_1: shows "\ \ \ = (cod \ \ \) \ (\ \ dom \)" using interchange composable_implies_arr comp_arr_dom comp_cod_arr hom_connected(2-3) by (metis comp_null(2)) lemma paste_2: shows "\ \ \ = (\ \ cod \) \ (dom \ \ \)" using interchange composable_implies_arr comp_arr_dom comp_cod_arr hom_connected(1,4) by (metis comp_null(2)) lemma whisker_left: assumes "seq \ \" and "ide f" shows "f \ (\ \ \) = (f \ \) \ (f \ \)" using assms interchange [of f f \ \] hom_connected by auto lemma whisker_right: assumes "seq \ \" and "ide f" shows "(\ \ \) \ f = (\ \ f) \ (\ \ f)" using assms interchange [of \ \ f f] hom_connected by auto subsection "Hom-Subcategories" definition left where "left \ \ \\. \ \ \ \ null" definition right where "right \ \ \\. \ \ \ \ null" lemma right_iff_left: shows "right \ \ \ left \ \" using right_def left_def by simp lemma left_respects_isomorphic: assumes "f \ g" shows "left f = left g" using assms isomorphic_implies_equicomposable left_def by auto lemma right_respects_isomorphic: assumes "f \ g" shows "right f = right g" using assms isomorphic_implies_equicomposable right_def by auto lemma left_iff_left_inv: assumes "iso \" shows "left \ \ \ left \ (inv \)" using assms left_def inv_in_hom hom_connected(2) hom_connected(4) [of \ "inv \"] by auto lemma right_iff_right_inv: assumes "iso \" shows "right \ \ \ right \ (inv \)" using assms right_def inv_in_hom hom_connected(1) hom_connected(3) [of "inv \" \] by auto lemma left_hom_is_subcategory: assumes "arr \" shows "subcategory V (left \)" proof (unfold left_def, unfold_locales) show "\f. \ \ f \ null \ arr f" using composable_implies_arr by simp show "\f. \ \ f \ null \ \ \ dom f \ null" using hom_connected(2) by simp show "\f. \ \ f \ null \ \ \ cod f \ null" using hom_connected(4) by auto show "\f g. \ \ \ f \ null; \ \ g \ null; cod f = dom g \ \ \ \ (g \ f) \ null" proof - fix f g assume f: "\ \ f \ null" and g: "\ \ g \ null" and fg: "cod f = dom g" show "\ \ (g \ f) \ null" using f g fg composable_implies_arr hom_connected(2) [of \ "g \ f"] hom_connected(2) by simp qed qed lemma right_hom_is_subcategory: assumes "arr \" shows "subcategory V (right \)" proof (unfold right_def, unfold_locales) show "\f. f \ \ \ null \ arr f" using composable_implies_arr by simp show "\f. f \ \ \ null \ dom f \ \ \ null" using hom_connected(1) by auto show "\f. f \ \ \ null \ cod f \ \ \ null" using hom_connected(3) by auto show "\f g. \ f \ \ \ null; g \ \ \ null; cod f = dom g \ \ (g \ f) \ \ \ null" proof - fix f g assume f: "f \ \ \ null" and g: "g \ \ \ null" and fg: "cod f = dom g" show "(g \ f) \ \ \ null" using f g fg composable_implies_arr hom_connected(1) [of "g \ f" \] hom_connected(1) by simp qed qed abbreviation Left where "Left a \ subcategory.comp V (left a)" abbreviation Right where "Right a \ subcategory.comp V (right a)" text \ We define operations of composition on the left or right with a fixed 1-cell, and show that such operations are functorial in case that 1-cell is horizontally self-composable. \ definition H\<^sub>L where "H\<^sub>L g \ \\. g \ \" definition H\<^sub>R where "H\<^sub>R f \ \\. \ \ f" (* TODO: Why do the following fail when I use @{thm ...} *) text \ Note that \match_3\ and \match_4\ are required for the next results. \ lemma endofunctor_H\<^sub>L: assumes "ide g" and "g \ g \ null" shows "endofunctor (Left g) (H\<^sub>L g)" proof - interpret L: subcategory V \left g\ using assms left_hom_is_subcategory by simp have *: "\\. L.arr \ \ H\<^sub>L g \ = g \ \" using assms H\<^sub>L_def by simp have preserves_arr: "\\. L.arr \ \ L.arr (H\<^sub>L g \)" using assms * L.arr_char left_def match_4 by force show "endofunctor L.comp (H\<^sub>L g)" proof show "\\. \ L.arr \ \ H\<^sub>L g \ = L.null" using assms L.arr_char L.null_char left_def H\<^sub>L_def by fastforce show "\\. L.arr \ \ L.arr (H\<^sub>L g \)" by fact fix \ assume "L.arr \" hence \: "L.arr \ \ arr \ \ g \ \ \ null" using assms L.arr_char composable_implies_arr left_def by metis show "L.dom (H\<^sub>L g \) = H\<^sub>L g (L.dom \)" using assms \ * L.arr_char L.dom_char preserves_arr hom_connected(2) left_def by simp show "L.cod (H\<^sub>L g \) = H\<^sub>L g (L.cod \)" using assms \ * L.arr_char L.cod_char preserves_arr hom_connected(4) left_def by simp next fix \ \ assume \\: "L.seq \ \" have \: "L.arr \" using \\ by (elim L.seqE, auto) have \: "L.arr \ \ arr \ \ in_hom \ (L.cod \) (L.cod \) \ left g \ \ g \ \ \ null" proof - have 1: "L.in_hom \ (L.cod \) (L.cod \)" using \\ by (elim L.seqE, auto) hence "arr \ \ left g \" using L.hom_char by blast thus ?thesis using assms 1 left_def by fastforce qed show "H\<^sub>L g (L.comp \ \) = L.comp (H\<^sub>L g \) (H\<^sub>L g \)" proof - have "H\<^sub>L g (L.comp \ \) = g \ (\ \ \)" using \ \ H\<^sub>L_def L.comp_def L.arr_char by fastforce also have "... = (g \ \) \ (g \ \)" using assms \ \ L.inclusion whisker_left L.arr_char by fastforce also have "... = L.comp (H\<^sub>L g \) (H\<^sub>L g \)" using assms \\ \ \ * preserves_arr L.arr_char L.dom_char L.cod_char L.comp_char L.inclusion H\<^sub>L_def left_def by auto finally show ?thesis by blast qed qed qed lemma endofunctor_H\<^sub>R: assumes "ide f" and "f \ f \ null" shows "endofunctor (Right f) (H\<^sub>R f)" proof - interpret R: subcategory V \right f\ using assms right_hom_is_subcategory by simp have *: "\\. R.arr \ \ H\<^sub>R f \ = \ \ f" using assms H\<^sub>R_def by simp have preserves_arr: "\\. R.arr \ \ R.arr (H\<^sub>R f \)" using assms * R.arr_char right_def match_3 by force show "endofunctor R.comp (H\<^sub>R f)" proof show "\\. \ R.arr \ \ H\<^sub>R f \ = R.null" using assms R.arr_char R.null_char right_def H\<^sub>R_def by fastforce show "\\. R.arr \ \ R.arr (H\<^sub>R f \)" by fact fix \ assume "R.arr \" hence \: "R.arr \ \ arr \ \ \ \ f \ null" using assms R.arr_char composable_implies_arr right_def by simp show "R.dom (H\<^sub>R f \) = H\<^sub>R f (R.dom \)" using assms \ * R.arr_char R.dom_char preserves_arr hom_connected(1) right_def by simp show "R.cod (H\<^sub>R f \) = H\<^sub>R f (R.cod \)" using assms \ * R.arr_char R.cod_char preserves_arr hom_connected(3) right_def by simp next fix \ \ assume \\: "R.seq \ \" have \: "R.arr \" using \\ by (elim R.seqE, auto) have \: "R.arr \ \ arr \ \ in_hom \ (R.cod \) (R.cod \) \ right f \ \ \ \ f \ null" proof - have 1: "R.in_hom \ (R.cod \) (R.cod \)" using \\ by (elim R.seqE, auto) hence "arr \ \ right f \" using R.hom_char by blast thus ?thesis using assms 1 right_def by fastforce qed show "H\<^sub>R f (R.comp \ \) = R.comp (H\<^sub>R f \) (H\<^sub>R f \)" proof - have "H\<^sub>R f (R.comp \ \) = (\ \ \) \ f" using \ \ H\<^sub>R_def R.comp_def R.arr_char by fastforce also have "... = (\ \ f) \ (\ \ f)" using assms \ \ R.inclusion whisker_right R.arr_char by fastforce also have "... = R.comp (H\<^sub>R f \) (H\<^sub>R f \)" using assms \\ \ \ * preserves_arr R.arr_char R.dom_char R.cod_char R.comp_char R.inclusion H\<^sub>R_def right_def by auto finally show ?thesis by blast qed qed qed end locale left_hom = weak_composition V H + S: subcategory V \left \\ for V :: "'a comp" (infixr "\" 55) and H :: "'a comp" (infixr "\" 53) and \ :: 'a + assumes arr_\: "arr \" begin no_notation in_hom ("\_ : _ \ _\") notation in_hom ("\_ : _ \ _\") notation S.comp (infixr "\\<^sub>S" 55) notation S.in_hom ("\_ : _ \\<^sub>S _\") lemma right_hcomp_closed [simp]: assumes "\\ : x \\<^sub>S y\" and "\\ : c \ d\" and "\ \ \ \ null" shows "\\ \ \ : x \ c \\<^sub>S y \ d\" proof show 1: "S.arr (\ \ \)" using assms arr_\ S.arr_char left_def match_4 by (elim S.in_homE, meson) show "S.dom (\ \ \) = x \ c" using assms 1 by force show "S.cod (\ \ \) = y \ d" using assms 1 by force qed lemma interchange: assumes "S.seq \ \" and "S.seq \ \" and "\ \ \ \ null" shows "(\ \\<^sub>S \) \ (\ \\<^sub>S \) = (\ \ \) \\<^sub>S (\ \ \)" proof - have 1: "\ \ \ \ null" using assms hom_connected(1) [of \ \] hom_connected(2) [of \ \] hom_connected(3-4) by force have "(\ \\<^sub>S \) \ (\ \\<^sub>S \) = (\ \ \) \ (\ \ \)" using assms S.comp_char S.seq_char by metis also have "... = (\ \ \) \ (\ \ \)" using assms interchange S.seq_char S.arr_char by simp also have "... = (\ \ \) \\<^sub>S (\ \ \)" proof - have "S.arr (\ \ \)" proof - have "\\ : dom \ \ cod \\" using assms S.in_hom_char by blast thus ?thesis using assms 1 right_hcomp_closed by blast qed moreover have "S.arr (\ \ \)" proof - have "\\ : dom \ \ cod \\" using assms S.in_hom_char by blast thus ?thesis using assms right_hcomp_closed [of \ "dom \" "cod \" \ "dom \" "cod \"] by fastforce qed moreover have "seq (\ \ \) (\ \ \)" using assms 1 S.in_hom_char by (metis (full_types) S.seq_char hcomp_simps\<^sub>W\<^sub>C(1-3) seqE seqI) ultimately show ?thesis using S.comp_char by auto qed finally show ?thesis by blast qed lemma inv_char: assumes "S.arr \" and "iso \" shows "S.inverse_arrows \ (inv \)" and "S.inv \ = inv \" proof - have 1: "S.arr (inv \)" proof - have "S.arr \" using assms by auto hence "\ \ \ \ null" using S.arr_char left_def by simp hence "\ \ cod \ \ null" using hom_connected(4) by blast hence "\ \ dom (inv \) \ null" using assms S.iso_char by simp hence "\ \ inv \ \ null" using hom_connected by blast thus "S.arr (inv \)" using S.arr_char left_def by force qed show "S.inv \ = inv \" using assms 1 S.inv_char S.iso_char by blast thus "S.inverse_arrows \ (inv \)" using assms 1 S.iso_char S.inv_is_inverse by metis qed lemma iso_char: assumes "S.arr \" shows "S.iso \ \ iso \" using assms S.iso_char inv_char by auto end locale right_hom = weak_composition V H + S: subcategory V \right \\ for V :: "'a comp" (infixr "\" 55) and H :: "'a comp" (infixr "\" 53) and \ :: 'a + assumes arr_\: "arr \" begin no_notation in_hom ("\_ : _ \ _\") notation in_hom ("\_ : _ \ _\") notation S.comp (infixr "\\<^sub>S" 55) notation S.in_hom ("\_ : _ \\<^sub>S _\") lemma left_hcomp_closed [simp]: assumes "\\ : x \\<^sub>S y\" and "\\ : c \ d\" and "\ \ \ \ null" shows "\\ \ \ : c \ x \\<^sub>S d \ y\" proof show 1: "S.arr (\ \ \)" using assms arr_\ S.arr_char right_def match_3 by (elim S.in_homE, meson) show "S.dom (\ \ \) = c \ x" using assms 1 by force show "S.cod (\ \ \) = d \ y" using assms 1 by force qed lemma interchange: assumes "S.seq \ \" and "S.seq \ \" and "\ \ \ \ null" shows "(\ \\<^sub>S \) \ (\ \\<^sub>S \) = (\ \ \) \\<^sub>S (\ \ \)" proof - have 1: "\ \ \ \ null" using assms hom_connected(1) [of \ \] hom_connected(2) [of \ \] hom_connected(3-4) by fastforce have "(\ \\<^sub>S \) \ (\ \\<^sub>S \) = (\ \ \) \ (\ \ \)" using assms S.comp_char S.seq_char by metis also have "... = (\ \ \) \ (\ \ \)" using assms interchange S.seq_char S.arr_char by simp also have "... = (\ \ \) \\<^sub>S (\ \ \)" proof - have "S.arr (\ \ \)" proof - have "\\ : dom \ \ cod \\" using assms S.in_hom_char by blast thus ?thesis using assms 1 left_hcomp_closed by blast qed moreover have "S.arr (\ \ \)" proof - have "\\ : dom \ \ cod \\" using assms S.in_hom_char by blast thus ?thesis using assms left_hcomp_closed [of \ "dom \" "cod \" \ "dom \" "cod \"] by fastforce qed moreover have "seq (\ \ \) (\ \ \)" using assms 1 S.in_hom_char by (metis (full_types) S.seq_char hcomp_simps\<^sub>W\<^sub>C(1-3) seqE seqI) ultimately show ?thesis using S.comp_char by auto qed finally show ?thesis by blast qed lemma inv_char: assumes "S.arr \" and "iso \" shows "S.inverse_arrows \ (inv \)" and "S.inv \ = inv \" proof - have 1: "S.arr (inv \)" proof - have "S.arr \" using assms by auto hence "\ \ \ \ null" using S.arr_char right_def by simp hence "cod \ \ \ \ null" using hom_connected(3) by blast hence "dom (inv \) \ \ \ null" using assms S.iso_char by simp hence "inv \ \ \ \ null" using hom_connected(1) by blast thus ?thesis using S.arr_char right_def by force qed show "S.inv \ = inv \" using assms 1 S.inv_char S.iso_char by blast thus "S.inverse_arrows \ (inv \)" using assms 1 S.iso_char S.inv_is_inverse by metis qed lemma iso_char: assumes "S.arr \" shows "S.iso \ \ iso \" using assms S.iso_char inv_char by auto end subsection "Weak Units" text \ We now define a \emph{weak unit} to be an arrow \a\ such that: \begin{enumerate} \item \a \ a\ is isomorphic to \a\ (and hence \a\ is a horizontally self-composable 1-cell). \item Horizontal composition on the left with \a\ is a fully faithful endofunctor of the subcategory of arrows that are composable on the left with \a\. \item Horizontal composition on the right with \a\ is fully faithful endofunctor of the subcategory of arrows that are composable on the right with \a\. \end{enumerate} \ context weak_composition begin definition weak_unit :: "'a \ bool" where "weak_unit a \ a \ a \ a \ fully_faithful_functor (Left a) (Left a) (H\<^sub>L a) \ fully_faithful_functor (Right a) (Right a) (H\<^sub>R a)" lemma weak_unit_self_composable [simp]: assumes "weak_unit a" shows "ide a" and "ide (a \ a)" and "a \ a \ null" proof - obtain \ where \: "\\ : a \ a \ a\ \ iso \" using assms weak_unit_def isomorphic_def by blast have 1: "arr \" using \ by blast show "ide a" using \ ide_cod by blast thus "ide (a \ a)" using \ ide_dom by force thus "a \ a \ null" using not_arr_null ideD(1) by metis qed lemma weak_unit_self_right: assumes "weak_unit a" shows "right a a" using assms weak_unit_self_composable right_def by simp lemma weak_unit_self_left: assumes "weak_unit a" shows "left a a" using assms weak_unit_self_composable left_def by simp lemma weak_unit_in_vhom: assumes "weak_unit a" shows "\a : a \ a\" using assms weak_unit_self_composable left_def by auto text \ If \a\ is a weak unit, then there exists a ``unit isomorphism'' \\\ : a \ a \ a\\. It need not be unique, but we may choose one arbitrarily. \ definition some_unit where "some_unit a \ SOME \. iso \ \ \\ : a \ a \ a\" lemma iso_some_unit: assumes "weak_unit a" shows "iso (some_unit a)" and "\some_unit a : a \ a \ a\" proof - let ?P = "\\. iso \ \ \\ : a \ a \ a\" have "\\. ?P \" using assms weak_unit_def by auto hence 1: "?P (some_unit a)" using someI_ex [of ?P] some_unit_def by simp show "iso (some_unit a)" using 1 by blast show "\some_unit a : a \ a \ a\" using 1 by blast qed text \ The \emph{sources} of an arbitrary arrow \\\ are the weak units that are composable with \\\ on the right. Similarly, the \emph{targets} of \\\ are the weak units that are composable with \\\ on the left. \ definition sources where "sources \ \ {a. weak_unit a \ \ \ a \ null}" lemma sourcesI [intro]: assumes "weak_unit a" and "\ \ a \ null" shows "a \ sources \" using assms sources_def by blast lemma sourcesD [dest]: assumes "a \ sources \" shows "ide a" and "weak_unit a" and "\ \ a \ null" using assms sources_def by auto definition targets where "targets \ \ {b. weak_unit b \ b \ \ \ null}" lemma targetsI [intro]: assumes "weak_unit b" and "b \ \ \ null" shows "b \ targets \" using assms targets_def by blast lemma targetsD [dest]: assumes "b \ targets \" shows "ide b" and "weak_unit b" and "b \ \ \ null" using assms targets_def by auto lemma sources_dom [simp]: assumes "arr \" shows "sources (dom \) = sources \" using assms hom_connected(1) by blast lemma sources_cod [simp]: assumes "arr \" shows "sources (cod \) = sources \" using assms hom_connected(3) by blast lemma targets_dom [simp]: assumes "arr \" shows "targets (dom \) = targets \" using assms hom_connected(2) by blast lemma targets_cod [simp]: assumes "arr \" shows "targets (cod \) = targets \" using assms hom_connected(4) by blast lemma weak_unit_iff_self_source: shows "weak_unit a \ a \ sources a" using weak_unit_self_composable by auto lemma weak_unit_iff_self_target: shows "weak_unit b \ b \ targets b" using weak_unit_self_composable by auto abbreviation (input) in_hhom\<^sub>W\<^sub>C ("\_ : _ \\<^sub>W\<^sub>C _\") where "in_hhom\<^sub>W\<^sub>C \ f g \ arr \ \ f \ sources \ \ g \ targets \" lemma sources_hcomp: assumes "\ \ \ \ null" shows "sources (\ \ \) = sources \" using assms match_1 match_3 null_agreement by blast lemma targets_hcomp: assumes "\ \ \ \ null" shows "targets (\ \ \) = targets \" using assms match_2 match_4 null_agreement by blast lemma H\<^sub>R_preserved_along_iso: assumes "weak_unit a" and "a \ a'" shows "endofunctor (Right a) (H\<^sub>R a')" proof - have a: "ide a \ weak_unit a" using assms isomorphic_def by auto have a': "ide a'" using assms isomorphic_def by auto (* TODO: The following interpretation re-introduces unwanted notation for "in_hom" *) interpret R: subcategory V \right a\ using a right_hom_is_subcategory by simp have *: "\\. R.arr \ \ H\<^sub>R a' \ = \ \ a'" using assms H\<^sub>R_def by simp have preserves_arr: "\\. R.arr \ \ R.arr (H\<^sub>R a' \)" using assms a' * R.arr_char right_def weak_unit_def weak_unit_self_composable isomorphic_implies_equicomposable R.ide_char match_3 hcomp_simps\<^sub>W\<^sub>C(1) null_agreement by metis show "endofunctor R.comp (H\<^sub>R a')" proof show "\\. \ R.arr \ \ H\<^sub>R a' \ = R.null" using assms R.arr_char R.null_char right_def H\<^sub>R_def null_agreement right_respects_isomorphic by metis fix \ assume "R.arr \" hence \: "R.arr \ \ arr \ \ right a \ \ right a' \ \ \ \ a \ null \ \ \ a' \ null" using assms R.arr_char right_respects_isomorphic composable_implies_arr null_agreement right_def by metis show "R.arr (H\<^sub>R a' \)" using \ preserves_arr by blast show "R.dom (H\<^sub>R a' \) = H\<^sub>R a' (R.dom \)" using a' \ * R.arr_char R.dom_char preserves_arr hom_connected(1) right_def by simp show "R.cod (H\<^sub>R a' \) = H\<^sub>R a' (R.cod \)" using a' \ * R.arr_char R.cod_char preserves_arr hom_connected(3) right_def by simp next fix \ \ assume \\: "R.seq \ \" have "R.arr \" using \\ by (elim R.seqE, auto) hence \: "R.arr \ \ arr \ \ right a \ \ right a' \ \ \ \ a \ null \ \ \ a' \ null" using assms R.arr_char right_respects_isomorphic composable_implies_arr null_agreement right_def by metis have "\ \ R.hom (R.cod \) (R.cod \)" using \\ by (elim R.seqE, auto) hence "\\ : R.cod \ \ R.cod \\ \ arr \ \ \ \ Collect (right a)" using R.hom_char by blast hence \: "\\ : R.cod \ \ R.cod \\ \ arr \ \ right a \ \ H \ a \ null \ right a' \ \ H \ a' \ null" using assms right_def right_respects_isomorphic isomorphic_implies_equicomposable by simp show "H\<^sub>R a' (R.comp \ \) = R.comp (H\<^sub>R a' \) (H\<^sub>R a' \)" proof - have 1: "R.arr (H\<^sub>R a' \)" using \ preserves_arr by blast have 2: "seq (\ \ a') (\ \ a')" using a' \ \ R.arr_char R.inclusion R.dom_char R.cod_char isomorphic_implies_equicomposable by auto show ?thesis proof - have "H\<^sub>R a' (R.comp \ \) = (\ \ \) \ a'" using \ \ H\<^sub>R_def R.comp_def by fastforce also have "... = (\ \ a') \ (\ \ a')" proof - have "seq \ \" using \ \ \\ by (elim R.seqE, auto) thus ?thesis using a' \ whisker_right right_def by blast qed also have "... = R.comp (H\<^sub>R a' \) (H\<^sub>R a' \)" using assms \ 1 2 preserves_arr R.comp_char R.inclusion H\<^sub>R_def by auto finally show ?thesis by blast qed qed qed qed lemma H\<^sub>L_preserved_along_iso: assumes "weak_unit a" and "a \ a'" shows "endofunctor (Left a) (H\<^sub>L a')" proof - have a: "ide a \ weak_unit a" using assms isomorphic_def by auto have a': "ide a'" using assms isomorphic_def by auto (* TODO: The following interpretation re-introduces unwanted notation for "in_hom" *) interpret L: subcategory V \left a\ using a left_hom_is_subcategory by simp have *: "\\. L.arr \ \ H\<^sub>L a' \ = a' \ \" using assms H\<^sub>L_def by simp have preserves_arr: "\\. L.arr \ \ L.arr (H\<^sub>L a' \)" using assms a' * L.arr_char left_def weak_unit_def weak_unit_self_composable isomorphic_implies_equicomposable L.ide_char match_4 hcomp_simps\<^sub>W\<^sub>C(1) null_agreement by metis show "endofunctor L.comp (H\<^sub>L a')" proof show "\\. \ L.arr \ \ H\<^sub>L a' \ = L.null" using assms L.arr_char L.null_char left_def H\<^sub>L_def null_agreement left_respects_isomorphic by metis fix \ assume "L.arr \" hence \: "L.arr \ \ arr \ \ left a \ \ left a' \ \ a \ \ \ null \ a' \ \ \ null" using assms L.arr_char left_respects_isomorphic composable_implies_arr null_agreement left_def by metis show "L.arr (H\<^sub>L a' \)" using \ preserves_arr by blast show "L.dom (H\<^sub>L a' \) = H\<^sub>L a' (L.dom \)" using a' \ * L.arr_char L.dom_char preserves_arr hom_connected(2) left_def by simp show "L.cod (H\<^sub>L a' \) = H\<^sub>L a' (L.cod \)" using a' \ * L.arr_char L.cod_char preserves_arr hom_connected(4) left_def by simp next fix \ \ assume \\: "L.seq \ \" have "L.arr \" using \\ by (elim L.seqE, auto) hence \: "L.arr \ \ arr \ \ left a \ \ left a' \ \ a \ \ \ null \ a' \ \ \ null" using assms L.arr_char left_respects_isomorphic composable_implies_arr null_agreement left_def by metis have "L.in_hom \ (L.cod \) (L.cod \)" using \\ by (elim L.seqE, auto) hence "\\ : L.cod \ \ L.cod \\ \ arr \ \ \ \ Collect (left a)" using L.hom_char by blast hence \: "\\ : L.cod \ \ L.cod \\ \ arr \ \ left a \ \ a \ \ \ null \ left a' \ \ a' \ \ \ null" using assms left_def left_respects_isomorphic isomorphic_implies_equicomposable by simp show "H\<^sub>L a' (L.comp \ \) = L.comp (H\<^sub>L a' \) (H\<^sub>L a' \)" proof - have 1: "L.arr (H\<^sub>L a' \)" using \ preserves_arr by blast have 2: "seq (a' \ \) (a' \ \)" using a' \ \ L.arr_char L.inclusion L.dom_char L.cod_char isomorphic_implies_equicomposable by auto have "H\<^sub>L a' (L.comp \ \) = a' \ (\ \ \)" using \ \ H\<^sub>L_def L.comp_def by fastforce also have "... = (a' \ \) \ (a' \ \)" proof - have "seq \ \" using \ \ \\ by (elim L.seqE, auto) thus ?thesis using a' \ whisker_left right_def by blast qed also have "... = L.comp (H\<^sub>L a' \) (H\<^sub>L a' \)" using assms \ 1 2 preserves_arr L.comp_char L.inclusion H\<^sub>L_def by auto finally show ?thesis by blast qed qed qed end subsection "Regularity" text \ We call a weak composition \emph{regular} if \f \ a \ f\ whenever \a\ is a source of 1-cell \f\, and \b \ f \ f\ whenever \b\ is a target of \f\. A consequence of regularity is that horizontal composability of 2-cells is fully determined by their sets of sources and targets. \ locale regular_weak_composition = weak_composition + assumes comp_ide_source: "\ a \ sources f; ide f \ \ f \ a \ f" and comp_target_ide: "\ b \ targets f; ide f \ \ b \ f \ f" begin lemma sources_determine_composability: assumes "a \ sources \" shows "\ \ \ \ null \ a \ \ \ null" proof - have *: "\\. ide \ \ a \ sources \ \ \ \ \ \ null \ a \ \ \ null" proof - fix \ assume \: "ide \ \ a \ sources \" show "\ \ \ \ null \ a \ \ \ null" proof assume \: "\ \ \ \ null" show "a \ \ \ null" using assms \ \ comp_ide_source isomorphic_implies_equicomposable match_1 by blast next assume \: "a \ \ \ null" show "\ \ \ \ null" using assms \ \ comp_ide_source isomorphic_implies_equicomposable match_3 by blast qed qed show ?thesis proof - have "arr \" using assms composable_implies_arr by auto thus ?thesis using assms * [of "dom \"] hom_connected(1) by auto qed qed lemma targets_determine_composability: assumes "b \ targets \" shows "\ \ \ \ null \ \ \ b \ null" proof - have *: "\\. ide \ \ b \ targets \ \ \ \ \ \ null \ \ \ b \ null" proof - fix \ assume \: "ide \ \ b \ targets \" show "\ \ \ \ null \ \ \ b \ null" proof assume \: "\ \ \ \ null" show "\ \ b \ null" using assms \ \ comp_target_ide isomorphic_implies_equicomposable match_2 by blast next assume \: "\ \ b \ null" show "\ \ \ \ null" using assms \ \ comp_target_ide isomorphic_implies_equicomposable match_4 by blast qed qed show ?thesis proof - have "arr \" using assms composable_implies_arr by auto thus ?thesis using assms * [of "dom \"] hom_connected(2) by auto qed qed lemma composable_if_connected: assumes "sources \ \ targets \ \ {}" shows "\ \ \ \ null" using assms targets_determine_composability by blast lemma connected_if_composable: assumes "\ \ \ \ null" shows "sources \ = targets \" using assms sources_determine_composability targets_determine_composability by blast lemma iso_hcomp\<^sub>R\<^sub>W\<^sub>C: assumes "iso \" and "iso \" and "sources \ \ targets \ \ {}" shows "iso (\ \ \)" and "inverse_arrows (\ \ \) (inv \ \ inv \)" proof - have \: "arr \ \ \\ : dom \ \ cod \\ \ iso \ \ \inv \ : cod \ \ dom \\" using assms inv_in_hom arr_iff_in_hom iso_is_arr by auto have \: "arr \ \ \\ : dom \ \ cod \\ \ iso \ \ \inv \ : cod \ \ dom \\" using assms inv_in_hom by blast have 1: "sources (inv \) \ targets (inv \) \ {}" proof - have "sources (inv \) \ targets (inv \) = sources \ \ targets \" proof - have "sources (inv \) \ targets (inv \) = sources (cod (inv \)) \ targets (cod (inv \))" using assms \ \ sources_cod targets_cod arr_inv by presburger also have "... = sources (dom \) \ targets (dom \)" using \ \ by simp also have "... = sources \ \ targets \" using \ \ sources_dom targets_dom by simp finally show ?thesis by blast qed thus ?thesis using assms by simp qed show "inverse_arrows (\ \ \) (inv \ \ inv \)" proof have "(inv \ \ inv \) \ (\ \ \) = dom \ \ dom \" using assms \ \ inv_in_hom inv_is_inverse comp_inv_arr interchange [of "inv \" \ "inv \" \] composable_if_connected by simp moreover have "ide (dom \ \ dom \)" using assms \ \ ide_hcomp\<^sub>W\<^sub>C composable_if_connected sources_dom targets_dom by auto ultimately show "ide ((inv \ \ inv \) \ (\ \ \))" by presburger have "(\ \ \) \ (inv \ \ inv \) = cod \ \ cod \" using assms 1 \ \ inv_in_hom inv_is_inverse comp_arr_inv interchange [of \ "inv \" \ "inv \"] composable_if_connected by simp moreover have "ide (cod \ \ cod \)" using assms \ \ ide_hcomp\<^sub>W\<^sub>C composable_if_connected sources_cod targets_cod by auto ultimately show "ide ((\ \ \) \ (inv \ \ inv \))" by presburger qed thus "iso (\ \ \)" by auto qed lemma inv_hcomp\<^sub>R\<^sub>W\<^sub>C: assumes "iso \" and "iso \" and "sources \ \ targets \ \ {}" shows "inv (\ \ \) = inv \ \ inv \" using assms iso_hcomp\<^sub>R\<^sub>W\<^sub>C(2) [of \ \] inverse_arrow_unique [of "H \ \"] inv_is_inverse by auto end subsection "Associativity" text \ An \emph{associative weak composition} consists of a weak composition that has been equipped with an \emph{associator} isomorphism: \\\[f, g, h] : (f \ g) \ h \ f \ g \ h\\ for each composable triple \(f, g, h)\ of 1-cells, subject to naturality and coherence conditions. \ locale associative_weak_composition = weak_composition + fixes \ :: "'a \ 'a \ 'a \ 'a" ("\[_, _, _]") assumes assoc_in_vhom\<^sub>A\<^sub>W\<^sub>C: "\ ide f; ide g; ide h; f \ g \ null; g \ h \ null \ \ \\[f, g, h] : (f \ g) \ h \ f \ g \ h\" and assoc_naturality\<^sub>A\<^sub>W\<^sub>C: "\ \ \ \ \ null; \ \ \ \ null \ \ \[cod \, cod \, cod \] \ ((\ \ \) \ \) = (\ \ \ \ \) \ \[dom \, dom \, dom \]" and iso_assoc\<^sub>A\<^sub>W\<^sub>C: "\ ide f; ide g; ide h; f \ g \ null; g \ h \ null \ \ iso \[f, g, h]" and pentagon\<^sub>A\<^sub>W\<^sub>C: "\ ide f; ide g; ide h; ide k; sources f \ targets g \ {}; sources g \ targets h \ {}; sources h \ targets k \ {} \ \ (f \ \[g, h, k]) \ \[f, g \ h, k] \ (\[f, g, h] \ k) = \[f, g, h \ k] \ \[f \ g, h, k]" begin lemma assoc_in_hom\<^sub>A\<^sub>W\<^sub>C: assumes "ide f" and "ide g" and "ide h" and "f \ g \ null" and "g \ h \ null" shows "sources \[f, g, h] = sources h" and "targets \[f, g, h] = targets f" and "\\[f, g, h] : (f \ g) \ h \ f \ g \ h\" proof - show 1: "\\[f, g, h] : (f \ g) \ h \ f \ g \ h\" using assms assoc_in_vhom\<^sub>A\<^sub>W\<^sub>C by simp show "sources \[f, g, h] = sources h" using assms 1 sources_dom [of "\[f, g, h]"] sources_hcomp match_3 by (elim in_homE, auto) show "targets \[f, g, h] = targets f" using assms 1 targets_cod [of "\[f, g, h]"] targets_hcomp match_4 by (elim in_homE, auto) qed lemma assoc_simps\<^sub>A\<^sub>W\<^sub>C [simp]: assumes "ide f" and "ide g" and "ide h" and "f \ g \ null" and "g \ h \ null" shows "arr \[f, g, h]" and "dom \[f, g, h] = (f \ g) \ h" and "cod \[f, g, h] = f \ g \ h" proof - have 1: "\\[f, g, h] : (f \ g) \ h \ f \ g \ h\" using assms assoc_in_hom\<^sub>A\<^sub>W\<^sub>C by auto show "arr \[f, g, h]" using 1 by auto show "dom \[f, g, h] = (f \ g) \ h" using 1 by auto show "cod \[f, g, h] = f \ g \ h" using 1 by auto qed lemma assoc'_in_hom\<^sub>A\<^sub>W\<^sub>C: assumes "ide f" and "ide g" and "ide h" and "f \ g \ null" and "g \ h \ null" shows "sources (inv \[f, g, h]) = sources h" and "targets (inv \[f, g, h]) = targets f" and "\inv \[f, g, h] : f \ g \ h \ (f \ g) \ h\" proof - show 1: "\inv \[f, g, h] : f \ g \ h \ (f \ g) \ h\" using assms assoc_in_hom\<^sub>A\<^sub>W\<^sub>C iso_assoc\<^sub>A\<^sub>W\<^sub>C inv_in_hom by auto show "sources (inv \[f, g, h]) = sources h" using assms 1 sources_hcomp [of "f \ g" h] sources_cod match_3 null_agreement by (elim in_homE, metis) show "targets (inv \[f, g, h]) = targets f" using assms 1 targets_hcomp [of f "g \ h"] targets_dom match_4 null_agreement by (elim in_homE, metis) qed lemma assoc'_simps\<^sub>A\<^sub>W\<^sub>C [simp]: assumes "ide f" and "ide g" and "ide h" and "f \ g \ null" and "g \ h \ null" shows "arr (inv \[f, g, h])" and "dom (inv \[f, g, h]) = f \ g \ h" and "cod (inv \[f, g, h]) = (f \ g) \ h" proof - have 1: "\inv \[f, g, h] : f \ g \ h \ (f \ g) \ h\" using assms assoc'_in_hom\<^sub>A\<^sub>W\<^sub>C by auto show "arr (inv \[f, g, h])" using 1 by auto show "dom (inv \[f, g, h]) = f \ g \ h" using 1 by auto show "cod (inv \[f, g, h]) = (f \ g) \ h" using 1 by auto qed lemma assoc'_naturality\<^sub>A\<^sub>W\<^sub>C: assumes "\ \ \ \ null" and "\ \ \ \ null" shows "inv \[cod \, cod \, cod \] \ (\ \ \ \ \) = ((\ \ \) \ \) \ inv \[dom \, dom \, dom \]" proof - have \\\: "arr \ \ arr \ \ arr \" using assms composable_implies_arr by simp have 0: "dom \ \ dom \ \ null \ dom \ \ dom \ \ null \ cod \ \ cod \ \ null \ cod \ \ cod \ \ null" using assms \\\ hom_connected by simp have 1: "\\ \ \ \ \ : dom \ \ dom \ \ dom \ \ cod \ \ cod \ \ cod \\" using assms match_4 by auto have 2: "\(\ \ \) \ \ : (dom \ \ dom \) \ dom \ \ (cod \ \ cod \) \ cod \\" using assms match_3 by auto have "(inv \[cod \, cod \, cod \] \ (\ \ \ \ \)) \ \[dom \, dom \, dom \] = (\ \ \) \ \" proof - have "(\ \ \) \ \ = (inv \[cod \, cod \, cod \] \ \[cod \, cod \, cod \]) \ ((\ \ \) \ \)" using 0 2 \\\ assoc_in_hom\<^sub>A\<^sub>W\<^sub>C iso_assoc\<^sub>A\<^sub>W\<^sub>C comp_inv_arr inv_is_inverse comp_cod_arr by auto also have "... = inv \[cod \, cod \, cod \] \ \[cod \, cod \, cod \] \ ((\ \ \) \ \)" using comp_assoc by auto also have "... = inv \[cod \, cod \, cod \] \ (\ \ \ \ \) \ \[dom \, dom \, dom \]" using assms \\\ 0 2 assoc_naturality\<^sub>A\<^sub>W\<^sub>C by presburger also have "... = (inv \[cod \, cod \, cod \] \ (\ \ \ \ \)) \ \[dom \, dom \, dom \]" using comp_assoc by auto finally show ?thesis by argo qed thus ?thesis using 0 1 2 \\\ iso_assoc\<^sub>A\<^sub>W\<^sub>C assoc'_in_hom\<^sub>A\<^sub>W\<^sub>C inv_in_hom invert_side_of_triangle(2) by auto qed end subsection "Unitors" text \ For an associative weak composition with a chosen unit isomorphism \\ : a \ a \ a\, where \a\ is a weak unit, horizontal composition on the right by \a\ is a fully faithful endofunctor \R\ of the subcategory of arrows composable on the right with \a\, and is consequently an endo-equivalence of that subcategory. This equivalence, together with the associator isomorphisms and unit isomorphism \\\, canonically associate, with each identity arrow \f\ composable on the right with \a\, a \emph{right unit} isomorphism \\\[f] : f \ a \ f\\. These isomorphisms are the components of a natural isomorphism from \R\ to the identity functor. \ locale right_hom_with_unit = associative_weak_composition V H \ + right_hom V H a for V :: "'a comp" (infixr "\" 55) and H :: "'a comp" (infixr "\" 53) and \ :: "'a \ 'a \ 'a \ 'a" ("\[_, _, _]") and \ :: 'a and a :: 'a + assumes weak_unit_a: "weak_unit a" and \_in_hom: "\\ : a \ a \ a\" and iso_\: "iso \" begin abbreviation R where "R \ H\<^sub>R a" interpretation R: endofunctor S.comp R using weak_unit_a weak_unit_self_composable endofunctor_H\<^sub>R by simp interpretation R: fully_faithful_functor S.comp S.comp R using weak_unit_a weak_unit_def by simp lemma fully_faithful_functor_R: shows "fully_faithful_functor S.comp S.comp R" .. definition runit ("\[_]") where "runit f \ THE \. \\ : R f \\<^sub>S f\ \ R \ = (f \ \) \\<^sub>S \[f, a, a]" lemma iso_unit: shows "S.iso \" and "\\ : a \ a \\<^sub>S a\" proof - show "\\ : a \ a \\<^sub>S a\" proof - have a: "weak_unit a \ S.ide a" using weak_unit_a S.ide_char S.arr_char right_def weak_unit_self_composable by metis moreover have "S.arr (a \ a)" using a S.ideD(1) R.preserves_arr H\<^sub>R_def by auto ultimately show ?thesis using a S.in_hom_char S.arr_char right_def \_in_hom by (metis S.ideD(1) hom_connected(3) in_homE) qed thus "S.iso \" using iso_\ iso_char by blast qed lemma characteristic_iso: assumes "S.ide f" shows "\\[f, a, a] : (f \ a) \ a \\<^sub>S f \ a \ a\" and "\f \ \ : f \ a \ a \\<^sub>S f \ a\" and "\(f \ \) \\<^sub>S \[f, a, a] : R (R f) \\<^sub>S R f\" and "S.iso ((f \ \) \\<^sub>S \[f, a, a])" proof - have f: "S.ide f \ ide f" using assms S.ide_char by simp have a: "weak_unit a \ ide a \ S.ide a" using weak_unit_a S.ide_char weak_unit_def S.arr_char right_def weak_unit_self_composable by metis have fa: "f \ a \ null \ (f \ a) \ a \ null \ ((f \ a) \ a) \ a \ null" proof - have "S.arr (f \ a) \ S.arr ((f \ a) \ a) \ S.arr (((f \ a) \ a) \ a)" using assms S.ideD(1) R.preserves_arr H\<^sub>R_def by auto thus ?thesis using S.not_arr_null by fastforce qed have aa: "a \ a \ null" using a S.ideD(1) R.preserves_arr H\<^sub>R_def S.not_arr_null by auto have ia_a: "\ \ a \ null" using weak_unit_a hom_connected(3) weak_unit_self_composable \_in_hom by blast have f_ia: "f \ \ \ null" using assms S.ide_char right_def S.arr_char hom_connected(4) \_in_hom by auto show assoc_in_hom: "\\[f, a, a] : (f \ a) \ a \\<^sub>S f \ a \ a\" using a f fa hom_connected(1) [of "\[f, a, a]" a] S.arr_char right_def match_3 match_4 S.in_hom_char by auto show 1: "\f \ \ : f \ a \ a \\<^sub>S f \ a\" using a f fa iso_unit by (simp add: f_ia ide_in_hom) moreover have "S.iso (f \ \)" using a f fa f_ia 1 VoV.arr_char VxV.inv_simp inv_in_hom hom_connected(2) [of f "inv \"] VoV.arr_char VoV.iso_char preserves_iso iso_char iso_\ by auto ultimately have unit_part: "\f \ \ : f \ a \ a \\<^sub>S f \ a\ \ S.iso (f \ \)" by blast show "S.iso ((f \ \) \\<^sub>S \[f, a, a])" using assms a f fa aa hom_connected(1) [of "\[f, a, a]" a] right_def iso_assoc\<^sub>A\<^sub>W\<^sub>C iso_char S.arr_char unit_part assoc_in_hom isos_compose using S.isos_compose S.seqI' by auto show "\(f \ \) \\<^sub>S \[f, a, a] : R (R f) \\<^sub>S R f\" unfolding H\<^sub>R_def using unit_part assoc_in_hom by blast qed lemma runit_char: assumes "S.ide f" shows "\\[f] : R f \\<^sub>S f\" and "R \[f] = (f \ \) \\<^sub>S \[f, a, a]" and "\!\. \\ : R f \\<^sub>S f\ \ R \ = (f \ \) \\<^sub>S \[f, a, a]" proof - let ?P = "\\. \\ : R f \\<^sub>S f\ \ R \ = (f \ \) \\<^sub>S \[f, a, a]" show "\!\. ?P \" proof - have "\\. ?P \" proof - have 1: "S.ide f" using assms S.ide_char S.arr_char by simp moreover have "S.ide (R f)" using 1 R.preserves_ide by simp ultimately show ?thesis using assms characteristic_iso(3) R.is_full by blast qed moreover have "\\ \'. ?P \ \ ?P \' \ \ = \'" proof fix \ show "\\'. ?P \ \ ?P \' \ \ = \'" using R.is_faithful [of \] by fastforce qed ultimately show ?thesis by blast qed hence "?P (THE \. ?P \)" using theI' [of ?P] by fastforce hence 1: "?P \[f]" unfolding runit_def by blast show "\\[f] : R f \\<^sub>S f\" using 1 by fast show "R \[f] = (f \ \) \\<^sub>S \[f, a, a]" using 1 by fast qed lemma iso_runit: assumes "S.ide f" shows "S.iso \[f]" using assms characteristic_iso(4) runit_char R.reflects_iso by metis lemma runit_eqI: assumes "\f : a \\<^sub>S b\" and "\\ : R f \\<^sub>S f\" and "R \ = ((f \ \) \\<^sub>S \[f, a, a])" shows "\ = \[f]" proof - have "S.ide f" using assms(2) S.ide_cod by auto thus ?thesis using assms runit_char [of f] by auto qed lemma runit_naturality: assumes "S.arr \" shows "\[S.cod \] \\<^sub>S R \ = \ \\<^sub>S \[S.dom \]" proof - have 1: "\\[S.cod \] \\<^sub>S R \ : R (S.dom \) \\<^sub>S S.cod \\" using assms runit_char(1) S.ide_cod by blast have 2: "S.par (\[S.cod \] \\<^sub>S R \) (\ \\<^sub>S \[S.dom \])" proof - have "\\ \\<^sub>S \[S.dom \] : R (S.dom \) \\<^sub>S S.cod \\" using assms S.ide_dom runit_char(1) by blast thus ?thesis using 1 by (elim S.in_homE, auto) qed moreover have "R (\[S.cod \] \\<^sub>S R \) = R (\ \\<^sub>S \[S.dom \])" proof - have 3: "\\ \ a \ a : S.dom \ \ a \ a \\<^sub>S S.cod \ \ a \ a\" using assms weak_unit_a R.preserves_hom H\<^sub>R_def S.arr_iff_in_hom S.arr_char by (metis match_4 weak_unit_in_vhom weak_unit_self_right S.in_hom_char left_hcomp_closed S.not_arr_null S.null_char) have 4: "R (\[S.cod \] \\<^sub>S R \) = R \[S.cod \] \\<^sub>S R (R \)" using assms 1 R.preserves_comp_2 by blast also have 5: "... = ((S.cod \ \ \) \\<^sub>S \[S.cod \, a, a]) \\<^sub>S ((\ \ a) \ a)" using assms R.preserves_arr runit_char S.ide_cod H\<^sub>R_def by auto also have 6: "... = (S.cod \ \ \) \\<^sub>S \[S.cod \, a, a] \\<^sub>S ((\ \ a) \ a)" using assms S.comp_assoc by simp also have "... = (S.cod \ \ \) \\<^sub>S (\ \ a \ a) \\<^sub>S \[S.dom \, a, a]" proof - have "(\ \ a \ a) \\<^sub>S \[S.dom \, a, a] = \[S.cod \, a, a] \\<^sub>S ((\ \ a) \ a)" proof - have "(\ \ a \ a) \\<^sub>S \[S.dom \, a, a] = (\ \ a \ a) \ \[S.dom \, a, a]" using assms 3 S.ide_dom characteristic_iso(1) S.in_hom_char S.comp_char [of "\ \ a \ a" "\[S.dom \, a, a]"] by fastforce also have "... = \[S.cod \, a, a] \ ((\ \ a) \ a)" proof - have "\ \ a \ null" using assms S.arr_char right_def by simp thus ?thesis using assms weak_unit_a assoc_naturality\<^sub>A\<^sub>W\<^sub>C [of \ a a] by fastforce qed also have "... = \[S.cod \, a, a] \\<^sub>S ((\ \ a) \ a)" using S.in_hom_char S.comp_char by (metis 2 4 5 6 R.preserves_arr S.seq_char) finally show ?thesis by blast qed thus ?thesis by argo qed also have "... = ((S.cod \ \ \) \\<^sub>S (\ \ a \ a)) \\<^sub>S \[S.dom \, a, a]" using S.comp_assoc by auto also have "... = ((\ \ a) \\<^sub>S (S.dom \ \ \)) \\<^sub>S \[S.dom \, a, a]" proof - have "\ \ a \ a \ null" using 3 S.not_arr_null by (elim S.in_homE, auto) moreover have "S.dom \ \ \ \ null" using assms S.not_arr_null by (metis S.dom_char \_in_hom calculation hom_connected(1-2) in_homE) ultimately have "(S.cod \ \ \) \\<^sub>S (\ \ a \ a) = (\ \ a) \\<^sub>S (S.dom \ \ \)" using assms weak_unit_a iso_unit S.comp_arr_dom S.comp_cod_arr interchange [of "S.cod \" \ \ "a \ a"] interchange [of \ "S.dom \" a \] by auto thus ?thesis by argo qed also have "... = (\ \ a) \\<^sub>S (S.dom \ \ \) \\<^sub>S \[S.dom \, a, a]" using S.comp_assoc by auto also have "... = R \ \\<^sub>S R \[S.dom \]" using assms runit_char(2) S.ide_dom H\<^sub>R_def by auto also have "... = R (\ \\<^sub>S \[S.dom \])" using assms S.arr_iff_in_hom [of \] runit_char(1) S.ide_dom by fastforce finally show ?thesis by blast qed ultimately show "\[S.cod \] \\<^sub>S (R \) = \ \\<^sub>S \[S.dom \]" using R.is_faithful by blast qed abbreviation \ where "\ \ \ if S.arr \ then \ \\<^sub>S \[S.dom \] else null" interpretation \: natural_transformation S.comp S.comp R S.map \ proof - interpret \: transformation_by_components S.comp S.comp R S.map runit using runit_char(1) runit_naturality by (unfold_locales, simp_all) have "\.map = \" using \.is_extensional \.map_def \.naturality \.map_simp_ide S.ide_dom S.ide_cod S.map_def by auto thus "natural_transformation S.comp S.comp R S.map \" using \.natural_transformation_axioms by auto qed lemma natural_transformation_\: shows "natural_transformation S.comp S.comp R S.map \" .. interpretation \: natural_isomorphism S.comp S.comp R S.map \ using S.ide_is_iso iso_runit runit_char(1) S.isos_compose by (unfold_locales, force) lemma natural_isomorphism_\: shows "natural_isomorphism S.comp S.comp R S.map \" .. interpretation R: equivalence_functor S.comp S.comp R using natural_isomorphism_\ R.isomorphic_to_identity_is_equivalence by blast lemma equivalence_functor_R: shows "equivalence_functor S.comp S.comp R" .. lemma runit_commutes_with_R: assumes "S.ide f" shows "\[R f] = R \[f]" proof - have "S.seq \[f] (R \[f])" using assms runit_char(1) R.preserves_hom [of "\[f]" "R f" f] by fastforce moreover have "S.seq \[f] \[R f]" using assms runit_char(1) [of f] runit_char(1) [of "R f"] by auto ultimately show ?thesis using assms runit_char(1) runit_naturality [of "\[f]"] iso_runit S.iso_is_section S.section_is_mono S.monoE [of "\[f]" "R \[f]" "\[R f]"] by force qed end text \ Symmetric results hold for the subcategory of all arrows composable on the left with a specified weak unit \b\. This yields the \emph{left unitors}. \ locale left_hom_with_unit = associative_weak_composition V H \ + left_hom V H b for V :: "'a comp" (infixr "\" 55) and H :: "'a comp" (infixr "\" 53) and \ :: "'a \ 'a \ 'a \ 'a" ("\[_, _, _]") and \ :: 'a and b :: 'a + assumes weak_unit_b: "weak_unit b" and \_in_hom: "\\ : b \ b \ b\" and iso_\: "iso \" begin abbreviation L where "L \ H\<^sub>L b" interpretation L: endofunctor S.comp L using weak_unit_b weak_unit_self_composable endofunctor_H\<^sub>L by simp interpretation L: fully_faithful_functor S.comp S.comp L using weak_unit_b weak_unit_def by simp lemma fully_faithful_functor_L: shows "fully_faithful_functor S.comp S.comp L" .. definition lunit ("\[_]") where "lunit f \ THE \. \\ : L f \\<^sub>S f\ \ L \ = (\ \ f) \\<^sub>S (inv \[b, b, f])" lemma iso_unit: shows "S.iso \" and "\\ : b \ b \\<^sub>S b\" proof - show "\\ : b \ b \\<^sub>S b\" proof - have b: "weak_unit b \ S.ide b" using weak_unit_b S.ide_char S.arr_char left_def weak_unit_self_composable by metis moreover have "S.arr (b \ b)" using b S.ideD(1) L.preserves_arr H\<^sub>L_def by auto ultimately show ?thesis using b S.in_hom_char S.arr_char left_def \_in_hom by (metis S.ideD(1) hom_connected(4) in_homE) qed thus "S.iso \" using iso_\ iso_char by blast qed lemma characteristic_iso: assumes "S.ide f" shows "\inv \[b, b, f] : b \ b \ f \\<^sub>S (b \ b) \ f\" and "\\ \ f : (b \ b) \ f \\<^sub>S b \ f\" and "\(\ \ f) \\<^sub>S inv \[b, b, f] : L (L f) \\<^sub>S L f\" and "S.iso ((\ \ f) \\<^sub>S inv \[b, b, f])" proof - have f: "S.ide f \ ide f" using assms S.ide_char by simp have b: "weak_unit b \ ide b \ S.ide b" using weak_unit_b S.ide_char weak_unit_def S.arr_char left_def weak_unit_self_composable by metis have bf: "b \ f \ null \ b \ b \ b \ f \ null" proof - have "S.arr (b \ f) \ S.arr (b \ b \ f) \ S.arr (b \ b \ b \ f)" using assms S.ideD(1) L.preserves_arr H\<^sub>L_def by auto thus ?thesis using S.not_arr_null by fastforce qed have bb: "b \ b \ null" proof - have "S.arr (b \ b)" using b S.ideD(1) L.preserves_arr H\<^sub>L_def by auto thus ?thesis using S.not_arr_null by fastforce qed have b_ib: "b \ \ \ null" using weak_unit_b hom_connected(4) weak_unit_self_composable \_in_hom by blast have ib_f: "\ \ f \ null" using assms S.ide_char left_def S.arr_char hom_connected(3) \_in_hom by auto show assoc_in_hom: "\inv \[b, b, f] : b \ b \ f \\<^sub>S (b \ b) \ f\" using b f bf bb hom_connected(2) [of b "inv \[b, b, f]"] left_def by (metis S.arrI S.cod_closed S.in_hom_char assoc'_in_hom\<^sub>A\<^sub>W\<^sub>C(3) assoc'_simps\<^sub>A\<^sub>W\<^sub>C(2-3)) show 1: "\\ \ f : (b \ b) \ f \\<^sub>S b \ f\" using b f bf by (simp add: ib_f ide_in_hom iso_unit(2)) moreover have "S.iso (\ \ f)" using b f bf ib_f 1 VoV.arr_char VxV.inv_simp inv_in_hom hom_connected(1) [of "inv \" f] VoV.arr_char VoV.iso_char preserves_iso iso_char iso_\ by auto ultimately have unit_part: "\\ \ f : (b \ b) \ f \\<^sub>S b \ f\ \ S.iso (\ \ f)" by blast show "S.iso ((\ \ f) \\<^sub>S inv \[b, b, f])" proof - have "S.iso (inv \[b, b, f])" using assms b f bf bb hom_connected(2) [of b "inv \[b, b, f]"] left_def iso_assoc\<^sub>A\<^sub>W\<^sub>C iso_inv_iso iso_char S.arr_char left_def by simp thus ?thesis using unit_part assoc_in_hom isos_compose by blast qed show "\(\ \ f) \\<^sub>S inv \[b, b, f] : L (L f) \\<^sub>S L f\" unfolding H\<^sub>L_def using unit_part assoc_in_hom by blast qed lemma lunit_char: assumes "S.ide f" shows "\\[f] : L f \\<^sub>S f\" and "L \[f] = (\ \ f) \\<^sub>S inv \[b, b, f]" and "\!\. \\ : L f \\<^sub>S f\ \ L \ = (\ \ f) \\<^sub>S inv \[b, b, f]" proof - let ?P = "\\. \\ : L f \\<^sub>S f\ \ L \ = (\ \ f) \\<^sub>S inv \[b, b, f]" show "\!\. ?P \" proof - have "\\. ?P \" proof - have 1: "S.ide f" using assms S.ide_char S.arr_char by simp moreover have "S.ide (L f)" using 1 L.preserves_ide by simp ultimately show ?thesis using assms characteristic_iso(3) L.is_full by blast qed moreover have "\\ \'. ?P \ \ ?P \' \ \ = \'" proof fix \ show "\\'. ?P \ \ ?P \' \ \ = \'" using L.is_faithful [of \] by fastforce qed ultimately show ?thesis by blast qed hence "?P (THE \. ?P \)" using theI' [of ?P] by fastforce hence 1: "?P \[f]" unfolding lunit_def by blast show "\\[f] : L f \\<^sub>S f\" using 1 by fast show "L \[f] = (\ \ f) \\<^sub>S inv \[b, b, f]" using 1 by fast qed lemma iso_lunit: assumes "S.ide f" shows "S.iso \[f]" using assms characteristic_iso(4) lunit_char L.reflects_iso by metis lemma lunit_eqI: assumes "\f : a \\<^sub>S b\" and "\\ : L f \\<^sub>S f\" and "L \ = ((\ \ f) \\<^sub>S inv \[b, b, f])" shows "\ = \[f]" proof - have "S.ide f" using assms(2) S.ide_cod by auto thus ?thesis using assms lunit_char [of f] by auto qed lemma lunit_naturality: assumes "S.arr \" shows "\[S.cod \] \\<^sub>S L \ = \ \\<^sub>S \[S.dom \]" proof - have 1: "\\[S.cod \] \\<^sub>S L \ : L (S.dom \) \\<^sub>S S.cod \\" using assms lunit_char(1) [of "S.cod \"] S.ide_cod by blast have "S.par (\[S.cod \] \\<^sub>S L \) (\ \\<^sub>S \[S.dom \])" proof - have "\\ \\<^sub>S \[S.dom \] : L (S.dom \) \\<^sub>S S.cod \\" using assms S.ide_dom lunit_char(1) by blast thus ?thesis using 1 by (elim S.in_homE, auto) qed moreover have "L (\[S.cod \] \\<^sub>S L \) = L (\ \\<^sub>S \[S.dom \])" proof - have 2: "\b \ b \ \ : b \ b \ S.dom \ \\<^sub>S b \ b \ S.cod \\" using assms weak_unit_b L.preserves_hom H\<^sub>L_def S.arr_iff_in_hom [of \] S.arr_char by simp have 3: "\(b \ b) \ \ : (b \ b) \ S.dom \ \\<^sub>S (b \ b) \ S.cod \\" using assms weak_unit_b L.preserves_hom H\<^sub>L_def S.arr_iff_in_hom S.arr_char by (metis match_3 weak_unit_in_vhom weak_unit_self_left S.in_hom_char S.not_arr_null S.null_char right_hcomp_closed) have "L (\[S.cod \] \\<^sub>S L \) = L \[S.cod \] \\<^sub>S L (L \)" using assms 1 L.preserves_comp_2 by blast also have "... = ((\ \ S.cod \) \\<^sub>S inv \[b, b, S.cod \]) \\<^sub>S (b \ b \ \)" using assms L.preserves_arr lunit_char S.ide_cod H\<^sub>L_def by auto also have "... = (\ \ S.cod \) \\<^sub>S inv \[b, b, S.cod \] \\<^sub>S (b \ b \ \)" using S.comp_assoc by auto also have "... = (\ \ S.cod \) \\<^sub>S ((b \ b) \ \) \\<^sub>S inv \[b, b, S.dom \]" proof - have "inv \[b, b, S.cod \] \\<^sub>S (b \ b \ \) = ((b \ b) \ \) \\<^sub>S inv \[b, b, S.dom \]" proof - have "((b \ b) \ \) \\<^sub>S inv \[b, b, S.dom \] = ((b \ b) \ \) \ inv \[b, b, S.dom \]" using assms 3 S.in_hom_char S.comp_char [of "(b \ b) \ \" "inv \[b, b, S.dom \]"] by (metis S.ide_dom characteristic_iso(1) ext) also have "... = inv \[b, b, S.cod \] \ (b \ b \ \)" proof - have "b \ \ \ null" using assms S.arr_char left_def by simp thus ?thesis using assms weak_unit_b assoc'_naturality\<^sub>A\<^sub>W\<^sub>C [of b b \] by fastforce qed also have "... = inv \[b, b, S.cod \] \\<^sub>S (b \ b \ \)" using assms 2 S.in_hom_char S.comp_char by (metis S.comp_simp S.ide_cod S.seqI' characteristic_iso(1)) finally show ?thesis by argo qed thus ?thesis by argo qed also have "... = ((\ \ S.cod \) \\<^sub>S ((b \ b) \ \)) \\<^sub>S inv \[b, b, S.dom \]" using S.comp_assoc by auto also have "... = ((b \ \) \\<^sub>S (\ \ S.dom \)) \\<^sub>S inv \[b, b, S.dom \]" proof - have "(b \ b) \ \ \ null" using 3 S.not_arr_null by (elim S.in_homE, auto) moreover have "\ \ S.dom \ \ null" using assms S.not_arr_null by (metis S.dom_char \_in_hom calculation hom_connected(1-2) in_homE) ultimately have "(\ \ S.cod \) \\<^sub>S ((b \ b) \ \) = (b \ \) \\<^sub>S (\ \ S.dom \)" using assms weak_unit_b iso_unit S.comp_arr_dom S.comp_cod_arr interchange [of \ "b \ b" "S.cod \" \ ] interchange [of b \ \ "S.dom \"] by auto thus ?thesis by argo qed also have "... = (b \ \) \\<^sub>S (\ \ S.dom \) \\<^sub>S inv \[b, b, S.dom \]" using S.comp_assoc by auto also have "... = L \ \\<^sub>S L \[S.dom \]" using assms lunit_char(2) S.ide_dom H\<^sub>L_def by auto also have "... = L (\ \\<^sub>S \[S.dom \])" using assms S.arr_iff_in_hom [of \] lunit_char(1) S.ide_dom S.seqI by fastforce finally show ?thesis by blast qed ultimately show "\[S.cod \] \\<^sub>S L \ = \ \\<^sub>S \[S.dom \]" using L.is_faithful by blast qed abbreviation \ where "\ \ \ if S.arr \ then \ \\<^sub>S \[S.dom \] else null" interpretation \: natural_transformation S.comp S.comp L S.map \ proof - interpret \: transformation_by_components S.comp S.comp L S.map lunit using lunit_char(1) lunit_naturality by (unfold_locales, simp_all) have "\.map = \" using \.is_extensional \.map_def \.naturality \.map_simp_ide S.ide_dom S.ide_cod S.map_def by auto thus "natural_transformation S.comp S.comp L S.map \" using \.natural_transformation_axioms by auto qed lemma natural_transformation_\: shows "natural_transformation S.comp S.comp L S.map \" .. interpretation \: natural_isomorphism S.comp S.comp L S.map \ using S.ide_is_iso iso_lunit lunit_char(1) S.isos_compose by (unfold_locales, force) lemma natural_isomorphism_\: shows "natural_isomorphism S.comp S.comp L S.map \" .. interpretation L: equivalence_functor S.comp S.comp L using natural_isomorphism_\ L.isomorphic_to_identity_is_equivalence by blast lemma equivalence_functor_L: shows "equivalence_functor S.comp S.comp L" .. lemma lunit_commutes_with_L: assumes "S.ide f" shows "\[L f] = L \[f]" proof - have "S.seq \[f] (L \[f])" using assms lunit_char(1) L.preserves_hom [of "\[f]" "L f" f] by fastforce moreover have "S.seq \[f] \[L f]" using assms lunit_char(1) [of f] lunit_char(1) [of "L f"] by auto ultimately show ?thesis using assms lunit_char(1) lunit_naturality [of "\[f]"] iso_lunit S.iso_is_section S.section_is_mono S.monoE [of "\[f]" "L \[f]" "\[L f]"] by force qed end subsection "Prebicategories" text \ A \emph{prebicategory} is an associative weak composition satisfying the additional assumption that every arrow has a source and a target. \ locale prebicategory = associative_weak_composition + assumes arr_has_source: "arr \ \ sources \ \ {}" and arr_has_target: "arr \ \ targets \ \ {}" begin lemma arr_iff_has_src: shows "arr \ \ sources \ \ {}" using arr_has_source composable_implies_arr by auto lemma arr_iff_has_trg: shows "arr \ \ targets \ \ {}" using arr_has_target composable_implies_arr by auto end text \ The horizontal composition of a prebicategory is regular. \ sublocale prebicategory \ regular_weak_composition V H proof show "\a f. a \ sources f \ ide f \ f \ a \ f" proof - fix a f assume a: "a \ sources f" and f: "ide f" interpret Right_a: subcategory V \right a\ using a right_hom_is_subcategory weak_unit_self_composable by force interpret Right_a: right_hom_with_unit V H \ \some_unit a\ a using a iso_some_unit by (unfold_locales, auto) show "f \ a \ f" proof - have "Right_a.ide f" using a f Right_a.ide_char Right_a.arr_char right_def by auto hence "Right_a.iso (Right_a.runit f) \ (Right_a.runit f) \ Right_a.hom (f \ a) f" using Right_a.iso_runit Right_a.runit_char(1) H\<^sub>R_def by simp hence "iso (Right_a.runit f) \ (Right_a.runit f) \ hom (f \ a) f" using Right_a.iso_char Right_a.hom_char by auto thus ?thesis using f isomorphic_def by auto qed qed show "\b f. b \ targets f \ ide f \ b \ f \ f" proof - fix b f assume b: "b \ targets f" and f: "ide f" interpret Left_b: subcategory V \left b\ using b left_hom_is_subcategory weak_unit_self_composable by force interpret Left_b: left_hom_with_unit V H \ \some_unit b\ b using b iso_some_unit by (unfold_locales, auto) show "b \ f \ f" proof - have "Left_b.ide f" using b f Left_b.ide_char Left_b.arr_char left_def by auto hence "Left_b.iso (Left_b.lunit f) \ (Left_b.lunit f) \ Left_b.hom (b \ f) f" using b f Left_b.iso_lunit Left_b.lunit_char(1) H\<^sub>L_def by simp hence "iso (Left_b.lunit f) \ (Left_b.lunit f) \ hom (b \ f) f" using Left_b.iso_char Left_b.hom_char by auto thus ?thesis using isomorphic_def by auto qed qed qed text \ The regularity allows us to show that, in a prebicategory, all sources of a given arrow are isomorphic, and similarly for targets. \ context prebicategory begin lemma sources_are_isomorphic: assumes "a \ sources \" and "a' \ sources \" shows "a \ a'" proof - have \: "arr \" using assms composable_implies_arr by auto have 0: "\f. \ ide f; a \ sources f; a' \ sources f \ \ a \ a'" proof - fix f assume f: "ide f" and a: "a \ sources f" and a': "a' \ sources f" have 1: "a \ a' \ null" using a a' f \ assms(1) sources_determine_composability sourcesD(2-3) by meson have 2: "a \ targets a' \ a' \ sources a" using assms 1 by blast show "a \ a'" using a a' 1 2 comp_ide_source comp_target_ide [of a a'] weak_unit_self_composable(1) [of a] weak_unit_self_composable(1) [of a'] isomorphic_transitive isomorphic_symmetric by blast qed have "ide (dom \) \ a \ sources (dom \) \ a' \ sources (dom \)" using assms \ sources_dom by auto thus ?thesis using 0 by auto qed lemma targets_are_isomorphic: assumes "b \ targets \" and "b' \ targets \" shows "b \ b'" proof - have \: "arr \" using assms composable_implies_arr by auto have 0: "\f. \ ide f; b \ targets f; b' \ targets f \ \ b \ b'" proof - fix f assume f: "ide f" and b: "b \ targets f" and b': "b' \ targets f" have 1: "b \ b' \ null" using b b' f \ assms(1) targets_determine_composability targetsD(2-3) by meson have 2: "b \ targets b' \ b' \ sources b" using assms 1 by blast show "b \ b'" using b b' 1 2 comp_ide_source comp_target_ide [of b b'] weak_unit_self_composable(1) [of b] weak_unit_self_composable(1) [of b'] isomorphic_transitive isomorphic_symmetric by blast qed have "ide (dom \) \ b \ targets (dom \) \ b' \ targets (dom \)" using assms \ targets_dom [of \] by auto thus ?thesis using 0 by auto qed text \ In fact, we now show that the sets of sources and targets of a 2-cell are isomorphism-closed, and hence are isomorphism classes. We first show that the notion ``weak unit'' is preserved under isomorphism. \ interpretation H: partial_magma H using is_partial_magma by auto lemma isomorphism_respects_weak_units: assumes "weak_unit a" and "a \ a'" shows "weak_unit a'" proof - obtain \ where \: "iso \ \ \\ : a \ a'\" using assms by auto interpret Left_a: subcategory V \left a\ using assms left_hom_is_subcategory by fastforce interpret Left_a: left_hom_with_unit V H \ \some_unit a\ a using assms iso_some_unit apply unfold_locales by auto interpret Right_a: subcategory V "right a" using assms right_hom_is_subcategory by fastforce interpret Right_a: right_hom_with_unit V H \ \some_unit a\ a using assms iso_some_unit apply unfold_locales by auto have a': "ide a' \ a \ a' \ null \ a' \ a \ null \ a' \ a' \ null \ \ \ a' \ null \ Left_a.ide a'" using assms \ weak_unit_self_composable hom_connected Left_a.ide_char Left_a.arr_char left_def apply auto apply (meson weak_unit_self_composable(3) isomorphic_implies_equicomposable) apply (meson weak_unit_self_composable(3) isomorphic_implies_equicomposable) apply (meson weak_unit_self_composable(3) isomorphic_implies_equicomposable) apply (metis weak_unit_self_composable(3) in_homE) by (meson weak_unit_self_composable(3) isomorphic_implies_equicomposable) have iso: "a' \ a' \ a'" proof - have 1: "Right a' = Right a" using assms right_respects_isomorphic by simp interpret Right_a': subcategory V \right a'\ using assms right_hom_is_subcategory by fastforce (* TODO: The previous interpretation brings in unwanted notation for in_hom. *) interpret Ra': endofunctor \Right a'\ \H\<^sub>R a'\ using assms a' endofunctor_H\<^sub>R by auto let ?\ = "Left_a.lunit a' \ inv (\ \ a')" have "iso ?\ \ \?\ : a' \ a' \ a'\" proof - have "iso (Left_a.lunit a') \ \Left_a.lunit a' : a \ a' \ a'\" using a' Left_a.lunit_char(1) Left_a.iso_lunit Left_a.iso_char Left_a.in_hom_char H\<^sub>L_def by auto moreover have "iso (\ \ a') \ \\ \ a' : a \ a' \ a' \ a'\" proof - have 1: "Right_a'.iso \ \ \ \ Right_a'.hom (Right_a'.dom \) (Right_a'.cod \)" using a' \ Right_a'.iso_char Right_a'.arr_char right_def right_iff_right_inv Right_a'.arr_iff_in_hom [of \] by simp have "Right_a'.iso (H\<^sub>R a' \) \ Right_a'.in_hom (H\<^sub>R a' \) (H\<^sub>R a' (Right_a'.dom \)) (H\<^sub>R a' (Right_a'.cod \))" using \ 1 Ra'.preserves_iso Ra'.preserves_hom Right_a'.iso_char Ra'.preserves_dom Ra'.preserves_cod Right_a'.arr_iff_in_hom [of "H\<^sub>R a' \"] by simp thus ?thesis using \ 1 Right_a'.in_hom_char Right_a'.iso_char H\<^sub>R_def by auto qed ultimately show ?thesis using isos_compose iso_inv_iso inv_in_hom by blast qed thus ?thesis using isomorphic_def by auto qed text \ We show that horizontal composition on the left and right by @{term a'} is naturally isomorphic to the identity functor. This follows from the fact that if @{term a} is isomorphic to @{term a'}, then horizontal composition with @{term a} is naturally isomorphic to horizontal composition with @{term a'}, hence the latter is naturally isomorphic to the identity if the former is. This is conceptually simple, but there are tedious composability details to handle. \ have 1: "Left a' = Left a \ Right a' = Right a" using assms left_respects_isomorphic right_respects_isomorphic by simp interpret L: fully_faithful_functor \Left a\ \Left a\ \H\<^sub>L a\ using assms weak_unit_def by simp interpret L': endofunctor \Left a\ \H\<^sub>L a'\ using a' 1 endofunctor_H\<^sub>L [of a'] by auto interpret \: natural_isomorphism \Left a\ \Left a\ \H\<^sub>L a\ \H\<^sub>L a'\ \H\<^sub>L \\ proof fix \ show "\ Left_a.arr \ \ H\<^sub>L \ \ = Left_a.null" using left_def \ H\<^sub>L_def hom_connected(1) Left_a.null_char null_agreement Left_a.arr_char by auto assume "Left_a.arr \" hence \: "Left_a.arr \ \ arr \ \ a \ \ \ null" using Left_a.arr_char left_def composable_implies_arr by simp have 2: "\ \ \ \ null" using assms \ \ Left_a.arr_char left_def hom_connected by auto show "Left_a.dom (H\<^sub>L \ \) = H\<^sub>L a (Left_a.dom \)" using assms 2 \ \ Left_a.arr_char left_def hom_connected(2) [of a \] weak_unit_self_composable match_4 Left_a.dom_char H\<^sub>L_def by auto show "Left_a.cod (H\<^sub>L \ \) = H\<^sub>L a' (Left_a.cod \)" using assms 2 \ \ Left_a.arr_char left_def hom_connected(2) [of a \] weak_unit_self_composable match_4 Left_a.cod_char H\<^sub>L_def by auto show "Left_a.comp (H\<^sub>L a' \) (H\<^sub>L \ (Left_a.dom \)) = H\<^sub>L \ \" proof - have "Left_a.comp (H\<^sub>L a' \) (H\<^sub>L \ (Left_a.dom \)) = Left_a.comp (a' \ \) (\ \ dom \)" using assms 1 2 \ \ Left_a.dom_char left_def H\<^sub>L_def by simp also have "... = (a' \ \) \ (\ \ dom \)" proof - have "Left_a.seq (a' \ \) (\ \ dom \)" proof (intro Left_a.seqI) show 3: "Left_a.arr (\ \ dom \)" using assms 2 \ \ Left_a.arr_char left_def by (metis H\<^sub>L_def L'.preserves_arr hcomp_simps\<^sub>W\<^sub>C(1) in_homE right_connected paste_1) show 4: "Left_a.arr (a' \ \)" using \ H\<^sub>L_def L'.preserves_arr by auto show "Left_a.dom (a' \ \) = Left_a.cod (\ \ dom \)" using a' \ \ 2 3 4 Left_a.dom_char Left_a.cod_char by (metis Left_a.seqE Left_a.seq_char hcomp_simps\<^sub>W\<^sub>C(1) in_homE paste_1) qed thus ?thesis using Left_a.comp_char Left_a.arr_char left_def by auto qed also have "... = a' \ \ \ \ \ dom \" using a' \ \ interchange hom_connected by auto also have "... = \ \ \" using \ \ comp_arr_dom comp_cod_arr by auto finally show ?thesis using H\<^sub>L_def by simp qed show "Left_a.comp (H\<^sub>L \ (Left_a.cod \)) (Left_a.L \) = H\<^sub>L \ \" proof - have "Left_a.comp (H\<^sub>L \ (Left_a.cod \)) (Left_a.L \) = Left_a.comp (\ \ cod \) (a \ \)" using assms 1 2 \ \ Left_a.cod_char left_def H\<^sub>L_def by simp also have "... = (\ \ cod \) \ (a \ \)" proof - have "Left_a.seq (\ \ cod \) (a \ \)" proof (intro Left_a.seqI) show 3: "Left_a.arr (\ \ cod \)" using \ \ 2 Left_a.arr_char left_def by (metis (no_types, lifting) H\<^sub>L_def L.preserves_arr hcomp_simps\<^sub>W\<^sub>C(1) in_homE right_connected paste_2) show 4: "Left_a.arr (a \ \)" using assms \ Left_a.arr_char left_def using H\<^sub>L_def L.preserves_arr by auto show "Left_a.dom (\ \ cod \) = Left_a.cod (a \ \)" using assms \ \ 2 3 4 Left_a.dom_char Left_a.cod_char by (metis Left_a.seqE Left_a.seq_char hcomp_simps\<^sub>W\<^sub>C(1) in_homE paste_2) qed thus ?thesis using Left_a.comp_char Left_a.arr_char left_def by auto qed also have "... = \ \ a \ cod \ \ \" using \ \ interchange hom_connected by auto also have "... = \ \ \" using \ \ comp_arr_dom comp_cod_arr by auto finally show ?thesis using H\<^sub>L_def by simp qed next fix \ assume \: "Left_a.ide \" have 1: "\ \ \ \ null" using assms \ \ Left_a.ide_char Left_a.arr_char left_def hom_connected by auto show "Left_a.iso (H\<^sub>L \ \)" proof - have "iso (\ \ \)" proof - have "a \ sources \ \ targets \" using assms \ \ 1 hom_connected weak_unit_self_composable Left_a.ide_char Left_a.arr_char left_def connected_if_composable by auto thus ?thesis using \ \ Left_a.ide_char ide_is_iso iso_hcomp\<^sub>R\<^sub>W\<^sub>C(1) by blast qed moreover have "left a (\ \ \)" using assms 1 \ weak_unit_self_composable hom_connected(2) [of a \] left_def match_4 null_agreement by auto ultimately show ?thesis using Left_a.iso_char Left_a.arr_char left_iff_left_inv Left_a.inv_char H\<^sub>L_def by simp qed qed interpret L': equivalence_functor \Left a'\ \Left a'\ \H\<^sub>L a'\ proof - have "naturally_isomorphic (Left a) (Left a) (H\<^sub>L a) Left_a.map" using assms Left_a.natural_isomorphism_\ naturally_isomorphic_def by blast moreover have "naturally_isomorphic (Left a) (Left a) (H\<^sub>L a) (H\<^sub>L a')" using naturally_isomorphic_def \.natural_isomorphism_axioms by blast ultimately have "naturally_isomorphic (Left a) (Left a) (H\<^sub>L a') (identity_functor.map (Left a))" using naturally_isomorphic_symmetric naturally_isomorphic_transitive by fast hence "naturally_isomorphic (Left a') (Left a') (H\<^sub>L a') (identity_functor.map (Left a'))" using 1 by auto thus "equivalence_functor (Left a') (Left a') (H\<^sub>L a')" using 1 L'.isomorphic_to_identity_is_equivalence naturally_isomorphic_def by fastforce qed text \ Now we do the same for \R'\. \ interpret R: fully_faithful_functor \Right a\ \Right a\ \H\<^sub>R a\ using assms weak_unit_def by simp interpret R': endofunctor \Right a\ \H\<^sub>R a'\ using a' 1 endofunctor_H\<^sub>R [of a'] by auto interpret \: natural_isomorphism \Right a\ \Right a\ \H\<^sub>R a\ \H\<^sub>R a'\ \H\<^sub>R \\ proof fix \ show "\ Right_a.arr \ \ H\<^sub>R \ \ = Right_a.null" using right_def \ H\<^sub>R_def hom_connected Right_a.null_char Right_a.arr_char by auto assume "Right_a.arr \" hence \: "Right_a.arr \ \ arr \ \ \ \ a \ null" using Right_a.arr_char right_def composable_implies_arr by simp have 2: "\ \ \ \ null" using assms \ \ Right_a.arr_char right_def hom_connected by auto show "Right_a.dom (H\<^sub>R \ \) = H\<^sub>R a (Right_a.dom \)" using assms 2 \ \ Right_a.arr_char right_def hom_connected(1) [of \ a] weak_unit_self_composable match_3 Right_a.dom_char H\<^sub>R_def by auto show "Right_a.cod (H\<^sub>R \ \) = H\<^sub>R a' (Right_a.cod \)" using assms 2 a' \ \ Right_a.arr_char right_def hom_connected(3) [of \ a] weak_unit_self_composable match_3 Right_a.cod_char H\<^sub>R_def by auto show "Right_a.comp (H\<^sub>R a' \) (H\<^sub>R \ (Right_a.dom \)) = H\<^sub>R \ \" proof - have "Right_a.comp (H\<^sub>R a' \) (H\<^sub>R \ (Right_a.dom \)) = Right_a.comp (\ \ a') (dom \ \ \)" using assms 1 2 \ \ Right_a.dom_char right_def H\<^sub>R_def by simp also have "... = (\ \ a') \ (dom \ \ \)" proof - have "Right_a.seq (\ \ a') (dom \ \ \)" proof (intro Right_a.seqI) show 3: "Right_a.arr (dom \ \ \)" using assms 2 \ \ Right_a.arr_char right_def by (metis H\<^sub>R_def R'.preserves_arr hcomp_simps\<^sub>W\<^sub>C(1) in_homE left_connected paste_2) show 4: "Right_a.arr (\ \ a')" using \ H\<^sub>R_def R'.preserves_arr by auto show "Right_a.dom (\ \ a') = Right_a.cod (dom \ \ \)" using a' \ \ 2 3 4 Right_a.dom_char Right_a.cod_char by (metis Right_a.seqE Right_a.seq_char hcomp_simps\<^sub>W\<^sub>C(1) in_homE paste_2) qed thus ?thesis using Right_a.comp_char Right_a.arr_char right_def by auto qed also have "... = \ \ dom \ \ a' \ \" using a' \ \ interchange hom_connected by auto also have "... = \ \ \" using \ \ comp_arr_dom comp_cod_arr by auto finally show ?thesis using H\<^sub>R_def by simp qed show "Right_a.comp (H\<^sub>R \ (Right_a.cod \)) (Right_a.R \) = H\<^sub>R \ \" proof - have "Right_a.comp (H\<^sub>R \ (Right_a.cod \)) (Right_a.R \) = Right_a.comp (cod \ \ \) (\ \ a)" using assms 1 2 \ \ Right_a.cod_char right_def H\<^sub>R_def by simp also have "... = (cod \ \ \) \ (\ \ a)" proof - have "Right_a.seq (cod \ \ \) (\ \ a)" proof (intro Right_a.seqI) show 3: "Right_a.arr (cod \ \ \)" using \ \ 2 Right_a.arr_char right_def by (metis (no_types, lifting) H\<^sub>R_def R.preserves_arr hcomp_simps\<^sub>W\<^sub>C(1) in_homE left_connected paste_1) show 4: "Right_a.arr (\ \ a)" using assms \ Right_a.arr_char right_def using H\<^sub>R_def R.preserves_arr by auto show "Right_a.dom (cod \ \ \) = Right_a.cod (\ \ a)" using assms \ \ 2 3 4 Right_a.dom_char Right_a.cod_char by (metis Right_a.seqE Right_a.seq_char hcomp_simps\<^sub>W\<^sub>C(1) in_homE paste_1) qed thus ?thesis using Right_a.comp_char Right_a.arr_char right_def by auto qed also have "... = cod \ \ \ \ \ \ a" using \ \ interchange hom_connected by auto also have "... = \ \ \" using \ \ comp_arr_dom comp_cod_arr by auto finally show ?thesis using H\<^sub>R_def by simp qed next fix \ assume \: "Right_a.ide \" have 1: "\ \ \ \ null" using assms \ \ Right_a.ide_char Right_a.arr_char right_def hom_connected by auto show "Right_a.iso (H\<^sub>R \ \)" proof - have "iso (\ \ \)" proof - have "a \ targets \ \ sources \" using assms \ \ 1 hom_connected weak_unit_self_composable Right_a.ide_char Right_a.arr_char right_def connected_if_composable by (metis (full_types) IntI targetsI) thus ?thesis using \ \ Right_a.ide_char ide_is_iso iso_hcomp\<^sub>R\<^sub>W\<^sub>C(1) by blast qed moreover have "right a (\ \ \)" using assms 1 \ weak_unit_self_composable hom_connected(1) [of \ a] right_def match_3 null_agreement by auto ultimately show ?thesis using Right_a.iso_char Right_a.arr_char right_iff_right_inv Right_a.inv_char H\<^sub>R_def by simp qed qed interpret R': equivalence_functor \Right a'\ \Right a'\ \H\<^sub>R a'\ proof - have "naturally_isomorphic (Right a) (Right a) (H\<^sub>R a) Right_a.map" using assms Right_a.natural_isomorphism_\ naturally_isomorphic_def by blast moreover have "naturally_isomorphic (Right a) (Right a) (H\<^sub>R a) (H\<^sub>R a')" using naturally_isomorphic_def \.natural_isomorphism_axioms by blast ultimately have "naturally_isomorphic (Right a) (Right a) (H\<^sub>R a') Right_a.map" using naturally_isomorphic_symmetric naturally_isomorphic_transitive by fast hence "naturally_isomorphic (Right a') (Right a') (H\<^sub>R a') (identity_functor.map (Right a'))" using 1 by auto thus "equivalence_functor (Right a') (Right a') (H\<^sub>R a')" using 1 R'.isomorphic_to_identity_is_equivalence naturally_isomorphic_def by fastforce qed show "weak_unit a'" using weak_unit_def iso L'.fully_faithful_functor_axioms R'.fully_faithful_functor_axioms by blast qed lemma sources_iso_closed: assumes "a \ sources \" and "a \ a'" shows "a' \ sources \" using assms isomorphism_respects_weak_units isomorphic_implies_equicomposable by blast lemma targets_iso_closed: assumes "a \ targets \" and "a \ a'" shows "a' \ targets \" using assms isomorphism_respects_weak_units isomorphic_implies_equicomposable by blast lemma sources_eqI: assumes "sources \ \ sources \ \ {}" shows "sources \ = sources \" using assms sources_iso_closed sources_are_isomorphic by blast lemma targets_eqI: assumes "targets \ \ targets \ \ {}" shows "targets \ = targets \" using assms targets_iso_closed targets_are_isomorphic by blast text \ The sets of sources and targets of a weak unit are isomorphism classes. \ lemma sources_char: assumes "weak_unit a" shows "sources a = {x. x \ a}" using assms sources_iso_closed weak_unit_iff_self_source sources_are_isomorphic isomorphic_symmetric by blast lemma targets_char: assumes "weak_unit a" shows "targets a = {x. x \ a}" using assms targets_iso_closed weak_unit_iff_self_target targets_are_isomorphic isomorphic_symmetric by blast end section "Horizontal Homs" text \ Here we define a locale that axiomatizes a (vertical) category \V\ that has been punctuated into ``horizontal homs'' by the choice of idempotent endofunctors \src\ and \trg\ that assign a specific ``source'' and ``target'' 1-cell to each of its arrows. The functors \src\ and \trg\ are also subject to further conditions that constrain how they commute with each other. \ locale horizontal_homs = category V + src: endofunctor V src + trg: endofunctor V trg for V :: "'a comp" (infixr "\" 55) and src :: "'a \ 'a" and trg :: "'a \ 'a" + assumes ide_src [simp]: "arr \ \ ide (src \)" and ide_trg [simp]: "arr \ \ ide (trg \)" and src_src [simp]: "arr \ \ src (src \) = src \" and trg_trg [simp]: "arr \ \ trg (trg \) = trg \" and trg_src [simp]: "arr \ \ trg (src \) = src \" and src_trg [simp]: "arr \ \ src (trg \) = trg \" begin no_notation in_hom ("\_ : _ \ _\") notation in_hom ("\_ : _ \ _\") text \ We define an \emph{object} to be an arrow that is its own source (or equivalently, its own target). \ definition obj where "obj a \ arr a \ src a = a" lemma obj_def': shows "obj a \ arr a \ trg a = a" using trg_src src_trg obj_def by metis lemma objE [elim]: assumes "obj a" and "\ ide a; src a = a; trg a = a \ \ T" shows T proof - have "ide a" using assms obj_def ide_src by metis moreover have "src a = a" using assms obj_def by simp moreover have "trg a = a" using assms obj_def' by simp ultimately show ?thesis using assms by simp qed (* TODO: Can't add "arr a" or "ide a" due to looping. *) - lemma obj_simps [simp]: + lemma obj_simps (* [simp] *): assumes "obj a" shows "src a = a" and "trg a = a" using assms by auto lemma obj_src [intro, simp]: assumes "arr \" shows "obj (src \)" using assms obj_def by auto lemma obj_trg [intro, simp]: assumes "arr \" shows "obj (trg \)" using assms obj_def by auto definition in_hhom ("\_ : _ \ _\") where "in_hhom \ a b \ arr \ \ src \ = a \ trg \ = b" abbreviation hhom where "hhom a b \ {\. \\ : a \ b\}" abbreviation (input) hseq\<^sub>H\<^sub>H where "hseq\<^sub>H\<^sub>H \ \\ \. arr \ \ arr \ \ src \ = trg \" lemma in_hhomI [intro, simp]: assumes "arr \" and "src \ = a" and "trg \ = b" shows "\\ : a \ b\" using assms in_hhom_def by auto lemma in_hhomE [elim]: assumes "\\ : a \ b\" and "\ arr \; obj a; obj b; src \ = a; trg \ = b \ \ T" shows "T" using assms in_hhom_def by auto (* * TODO: I tried removing the second assertion here, thinking that it should already * be covered by the category locale, but in fact it breaks some proofs in * SpanBicategory that ought to be trivial. So it seems that the presence of * this introduction rule adds something, and I should consider whether this rule * should be added to the category locale. *) lemma ide_in_hom [intro]: assumes "ide f" shows "\f : src f \ trg f\" and "\f : f \ f\" using assms by auto lemma src_dom [simp]: assumes "arr \" shows "src (dom \) = src \" using assms src.preserves_dom [of \] by simp lemma src_cod [simp]: assumes "arr \" shows "src (cod \) = src \" using assms src.preserves_cod [of \] by simp lemma trg_dom [simp]: assumes "arr \" shows "trg (dom \) = trg \" using assms trg.preserves_dom [of \] by simp lemma trg_cod [simp]: assumes "arr \" shows "trg (cod \) = trg \" using assms trg.preserves_cod [of \] by simp (* * TODO: In theory, the following simps should already be available from the fact * that src and trg are endofunctors. But they seem not to get used. *) lemma dom_src [simp]: assumes "arr \" shows "dom (src \) = src \" using assms by simp lemma cod_src [simp]: assumes "arr \" shows "cod (src \) = src \" using assms by simp lemma dom_trg [simp]: assumes "arr \" shows "dom (trg \) = trg \" using assms by simp lemma cod_trg [simp]: assumes "arr \" shows "cod (trg \) = trg \" using assms by simp lemma vcomp_in_hhom [intro, simp]: assumes "seq \ \" and "src \ = a" and "trg \ = b" shows "\\ \ \ : a \ b\" using assms src_cod [of "\ \ \"] trg_cod [of "\ \ \"] by auto lemma src_vcomp [simp]: assumes "seq \ \" shows "src (\ \ \) = src \" using assms src_cod [of "\ \ \"] by auto lemma trg_vcomp [simp]: assumes "seq \ \" shows "trg (\ \ \) = trg \" using assms trg_cod [of "\ \ \"] by auto lemma vseq_implies_hpar: assumes "seq \ \" shows "src \ = src \" and "trg \ = trg \" using assms src_dom [of "\ \ \"] trg_dom [of "\ \ \"] src_cod [of "\ \ \"] trg_cod [of "\ \ \"] by auto lemma vconn_implies_hpar: assumes "\\ : f \ g\" shows "src \ = src f" and "trg \ = trg f" and "src g = src f" and "trg g = trg f" using assms by auto lemma src_inv [simp]: assumes "iso \" shows "src (inv \) = src \" using assms inv_in_hom iso_is_arr src_dom src_cod iso_inv_iso dom_inv by metis lemma trg_inv [simp]: assumes "iso \" shows "trg (inv \) = trg \" using assms inv_in_hom iso_is_arr trg_dom trg_cod iso_inv_iso cod_inv by metis lemma inv_in_hhom [intro, simp]: assumes "iso \" and "src \ = a" and "trg \ = b" shows "\inv \ : a \ b\" using assms iso_is_arr by simp lemma hhom_is_subcategory: shows "subcategory V (\\. \\ : a \ b\)" using src_dom trg_dom src_cod trg_cod by (unfold_locales, auto) lemma isomorphic_objects_are_equal: assumes "obj a" and "obj b" and "a \ b" shows "a = b" using assms isomorphic_def by (metis arr_inv dom_inv in_homE objE src_dom src_inv) text \ Having the functors \src\ and \trg\ allows us to form categories VV and VVV of formally horizontally composable pairs and triples of arrows. \ interpretation VxV: product_category V V .. interpretation VV: subcategory VxV.comp \\\\. hseq\<^sub>H\<^sub>H (fst \\) (snd \\)\ by (unfold_locales, auto) lemma subcategory_VV: shows "subcategory VxV.comp (\\\. hseq\<^sub>H\<^sub>H (fst \\) (snd \\))" .. interpretation VxVxV: product_category V VxV.comp .. interpretation VVV: subcategory VxVxV.comp \\\\\. arr (fst \\\) \ VV.arr (snd \\\) \ src (fst \\\) = trg (fst (snd \\\))\ using VV.arr_char by (unfold_locales, auto) lemma subcategory_VVV: shows "subcategory VxVxV.comp (\\\\. arr (fst \\\) \ VV.arr (snd \\\) \ src (fst \\\) = trg (fst (snd \\\)))" .. end subsection "Prebicategories with Homs" text \ A \emph{weak composition with homs} consists of a weak composition that is equipped with horizontal homs in such a way that the chosen source and target of each 2-cell \\\ in fact lie in the set of sources and targets, respectively, of \\\, such that horizontal composition respects the chosen sources and targets, and such that if 2-cells \\\ and \\\ are horizontally composable, then the chosen target of \\\ coincides with the chosen source of \\\. \ locale weak_composition_with_homs = weak_composition + horizontal_homs + assumes src_in_sources: "arr \ \ src \ \ sources \" and trg_in_targets: "arr \ \ trg \ \ targets \" and src_hcomp': "\ \ \ \ null \ src (\ \ \) = src \" and trg_hcomp': "\ \ \ \ null \ trg (\ \ \) = trg \" and seq_if_composable: "\ \ \ \ null \ src \ = trg \" locale prebicategory_with_homs = prebicategory + weak_composition_with_homs begin lemma composable_char\<^sub>P\<^sub>B\<^sub>H: shows "\ \ \ \ null \ arr \ \ arr \ \ src \ = trg \" proof show "arr \ \ arr \ \ src \ = trg \ \ \ \ \ \ null" using trg_in_targets src_in_sources composable_if_connected by (metis sourcesD(3) targets_determine_composability) show "\ \ \ \ null \ arr \ \ arr \ \ src \ = trg \" using seq_if_composable composable_implies_arr by auto qed lemma hcomp_in_hom\<^sub>P\<^sub>B\<^sub>H: assumes "\\ : a \\<^sub>W\<^sub>C b\" and "\\ : b \\<^sub>W\<^sub>C c\" shows "\\ \ \ : a \\<^sub>W\<^sub>C c\" and "\\ \ \ : dom \ \ dom \ \ cod \ \ cod \\" proof - show "\\ \ \ : a \\<^sub>W\<^sub>C c\" using assms sources_determine_composability sources_hcomp targets_hcomp by auto thus "\\ \ \ : dom \ \ dom \ \ cod \ \ cod \\" using assms by auto qed text \ In a prebicategory with homs, if \a\ is an object (i.e. \src a = a\ and \trg a = a\), then \a\ is a weak unit. The converse need not hold: there can be weak units that the \src\ and \trg\ mappings send to other 1-cells in the same isomorphism class. \ lemma obj_is_weak_unit: assumes "obj a" shows "weak_unit a" proof - have "a \ sources a" using assms objE src_in_sources ideD(1) by metis thus ?thesis by auto qed end subsection "Choosing Homs" text \ Every prebicategory extends to a prebicategory with homs, by choosing an arbitrary representative of each isomorphism class of weak units to serve as an object. ``The source'' of a 2-cell is defined to be the chosen representative of the set of all its sources (which is an isomorphism class), and similarly for ``the target''. \ context prebicategory begin definition rep where "rep f \ SOME f'. f' \ { f'. f \ f' }" definition some_src where "some_src \ \ if arr \ then rep (SOME a. a \ sources \) else null" definition some_trg where "some_trg \ \ if arr \ then rep (SOME b. b \ targets \) else null" lemma isomorphic_ide_rep: assumes "ide f" shows "f \ rep f" proof - have "\f'. f' \ { f'. f \ f' }" using assms isomorphic_reflexive by blast thus ?thesis using rep_def someI_ex by simp qed lemma rep_rep: assumes "ide f" shows "rep (rep f) = rep f" proof - have "rep f \ { f'. f \ f' }" using assms isomorphic_ide_rep by simp have "{ f'. f \ f' } = { f'. rep f \ f' }" proof - have "\f'. f \ f' \ rep f \ f'" proof fix f' assume f': "f \ f'" show "rep f \ f'" proof - obtain \ where \: "\ \ hom f f' \ iso \" using f' by auto obtain \ where \: "\ \ hom f (rep f) \ iso \" using assms isomorphic_ide_rep by blast have "inv \ \ hom (rep f) f \ iso (inv \)" using \ iso_inv_iso inv_in_hom by simp hence "iso (V \ (inv \)) \ V \ (inv \) \ hom (rep f) f'" using \ isos_compose by auto thus ?thesis using isomorphic_def by auto qed next fix f' assume f': "rep f \ f'" show "f \ f'" using assms f' isomorphic_ide_rep isos_compose isomorphic_def by blast qed thus ?thesis by auto qed hence "rep (rep f) = (SOME f'. f' \ { f'. f \ f' })" using assms rep_def by fastforce also have "... = rep f" using assms rep_def by simp finally show ?thesis by simp qed lemma some_src_in_sources: assumes "arr \" shows "some_src \ \ sources \" proof - have 1: "(SOME a. a \ sources \) \ sources \" using assms arr_iff_has_src someI_ex [of "\a. a \ sources \"] by blast moreover have "ide (SOME a. a \ sources \)" using 1 weak_unit_self_composable by auto ultimately show ?thesis using assms 1 some_src_def sources_iso_closed [of "SOME a. a \ sources \" \] isomorphic_ide_rep [of "SOME a. a \ sources \"] by metis qed lemma some_trg_in_targets: assumes "arr \" shows "some_trg \ \ targets \" proof - have 1: "(SOME a. a \ targets \) \ targets \" using assms arr_iff_has_trg someI_ex [of "\a. a \ targets \"] by blast moreover have "ide (SOME a. a \ targets \)" using 1 weak_unit_self_composable by auto ultimately show ?thesis using assms 1 some_trg_def targets_iso_closed [of "SOME a. a \ targets \" \] isomorphic_ide_rep [of "SOME a. a \ targets \"] by presburger qed lemma some_src_dom: assumes "arr \" shows "some_src (dom \) = some_src \" using assms some_src_def sources_dom by simp lemma some_src_cod: assumes "arr \" shows "some_src (cod \) = some_src \" using assms some_src_def sources_cod by simp lemma some_trg_dom: assumes "arr \" shows "some_trg (dom \) = some_trg \" using assms some_trg_def targets_dom by simp lemma some_trg_cod: assumes "arr \" shows "some_trg (cod \) = some_trg \" using assms some_trg_def targets_cod by simp lemma ide_some_src: assumes "arr \" shows "ide (some_src \)" using assms some_src_in_sources weak_unit_self_composable by blast lemma ide_some_trg: assumes "arr \" shows "ide (some_trg \)" using assms some_trg_in_targets weak_unit_self_composable by blast lemma some_src_composable: assumes "arr \" shows "\ \ \ \ null \ some_src \ \ \ \ null" using assms some_src_in_sources sources_determine_composability by blast lemma some_trg_composable: assumes "arr \" shows "\ \ \ \ null \ \ \ some_trg \ \ null" using assms some_trg_in_targets targets_determine_composability by blast lemma sources_some_src: assumes "arr \" shows "sources (some_src \) = sources \" using assms sources_determine_composability some_src_in_sources by blast lemma targets_some_trg: assumes "arr \" shows "targets (some_trg \) = targets \" using assms targets_determine_composability some_trg_in_targets by blast lemma src_some_src: assumes "arr \" shows "some_src (some_src \) = some_src \" using assms some_src_def ide_some_src sources_some_src by force lemma trg_some_trg: assumes "arr \" shows "some_trg (some_trg \) = some_trg \" using assms some_trg_def ide_some_trg targets_some_trg by force lemma sources_char': assumes "arr \" shows "a \ sources \ \ some_src \ \ a" using assms some_src_in_sources sources_iso_closed sources_are_isomorphic by meson lemma targets_char': assumes "arr \" shows "a \ targets \ \ some_trg \ \ a" using assms some_trg_in_targets targets_iso_closed targets_are_isomorphic by blast text \ An arbitrary choice of sources and targets in a prebicategory results in a notion of formal composability that coincides with the actual horizontal composability of the prebicategory. \ lemma composable_char\<^sub>P\<^sub>B: shows "\ \ \ \ null \ arr \ \ arr \ \ some_src \ = some_trg \" proof assume \\: "\ \ \ \ null" show "arr \ \ arr \ \ some_src \ = some_trg \" using \\ composable_implies_arr connected_if_composable some_src_def some_trg_def by force next assume \\: "arr \ \ arr \ \ some_src \ = some_trg \" show "\ \ \ \ null" using \\ some_src_in_sources some_trg_composable by force qed text \ A 1-cell is its own source if and only if it is its own target. \ lemma self_src_iff_self_trg: assumes "ide a" shows "a = some_src a \ a = some_trg a" proof assume a: "a = some_src a" have "weak_unit a \ a \ a \ null" using assms a some_src_in_sources [of a] by force thus "a = some_trg a" using a composable_char\<^sub>P\<^sub>B by simp next assume a: "a = some_trg a" have "weak_unit a \ a \ a \ null" using assms a some_trg_in_targets [of a] by force thus "a = some_src a" using a composable_char\<^sub>P\<^sub>B by simp qed lemma some_trg_some_src: assumes "arr \" shows "some_trg (some_src \) = some_src \" using assms ide_some_src some_src_def some_trg_def some_src_in_sources sources_char targets_char sources_some_src by force lemma src_some_trg: assumes "arr \" shows "some_src (some_trg \) = some_trg \" using assms ide_some_trg some_src_def some_trg_def some_trg_in_targets sources_char targets_char targets_some_trg by force lemma some_src_eqI: assumes "a \ sources \" and "some_src a = a" shows "some_src \ = a" proof - have 1: "arr \ \ arr a" using assms composable_implies_arr by auto have "some_src \ = rep (SOME x. x \ sources \)" using assms 1 some_src_def by simp also have "... = rep (SOME x. some_src \ \ x)" using assms 1 sources_char' by simp also have "... = rep (SOME x. some_src a \ x)" using assms 1 some_src_in_sources sources_are_isomorphic isomorphic_symmetric isomorphic_transitive by metis also have "... = rep (SOME x. x \ sources a)" using assms 1 sources_char' by auto also have "... = some_src a" using assms 1 some_src_def by simp also have "... = a" using assms by auto finally show ?thesis by simp qed lemma some_trg_eqI: assumes "b \ targets \" and "some_trg b = b" shows "some_trg \ = b" proof - have 1: "arr \ \ arr b" using assms composable_implies_arr by auto have "some_trg \ = rep (SOME x. x \ targets \)" using assms 1 some_trg_def by simp also have "... = rep (SOME x. some_trg \ \ x)" using assms 1 targets_char' by simp also have "... = rep (SOME x. some_trg b \ x)" using assms 1 some_trg_in_targets targets_are_isomorphic isomorphic_symmetric isomorphic_transitive by metis also have "... = rep (SOME x. x \ targets b)" using assms 1 targets_char' by auto also have "... = some_trg b" using assms 1 some_trg_def by simp also have "... = b" using assms by auto finally show ?thesis by simp qed lemma some_src_comp: assumes "\ \ \ \ null" shows "some_src (\ \ \) = some_src \" proof (intro some_src_eqI [of "some_src \" "\ \ \"]) show "some_src (some_src \) = some_src \" using assms src_some_src composable_implies_arr by simp show "some_src \ \ sources (H \ \)" using assms some_src_in_sources composable_char\<^sub>P\<^sub>B match_3 [of \ "some_src \"] by (simp add: sources_hcomp) qed lemma some_trg_comp: assumes "\ \ \ \ null" shows "some_trg (\ \ \) = some_trg \" proof (intro some_trg_eqI [of "some_trg \" "\ \ \"]) show "some_trg (some_trg \) = some_trg \" using assms trg_some_trg composable_implies_arr by simp show "some_trg \ \ targets (H \ \)" using assms some_trg_in_targets composable_char\<^sub>P\<^sub>B match_4 [of \ \ "some_trg \"] by (simp add: targets_hcomp) qed text \ The mappings that take an arrow to its chosen source or target are endofunctors of the vertical category, which commute with each other in the manner required for horizontal homs. \ interpretation S: endofunctor V some_src apply unfold_locales using some_src_def apply simp using ide_some_src apply simp using some_src_dom ide_some_src apply simp using some_src_cod ide_some_src apply simp proof - fix \ \ assume \\: "seq \ \" show "some_src (\ \ \) = some_src \ \ some_src \" using \\ some_src_dom [of "\ \ \"] some_src_dom some_src_cod [of "\ \ \"] some_src_cod ide_some_src by auto qed interpretation T: endofunctor V some_trg apply unfold_locales using some_trg_def apply simp using ide_some_trg apply simp using some_trg_dom ide_some_trg apply simp using some_trg_cod ide_some_trg apply simp proof - fix \ \ assume \\: "seq \ \" show "some_trg (\ \ \) = some_trg \ \ some_trg \" using \\ some_trg_dom [of "\ \ \"] some_trg_dom some_trg_cod [of "\ \ \"] some_trg_cod ide_some_trg by auto qed interpretation weak_composition_with_homs V H some_src some_trg apply unfold_locales using some_src_in_sources some_trg_in_targets src_some_src trg_some_trg src_some_trg some_trg_some_src some_src_comp some_trg_comp composable_char\<^sub>P\<^sub>B ide_some_src ide_some_trg by simp_all proposition extends_to_weak_composition_with_homs: shows "weak_composition_with_homs V H some_src some_trg" .. proposition extends_to_prebicategory_with_homs: shows "prebicategory_with_homs V H \ some_src some_trg" .. end subsection "Choosing Units" text \ A \emph{prebicategory with units} is a prebicategory equipped with a choice, for each weak unit \a\, of a ``unit isomorphism'' \\\[a] : a \ a \ a\\. \ locale prebicategory_with_units = prebicategory V H \ + weak_composition V H for V :: "'a comp" (infixr "\" 55) and H :: "'a comp" (infixr "\" 53) and \ :: "'a \ 'a \ 'a \ 'a" ("\[_, _, _]") and \ :: "'a \ 'a" ("\[_]") + assumes unit_in_vhom\<^sub>P\<^sub>B\<^sub>U: "weak_unit a \ \\[a] : a \ a \ a\" and iso_unit\<^sub>P\<^sub>B\<^sub>U: "weak_unit a \ iso \[a]" begin lemma unit_in_hom\<^sub>P\<^sub>B\<^sub>U: assumes "weak_unit a" shows "\\[a] : a \\<^sub>W\<^sub>C a\" and "\\[a] : a \ a \ a\" proof - show 1: "\\[a] : a \ a \ a\" using assms unit_in_vhom\<^sub>P\<^sub>B\<^sub>U by auto show "\\[a] : a \\<^sub>W\<^sub>C a\" using assms 1 weak_unit_iff_self_source weak_unit_iff_self_target sources_cod [of "\[a]"] targets_cod [of "\[a]"] by (elim in_homE, auto) qed lemma unit_simps [simp]: assumes "weak_unit a" shows "arr \[a]" and "dom \[a] = a \ a" and "cod \[a] = a" using assms unit_in_vhom\<^sub>P\<^sub>B\<^sub>U by auto end text \ Every prebicategory extends to a prebicategory with units, simply by choosing the unit isomorphisms arbitrarily. \ context prebicategory begin proposition extends_to_prebicategory_with_units: shows "prebicategory_with_units V H \ some_unit" using iso_some_unit by (unfold_locales, auto) end subsection "Horizontal Composition" text \ The following locale axiomatizes a (vertical) category \V\ with horizontal homs, which in addition has been equipped with a functorial operation \H\ of horizontal composition from \VV\ to \V\, assumed to preserve source and target. \ locale horizontal_composition = horizontal_homs V src trg + VxV: product_category V V + VV: subcategory VxV.comp \\\\. hseq\<^sub>H\<^sub>H (fst \\) (snd \\)\ + H: "functor" VV.comp V \\\\. H (fst \\) (snd \\)\ for V :: "'a comp" (infixr "\" 55) and H :: "'a \ 'a \ 'a" (infixr "\" 53) and src :: "'a \ 'a" and trg :: "'a \ 'a" + assumes src_hcomp: "arr (\ \ \) \ src (\ \ \) = src \" and trg_hcomp: "arr (\ \ \) \ trg (\ \ \) = trg \" begin (* TODO: Why does this get re-introduced? *) no_notation in_hom ("\_ : _ \ _\") text \ \H\ is a partial magma, which shares its null with \V\. \ lemma is_partial_magma: shows "partial_magma H" and "partial_magma.null H = null" proof - have 1: "\f. null \ f = null \ f \ null = null" using H.is_extensional VV.arr_char not_arr_null by auto interpret H: partial_magma H proof show "\!n. \f. n \ f = n \ f \ n = n" proof show "\f. null \ f = null \ f \ null = null" by fact show "\n. \f. n \ f = n \ f \ n = n \ n = null" using 1 VV.arr_char H.is_extensional not_arr_null by metis qed qed show "partial_magma H" .. show "H.null = null" using 1 H.null_def the1_equality [of "\n. \f. n \ f = n \ f \ n = n"] by metis qed text \ \textbf{Note:} The following is ``almost'' \H.seq\, but for that we would need \H.arr = V.arr\. This would be unreasonable to expect, in general, as the definition of \H.arr\ is based on ``strict'' units rather than weak units. Later we will show that we do have \H.arr = V.arr\ if the vertical category is discrete. \ abbreviation hseq where "hseq \ \ \ arr (\ \ \)" lemma hseq_char: shows "hseq \ \ \ arr \ \ arr \ \ src \ = trg \" proof - have "hseq \ \ \ VV.arr (\, \)" using H.is_extensional H.preserves_arr by force also have "... \ arr \ \ arr \ \ src \ = trg \" using VV.arr_char by force finally show ?thesis by blast qed lemma hseq_char': shows "hseq \ \ \ \ \ \ \ null" using VV.arr_char H.preserves_arr H.is_extensional hseq_char [of \ \] by auto lemma hseqI' [simp]: assumes "arr \" and "arr \" and "src \ = trg \" shows "hseq \ \" using assms hseq_char by simp lemma hseqI [intro]: assumes "\\ : a \ b\" and "\\ : b \ c\" shows "hseq \ \" using assms hseq_char by auto lemma hseqE [elim]: assumes "hseq \ \" and "arr \ \ arr \ \ src \ = trg \ \ T" shows "T" using assms hseq_char by simp lemma hcomp_simps [simp]: assumes "hseq \ \" shows "src (\ \ \) = src \" and "trg (\ \ \) = trg \" and "dom (\ \ \) = dom \ \ dom \" and "cod (\ \ \) = cod \ \ cod \" using assms VV.arr_char src_hcomp apply blast using assms VV.arr_char trg_hcomp apply blast using assms VV.arr_char H.preserves_dom apply force using assms VV.arr_char H.preserves_cod by force lemma ide_hcomp [intro, simp]: assumes "ide \" and "ide \" and "src \ = trg \" shows "ide (\ \ \)" using assms VV.ide_char VV.arr_char H.preserves_ide [of "(\, \)"] by auto - lemma hcomp_in_hhom [intro, simp]: + lemma hcomp_in_hhom [intro]: assumes "\\ : a \ b\" and "\\ : b \ c\" shows "\\ \ \ : a \ c\" using assms hseq_char by fastforce + lemma hcomp_in_hhom' (* [simp] *): + assumes "arr \" and "arr \" and "src \ = a" and "trg \ = c" and "src \ = trg \" + shows "\\ \ \ : a \ c\" + using assms hseq_char by fastforce + lemma hcomp_in_hhomE [elim]: assumes "\\ \ \ : a \ c\" and "\ arr \; arr \; src \ = trg \; src \ = a; trg \ = c \ \ T" shows T using assms in_hhom_def by fastforce lemma hcomp_in_vhom [intro]: assumes "\\ : f \ g\" and "\\ : h \ k\" and "src h = trg f" shows "\\ \ \ : h \ f \ k \ g\" using assms by fastforce lemma hcomp_in_vhom' [simp]: assumes "hseq \ \" and "dom \ = f" and "dom \ = h" and "cod \ = g" and "cod \ = k" assumes "\\ : f \ g\" and "\\ : h \ k\" and "src h = trg f" shows "\\ \ \ : h \ f \ k \ g\" using assms by fastforce lemma hcomp_in_vhomE [elim]: assumes "\\ \ \ : f \ g\" and "\ arr \; arr \; src \ = trg \; src \ = src f; src \ = src g; trg \ = trg f; trg \ = trg g \ \ T" shows T using assms in_hom_def by (metis in_homE hseqE src_cod src_dom src_hcomp trg_cod trg_dom trg_hcomp) text \ A horizontal composition yields a weak composition by simply forgetting the \src\ and \trg\ functors. \ lemma match_1: assumes "\ \ \ \ null" and "(\ \ \) \ \ \ null" shows "\ \ \ \ null" using assms H.is_extensional not_arr_null VV.arr_char hseq_char hseq_char' by auto lemma match_2: assumes "\ \ (\ \ \) \ null" and "\ \ \ \ null" shows "\ \ \ \ null" using assms H.is_extensional not_arr_null VV.arr_char hseq_char hseq_char' by auto lemma match_3: assumes "\ \ \ \ null" and "\ \ \ \ null" shows "(\ \ \) \ \ \ null" using assms H.is_extensional not_arr_null VV.arr_char hseq_char hseq_char' by auto lemma match_4: assumes "\ \ \ \ null" and "\ \ \ \ null" shows "\ \ (\ \ \) \ null" using assms H.is_extensional not_arr_null VV.arr_char hseq_char hseq_char' by auto lemma left_connected: assumes "seq \ \'" shows "\ \ \ \ null \ \' \ \ \ null" using assms H.is_extensional not_arr_null VV.arr_char hseq_char' by (metis hseq_char seqE vseq_implies_hpar(1)) lemma right_connected: assumes "seq \ \'" shows "H \ \ \ null \ H \ \' \ null" using assms H.is_extensional not_arr_null VV.arr_char hseq_char' by (metis hseq_char seqE vseq_implies_hpar(2)) proposition is_weak_composition: shows "weak_composition V H" proof - have 1: "(\\\. fst \\ \ snd \\ \ null) = (\\\. arr (fst \\) \ arr (snd \\) \ src (fst \\) = trg (snd \\))" using hseq_char' by auto interpret VoV: subcategory VxV.comp \\\\. fst \\ \ snd \\ \ null\ using 1 VV.subcategory_axioms by simp interpret H: "functor" VoV.comp V \\\\. fst \\ \ snd \\\ using H.functor_axioms 1 by simp show ?thesis using match_1 match_2 match_3 match_4 left_connected right_connected by (unfold_locales, metis) qed interpretation weak_composition V H using is_weak_composition by auto text \ It can be shown that \arr ((\ \ \) \ (\ \ \)) \ (\ \ \) \ (\ \ \) = (\ \ \) \ (\ \ \)\. However, we do not have \arr ((\ \ \) \ (\ \ \)) \ (\ \ \) \ (\ \ \) = (\ \ \) \ (\ \ \)\, because it does not follow from \arr ((\ \ \) \ (\ \ \))\ that \dom \ = cod \\ and \dom \ = cod \\, only that \dom \ \ dom \ = cod \ \ cod \\. So we don't get interchange unconditionally. \ lemma interchange: assumes "seq \ \" and "seq \ \" shows "(\ \ \) \ (\ \ \) = (\ \ \) \ (\ \ \)" using assms interchange by simp lemma whisker_right: assumes "ide f" and "seq \ \" shows "(\ \ \) \ f = (\ \ f) \ (\ \ f)" using assms whisker_right by simp lemma whisker_left: assumes "ide f" and "seq \ \" shows "f \ (\ \ \) = (f \ \) \ (f \ \)" using assms whisker_left by simp lemma inverse_arrows_hcomp: assumes "iso \" and "iso \" and "src \ = trg \" shows "inverse_arrows (\ \ \) (inv \ \ inv \)" proof - show "inverse_arrows (\ \ \) (inv \ \ inv \)" proof show "ide ((inv \ \ inv \) \ (\ \ \))" proof - have "(inv \ \ inv \) \ (\ \ \) = dom \ \ dom \" using assms interchange iso_is_arr comp_inv_arr' by (metis arr_dom) thus ?thesis using assms iso_is_arr by simp qed show "ide ((\ \ \) \ (inv \ \ inv \))" proof - have "(\ \ \) \ (inv \ \ inv \) = cod \ \ cod \" using assms interchange iso_is_arr comp_arr_inv' by (metis arr_cod) thus ?thesis using assms iso_is_arr by simp qed qed qed lemma iso_hcomp [intro, simp]: assumes "iso \" and "iso \" and "src \ = trg \" shows "iso (\ \ \)" using assms inverse_arrows_hcomp by auto lemma isomorphic_implies_ide: assumes "f \ g" shows "ide f" and "ide g" using assms isomorphic_def by auto lemma hcomp_ide_isomorphic: assumes "ide f" and "g \ h" and "src f = trg g" shows "f \ g \ f \ h" proof - obtain \ where \: "iso \ \ \\ : g \ h\" using assms isomorphic_def by auto have "iso (f \ \) \ \f \ \ : f \ g \ f \ h\" using assms \ iso_hcomp by auto thus ?thesis using isomorphic_def by auto qed lemma hcomp_isomorphic_ide: assumes "f \ g" and "ide h" and "src f = trg h" shows "f \ h \ g \ h" proof - obtain \ where \: "iso \ \ \\ : f \ g\" using assms isomorphic_def by auto have "iso (\ \ h) \ \\ \ h : f \ h \ g \ h\" using assms \ iso_hcomp by auto thus ?thesis using isomorphic_def by auto qed lemma isomorphic_implies_hpar: assumes "f \ f'" shows "ide f" and "ide f'" and "src f = src f'" and "trg f = trg f'" using assms isomorphic_def by auto lemma inv_hcomp [simp]: assumes "iso \" and "iso \" and "src \ = trg \" shows "inv (\ \ \) = inv \ \ inv \" using assms inverse_arrow_unique [of "\ \ \"] inv_is_inverse inverse_arrows_hcomp by auto interpretation VxVxV: product_category V VxV.comp .. interpretation VVV: subcategory VxVxV.comp \\\\\. arr (fst \\\) \ VV.arr (snd \\\) \ src (fst \\\) = trg (fst (snd \\\))\ using subcategory_VVV by auto text \ The following define the two ways of using horizontal composition to compose three arrows. \ definition HoHV where "HoHV \ \ if VVV.arr \ then (fst \ \ fst (snd \)) \ snd (snd \) else null" definition HoVH where "HoVH \ \ if VVV.arr \ then fst \ \ fst (snd \) \ snd (snd \) else null" lemma functor_HoHV: shows "functor VVV.comp V HoHV" apply unfold_locales using VVV.arr_char VV.arr_char VVV.dom_char VVV.cod_char VVV.comp_char HoHV_def apply auto[4] proof - fix f g assume fg: "VVV.seq g f" show "HoHV (VVV.comp g f) = HoHV g \ HoHV f" proof - have "VxVxV.comp g f = (fst g \ fst f, fst (snd g) \ fst (snd f), snd (snd g) \ snd (snd f))" using fg VVV.seq_char VVV.arr_char VV.arr_char VxVxV.comp_char VxV.comp_char by (metis (no_types, lifting) VxV.seqE VxVxV.seqE) hence "HoHV (VVV.comp g f) = (fst g \ fst f \ fst (snd g) \ fst (snd f)) \ snd (snd g) \ snd (snd f)" using HoHV_def VVV.comp_simp fg by auto also have "... = ((fst g \ fst (snd g)) \ snd (snd g)) \ ((fst f \ fst (snd f)) \ snd (snd f))" using fg VVV.seq_char VVV.arr_char VV.arr_char interchange by (metis (no_types, lifting) VxV.seqE VxVxV.seqE hseqI' src_vcomp trg_vcomp) also have "... = HoHV g \ HoHV f" using HoHV_def fg by auto finally show ?thesis by simp qed qed lemma functor_HoVH: shows "functor VVV.comp V HoVH" apply unfold_locales using VVV.arr_char VV.arr_char VVV.dom_char VVV.cod_char VVV.comp_char HoHV_def HoVH_def apply auto[4] proof - fix f g assume fg: "VVV.seq g f" show "HoVH (VVV.comp g f) = HoVH g \ HoVH f" proof - have "VxVxV.comp g f = (fst g \ fst f, fst (snd g) \ fst (snd f), snd (snd g) \ snd (snd f))" using fg VVV.seq_char VVV.arr_char VV.arr_char VxVxV.comp_char VxV.comp_char by (metis (no_types, lifting) VxV.seqE VxVxV.seqE) hence "HoVH (VVV.comp g f) = fst g \ fst f \ fst (snd g) \ fst (snd f) \ snd (snd g) \ snd (snd f)" using HoVH_def VVV.comp_simp fg by auto also have "... = (fst g \ fst (snd g) \ snd (snd g)) \ (fst f \ fst (snd f) \ snd (snd f))" using fg VVV.seq_char VVV.arr_char VV.arr_char interchange by (metis (no_types, lifting) VxV.seqE VxVxV.seqE hseqI' src_vcomp trg_vcomp) also have "... = HoVH g \ HoVH f" using fg VVV.seq_char VVV.arr_char HoVH_def VVV.comp_char VV.arr_char by (metis (no_types, lifting)) finally show ?thesis by simp qed qed text \ The following define horizontal composition of an arrow on the left by its target and on the right by its source. \ abbreviation L where "L \ \\. if arr \ then trg \ \ \ else null" abbreviation R where "R \ \\. if arr \ then \ \ src \ else null" lemma endofunctor_L: shows "endofunctor V L" using vseq_implies_hpar(2) whisker_left by (unfold_locales, auto) lemma endofunctor_R: shows "endofunctor V R" using vseq_implies_hpar(1) whisker_right by (unfold_locales, auto) end end diff --git a/thys/Bicategory/Strictness.thy b/thys/Bicategory/Strictness.thy --- a/thys/Bicategory/Strictness.thy +++ b/thys/Bicategory/Strictness.thy @@ -1,4491 +1,4491 @@ (* Title: Strictness Author: Eugene W. Stark , 2019 Maintainer: Eugene W. Stark *) section "Strictness" theory Strictness imports Category3.ConcreteCategory Pseudofunctor CanonicalIsos begin text \ In this section we consider bicategories in which some or all of the canonical isomorphisms are assumed to be identities. A \emph{normal} bicategory is one in which the unit isomorphisms are identities, so that unit laws for horizontal composition are satisfied ``on the nose''. A \emph{strict} bicategory (also known as a \emph{2-category}) is a bicategory in which both the unit and associativity isomoprhisms are identities, so that horizontal composition is strictly associative as well as strictly unital. From any given bicategory \B\ we may construct a related strict bicategory \S\, its \emph{strictification}, together with a pseudofunctor that embeds \B\ in \S\. The Strictness Theorem states that this pseudofunctor is an equivalence pseudofunctor, so that bicategory \B\ is biequivalent to its strictification. The Strictness Theorem is often used informally to justify suppressing canonical isomorphisms; which amounts to proving a theorem about 2-categories and asserting that it holds for all bicategories. Here we are working formally, so we can't just wave our hands and mutter something about the Strictness Theorem when we want to avoid dealing with units and associativities. However, in cases where we can establish that the property we would like to prove is reflected by the embedding of a bicategory in its strictification, then we can formally apply the Strictness Theorem to generalize to all bicategories a result proved for 2-categories. We will apply this approach here to simplify the proof of some facts about internal equivalences in a bicategory. \ subsection "Normal and Strict Bicategories" text \ A \emph{normal} bicategory is one in which the unit isomorphisms are identities, so that unit laws for horizontal composition are satisfied ``on the nose''. \ locale normal_bicategory = bicategory + assumes strict_lunit: "\f. ide f \ \[f] = f" and strict_runit: "\f. ide f \ \[f] = f" begin lemma strict_unit: assumes "obj a" shows "ide \[a]" using assms strict_runit unitor_coincidence(2) [of a] by auto lemma strict_lunit': assumes "ide f" shows "\\<^sup>-\<^sup>1[f] = f" using assms strict_lunit by simp lemma strict_runit': assumes "ide f" shows "\\<^sup>-\<^sup>1[f] = f" using assms strict_runit by simp lemma hcomp_obj_arr: assumes "obj b" and "arr f" and "b = trg f" shows "b \ f = f" using assms strict_lunit by (metis comp_arr_dom comp_ide_arr ide_cod ide_dom lunit_naturality) lemma hcomp_arr_obj: assumes "arr f" and "obj a" and "src f = a" shows "f \ a = f" using assms strict_runit by (metis comp_arr_dom comp_ide_arr ide_cod ide_dom runit_naturality) end text \ A \emph{strict} bicategory is a normal bicategory in which the associativities are also identities, so that associativity of horizontal composition holds ``on the nose''. \ locale strict_bicategory = normal_bicategory + assumes strict_assoc: "\f g h. \ide f; ide g; ide h; src f = trg g; src g = trg h\ \ ide \[f, g, h]" begin lemma strict_assoc': assumes "ide f" and "ide g" and "ide h" and "src f = trg g" and "src g = trg h" shows "ide \\<^sup>-\<^sup>1[f, g, h]" using assms strict_assoc by simp lemma hcomp_assoc: shows "(\ \ \) \ \ = \ \ \ \ \" proof (cases "hseq \ \ \ hseq \ \") show "\ (hseq \ \ \ hseq \ \) \ ?thesis" by (metis hseqE hseq_char' match_1 match_2) show "hseq \ \ \ hseq \ \ \ ?thesis" proof - assume 1: "hseq \ \ \ hseq \ \" have 2: "arr \ \ arr \ \ arr \ \ src \ = trg \ \ src \ = trg \" using 1 by blast have "(\ \ \) \ \ = \[cod \, cod \, cod \] \ ((\ \ \) \ \)" using 1 assoc_in_hom strict_assoc comp_cod_arr assoc_simps(4) hseq_char by simp also have "... = (\ \ \ \ \) \ \[dom \, dom \, dom \]" using 1 assoc_naturality by auto also have "... = \ \ \ \ \" using 2 assoc_in_hom [of "dom \" "dom \" "dom \"] strict_assoc comp_arr_dom by auto finally show ?thesis by simp qed qed text \ In a strict bicategory, every canonical isomorphism is an identity. \ interpretation bicategorical_language .. interpretation E: self_evaluation_map V H \ \ src trg .. notation E.eval ("\_\") lemma ide_eval_Can: assumes "Can t" shows "ide \t\" proof - have 1: "\u1 u2. \ ide \u1\; ide \u2\; Arr u1; Arr u2; Dom u1 = Cod u2 \ \ ide (\u1\ \ \u2\)" by (metis (no_types, lifting) E.eval_simps'(4-5) comp_ide_self ide_char) have "\u. Can u \ ide \u\" proof - fix u show "Can u \ ide \u\" (* TODO: Rename \_ide_simp \_ide_simp to \_ide_eq \_ide_eq *) using 1 \_def \'_def strict_lunit strict_runit strict_assoc strict_assoc' \_ide_simp \_ide_simp Can_implies_Arr comp_ide_arr E.eval_simps'(2-3) apply (induct u) by auto qed thus ?thesis using assms by simp qed lemma ide_can: assumes "Ide f" and "Ide g" and "\<^bold>\f\<^bold>\ = \<^bold>\g\<^bold>\" shows "ide (can g f)" proof - have "Can (Inv (g\<^bold>\) \<^bold>\ f\<^bold>\)" using assms Can_red Can_Inv red_in_Hom Inv_in_Hom by simp thus ?thesis using assms can_def ide_eval_Can by presburger qed end subsection "Strictification" (* * TODO: Perhaps change the typeface used for a symbol that stands for a bicategory; * for example, to avoid the clashes here between B used as the name of a bicategory * and B used to denote a syntactic identity term. *) text \ The Strictness Theorem asserts that every bicategory is biequivalent to a strict bicategory. More specifically, it shows how to construct, given an arbitrary bicategory, a strict bicategory (its \emph{strictification}) that is biequivalent to it. Consequently, given a property \P\ of bicategories that is ``bicategorical'' (\emph{i.e.}~respects biequivalence), if we want to show that \P\ holds for a bicategory \B\ then it suffices to show that \P\ holds for the strictification of \B\, and if we want to show that \P\ holds for all bicategories, it is sufficient to show that it holds for all strict bicategories. This is very useful, because it becomes quite tedious, even with the aid of a proof assistant, to do ``diagram chases'' with all the units and associativities fully spelled out. Given a bicategory \B\, the strictification \S\ of \B\ may be constructed as the bicategory whose arrows are triples \(A, B, \)\, where \X\ and \Y\ are ``normal identity terms'' (essentially, nonempty horizontally composable lists of 1-cells of \B\) having the same syntactic source and target, and \\\ : \X\ \ \Y\\\ in \B\. Vertical composition in \S\ is given by composition of the underlying arrows in \B\. Horizontal composition in \S\ is given by \(A, B, \) \ (A', B', \') = (AA', BB', \)\, where \AA'\ and \BB'\ denote concatenations of lists and where \\\ is defined as the composition \can BB' (B \<^bold>\ B') \ (\ \ \') \ can (A \<^bold>\ A') AA'\, where \can (A \<^bold>\ A') AA'\ and \can BB' (B \<^bold>\ B')\ are canonical isomorphisms in \B\. The canonical isomorphism \can (A \<^bold>\ A') AA'\ corresponds to taking a pair of lists \A \<^bold>\ A'\ and ``shifting the parentheses to the right'' to obtain a single list \AA'\. The canonical isomorphism can \BB' (B \<^bold>\ B')\ corresponds to the inverse rearrangement. The bicategory \B\ embeds into its strictification \S\ via the functor \UP\ that takes each arrow \\\ of \B\ to \(\<^bold>\dom \\<^bold>\, \<^bold>\cod \\<^bold>\, \)\, where \\<^bold>\dom \\<^bold>\\ and \\<^bold>\cod \\<^bold>\\ denote one-element lists. This mapping extends to a pseudofunctor. There is also a pseudofunctor \DN\, which maps \(A, B, \)\ in \S\ to \\\ in \B\; this is such that \DN o UP\ is the identity on \B\ and \UP o DN\ is equivalent to the identity on \S\, so we obtain a biequivalence between \B\ and \S\. It seems difficult to find references that explicitly describe a strictification construction in elementary terms like this (in retrospect, it ought to have been relatively easy to rediscover such a construction, but my thinking got off on the wrong track). One reference that I did find useful was \cite{unapologetic-strictification}, which discusses strictification for monoidal categories. \ locale strictified_bicategory = B: bicategory V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B for V\<^sub>B :: "'a comp" (infixr "\\<^sub>B" 55) and H\<^sub>B :: "'a \ 'a \ 'a" (infixr "\\<^sub>B" 53) and \\<^sub>B :: "'a \ 'a \ 'a \ 'a" ("\\<^sub>B[_, _, _]") and \\<^sub>B :: "'a \ 'a" ("\\<^sub>B[_]") and src\<^sub>B :: "'a \ 'a" and trg\<^sub>B :: "'a \ 'a" begin sublocale E: self_evaluation_map V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B .. notation B.in_hhom ("\_ : _ \\<^sub>B _\") notation B.in_hom ("\_ : _ \\<^sub>B _\") notation E.eval ("\_\") notation E.Nmlize ("\<^bold>\_\<^bold>\") text \ The following gives the construction of a bicategory whose arrows are triples \(A, B, \)\, where \Nml A \ Ide A\, \Nml B \ Ide B\, \Src A = Src B\, \Trg A = Trg B\, and \\ : \A\ \ \B\\. We use @{locale concrete_category} to construct the vertical composition, so formally the arrows of the bicategory will be of the form \MkArr A B \\. \ text \ The 1-cells of the bicategory correspond to normal, identity terms \A\ in the bicategorical language associated with \B\. \ abbreviation IDE where "IDE \ {A. E.Nml A \ E.Ide A}" text \ If terms \A\ and \B\ determine 1-cells of the strictification and have a common source and target, then the 2-cells between these 1-cells correspond to arrows \\\ of the underlying bicategory such that \\\ : \A\ \\<^sub>B \B\\\. \ abbreviation HOM where "HOM A B \ {\. E.Src A = E.Src B \ E.Trg A = E.Trg B \ \\ : \A\ \\<^sub>B \B\\}" text \ The map taking term \A \ OBJ\ to its evaluation \\A\ \ HOM A A\ defines the embedding of 1-cells as identity 2-cells. \ abbreviation EVAL where "EVAL \ E.eval" sublocale concrete_category IDE HOM EVAL \\_ _ _ \ \. \ \\<^sub>B \\ using E.ide_eval_Ide B.comp_arr_dom B.comp_cod_arr B.comp_assoc by (unfold_locales, auto) lemma is_concrete_category: shows "concrete_category IDE HOM EVAL (\_ _ _ \ \. \ \\<^sub>B \)" .. abbreviation vcomp (infixr "\" 55) where "vcomp \ COMP" lemma arr_char: shows "arr F \ E.Nml (Dom F) \ E.Ide (Dom F) \ E.Nml (Cod F) \ E.Ide (Cod F) \ E.Src (Dom F) = E.Src (Cod F) \ E.Trg (Dom F) = E.Trg (Cod F) \ \Map F : \Dom F\ \\<^sub>B \Cod F\\ \ F \ Null" using arr_char by auto lemma arrI (* [intro] *): assumes "E.Nml (Dom F)" and "E.Ide (Dom F)" and "E.Nml (Cod F)" and "E.Ide (Cod F)" and "E.Src (Dom F) = E.Src (Cod F)" and "E.Trg (Dom F) = E.Trg (Cod F)" and "\Map F : \Dom F\ \\<^sub>B \Cod F\\" and "F \ Null" shows "arr F" using assms arr_char by blast lemma arrE [elim]: assumes "arr F" shows "(\ E.Nml (Dom F); E.Ide (Dom F); E.Nml (Cod F); E.Ide (Cod F); E.Src (Dom F) = E.Src (Cod F); E.Trg (Dom F) = E.Trg (Cod F); \Map F : \Dom F\ \\<^sub>B \Cod F\\; F \ Null \ \ T) \ T" using assms arr_char by simp lemma ide_char: shows "ide F \ endo F \ B.ide (Map F)" proof show "ide F \ endo F \ B.ide (Map F)" using ide_char by (simp add: E.ide_eval_Ide) show "endo F \ B.ide (Map F) \ ide F" by (metis (no_types, lifting) B.ide_char B.in_homE arr_char ide_char mem_Collect_eq seq_char) qed lemma ideI [intro]: assumes "arr F" and "Dom F = Cod F" and "B.ide (Map F)" shows "ide F" using assms ide_char dom_char cod_char seq_char by presburger lemma ideE [elim]: assumes "ide F" shows "(\ arr F; Dom F = Cod F; B.ide (Map F); Map F = \Dom F\; Map F = \Cod F\ \ \ T) \ T" proof - assume 1: "\ arr F; Dom F = Cod F; B.ide (Map F); Map F = \Dom F\; Map F = \Cod F\ \ \ T" show T proof - have "arr F" using assms by auto moreover have "Dom F = Cod F" using assms ide_char dom_char cod_char by (metis (no_types, lifting) Dom_cod calculation ideD(3)) moreover have "B.ide (Map F)" using assms ide_char by blast moreover have "Map F = \Dom F\" using assms ide_char dom_char Map_ide(1) by blast ultimately show T using 1 by simp qed qed text \ Source and target are defined by the corresponding syntactic operations on terms. \ definition src where "src F \ if arr F then MkIde (E.Src (Dom F)) else null" definition trg where "trg F \ if arr F then MkIde (E.Trg (Dom F)) else null" lemma src_simps [simp]: assumes "arr F" shows "Dom (src F) = E.Src (Dom F)" and "Cod (src F) = E.Src (Dom F)" and "Map (src F) = \E.Src (Dom F)\" using assms src_def arr_char by auto lemma trg_simps [simp]: assumes "arr F" shows "Dom (trg F) = E.Trg (Dom F)" and "Cod (trg F) = E.Trg (Dom F)" and "Map (trg F) = \E.Trg (Dom F)\" using assms trg_def arr_char by auto interpretation src: endofunctor vcomp src using src_def comp_char apply (unfold_locales) apply auto[4] proof - show "\g f. seq g f \ src (g \ f) = src g \ src f" proof - fix g f assume gf: "seq g f" have "src (g \ f) = MkIde (E.Src (Dom (g \ f)))" using gf src_def comp_char by simp also have "... = MkIde (E.Src (Dom f))" using gf by (simp add: seq_char) also have "... = MkIde (E.Src (Dom g)) \ MkIde (E.Src (Dom f))" using gf seq_char by auto also have "... = src g \ src f" using gf src_def comp_char by auto finally show "src (g \ f) = src g \ src f" by blast qed qed interpretation trg: endofunctor vcomp trg using trg_def comp_char apply (unfold_locales) apply auto[4] proof - show "\g f. seq g f \ trg (g \ f) = trg g \ trg f" proof - fix g f assume gf: "seq g f" have "trg (g \ f) = MkIde (E.Trg (Dom (g \ f)))" using gf trg_def comp_char by simp also have "... = MkIde (E.Trg (Dom f))" using gf by (simp add: seq_char) also have "... = MkIde (E.Trg (Dom g)) \ MkIde (E.Trg (Dom f))" using gf seq_char by auto also have "... = trg g \ trg f" using gf trg_def comp_char by auto finally show "trg (g \ f) = trg g \ trg f" by blast qed qed interpretation horizontal_homs vcomp src trg using src_def trg_def Cod_in_Obj Map_in_Hom by (unfold_locales, auto) notation in_hhom ("\_ : _ \ _\") definition hcomp (infixr "\" 53) where "\ \ \ \ if arr \ \ arr \ \ src \ = trg \ then MkArr (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \) \\<^sub>B (Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \)) else null" lemma arr_hcomp: assumes "arr \" and "arr \" and "src \ = trg \" shows "arr (\ \ \)" proof - have 1: "E.Ide (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \ E.Nml (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \ E.Ide (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) \ E.Nml (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \)" using assms arr_char src_def trg_def E.Ide_HcompNml E.Nml_HcompNml(1) by auto moreover have "\B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \) \\<^sub>B (Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) : \Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \\ \\<^sub>B \Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \\\" proof - have "\Map \ \\<^sub>B Map \ : \Dom \ \<^bold>\ Dom \\ \\<^sub>B \Cod \ \<^bold>\ Cod \\\" using assms arr_char dom_char cod_char src_def trg_def E.eval_simps'(2-3) by auto moreover have "\B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) : \Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \\ \\<^sub>B \Dom \ \<^bold>\ Dom \\\ \ \B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \) : \Cod \ \<^bold>\ Cod \\ \\<^sub>B \Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \\\" using assms 1 arr_char src_def trg_def apply (intro conjI B.in_homI) by auto ultimately show ?thesis by auto qed moreover have "E.Src (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) = E.Src (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) \ E.Trg (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) = E.Trg (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \)" using assms arr_char src_def trg_def by (simp add: E.Src_HcompNml E.Trg_HcompNml) ultimately show ?thesis unfolding hcomp_def using assms by (intro arrI, auto) qed lemma src_hcomp [simp]: assumes "arr \" and "arr \" and "src \ = trg \" shows "src (\ \ \) = src \" using assms arr_char hcomp_def src_def trg_def arr_hcomp E.Src_HcompNml by simp lemma trg_hcomp [simp]: assumes "arr \" and "arr \" and "src \ = trg \" shows "trg (hcomp \ \) = trg \" using assms arr_char hcomp_def src_def trg_def arr_hcomp E.Trg_HcompNml by simp lemma hseq_char: shows "arr (\ \ \) \ arr \ \ arr \ \ src \ = trg \" using arr_hcomp hcomp_def by simp lemma Dom_hcomp [simp]: assumes "arr \" and "arr \" and "src \ = trg \" shows "Dom (\ \ \) = Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \" using assms hcomp_def [of \ \] by simp lemma Cod_hcomp [simp]: assumes "arr \" and "arr \" and "src \ = trg \" shows "Cod (\ \ \) = Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \" using assms hcomp_def [of \ \] by simp lemma Map_hcomp [simp]: assumes "arr \" and "arr \" and "src \ = trg \" shows "Map (\ \ \) = B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \) \\<^sub>B (Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \)" using assms hcomp_def [of \ \] by simp interpretation VxV: product_category vcomp vcomp .. interpretation VV: subcategory VxV.comp \\\\. arr (fst \\) \ arr (snd \\) \ src (fst \\) = trg (snd \\)\ using subcategory_VV by simp interpretation H: "functor" VV.comp vcomp \\\\. hcomp (fst \\) (snd \\)\ proof show "\f. \ VV.arr f \ fst f \ snd f = null" using hcomp_def by auto show A: "\f. VV.arr f \ arr (fst f \ snd f)" using VV.arrE hseq_char by blast show "\f. VV.arr f \ dom (fst f \ snd f) = fst (VV.dom f) \ snd (VV.dom f)" proof - fix f assume f: "VV.arr f" have "dom (fst f \ snd f) = MkIde (Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f))" using f VV.arrE [of f] dom_char arr_hcomp hcomp_def by simp also have "... = fst (VV.dom f) \ snd (VV.dom f)" proof - have "hcomp (fst (VV.dom f)) (snd (VV.dom f)) = MkArr (Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f)) (Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f)) (B.can (Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f)) (Dom (fst f) \<^bold>\ Dom (snd f)) \\<^sub>B (\Dom (fst f)\ \\<^sub>B \Dom (snd f)\) \\<^sub>B B.can (Dom (fst f) \<^bold>\ Dom (snd f)) (Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f)))" using f VV.arrE [of f] arr_hcomp hcomp_def by simp moreover have "B.can (Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f)) (Dom (fst f) \<^bold>\ Dom (snd f)) \\<^sub>B (\Dom (fst f)\ \\<^sub>B \Dom (snd f)\) \\<^sub>B B.can (Dom (fst f) \<^bold>\ Dom (snd f)) (Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f)) = \Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f)\" proof - have 1: "E.Ide (Dom (fst f) \<^bold>\ Dom (snd f))" using f VV.arr_char arr_char dom_char apply simp by (metis (no_types, lifting) src_simps(1) trg_simps(1)) have 2: "E.Ide (Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f))" using f VV.arr_char arr_char dom_char apply simp by (metis (no_types, lifting) E.Ide_HcompNml src_simps(1) trg_simps(1)) have 3: "\<^bold>\Dom (fst f) \<^bold>\ Dom (snd f)\<^bold>\ = \<^bold>\Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f)\<^bold>\" using f VV.arr_char arr_char dom_char apply simp by (metis (no_types, lifting) E.Nml_HcompNml(1) E.Nmlize_Nml src_simps(1) trg_simps(1)) have "(\Dom (fst f)\ \\<^sub>B \Dom (snd f)\) \\<^sub>B B.can (Dom (fst f) \<^bold>\ Dom (snd f)) (Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f)) = B.can (Dom (fst f) \<^bold>\ Dom (snd f)) (Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f))" proof - have "B.in_hom (B.can (Dom (fst f) \<^bold>\ Dom (snd f)) (Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f))) \Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f)\ (\Dom (fst f)\ \\<^sub>B \Dom (snd f)\)" using 1 2 3 f VV.arr_char arr_char B.can_in_hom [of "Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f)" "Dom (fst f) \<^bold>\ Dom (snd f)"] by simp thus ?thesis using B.comp_cod_arr by auto qed thus ?thesis using 1 2 3 f VV.arr_char B.can_Ide_self B.vcomp_can by simp qed ultimately show ?thesis by simp qed finally show "dom (fst f \ snd f) = fst (VV.dom f) \ snd (VV.dom f)" by simp qed show "\f. VV.arr f \ cod (fst f \ snd f) = fst (VV.cod f) \ snd (VV.cod f)" proof - fix f assume f: "VV.arr f" have "cod (fst f \ snd f) = MkIde (Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f))" using f VV.arrE [of f] cod_char arr_hcomp hcomp_def by simp also have "... = fst (VV.cod f) \ snd (VV.cod f)" proof - have "hcomp (fst (VV.cod f)) (snd (VV.cod f)) = MkArr (Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f)) (Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f)) (B.can (Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f)) (Cod (fst f) \<^bold>\ Cod (snd f)) \\<^sub>B (\Cod (fst f)\ \\<^sub>B \Cod (snd f)\) \\<^sub>B B.can (Cod (fst f) \<^bold>\ Cod (snd f)) (Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f)))" using f VV.arrE [of f] arr_hcomp hcomp_def by simp moreover have "B.can (Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f)) (Cod (fst f) \<^bold>\ Cod (snd f)) \\<^sub>B (\Cod (fst f)\ \\<^sub>B \Cod (snd f)\) \\<^sub>B B.can (Cod (fst f) \<^bold>\ Cod (snd f)) (Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f)) = \Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f)\" proof - have 1: "E.Ide (Cod (fst f) \<^bold>\ Cod (snd f))" using f VV.arr_char arr_char dom_char apply simp by (metis (no_types, lifting) src_simps(1) trg_simps(1)) have 2: "E.Ide (Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f))" using f VV.arr_char arr_char dom_char apply simp by (metis (no_types, lifting) E.Ide_HcompNml src_simps(1) trg_simps(1)) have 3: "\<^bold>\Cod (fst f) \<^bold>\ Cod (snd f)\<^bold>\ = \<^bold>\Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f)\<^bold>\" using f VV.arr_char arr_char dom_char apply simp by (metis (no_types, lifting) E.Nml_HcompNml(1) E.Nmlize_Nml src_simps(1) trg_simps(1)) have "(\Cod (fst f)\ \\<^sub>B \Cod (snd f)\) \\<^sub>B B.can (Cod (fst f) \<^bold>\ Cod (snd f)) (Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f)) = B.can (Cod (fst f) \<^bold>\ Cod (snd f)) (Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f))" proof - have "B.in_hom (B.can (Cod (fst f) \<^bold>\ Cod (snd f)) (Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f))) \Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f)\ (\Cod (fst f)\ \\<^sub>B \Cod (snd f)\)" using 1 2 3 f VV.arr_char arr_char B.can_in_hom [of "Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f)" "Cod (fst f) \<^bold>\ Cod (snd f)"] by simp thus ?thesis using B.comp_cod_arr by auto qed thus ?thesis using 1 2 3 f VV.arr_char B.can_Ide_self B.vcomp_can by simp qed ultimately show ?thesis by simp qed finally show "cod (fst f \ snd f) = fst (VV.cod f) \ snd (VV.cod f)" by simp qed show "\g f. VV.seq g f \ fst (VV.comp g f) \ snd (VV.comp g f) = (fst g \ snd g) \ (fst f \ snd f)" proof - fix f g assume fg: "VV.seq g f" have f: "arr (fst f) \ arr (snd f) \ src (fst f) = trg (snd f)" using fg VV.seq_char VV.arr_char by simp have g: "arr (fst g) \ arr (snd g) \ src (fst g) = trg (snd g)" using fg VV.seq_char VV.arr_char by simp have 1: "arr (fst (VV.comp g f)) \ arr (snd (VV.comp g f)) \ src (fst (VV.comp g f)) = trg (snd (VV.comp g f))" using fg VV.arrE by blast have 0: "VV.comp g f = (fst g \ fst f, snd g \ snd f)" using fg 1 VV.comp_char VxV.comp_char by (metis (no_types, lifting) VV.seq_char VxV.seqE) let ?X = "MkArr (Dom (fst (VV.comp g f)) \<^bold>\\<^bold>\\<^bold>\ Dom (snd (VV.comp g f))) (Cod (fst (VV.comp g f)) \<^bold>\\<^bold>\\<^bold>\ Cod (snd (VV.comp g f))) (B.can (Cod (fst (VV.comp g f)) \<^bold>\\<^bold>\\<^bold>\ Cod (snd (VV.comp g f))) (Cod (fst (VV.comp g f)) \<^bold>\ Cod (snd (VV.comp g f))) \\<^sub>B (Map (fst (VV.comp g f)) \\<^sub>B Map (snd (VV.comp g f))) \\<^sub>B B.can (Dom (fst (VV.comp g f)) \<^bold>\ Dom (snd (VV.comp g f))) (Dom (fst (VV.comp g f)) \<^bold>\\<^bold>\\<^bold>\ Dom (snd (VV.comp g f))))" have 2: "fst (VV.comp g f) \ snd (VV.comp g f) = ?X" unfolding hcomp_def using 1 by simp also have "... = (fst g \ snd g) \ (fst f \ snd f)" proof - let ?GG = "MkArr (Dom (fst g) \<^bold>\\<^bold>\\<^bold>\ Dom (snd g)) (Cod (fst g) \<^bold>\\<^bold>\\<^bold>\ Cod (snd g)) (B.can (Cod (fst g) \<^bold>\\<^bold>\\<^bold>\ Cod (snd g)) (Cod (fst g) \<^bold>\ Cod (snd g)) \\<^sub>B (Map (fst g) \\<^sub>B Map (snd g)) \\<^sub>B B.can (Dom (fst g) \<^bold>\ Dom (snd g)) (Dom (fst g) \<^bold>\\<^bold>\\<^bold>\ Dom (snd g)))" let ?FF = "MkArr (Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f)) (Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f)) (B.can (Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f)) (Cod (fst f) \<^bold>\ Cod (snd f)) \\<^sub>B (Map (fst f) \\<^sub>B Map (snd f)) \\<^sub>B B.can (Dom (fst f) \<^bold>\ Dom (snd f)) (Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f)))" have 4: "arr ?FF \ arr ?GG \ Dom ?GG = Cod ?FF" proof - have "arr ?FF \ arr ?GG" using f g fg VV.arr_char VV.seqE hcomp_def A by presburger thus ?thesis using 0 1 by (simp add: fg seq_char) qed have "(fst g \ snd g) \ (fst f \ snd f) = ?GG \ ?FF" unfolding hcomp_def using 1 f g fg VV.arr_char VV.seqE by simp also have "... = ?X" proof (intro arr_eqI) show "seq ?GG ?FF" using fg 4 seq_char by blast show "arr ?X" using fg 1 arr_hcomp hcomp_def by simp show "Dom (?GG \ ?FF) = Dom ?X" using fg 0 1 4 seq_char by simp show "Cod (?GG \ ?FF) = Cod ?X" using fg 0 1 4 seq_char by simp show "Map (?GG \ ?FF) = Map ?X" proof - have "Map (?GG \ ?FF) = Map ?GG \\<^sub>B Map ?FF" using 4 by auto also have "... = (B.can (Cod (fst g) \<^bold>\\<^bold>\\<^bold>\ Cod (snd g)) (Cod (fst g) \<^bold>\ Cod (snd g)) \\<^sub>B (Map (fst g) \\<^sub>B Map (snd g)) \\<^sub>B B.can (Dom (fst g) \<^bold>\ Dom (snd g)) (Dom (fst g) \<^bold>\\<^bold>\\<^bold>\ Dom (snd g))) \\<^sub>B (B.can (Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f)) (Cod (fst f) \<^bold>\ Cod (snd f)) \\<^sub>B (Map (fst f) \\<^sub>B Map (snd f)) \\<^sub>B B.can (Dom (fst f) \<^bold>\ Dom (snd f)) (Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f)))" using fg by simp also have "... = B.can (Cod (fst g) \<^bold>\\<^bold>\\<^bold>\ Cod (snd g)) (Cod (fst g) \<^bold>\ Cod (snd g)) \\<^sub>B ((Map (fst g) \\<^sub>B Map (snd g)) \\<^sub>B (Map (fst f) \\<^sub>B Map (snd f))) \\<^sub>B B.can (Dom (fst f) \<^bold>\ Dom (snd f)) (Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f))" proof - have "(B.can (Cod (fst g) \<^bold>\\<^bold>\\<^bold>\ Cod (snd g)) (Cod (fst g) \<^bold>\ Cod (snd g)) \\<^sub>B (Map (fst g) \\<^sub>B Map (snd g)) \\<^sub>B B.can (Dom (fst g) \<^bold>\ Dom (snd g)) (Dom (fst g) \<^bold>\\<^bold>\\<^bold>\ Dom (snd g))) \\<^sub>B (B.can (Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f)) (Cod (fst f) \<^bold>\ Cod (snd f)) \\<^sub>B (Map (fst f) \\<^sub>B Map (snd f)) \\<^sub>B B.can (Dom (fst f) \<^bold>\ Dom (snd f)) (Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f))) = B.can (Cod (fst g) \<^bold>\\<^bold>\\<^bold>\ Cod (snd g)) (Cod (fst g) \<^bold>\ Cod (snd g)) \\<^sub>B ((Map (fst g) \\<^sub>B Map (snd g)) \\<^sub>B (B.can (Dom (fst g) \<^bold>\ Dom (snd g)) (Dom (fst g) \<^bold>\\<^bold>\\<^bold>\ Dom (snd g)) \\<^sub>B B.can (Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f)) (Cod (fst f) \<^bold>\ Cod (snd f))) \\<^sub>B (Map (fst f) \\<^sub>B Map (snd f))) \\<^sub>B B.can (Dom (fst f) \<^bold>\ Dom (snd f)) (Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f))" using B.comp_assoc by simp also have "... = B.can (Cod (fst g) \<^bold>\\<^bold>\\<^bold>\ Cod (snd g)) (Cod (fst g) \<^bold>\ Cod (snd g)) \\<^sub>B ((Map (fst g) \\<^sub>B Map (snd g)) \\<^sub>B (Map (fst f) \\<^sub>B Map (snd f))) \\<^sub>B B.can (Dom (fst f) \<^bold>\ Dom (snd f)) (Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f))" proof - have "(B.can (Dom (fst g) \<^bold>\ Dom (snd g)) (Dom (fst g) \<^bold>\\<^bold>\\<^bold>\ Dom (snd g))) \\<^sub>B (B.can (Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f)) (Cod (fst f) \<^bold>\ Cod (snd f))) = \Cod (fst f) \<^bold>\ Cod (snd f)\" proof - have "(B.can (Dom (fst g) \<^bold>\ Dom (snd g)) (Dom (fst g) \<^bold>\\<^bold>\\<^bold>\ Dom (snd g))) \\<^sub>B (B.can (Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f)) (Cod (fst f) \<^bold>\ Cod (snd f))) = B.can (Dom (fst g) \<^bold>\ Dom (snd g)) (Cod (fst f) \<^bold>\ Cod (snd f))" proof - have "E.Ide (Dom (fst g) \<^bold>\ Dom (snd g))" using g arr_char apply simp by (metis (no_types, lifting) src_simps(1) trg_simps(1)) moreover have "E.Ide (Dom (fst g) \<^bold>\\<^bold>\\<^bold>\ Dom (snd g))" using g arr_char apply simp by (metis (no_types, lifting) E.Ide_HcompNml src_simps(1) trg_simps(1)) moreover have "E.Ide (Cod (fst f) \<^bold>\ Cod (snd f))" using f arr_char apply simp by (metis (no_types, lifting) src_simps(1) trg_simps(1)) moreover have "\<^bold>\Dom (fst g) \<^bold>\ Dom (snd g)\<^bold>\ = \<^bold>\Dom (fst g) \<^bold>\\<^bold>\\<^bold>\ Dom (snd g)\<^bold>\" using g apply simp by (metis (no_types, lifting) E.Nml_HcompNml(1) E.Nmlize_Nml arrE src_simps(1) trg_simps(1)) moreover have "\<^bold>\Dom (fst g) \<^bold>\\<^bold>\\<^bold>\ Dom (snd g)\<^bold>\ = \<^bold>\Cod (fst f) \<^bold>\ Cod (snd f)\<^bold>\" using g apply simp by (metis (no_types, lifting) "0" "1" E.Nmlize.simps(3) calculation(4) fst_conv seq_char snd_conv) moreover have "Dom (fst g) \<^bold>\\<^bold>\\<^bold>\ Dom (snd g) = Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f)" using 0 1 by (simp add: seq_char) ultimately show ?thesis using B.vcomp_can by simp qed also have "... = \Cod (fst f) \<^bold>\ Cod (snd f)\" proof - have "Dom (fst g) \<^bold>\ Dom (snd g) = Cod (fst f) \<^bold>\ Cod (snd f)" using 0 f g fg seq_char VV.seq_char VV.arr_char by simp thus ?thesis using f B.can_Ide_self [of "Dom (fst f) \<^bold>\ Dom (snd f)"] apply simp by (metis (no_types, lifting) B.can_Ide_self E.eval.simps(3) E.Ide.simps(3) arr_char src_simps(2) trg_simps(2)) qed finally show ?thesis by simp qed hence "(B.can (Dom (fst g) \<^bold>\ Dom (snd g)) (Dom (fst g) \<^bold>\\<^bold>\\<^bold>\ Dom (snd g)) \\<^sub>B B.can (Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f)) (Cod (fst f) \<^bold>\ Cod (snd f))) \\<^sub>B (Map (fst f) \\<^sub>B Map (snd f)) = \Cod (fst f) \<^bold>\ Cod (snd f)\ \\<^sub>B (Map (fst f) \\<^sub>B Map (snd f))" by simp also have "... = Map (fst f) \\<^sub>B Map (snd f)" proof - have 1: "\p. arr p \ map (cod p) \ map p = map p" by blast have 3: "\Cod (fst f)\ \\<^sub>B Map (fst f) = Map (map (cod (fst f)) \ map (fst f))" by (simp add: f) have 4: "map (cod (fst f)) \ map (fst f) = fst f" using 1 f map_simp by simp show ?thesis proof - have 2: "\Cod (snd f)\ \\<^sub>B Map (snd f) = Map (snd f)" proof - have "\Cod (snd f)\ \\<^sub>B Map (snd f) = Map (map (cod (snd f)) \ map (snd f))" by (simp add: f) moreover have "map (cod (snd f)) \ map (snd f) = snd f" using 1 f map_simp by simp ultimately show ?thesis by presburger qed have "B.seq \Cod (snd f)\ (Map (snd f))" using f 2 by auto moreover have "B.seq \Cod (fst f)\ (Map (fst f))" using 4 f 3 by auto moreover have "\Cod (fst f)\ \\<^sub>B Map (fst f) \\<^sub>B \Cod (snd f)\ \\<^sub>B Map (snd f) = Map (fst f) \\<^sub>B Map (snd f)" using 2 3 4 by presburger ultimately show ?thesis by (simp add: B.interchange) qed qed finally have "(B.can (Dom (fst g) \<^bold>\ Dom (snd g)) (Dom (fst g) \<^bold>\\<^bold>\\<^bold>\ Dom (snd g)) \\<^sub>B B.can (Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f)) (Cod (fst f) \<^bold>\ Cod (snd f))) \\<^sub>B (Map (fst f) \\<^sub>B Map (snd f)) = Map (fst f) \\<^sub>B Map (snd f)" by simp thus ?thesis using fg B.comp_cod_arr by simp qed finally show ?thesis by simp qed also have "... = B.can (Cod (fst g) \<^bold>\\<^bold>\\<^bold>\ Cod (snd g)) (Cod (fst g) \<^bold>\ Cod (snd g)) \\<^sub>B (Map (fst g \ fst f) \\<^sub>B Map (snd g \ snd f)) \\<^sub>B B.can (Dom (fst f) \<^bold>\ Dom (snd f)) (Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f))" proof - have 2: "Dom (fst g) = Cod (fst f)" using 0 f g fg VV.seq_char [of g f] VV.arr_char arr_char seq_char by (metis (no_types, lifting) fst_conv) hence "Map (fst g \ fst f) = Map (fst g) \\<^sub>B Map (fst f)" using f g Map_comp [of "fst f" "fst g"] by simp moreover have "B.seq (Map (fst g)) (Map (fst f)) \ B.seq (Map (snd g)) (Map (snd f))" using f g 0 1 2 arr_char by (metis (no_types, lifting) B.seqI' prod.sel(2) seq_char) ultimately show ?thesis using 0 1 seq_char Map_comp B.interchange by auto qed also have "... = Map ?X" using fg 0 1 by (simp add: seq_char) finally show ?thesis by simp qed qed finally show ?thesis by simp qed finally show "fst (VV.comp g f) \ snd (VV.comp g f) = (fst g \ snd g) \ (fst f \ snd f)" by simp qed qed interpretation H: horizontal_composition vcomp hcomp src trg using hseq_char by (unfold_locales, auto) lemma hcomp_assoc: assumes "arr \" and "arr \" and "arr \" and "src \ = trg \" and "src \ = trg \" shows "(\ \ \) \ \ = \ \ \ \ \" proof (intro arr_eqI) have \\: "\Map \ \\<^sub>B Map \ : \Dom \ \<^bold>\ Dom \\ \\<^sub>B \Cod \ \<^bold>\ Cod \\\" using assms src_def trg_def arr_char by (auto simp add: E.eval_simps'(2-3) Pair_inject) have \\: "\Map \ \\<^sub>B Map \ : \Dom \ \<^bold>\ Dom \\ \\<^sub>B \Cod \ \<^bold>\ Cod \\\" using assms src_def trg_def arr_char by (auto simp add: E.eval_simps'(2-3) Pair_inject) show "H.hseq (\ \ \) \" using assms \\ \\ by auto show "H.hseq \ (\ \ \)" using assms \\ \\ by auto show "Dom ((\ \ \) \ \) = Dom (\ \ \ \ \)" proof - have "E.Nml (Dom \) \ E.Nml (Dom \) \ E.Nml (Dom \)" using assms by blast moreover have "E.Src (Dom \) = E.Trg (Dom \) \ E.Src (Dom \) = E.Trg (Dom \)" using assms \\ \\ by (metis (no_types, lifting) src_simps(2) trg_simps(2)) ultimately show ?thesis using assms \\ \\ E.HcompNml_assoc by simp qed show "Cod ((\ \ \) \ \) = Cod (\ \ \ \ \)" proof - have "E.Nml (Cod \) \ E.Nml (Cod \) \ E.Nml (Cod \)" using assms by blast moreover have "E.Src (Cod \) = E.Trg (Cod \) \ E.Src (Cod \) = E.Trg (Cod \)" using assms \\ \\ by (metis (no_types, lifting) arrE src_simps(2) trg_simps(2)) ultimately show ?thesis using assms \\ \\ E.HcompNml_assoc by simp qed show "Map ((\ \ \) \ \) = Map (\ \ \ \ \)" proof - have "Map ((\ \ \) \ \) = B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \ \<^bold>\ Cod \) \\<^sub>B (Map \ \\<^sub>B Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \)" proof - have 1: "Map ((\ \ \) \ \) = B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) ((Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) \<^bold>\ Cod \) \\<^sub>B (B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \) \\<^sub>B (Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \\<^sub>B Map \) \\<^sub>B B.can ((Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \)" using assms \\ \\ hcomp_def E.HcompNml_assoc src_def trg_def arr_char E.Nml_HcompNml E.Ide_HcompNml by auto (* 5 sec *) also have "... = B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) ((Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) \<^bold>\ Cod \) \\<^sub>B (B.can ((Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) \<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \ \<^bold>\ Cod \) \\<^sub>B (Map \ \\<^sub>B Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \ \<^bold>\ Dom \) ((Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \<^bold>\ Dom \)) \\<^sub>B B.can ((Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \)" proof - have "B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \) \\<^sub>B (Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \\<^sub>B Map \ = B.can ((Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) \<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \ \<^bold>\ Cod \) \\<^sub>B (Map \ \\<^sub>B Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \ \<^bold>\ Dom \) ((Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \<^bold>\ Dom \)" proof - have "B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \) \\<^sub>B (Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \\<^sub>B Map \ = (B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \) \\<^sub>B B.can (Cod \) (Cod \)) \\<^sub>B ((Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \\<^sub>B Map \)" proof - have "B.seq (B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \)) ((Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \))" by (metis (no_types, lifting) B.arrI Map_hcomp arrE arr_hcomp assms(1) assms(2) assms(4)) moreover have "B.seq (B.can (Cod \) (Cod \)) (Map \)" using B.can_in_hom assms(3) by blast moreover have "B.ide (B.can (Cod \) (Cod \))" using B.can_Ide_self E.ide_eval_Ide arr_char assms(3) by presburger ultimately show ?thesis by (metis (no_types) B.comp_ide_arr B.interchange) qed also have "... = (B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \) \\<^sub>B B.can (Cod \) (Cod \)) \\<^sub>B ((Map \ \\<^sub>B Map \) \\<^sub>B Map \) \\<^sub>B (B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \\<^sub>B B.can (Dom \) (Dom \))" proof - have "B.seq (Map \ \\<^sub>B Map \) (B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \))" by (metis (no_types, lifting) B.arrI B.comp_null(2) B.ext Map_hcomp arrE arr_hcomp assms(1) assms(2) assms(4)) moreover have "B.seq (Map \) (B.can (Dom \) (Dom \))" using assms(3) by fastforce ultimately show ?thesis using B.interchange by (metis (no_types, lifting) B.can_Ide_self B.comp_arr_ide E.ide_eval_Ide arrE assms(3)) qed also have "... = (B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \) \\<^sub>B B.can (Cod \) (Cod \)) \\<^sub>B (B.can ((Cod \ \<^bold>\ Cod \) \<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \ \<^bold>\ Cod \) \\<^sub>B (Map \ \\<^sub>B Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \ \<^bold>\ Dom \) ((Dom \ \<^bold>\ Dom \) \<^bold>\ Dom \)) \\<^sub>B (B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \\<^sub>B B.can (Dom \) (Dom \))" proof - have "(Map \ \\<^sub>B Map \) \\<^sub>B Map \ = B.\' \Cod \\ \Cod \\ \Cod \\ \\<^sub>B (Map \ \\<^sub>B Map \ \\<^sub>B Map \) \\<^sub>B \\<^sub>B \Dom \\ \Dom \\ \Dom \\" using B.hcomp_reassoc(1) by (metis (no_types, lifting) B.hcomp_in_vhomE B.in_homE \\ \\ arrE assms(1) assms(2) assms(3)) also have "... = B.can ((Cod \ \<^bold>\ Cod \) \<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \ \<^bold>\ Cod \) \\<^sub>B (Map \ \\<^sub>B Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \ \<^bold>\ Dom \) ((Dom \ \<^bold>\ Dom \) \<^bold>\ Dom \)" using assms arr_char src_def trg_def arr_char B.canE_associator by simp finally show ?thesis by simp qed also have "... = ((B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \) \\<^sub>B B.can (Cod \) (Cod \)) \\<^sub>B (B.can ((Cod \ \<^bold>\ Cod \) \<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \ \<^bold>\ Cod \))) \\<^sub>B (Map \ \\<^sub>B Map \ \\<^sub>B Map \) \\<^sub>B (B.can (Dom \ \<^bold>\ Dom \ \<^bold>\ Dom \) ((Dom \ \<^bold>\ Dom \) \<^bold>\ Dom \) \\<^sub>B (B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \\<^sub>B B.can (Dom \) (Dom \)))" using B.comp_assoc by simp also have "... = B.can ((Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) \<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \ \<^bold>\ Cod \) \\<^sub>B (Map \ \\<^sub>B Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \ \<^bold>\ Dom \) ((Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \<^bold>\ Dom \)" proof - have "(B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \) \\<^sub>B B.can (Cod \) (Cod \)) \\<^sub>B (B.can ((Cod \ \<^bold>\ Cod \) \<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \ \<^bold>\ Cod \)) = B.can ((Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) \<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \ \<^bold>\ Cod \)" proof - have "E.Ide (Cod \ \<^bold>\ Cod \)" by (metis (no_types, lifting) E.Ide.simps(3) arrE assms(1-2,4) src_simps(1) trg_simps(1)) moreover have "E.Ide (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \)" using E.Ide_HcompNml assms(1) assms(2) calculation by auto moreover have "\<^bold>\Cod \ \<^bold>\ Cod \\<^bold>\ = \<^bold>\Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \\<^bold>\" using E.Nml_HcompNml(1) assms(1) assms(2) calculation(1) by fastforce moreover have "E.Src (Cod \ \<^bold>\ Cod \) = E.Trg (Cod \)" by (metis (no_types, lifting) E.Src.simps(3) arrE assms(2-3,5) src_simps(2) trg_simps(2)) moreover have "E.Src (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) = E.Trg (Cod \)" using E.Src_HcompNml assms(1) assms(2) calculation(1) calculation(4) by fastforce moreover have "\<^bold>\Cod \ \<^bold>\ Cod \ \<^bold>\ Cod \\<^bold>\ = \<^bold>\(Cod \ \<^bold>\ Cod \) \<^bold>\ Cod \\<^bold>\" by (metis (no_types, lifting) E.Arr.simps(3) E.Nmlize_Hcomp_Hcomp E.Nmlize_Hcomp_Hcomp' E.Ide_implies_Arr E.Src.simps(3) arrE assms(3) calculation(1) calculation(4)) ultimately show ?thesis using assms(3) B.hcomp_can B.vcomp_can by auto qed moreover have "B.can (Dom \ \<^bold>\ Dom \ \<^bold>\ Dom \) ((Dom \ \<^bold>\ Dom \) \<^bold>\ Dom \) \\<^sub>B (B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \\<^sub>B B.can (Dom \) (Dom \)) = B.can (Dom \ \<^bold>\ Dom \ \<^bold>\ Dom \) ((Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \<^bold>\ Dom \)" proof - have "E.Ide (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \)" by (metis (no_types, lifting) E.Ide_HcompNml arrE assms(1-2,4) src_simps(2) trg_simps(2)) moreover have "E.Ide (Dom \ \<^bold>\ Dom \)" by (metis (no_types, lifting) E.Ide.simps(3) arrE assms(1-2,4) src_simps(1) trg_simps(1)) moreover have "\<^bold>\Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \\<^bold>\ = \<^bold>\Dom \ \<^bold>\ Dom \\<^bold>\" using E.Nml_HcompNml(1) assms(1-2) calculation(2) by fastforce moreover have "E.Src (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) = E.Trg (Dom \)" by (metis (no_types, lifting) E.Ide.simps(3) E.Src_HcompNml arrE assms(1-3,5) calculation(2) src_simps(2) trg_simps(2)) moreover have "E.Src (Dom \ \<^bold>\ Dom \) = E.Trg (Dom \)" using E.Src_HcompNml assms(1-2) calculation(2) calculation(4) by fastforce moreover have "E.Ide ((Dom \ \<^bold>\ Dom \) \<^bold>\ Dom \)" using E.Ide.simps(3) assms(3) calculation(2) calculation(5) by blast moreover have "\<^bold>\(Dom \ \<^bold>\ Dom \) \<^bold>\ Dom \\<^bold>\ = \<^bold>\Dom \ \<^bold>\ Dom \ \<^bold>\ Dom \\<^bold>\" using E.Nmlize_Hcomp_Hcomp calculation(6) by auto ultimately show ?thesis using assms(3) B.hcomp_can B.vcomp_can by auto qed ultimately show ?thesis by simp qed finally show ?thesis by simp qed thus ?thesis by simp qed also have "... = (B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) ((Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) \<^bold>\ Cod \) \\<^sub>B B.can ((Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) \<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \ \<^bold>\ Cod \)) \\<^sub>B (Map \ \\<^sub>B Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \ \<^bold>\ Dom \) ((Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \<^bold>\ Dom \) \\<^sub>B B.can ((Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \)" using B.comp_assoc by simp also have "... = B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \ \<^bold>\ Cod \) \\<^sub>B (Map \ \\<^sub>B Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \)" proof - have "B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) ((Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) \<^bold>\ Cod \) \\<^sub>B B.can ((Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) \<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \ \<^bold>\ Cod \) = B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \ \<^bold>\ Cod \)" proof - have "E.Ide (Cod \ \<^bold>\ Cod \ \<^bold>\ Cod \)" using assms src_def trg_def by fastforce moreover have "E.Ide ((Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) \<^bold>\ Cod \)" using assms arr_char src_def trg_def E.Ide_HcompNml E.Src_HcompNml by auto moreover have "E.Ide (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \)" using assms arr_char src_def trg_def by (simp add: E.Nml_HcompNml(1) E.Ide_HcompNml E.Trg_HcompNml) moreover have "\<^bold>\Cod \ \<^bold>\ Cod \ \<^bold>\ Cod \\<^bold>\ = \<^bold>\(Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) \<^bold>\ Cod \\<^bold>\" using assms arr_char src_def trg_def E.Nml_HcompNml E.HcompNml_assoc by simp moreover have "\<^bold>\(Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) \<^bold>\ Cod \\<^bold>\ = \<^bold>\Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \\<^bold>\" using assms arr_char src_def trg_def E.Nml_HcompNml E.HcompNml_assoc by simp ultimately show ?thesis by simp qed moreover have "B.can (Dom \ \<^bold>\ Dom \ \<^bold>\ Dom \) ((Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \<^bold>\ Dom \) \\<^sub>B B.can ((Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) = B.can (Dom \ \<^bold>\ Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \)" proof - have "E.Ide (Dom \ \<^bold>\ Dom \ \<^bold>\ Dom \)" using assms src_def trg_def by fastforce moreover have "E.Ide ((Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \<^bold>\ Dom \)" using assms arr_char src_def trg_def E.Ide_HcompNml E.Src_HcompNml by auto moreover have "E.Ide (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \)" using assms arr_char src_def trg_def by (simp add: E.Nml_HcompNml(1) E.Ide_HcompNml E.Trg_HcompNml) moreover have "\<^bold>\Dom \ \<^bold>\ Dom \ \<^bold>\ Dom \\<^bold>\ = \<^bold>\(Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \<^bold>\ Dom \\<^bold>\" using assms arr_char src_def trg_def E.Nml_HcompNml E.HcompNml_assoc by simp moreover have "\<^bold>\(Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \<^bold>\ Dom \\<^bold>\ = \<^bold>\Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \\<^bold>\" using assms arr_char src_def trg_def E.Nml_HcompNml E.HcompNml_assoc by simp ultimately show ?thesis by simp qed ultimately show ?thesis by simp qed finally show ?thesis by simp qed also have "... = Map (\ \ \ \ \)" proof - have 1: "Map (\ \ \ \ \) = B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) \\<^sub>B (Map \ \\<^sub>B B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \) \\<^sub>B (Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \)) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \)" using assms Map_hcomp [of \ "\ \ \"] Map_hcomp [of \ \] by simp also have "... = B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) \\<^sub>B ((B.can (Cod \) (Cod \) \\<^sub>B B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \)) \\<^sub>B (Map \ \\<^sub>B Map \ \\<^sub>B Map \) \\<^sub>B (B.can (Dom \) (Dom \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \))) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \)" proof - have "Map \ \\<^sub>B B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \) \\<^sub>B (Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) = (B.can (Cod \) (Cod \) \\<^sub>B B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \)) \\<^sub>B (Map \ \\<^sub>B (Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \))" using assms B.interchange B.comp_cod_arr by (metis (no_types, lifting) B.can_Ide_self B.in_homE Map_hcomp arrE hseq_char) also have "... = (B.can (Cod \) (Cod \) \\<^sub>B B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \)) \\<^sub>B (Map \ \\<^sub>B Map \ \\<^sub>B Map \) \\<^sub>B (B.can (Dom \) (Dom \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \))" using assms B.interchange B.comp_arr_dom [of "Map \" "B.can (Dom \) (Dom \)"] by (metis (no_types, lifting) B.can_Ide_self B.comp_null(2) B.ext B.in_homE Map_hcomp arrE hseq_char) finally have "Map \ \\<^sub>B B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \) \\<^sub>B (Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) = (B.can (Cod \) (Cod \) \\<^sub>B B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \)) \\<^sub>B (Map \ \\<^sub>B Map \ \\<^sub>B Map \) \\<^sub>B (B.can (Dom \) (Dom \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \))" by simp thus ?thesis by simp qed also have "... = (B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) \\<^sub>B (B.can (Cod \) (Cod \) \\<^sub>B B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \))) \\<^sub>B (Map \ \\<^sub>B Map \ \\<^sub>B Map \) \\<^sub>B ((B.can (Dom \) (Dom \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \)) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \))" using B.comp_assoc by simp also have "... = B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \ \<^bold>\ Cod \) \\<^sub>B (Map \ \\<^sub>B Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \)" using assms \\ \\ E.HcompNml_assoc src_def trg_def arr_char E.Src_HcompNml E.Trg_HcompNml E.Nml_HcompNml E.Ide_HcompNml by simp finally show ?thesis by simp qed ultimately show ?thesis by metis qed qed lemma obj_char: shows "obj a \ endo a \ E.Obj (Dom a) \ Map a = \Dom a\" proof assume a: "obj a" show "endo a \ E.Obj (Dom a) \ Map a = \Dom a\" proof (intro conjI) show "endo a" using a ide_char by blast show "E.Obj (Dom a)" using a ide_char src_def by (metis (no_types, lifting) E.Ide_implies_Arr E.Obj_Trg arrE obj_def trg_simps(1) trg_src) show "Map a = \Dom a\" using a ide_char src_def by blast qed next assume a: "endo a \ E.Obj (Dom a) \ Map a = \Dom a\" show "obj a" proof - have "arr a" using a by auto moreover have "src a = a" using a E.Obj_in_Hom(1) seq_char by (intro arr_eqI, auto) ultimately show ?thesis using obj_def by simp qed qed lemma hcomp_obj_self: assumes "obj a" shows "a \ a = a" proof (intro arr_eqI) show "H.hseq a a" using assms by auto show "arr a" using assms by auto show 1: "Dom (a \ a) = Dom a" unfolding hcomp_def using assms arr_char E.HcompNml_Trg_Nml apply simp by (metis (no_types, lifting) objE obj_def trg_simps(1)) show 2: "Cod (a \ a) = Cod a" unfolding hcomp_def using assms 1 arr_char E.HcompNml_Trg_Nml apply simp by (metis (no_types, lifting) Dom_hcomp ideE objE) show "Map (a \ a) = Map a" proof - have "Map (a \ a) = B.can (Cod a \<^bold>\\<^bold>\\<^bold>\ Cod a) (Cod a \<^bold>\ Cod a) \\<^sub>B (Map a \\<^sub>B Map a) \\<^sub>B B.can (Dom a \<^bold>\ Dom a) (Dom a \<^bold>\\<^bold>\\<^bold>\ Dom a)" using assms Map_hcomp by auto also have "... = B.can (Dom a) (Dom a \<^bold>\ Dom a) \\<^sub>B (\Dom a\ \\<^sub>B \Dom a\) \\<^sub>B B.can (Dom a \<^bold>\ Dom a) (Dom a)" proof - have "Dom a \<^bold>\\<^bold>\\<^bold>\ Dom a = Dom a" using assms obj_char arr_char E.HcompNml_Trg_Nml by (metis (no_types, lifting) ideE objE obj_def' trg_simps(2)) moreover have "Cod a = Dom a" using assms obj_char arr_char dom_char cod_char objE ide_char' by (metis (no_types, lifting) src_simps(1) src_simps(2)) moreover have "Map a = \Dom a\" using assms obj_char by simp ultimately show ?thesis by simp qed also have "... = B.can (Dom a) (Dom a \<^bold>\ Dom a) \\<^sub>B B.can (Dom a \<^bold>\ Dom a) (Dom a)" using assms obj_char arr_char B.comp_cod_arr E.ide_eval_Ide B.can_in_hom by (metis (no_types, lifting) H.ide_hcomp obj_def obj_def' calculation B.comp_ide_arr B.ide_hcomp B.hseqE B.ideD(1) ide_char B.seqE) also have "... = \Dom a\" using assms 1 2 obj_char arr_char B.vcomp_can calculation H.ide_hcomp ideE objE by (metis (no_types, lifting)) also have "... = Map a" using assms obj_char by simp finally show ?thesis by simp qed qed lemma hcomp_ide_src: assumes "ide f" shows "f \ src f = f" proof (intro arr_eqI) show "H.hseq f (src f)" using assms by simp show "arr f" using assms by simp show 1: "Dom (f \ src f) = Dom f" unfolding hcomp_def using assms apply simp using assms ide_char arr_char E.HcompNml_Nml_Src by (metis (no_types, lifting) ideD(1)) show "Cod (f \ src f) = Cod f" unfolding hcomp_def using assms apply simp using assms ide_char arr_char E.HcompNml_Nml_Src by (metis (no_types, lifting) ideD(1)) show "Map (f \ src f) = Map f" proof - have "Map (f \ src f) = B.can (Cod f \<^bold>\\<^bold>\\<^bold>\ Cod (src f)) (Cod f \<^bold>\ Cod (src f)) \\<^sub>B (Map f \\<^sub>B Map (src f)) \\<^sub>B B.can (Dom f \<^bold>\ Dom (src f)) (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom (src f))" unfolding hcomp_def using assms by simp also have "... = B.can (Dom f) (Dom f \<^bold>\ E.Src (Dom f)) \\<^sub>B (\Dom f\ \\<^sub>B \E.Src (Dom f)\) \\<^sub>B B.can (Dom f \<^bold>\ E.Src (Dom f)) (Dom f)" using assms arr_char E.HcompNml_Nml_Src by fastforce also have "... = B.can (Dom f) (Dom f \<^bold>\ E.Src (Dom f)) \\<^sub>B B.can (Dom f \<^bold>\ E.Src (Dom f)) (Dom f)" proof - have "\B.can (Dom f \<^bold>\ E.Src (Dom f)) (Dom f) : \Dom f\ \\<^sub>B \Dom f\ \\<^sub>B \E.Src (Dom f)\\" using assms ide_char arr_char B.can_in_hom by (metis (no_types, lifting) B.canE_unitor(3) B.runit'_in_vhom E.eval_simps(2) E.Ide_implies_Arr ideE) thus ?thesis using B.comp_cod_arr by auto qed also have "... = \Dom f\" using assms 1 ide_char arr_char by (metis (no_types, lifting) H.ide_hcomp calculation ideE ide_src obj_def' obj_src) also have "... = Map f" using assms by auto finally show ?thesis by simp qed qed lemma hcomp_trg_ide: assumes "ide f" shows "trg f \ f = f" proof (intro arr_eqI) show "H.hseq (trg f) f" using assms by auto show "arr f" using assms by auto show 1: "Dom (trg f \ f) = Dom f" unfolding hcomp_def using assms apply simp using assms ide_char arr_char E.HcompNml_Trg_Nml by (metis (no_types, lifting) ideD(1)) show "Cod (trg f \ f) = Cod f" unfolding hcomp_def using assms apply simp using assms ide_char arr_char E.HcompNml_Trg_Nml by (metis (no_types, lifting) ideD(1)) show "Map (trg f \ f) = Map f" proof - have "Map (trg f \ f) = B.can (Cod (trg f) \<^bold>\\<^bold>\\<^bold>\ Cod f) (Cod (trg f) \<^bold>\ Cod f) \\<^sub>B (Map (trg f) \\<^sub>B Map f) \\<^sub>B B.can (Dom (trg f) \<^bold>\ Dom f) (Dom (trg f) \<^bold>\\<^bold>\\<^bold>\ Dom f)" unfolding hcomp_def using assms by simp also have "... = B.can (Dom f) (E.Trg (Dom f) \<^bold>\ Dom f) \\<^sub>B (\E.Trg (Dom f)\ \\<^sub>B \Dom f\) \\<^sub>B B.can (E.Trg (Dom f) \<^bold>\ Dom f) (Dom f)" using assms arr_char E.HcompNml_Trg_Nml by fastforce also have "... = B.can (Dom f) (E.Trg (Dom f) \<^bold>\ Dom f) \\<^sub>B B.can (E.Trg (Dom f) \<^bold>\ Dom f) (Dom f)" proof - have "\B.can (E.Trg (Dom f) \<^bold>\ Dom f) (Dom f) : \Dom f\ \\<^sub>B \E.Trg (Dom f)\ \\<^sub>B \Dom f\\" using assms ide_char arr_char B.can_in_hom by (metis (no_types, lifting) B.canE_unitor(4) B.lunit'_in_vhom E.Nml_implies_Arr E.eval_simps'(3) ideE) thus ?thesis using B.comp_cod_arr by auto qed also have "... = \Dom f\" using assms 1 ide_char arr_char by (metis (no_types, lifting) H.ide_hcomp Map_ide(1) calculation ideD(1) src_trg trg.preserves_ide) also have "... = Map f" using assms by auto finally show ?thesis by simp qed qed interpretation L: endofunctor vcomp H.L using H.endofunctor_L by auto interpretation R: endofunctor vcomp H.R using H.endofunctor_R by auto interpretation L: full_functor vcomp vcomp H.L proof fix a a' g assume a: "ide a" and a': "ide a'" assume g: "in_hom g (H.L a') (H.L a)" have a_eq: "a = MkIde (Dom a)" using a dom_char [of a] by simp have a'_eq: "a' = MkIde (Dom a')" using a' dom_char [of a'] by simp have 1: "Cod g = Dom a" proof - have "Dom (H.L a) = Dom a" proof - have "Dom (H.L a) = E.Trg (Dom a) \<^bold>\\<^bold>\\<^bold>\ Dom a" using a trg_def hcomp_def apply simp by (metis (no_types, lifting) ideE src_trg trg.preserves_reflects_arr) also have "... = Dom a" using a arr_char E.Trg_HcompNml by (metis (no_types, lifting) E.HcompNml_Trg_Nml ideD(1)) finally show ?thesis by simp qed thus ?thesis using g cod_char [of g] by (metis (no_types, lifting) Dom_cod in_homE) qed have 2: "Dom g = Dom a'" proof - have "Dom (H.L a') = Dom a'" proof - have "Dom (H.L a') = E.Trg (Dom a') \<^bold>\\<^bold>\\<^bold>\ Dom a'" using a' trg_def hcomp_def apply simp by (metis (no_types, lifting) ideE src_trg trg.preserves_reflects_arr) also have "... = Dom a'" using a' arr_char E.Trg_HcompNml by (metis (no_types, lifting) E.HcompNml_Trg_Nml ideD(1)) finally show ?thesis by simp qed thus ?thesis using g dom_char [of g] by (metis (no_types, lifting) Dom_dom in_homE) qed let ?f = "MkArr (Dom a') (Cod a) (Map g)" have f: "in_hom ?f a' a" proof (intro in_homI) show 3: "arr (MkArr (Dom a') (Cod a) (Map g))" proof (intro arr_MkArr [of "Dom a'" "Cod a" "Map g"]) show "Dom a' \ IDE" using a' ide_char arr_char by blast show "Cod a \ IDE" using a ide_char arr_char by blast show "Map g \ HOM (Dom a') (Cod a)" proof show "E.Src (Dom a') = E.Src (Cod a) \ E.Trg (Dom a') = E.Trg (Cod a) \ \Map g : \Dom a'\ \\<^sub>B \Cod a\\" using a a' a_eq g 1 2 ide_char arr_char src_def trg_def trg_hcomp by (metis (no_types, lifting) Cod.simps(1) in_homE) qed qed show "dom (MkArr (Dom a') (Cod a) (Map g)) = a'" using a a' 3 dom_char by auto show "cod (MkArr (Dom a') (Cod a) (Map g)) = a" using a a' 3 cod_char by auto qed moreover have "H.L ?f = g" proof - have "H.L ?f = trg (MkArr (Dom a') (Cod a) (Map g)) \ MkArr (Dom a') (Cod a) (Map g)" using f by auto also have "... = MkIde (E.Trg (Cod a)) \ MkArr (Dom a') (Cod a) (Map g)" using a a' f trg_def [of a] vconn_implies_hpar by auto also have "... = MkArr (E.Trg (Cod a) \<^bold>\\<^bold>\\<^bold>\ Dom a') (E.Trg (Cod a) \<^bold>\\<^bold>\\<^bold>\ Cod a) (B.can (E.Trg (Cod a) \<^bold>\\<^bold>\\<^bold>\ Cod a) (E.Trg (Cod a) \<^bold>\ Cod a) \\<^sub>B (\E.Trg (Cod a)\ \\<^sub>B Map g) \\<^sub>B B.can (E.Trg (Cod a) \<^bold>\ Dom a') (E.Trg (Cod a) \<^bold>\\<^bold>\\<^bold>\ Dom a'))" using hcomp_def apply simp by (metis (no_types, lifting) Cod.simps(1) arrE f in_homE src_trg trg.preserves_arr trg_def) also have "... = MkArr (Dom a') (Cod a) (B.can (Cod a) (E.Trg (Cod a) \<^bold>\ Cod a) \\<^sub>B (trg\<^sub>B \Cod a\ \\<^sub>B Map g) \\<^sub>B B.can (E.Trg (Cod a) \<^bold>\ Dom a') (Dom a'))" proof - have "E.Trg (Cod a) \<^bold>\\<^bold>\\<^bold>\ Dom a' = Dom a'" using a a' arr_char E.HcompNml_Trg_Nml by (metis (no_types, lifting) f ideE trg_simps(1) vconn_implies_hpar(4)) moreover have "E.Trg (Cod a) \<^bold>\\<^bold>\\<^bold>\ Cod a = Cod a" using a a' arr_char E.HcompNml_Trg_Nml by blast moreover have "\E.Trg (Cod a)\ = trg\<^sub>B \Cod a\" using a a' arr_char E.eval_simps'(3) by fastforce ultimately show ?thesis by simp qed also have "... = MkArr (Dom a') (Cod a) (B.lunit \Cod a\ \\<^sub>B (trg\<^sub>B \Cod a\ \\<^sub>B Map g) \\<^sub>B B.lunit' \Dom a'\)" proof - have "E.Trg (Cod a) = E.Trg (Dom a')" using a a' a_eq g ide_char arr_char src_def trg_def trg_hcomp \Cod g = Dom a\ \Dom g = Dom a'\ by (metis (no_types, lifting) Cod.simps(1) in_homE) moreover have "B.can (Cod a) (E.Trg (Cod a) \<^bold>\ Cod a) = B.lunit \Cod a\" using a ide_char arr_char B.canE_unitor(2) by blast moreover have "B.can (E.Trg (Dom a') \<^bold>\ Dom a') (Dom a') = B.lunit' \Dom a'\" using a' ide_char arr_char B.canE_unitor(4) by blast ultimately show ?thesis by simp qed also have "... = MkArr (Dom g) (Cod g) (Map g)" proof - have "src\<^sub>B \Cod a\ = src\<^sub>B (Map g)" using a f g ide_char arr_char src_def B.comp_cod_arr by (metis (no_types, lifting) B.vconn_implies_hpar(1) B.vconn_implies_hpar(3) Cod.simps(1) Map.simps(1) in_homE) moreover have "B.lunit \Cod g\ \\<^sub>B (trg\<^sub>B (Map g) \\<^sub>B Map g) \\<^sub>B B.lunit' \Dom g\ = Map g" proof - have "B.lunit \Cod g\ \\<^sub>B (trg\<^sub>B (Map g) \\<^sub>B Map g) \\<^sub>B B.lunit' \Dom g\ = B.lunit \Cod g\ \\<^sub>B B.lunit' \Cod g\ \\<^sub>B Map g" using g ide_char arr_char B.lunit'_naturality by (metis (no_types, lifting) partial_magma_axioms B.in_homE partial_magma.arrI) also have "... = (B.lunit \Cod g\ \\<^sub>B B.lunit' \Cod g\) \\<^sub>B Map g" using B.comp_assoc by simp also have "... = Map g" using g arr_char E.ide_eval_Ide B.comp_arr_inv' B.comp_cod_arr by fastforce finally show ?thesis by simp qed ultimately have "B.lunit \Cod a\ \\<^sub>B (trg\<^sub>B \Cod a\ \\<^sub>B Map g) \\<^sub>B B.lunit' \Dom a'\ = Map g" using a a' 1 2 f g hcomp_def dom_char cod_char by (metis (no_types, lifting) B.comp_null(2) B.ext B.lunit_simps(2) B.lunit_simps(3) B.src.preserves_reflects_arr B.trg_vcomp B.vseq_implies_hpar(1) ideE) thus ?thesis using a 1 2 by auto qed also have "... = g" using g MkArr_Map by blast finally show ?thesis by simp qed ultimately show "\f. in_hom f a' a \ H.L f = g" by blast qed interpretation R: full_functor vcomp vcomp H.R proof fix a a' g assume a: "ide a" and a': "ide a'" assume g: "in_hom g (H.R a') (H.R a)" have a_eq: "a = MkIde (Dom a)" using a dom_char [of a] by simp have a'_eq: "a' = MkIde (Dom a')" using a' dom_char [of a'] by simp have 1: "Cod g = Dom a" proof - have "Dom (H.R a) = Dom a" proof - have "Dom (H.R a) = Dom a \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom a)" using a src_def hcomp_def apply simp by (metis (no_types, lifting) ideE trg_src src.preserves_reflects_arr) also have "... = Dom a" using a arr_char E.Src_HcompNml by (metis (no_types, lifting) E.HcompNml_Nml_Src ideD(1)) finally show ?thesis by simp qed thus ?thesis using g cod_char [of g] by (metis (no_types, lifting) Dom_cod in_homE) qed have 2: "Dom g = Dom a'" proof - have "Dom (H.R a') = Dom a'" proof - have "Dom (H.R a') = Dom a' \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom a')" using a' src_def hcomp_def apply simp by (metis (no_types, lifting) ideE trg_src src.preserves_reflects_arr) also have "... = Dom a'" using a' arr_char E.Src_HcompNml by (metis (no_types, lifting) E.HcompNml_Nml_Src ideD(1)) finally show ?thesis by simp qed thus ?thesis using g dom_char [of g] by (metis (no_types, lifting) Dom_dom in_homE) qed let ?f = "MkArr (Dom a') (Cod a) (Map g)" have f: "in_hom ?f a' a" proof (intro in_homI) show 3: "arr (MkArr (Dom a') (Cod a) (Map g))" proof (intro arr_MkArr [of "Dom a'" "Cod a" "Map g"]) show "Dom a' \ IDE" using a' ide_char arr_char by blast show "Cod a \ IDE" using a ide_char arr_char by blast show "Map g \ HOM (Dom a') (Cod a)" proof show "E.Src (Dom a') = E.Src (Cod a) \ E.Trg (Dom a') = E.Trg (Cod a) \ \Map g : \Dom a'\ \\<^sub>B \Cod a\\" using a a' a_eq g 1 2 ide_char arr_char src_def trg_def trg_hcomp by (metis (no_types, lifting) Cod.simps(1) in_homE) qed qed show "dom (MkArr (Dom a') (Cod a) (Map g)) = a'" using a a' 3 dom_char by auto show "cod (MkArr (Dom a') (Cod a) (Map g)) = a" using a a' 3 cod_char by auto qed moreover have "H.R ?f = g" proof - have "H.R ?f = MkArr (Dom a') (Cod a) (Map g) \ src (MkArr (Dom a') (Cod a) (Map g))" using f by auto also have "... = MkArr (Dom a') (Cod a) (Map g) \ MkIde (E.Src (Cod a))" using a a' f src_def [of a] vconn_implies_hpar by auto also have "... = MkArr (Dom a' \<^bold>\\<^bold>\\<^bold>\ E.Src (Cod a)) (Cod a \<^bold>\\<^bold>\\<^bold>\ E.Src (Cod a)) (B.can (Cod a \<^bold>\\<^bold>\\<^bold>\ E.Src (Cod a)) (Cod a \<^bold>\ E.Src (Cod a)) \\<^sub>B (Map g \\<^sub>B \E.Src (Cod a)\) \\<^sub>B B.can (Dom a' \<^bold>\ E.Src (Cod a)) (Dom a' \<^bold>\\<^bold>\\<^bold>\ E.Src (Cod a)))" using hcomp_def apply simp by (metis (no_types, lifting) Cod_cod arrE f in_homE trg_src src.preserves_arr src_def) also have "... = MkArr (Dom a') (Cod a) (B.can (Cod a) (Cod a \<^bold>\ E.Src (Cod a)) \\<^sub>B (Map g \\<^sub>B src\<^sub>B \Cod a\) \\<^sub>B B.can (Dom a' \<^bold>\ E.Src (Cod a)) (Dom a'))" proof - have "Dom a' \<^bold>\\<^bold>\\<^bold>\ E.Src (Cod a) = Dom a'" using a a' arr_char E.HcompNml_Nml_Src by (metis (no_types, lifting) f ideE src_simps(1) vconn_implies_hpar(3)) moreover have "Cod a \<^bold>\\<^bold>\\<^bold>\ E.Src (Cod a) = Cod a" using a a' arr_char E.HcompNml_Nml_Src by blast moreover have "\E.Src (Cod a)\ = src\<^sub>B \Cod a\" using a a' arr_char E.eval_simps'(2) by fastforce ultimately show ?thesis by simp qed also have "... = MkArr (Dom a') (Cod a) (B.runit \Cod a\ \\<^sub>B (Map g \\<^sub>B src\<^sub>B \Cod a\) \\<^sub>B B.runit' \Dom a'\)" proof - have "E.Src (Cod a) = E.Src (Dom a')" using a a' g ide_char arr_char src_def trg_def src_hcomp by (metis (no_types, lifting) Cod_dom f ideE in_homE src_cod src_simps(1)) moreover have "B.can (Cod a) (Cod a \<^bold>\ E.Src (Cod a)) = B.runit \Cod a\" using a ide_char arr_char B.canE_unitor(1) by blast moreover have "B.can (Dom a' \<^bold>\ E.Src (Dom a')) (Dom a') = B.runit' \Dom a'\" using a' ide_char arr_char B.canE_unitor(3) by blast ultimately show ?thesis by simp qed also have "... = MkArr (Dom g) (Cod g) (Map g)" proof - have "src\<^sub>B \Cod a\ = src\<^sub>B (Map g)" using a f g ide_char arr_char src_def B.comp_cod_arr by (metis (no_types, lifting) B.vconn_implies_hpar(1) B.vconn_implies_hpar(3) Cod.simps(1) Map.simps(1) in_homE) moreover have "B.runit \Cod g\ \\<^sub>B (Map g \\<^sub>B src\<^sub>B (Map g)) \\<^sub>B B.runit' \Dom g\ = Map g" proof - have "B.runit \Cod g\ \\<^sub>B (Map g \\<^sub>B src\<^sub>B (Map g)) \\<^sub>B B.runit' \Dom g\ = B.runit \Cod g\ \\<^sub>B B.runit'\Cod g\ \\<^sub>B Map g" using g ide_char arr_char B.runit'_naturality [of "Map g"] by (metis (no_types, lifting) partial_magma_axioms B.in_homE partial_magma.arrI) also have "... = (B.runit \Cod g\ \\<^sub>B B.runit' \Cod g\) \\<^sub>B Map g" using B.comp_assoc by simp also have "... = Map g" using g arr_char E.ide_eval_Ide B.comp_arr_inv' B.comp_cod_arr by fastforce finally show ?thesis by simp qed ultimately have "B.runit \Cod a\ \\<^sub>B (Map g \\<^sub>B src\<^sub>B \Cod a\) \\<^sub>B B.runit' \Dom a'\ = Map g" using a a' 1 2 f g hcomp_def dom_char cod_char by (metis (no_types, lifting) ideE) thus ?thesis using a 1 2 by auto qed also have "... = g" using g MkArr_Map by blast finally show ?thesis by simp qed ultimately show "\f. in_hom f a' a \ H.R f = g" by blast qed interpretation L: faithful_functor vcomp vcomp H.L proof fix f f' assume par: "par f f'" and eq: "H.L f = H.L f'" show "f = f'" proof (intro arr_eqI) have 1: "Dom f = Dom f' \ Cod f = Cod f'" using par dom_char cod_char by auto show "arr f" using par by simp show "arr f'" using par by simp show 2: "Dom f = Dom f'" and 3: "Cod f = Cod f'" using 1 by auto show "Map f = Map f'" proof - have "B.L (Map f) = trg\<^sub>B (Map f) \\<^sub>B Map f" using par by auto also have "... = trg\<^sub>B (Map f') \\<^sub>B Map f'" proof - have "\E.Trg (Dom f)\ \\<^sub>B Map f = \E.Trg (Dom f')\ \\<^sub>B Map f'" proof - have A: "\B.can (E.Trg (Dom f) \<^bold>\ Dom f) (E.Trg (Dom f) \<^bold>\\<^bold>\\<^bold>\ Dom f) : \E.Trg (Dom f) \<^bold>\\<^bold>\\<^bold>\ Dom f\ \\<^sub>B \E.Trg (Dom f)\ \\<^sub>B \Dom f\\" using par arr_char B.can_in_hom E.Ide_HcompNml E.Ide_Nmlize_Ide E.Nml_Trg E.Nmlize_Nml E.HcompNml_Trg_Nml src_def trg_def by (metis (no_types, lifting) E.eval_simps(3) E.ide_eval_Ide E.Ide_implies_Arr B.canE_unitor(4) B.lunit'_in_vhom) have B: "\B.can (E.Trg (Dom f) \<^bold>\\<^bold>\\<^bold>\ Cod f) (E.Trg (Dom f) \<^bold>\ Cod f) : \E.Trg (Dom f)\ \\<^sub>B \Cod f\ \\<^sub>B \E.Trg (Dom f) \<^bold>\\<^bold>\\<^bold>\ Cod f\\" using par arr_char B.can_in_hom E.Ide_HcompNml E.Ide_Nmlize_Ide E.Nml_Trg E.Nmlize_Nml E.HcompNml_Trg_Nml src_def trg_def by (metis (no_types, lifting) E.Nmlize.simps(3) E.eval.simps(3) E.Ide.simps(3) E.Ide_implies_Arr E.Src_Trg trg.preserves_arr trg_simps(2)) have C: "\\E.Trg (Dom f)\ \\<^sub>B Map f : \E.Trg (Dom f)\ \\<^sub>B \Dom f\ \\<^sub>B \E.Trg (Dom f)\ \\<^sub>B \Cod f\\" using par arr_char by (metis (no_types, lifting) E.eval_simps'(1) E.eval_simps(3) E.ide_eval_Ide E.Ide_implies_Arr E.Obj_Trg E.Obj_implies_Ide B.hcomp_in_vhom B.ide_in_hom(2) B.src_trg) have 3: "(\E.Trg (Dom f)\ \\<^sub>B Map f) \\<^sub>B B.can (E.Trg (Dom f) \<^bold>\ Dom f) (E.Trg (Dom f) \<^bold>\\<^bold>\\<^bold>\ Dom f) = (\E.Trg (Dom f')\ \\<^sub>B Map f') \\<^sub>B B.can (E.Trg (Dom f') \<^bold>\ Dom f') (E.Trg (Dom f') \<^bold>\\<^bold>\\<^bold>\ Dom f')" proof - have 2: "B.can (E.Trg (Dom f) \<^bold>\\<^bold>\\<^bold>\ Cod f) (E.Trg (Dom f) \<^bold>\ Cod f) \\<^sub>B (\E.Trg (Dom f)\ \\<^sub>B Map f) \\<^sub>B B.can (E.Trg (Dom f) \<^bold>\ Dom f) (E.Trg (Dom f) \<^bold>\\<^bold>\\<^bold>\ Dom f) = B.can (E.Trg (Dom f') \<^bold>\\<^bold>\\<^bold>\ Cod f') (E.Trg (Dom f') \<^bold>\ Cod f') \\<^sub>B (\E.Trg (Dom f')\ \\<^sub>B Map f') \\<^sub>B B.can (E.Trg (Dom f') \<^bold>\ Dom f') (E.Trg (Dom f') \<^bold>\\<^bold>\\<^bold>\ Dom f')" using par eq hcomp_def trg_def src_trg trg.preserves_arr Map_hcomp trg_simps(1) trg_simps(2) trg_simps(3) by auto have "B.mono (B.can (E.Trg (Dom f) \<^bold>\\<^bold>\\<^bold>\ Cod f) (E.Trg (Dom f) \<^bold>\ Cod f))" using par arr_char B.inverse_arrows_can B.iso_is_section B.section_is_mono src_def trg_def E.Nmlize_Nml E.HcompNml_Trg_Nml E.Ide_implies_Arr trg.preserves_arr trg_simps(1) by auto moreover have "B.seq (B.can (E.Trg (Dom f) \<^bold>\\<^bold>\\<^bold>\ Cod f) (E.Trg (Dom f) \<^bold>\ Cod f)) ((\E.Trg (Dom f)\ \\<^sub>B Map f) \\<^sub>B B.can (E.Trg (Dom f) \<^bold>\ Dom f) (E.Trg (Dom f) \<^bold>\\<^bold>\\<^bold>\ Dom f))" using A B C by auto moreover have "B.seq (B.can (E.Trg (Dom f) \<^bold>\\<^bold>\\<^bold>\ Cod f) (E.Trg (Dom f) \<^bold>\ Cod f)) ((\E.Trg (Dom f')\ \\<^sub>B Map f') \\<^sub>B B.can (E.Trg (Dom f') \<^bold>\ Dom f') (E.Trg (Dom f') \<^bold>\\<^bold>\\<^bold>\ Dom f'))" using par 1 2 arr_char calculation(2) by auto moreover have "B.can (E.Trg (Dom f) \<^bold>\\<^bold>\\<^bold>\ Cod f) (E.Trg (Dom f) \<^bold>\ Cod f) = B.can (E.Trg (Dom f') \<^bold>\\<^bold>\\<^bold>\ Cod f') (E.Trg (Dom f') \<^bold>\ Cod f')" using par 1 arr_char by simp ultimately show ?thesis using 2 B.monoE cod_char by auto qed show ?thesis proof - have "B.epi (B.can (E.Trg (Dom f) \<^bold>\ Dom f) (E.Trg (Dom f) \<^bold>\\<^bold>\\<^bold>\ Dom f))" using par arr_char B.inverse_arrows_can B.iso_is_retraction B.retraction_is_epi E.Nmlize_Nml E.HcompNml_Trg_Nml src_def trg_def E.Ide_implies_Arr by (metis (no_types, lifting) E.Nmlize.simps(3) E.Ide.simps(3) E.Src_Trg trg.preserves_arr trg_simps(1)) moreover have "B.seq (\E.Trg (Dom f)\ \\<^sub>B Map f) (B.can (E.Trg (Dom f) \<^bold>\ Dom f) (E.Trg (Dom f) \<^bold>\\<^bold>\\<^bold>\ Dom f))" using A C by auto moreover have "B.seq (\E.Trg (Dom f')\ \\<^sub>B Map f') (B.can (E.Trg (Dom f) \<^bold>\ Dom f) (E.Trg (Dom f) \<^bold>\\<^bold>\\<^bold>\ Dom f))" using 1 3 calculation(2) by auto ultimately show ?thesis using par 1 3 arr_char B.epiE by simp qed qed moreover have "trg\<^sub>B (Map f) = \E.Trg (Dom f)\ \ trg\<^sub>B (Map f') = \E.Trg (Dom f')\" using par arr_char trg_def E.Ide_implies_Arr B.comp_arr_dom B.vseq_implies_hpar(2) E.eval_simps(3) by (metis (no_types, lifting) B.vconn_implies_hpar(2)) ultimately show ?thesis by simp qed also have "... = B.L (Map f')" using par B.hseqE B.hseq_char' by auto finally have "B.L (Map f) = B.L (Map f')" by simp thus ?thesis using 2 3 par arr_char B.L.is_faithful by (metis (no_types, lifting) B.in_homE) qed qed qed interpretation R: faithful_functor vcomp vcomp H.R proof fix f f' assume par: "par f f'" and eq: "H.R f = H.R f'" show "f = f'" proof (intro arr_eqI) have 1: "Dom f = Dom f' \ Cod f = Cod f'" using par dom_char cod_char by auto show "arr f" using par by simp show "arr f'" using par by simp show 2: "Dom f = Dom f'" and 3: "Cod f = Cod f'" using 1 by auto show "Map f = Map f'" proof - have "B.R (Map f) = Map f \\<^sub>B src\<^sub>B (Map f)" using par apply simp by (metis B.hseqE B.hseq_char') also have "... = Map f' \\<^sub>B src\<^sub>B (Map f')" proof - have "Map f \\<^sub>B \E.Src (Dom f)\ = Map f' \\<^sub>B \E.Src (Dom f')\" proof - have 2: "E.Ide (Cod f \<^bold>\ E.Src (Dom f))" using par arr_char src.preserves_arr by auto hence 3: "E.Ide (Cod f \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f))" using par arr_char E.Nml_Src E.Ide_HcompNml calculation by auto have 4: "\<^bold>\Cod f \<^bold>\ E.Src (Dom f)\<^bold>\ = \<^bold>\Cod f \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f)\<^bold>\" using par arr_char by (simp add: E.Nml_HcompNml(1)) have A: "\B.can (Dom f \<^bold>\ E.Src (Dom f)) (Dom f \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f)) : \Dom f \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f)\ \\<^sub>B \Dom f\ \\<^sub>B \E.Src (Dom f)\\" using par arr_char B.can_in_hom E.Ide_HcompNml E.Ide_Nmlize_Ide E.Nml_Src E.Nmlize_Nml E.HcompNml_Nml_Src src_def trg_def by (metis (no_types, lifting) E.eval_simps(2) E.ide_eval_Ide E.Ide_implies_Arr B.canE_unitor(3) B.runit'_in_vhom) have B: "\B.can (Cod f \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f)) (Cod f \<^bold>\ E.Src (Dom f)) : \Cod f\ \\<^sub>B \E.Src (Dom f)\ \\<^sub>B \Cod f \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f)\\" using 2 3 4 B.can_in_hom [of "Cod f \<^bold>\ E.Src (Dom f)" "Cod f \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f)"] by simp have C: "\Map f \\<^sub>B \E.Src (Dom f)\ : \Dom f\ \\<^sub>B \E.Src (Dom f)\ \\<^sub>B \Cod f\ \\<^sub>B \E.Src (Dom f)\\" using par arr_char E.Ide_Nmlize_Ide E.Nml_Trg E.Nmlize_Nml E.HcompNml_Trg_Nml src_def trg_def E.ide_eval_Ide E.Ide_implies_Arr E.Obj_implies_Ide apply (intro B.hcomp_in_vhom) apply (simp add: B.ide_in_hom(2)) apply simp by (metis (no_types, lifting) A B.ideD(1) B.not_arr_null B.seq_if_composable B.src.preserves_reflects_arr B.vconn_implies_hpar(3) E.HcompNml_Nml_Src) have 5: "(Map f \\<^sub>B \E.Src (Dom f)\) \\<^sub>B B.can (Dom f \<^bold>\ E.Src (Dom f)) (Dom f \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f)) = (Map f' \\<^sub>B \E.Src (Dom f')\) \\<^sub>B B.can (Dom f' \<^bold>\ E.Src (Dom f')) (Dom f' \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f'))" proof - have 6: "B.can (Cod f \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f)) (Cod f \<^bold>\ E.Src (Dom f)) \\<^sub>B (Map f \\<^sub>B \E.Src (Dom f)\) \\<^sub>B B.can (Dom f \<^bold>\ E.Src (Dom f)) (Dom f \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f)) = B.can (Cod f' \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f')) (Cod f' \<^bold>\ E.Src (Dom f')) \\<^sub>B (Map f' \\<^sub>B \E.Src (Dom f')\) \\<^sub>B B.can (Dom f' \<^bold>\ E.Src (Dom f')) (Dom f' \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f'))" using par eq hcomp_def src_def trg_src src.preserves_arr Map_hcomp src_simps(1) src_simps(2) src_simps(3) by auto have "B.mono (B.can (Cod f \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f)) (Cod f \<^bold>\ E.Src (Dom f)))" using 2 3 4 B.inverse_arrows_can(1) B.iso_is_section B.section_is_mono by simp moreover have "B.seq (B.can (Cod f \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f)) (Cod f \<^bold>\ E.Src (Dom f))) ((Map f \\<^sub>B \E.Src (Dom f)\) \\<^sub>B B.can (Dom f \<^bold>\ E.Src (Dom f)) (Dom f \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f)))" using A B C by auto moreover have "B.seq (B.can (Cod f \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f)) (Cod f \<^bold>\ E.Src (Dom f))) ((Map f' \\<^sub>B \E.Src (Dom f')\) \\<^sub>B B.can (Dom f' \<^bold>\ E.Src (Dom f')) (Dom f' \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f')))" using par 1 6 arr_char calculation(2) by auto moreover have "B.can (Cod f \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f)) (Cod f \<^bold>\ E.Src (Dom f)) = B.can (Cod f' \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f')) (Cod f' \<^bold>\ E.Src (Dom f'))" using par 1 arr_char by simp ultimately show ?thesis using 6 B.monoE cod_char by auto qed show ?thesis proof - have "B.epi (B.can (Dom f \<^bold>\ E.Src (Dom f)) (Dom f \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f)))" using 2 3 4 B.inverse_arrows_can(1) B.iso_is_retraction B.retraction_is_epi by (metis (no_types, lifting) E.Nml_Src E.Nmlize.simps(3) E.Nmlize_Nml E.HcompNml_Nml_Src E.Ide.simps(3) par arrE) moreover have "B.seq (Map f \\<^sub>B \E.Src (Dom f)\) (B.can (Dom f \<^bold>\ E.Src (Dom f)) (Dom f \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f)))" using A C by auto moreover have "B.seq (Map f' \\<^sub>B \E.Src (Dom f')\) (B.can (Dom f \<^bold>\ E.Src (Dom f)) (Dom f \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f)))" using 1 5 calculation(2) by auto ultimately show ?thesis using par 1 5 arr_char B.epiE by simp qed qed moreover have "src\<^sub>B (Map f) = \E.Src (Dom f)\ \ src\<^sub>B (Map f') = \E.Src (Dom f')\" using par arr_char src_def by (metis (no_types, lifting) B.vconn_implies_hpar(1) E.Nml_implies_Arr E.eval_simps(2)) ultimately show ?thesis by simp qed also have "... = B.R (Map f')" using par B.hseqE B.hseq_char' by auto finally have "B.R (Map f) = B.R (Map f')" by simp thus ?thesis using 2 3 par arr_char B.R.is_faithful by (metis (no_types, lifting) B.in_homE) qed qed qed interpretation VxVxV: product_category vcomp VxV.comp .. interpretation VVV: subcategory VxVxV.comp \\\\\. arr (fst \\\) \ VV.arr (snd \\\) \ src (fst \\\) = trg (fst (snd \\\))\ using subcategory_VVV by auto interpretation HoHV: "functor" VVV.comp vcomp H.HoHV using H.functor_HoHV by auto interpretation HoVH: "functor" VVV.comp vcomp H.HoVH using H.functor_HoVH by auto definition \ where "\ \ \ \ \ if VVV.arr (\, \, \) then hcomp \ (hcomp \ \) else null" interpretation natural_isomorphism VVV.comp vcomp H.HoHV H.HoVH \\\\\. \ (fst \\\) (fst (snd \\\)) (snd (snd \\\))\ proof show "\\\\. \ VVV.arr \\\ \ \ (fst \\\) (fst (snd \\\)) (snd (snd \\\)) = null" using \_def by simp show "\\\\. VVV.arr \\\ \ dom (\ (fst \\\) (fst (snd \\\)) (snd (snd \\\))) = H.HoHV (VVV.dom \\\)" using VVV.arr_char VV.arr_char \_def hcomp_assoc H.HoHV_def by force show 1: "\\\\. VVV.arr \\\ \ cod (\ (fst \\\) (fst (snd \\\)) (snd (snd \\\))) = H.HoVH (VVV.cod \\\)" using VVV.arr_char VV.arr_char \_def H.HoVH_def by force show "\\\\. VVV.arr \\\ \ H.HoVH \\\ \ \ (fst (VVV.dom \\\)) (fst (snd (VVV.dom \\\))) (snd (snd (VVV.dom \\\))) = \ (fst \\\) (fst (snd \\\)) (snd (snd \\\))" using \_def HoVH.is_natural_1 H.HoVH_def by auto show "\\\\. VVV.arr \\\ \ \ (fst (VVV.cod \\\)) (fst (snd (VVV.cod \\\))) (snd (snd (VVV.cod \\\))) \ H.HoHV \\\ = \ (fst \\\) (fst (snd \\\)) (snd (snd \\\))" proof - fix \\\ assume \\\: "VVV.arr \\\" have "H.HoHV \\\ = \ (fst \\\) (fst (snd \\\)) (snd (snd \\\))" unfolding \_def H.HoHV_def using \\\ HoHV.preserves_cod hcomp_assoc VVV.arr_char VV.arr_char by simp thus "\ (fst (VVV.cod \\\)) (fst (snd (VVV.cod \\\))) (snd (snd (VVV.cod \\\))) \ H.HoHV \\\ = \ (fst \\\) (fst (snd \\\)) (snd (snd \\\))" using 1 \\\ comp_cod_arr \_def by (metis (no_types, lifting) H.HoVH_def HoHV.preserves_arr prod.collapse) qed show "\fgh. VVV.ide fgh \ iso (\ (fst fgh) (fst (snd fgh)) (snd (snd fgh)))" using \_def HoVH.preserves_ide H.HoVH_def by auto qed definition \ where "\ \ \a. a" sublocale bicategory vcomp hcomp \ \ src trg using hcomp_obj_self \_def hcomp_assoc VVV.arr_char VV.arr_char apply unfold_locales by (auto simp add: \_def ide_in_hom(2)) lemma is_bicategory: shows "bicategory vcomp hcomp \ \ src trg" .. sublocale strict_bicategory vcomp hcomp \ \ src trg proof show "\fgh. ide fgh \ lunit fgh = fgh" proof - fix fgh assume fgh: "ide fgh" have "fgh = lunit fgh" proof (intro lunit_eqI) show "ide fgh" using fgh by simp show "\fgh : trg fgh \ fgh \ fgh\" using fgh hcomp_def hcomp_trg_ide by auto show "trg fgh \ fgh = (\ (trg fgh) \ fgh) \ \' (trg fgh) (trg fgh) fgh" proof - have "(\ (trg fgh) \ fgh) \ \' (trg fgh) (trg fgh) fgh = (trg fgh \ fgh) \ \' (trg fgh) (trg fgh) fgh" using fgh \_def by metis also have "... = (trg fgh \ fgh) \ (trg fgh \ trg fgh \ fgh)" using fgh \_def by fastforce also have "... = trg fgh \ fgh" using fgh hcomp_obj_self hcomp_assoc by (simp add: hcomp_trg_ide) finally show ?thesis by simp qed qed thus "lunit fgh = fgh" by simp qed show "\fgh. ide fgh \ runit fgh = fgh" proof - fix fgh assume fgh: "ide fgh" have "fgh = runit fgh" proof (intro runit_eqI) show "ide fgh" using fgh by simp show "\fgh : fgh \ src fgh \ fgh\" using fgh hcomp_def hcomp_ide_src by auto show "fgh \ src fgh = (fgh \ \ (src fgh)) \ \ fgh (src fgh) (src fgh)" proof - have "(fgh \ \ (src fgh)) \ \ fgh (src fgh) (src fgh) = (fgh \ src fgh) \ \ fgh (src fgh) (src fgh)" using fgh \_def by metis also have "... = (fgh \ src fgh) \ (fgh \ src fgh \ src fgh)" using fgh \_def by fastforce also have "... = fgh \ src fgh" using fgh comp_arr_dom hcomp_obj_self by simp finally show ?thesis by simp qed qed thus "runit fgh = fgh" by simp qed show "\f g h. \ ide f; ide g; ide h; src f = trg g; src g = trg h \ \ ide (\ f g h)" using \_def VV.arr_char VVV.arr_char by auto qed theorem is_strict_bicategory: shows "strict_bicategory vcomp hcomp \ \ src trg" .. subsection "The Strictness Theorem" text \ The Strictness Theorem asserts: ``Every bicategory is biequivalent to a strict bicategory.'' This amounts to an equivalent (and perhaps more desirable) formulation of the Coherence Theorem. In this section we prove the Strictness Theorem by constructing an equivalence pseudofunctor from a bicategory to its strictification. \ lemma iso_char: shows "iso \ \ arr \ \ B.iso (Map \)" and "iso \ \ inv \ = MkArr (Cod \) (Dom \) (B.inv (Map \))" proof - have 1: "iso \ \ arr \ \ B.iso (Map \)" proof - assume \: "iso \" obtain \ where \: "inverse_arrows \ \" using \ by auto have "B.inverse_arrows (Map \) (Map \)" proof show "B.ide (Map \ \\<^sub>B Map \)" proof - have "Map \ \\<^sub>B Map \ = Map (\ \ \)" using \ \ inverse_arrows_def Map_comp arr_char seq_char by (metis (no_types, lifting) ide_compE) moreover have "B.ide ..." using \ ide_char by blast ultimately show ?thesis by simp qed show "B.ide (Map \ \\<^sub>B Map \)" proof - have "Map \ \\<^sub>B Map \ = Map (\ \ \)" using \ \ inverse_arrows_def comp_char [of \ \] by simp moreover have "B.ide ..." using \ ide_char by blast ultimately show ?thesis by simp qed qed thus "arr \ \ B.iso (Map \)" using \ by auto qed let ?\ = "MkArr (Cod \) (Dom \) (B.inv (Map \))" have 2: "arr \ \ B.iso (Map \) \ iso \ \ inv \ = ?\" proof assume \: "arr \ \ B.iso (Map \)" have \: "\?\ : cod \ \ dom \\" using \ arr_char dom_char cod_char by auto have 4: "inverse_arrows \ ?\" proof show "ide (?\ \ \)" proof - have "?\ \ \ = dom \" using \ \ MkArr_Map comp_char seq_char B.comp_inv_arr' dom_char by auto thus ?thesis using \ by simp qed show "ide (\ \ ?\)" proof - have "\ \ ?\ = cod \" using \ \ MkArr_Map comp_char seq_char B.comp_arr_inv' cod_char by auto thus ?thesis using \ by simp qed qed thus "iso \" by auto show "inv \ = ?\" using 4 inverse_unique by simp qed have 3: "arr \ \ B.iso (Map \) \ iso \" using 2 by simp show "iso \ \ arr \ \ B.iso (Map \)" using 1 3 by blast show "iso \ \ inv \ = ?\" using 1 2 by blast qed text \ We next define a map \UP\ from the given bicategory \B\ to its strictification, and show that it is an equivalence pseudofunctor. The following auxiliary definition is not logically necessary, but it provides some terms that can be the targets of simplification rules and thereby enables some proofs to be done by simplification that otherwise could not be. Trying to eliminate it breaks some short proofs below, so I have kept it. \ definition UP\<^sub>0 where "UP\<^sub>0 a \ if B.obj a then MkIde \<^bold>\a\<^bold>\\<^sub>0 else null" lemma obj_UP\<^sub>0 [simp]: assumes "B.obj a" shows "obj (UP\<^sub>0 a)" using assms UP\<^sub>0_def ide_MkIde [of "\<^bold>\a\<^bold>\\<^sub>0"] src_def obj_def by simp lemma UP\<^sub>0_in_hom [intro]: assumes "B.obj a" shows "\UP\<^sub>0 a : UP\<^sub>0 a \ UP\<^sub>0 a\" and "\UP\<^sub>0 a : UP\<^sub>0 a \ UP\<^sub>0 a\" using assms obj_UP\<^sub>0 by blast+ lemma UP\<^sub>0_simps [simp]: assumes "B.obj a" shows "ide (UP\<^sub>0 a)" "arr (UP\<^sub>0 a)" and "src (UP\<^sub>0 a) = UP\<^sub>0 a" and "trg (UP\<^sub>0 a) = UP\<^sub>0 a" and "dom (UP\<^sub>0 a) = UP\<^sub>0 a" and "cod (UP\<^sub>0 a) = UP\<^sub>0 a" using assms obj_UP\<^sub>0 apply blast using assms obj_UP\<^sub>0 apply blast - using assms obj_UP\<^sub>0 + using assms obj_UP\<^sub>0 obj_simps apply simp_all using ideD(2) obj_UP\<^sub>0 apply blast using ideD(3) obj_UP\<^sub>0 by blast definition UP where "UP \ \ if B.arr \ then MkArr \<^bold>\B.dom \\<^bold>\ \<^bold>\B.cod \\<^bold>\ \ else null" lemma Dom_UP [simp]: assumes "B.arr \" shows "Dom (UP \) = \<^bold>\B.dom \\<^bold>\" using assms UP_def by fastforce lemma Cod_UP [simp]: assumes "B.arr \" shows "Cod (UP \) = \<^bold>\B.cod \\<^bold>\" using assms UP_def by fastforce lemma Map_UP [simp]: assumes "B.arr \" shows "Map (UP \) = \" using assms UP_def by fastforce lemma arr_UP: assumes "B.arr \" shows "arr (UP \)" using assms UP_def by (intro arrI, fastforce+) lemma UP_in_hom [intro]: assumes "B.arr \" shows "\UP \ : UP\<^sub>0 (src\<^sub>B \) \ UP\<^sub>0 (trg\<^sub>B \)\" and "\UP \ : UP (B.dom \) \ UP (B.cod \)\" using assms arr_UP UP_def UP\<^sub>0_def dom_char cod_char src_def trg_def by auto lemma UP_simps [simp]: assumes "B.arr \" shows "arr (UP \)" and "src (UP \) = UP\<^sub>0 (src\<^sub>B \)" and "trg (UP \) = UP\<^sub>0 (trg\<^sub>B \)" and "dom (UP \) = UP (B.dom \)" and "cod (UP \) = UP (B.cod \)" using assms arr_UP UP_in_hom by auto interpretation UP: "functor" V\<^sub>B vcomp UP using UP_def arr_UP UP_simps(4-5) apply unfold_locales apply auto[4] using arr_UP UP_def comp_char seq_char by auto interpretation UP: weak_arrow_of_homs V\<^sub>B src\<^sub>B trg\<^sub>B vcomp src trg UP proof fix \ assume \: "B.arr \" show "isomorphic (UP (src\<^sub>B \)) (src (UP \))" proof - let ?\ = "MkArr \<^bold>\src\<^sub>B \\<^bold>\ \<^bold>\src\<^sub>B \\<^bold>\\<^sub>0 (src\<^sub>B \)" have \: "\?\ : UP (src\<^sub>B \) \ src (UP \)\" proof show 1: "arr ?\" using \ by (intro arrI, auto) show "dom ?\ = UP (src\<^sub>B \)" using \ 1 dom_char UP_def by simp show "cod ?\ = src (UP \)" using \ 1 cod_char src_def by auto qed have "iso ?\" using \ \ iso_char src_def by auto thus ?thesis using \ isomorphic_def by auto qed show "isomorphic (UP (trg\<^sub>B \)) (trg (UP \))" proof - let ?\ = "MkArr \<^bold>\trg\<^sub>B \\<^bold>\ \<^bold>\trg\<^sub>B \\<^bold>\\<^sub>0 (trg\<^sub>B \)" have \: "\?\ : UP (trg\<^sub>B \) \ trg (UP \)\" proof show 1: "arr ?\" using \ by (intro arrI, auto) show "dom ?\ = UP (trg\<^sub>B \)" using \ 1 dom_char UP_def by simp show "cod ?\ = trg (UP \)" using \ 1 cod_char trg_def by auto qed have "iso ?\" using \ \ iso_char trg_def by auto thus ?thesis using \ isomorphic_def by auto qed qed interpretation "functor" B.VV.comp VV.comp UP.FF using UP.functor_FF by auto interpretation HoUP_UP: composite_functor B.VV.comp VV.comp vcomp UP.FF \\\\. hcomp (fst \\) (snd \\)\ .. interpretation UPoH: composite_functor B.VV.comp V\<^sub>B vcomp \\\\. fst \\ \\<^sub>B snd \\\ UP .. abbreviation \\<^sub>o where "\\<^sub>o fg \ MkArr (\<^bold>\fst fg\<^bold>\ \<^bold>\ \<^bold>\snd fg\<^bold>\) \<^bold>\fst fg \\<^sub>B snd fg\<^bold>\ (fst fg \\<^sub>B snd fg)" interpretation \: transformation_by_components B.VV.comp vcomp HoUP_UP.map UPoH.map \\<^sub>o proof fix fg assume fg: "B.VV.ide fg" show "\\\<^sub>o fg : HoUP_UP.map fg \ UPoH.map fg\" using fg arr_char dom_char cod_char B.VV.ide_char B.VV.arr_char UP.FF_def UP_def hcomp_def B.can_Ide_self src_def trg_def apply (intro in_homI) by auto next fix \\ assume \\: "B.VV.arr \\" show "\\<^sub>o (B.VV.cod \\) \ HoUP_UP.map \\ = UPoH.map \\ \ \\<^sub>o (B.VV.dom \\)" proof - have "\\<^sub>o (B.VV.cod \\) \ HoUP_UP.map \\ = MkArr (\<^bold>\B.dom (fst \\)\<^bold>\ \<^bold>\ \<^bold>\B.dom (snd \\)\<^bold>\) (\<^bold>\B.cod (fst \\) \\<^sub>B B.cod (snd \\)\<^bold>\) (fst \\ \\<^sub>B snd \\)" proof - have "\\<^sub>o (B.VV.cod \\) \ HoUP_UP.map \\ = MkArr (\<^bold>\B.cod (fst \\)\<^bold>\ \<^bold>\ \<^bold>\B.cod (snd \\)\<^bold>\) (\<^bold>\B.cod (fst \\) \\<^sub>B B.cod (snd \\)\<^bold>\) (B.cod (fst \\) \\<^sub>B B.cod (snd \\)) \ MkArr (\<^bold>\B.dom (fst \\)\<^bold>\ \<^bold>\ \<^bold>\B.dom (snd \\)\<^bold>\) (\<^bold>\B.cod (fst \\)\<^bold>\ \<^bold>\ \<^bold>\B.cod (snd \\)\<^bold>\) (fst \\ \\<^sub>B snd \\)" using \\ B.VV.arr_char arr_char UP.FF_def hcomp_def UP_def src_def trg_def B.can_in_hom B.can_Ide_self B.comp_arr_dom B.comp_cod_arr by auto also have "... = MkArr (\<^bold>\B.dom (fst \\)\<^bold>\ \<^bold>\ \<^bold>\B.dom (snd \\)\<^bold>\) (\<^bold>\B.cod (fst \\) \\<^sub>B B.cod (snd \\)\<^bold>\) ((B.cod (fst \\) \\<^sub>B B.cod (snd \\)) \\<^sub>B (fst \\ \\<^sub>B snd \\))" using \\ B.VV.arr_char by (intro comp_MkArr arr_MkArr, auto) also have "... = MkArr (\<^bold>\B.dom (fst \\)\<^bold>\ \<^bold>\ \<^bold>\B.dom (snd \\)\<^bold>\) (\<^bold>\B.cod (fst \\) \\<^sub>B B.cod (snd \\)\<^bold>\) (fst \\ \\<^sub>B snd \\)" using \\ B.VV.arr_char B.comp_cod_arr by auto finally show ?thesis by simp qed also have "... = UPoH.map \\ \ \\<^sub>o (B.VV.dom \\)" proof - have "UPoH.map \\ \ \\<^sub>o (B.VV.dom \\) = MkArr (\<^bold>\B.dom (fst \\) \\<^sub>B B.dom (snd \\)\<^bold>\) (\<^bold>\B.cod (fst \\) \\<^sub>B B.cod (snd \\)\<^bold>\) (fst \\ \\<^sub>B snd \\) \ MkArr (\<^bold>\B.dom (fst \\)\<^bold>\ \<^bold>\ \<^bold>\B.dom (snd \\)\<^bold>\) (\<^bold>\B.dom (fst \\) \\<^sub>B B.dom (snd \\)\<^bold>\) (B.dom (fst \\) \\<^sub>B B.dom (snd \\))" using \\ B.VV.arr_char arr_char UP.FF_def hcomp_def UP_def src_def trg_def B.can_in_hom B.can_Ide_self B.comp_arr_dom B.comp_cod_arr by auto also have "... = MkArr (\<^bold>\B.dom (fst \\)\<^bold>\ \<^bold>\ \<^bold>\B.dom (snd \\)\<^bold>\) (\<^bold>\B.cod (fst \\) \\<^sub>B B.cod (snd \\)\<^bold>\) ((fst \\ \\<^sub>B snd \\) \\<^sub>B (B.dom (fst \\) \\<^sub>B B.dom (snd \\)))" using \\ B.VV.arr_char arr_MkArr apply (intro comp_MkArr arr_MkArr) by auto also have "... = MkArr (\<^bold>\B.dom (fst \\)\<^bold>\ \<^bold>\ \<^bold>\B.dom (snd \\)\<^bold>\) (\<^bold>\B.cod (fst \\) \\<^sub>B B.cod (snd \\)\<^bold>\) (fst \\ \\<^sub>B snd \\)" using \\ B.VV.arr_char B.comp_arr_dom by auto finally show ?thesis by simp qed finally show ?thesis by simp qed qed abbreviation \ where "\ \ \.map" lemma \_in_hom [intro]: assumes "B.arr (fst \\)" and "B.arr (snd \\)" and "src\<^sub>B (fst \\) = trg\<^sub>B (snd \\)" shows "\\ \\ : UP\<^sub>0 (src\<^sub>B (snd \\)) \ UP\<^sub>0 (trg\<^sub>B (fst \\))\" and "\\ \\ : UP (B.dom (fst \\)) \ UP (B.dom (snd \\)) \ UP (B.cod (fst \\) \\<^sub>B B.cod (snd \\))\" proof - let ?\ = "fst \\" and ?\ = "snd \\" show 1: "\\ \\ : UP (B.dom ?\) \ UP (B.dom ?\) \ UP (B.cod ?\ \\<^sub>B B.cod ?\)\" proof show "arr (\ \\)" using assms by auto show "dom (\ \\) = UP (B.dom ?\) \ UP (B.dom ?\)" proof - have "B.VV.in_hom (?\, ?\) (B.dom ?\, B.dom ?\) (B.cod ?\, B.cod ?\)" using assms B.VV.in_hom_char B.VV.arr_char by auto hence "dom (\ \\) = HoUP_UP.map (B.dom ?\, B.dom ?\)" by auto also have "... = UP (B.dom ?\) \ UP (B.dom ?\)" using assms UP.FF_def by fastforce finally show ?thesis by simp qed show "cod (\ \\) = UP (B.cod ?\ \\<^sub>B B.cod ?\)" using assms B.VV.in_hom_char B.VV.arr_char by auto qed show "\\ \\ : UP\<^sub>0 (src\<^sub>B ?\) \ UP\<^sub>0 (trg\<^sub>B ?\)\" using assms 1 src_dom [of "\ \\"] trg_dom [of "\ \\"] by auto qed lemma \_simps [simp]: assumes "B.arr (fst \\)" and "B.arr (snd \\)" and "src\<^sub>B (fst \\) = trg\<^sub>B (snd \\)" shows "arr (\ \\)" and "src (\ \\) = UP\<^sub>0 (src\<^sub>B (snd \\))" and "trg (\ \\) = UP\<^sub>0 (trg\<^sub>B (fst \\))" and "dom (\ \\) = UP (B.dom (fst \\)) \ UP (B.dom (snd \\))" and "cod (\ \\) = UP (B.cod (fst \\) \\<^sub>B B.cod (snd \\))" using assms \_in_hom [of \\] by auto lemma \_ide_simps [simp]: assumes "B.ide (fst fg)" and "B.ide (snd fg)" and "src\<^sub>B (fst fg) = trg\<^sub>B (snd fg)" shows "Dom (\ fg) = \<^bold>\fst fg\<^bold>\ \<^bold>\ \<^bold>\snd fg\<^bold>\" and "Cod (\ fg) = \<^bold>\fst fg \\<^sub>B snd fg\<^bold>\" and "Map (\ fg) = fst fg \\<^sub>B snd fg" using assms B.VV.ide_char B.VV.arr_char by auto interpretation \: natural_isomorphism B.VV.comp vcomp HoUP_UP.map UPoH.map \ proof fix fg assume fg: "B.VV.ide fg" have "arr (\ fg)" using fg \.preserves_reflects_arr [of fg] by simp thus "iso (\ fg)" using fg iso_char by simp qed lemma \_ide_simp: assumes "B.ide f" and "B.ide g" and "src\<^sub>B f = trg\<^sub>B g" shows "\ (f, g) = MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) \<^bold>\f \\<^sub>B g\<^bold>\ (f \\<^sub>B g)" using assms B.VV.ide_char B.VV.arr_char by simp lemma \'_ide_simp: assumes "B.ide f" and "B.ide g" and "src\<^sub>B f = trg\<^sub>B g" shows "inv (\ (f, g)) = MkArr \<^bold>\f \\<^sub>B g\<^bold>\ (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) (f \\<^sub>B g)" using assms \_ide_simp iso_char \.components_are_iso [of "(f, g)"] B.VV.ide_char B.VV.arr_char by simp interpretation UP: pseudofunctor V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B vcomp hcomp \ \ src trg UP \ proof fix f g h assume f: "B.ide f" and g: "B.ide g" and h: "B.ide h" and fg: "src\<^sub>B f = trg\<^sub>B g" and gh: "src\<^sub>B g = trg\<^sub>B h" show "UP \\<^sub>B[f, g, h] \ \ (f \\<^sub>B g, h) \ (\ (f, g) \ UP h) = \ (f, g \\<^sub>B h) \ (UP f \ \ (g, h)) \ \ (UP f) (UP g) (UP h)" proof - have "UP \\<^sub>B[f, g, h] \ \ (f \\<^sub>B g, h) \ (\ (f, g) \ UP h) = MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) \<^bold>\f \\<^sub>B g \\<^sub>B h\<^bold>\ (f \\<^sub>B g \\<^sub>B h)" proof - have 1: "UP.hseq\<^sub>D (MkIde \<^bold>\f\<^bold>\) (MkIde \<^bold>\g\<^bold>\)" using f g fg hseq_char src_def trg_def arr_char by auto have 2: "UP.hseq\<^sub>D (MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) \<^bold>\f \\<^sub>B g\<^bold>\ (f \\<^sub>B g) \ MkIde (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\)) (MkIde \<^bold>\h\<^bold>\)" proof - have "MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) \<^bold>\f \\<^sub>B g\<^bold>\ (f \\<^sub>B g) \ MkIde (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) = MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) \<^bold>\f \\<^sub>B g\<^bold>\ (f \\<^sub>B g)" proof - have "MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) \<^bold>\f \\<^sub>B g\<^bold>\ (f \\<^sub>B g) \ MkIde (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) = MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) \<^bold>\f \\<^sub>B g\<^bold>\ (f \\<^sub>B g) \ MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) (f \\<^sub>B g)" using f g fg by simp also have "... = MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) \<^bold>\f \\<^sub>B g\<^bold>\ ((f \\<^sub>B g) \\<^sub>B (f \\<^sub>B g))" using f g fg by (intro comp_MkArr arr_MkArr, auto) also have "... = MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) \<^bold>\f \\<^sub>B g\<^bold>\ (f \\<^sub>B g)" using f g fg by simp finally show ?thesis by blast qed thus ?thesis using f g h fg gh arr_char src_def trg_def by auto qed have "UP \\<^sub>B[f, g, h] = MkArr \<^bold>\(f \\<^sub>B g) \\<^sub>B h\<^bold>\ \<^bold>\f \\<^sub>B g \\<^sub>B h\<^bold>\ \\<^sub>B[f, g, h]" using f g h fg gh UP_def B.HoHV_def B.HoVH_def B.VVV.arr_char B.VV.arr_char B.VVV.dom_char B.VVV.cod_char by simp moreover have "\ (f \\<^sub>B g, h) = MkArr (\<^bold>\f \\<^sub>B g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) \<^bold>\(f \\<^sub>B g) \\<^sub>B h\<^bold>\ ((f \\<^sub>B g) \\<^sub>B h) \ MkArr (\<^bold>\f \\<^sub>B g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) (\<^bold>\f \\<^sub>B g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) ((f \\<^sub>B g) \\<^sub>B h)" using f g h fg gh \.map_simp_ide \.map_def UP.FF_def UP_def hcomp_def B.VV.arr_char B.can_Ide_self B.comp_arr_dom B.comp_cod_arr src_def trg_def apply simp by (metis (no_types, lifting) B.ide_hcomp B.ide_char arr_UP) moreover have "\ (f, g) \ UP h = MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) (\<^bold>\f \\<^sub>B g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) (B.inv \\<^sub>B[f, g, h])" proof - have "MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) (B.can (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) \\<^sub>B (f \\<^sub>B g) \\<^sub>B B.can (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\)) = MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) (f \\<^sub>B g)" using f g fg B.can_Ide_self B.comp_arr_dom B.comp_cod_arr by simp moreover have "MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) \<^bold>\f \\<^sub>B g\<^bold>\ (f \\<^sub>B g) \ MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) (f \\<^sub>B g) = MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) \<^bold>\f \\<^sub>B g\<^bold>\ (f \\<^sub>B g)" proof - have "MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) \<^bold>\f \\<^sub>B g\<^bold>\ (f \\<^sub>B g) \ MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) (f \\<^sub>B g) = MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) \<^bold>\f \\<^sub>B g\<^bold>\ ((f \\<^sub>B g) \\<^sub>B (f \\<^sub>B g))" using f g fg arr_MkArr by (intro comp_MkArr arr_MkArr) auto also have "... = MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) \<^bold>\f \\<^sub>B g\<^bold>\ (f \\<^sub>B g)" using f g fg by simp finally show ?thesis by blast qed moreover have "B.can ((\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) \<^bold>\ \<^bold>\h\<^bold>\) (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) = B.inv \\<^sub>B[f, g, h]" using f g h fg gh B.canI_associator_0 B.inverse_arrows_can by simp ultimately show ?thesis using 1 2 f g h fg gh \.map_def UP_def hcomp_def UP.FF_def B.VV.arr_char trg_def B.can_Ide_self B.comp_cod_arr by simp qed ultimately have "UP \\<^sub>B[f, g, h] \ \ (f \\<^sub>B g, h) \ (\ (f, g) \ UP h) = MkArr \<^bold>\(f \\<^sub>B g) \\<^sub>B h\<^bold>\ \<^bold>\f \\<^sub>B g \\<^sub>B h\<^bold>\ \\<^sub>B[f, g, h] \ MkArr (\<^bold>\f \\<^sub>B g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) \<^bold>\(f \\<^sub>B g) \\<^sub>B h\<^bold>\ ((f \\<^sub>B g) \\<^sub>B h) \ MkArr (\<^bold>\f \\<^sub>B g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) (\<^bold>\f \\<^sub>B g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) ((f \\<^sub>B g) \\<^sub>B h) \ MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) (\<^bold>\f \\<^sub>B g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) (B.inv \\<^sub>B[f, g, h])" using comp_assoc by presburger also have "... = MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) \<^bold>\f \\<^sub>B g \\<^sub>B h\<^bold>\ (\\<^sub>B[f, g, h] \\<^sub>B ((f \\<^sub>B g) \\<^sub>B h) \\<^sub>B ((f \\<^sub>B g) \\<^sub>B h) \\<^sub>B B.inv \\<^sub>B[f, g, h])" proof - have "Arr (MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) (\<^bold>\f \\<^sub>B g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) (B.inv \\<^sub>B[f, g, h])) \ Arr (MkArr (\<^bold>\f \\<^sub>B g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) (\<^bold>\f \\<^sub>B g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) ((f \\<^sub>B g) \\<^sub>B h)) \ Arr (MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) (\<^bold>\f \\<^sub>B g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) (((f \\<^sub>B g) \\<^sub>B h) \\<^sub>B B.inv \\<^sub>B[f, g, h])) \ Arr (MkArr (\<^bold>\f \\<^sub>B g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) \<^bold>\(f \\<^sub>B g) \\<^sub>B h\<^bold>\ ((f \\<^sub>B g) \\<^sub>B h)) \ Arr (MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) \<^bold>\(f \\<^sub>B g) \\<^sub>B h\<^bold>\ (((f \\<^sub>B g) \\<^sub>B h) \\<^sub>B ((f \\<^sub>B g) \\<^sub>B h) \\<^sub>B B.inv \\<^sub>B[f, g, h])) \ Arr (MkArr \<^bold>\(f \\<^sub>B g) \\<^sub>B h\<^bold>\ \<^bold>\f \\<^sub>B g \\<^sub>B h\<^bold>\ \\<^sub>B[f, g, h])" using f g h fg gh B.\.preserves_hom B.HoHV_def B.HoVH_def by auto thus ?thesis using f g h fg gh comp_def B.comp_arr_dom B.comp_cod_arr by simp qed also have "... = MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) \<^bold>\f \\<^sub>B g \\<^sub>B h\<^bold>\ (f \\<^sub>B g \\<^sub>B h)" using B.comp_cod_arr B.comp_arr_inv' by (auto simp add: f fg g gh h) finally show ?thesis by simp qed also have "... = \ (f, g \\<^sub>B h) \ (UP f \ \ (g, h)) \ \ (UP f) (UP g) (UP h)" proof - have "\ (f, g \\<^sub>B h) \ (UP f \ \ (g, h)) \ \ (UP f) (UP g) (UP h) = \ (f, g \\<^sub>B h) \ (MkIde \<^bold>\f\<^bold>\ \ \ (g, h)) \ (MkIde \<^bold>\f\<^bold>\ \ MkIde \<^bold>\g\<^bold>\ \ MkIde \<^bold>\h\<^bold>\)" using f g h fg gh VVV.arr_char VV.arr_char arr_char src_def trg_def UP_def \_def by auto also have "... = MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g \\<^sub>B h\<^bold>\) \<^bold>\f \\<^sub>B g \\<^sub>B h\<^bold>\ (f \\<^sub>B g \\<^sub>B h) \ MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g \\<^sub>B h\<^bold>\) (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g \\<^sub>B h\<^bold>\) (f \\<^sub>B g \\<^sub>B h) \ MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g \\<^sub>B h\<^bold>\) (f \\<^sub>B g \\<^sub>B h) \ MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) (f \\<^sub>B g \\<^sub>B h)" proof - have "\ (f, g \\<^sub>B h) = MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g \\<^sub>B h\<^bold>\) \<^bold>\f \\<^sub>B g \\<^sub>B h\<^bold>\ (f \\<^sub>B g \\<^sub>B h) \ MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g \\<^sub>B h\<^bold>\) (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g \\<^sub>B h\<^bold>\) (f \\<^sub>B g \\<^sub>B h)" using f g h fg gh \.map_simp_ide \.map_def UP.FF_def UP_def hcomp_def B.VV.arr_char B.can_Ide_self B.comp_arr_dom B.comp_cod_arr src_def trg_def arr_char apply simp_all by blast moreover have "\ (g, h) = MkArr (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) \<^bold>\g \\<^sub>B h\<^bold>\ (g \\<^sub>B h)" using f g h fg gh \.map_def UP.FF_def UP_def hcomp_def B.VV.arr_char B.can_Ide_self src_def trg_def arr_char by auto moreover have "MkIde \<^bold>\f\<^bold>\ \ MkArr (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) \<^bold>\g \\<^sub>B h\<^bold>\ (g \\<^sub>B h) = MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g \\<^sub>B h\<^bold>\) (f \\<^sub>B g \\<^sub>B h)" using f g h fg gh hcomp_def arr_char src_def trg_def B.can_Ide_self B.comp_arr_dom B.comp_cod_arr by auto moreover have "MkIde \<^bold>\f\<^bold>\ \ MkIde \<^bold>\g\<^bold>\ \ MkIde \<^bold>\h\<^bold>\ = MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) (f \\<^sub>B g \\<^sub>B h)" proof - have "\f : f \\<^sub>B f\ \ \g : g \\<^sub>B g\ \ \h : h \\<^sub>B h\" using f g h by auto thus ?thesis using f g h fg gh hcomp_def arr_char src_def trg_def B.can_Ide_self B.comp_arr_dom B.comp_cod_arr by simp qed ultimately show ?thesis using comp_assoc by auto qed also have "... = MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) \<^bold>\f \\<^sub>B g \\<^sub>B h\<^bold>\ (f \\<^sub>B g \\<^sub>B h)" proof - have "Arr (MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) (f \\<^sub>B g \\<^sub>B h)) \ Arr (MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g \\<^sub>B h\<^bold>\) (f \\<^sub>B g \\<^sub>B h)) \ Arr (MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g \\<^sub>B h\<^bold>\) (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g \\<^sub>B h\<^bold>\) (f \\<^sub>B g \\<^sub>B h)) \ Arr (MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g \\<^sub>B h\<^bold>\) \<^bold>\f \\<^sub>B g \\<^sub>B h\<^bold>\ (f \\<^sub>B g \\<^sub>B h))" using f g h fg gh by auto thus ?thesis using f g h fg gh comp_def by auto qed finally show ?thesis by simp qed finally show ?thesis by blast qed qed lemma UP_is_pseudofunctor: shows "pseudofunctor V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B vcomp hcomp \ \ src trg UP \" .. lemma UP_map\<^sub>0_obj [simp]: assumes "B.obj a" shows "UP.map\<^sub>0 a = UP\<^sub>0 a" using assms UP.map\<^sub>0_def by auto interpretation UP: full_functor V\<^sub>B vcomp UP proof fix \ f g assume f: "B.ide f" and g: "B.ide g" assume \: "\\ : UP f \ UP g\" show "\\. \\ : f \\<^sub>B g\ \ UP \ = \" proof - have 1: "\Map \ : f \\<^sub>B g\" using f g \ UP_def arr_char in_hom_char by auto moreover have "UP (Map \) = \" proof - have "\ = MkArr (Dom \) (Cod \) (Map \)" using \ MkArr_Map by auto also have "... = UP (Map \)" using f g \ 1 UP_def arr_char dom_char cod_char apply simp by (metis (no_types, lifting) B.in_homE Dom.simps(1) in_homE) finally show ?thesis by auto qed ultimately show ?thesis by blast qed qed interpretation UP: faithful_functor V\<^sub>B vcomp UP using arr_char UP_def by (unfold_locales, simp_all) interpretation UP: fully_faithful_functor V\<^sub>B vcomp UP .. lemma UP_is_fully_faithful_functor: shows "fully_faithful_functor V\<^sub>B vcomp UP" .. no_notation B.in_hom ("\_ : _ \\<^sub>B _\") (* Inherited from functor, I think. *) lemma Map_reflects_hhom: assumes "B.obj a" and "B.obj b" and "ide g" and "\g : UP.map\<^sub>0 a \ UP.map\<^sub>0 b\" shows "\Map g : a \\<^sub>B b\" proof have 1: "B.ide (Map g)" using assms ide_char by blast show "B.arr (Map g)" using 1 by simp show "src\<^sub>B (Map g) = a" proof - have "src\<^sub>B (Map g) = Map (src g)" using assms src_def apply simp by (metis (no_types, lifting) E.eval_simps(2) E.Ide_implies_Arr arr_char ideE) also have "... = Map (UP.map\<^sub>0 a)" using assms by (metis (no_types, lifting) in_hhomE) also have "... = a" using assms UP.map\<^sub>0_def UP_def [of a] src_def by auto finally show ?thesis by simp qed show "trg\<^sub>B (Map g) = b" proof - have "trg\<^sub>B (Map g) = Map (trg g)" using assms trg_def apply simp by (metis (no_types, lifting) E.eval_simps(3) E.Ide_implies_Arr arr_char ideE) also have "... = Map (UP.map\<^sub>0 b)" using assms by (metis (no_types, lifting) in_hhomE) also have "... = b" using assms UP.map\<^sub>0_def UP_def [of b] src_def by auto finally show ?thesis by simp qed qed lemma eval_Dom_ide [simp]: assumes "ide g" shows "\Dom g\ = Map g" using assms dom_char ideD by auto lemma Cod_ide: assumes "ide f" shows "Cod f = Dom f" using assms dom_char by auto lemma Map_preserves_objects: assumes "obj a" shows "B.obj (Map a)" proof - have "src\<^sub>B (Map a) = Map (src a)" using assms src_def apply simp using E.eval_simps(2) E.Ide_implies_Arr arr_char ideE by (metis (no_types, lifting) objE) also have 1: "... = \E.Src (Dom a)\" using assms src_def by auto also have "... = \\<^bold>\Map a\<^bold>\\<^sub>0\" using assms B.src.is_extensional 1 by force also have "... = Map a" using assms by auto finally have "src\<^sub>B (Map a) = Map a" by simp moreover have "B.arr (Map a)" using assms B.ideD arr_char by auto ultimately show ?thesis using B.obj_def by simp qed interpretation UP: equivalence_pseudofunctor V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B vcomp hcomp \ \ src trg UP \ proof (* UP is full, hence locally full. *) show "\f f' \. \ B.ide f; B.ide f'; src\<^sub>B f = src\<^sub>B f'; trg\<^sub>B f = trg\<^sub>B f'; \\ : UP f \ UP f'\ \ \ \\. \\ : f \\<^sub>B f'\ \ UP \ = \" using UP.is_full by simp (* UP is essentially surjective up to equivalence on objects. *) show "\b. obj b \ \a. B.obj a \ equivalent_objects (UP.map\<^sub>0 a) b" proof - fix b assume b: "obj b" have 1: "B.obj (Map b)" using b Map_preserves_objects by simp have 3: "UP.map\<^sub>0 (Map b) = MkArr \<^bold>\Map b\<^bold>\\<^sub>0 \<^bold>\Map b\<^bold>\\<^sub>0 (Map b)" using b 1 UP.map\<^sub>0_def [of "Map b"] UP_def src_def arr_char by auto have 4: "b = MkArr (Dom b) (Dom b) (Map b)" using b objE eval_Dom_ide by (metis (no_types, lifting) dom_char ideD(2) obj_def) let ?\ = "MkArr \<^bold>\Map b\<^bold>\\<^sub>0 (Dom b) (Map b)" have \: "arr ?\" proof - have 2: "E.Obj (Dom b)" proof - have "Dom b = Dom (src b)" using b obj_def by simp moreover have "Dom (src b) = E.Src (Dom b)" using b obj_def src_def arr_char by simp moreover have "E.Obj (E.Src (Dom b))" using b obj_def src_def arr_char arr_def E.Obj_Src by simp ultimately show ?thesis by simp qed have "E.Nml \<^bold>\Map b\<^bold>\\<^sub>0 \ E.Ide \<^bold>\Map b\<^bold>\\<^sub>0" using 1 by auto moreover have "E.Nml (Dom b) \ E.Ide (Dom b)" using b arr_char [of b] by auto moreover have "E.Src \<^bold>\Map b\<^bold>\\<^sub>0 = E.Src (Dom b)" using b 1 2 B.obj_def obj_char by (cases "Dom b", simp_all) moreover have "E.Trg \<^bold>\Map b\<^bold>\\<^sub>0 = E.Trg (Dom b)" using b 1 2 B.obj_def obj_char by (cases "Dom b", simp_all) moreover have "\Map b : \\<^bold>\Map b\<^bold>\\<^sub>0\ \\<^sub>B \Dom b\\" using b 1 by (elim objE, auto) ultimately show ?thesis using arr_char \E.Nml \<^bold>\Map b\<^bold>\\<^sub>0 \ E.Ide \<^bold>\Map b\<^bold>\\<^sub>0\ by auto qed hence "iso ?\" using 1 iso_char by auto moreover have "dom ?\ = UP.map\<^sub>0 (Map b)" using \ dom_char b 1 3 B.objE UP.map\<^sub>0_def UP_def src_def by auto moreover have "cod ?\ = b" using \ cod_char b 4 1 by auto ultimately have "isomorphic (UP.map\<^sub>0 (Map b)) b" using \ 3 4 isomorphic_def by blast moreover have 5: "obj (UP.map\<^sub>0 (Map b))" using 1 UP.map\<^sub>0_simps(2) by simp ultimately have 6: "UP.map\<^sub>0 (Map b) = b" using b isomorphic_objects_are_equal by simp have "equivalent_objects (UP.map\<^sub>0 (Map b)) b" using b 6 equivalent_objects_reflexive [of b] by simp thus "\a. B.obj a \ equivalent_objects (UP.map\<^sub>0 a) b" using 1 6 by auto qed (* UP is locally essentially surjective. *) show "\a b g. \ B.obj a; B.obj b; \g : UP.map\<^sub>0 a \ UP.map\<^sub>0 b\; ide g \ \ \f. \f : a \\<^sub>B b\ \ B.ide f \ isomorphic (UP f) g" proof - fix a b g assume a: "B.obj a" and b: "B.obj b" assume g_in_hhom: "\g : UP.map\<^sub>0 a \ UP.map\<^sub>0 b\" assume ide_g: "ide g" have 1: "B.ide (Map g)" using ide_g ide_char by blast have "arr (UP a)" using a by auto have "arr (UP b)" using b by auto have Map_g_eq: "Map g = \Dom g\" using ide_g by simp have Map_g_in_hhom: "\Map g : a \\<^sub>B b\" using a b ide_g g_in_hhom Map_reflects_hhom by simp let ?\ = "MkArr \<^bold>\Map g\<^bold>\ (Dom g) (Map g)" have \: "arr ?\" proof - have "\Map ?\ : \Dom ?\\ \\<^sub>B \Cod ?\\\" using 1 Map_g_eq by auto moreover have "E.Ide \<^bold>\Map g\<^bold>\ \ E.Nml \<^bold>\Map g\<^bold>\" using 1 by simp moreover have "E.Ide (Dom g) \ E.Nml (Dom g)" using ide_g arr_char ide_char by blast moreover have "E.Src \<^bold>\Map g\<^bold>\ = E.Src (Dom g)" using ide_g g_in_hhom src_def Map_g_in_hhom by (metis (no_types, lifting) B.ideD(2) B.in_hhom_def B.objE B.obj_def' Dom.simps(1) E.Src.simps(2) UP.map\<^sub>0_def \arr (UP a)\ a in_hhomE UP_def) moreover have "E.Trg \<^bold>\Map g\<^bold>\ = E.Trg (Dom g)" proof - have "E.Trg \<^bold>\Map g\<^bold>\ = \<^bold>\b\<^bold>\\<^sub>0" using Map_g_in_hhom by auto also have "... = E.Trg (Dom g)" proof - have "E.Trg (Dom g) = Dom (trg g)" using ide_g trg_def by simp also have "... = Dom (UP.map\<^sub>0 b)" using g_in_hhom by auto also have "... = \<^bold>\b\<^bold>\\<^sub>0" using b \arr (UP b)\ UP.map\<^sub>0_def src_def UP_def B.objE by auto finally show ?thesis by simp qed finally show ?thesis by simp qed ultimately show ?thesis using arr_char by simp qed have "\?\ : UP (Map g) \ g\" proof show "arr ?\" using \ by simp show "dom ?\ = UP (Map g)" using \ 1 dom_char UP_def by simp show "cod ?\ = g" proof - have "cod ?\ = MkArr (Dom g) (Dom g) (Map g)" using ide_g cod_char Map_g_eq \ by auto moreover have "Dom g = Cod g" using ide_g Cod_ide by simp ultimately have "cod ?\ = MkArr (Dom g) (Cod g) (Map g)" by simp thus ?thesis by (metis (no_types, lifting) "1" B.comp_ide_self \Dom g = Cod g\ comp_cod_arr ideD(1) ideD(3) ide_g comp_char) qed qed moreover have "iso ?\" using \ 1 iso_char by simp ultimately have "isomorphic (UP (Map g)) g" using isomorphic_def by auto thus "\f. \f : a \\<^sub>B b\ \ B.ide f \ isomorphic (UP f) g" using 1 Map_g_in_hhom by auto qed qed theorem UP_is_equivalence_pseudofunctor: shows "equivalence_pseudofunctor V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B vcomp hcomp \ \ src trg UP \" .. text \ Next, we work out the details of the equivalence pseudofunctor \DN\ in the converse direction. \ definition DN where "DN \ \ if arr \ then Map \ else B.null" lemma DN_in_hom [intro]: assumes "arr \" shows "\DN \ : DN (src \) \\<^sub>B DN (trg \)\" and "\DN \ : DN (dom \) \\<^sub>B DN (cod \)\" using assms DN_def arr_char [of \] B.vconn_implies_hpar(1-2) E.eval_in_hom(1) B.in_hhom_def by auto lemma DN_simps [simp]: assumes "arr \" shows "B.arr (DN \)" and "src\<^sub>B (DN \) = DN (src \)" and "trg\<^sub>B (DN \) = DN (trg \)" and "B.dom (DN \) = DN (dom \)" and "B.cod (DN \) = DN (cod \)" using assms DN_in_hom by auto interpretation "functor" vcomp V\<^sub>B DN using DN_def seqE Map_comp seq_char by (unfold_locales, auto) interpretation DN: weak_arrow_of_homs vcomp src trg V\<^sub>B src\<^sub>B trg\<^sub>B DN proof fix \ assume \: "arr \" show "B.isomorphic (DN (src \)) (src\<^sub>B (DN \))" proof - have "DN (src \) = src\<^sub>B (DN \)" using \ DN_def arr_char E.eval_simps(2) E.Ide_implies_Arr apply simp by (metis (no_types, lifting) B.vconn_implies_hpar(1) E.Nml_implies_Arr ideE ide_src src_simps(3)) moreover have "B.ide (DN (src \))" using \ by simp ultimately show ?thesis using \ B.isomorphic_reflexive by auto qed show "B.isomorphic (DN (trg \)) (trg\<^sub>B (DN \))" proof - have "DN (trg \) = trg\<^sub>B (DN \)" using \ DN_def arr_char E.eval_simps(3) E.Ide_implies_Arr apply simp by (metis (no_types, lifting) B.vconn_implies_hpar(2) E.Nml_implies_Arr ideE ide_trg trg_simps(3)) moreover have "B.ide (DN (trg \))" using \ by simp ultimately show ?thesis using B.isomorphic_reflexive by auto qed qed interpretation "functor" VV.comp B.VV.comp DN.FF using DN.functor_FF by auto interpretation HoDN_DN: composite_functor VV.comp B.VV.comp V\<^sub>B DN.FF \\\\. H\<^sub>B (fst \\) (snd \\)\ .. interpretation DNoH: composite_functor VV.comp vcomp V\<^sub>B \\\\. fst \\ \ snd \\\ DN .. abbreviation \\<^sub>o where "\\<^sub>o fg \ B.can (Dom (fst fg) \<^bold>\\<^bold>\\<^bold>\ Dom (snd fg)) (Dom (fst fg) \<^bold>\ Dom (snd fg))" abbreviation \\<^sub>o' where "\\<^sub>o' fg \ B.can (Dom (fst fg) \<^bold>\ Dom (snd fg)) (Dom (fst fg) \<^bold>\\<^bold>\\<^bold>\ Dom (snd fg))" lemma \\<^sub>o_in_hom: assumes "VV.ide fg" shows "\\\<^sub>o fg : Map (fst fg) \\<^sub>B Map (snd fg) \\<^sub>B \Dom (fst fg) \<^bold>\\<^bold>\\<^bold>\ Dom (snd fg)\\" and "\\\<^sub>o' fg : \Dom (fst fg) \<^bold>\\<^bold>\\<^bold>\ Dom (snd fg)\ \\<^sub>B Map (fst fg) \\<^sub>B Map (snd fg)\" and "B.inverse_arrows (\\<^sub>o fg) (\\<^sub>o' fg)" proof - have 1: "E.Ide (Dom (fst fg) \<^bold>\ Dom (snd fg))" unfolding E.Ide.simps(3) apply (intro conjI) using assms VV.ide_char VV.arr_char arr_char apply simp using VV.arr_char VV.ideD(1) assms apply blast by (metis (no_types, lifting) VV.arrE VV.ideD(1) assms src_simps(1) trg_simps(1)) have 2: "E.Ide (Dom (fst fg) \<^bold>\\<^bold>\\<^bold>\ Dom (snd fg))" using 1 by (meson E.Ide.simps(3) E.Ide_HcompNml VV.arr_char VV.ideD(1) arr_char assms) have 3: "\<^bold>\Dom (fst fg) \<^bold>\ Dom (snd fg)\<^bold>\ = \<^bold>\Dom (fst fg) \<^bold>\\<^bold>\\<^bold>\ Dom (snd fg)\<^bold>\" by (metis (no_types, lifting) E.Ide.simps(3) E.Nml_HcompNml(1) E.Nmlize.simps(3) E.Nmlize_Nml VV.arr_char VV.ideD(1) arr_char assms 1) have 4: "\Dom (fst fg) \<^bold>\ Dom (snd fg)\ = Map (fst fg) \\<^sub>B Map (snd fg)" using assms VV.ide_char VV.arr_char arr_char by simp show "\\\<^sub>o fg : Map (fst fg) \\<^sub>B Map (snd fg) \\<^sub>B \Dom (fst fg) \<^bold>\\<^bold>\\<^bold>\ Dom (snd fg)\\" using 1 2 3 4 by auto show "\\\<^sub>o' fg : \Dom (fst fg) \<^bold>\\<^bold>\\<^bold>\ Dom (snd fg)\ \\<^sub>B Map (fst fg) \\<^sub>B Map (snd fg)\" using 1 2 3 4 by auto show "B.inverse_arrows (\\<^sub>o fg) (\\<^sub>o' fg)" using 1 2 3 B.inverse_arrows_can by blast qed interpretation \: transformation_by_components VV.comp V\<^sub>B HoDN_DN.map DNoH.map \\<^sub>o proof fix fg assume fg: "VV.ide fg" have 1: "\Dom (fst fg) \<^bold>\ Dom (snd fg)\ = Map (fst fg) \\<^sub>B Map (snd fg)" using fg VV.ide_char VV.arr_char arr_char by simp show "\\\<^sub>o fg : HoDN_DN.map fg \\<^sub>B DNoH.map fg\" proof show "B.arr (\\<^sub>o fg)" using fg \\<^sub>o_in_hom by blast show "B.dom (\\<^sub>o fg) = HoDN_DN.map fg" proof - have "B.dom (\\<^sub>o fg) = Map (fst fg) \\<^sub>B Map (snd fg)" using fg \\<^sub>o_in_hom by blast also have "... = HoDN_DN.map fg" using fg DN.FF_def DN_def VV.arr_char src_def trg_def VV.ide_char by auto finally show ?thesis by simp qed show "B.cod (\\<^sub>o fg) = DNoH.map fg" proof - have "B.cod (\\<^sub>o fg) = \Dom (fst fg) \<^bold>\\<^bold>\\<^bold>\ Dom (snd fg)\" using fg \\<^sub>o_in_hom by blast also have "... = DNoH.map fg" proof - have "DNoH.map fg = B.can (Cod (fst fg) \<^bold>\\<^bold>\\<^bold>\ Cod (snd fg)) (Cod (fst fg) \<^bold>\ Cod (snd fg)) \\<^sub>B (Map (fst fg) \\<^sub>B Map (snd fg)) \\<^sub>B B.can (Dom (fst fg) \<^bold>\ Dom (snd fg)) (Dom (fst fg) \<^bold>\\<^bold>\\<^bold>\ Dom (snd fg))" using fg DN_def Map_hcomp VV.arr_char apply simp using VV.ideD(1) by blast also have "... = B.can (Cod (fst fg) \<^bold>\\<^bold>\\<^bold>\ Cod (snd fg)) (Cod (fst fg) \<^bold>\ Cod (snd fg)) \\<^sub>B B.can (Dom (fst fg) \<^bold>\ Dom (snd fg)) (Dom (fst fg) \<^bold>\\<^bold>\\<^bold>\ Dom (snd fg))" proof - have "(Map (fst fg) \\<^sub>B Map (snd fg)) \\<^sub>B B.can (Dom (fst fg) \<^bold>\ Dom (snd fg)) (Dom (fst fg) \<^bold>\\<^bold>\\<^bold>\ Dom (snd fg)) = B.can (Dom (fst fg) \<^bold>\ Dom (snd fg)) (Dom (fst fg) \<^bold>\\<^bold>\\<^bold>\ Dom (snd fg))" using fg 1 \\<^sub>o_in_hom B.comp_cod_arr by blast thus ?thesis by simp qed also have "... = \Dom (fst fg) \<^bold>\\<^bold>\\<^bold>\ Dom (snd fg)\" proof - have "B.can (Cod (fst fg) \<^bold>\\<^bold>\\<^bold>\ Cod (snd fg)) (Cod (fst fg) \<^bold>\ Cod (snd fg)) = \\<^sub>o fg" using fg VV.ide_char Cod_ide by simp thus ?thesis using fg 1 \\<^sub>o_in_hom [of fg] B.comp_arr_inv' by fastforce qed finally show ?thesis by simp qed finally show ?thesis by blast qed qed next show "\f. VV.arr f \ \\<^sub>o (VV.cod f) \\<^sub>B HoDN_DN.map f = DNoH.map f \\<^sub>B \\<^sub>o (VV.dom f)" proof - fix \\ assume \\: "VV.arr \\" show "\\<^sub>o (VV.cod \\) \\<^sub>B HoDN_DN.map \\ = DNoH.map \\ \\<^sub>B \\<^sub>o (VV.dom \\)" proof - have 1: "E.Ide (Dom (fst \\) \<^bold>\ Dom (snd \\))" unfolding E.Ide.simps(3) apply (intro conjI) using \\ VV.ide_char VV.arr_char arr_char apply simp using VV.arr_char VV.ideD(1) \\ apply blast by (metis (no_types, lifting) VV.arrE \\ src_simps(1) trg_simps(1)) have 2: "E.Ide (Dom (fst \\) \<^bold>\\<^bold>\\<^bold>\ Dom (snd \\))" using 1 by (meson E.Ide.simps(3) E.Ide_HcompNml VV.arr_char VV.ideD(1) arr_char \\) have 3: "\<^bold>\Dom (fst \\) \<^bold>\ Dom (snd \\)\<^bold>\ = \<^bold>\Dom (fst \\) \<^bold>\\<^bold>\\<^bold>\ Dom (snd \\)\<^bold>\" by (metis (no_types, lifting) E.Ide.simps(3) E.Nml_HcompNml(1) E.Nmlize.simps(3) E.Nmlize_Nml VV.arr_char arr_char \\ 1) have 4: "E.Ide (Cod (fst \\) \<^bold>\ Cod (snd \\))" unfolding E.Ide.simps(3) apply (intro conjI) using \\ VV.ide_char VV.arr_char arr_char apply simp using VV.arr_char VV.ideD(1) \\ apply blast by (metis (no_types, lifting) "1" E.Ide.simps(3) VV.arrE \\ arrE) have 5: "E.Ide (Cod (fst \\) \<^bold>\\<^bold>\\<^bold>\ Cod (snd \\))" using 4 by (meson E.Ide.simps(3) E.Ide_HcompNml VV.arr_char VV.ideD(1) arr_char \\) have 6: "\<^bold>\Cod (fst \\) \<^bold>\ Cod (snd \\)\<^bold>\ = \<^bold>\Cod (fst \\) \<^bold>\\<^bold>\\<^bold>\ Cod (snd \\)\<^bold>\" by (metis (no_types, lifting) E.Ide.simps(3) E.Nml_HcompNml(1) E.Nmlize.simps(3) E.Nmlize_Nml VV.arr_char arr_char \\ 1) have A: "\\\<^sub>o' \\ : \Dom (fst \\) \<^bold>\\<^bold>\\<^bold>\ Dom (snd \\)\ \\<^sub>B \Dom (fst \\) \<^bold>\ Dom (snd \\)\\" using 1 2 3 by auto have B: "\Map (fst \\) \\<^sub>B Map (snd \\) : \Dom (fst \\) \<^bold>\ Dom (snd \\)\ \\<^sub>B \Cod (fst \\) \<^bold>\ Cod (snd \\)\\" using \\ VV.arr_char arr_char src_def trg_def E.Nml_implies_Arr E.eval_simps'(2-3) by auto have C: "\B.can (Cod (fst \\) \<^bold>\\<^bold>\\<^bold>\ Cod (snd \\)) (Cod (fst \\) \<^bold>\ Cod (snd \\)) : \Cod (fst \\) \<^bold>\ Cod (snd \\)\ \\<^sub>B \Cod (fst \\) \<^bold>\\<^bold>\\<^bold>\ Cod (snd \\)\\" using 4 5 6 by auto have "\\<^sub>o (VV.cod \\) \\<^sub>B HoDN_DN.map \\ = B.can (Cod (fst \\) \<^bold>\\<^bold>\\<^bold>\ Cod (snd \\)) (Cod (fst \\) \<^bold>\ Cod (snd \\)) \\<^sub>B (Map (fst \\) \\<^sub>B Map (snd \\))" using \\ VV.arr_char VV.cod_char arr_char src_def trg_def cod_char DN.FF_def DN_def by auto also have "... = B.can (Cod (fst \\) \<^bold>\\<^bold>\\<^bold>\ Cod (snd \\)) (Cod (fst \\) \<^bold>\ Cod (snd \\)) \\<^sub>B (Map (fst \\) \\<^sub>B Map (snd \\)) \\<^sub>B \\<^sub>o' \\ \\<^sub>B \\<^sub>o \\" using B B.comp_assoc \\ VV.arr_char arr_char src_def trg_def B.inverse_arrows_can E.Ide_HcompNml E.Nmlize_Nml E.Nml_HcompNml(1) B.can_Ide_self B.comp_arr_dom by auto also have "... = DNoH.map \\ \\<^sub>B \\<^sub>o (VV.dom \\)" proof - have "DNoH.map \\ \\<^sub>B \\<^sub>o (VV.dom \\) = B.can (Cod (fst \\) \<^bold>\\<^bold>\\<^bold>\ Cod (snd \\)) (Cod (fst \\) \<^bold>\ Cod (snd \\)) \\<^sub>B (Map (fst \\) \\<^sub>B Map (snd \\)) \\<^sub>B \\<^sub>o' \\ \\<^sub>B \\<^sub>o (VV.dom \\)" using \\ DN_def VV.arr_char B.comp_assoc by simp moreover have "\\<^sub>o (VV.dom \\) = \\<^sub>o \\" using \\ VV.dom_char VV.arr_char by auto ultimately show ?thesis using B.comp_assoc by simp qed finally show ?thesis by blast qed qed qed abbreviation \ where "\ \ \.map" interpretation \: natural_isomorphism VV.comp V\<^sub>B HoDN_DN.map DNoH.map \ proof show "\fg. VV.ide fg \ B.iso (\ fg)" proof - fix fg assume fg: "VV.ide fg" have "B.inverse_arrows (\\<^sub>o fg) (\\<^sub>o' fg)" using fg \\<^sub>o_in_hom by simp thus "B.iso (\ fg)" using fg B.iso_def \.map_simp_ide by auto qed qed no_notation B.in_hom ("\_ : _ \\<^sub>B _\") lemma \_in_hom [intro]: assumes "arr (fst \\)" and "arr (snd \\)" and "src (fst \\) = trg (snd \\)" shows "\\ \\ : DN (src (snd \\)) \\<^sub>B DN (trg (fst \\))\" and "\\ \\ : DN (dom (fst \\)) \\<^sub>B DN (dom (snd \\)) \\<^sub>B DN (cod (fst \\) \ cod (snd \\))\" proof - have 1: "VV.arr \\" using assms VV.arr_char by simp show 2: "\\ \\ : DN (dom (fst \\)) \\<^sub>B DN (dom (snd \\)) \\<^sub>B DN (cod (fst \\) \ cod (snd \\))\" proof - have "HoDN_DN.map (VV.dom \\) = DN (dom (fst \\)) \\<^sub>B DN (dom (snd \\))" using assms 1 DN.FF_def by auto moreover have "DNoH.map (VV.cod \\) = DN (cod (fst \\) \ cod (snd \\))" using assms 1 by simp ultimately show ?thesis using assms 1 \.preserves_hom by auto qed show "\\ \\ : DN (src (snd \\)) \\<^sub>B DN (trg (fst \\))\" using assms 1 2 B.src_dom [of "\ \\"] B.trg_dom [of "\ \\"] by auto qed lemma \_simps [simp]: assumes "arr (fst \\)" and "arr (snd \\)" and "src (fst \\) = trg (snd \\)" shows "B.arr (\ \\)" and "src\<^sub>B (\ \\) = DN (src (snd \\))" and "trg\<^sub>B (\ \\) = DN (trg (fst \\))" and "B.dom (\ \\) = DN (dom (fst \\)) \\<^sub>B DN (dom (snd \\))" and "B.cod (\ \\) = DN (cod (fst \\) \ cod (snd \\))" proof show "VV.arr \\" using assms by blast have 1: "\\ \\ : DN (src (snd \\)) \\<^sub>B DN (trg (fst \\))\" using assms by blast show "src\<^sub>B (\ \\) = DN (src (snd \\))" using 1 by fast show "trg\<^sub>B (\ \\) = DN (trg (fst \\))" using 1 by fast have 2: "\\ \\ : DN (dom (fst \\)) \\<^sub>B DN (dom (snd \\)) \\<^sub>B DN (cod (fst \\) \ cod (snd \\))\" using assms by blast show "B.dom (\ \\) = DN (dom (fst \\)) \\<^sub>B DN (dom (snd \\))" using 2 by fast show "B.cod (\ \\) = DN (cod (fst \\) \ cod (snd \\))" using 2 by fast qed interpretation DN: pseudofunctor vcomp hcomp \ \ src trg V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B DN \ proof show "\f g h. \ ide f; ide g; ide h; src f = trg g; src g = trg h \ \ DN (\ f g h) \\<^sub>B \ (f \ g, h) \\<^sub>B (\ (f, g) \\<^sub>B DN h) = \ (f, g \ h) \\<^sub>B (DN f \\<^sub>B \ (g, h)) \\<^sub>B \\<^sub>B[DN f, DN g, DN h]" proof - fix f g h assume f: "ide f" and g: "ide g" and h: "ide h" and fg: "src f = trg g" and gh: "src g = trg h" show "DN (\ f g h) \\<^sub>B \ (f \ g, h) \\<^sub>B (\ (f, g) \\<^sub>B DN h) = \ (f, g \ h) \\<^sub>B (DN f \\<^sub>B \ (g, h)) \\<^sub>B \\<^sub>B[DN f, DN g, DN h]" proof - have 1: "E.Trg (Dom g) = E.Trg (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) \ \E.Trg (Dom g)\ = \E.Trg (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h)\" using f g h fg gh arr_char src_def trg_def E.Trg_HcompNml by (metis (no_types, lifting) ideD(1) src_simps(2) trg_simps(2)) have 2: "arr (MkArr (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Cod f \<^bold>\\<^bold>\\<^bold>\ Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) (B.can (Cod f \<^bold>\\<^bold>\\<^bold>\ Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) (Cod f \<^bold>\ Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) \\<^sub>B (Map f \\<^sub>B B.can (Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) (Cod g \<^bold>\ Cod h) \\<^sub>B (Map g \\<^sub>B Map h) \\<^sub>B B.can (Dom g \<^bold>\ Dom h) (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h)) \\<^sub>B B.can (Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h)))" proof - have "\B.can (Cod f \<^bold>\\<^bold>\\<^bold>\ Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) (Cod f \<^bold>\ Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) \\<^sub>B (Map f \\<^sub>B B.can (Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) (Cod g \<^bold>\ Cod h) \\<^sub>B (Map g \\<^sub>B Map h) \\<^sub>B B.can (Dom g \<^bold>\ Dom h) (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h)) \\<^sub>B B.can (Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) : EVAL (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) \\<^sub>B EVAL (Cod f \<^bold>\\<^bold>\\<^bold>\ Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h)\" proof (intro B.comp_in_homI) show 2: "\B.can (Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) : EVAL (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) \\<^sub>B EVAL (Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h)\" using f g h fg gh 1 apply (intro B.can_in_hom) apply (metis (no_types, lifting) E.Ide_HcompNml E.Nml_HcompNml(1) arr_char ideD(1) src_simps(1) trg_simps(1)) apply (metis (no_types, lifting) E.Ide.simps(3) E.Ide_HcompNml ideD(1) arr_char src_simps(1) trg_simps(1)) by (metis (no_types, lifting) E.Nml_HcompNml(1) E.Nmlize.simps(3) E.Nmlize_Nml ideD(1) arr_char src_simps(1) trg_simps(1)) show "\B.can (Cod f \<^bold>\\<^bold>\\<^bold>\ Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) (Cod f \<^bold>\ Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) : EVAL (Cod f \<^bold>\ Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) \\<^sub>B EVAL (Cod f \<^bold>\\<^bold>\\<^bold>\ Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h)\" proof - have "E.Ide (Cod f \<^bold>\ Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h)" using f g h fg gh 1 Cod_ide E.Ide_HcompNml arr_char apply simp by (metis (no_types, lifting) ideD(1) src_simps(1) trg_simps(1)) moreover have "E.Ide (Cod f \<^bold>\\<^bold>\\<^bold>\ Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h)" using f g h fg gh by (metis (no_types, lifting) E.Ide.simps(3) E.Ide_HcompNml E.Nml_HcompNml(1) arr_char calculation ideD(1) src_simps(1) trg_simps(1)) moreover have "E.Nmlize (Cod f \<^bold>\ Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) = E.Nmlize (Cod f \<^bold>\\<^bold>\\<^bold>\ Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h)" using f g h fg gh by (metis (no_types, lifting) E.Ide.simps(3) E.Nml_HcompNml(1) E.Nmlize.simps(3) E.Nmlize_Nml arr_char calculation(1) ideD(1) src_simps(1) trg_simps(1)) ultimately show ?thesis using B.can_in_hom [of "Cod f \<^bold>\ Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h" "Cod f \<^bold>\\<^bold>\\<^bold>\ Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h"] by blast qed show "\Map f \\<^sub>B B.can (Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) (Cod g \<^bold>\ Cod h) \\<^sub>B (Map g \\<^sub>B Map h) \\<^sub>B B.can (Dom g \<^bold>\ Dom h) (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) : EVAL (Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) \\<^sub>B EVAL (Cod f \<^bold>\ Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h)\" using f g h fg gh B.can_in_hom apply simp proof (intro B.hcomp_in_vhom B.comp_in_homI) show 1: "\B.can (Dom g \<^bold>\ Dom h) (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) : EVAL (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) \\<^sub>B EVAL (Dom g \<^bold>\ Dom h)\" using g h gh B.can_in_hom by (metis (no_types, lifting) E.Ide.simps(3) E.Ide_HcompNml E.Nml_HcompNml(1) E.Nmlize.simps(3) E.Nmlize_Nml arr_char ideD(1) src_simps(1) trg_simps(1)) show "\B.can (Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) (Cod g \<^bold>\ Cod h) : EVAL (Cod g \<^bold>\ Cod h) \\<^sub>B EVAL (Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h)\" using g h gh B.can_in_hom by (metis (no_types, lifting) Cod_ide E.Ide.simps(3) E.Ide_HcompNml E.Nml_HcompNml(1) E.Nmlize.simps(3) E.Nmlize_Nml arr_char ideD(1) src_simps(2) trg_simps(2)) show "\Map g \\<^sub>B Map h : EVAL (Dom g \<^bold>\ Dom h) \\<^sub>B EVAL (Cod g \<^bold>\ Cod h)\" using g h gh 1 Map_in_Hom B.hcomp_in_vhom B.not_arr_null B.seq_if_composable B.trg.is_extensional B.trg.preserves_hom B.vconn_implies_hpar(2) B.vconn_implies_hpar(4) Cod_ide E.eval.simps(3) Map_ide(1) arr_char ideD(1) by (metis (no_types, lifting)) show "\Map f : Map f \\<^sub>B EVAL (Cod f)\" using f arr_char Cod_ide by auto show "src\<^sub>B (Map f) = trg\<^sub>B \Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h\" using f g h fg gh 1 2 src_def trg_def B.arrI B.hseqE B.not_arr_null B.trg.is_extensional B.trg.preserves_hom B.vconn_implies_hpar(2) B.vconn_implies_hpar(4) E.eval.simps(3) by (metis (no_types, lifting) Map_ide(1)) qed qed thus ?thesis using f g h fg gh arr_char src_def trg_def E.Nml_HcompNml E.Ide_HcompNml ideD(1) apply (intro arr_MkArr) by auto qed have 3: "E.Ide (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h)" using g h gh ide_char arr_char src_def trg_def E.Ide_HcompNml Cod_ide by (metis (no_types, lifting) ideD(1) src_simps(2) trg_simps(2)) have 4: "E.Ide (Dom g \<^bold>\ Dom h)" using g h gh ide_char arr_char src_def trg_def Cod_ide by (metis (no_types, lifting) E.Ide.simps(3) arrE ideD(1) src_simps(2) trg_simps(2)) have 5: "E.Nmlize (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) = E.Nmlize (Dom g \<^bold>\ Dom h)" using g h gh ide_char arr_char src_def trg_def E.Nml_HcompNml by (metis (no_types, lifting) 4 E.Ide.simps(3) E.Nmlize.simps(3) E.Nmlize_Nml ideD(1)) have 6: "E.Ide (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h)" using f g h fg gh arr_char src_def trg_def by (metis (no_types, lifting) 1 E.Nml_HcompNml(1) E.Ide_HcompNml ideD(1) src_simps(2) trg_simps(2)) have 7: "E.Ide (Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h)" using f g h fg gh arr_char src_def trg_def by (metis (no_types, lifting) 1 3 E.Ide.simps(3) ideD(1) src_simps(2) trg_simps(2)) have 8: "E.Nmlize (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) = E.Nmlize (Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h)" using f g h fg gh arr_char src_def trg_def 7 E.Nml_HcompNml(1) ideD(1) by auto have "DN (\ f g h) \\<^sub>B \ (f \ g, h) \\<^sub>B (\ (f, g) \\<^sub>B DN h) = B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) ((Dom f \<^bold>\ Dom g) \<^bold>\ Dom h)" proof - have 9: "VVV.arr (f, g, h)" using f g h fg gh VVV.arr_char VV.arr_char arr_char ideD by simp have 10: "VV.ide (f, g)" using f g fg VV.ide_char by auto have 11: "VV.ide (hcomp f g, h)" using f g h fg gh VV.ide_char VV.arr_char by simp have 12: "arr (MkArr (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) (B.can (Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) (Cod g \<^bold>\ Cod h) \\<^sub>B (Map g \\<^sub>B Map h) \\<^sub>B B.can (Dom g \<^bold>\ Dom h) (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h)))" proof (intro arr_MkArr) show "Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h \ IDE" using g h gh by (metis (no_types, lifting) 3 E.Nml_HcompNml(1) arr_char ideD(1) mem_Collect_eq src_simps(2) trg_simps(2)) show "Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h \ IDE" using g h gh Cod_ide \Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h \ IDE\ by auto show "B.can (Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) (Cod g \<^bold>\ Cod h) \\<^sub>B (Map g \\<^sub>B Map h) \\<^sub>B B.can (Dom g \<^bold>\ Dom h) (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) \ HOM (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h)" proof show "E.Src (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) = E.Src (Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) \ E.Trg (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) = E.Trg (Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) \ \B.can (Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) (Cod g \<^bold>\ Cod h) \\<^sub>B (Map g \\<^sub>B Map h) \\<^sub>B B.can (Dom g \<^bold>\ Dom h) (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) : \Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h\ \\<^sub>B \Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h\\" proof (intro conjI) show "E.Src (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) = E.Src (Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h)" using g h gh Cod_ide by simp show "E.Trg (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) = E.Trg (Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h)" using g h gh Cod_ide by simp show "\B.can (Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) (Cod g \<^bold>\ Cod h) \\<^sub>B (Map g \\<^sub>B Map h) \\<^sub>B B.can (Dom g \<^bold>\ Dom h) (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) : \Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h\ \\<^sub>B \Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h\\" proof (intro B.comp_in_homI) show "\B.can (Dom g \<^bold>\ Dom h) (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) : \Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h\ \\<^sub>B \Dom g \<^bold>\ Dom h\\" using 3 4 5 by blast show "\Map g \\<^sub>B Map h : \Dom g \<^bold>\ Dom h\ \\<^sub>B \Cod g \<^bold>\ Cod h\\" using g h gh by (metis (no_types, lifting) 4 B.ide_in_hom(2) Cod_ide E.eval.simps(3) E.ide_eval_Ide Map_ide(2)) show "\B.can (Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) (Cod g \<^bold>\ Cod h) : \Cod g \<^bold>\ Cod h\ \\<^sub>B \Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h\\" using 3 4 5 Cod_ide g h by auto qed qed qed qed have "DN (\ f g h) = \Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h\" proof - have "DN (\ f g h) = (B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) \\<^sub>B ((Map f \\<^sub>B B.can (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom g \<^bold>\ Dom h) \\<^sub>B (Map g \\<^sub>B Map h) \\<^sub>B B.can (Dom g \<^bold>\ Dom h) (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h))) \\<^sub>B B.can (Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h))" using f g h fg gh 1 2 9 10 11 12 DN_def \_def hcomp_def src_def trg_def B.comp_assoc Cod_ide by simp also have "... = B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) \\<^sub>B (Map f \\<^sub>B \Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h\) \\<^sub>B B.can (Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h)" proof - have "\B.can (Dom g \<^bold>\ Dom h) (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) : \Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h\ \\<^sub>B Map g \\<^sub>B Map h\" using g h 3 4 5 B.can_in_hom [of "Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h" "Dom g \<^bold>\ Dom h"] by simp hence "Map f \\<^sub>B B.can (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom g \<^bold>\ Dom h) \\<^sub>B (Map g \\<^sub>B Map h) \\<^sub>B B.can (Dom g \<^bold>\ Dom h) (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) = Map f \\<^sub>B B.can (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom g \<^bold>\ Dom h) \\<^sub>B B.can (Dom g \<^bold>\ Dom h) (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h)" using B.comp_cod_arr by auto also have "... = Map f \\<^sub>B \Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h\" using f g h fg gh arr_char src_def trg_def B.vcomp_can B.can_Ide_self using 3 4 5 by auto finally have "Map f \\<^sub>B B.can (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom g \<^bold>\ Dom h) \\<^sub>B (Map g \\<^sub>B Map h) \\<^sub>B B.can (Dom g \<^bold>\ Dom h) (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) = Map f \\<^sub>B \Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h\" by simp thus ?thesis by simp qed also have "... = B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) \\<^sub>B B.can (Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h)" proof - have "\B.can (Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) : \Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h\ \\<^sub>B Map f \\<^sub>B \Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h\\" using f g h 6 7 8 B.can_in_hom [of "Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h" "Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h"] by simp hence "(Map f \\<^sub>B \Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h\) \\<^sub>B B.can (Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) = B.can (Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h)" using B.comp_cod_arr by auto thus ?thesis by simp qed also have "... = B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h)" using f g h fg gh src_def trg_def B.vcomp_can using 6 7 8 by auto also have "... = \Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h\" using f g h fg gh src_def trg_def B.can_Ide_self using 6 by blast finally show ?thesis by simp qed have "DN (\ f g h) \\<^sub>B \ (f \ g, h) \\<^sub>B (\ (f, g) \\<^sub>B DN h) = B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\\<^bold>\\<^bold>\ Dom h) \\<^sub>B B.can ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\\<^bold>\\<^bold>\ Dom h) ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\ Dom h) \\<^sub>B (B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) (Dom f \<^bold>\ Dom g) \\<^sub>B Map h)" proof - have "DN (\ f g h) = B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\\<^bold>\\<^bold>\ Dom h)" using f g h fg gh DN_def 1 4 6 7 B.can_Ide_self E.HcompNml_assoc E.Ide.simps(3) \DN (\ f g h) = \Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h\\ ide_char by (metis (no_types, lifting) arr_char ideD(1)) thus ?thesis using f g h fg gh 1 2 4 5 6 7 8 9 10 11 12 DN_def \_def \.map_simp_ide hcomp_def src_def trg_def Cod_ide by simp qed also have "... = (B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\\<^bold>\\<^bold>\ Dom h) \\<^sub>B B.can ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\\<^bold>\\<^bold>\ Dom h) ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\ Dom h)) \\<^sub>B (B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) (Dom f \<^bold>\ Dom g) \\<^sub>B Map h)" using B.comp_assoc by simp also have "... = B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\ Dom h) \\<^sub>B B.can ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\ Dom h) ((Dom f \<^bold>\ Dom g) \<^bold>\ Dom h)" proof - have "B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) (Dom f \<^bold>\ Dom g) \\<^sub>B Map h = B.can ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\ Dom h) ((Dom f \<^bold>\ Dom g) \<^bold>\ Dom h)" proof - have "B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) (Dom f \<^bold>\ Dom g) \\<^sub>B Map h = B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) (Dom f \<^bold>\ Dom g) \\<^sub>B B.can (Dom h) (Dom h)" using h B.can_Ide_self by fastforce also have "... = B.can ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\ Dom h) ((Dom f \<^bold>\ Dom g) \<^bold>\ Dom h)" using f g h 1 4 7 arr_char E.Nml_HcompNml(1) E.Src_HcompNml B.hcomp_can [of "Dom f \<^bold>\ Dom g" "Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g" "Dom h" "Dom h"] by (metis (no_types, lifting) E.Nmlize.simps(3) E.Nmlize_Nml E.Ide.simps(3) E.Ide_HcompNml E.Src.simps(3) arrE ideD(1)) finally show ?thesis by simp qed moreover have "B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\\<^bold>\\<^bold>\ Dom h) \\<^sub>B B.can ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\\<^bold>\\<^bold>\ Dom h) ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\ Dom h) = B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\ Dom h)" proof - have "E.Ide ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\ Dom h)" using f g h 1 4 7 by (metis (no_types, lifting) E.Ide.simps(3) E.Ide_HcompNml E.Src_HcompNml arrE ideD(1)) moreover have "E.Ide ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\\<^bold>\\<^bold>\ Dom h)" using f g h 1 4 7 E.Ide_HcompNml E.Nml_HcompNml(1) arr_char calculation ideD(1) by auto moreover have "E.Ide (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h)" using f g h 1 4 6 by blast moreover have "E.Nmlize ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\ Dom h) = E.Nmlize ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\\<^bold>\\<^bold>\ Dom h)" using f g h 1 4 7 by (metis (no_types, lifting) E.Nml_HcompNml(1) E.Nmlize.simps(3) E.Nmlize_Nml E.Ide.simps(3) arrE calculation(1) ideD(1)) moreover have "E.Nmlize ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\\<^bold>\\<^bold>\ Dom h) = E.Nmlize (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h)" using f g h 1 4 7 E.HcompNml_assoc by fastforce ultimately show ?thesis using B.vcomp_can by simp qed ultimately show ?thesis by simp qed also have "... = B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) ((Dom f \<^bold>\ Dom g) \<^bold>\ Dom h)" proof - have "E.Ide ((Dom f \<^bold>\ Dom g) \<^bold>\ Dom h)" using 1 4 7 by simp moreover have "E.Ide ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\ Dom h)" using f g 1 4 7 by (metis (no_types, lifting) E.Ide.simps(3) E.Ide_HcompNml E.Src_HcompNml arrE ideD(1)) moreover have "E.Ide (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h)" using f g h 1 8 6 7 by blast moreover have "E.Nmlize ((Dom f \<^bold>\ Dom g) \<^bold>\ Dom h) = E.Nmlize ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\ Dom h)" using f g h 1 4 7 E.Nml_HcompNml(1) by fastforce moreover have "E.Nmlize ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\ Dom h) = E.Nmlize (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h)" using f g h 1 4 7 by (metis (no_types, lifting) E.Nml_HcompNml(1) E.Nmlize.simps(3) E.Nmlize_Nml E.HcompNml_assoc E.Ide.simps(3) arrE ideD(1)) ultimately show ?thesis using B.vcomp_can by simp qed finally show ?thesis by simp qed also have "... = \ (f, g \ h) \\<^sub>B (DN f \\<^sub>B \ (g, h)) \\<^sub>B \\<^sub>B[DN f, DN g, DN h]" proof - have "\ (f, g \ h) \\<^sub>B (DN f \\<^sub>B \ (g, h)) \\<^sub>B \\<^sub>B[DN f, DN g, DN h] = (\ (f, g \ h) \\<^sub>B (DN f \\<^sub>B \ (g, h))) \\<^sub>B \\<^sub>B[DN f, DN g, DN h]" using B.comp_assoc by simp also have "... = B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom f \<^bold>\ Dom g \<^bold>\ Dom h) \\<^sub>B B.can (Dom f \<^bold>\ Dom g \<^bold>\ Dom h) ((Dom f \<^bold>\ Dom g) \<^bold>\ Dom h)" proof - have "\ (f, g \ h) \\<^sub>B (DN f \\<^sub>B \ (g, h)) = B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom f \<^bold>\ Dom g \<^bold>\ Dom h)" proof - have "\ (f, g \ h) \\<^sub>B (DN f \\<^sub>B \ (g, h)) = B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) \\<^sub>B (Map f \\<^sub>B B.can (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom g \<^bold>\ Dom h))" proof - have "VV.ide (g, h)" using g h gh VV.ide_char VV.arr_char by simp moreover have "VV.ide (f, hcomp g h)" using f g h fg gh VV.ide_char VV.arr_char by simp ultimately show ?thesis using f g h fg gh \.map_simp_ide DN_def hcomp_def src_def trg_def by simp qed also have "... = B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) \\<^sub>B (B.can (Dom f) (Dom f) \\<^sub>B B.can (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom g \<^bold>\ Dom h))" proof - have "Map f = B.can (Dom f) (Dom f)" using f arr_char B.can_Ide_self [of "Dom f"] Map_ide by (metis (no_types, lifting) ide_char') thus ?thesis by simp qed also have "... = B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) \\<^sub>B B.can (Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom f \<^bold>\ Dom g \<^bold>\ Dom h)" using 1 4 5 7 B.hcomp_can by auto also have "... = B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom f \<^bold>\ Dom g \<^bold>\ Dom h)" using 1 4 5 6 7 8 B.vcomp_can by auto finally show ?thesis by simp qed moreover have "\\<^sub>B[DN f, DN g, DN h] = B.can (Dom f \<^bold>\ Dom g \<^bold>\ Dom h) ((Dom f \<^bold>\ Dom g) \<^bold>\ Dom h)" using f g h 1 4 7 DN_def B.canE_associator(1) Map_ide by auto ultimately show ?thesis by simp qed also have "... = B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) ((Dom f \<^bold>\ Dom g) \<^bold>\ Dom h)" using 1 4 5 6 7 8 E.Nmlize_Hcomp_Hcomp B.vcomp_can [of "(Dom f \<^bold>\ Dom g) \<^bold>\ Dom h" "Dom f \<^bold>\ Dom g \<^bold>\ Dom h" "Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h"] by simp finally show ?thesis by simp qed finally show ?thesis by blast qed qed qed lemma DN_is_pseudofunctor: shows "pseudofunctor vcomp hcomp \ \ src trg V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B DN \" .. interpretation faithful_functor vcomp V\<^sub>B DN proof fix \ \' assume par: "par \ \'" and eq: "DN \ = DN \'" show "\ = \'" proof (intro arr_eqI) show "arr \" using par by simp show "arr \'" using par by simp show "Dom \ = Dom \'" using par arr_char dom_char by force show "Cod \ = Cod \'" using par arr_char cod_char by force show "Map \ = Map \'" using par eq DN_def by simp qed qed no_notation B.in_hom ("\_ : _ \\<^sub>B _\") lemma DN_UP: assumes "B.arr \" shows "DN (UP \) = \" using assms UP_def DN_def arr_UP by auto interpretation DN: equivalence_pseudofunctor vcomp hcomp \ \ src trg V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B DN \ proof (* DN is locally (but not globally) full. *) show "\f f' \. \ ide f; ide f'; src f = src f'; trg f = trg f'; \\ : DN f \\<^sub>B DN f'\ \ \ \\. \\ : f \ f'\ \ DN \ = \" proof - fix f f' \ assume f: "ide f" and f': "ide f'" and eq_src: "src f = src f'" and eq_trg: "trg f = trg f'" and \: "\\ : DN f \\<^sub>B DN f'\" show "\\. \\ : f \ f'\ \ DN \ = \" proof - let ?\ = "MkArr (Dom f) (Dom f') \" have \: "\?\ : f \ f'\" proof have "Map f = \Dom f\" using f by simp have "Map f' = \Dom f'\" using f' by simp have "Dom f' = Cod f'" using f' Cod_ide by simp show \: "arr ?\" proof - have "E.Nml (Dom ?\) \ E.Ide (Dom ?\)" proof - have "E.Nml (Dom f) \ E.Ide (Dom f)" using f ide_char arr_char by blast thus ?thesis using f by simp qed moreover have "E.Nml (Cod ?\) \ E.Ide (Cod ?\)" proof - have "E.Nml (Dom f') \ E.Ide (Dom f')" using f' ide_char arr_char by blast thus ?thesis using f' by simp qed moreover have "E.Src (Dom ?\) = E.Src (Cod ?\)" using f f' \ arr_char src_def eq_src ideD(1) by auto moreover have "E.Trg (Dom ?\) = E.Trg (Cod ?\)" using f f' \ arr_char trg_def eq_trg ideD(1) by auto moreover have "\Map ?\ : \Dom ?\\ \\<^sub>B \Cod ?\\\" proof - have "\\ : \Dom f\ \\<^sub>B \Dom f'\\" using f f' \ ide_char arr_char DN_def Cod_ide Map_ide by (metis (no_types, lifting) ideD(1)) thus ?thesis by simp qed ultimately show ?thesis using f f' \ ide_char arr_char by blast qed show "dom ?\ = f" using f \ dom_char MkArr_Map MkIde_Dom' by simp show "cod ?\ = f'" proof - have "cod ?\ = MkIde (Dom f')" using \ cod_char by simp also have "... = MkArr (Dom f') (Cod f') (Map f')" using f' by auto also have "... = f'" using f' MkArr_Map by simp finally show ?thesis by simp qed qed moreover have "DN ?\ = \" using \ DN_def by auto ultimately show ?thesis by blast qed qed (* DN is essentially surjective up to equivalence on objects. *) show "\a'. B.obj a' \ \a. obj a \ B.equivalent_objects (DN.map\<^sub>0 a) a'" proof - fix a' assume a': "B.obj a'" have "obj (UP.map\<^sub>0 a')" using a' UP.map\<^sub>0_simps(1) by simp moreover have "B.equivalent_objects (DN.map\<^sub>0 (UP.map\<^sub>0 a')) a'" proof - have "arr (MkArr \<^bold>\a'\<^bold>\ \<^bold>\a'\<^bold>\ a')" using a' UP_def [of a'] UP.preserves_reflects_arr [of a'] by auto moreover have "arr (MkArr \<^bold>\a'\<^bold>\\<^sub>0 \<^bold>\a'\<^bold>\\<^sub>0 a')" using a' arr_char B.obj_def by auto ultimately have "DN.map\<^sub>0 (UP.map\<^sub>0 a') = a'" using a' UP.map\<^sub>0_def UP_def DN.map\<^sub>0_def DN_def src_def UP.map\<^sub>0_simps(1) by auto thus ?thesis using a' B.equivalent_objects_reflexive by simp qed ultimately show "\a. obj a \ B.equivalent_objects (DN.map\<^sub>0 a) a'" by blast qed (* DN is locally essentially surjective. *) show "\a b g. \ obj a; obj b; \g : DN.map\<^sub>0 a \\<^sub>B DN.map\<^sub>0 b\; B.ide g \ \ \f. \f : a \ b\ \ ide f \ B.isomorphic (DN f) g" proof - fix a b g assume a: "obj a" and b: "obj b" and g: "\g : DN.map\<^sub>0 a \\<^sub>B DN.map\<^sub>0 b\" and ide_g: "B.ide g" have "ide (UP g)" using ide_g UP.preserves_ide by simp moreover have "B.isomorphic (DN (UP g)) g" using ide_g DN_UP B.isomorphic_reflexive by simp moreover have "\UP g : a \ b\" proof show "arr (UP g)" using g UP.preserves_reflects_arr by auto show "src (UP g) = a" proof - have "src (UP g) = MkArr \<^bold>\src\<^sub>B g\<^bold>\\<^sub>0 \<^bold>\src\<^sub>B g\<^bold>\\<^sub>0 (src\<^sub>B g)" using ide_g src_def UP_def UP.preserves_reflects_arr [of g] B.ideD(1) by simp also have "... = a" proof - have "src\<^sub>B g = src\<^sub>B (DN.map\<^sub>0 a)" using a g B.in_hhom_def by simp also have "... = Map a" - using a Map_preserves_objects DN.map\<^sub>0_def DN_def B.src_src by auto + using a Map_preserves_objects DN.map\<^sub>0_def DN_def B.src_src B.obj_simps + by auto finally have "src\<^sub>B g = Map a" by simp hence "MkArr \<^bold>\src\<^sub>B g\<^bold>\\<^sub>0 \<^bold>\src\<^sub>B g\<^bold>\\<^sub>0 (src\<^sub>B g) = MkArr \<^bold>\Map a\<^bold>\\<^sub>0 \<^bold>\Map a\<^bold>\\<^sub>0 (Map a)" by simp also have "... = a" using a obj_char apply (cases "Dom a", simp_all) by (metis (no_types, lifting) B.obj_def' a comp_ide_arr dom_char dom_eqI objE) finally show ?thesis by simp qed finally show ?thesis by simp qed show "trg (UP g) = b" proof - have "trg (UP g) = MkArr \<^bold>\trg\<^sub>B g\<^bold>\\<^sub>0 \<^bold>\trg\<^sub>B g\<^bold>\\<^sub>0 (trg\<^sub>B g)" using ide_g trg_def UP_def UP.preserves_reflects_arr [of g] B.ideD(1) by simp also have "... = b" proof - have "trg\<^sub>B g = trg\<^sub>B (DN.map\<^sub>0 b)" using b g B.in_hhom_def by simp also have "... = Map b" - using b Map_preserves_objects DN.map\<^sub>0_def DN_def B.src_src by auto + using b Map_preserves_objects DN.map\<^sub>0_def DN_def B.src_src B.obj_simps + by auto finally have "trg\<^sub>B g = Map b" by simp hence "MkArr \<^bold>\trg\<^sub>B g\<^bold>\\<^sub>0 \<^bold>\trg\<^sub>B g\<^bold>\\<^sub>0 (trg\<^sub>B g) = MkArr \<^bold>\Map b\<^bold>\\<^sub>0 \<^bold>\Map b\<^bold>\\<^sub>0 (Map b)" by simp also have "... = b" using b obj_char apply (cases "Dom b", simp_all) by (metis (no_types, lifting) B.obj_def' b comp_ide_arr dom_char dom_eqI objE) finally show ?thesis by simp qed finally show ?thesis by simp qed qed ultimately show "\f. \f : a \ b\ \ ide f \ B.isomorphic (DN f) g" by blast qed qed theorem DN_is_equivalence_pseudofunctor: shows "equivalence_pseudofunctor vcomp hcomp \ \ src trg V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B DN \" .. text \ The following gives an explicit formula for a component of the unit isomorphism of the pseudofunctor \UP\ from a bicategory to its strictification. It is not currently being used -- I originally proved it in order to establish something that I later proved in a more abstract setting -- but it might be useful at some point. \ interpretation L: bicategorical_language V\<^sub>B src\<^sub>B trg\<^sub>B .. interpretation E: evaluation_map V\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B \\\. if B.arr \ then \ else B.null\ using B.src.is_extensional B.trg.is_extensional by (unfold_locales, auto) notation E.eval ("\_\") interpretation UP: equivalence_pseudofunctor V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B vcomp hcomp \ \ src trg UP \ using UP_is_equivalence_pseudofunctor by auto lemma UP_\_char: assumes "B.obj a" shows "UP.\ a = MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ a" proof - have " MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ a = UP.\ a" proof (intro UP.\_eqI) show "B.obj a" using assms by simp have 0: "\a : a \\<^sub>B a\" using assms by auto have 1: "arr (MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ a)" apply (unfold arr_char, intro conjI) using assms by auto have 2: "arr (MkArr \<^bold>\a\<^bold>\ \<^bold>\a\<^bold>\ a)" apply (unfold arr_char, intro conjI) using assms by auto have 3: "arr (MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\\<^sub>0 a)" apply (unfold arr_char, intro conjI) using assms by auto show "\MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ a : UP.map\<^sub>0 a \ UP a\" proof show "arr (MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ a)" by fact show "dom (MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ a) = UP.map\<^sub>0 a" using assms 1 2 dom_char UP.map\<^sub>0_def UP_def src_def by auto show "cod (MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ a) = UP a" using assms 1 2 cod_char UP.map\<^sub>0_def UP_def src_def by auto qed show "iso (MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ a)" using assms 1 iso_char by auto show "MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ a \ \ (UP.map\<^sub>0 a) = (UP \\<^sub>B[a] \ \ (a, a)) \ (MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ a \ MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ a)" proof - have "MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ a \ \ (UP.map\<^sub>0 a) = MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ a" unfolding \_def UP.map\<^sub>0_def UP_def using assms 0 1 2 src_def by auto also have "... = (UP \\<^sub>B[a] \ \ (a, a)) \ (MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ a \ MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ a)" proof - have "(UP \\<^sub>B[a] \ \ (a, a)) \ (MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ a \ MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ a) = (MkArr \<^bold>\a \\<^sub>B a\<^bold>\ \<^bold>\a\<^bold>\ \\<^sub>B[a] \ MkArr (\<^bold>\a\<^bold>\ \<^bold>\ \<^bold>\a\<^bold>\) \<^bold>\a \\<^sub>B a\<^bold>\ (a \\<^sub>B a)) \ (MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ a \ MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ a)" using assms UP_def \_ide_simp by auto also have "... = (MkArr \<^bold>\a \\<^sub>B a\<^bold>\ \<^bold>\a\<^bold>\ \\<^sub>B[a] \ MkArr (\<^bold>\a\<^bold>\ \<^bold>\ \<^bold>\a\<^bold>\) \<^bold>\a \\<^sub>B a\<^bold>\ (a \\<^sub>B a)) \ MkArr \<^bold>\a\<^bold>\\<^sub>0 (\<^bold>\a\<^bold>\ \<^bold>\ \<^bold>\a\<^bold>\) (B.runit' a)" using assms 0 1 2 3 hcomp_def B.comp_cod_arr src_def trg_def B.can_Ide_self B.canE_unitor [of "\<^bold>\a\<^bold>\\<^sub>0"] B.comp_cod_arr by auto also have "... = MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ ((\\<^sub>B[a] \\<^sub>B (a \\<^sub>B a)) \\<^sub>B B.runit' a)" proof - have "MkArr \<^bold>\a \\<^sub>B a\<^bold>\ \<^bold>\a\<^bold>\ \\<^sub>B[a] \ MkArr (\<^bold>\a\<^bold>\ \<^bold>\ \<^bold>\a\<^bold>\) \<^bold>\a \\<^sub>B a\<^bold>\ (a \\<^sub>B a) = MkArr (\<^bold>\a\<^bold>\ \<^bold>\ \<^bold>\a\<^bold>\) \<^bold>\a\<^bold>\ (\\<^sub>B[a] \\<^sub>B (a \\<^sub>B a))" using assms - apply (intro comp_MkArr arr_MkArr) - apply auto - by fastforce + by (intro comp_MkArr arr_MkArr) auto moreover have "MkArr (\<^bold>\a\<^bold>\ \<^bold>\ \<^bold>\a\<^bold>\) \<^bold>\a\<^bold>\ (\\<^sub>B[a] \\<^sub>B (a \\<^sub>B a)) \ MkArr \<^bold>\a\<^bold>\\<^sub>0 (\<^bold>\a\<^bold>\ \<^bold>\ \<^bold>\a\<^bold>\) (B.runit' a) = MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ ((\\<^sub>B[a] \\<^sub>B (a \\<^sub>B a)) \\<^sub>B B.runit' a)" using assms 0 B.comp_arr_dom by (intro comp_MkArr arr_MkArr, auto) ultimately show ?thesis by argo qed also have "... = MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ a" using assms B.comp_arr_dom B.comp_arr_inv' B.iso_unit B.unitor_coincidence(2) by simp finally show ?thesis by argo qed finally show ?thesis by simp qed qed thus ?thesis by simp qed end subsection "Pseudofunctors into a Strict Bicategory" text \ In the special case of a pseudofunctor into a strict bicategory, we can obtain explicit formulas for the images of the units and associativities under the pseudofunctor, which only involve the structure maps of the pseudofunctor, since the units and associativities in the target bicategory are all identities. This is useful in applying strictification. \ locale pseudofunctor_into_strict_bicategory = pseudofunctor + D: strict_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D begin lemma image_of_unitor: assumes "C.ide g" shows "F \\<^sub>C[g] = (D.inv (\ (trg\<^sub>C g)) \\<^sub>D F g) \\<^sub>D D.inv (\ (trg\<^sub>C g, g))" and "F \\<^sub>C[g] = (F g \\<^sub>D D.inv (\ (src\<^sub>C g))) \\<^sub>D D.inv (\ (g, src\<^sub>C g))" and "F (C.lunit' g) = \ (trg\<^sub>C g, g) \\<^sub>D (\ (trg\<^sub>C g) \\<^sub>D F g)" and "F (C.runit' g) = \ (g, src\<^sub>C g) \\<^sub>D (F g \\<^sub>D \ (src\<^sub>C g))" proof - show A: "F \\<^sub>C[g] = (D.inv (\ (trg\<^sub>C g)) \\<^sub>D F g) \\<^sub>D D.inv (\ (trg\<^sub>C g, g))" proof - have 1: "\(D.inv (\ (trg\<^sub>C g)) \\<^sub>D F g) \\<^sub>D D.inv (\ (trg\<^sub>C g, g)) : F (trg\<^sub>C g \\<^sub>C g) \\<^sub>D F g\" proof show "\D.inv (\ (trg\<^sub>C g)) \\<^sub>D F g : F (trg\<^sub>C g) \\<^sub>D F g \\<^sub>D F g\" using assms \_char by (auto simp add: D.hcomp_obj_arr) show "\D.inv (\ (trg\<^sub>C g, g)) : F (trg\<^sub>C g \\<^sub>C g) \\<^sub>D F (trg\<^sub>C g) \\<^sub>D F g\" using assms \_in_hom(2) D.inv_is_inverse by simp qed have "(D.inv (\ (trg\<^sub>C g)) \\<^sub>D F g) \\<^sub>D D.inv (\ (trg\<^sub>C g, g)) = F g \\<^sub>D (D.inv (\ (trg\<^sub>C g)) \\<^sub>D F g) \\<^sub>D D.inv (\ (trg\<^sub>C g, g))" using 1 D.comp_cod_arr by auto also have "... = (F \\<^sub>C[g] \\<^sub>D \ (trg\<^sub>C g, g) \\<^sub>D (\ (trg\<^sub>C g) \\<^sub>D F g)) \\<^sub>D (D.inv (\ (trg\<^sub>C g)) \\<^sub>D F g) \\<^sub>D D.inv (\ (trg\<^sub>C g, g))" using assms lunit_coherence [of g] D.strict_lunit by simp also have "... = F \\<^sub>C[g] \\<^sub>D \ (trg\<^sub>C g, g) \\<^sub>D ((\ (trg\<^sub>C g) \\<^sub>D F g) \\<^sub>D (D.inv (\ (trg\<^sub>C g)) \\<^sub>D F g)) \\<^sub>D D.inv (\ (trg\<^sub>C g, g))" using D.comp_assoc by simp also have "... = F \\<^sub>C[g]" proof - have "(\ (trg\<^sub>C g) \\<^sub>D F g) \\<^sub>D (D.inv (\ (trg\<^sub>C g)) \\<^sub>D F g) = F (trg\<^sub>C g) \\<^sub>D F g" using assms \_char D.whisker_right by (metis C.ideD(1) C.obj_trg C.trg.preserves_reflects_arr D.comp_arr_inv' \_simps(5) preserves_arr preserves_ide) moreover have "\ (trg\<^sub>C g, g) \\<^sub>D D.inv (\ (trg\<^sub>C g, g)) = F (trg\<^sub>C g \\<^sub>C g)" using assms D.comp_arr_inv D.inv_is_inverse by simp ultimately show ?thesis using assms D.comp_arr_dom D.comp_cod_arr \_char by auto qed finally show ?thesis by simp qed show B: "F \\<^sub>C[g] = (F g \\<^sub>D D.inv (\ (src\<^sub>C g))) \\<^sub>D D.inv (\ (g, src\<^sub>C g))" proof - have 1: "\(F g \\<^sub>D D.inv (\ (src\<^sub>C g))) \\<^sub>D D.inv (\ (g, src\<^sub>C g)) : F (g \\<^sub>C src\<^sub>C g) \\<^sub>D F g\" proof show "\F g \\<^sub>D D.inv (\ (src\<^sub>C g)) : F g \\<^sub>D F (src\<^sub>C g) \\<^sub>D F g\" using assms \_char by (auto simp add: D.hcomp_arr_obj) show "\D.inv (\ (g, src\<^sub>C g)) : F (g \\<^sub>C src\<^sub>C g) \\<^sub>D F g \\<^sub>D F (src\<^sub>C g)\" using assms \_in_hom(2) by simp qed have "(F g \\<^sub>D D.inv (\ (src\<^sub>C g))) \\<^sub>D D.inv (\ (g, src\<^sub>C g)) = F g \\<^sub>D (F g \\<^sub>D D.inv (\ (src\<^sub>C g))) \\<^sub>D D.inv (\ (g, src\<^sub>C g))" using 1 D.comp_cod_arr by auto also have "... = (F \\<^sub>C[g] \\<^sub>D \ (g, src\<^sub>C g) \\<^sub>D (F g \\<^sub>D \ (src\<^sub>C g))) \\<^sub>D (F g \\<^sub>D D.inv (\ (src\<^sub>C g))) \\<^sub>D D.inv (\ (g, src\<^sub>C g))" using assms runit_coherence [of g] D.strict_runit by simp also have "... = F \\<^sub>C[g] \\<^sub>D (\ (g, src\<^sub>C g) \\<^sub>D ((F g \\<^sub>D \ (src\<^sub>C g)) \\<^sub>D (F g \\<^sub>D D.inv (\ (src\<^sub>C g))))) \\<^sub>D D.inv (\ (g, src\<^sub>C g))" using D.comp_assoc by simp also have "... = F \\<^sub>C[g]" proof - have "(F g \\<^sub>D \ (src\<^sub>C g)) \\<^sub>D (F g \\<^sub>D D.inv (\ (src\<^sub>C g))) = F g \\<^sub>D F (src\<^sub>C g)" using assms D.whisker_left \_char by (metis C.ideD(1) C.obj_src C.src.preserves_ide D.comp_arr_inv' D.ideD(1) \_simps(5) preserves_ide) moreover have "\ (g, src\<^sub>C g) \\<^sub>D D.inv (\ (g, src\<^sub>C g)) = F (g \\<^sub>C src\<^sub>C g)" using assms D.comp_arr_inv D.inv_is_inverse by simp ultimately show ?thesis using assms D.comp_arr_dom D.comp_cod_arr \_char \_in_hom(2) [of g "src\<^sub>C g"] by auto qed finally show ?thesis by simp qed show "F (C.lunit' g) = \ (trg\<^sub>C g, g) \\<^sub>D (\ (trg\<^sub>C g) \\<^sub>D F g)" proof - have "F (C.lunit' g) = D.inv (F \\<^sub>C[g])" using assms C.iso_lunit preserves_inv by simp also have "... = D.inv ((D.inv (\ (trg\<^sub>C g)) \\<^sub>D F g) \\<^sub>D D.inv (\ (trg\<^sub>C g, g)))" using A by simp also have "... = \ (trg\<^sub>C g, g) \\<^sub>D (\ (trg\<^sub>C g) \\<^sub>D F g)" proof - have "D.iso (D.inv (\ (trg\<^sub>C g, g))) \ D.inv (D.inv (\ (trg\<^sub>C g, g))) = \ (trg\<^sub>C g, g)" using assms D.iso_inv_iso by simp moreover have "D.iso (D.inv (\ (trg\<^sub>C g)) \\<^sub>D F g) \ D.inv (D.inv (\ (trg\<^sub>C g)) \\<^sub>D F g) = \ (trg\<^sub>C g) \\<^sub>D F g" using assms \_char D.iso_inv_iso by simp moreover have "D.seq (D.inv (\ (trg\<^sub>C g)) \\<^sub>D F g) (D.inv (\ (trg\<^sub>C g, g)))" using assms \_char by (metis A C.lunit_simps(1) preserves_arr) ultimately show ?thesis using D.inv_comp by simp qed finally show ?thesis by simp qed show "F (C.runit' g) = \ (g, src\<^sub>C g) \\<^sub>D (F g \\<^sub>D \ (src\<^sub>C g))" proof - have "F (C.runit' g) = D.inv (F \\<^sub>C[g])" using assms C.iso_runit preserves_inv by simp also have "... = D.inv ((F g \\<^sub>D D.inv (\ (src\<^sub>C g))) \\<^sub>D D.inv (\ (g, src\<^sub>C g)))" using B by simp also have "... = \ (g, src\<^sub>C g) \\<^sub>D (F g \\<^sub>D \ (src\<^sub>C g))" proof - have "D.iso (D.inv (\ (g, src\<^sub>C g))) \ D.inv (D.inv (\ (g, src\<^sub>C g))) = \ (g, src\<^sub>C g)" using assms D.iso_inv_iso by simp moreover have "D.iso (F g \\<^sub>D D.inv (\ (src\<^sub>C g))) \ D.inv (F g \\<^sub>D D.inv (\ (src\<^sub>C g))) = F g \\<^sub>D \ (src\<^sub>C g)" using assms \_char D.iso_inv_iso by simp moreover have "D.seq (F g \\<^sub>D D.inv (\ (src\<^sub>C g))) (D.inv (\ (g, src\<^sub>C g)))" using assms \_char by (metis B C.runit_simps(1) preserves_arr) ultimately show ?thesis using D.inv_comp by simp qed finally show ?thesis by simp qed qed lemma image_of_associator: assumes "C.ide f" and "C.ide g" and "C.ide h" and "src\<^sub>C f = trg\<^sub>C g" and "src\<^sub>C g = trg\<^sub>C h" shows "F \\<^sub>C[f, g, h] = \ (f, g \\<^sub>C h) \\<^sub>D (F f \\<^sub>D \ (g, h)) \\<^sub>D (D.inv (\ (f, g)) \\<^sub>D F h) \\<^sub>D D.inv (\ (f \\<^sub>C g, h))" and "F (C.\' f g h) = \ (f \\<^sub>C g, h) \\<^sub>D (\ (f, g) \\<^sub>D F h) \\<^sub>D (F f \\<^sub>D D.inv (\ (g, h))) \\<^sub>D D.inv (\ (f, g \\<^sub>C h))" proof - show 1: "F \\<^sub>C[f, g, h] = \ (f, g \\<^sub>C h) \\<^sub>D (F f \\<^sub>D \ (g, h)) \\<^sub>D (D.inv (\ (f, g)) \\<^sub>D F h) \\<^sub>D D.inv (\ (f \\<^sub>C g, h))" proof - have 2: "D.seq (\ (f, g \\<^sub>C h)) ((F f \\<^sub>D \ (g, h)) \\<^sub>D \\<^sub>D[F f, F g, F h])" proof (intro D.seqI) show "D.arr \\<^sub>D[F f, F g, F h]" using assms D.assoc_in_hom(2) [of "F f" "F g" "F h"] by auto show "D.hseq (F f) (\ (g, h))" using assms by fastforce show "D.dom (F f \\<^sub>D \ (g, h)) = D.cod \\<^sub>D[F f, F g, F h]" using assms \D.hseq (F f) (\ (g, h))\ \_simps(1) \_simps(4) by auto show "D.arr (\ (f, g \\<^sub>C h))" using assms by auto show "D.dom (\ (f, g \\<^sub>C h)) = D.cod ((F f \\<^sub>D \ (g, h)) \\<^sub>D \\<^sub>D[F f, F g, F h])" using assms \D.hseq (F f) (\ (g, h))\ \_simps(1) \_simps(4) by auto qed moreover have 3: "F \\<^sub>C[f, g, h] \\<^sub>D \ (f \\<^sub>C g, h) \\<^sub>D (\ (f, g) \\<^sub>D F h) = \ (f, g \\<^sub>C h) \\<^sub>D (F f \\<^sub>D \ (g, h)) \\<^sub>D \\<^sub>D[F f, F g, F h]" using assms assoc_coherence [of f g h] by blast moreover have "D.iso (\ (f \\<^sub>C g, h) \\<^sub>D (\ (f, g) \\<^sub>D F h))" proof - have "D.iso (\ (f \\<^sub>C g, h)) \ D.iso (\ (f, g)) \ D.iso (F h)" using assms by simp moreover have "D.seq (\ (f \\<^sub>C g, h)) (\ (f, g) \\<^sub>D F h)" proof (intro D.seqI) show "D.hseq (\ (f, g)) (F h)" using assms C.VV.arr_char by simp show "D.arr (\ (f \\<^sub>C g, h))" using assms C.VV.arr_char by simp show "D.dom (\ (f \\<^sub>C g, h)) = D.cod (\ (f, g) \\<^sub>D F h)" using assms 2 3 by (metis D.seqE) qed ultimately show ?thesis using assms D.isos_compose by simp qed ultimately have "F \\<^sub>C[f, g, h] = (\ (f, g \\<^sub>C h) \\<^sub>D ((F f \\<^sub>D \ (g, h)) \\<^sub>D \\<^sub>D[F f, F g, F h])) \\<^sub>D D.inv (\ (f \\<^sub>C g, h) \\<^sub>D (\ (f, g) \\<^sub>D F h))" using D.invert_side_of_triangle(2) [of "\ (f, g \\<^sub>C h) \\<^sub>D (F f \\<^sub>D \ (g, h)) \\<^sub>D \\<^sub>D[F f, F g, F h]" "F \\<^sub>C[f, g, h]" "\ (f \\<^sub>C g, h) \\<^sub>D (\ (f, g) \\<^sub>D F h)"] by presburger also have "... = \ (f, g \\<^sub>C h) \\<^sub>D (F f \\<^sub>D \ (g, h)) \\<^sub>D (D.inv (\ (f, g)) \\<^sub>D F h) \\<^sub>D D.inv (\ (f \\<^sub>C g, h))" proof - have "D.inv (\ (f \\<^sub>C g, h) \\<^sub>D (\ (f, g) \\<^sub>D F h)) = (D.inv (\ (f, g)) \\<^sub>D F h) \\<^sub>D D.inv (\ (f \\<^sub>C g, h))" proof - have "D.iso (\ (f \\<^sub>C g, h))" using assms D.iso_inv_iso by simp moreover have "D.iso (\ (f, g) \\<^sub>D F h) \ D.inv (\ (f, g) \\<^sub>D F h) = D.inv (\ (f, g)) \\<^sub>D F h" using assms D.iso_inv_iso by simp moreover have "D.seq (\ (f \\<^sub>C g, h)) (\ (f, g) \\<^sub>D F h)" using assms by fastforce ultimately show ?thesis using D.inv_comp by simp qed moreover have "(F f \\<^sub>D \ (g, h)) \\<^sub>D \\<^sub>D[F f, F g, F h] = (F f \\<^sub>D \ (g, h))" using assms D.comp_arr_dom D.assoc_in_hom [of "F f" "F g" "F h"] \_in_hom by (metis "2" "3" D.comp_arr_ide D.hseq_char D.seqE D.strict_assoc \_simps(2) \_simps(3) preserves_ide) ultimately show ?thesis using D.comp_assoc by simp qed finally show ?thesis by simp qed show "F (C.\' f g h) = \ (f \\<^sub>C g, h) \\<^sub>D (\ (f, g) \\<^sub>D F h) \\<^sub>D (F f \\<^sub>D D.inv (\ (g, h))) \\<^sub>D D.inv (\ (f, g \\<^sub>C h))" proof - have "F (C.\' f g h) = D.inv (F \\<^sub>C[f, g, h])" using assms preserves_inv by simp also have "... = D.inv (\ (f, g \\<^sub>C h) \\<^sub>D (F f \\<^sub>D \ (g, h)) \\<^sub>D (D.inv (\ (f, g)) \\<^sub>D F h) \\<^sub>D D.inv (\ (f \\<^sub>C g, h)))" using 1 by simp also have "... = \ (f \\<^sub>C g, h) \\<^sub>D (\ (f, g) \\<^sub>D F h) \\<^sub>D (F f \\<^sub>D D.inv (\ (g, h))) \\<^sub>D D.inv (\ (f, g \\<^sub>C h))" using assms C.VV.arr_char D.iso_inv_iso FF_def D.hcomp_assoc D.comp_assoc (* OK, this is pretty cool, but not as cool as if I didn't have to direct it. *) by (simp add: D.inv_comp D.isos_compose) finally show ?thesis by simp qed qed end subsection "Internal Equivalences in a Strict Bicategory" text \ In this section we prove a useful fact about internal equivalences in a strict bicategory: namely, that if the ``right'' triangle identity holds for such an equivalence then the ``left'' does, as well. Later we will dualize this property, and use strictification to extend it to arbitrary bicategories. \ locale equivalence_in_strict_bicategory = strict_bicategory + equivalence_in_bicategory begin lemma triangle_right_implies_left: shows "(g \ \) \ (\ \ g) = g \ (\ \ f) \ (f \ \) = f" proof - text \ The formal proof here was constructed following the string diagram sketch below, which appears in \cite{nlab-zigzag-diagram} (see it also in context in \cite{nlab-adjoint-equivalence}). The diagram is reproduced here by permission of its author, Mike Shulman, who says (private communication): ``Just don't give the impression that the proof itself is due to me, because it's not. I don't know who first gave that proof; it's probably folklore.'' \begin{figure}[h] \includegraphics[width=6.5in]{triangle_right_implies_left.png} \end{figure} \ assume 1: "(g \ \) \ (\ \ g) = g" have 2: "(inv \ \ g) \ (g \ inv \) = g" proof - have "(inv \ \ g) \ (g \ inv \) = inv ((g \ \) \ (\ \ g))" using antipar inv_comp hcomp_assoc by simp also have "... = g" using 1 by simp finally show ?thesis by blast qed have "(\ \ f) \ (f \ \) = (\ \ f) \ (f \ (inv \ \ g) \ (g \ inv \) \ f) \ (f \ \)" proof - have "(f \ (inv \ \ g) \ (g \ inv \) \ f) \ (f \ \) = f \ \" using 2 ide_left ide_right antipar whisker_left by (metis comp_cod_arr unit_simps(1) unit_simps(3)) thus ?thesis by simp qed also have "... = (\ \ f) \ (f \ (inv \ \ g) \ (g \ inv \) \ f) \ (f \ \) \ (inv \ \ \)" proof - have "inv \ \ \ = src f" by (simp add: comp_inv_arr') thus ?thesis by (metis antipar(1) antipar(2) arrI calculation comp_ide_arr hcomp_arr_obj ideD(1) ide_left ide_right obj_src seqE strict_assoc' triangle_in_hom(1) vconn_implies_hpar(1)) qed also have "... = (\ \ (f \ (inv \ \ g) \ (g \ inv \)) \ ((f \ inv \) \ (f \ \))) \ (f \ \)" using ide_left ide_right antipar unit_is_iso by (metis "2" arr_inv calculation comp_arr_dom comp_inv_arr' counit_simps(1) counit_simps(2) dom_inv hcomp_arr_obj ideD(1) unit_simps(1) unit_simps(2) unit_simps(5) obj_trg seqI whisker_left) also have "... = (\ \ (f \ (inv \ \ g) \ (g \ inv \)) \ ((f \ inv \) \ ((inv \ \ f) \ (\ \ f)) \ (f \ \))) \ (f \ \)" proof - have "(inv \ \ f) \ (\ \ f) = (f \ g) \ f" using whisker_right [of f "inv \" \] counit_in_hom by (simp add: antipar(1) comp_inv_arr') thus ?thesis using hcomp_assoc comp_arr_dom by (metis comp_cod_arr ide_left local.unit_simps(1) local.unit_simps(3) whisker_left) qed also have "... = (((\ \ (f \ (inv \ \ g) \ (g \ inv \))) \ (f \ g)) \ (((f \ inv \) \ (inv \ \ f)) \ (\ \ f) \ (f \ \))) \ (f \ \)" using ide_left ide_right antipar comp_assoc whisker_right comp_cod_arr by (metis "2" comp_arr_dom counit_simps(1-2)) also have "... = (((\ \ (f \ (inv \ \ g) \ (g \ inv \))) \ ((f \ inv \) \ (inv \ \ f))) \ ((f \ g) \ (\ \ f) \ (f \ \))) \ (f \ \)" proof - have 3: "seq (\ \ (f \ (inv \ \ g) \ (g \ inv \))) (f \ g)" using 2 antipar by auto moreover have 4: "seq ((f \ inv \) \ (inv \ \ f)) ((\ \ f) \ (f \ \))" using antipar hcomp_assoc by auto ultimately show ?thesis using interchange by simp qed also have "... = ((\ \ (f \ (inv \ \ g) \ (g \ inv \))) \ ((f \ inv \) \ (inv \ \ f))) \ ((f \ g) \ (\ \ f) \ (f \ \)) \ (f \ \)" using comp_assoc by presburger also have "... = ((\ \ (f \ inv \) \ (inv \ \ f)) \ ((f \ (inv \ \ g) \ (g \ inv \)) \ f)) \ (f \ (g \ \) \ (\ \ g) \ f) \ (f \ \)" proof - have "(\ \ (f \ (inv \ \ g) \ (g \ inv \))) \ ((f \ inv \) \ (inv \ \ f)) = (\ \ (f \ inv \) \ (inv \ \ f)) \ ((f \ (inv \ \ g) \ (g \ inv \)) \ f)" proof - have "seq \ (f \ (inv \ \ g) \ (g \ inv \))" using 2 antipar by simp moreover have "seq ((f \ inv \) \ (inv \ \ f)) f" using antipar hcomp_assoc hcomp_obj_arr by auto ultimately show ?thesis using comp_assoc comp_arr_dom hcomp_obj_arr interchange [of \ "f \ (inv \ \ g) \ (g \ inv \)" "(f \ inv \) \ (inv \ \ f)" f] by simp qed moreover have "((f \ g) \ (\ \ f) \ (f \ \)) \ (f \ \) = (f \ (g \ \) \ (\ \ g) \ f) \ (f \ \)" proof - have "((f \ g) \ (\ \ f) \ (f \ \)) \ (f \ \) = (f \ g \ \ \ f) \ (f \ (g \ f) \ \) \ (f \ \ \ src f)" using antipar comp_assoc hcomp_assoc whisker_left hcomp_arr_obj by simp also have "... = (f \ g \ \ \ f) \ (f \ ((g \ f) \ \) \ (\ \ src f))" using antipar comp_arr_dom whisker_left hcomp_arr_obj by simp also have "... = (f \ g \ \ \ f) \ (f \ \ \ \)" using antipar comp_arr_dom comp_cod_arr hcomp_arr_obj interchange [of "g \ f" \ \ "src f"] by simp also have "... = (f \ g \ \ \ f) \ (f \ \ \ g \ f) \ (f \ src f \ \)" using antipar comp_arr_dom comp_cod_arr whisker_left interchange [of \ "src f" "g \ f" \] by simp also have "... = ((f \ g \ \ \ f) \ (f \ \ \ g \ f)) \ (f \ \)" using antipar comp_assoc by (simp add: hcomp_obj_arr) also have "... = (f \ (g \ \) \ (\ \ g) \ f) \ (f \ \)" using antipar comp_assoc whisker_left whisker_right hcomp_assoc by simp finally show ?thesis by blast qed ultimately show ?thesis by simp qed also have "... = (\ \ (f \ inv \) \ (inv \ \ f)) \ ((f \ (inv \ \ g) \ (g \ inv \) \ f) \ (f \ (g \ \) \ (\ \ g) \ f)) \ (f \ \)" using comp_assoc hcomp_assoc antipar(1) antipar(2) by auto also have "... = (\ \ (f \ inv \) \ (inv \ \ f)) \ ((f \ (inv \ \ g) \ (g \ inv \) \ (g \ \) \ (\ \ g) \ f)) \ (f \ \)" using ide_left ide_right antipar comp_cod_arr comp_assoc whisker_left by (metis "1" "2" comp_ide_self unit_simps(1) unit_simps(3)) also have "... = (\ \ (f \ inv \) \ (inv \ \ f)) \ (f \ \)" proof - have "(inv \ \ g) \ (g \ inv \) \ (g \ \) \ (\ \ g) = g" using ide_left ide_right antipar comp_arr_dom comp_assoc by (metis "1" "2" comp_ide_self) thus ?thesis using antipar comp_cod_arr by simp qed also have "... = (f \ inv \) \ ((inv \ \ f) \ (\ \ f)) \ (f \ \)" proof - have "(\ \ (f \ inv \) \ (inv \ \ f)) \ (f \ \) = (trg f \ \ \ (f \ inv \) \ (inv \ \ f)) \ (f \ \)" using hcomp_obj_arr comp_cod_arr by simp also have "... = ((trg f \ f \ inv \) \ (\ \ inv \ \ f)) \ (f \ \)" using antipar hcomp_arr_obj hcomp_assoc interchange by auto also have "... = (f \ inv \) \ ((inv \ \ f) \ (\ \ f)) \ (f \ \)" proof - have "(inv \ \ f) \ (\ \ f) = (trg f \ inv \ \ f) \ (\ \ trg f \ f)" using hseqI' by (simp add: hcomp_obj_arr) also have "... = \ \ inv \ \ f" using antipar comp_arr_dom comp_cod_arr interchange [of "trg f" \ "inv \ \ f" "trg f \ f"] by force finally have "(inv \ \ f) \ (\ \ f) = \ \ inv \ \ f" by simp moreover have "trg f \ f \ inv \ = f \ inv \" using hcomp_obj_arr [of "trg f" "f \ inv \"] by fastforce ultimately have "((trg f \ f \ inv \) \ (\ \ inv \ \ f)) \ (f \ \) = ((f \ inv \) \ ((inv \ \ f) \ (\ \ f))) \ (f \ \)" by simp thus ?thesis using comp_assoc by simp qed finally show ?thesis by simp qed also have "... = f \ inv \ \ \" proof - have "(inv \ \ f) \ (\ \ f) = f \ g \ f" using ide_left ide_right antipar counit_is_iso whisker_right hcomp_assoc by (metis comp_arr_dom comp_inv_arr' counit_simps(1) counit_simps(2) seqE) thus ?thesis using ide_left ide_right antipar unit_is_iso comp_cod_arr by (metis arr_inv dom_inv unit_simps(1) unit_simps(3) seqI whisker_left) qed also have "... = f \ src f" using antipar by (simp add: comp_inv_arr') also have "... = f" using hcomp_arr_obj by simp finally show ?thesis by simp qed end text \ Now we use strictification to generalize the preceding result to arbitrary bicategories. I originally attempted to generalize this proof directly from the strict case, by filling in the necessary canonical isomorphisms, but it turned out to be too daunting. The proof using strictification is still fairly tedious, but it is manageable. \ context equivalence_in_bicategory begin interpretation S: strictified_bicategory V H \ \ src trg .. notation S.vcomp (infixr "\\<^sub>S" 55) notation S.hcomp (infixr "\\<^sub>S" 53) notation S.in_hom ("\_ : _ \\<^sub>S _\") notation S.in_hhom ("\_ : _ \\<^sub>S _\") interpretation UP: equivalence_pseudofunctor V H \ \ src trg S.vcomp S.hcomp S.\ S.\ S.src S.trg S.UP S.\ using S.UP_is_equivalence_pseudofunctor by auto interpretation UP: pseudofunctor_into_strict_bicategory V H \ \ src trg S.vcomp S.hcomp S.\ S.\ S.src S.trg S.UP S.\ .. interpretation E: equivalence_in_bicategory S.vcomp S.hcomp S.\ S.\ S.src S.trg \S.UP f\ \S.UP g\ \S.inv (S.\ (g, f)) \\<^sub>S S.UP \ \\<^sub>S UP.\ (src f)\ \S.inv (UP.\ (trg f)) \\<^sub>S S.UP \ \\<^sub>S S.\ (f, g)\ using equivalence_in_bicategory_axioms UP.preserves_equivalence [of f g \ \] by auto interpretation E: equivalence_in_strict_bicategory S.vcomp S.hcomp S.\ S.\ S.src S.trg \S.UP f\ \S.UP g\ \S.inv (S.\ (g, f)) \\<^sub>S S.UP \ \\<^sub>S UP.\ (src f)\ \S.inv (UP.\ (trg f)) \\<^sub>S S.UP \ \\<^sub>S S.\ (f, g)\ .. lemma UP_triangle: shows "S.UP ((g \ \) \ \[g, f, g] \ (\ \ g)) = S.\ (g, src g) \\<^sub>S (S.UP g \\<^sub>S S.UP \ \\<^sub>S S.\ (f, g)) \\<^sub>S (S.inv (S.\ (g, f)) \\<^sub>S S.UP \ \\<^sub>S S.UP g) \\<^sub>S S.inv (S.\ (trg g, g))" and "S.UP (\\<^sup>-\<^sup>1[g] \ \[g]) = (S.\ (g, src g) \\<^sub>S (S.UP g \\<^sub>S UP.\ (src g))) \\<^sub>S (S.inv (UP.\ (trg g)) \\<^sub>S S.UP g) \\<^sub>S S.inv (S.\ (trg g, g))" and "S.UP ((\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \)) = S.\ (trg f, f) \\<^sub>S (S.UP \ \\<^sub>S S.\ (f, g) \\<^sub>S S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.\ (g, f)) \\<^sub>S S.UP \) \\<^sub>S S.inv (S.\ (f, src f))" and "S.UP (\\<^sup>-\<^sup>1[f] \ \[f]) = (S.\ (trg f, f) \\<^sub>S (UP.\ (trg f) \\<^sub>S S.UP f)) \\<^sub>S (S.UP f \\<^sub>S S.inv (UP.\ (src f))) \\<^sub>S S.inv (S.\ (f, src f))" and "(g \ \) \ \[g, f, g] \ (\ \ g) = \\<^sup>-\<^sup>1[g] \ \[g] \ S.\ (trg f, f) \\<^sub>S (S.UP \ \\<^sub>S S.\ (f, g) \\<^sub>S S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.\ (g, f)) \\<^sub>S S.UP \) \\<^sub>S S.inv (S.\ (f, src f)) = (S.\ (trg f, f) \\<^sub>S (UP.\ (trg f) \\<^sub>S S.UP f)) \\<^sub>S (S.UP f \\<^sub>S S.inv (UP.\ (src f))) \\<^sub>S S.inv (S.\ (f, src f))" proof - show T1: "S.UP ((g \ \) \ \[g, f, g] \ (\ \ g)) = S.\ (g, src g) \\<^sub>S (S.UP g \\<^sub>S S.UP \ \\<^sub>S S.\ (f, g)) \\<^sub>S (S.inv (S.\ (g, f)) \\<^sub>S S.UP \ \\<^sub>S S.UP g) \\<^sub>S S.inv (S.\ (trg g, g))" proof - have "S.UP ((g \ \) \ \[g, f, g] \ (\ \ g)) = S.UP (g \ \) \\<^sub>S S.UP \[g, f, g] \\<^sub>S S.UP (\ \ g)" using antipar by simp also have "... = (S.\ (g, src g) \\<^sub>S (S.UP g \\<^sub>S S.UP \) \\<^sub>S ((S.inv (S.\ (g, f \ g)) \\<^sub>S S.\ (g, f \ g)) \\<^sub>S (S.UP g \\<^sub>S S.\ (f, g))) \\<^sub>S (((S.inv (S.\ (g, f)) \\<^sub>S S.UP g) \\<^sub>S (S.inv (S.\ (g \ f, g)))) \\<^sub>S S.\ (g \ f, g)) \\<^sub>S (S.UP \ \\<^sub>S S.UP g)) \\<^sub>S S.inv (S.\ (trg g, g))" proof - have "S.UP \[g, f, g] = S.\ (g, f \ g) \\<^sub>S (S.UP g \\<^sub>S S.\ (f, g)) \\<^sub>S (S.inv (S.\ (g, f)) \\<^sub>S S.UP g) \\<^sub>S S.inv (S.\ (g \ f, g))" using ide_left ide_right antipar UP.image_of_associator [of g f g] by simp moreover have "S.UP (g \ \) = S.\ (g, src g) \\<^sub>S (S.UP g \\<^sub>S S.UP \) \\<^sub>S S.inv (S.\ (g, f \ g))" proof - have "S.seq (S.\ (g, src g)) (S.UP g \\<^sub>S S.UP \)" using antipar UP.FF_def UP.\_in_hom [of g "src g"] apply (intro S.seqI) by auto moreover have "S.UP (g \ \) \\<^sub>S S.\ (g, f \ g) = S.\ (g, src g) \\<^sub>S (S.UP g \\<^sub>S S.UP \)" using antipar UP.\.naturality [of "(g, \)"] UP.FF_def VV.arr_char by simp moreover have "S.iso (S.\ (g, f \ g))" using antipar by simp ultimately show ?thesis using antipar S.comp_assoc S.invert_side_of_triangle(2) [of "S.\ (g, src g) \\<^sub>S (S.UP g \\<^sub>S S.UP \)" "S.UP (g \ \)" "S.\ (g, f \ g)"] by simp qed moreover have "S.UP (\ \ g) = (S.\ (g \ f, g) \\<^sub>S (S.UP \ \\<^sub>S S.UP g)) \\<^sub>S S.inv (S.\ (trg g, g))" proof - have "S.UP (\ \ g) \\<^sub>S S.\ (trg g, g) = S.\ (g \ f, g) \\<^sub>S (S.UP \ \\<^sub>S S.UP g)" using antipar UP.\.naturality [of "(\, g)"] UP.FF_def VV.arr_char by simp moreover have "S.seq (S.\ (g \ f, g)) (S.UP \ \\<^sub>S S.UP g)" using antipar UP.\_in_hom(2) by (intro S.seqI, auto) ultimately show ?thesis using antipar S.invert_side_of_triangle(2) by simp qed ultimately show ?thesis using S.comp_assoc by simp qed also have "... = S.\ (g, src g) \\<^sub>S ((S.UP g \\<^sub>S S.UP \) \\<^sub>S (S.UP g \\<^sub>S S.\ (f, g))) \\<^sub>S ((S.inv (S.\ (g, f)) \\<^sub>S S.UP g) \\<^sub>S (S.UP \ \\<^sub>S S.UP g)) \\<^sub>S S.inv (S.\ (trg g, g))" proof - have "(S.inv (S.\ (g \ f, g)) \\<^sub>S S.\ (g \ f, g)) \\<^sub>S (S.UP \ \\<^sub>S S.UP g) = (S.UP \ \\<^sub>S S.UP g)" using antipar S.comp_inv_arr' S.comp_cod_arr by auto moreover have "(S.inv (S.\ (g, f \ g)) \\<^sub>S S.\ (g, f \ g)) \\<^sub>S (S.UP g \\<^sub>S S.\ (f, g)) = (S.UP g \\<^sub>S S.\ (f, g))" proof - have "S.inv (S.\ (g, f \ g)) \\<^sub>S S.\ (g, f \ g) = S.UP g \\<^sub>S S.UP (f \ g)" using antipar S.comp_inv_arr' UP.\_in_hom by auto thus ?thesis using antipar VV.arr_char S.comp_cod_arr by simp qed ultimately show ?thesis using S.comp_assoc by simp qed also have "... = S.\ (g, src g) \\<^sub>S (S.UP g \\<^sub>S S.UP \ \\<^sub>S S.\ (f, g)) \\<^sub>S (S.inv (S.\ (g, f)) \\<^sub>S S.UP \ \\<^sub>S S.UP g) \\<^sub>S S.inv (S.\ (trg g, g))" using antipar VV.arr_char S.whisker_left S.whisker_right by auto finally show ?thesis by simp qed show T2: "S.UP (\\<^sup>-\<^sup>1[g] \ \[g]) = (S.\ (g, src g) \\<^sub>S (S.UP g \\<^sub>S UP.\ (src g))) \\<^sub>S (S.inv (UP.\ (trg g)) \\<^sub>S S.UP g) \\<^sub>S S.inv (S.\ (trg g, g))" using UP.image_of_unitor by simp show "S.UP ((\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \)) = S.\ (trg f, f) \\<^sub>S (S.UP \ \\<^sub>S S.\ (f, g) \\<^sub>S S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.\ (g, f)) \\<^sub>S S.UP \) \\<^sub>S S.inv (S.\ (f, src f))" proof - have "S.UP ((\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \)) = S.UP (\ \ f) \\<^sub>S S.UP \\<^sup>-\<^sup>1[f, g, f] \\<^sub>S S.UP (f \ \)" using antipar by simp also have "... = S.\ (trg f, f) \\<^sub>S (S.UP \ \\<^sub>S S.UP f) \\<^sub>S (S.inv (S.\ (f \ g, f)) \\<^sub>S S.\ (f \ g, f) \\<^sub>S (S.\ (f, g) \\<^sub>S S.UP f)) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.\ (g, f))) \\<^sub>S (S.inv (S.\ (f, g \ f)) \\<^sub>S S.\ (f, g \ f) \\<^sub>S (S.UP f \\<^sub>S S.UP \)) \\<^sub>S S.inv (S.\ (f, src f))" proof - have "S.UP \\<^sup>-\<^sup>1[f, g, f] = S.\ (f \ g, f) \\<^sub>S (S.\ (f, g) \\<^sub>S S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.\ (g, f))) \\<^sub>S S.inv (S.\ (f, g \ f))" using ide_left ide_right antipar UP.image_of_associator by simp moreover have "S.UP (\ \ f) = S.\ (trg f, f) \\<^sub>S (S.UP \ \\<^sub>S S.UP f) \\<^sub>S S.inv (S.\ (f \ g, f))" proof - have "S.seq (S.\ (trg f, f)) (S.UP \ \\<^sub>S S.UP f)" using antipar UP.FF_def VV.ide_char VV.arr_char UP.\_in_hom [of "trg f" f] apply (intro S.seqI) by auto moreover have "S.\ (trg f, f) \\<^sub>S (S.UP \ \\<^sub>S S.UP f) = S.UP (\ \ f) \\<^sub>S S.\ (f \ g, f)" using antipar UP.\.naturality [of "(\, f)"] UP.FF_def VV.arr_char by simp moreover have "S.iso (S.\ (f \ g, f))" using antipar by simp ultimately show ?thesis using antipar S.comp_assoc S.invert_side_of_triangle(2) [of "S.\ (trg f, f) \\<^sub>S (S.UP \ \\<^sub>S S.UP f)" "S.UP (\ \ f)" "S.\ (f \ g, f)"] by metis qed moreover have "S.UP (f \ \) = (S.\ (f, g \ f) \\<^sub>S (S.UP f \\<^sub>S S.UP \)) \\<^sub>S S.inv (S.\ (f, src f))" proof - have "S.\ (f, g \ f) \\<^sub>S (S.UP f \\<^sub>S S.UP \) = S.UP (f \ \) \\<^sub>S S.\ (f, src f)" using antipar UP.\.naturality [of "(f, \)"] UP.FF_def VV.arr_char by simp moreover have "S.seq (S.\ (f, g \ f)) (S.UP f \\<^sub>S S.UP \)" using antipar by (intro S.seqI, auto) ultimately show ?thesis using antipar S.invert_side_of_triangle(2) by auto qed ultimately show ?thesis using S.comp_assoc by simp qed also have "... = S.\ (trg f, f) \\<^sub>S ((S.UP \ \\<^sub>S S.UP f) \\<^sub>S (S.\ (f, g) \\<^sub>S S.UP f)) \\<^sub>S ((S.UP f \\<^sub>S S.inv (S.\ (g, f))) \\<^sub>S (S.UP f \\<^sub>S S.UP \)) \\<^sub>S S.inv (S.\ (f, src f))" proof - have "(S.inv (S.\ (f \ g, f)) \\<^sub>S S.\ (f \ g, f)) \\<^sub>S (S.\ (f, g) \\<^sub>S S.UP f) = (S.\ (f, g) \\<^sub>S S.UP f)" using antipar S.comp_cod_arr VV.arr_char S.comp_inv_arr' by auto moreover have "(S.inv (S.\ (f, g \ f)) \\<^sub>S S.\ (f, g \ f)) \\<^sub>S (S.UP f \\<^sub>S S.UP \) = (S.UP f \\<^sub>S S.UP \)" using antipar S.comp_inv_arr' S.comp_cod_arr by auto ultimately show ?thesis using S.comp_assoc by simp qed also have "... = S.\ (trg f, f) \\<^sub>S (S.UP \ \\<^sub>S S.\ (f, g) \\<^sub>S S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.\ (g, f)) \\<^sub>S S.UP \) \\<^sub>S S.inv (S.\ (f, src f))" using antipar VV.arr_char S.whisker_left S.whisker_right by auto finally show ?thesis by simp qed show "S.UP (\\<^sup>-\<^sup>1[f] \ \[f]) = (S.\ (trg f, f) \\<^sub>S (UP.\ (trg f) \\<^sub>S S.UP f)) \\<^sub>S (S.UP f \\<^sub>S S.inv (UP.\ (src f))) \\<^sub>S S.inv (S.\ (f, src f))" using UP.image_of_unitor by simp show "(g \ \) \ \[g, f, g] \ (\ \ g) = \\<^sup>-\<^sup>1[g] \ \[g] \ S.\ (trg f, f) \\<^sub>S (S.UP \ \\<^sub>S S.\ (f, g) \\<^sub>S S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.\ (g, f)) \\<^sub>S S.UP \) \\<^sub>S S.inv (S.\ (f, src f)) = (S.\ (trg f, f) \\<^sub>S (UP.\ (trg f) \\<^sub>S S.UP f)) \\<^sub>S (S.UP f \\<^sub>S S.inv (UP.\ (src f))) \\<^sub>S S.inv (S.\ (f, src f))" proof - assume A: "(g \ \) \ \[g, f, g] \ (\ \ g) = \\<^sup>-\<^sup>1[g] \ \[g]" have B: "(S.UP g \\<^sub>S S.inv (UP.\ (src g)) \\<^sub>S S.UP \ \\<^sub>S S.\ (f, g)) \\<^sub>S (S.inv (S.\ (g, f)) \\<^sub>S S.UP \ \\<^sub>S UP.\ (trg g) \\<^sub>S S.UP g) = S.UP g" proof - show "(S.UP g \\<^sub>S S.inv (UP.\ (src g)) \\<^sub>S S.UP \ \\<^sub>S S.\ (f, g)) \\<^sub>S (S.inv (S.\ (g, f)) \\<^sub>S S.UP \ \\<^sub>S UP.\ (trg g) \\<^sub>S S.UP g) = S.UP g" proof - have 2: "S.\ (g, src g) \\<^sub>S (S.UP g \\<^sub>S S.UP \ \\<^sub>S S.\ (f, g)) \\<^sub>S (S.inv (S.\ (g, f)) \\<^sub>S S.UP \ \\<^sub>S S.UP g) \\<^sub>S S.inv (S.\ (trg g, g)) = (S.\ (g, src g) \\<^sub>S (S.UP g \\<^sub>S UP.\ (src g))) \\<^sub>S (S.inv (UP.\ (trg g)) \\<^sub>S S.UP g) \\<^sub>S S.inv (S.\ (trg g, g))" using A T1 T2 by simp show ?thesis proof - have 8: "S.seq (S.\ (g, src g)) ((S.UP g \\<^sub>S S.UP \ \\<^sub>S S.\ (f, g)) \\<^sub>S (S.inv (S.\ (g, f)) \\<^sub>S S.UP \ \\<^sub>S S.UP g) \\<^sub>S S.inv (S.\ (trg g, g)))" using antipar VV.arr_char S.hcomp_assoc apply (intro S.seqI) by auto have 7: "S.seq (S.\ (g, src g)) ((S.UP g \\<^sub>S UP.\ (src g)) \\<^sub>S (S.inv (UP.\ (trg g)) \\<^sub>S S.UP g) \\<^sub>S S.inv (S.\ (trg g, g)))" using antipar 2 8 S.comp_assoc by auto have 5: "(S.UP g \\<^sub>S S.UP \ \\<^sub>S S.\ (f, g)) \\<^sub>S (S.inv (S.\ (g, f)) \\<^sub>S S.UP \ \\<^sub>S S.UP g) = (S.UP g \\<^sub>S UP.\ (src g)) \\<^sub>S (S.inv (UP.\ (trg g)) \\<^sub>S S.UP g)" proof - have "((S.UP g \\<^sub>S S.UP \ \\<^sub>S S.\ (f, g)) \\<^sub>S (S.inv (S.\ (g, f)) \\<^sub>S S.UP \ \\<^sub>S S.UP g)) \\<^sub>S S.inv (S.\ (trg g, g)) = ((S.UP g \\<^sub>S UP.\ (src g)) \\<^sub>S (S.inv (UP.\ (trg g)) \\<^sub>S S.UP g)) \\<^sub>S S.inv (S.\ (trg g, g))" proof - have "S.mono (S.\ (g, src g))" using antipar S.iso_is_section S.section_is_mono by simp thus ?thesis using 2 8 7 S.monoE S.comp_assoc by presburger qed moreover have "S.epi (S.inv (S.\ (trg g, g)))" using antipar S.iso_is_retraction S.retraction_is_epi S.iso_inv_iso by simp moreover have "S.seq ((S.UP g \\<^sub>S S.UP \ \\<^sub>S S.\ (f, g)) \\<^sub>S (S.inv (S.\ (g, f)) \\<^sub>S S.UP \ \\<^sub>S S.UP g)) (S.inv (S.\ (trg g, g)))" using S.comp_assoc S.seq_char 8 by presburger moreover have "S.seq ((S.UP g \\<^sub>S UP.\ (src g)) \\<^sub>S (S.inv (UP.\ (trg g)) \\<^sub>S S.UP g)) (S.inv (S.\ (trg g, g)))" using antipar calculation(1) calculation(3) by presburger ultimately show ?thesis using 2 S.epiE by blast qed have 6: "S.seq (S.UP g \\<^sub>S S.UP \ \\<^sub>S S.\ (f, g)) (S.inv (S.\ (g, f)) \\<^sub>S S.UP \ \\<^sub>S S.UP g)" using antipar VV.arr_char S.hcomp_assoc by auto have 3: "(S.UP g \\<^sub>S S.inv (UP.\ (src g))) \\<^sub>S (S.UP g \\<^sub>S S.UP \ \\<^sub>S S.\ (f, g)) \\<^sub>S (S.inv (S.\ (g, f)) \\<^sub>S S.UP \ \\<^sub>S S.UP g) = (S.inv (UP.\ (trg g)) \\<^sub>S S.UP g)" proof - have "S.seq (S.UP g \\<^sub>S S.UP \ \\<^sub>S S.\ (f, g)) (S.inv (S.\ (g, f)) \\<^sub>S S.UP \ \\<^sub>S S.UP g)" using 6 by simp moreover have "(S.UP g \\<^sub>S UP.\ (src g)) \\<^sub>S (S.inv (UP.\ (trg g)) \\<^sub>S S.UP g) = (S.UP g \\<^sub>S S.UP \ \\<^sub>S S.\ (f, g)) \\<^sub>S (S.inv (S.\ (g, f)) \\<^sub>S S.UP \ \\<^sub>S S.UP g)" using 5 by argo moreover have "S.iso (S.UP g \\<^sub>S UP.\ (src g))" using antipar UP.\_char S.UP_map\<^sub>0_obj by simp moreover have "S.inv (S.UP g \\<^sub>S UP.\ (src g)) = S.UP g \\<^sub>S S.inv (UP.\ (src g))" using antipar UP.\_char S.UP_map\<^sub>0_obj by simp ultimately show ?thesis using S.invert_side_of_triangle(1) [of "(S.UP g \\<^sub>S S.UP \ \\<^sub>S S.\ (f, g)) \\<^sub>S (S.inv (S.\ (g, f)) \\<^sub>S S.UP \ \\<^sub>S S.UP g)" "S.UP g \\<^sub>S UP.\ (src g)" "S.inv (UP.\ (trg g)) \\<^sub>S S.UP g"] by presburger qed have 4: "(((S.UP g \\<^sub>S S.inv (UP.\ (src g))) \\<^sub>S (S.UP g \\<^sub>S S.UP \ \\<^sub>S S.\ (f, g))) \\<^sub>S ((S.inv (S.\ (g, f)) \\<^sub>S S.UP \ \\<^sub>S S.UP g)) \\<^sub>S (UP.\ (trg g) \\<^sub>S S.UP g)) = S.UP g" proof - have "(((S.UP g \\<^sub>S S.inv (UP.\ (src g))) \\<^sub>S (S.UP g \\<^sub>S S.UP \ \\<^sub>S S.\ (f, g))) \\<^sub>S ((S.inv (S.\ (g, f)) \\<^sub>S S.UP \ \\<^sub>S S.UP g)) \\<^sub>S (UP.\ (trg g) \\<^sub>S S.UP g)) = (((S.UP g \\<^sub>S S.inv (UP.\ (src g))) \\<^sub>S (S.UP g \\<^sub>S S.UP \ \\<^sub>S S.\ (f, g))) \\<^sub>S ((S.inv (S.\ (g, f)) \\<^sub>S S.UP \ \\<^sub>S S.UP g))) \\<^sub>S (UP.\ (trg g) \\<^sub>S S.UP g)" using S.comp_assoc by simp also have "... = (S.inv (UP.\ (trg g)) \\<^sub>S S.UP g) \\<^sub>S (UP.\ (trg g) \\<^sub>S S.UP g)" using 3 S.comp_assoc by auto also have "... = S.inv (UP.\ (trg g)) \\<^sub>S UP.\ (trg g) \\<^sub>S S.UP g" using UP.\_char(2) S.whisker_right by auto also have "... = UP.map\<^sub>0 (trg g) \\<^sub>S S.UP g" using UP.\_char [of "trg g"] S.comp_inv_arr S.inv_is_inverse by simp also have "... = S.UP g" using S.UP_map\<^sub>0_obj by (simp add: S.hcomp_obj_arr) finally show ?thesis by blast qed thus ?thesis using antipar S.whisker_left S.whisker_right UP.\_char S.comp_assoc by simp qed qed qed show "S.\ (trg f, f) \\<^sub>S (S.UP \ \\<^sub>S S.\ (f, g) \\<^sub>S S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.\ (g, f)) \\<^sub>S S.UP \) \\<^sub>S S.inv (S.\ (f, src f)) = (S.\ (trg f, f) \\<^sub>S (UP.\ (trg f) \\<^sub>S S.UP f)) \\<^sub>S (S.UP f \\<^sub>S S.inv (UP.\ (src f))) \\<^sub>S S.inv (S.\ (f, src f))" proof - have "(S.UP \ \\<^sub>S S.\ (f, g) \\<^sub>S S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.\ (g, f)) \\<^sub>S S.UP \) = (UP.\ (trg f) \\<^sub>S S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (UP.\ (src f)))" proof - have 2: "(S.inv (UP.\ (trg f)) \\<^sub>S S.UP f) \\<^sub>S ((S.UP \ \\<^sub>S S.\ (f, g) \\<^sub>S S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.\ (g, f)) \\<^sub>S S.UP \)) \\<^sub>S (S.UP f \\<^sub>S UP.\ (src f)) = S.UP f" proof - have "S.UP f = (S.inv (UP.\ (trg f)) \\<^sub>S S.UP \ \\<^sub>S S.\ (f, g) \\<^sub>S S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.\ (g, f)) \\<^sub>S S.UP \ \\<^sub>S UP.\ (src f))" using B antipar E.triangle_right_implies_left by argo also have "... = (S.inv (UP.\ (trg f)) \\<^sub>S S.UP f) \\<^sub>S ((S.UP \ \\<^sub>S S.\ (f, g) \\<^sub>S S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.\ (g, f)) \\<^sub>S S.UP \)) \\<^sub>S (S.UP f \\<^sub>S UP.\ (src f))" proof - have "S.inv (UP.\ (trg f)) \\<^sub>S S.UP \ \\<^sub>S S.\ (f, g) \\<^sub>S S.UP f = (S.inv (UP.\ (trg f)) \\<^sub>S S.UP f) \\<^sub>S (S.UP \ \\<^sub>S S.\ (f, g) \\<^sub>S S.UP f)" using UP.\_char S.whisker_right by simp moreover have "S.UP f \\<^sub>S S.inv (S.\ (g, f)) \\<^sub>S S.UP \ \\<^sub>S UP.\ (src f) = (S.UP f \\<^sub>S S.inv (S.\ (g, f)) \\<^sub>S S.UP \) \\<^sub>S (S.UP f \\<^sub>S UP.\ (src f))" using antipar UP.\_char S.whisker_left S.comp_assoc by simp ultimately show ?thesis using S.comp_assoc by presburger qed finally show ?thesis by argo qed show ?thesis proof - have "((S.UP \ \\<^sub>S S.\ (f, g) \\<^sub>S S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.\ (g, f)) \\<^sub>S S.UP \)) \\<^sub>S (S.UP f \\<^sub>S UP.\ (src f)) = (UP.\ (trg f) \\<^sub>S S.UP f)" proof - have "S.inv (S.inv (UP.\ (trg f)) \\<^sub>S S.UP f) \\<^sub>S S.UP f = UP.\ (trg f) \\<^sub>S S.UP f" using UP.\_char S.iso_inv_iso S.comp_arr_dom S.UP_map\<^sub>0_obj by (simp add: S.hcomp_obj_arr) moreover have "S.arr (S.UP f)" by simp moreover have "S.iso (S.inv (UP.\ (trg f)) \\<^sub>S S.UP f)" using S.iso_inv_iso S.UP_map\<^sub>0_obj by (simp add: UP.\_char(2)) ultimately show ?thesis using 2 S.invert_side_of_triangle(1) [of "S.UP f" "S.inv (UP.\ (trg f)) \\<^sub>S S.UP f" "((S.UP \ \\<^sub>S S.\ (f, g) \\<^sub>S S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.\ (g, f)) \\<^sub>S S.UP \)) \\<^sub>S (S.UP f \\<^sub>S UP.\ (src f))"] by presburger qed moreover have "S.hseq (UP.\ (trg f)) (S.UP f) \ S.iso (S.UP f \\<^sub>S UP.\ (src f)) \ S.inv (S.UP f \\<^sub>S UP.\ (src f)) = S.UP f \\<^sub>S S.inv (UP.\ (src f))" using UP.\_char S.UP_map\<^sub>0_obj by simp ultimately show ?thesis using S.invert_side_of_triangle(2) [of "UP.\ (trg f) \\<^sub>S S.UP f" "(S.UP \ \\<^sub>S S.\ (f, g) \\<^sub>S S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.\ (g, f)) \\<^sub>S S.UP \)" "S.UP f \\<^sub>S UP.\ (src f)"] by presburger qed qed hence "S.\ (trg f, f) \\<^sub>S ((S.UP \ \\<^sub>S S.\ (f, g) \\<^sub>S S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.\ (g, f)) \\<^sub>S S.UP \)) \\<^sub>S S.inv (S.\ (f, src f)) = S.\ (trg f, f) \\<^sub>S ((UP.\ (trg f) \\<^sub>S S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (UP.\ (src f)))) \\<^sub>S S.inv (S.\ (f, src f))" by simp thus ?thesis using S.comp_assoc by simp qed qed qed lemma triangle_right_implies_left: assumes "(g \ \) \ \[g, f, g] \ (\ \ g) = \\<^sup>-\<^sup>1[g] \ \[g]" shows "(\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \) = \\<^sup>-\<^sup>1[f] \ \[f]" proof - have "par ((\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \)) (\\<^sup>-\<^sup>1[f] \ \[f])" using antipar by simp moreover have "S.UP ((\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \)) = S.UP (\\<^sup>-\<^sup>1[f] \ \[f])" using assms UP_triangle(3-5) by simp ultimately show "(\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \) = \\<^sup>-\<^sup>1[f] \ \[f]" using UP.is_faithful by blast qed text \ We \emph{really} don't want to go through the ordeal of proving a dual version of \UP_triangle(5)\, do we? So let's be smart and dualize via the opposite bicategory. \ lemma triangle_left_implies_right: assumes "(\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \) = \\<^sup>-\<^sup>1[f] \ \[f]" shows "(g \ \) \ \[g, f, g] \ (\ \ g) = \\<^sup>-\<^sup>1[g] \ \[g]" proof - interpret Cop: op_bicategory V H \ \ src trg .. interpret Eop: equivalence_in_bicategory V Cop.H Cop.\ \ Cop.src Cop.trg g f \ \ using antipar unit_in_hom counit_in_hom iso_inv_iso by (unfold_locales, simp_all) have "(\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \) = \\<^sup>-\<^sup>1[f] \ \[f] \ (g \ \) \ \[g, f, g] \ (\ \ g) = \\<^sup>-\<^sup>1[g] \ \[g]" using antipar Cop.lunit_ide_simp Cop.runit_ide_simp Cop.assoc_ide_simp VVV.ide_char VVV.arr_char VV.arr_char Eop.triangle_right_implies_left by simp thus ?thesis using assms by simp qed lemma triangle_left_iff_right: shows "(\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \) = \\<^sup>-\<^sup>1[f] \ \[f] \ (g \ \) \ \[g, f, g] \ (\ \ g) = \\<^sup>-\<^sup>1[g] \ \[g]" using triangle_left_implies_right triangle_right_implies_left by auto end text \ We might as well specialize the dual result back to the strict case while we are at it. \ context equivalence_in_strict_bicategory begin lemma triangle_left_iff_right: shows "(\ \ f) \ (f \ \) = f \ (g \ \) \ (\ \ g) = g" proof - have "(\ \ f) \ (f \ \) = f \ (\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \) = \\<^sup>-\<^sup>1[f] \ \[f]" proof - have "\\<^sup>-\<^sup>1[f] \ \[f] = f" using strict_lunit strict_runit by simp moreover have "(\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \) = (\ \ f) \ (f \ \)" using antipar strict_assoc assoc'_in_hom(2) [of f g f] comp_cod_arr by auto ultimately show ?thesis by simp qed also have "... \ (g \ \) \ \[g, f, g] \ (\ \ g) = \\<^sup>-\<^sup>1[g] \ \[g]" using triangle_left_iff_right by blast also have "... \ (g \ \) \ (\ \ g) = g" proof - have "\\<^sup>-\<^sup>1[g] \ \[g] = g" using strict_lunit strict_runit by simp moreover have "(g \ \) \ \[g, f, g] \ (\ \ g) = (g \ \) \ (\ \ g)" using antipar strict_assoc assoc_in_hom(2) [of g f g] comp_cod_arr by auto ultimately show ?thesis by simp qed finally show ?thesis by simp qed end end

\<^sub>1[?R, ?w1]" let ?\ = "\?R \r0, r0\ ?R\" have P: "?P = ru.apex" using ru.apex_composite by auto have Chn_\: "\?\ : ?V \\<^sub>C ?P\" using P Chn_in_hom \ by force have p0\: "\?p0 \ ?\ : ?V \\<^sub>C ?U\" using Chn_\ ru.legs_form_cospan by auto have w1: "\?w1 : ?V \\<^sub>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 \\ = p\<^sub>0\\, but taking this literally and trying to define \\\ so that \Chn \ = ?p\<^sub>0 \ ?\\, does not yield the required \dom \ = ?R \\ ?w1\. We need \\Chn \ : ?R \\ ?w1 \\<^sub>C ?U\\, so we have to compose with a projection, which leads to: \Chn \ = ?p0 \ ?\ \ \