diff --git a/thys/Category3/Adjunction.thy b/thys/Category3/Adjunction.thy --- a/thys/Category3/Adjunction.thy +++ b/thys/Category3/Adjunction.thy @@ -1,3187 +1,3201 @@ (* Title: Adjunction Author: Eugene W. Stark , 2016 Maintainer: Eugene W. Stark *) chapter Adjunction theory Adjunction imports Yoneda begin text\ This theory defines the notions of adjoint functor and adjunction in various ways and establishes their equivalence. The notions ``left adjoint functor'' and ``right adjoint functor'' are defined in terms of universal arrows. ``Meta-adjunctions'' are defined in terms of natural bijections between hom-sets, where the notion of naturality is axiomatized directly. ``Hom-adjunctions'' formalize the notion of adjunction in terms of natural isomorphisms of hom-functors. ``Unit-counit adjunctions'' define adjunctions in terms of functors equipped with unit and counit natural transformations that satisfy the usual ``triangle identities.'' The \adjunction\ locale is defined as the grand unification of all the definitions, and includes formulas that connect the data from each of them. It is shown that each of the definitions induces an interpretation of the \adjunction\ locale, so that all the definitions are essentially equivalent. Finally, it is shown that right adjoint functors are unique up to natural isomorphism. The reference \cite{Wikipedia-Adjoint-Functors} was useful in constructing this theory. \ section "Left Adjoint Functor" text\ ``@{term e} is an arrow from @{term "F x"} to @{term y}.'' \ locale arrow_from_functor = C: category C + D: category D + F: "functor" D C F for D :: "'d comp" (infixr "\\<^sub>D" 55) and C :: "'c comp" (infixr "\\<^sub>C" 55) and F :: "'d \ 'c" and x :: 'd and y :: 'c and e :: 'c + assumes arrow: "D.ide x \ C.in_hom e (F x) y" begin notation C.in_hom ("\_ : _ \\<^sub>C _\") notation D.in_hom ("\_ : _ \\<^sub>D _\") text\ ``@{term g} is a @{term[source=true] D}-coextension of @{term f} along @{term e}.'' \ definition is_coext :: "'d \ 'c \ 'd \ bool" where "is_coext x' f g \ \g : x' \\<^sub>D x\ \ f = e \\<^sub>C F g" end text\ ``@{term e} is a terminal arrow from @{term "F x"} to @{term y}.'' \ locale terminal_arrow_from_functor = arrow_from_functor D C F x y e for D :: "'d comp" (infixr "\\<^sub>D" 55) and C :: "'c comp" (infixr "\\<^sub>C" 55) and F :: "'d \ 'c" and x :: 'd and y :: 'c and e :: 'c + assumes is_terminal: "arrow_from_functor D C F x' y f \ (\!g. is_coext x' f g)" begin definition the_coext :: "'d \ 'c \ 'd" where "the_coext x' f = (THE g. is_coext x' f g)" lemma the_coext_prop: assumes "arrow_from_functor D C F x' y f" shows "\the_coext x' f : x' \\<^sub>D x\" and "f = e \\<^sub>C F (the_coext x' f)" by (metis assms is_coext_def is_terminal the_coext_def the_equality)+ lemma the_coext_unique: assumes "arrow_from_functor D C F x' y f" and "is_coext x' f g" shows "g = the_coext x' f" using assms is_terminal the_coext_def the_equality by metis end text\ A left adjoint functor is a functor \F: D \ C\ that enjoys the following universal coextension property: for each object @{term y} of @{term C} there exists an object @{term x} of @{term D} and an arrow \e \ C.hom (F x) y\ such that for any arrow \f \ C.hom (F x') y\ there exists a unique \g \ D.hom x' x\ such that @{term "f = C e (F g)"}. \ locale left_adjoint_functor = C: category C + D: category D + "functor" D C F for D :: "'d comp" (infixr "\\<^sub>D" 55) and C :: "'c comp" (infixr "\\<^sub>C" 55) and F :: "'d \ 'c" + assumes ex_terminal_arrow: "C.ide y \ (\x e. terminal_arrow_from_functor D C F x y e)" begin notation C.in_hom ("\_ : _ \\<^sub>C _\") notation D.in_hom ("\_ : _ \\<^sub>D _\") end section "Right Adjoint Functor" text\ ``@{term e} is an arrow from @{term x} to @{term "G y"}.'' \ locale arrow_to_functor = C: category C + D: category D + G: "functor" C D G for C :: "'c comp" (infixr "\\<^sub>C" 55) and D :: "'d comp" (infixr "\\<^sub>D" 55) and G :: "'c \ 'd" and x :: 'd and y :: 'c and e :: 'd + assumes arrow: "C.ide y \ D.in_hom e x (G y)" begin notation C.in_hom ("\_ : _ \\<^sub>C _\") notation D.in_hom ("\_ : _ \\<^sub>D _\") text\ ``@{term f} is a @{term[source=true] C}-extension of @{term g} along @{term e}.'' \ definition is_ext :: "'c \ 'd \ 'c \ bool" where "is_ext y' g f \ \f : y \\<^sub>C y'\ \ g = G f \\<^sub>D e" end text\ ``@{term e} is an initial arrow from @{term x} to @{term "G y"}.'' \ locale initial_arrow_to_functor = arrow_to_functor C D G x y e for C :: "'c comp" (infixr "\\<^sub>C" 55) and D :: "'d comp" (infixr "\\<^sub>D" 55) and G :: "'c \ 'd" and x :: 'd and y :: 'c and e :: 'd + assumes is_initial: "arrow_to_functor C D G x y' g \ (\!f. is_ext y' g f)" begin definition the_ext :: "'c \ 'd \ 'c" where "the_ext y' g = (THE f. is_ext y' g f)" lemma the_ext_prop: assumes "arrow_to_functor C D G x y' g" shows "\the_ext y' g : y \\<^sub>C y'\" and "g = G (the_ext y' g) \\<^sub>D e" by (metis assms is_initial is_ext_def the_equality the_ext_def)+ lemma the_ext_unique: assumes "arrow_to_functor C D G x y' g" and "is_ext y' g f" shows "f = the_ext y' g" using assms is_initial the_ext_def the_equality by metis end text\ A right adjoint functor is a functor \G: C \ D\ that enjoys the following universal extension property: for each object @{term x} of @{term D} there exists an object @{term y} of @{term C} and an arrow \e \ D.hom x (G y)\ such that for any arrow \g \ D.hom x (G y')\ there exists a unique \f \ C.hom y y'\ such that @{term "h = D e (G f)"}. \ locale right_adjoint_functor = C: category C + D: category D + "functor" C D G for C :: "'c comp" (infixr "\\<^sub>C" 55) and D :: "'d comp" (infixr "\\<^sub>D" 55) and G :: "'c \ 'd" + assumes initial_arrows_exist: "D.ide x \ (\y e. initial_arrow_to_functor C D G x y e)" begin notation C.in_hom ("\_ : _ \\<^sub>C _\") notation D.in_hom ("\_ : _ \\<^sub>D _\") end section "Various Definitions of Adjunction" subsection "Meta-Adjunction" text\ A ``meta-adjunction'' consists of a functor \F: D \ C\, a functor \G: C \ D\, and for each object @{term x} of @{term C} and @{term y} of @{term D} a bijection between \C.hom (F y) x\ to \D.hom y (G x)\ which is natural in @{term x} and @{term y}. The naturality is easy to express at the meta-level without having to resort to the formal baggage of ``set category,'' ``hom-functor,'' and ``natural isomorphism,'' hence the name. \ locale meta_adjunction = C: category C + D: category D + F: "functor" D C F + G: "functor" C D G for C :: "'c comp" (infixr "\\<^sub>C" 55) and D :: "'d comp" (infixr "\\<^sub>D" 55) and F :: "'d \ 'c" and G :: "'c \ 'd" and \ :: "'d \ 'c \ 'd" and \ :: "'c \ 'd \ 'c" + assumes \_in_hom: "\ D.ide y; C.in_hom f (F y) x \ \ D.in_hom (\ y f) y (G x)" and \_in_hom: "\ C.ide x; D.in_hom g y (G x) \ \ C.in_hom (\ x g) (F y) x" and \_\: "\ D.ide y; C.in_hom f (F y) x \ \ \ x (\ y f) = f" and \_\: "\ C.ide x; D.in_hom g y (G x) \ \ \ y (\ x g) = g" and \_naturality: "\ C.in_hom f x x'; D.in_hom g y' y; C.in_hom h (F y) x \ \ \ y' (f \\<^sub>C h \\<^sub>C F g) = G f \\<^sub>D \ y h \\<^sub>D g" begin notation C.in_hom ("\_ : _ \\<^sub>C _\") notation D.in_hom ("\_ : _ \\<^sub>D _\") text\ The naturality of @{term \} is a consequence of the naturality of @{term \} and the other assumptions. \ lemma \_naturality: assumes f: "\f : x \\<^sub>C x'\" and g: "\g : y' \\<^sub>D y\" and h: "\h : y \\<^sub>D G x\" shows "f \\<^sub>C \ x h \\<^sub>C F g = \ x' (G f \\<^sub>D h \\<^sub>D g)" using f g h \_naturality \_in_hom C.ide_dom D.ide_dom D.in_homE \_\ \_\ by (metis C.comp_in_homI' F.preserves_hom C.in_homE D.in_homE) lemma respects_natural_isomorphism: assumes "natural_isomorphism D C F' F \" and "natural_isomorphism C D G G' \" shows "meta_adjunction C D F' G' (\y f. \ (C.cod f) \\<^sub>D \ y (f \\<^sub>C inverse_transformation.map D C F \ y)) (\x g. \ x ((inverse_transformation.map C D G' \ x) \\<^sub>D g) \\<^sub>C \ (D.dom g))" proof - interpret \: natural_isomorphism D C F' F \ using assms(1) by simp interpret \': inverse_transformation D C F' F \ .. interpret \: natural_isomorphism C D G G' \ using assms(2) by simp interpret \': inverse_transformation C D G G' \ .. let ?\' = "\y f. \ (C.cod f) \\<^sub>D \ y (f \\<^sub>C \'.map y)" let ?\' = "\x g. \ x (\'.map x \\<^sub>D g) \\<^sub>C \ (D.dom g)" show "meta_adjunction C D F' G' ?\' ?\'" proof show "\y f x. \D.ide y; \f : F' y \\<^sub>C x\\ \ \\ (C.cod f) \\<^sub>D \ y (f \\<^sub>C \'.map y) : y \\<^sub>D G' x\" proof - fix x y f assume y: "D.ide y" and f: "\f : F' y \\<^sub>C x\" show "\\ (C.cod f) \\<^sub>D \ y (f \\<^sub>C \'.map y) : y \\<^sub>D G' x\" proof (intro D.comp_in_homI) show "\\ (C.cod f) : G x \\<^sub>D G' x\" using f by fastforce show "\\ y (f \\<^sub>C \'.map y) : y \\<^sub>D G x\" using f y \_in_hom by auto qed qed show "\x g y. \C.ide x; \g : y \\<^sub>D G' x\\ \ \\ x (\'.map x \\<^sub>D g) \\<^sub>C \ (D.dom g) : F' y \\<^sub>C x\" proof - fix x y g assume x: "C.ide x" and g: "\g : y \\<^sub>D G' x\" show "\\ x (\'.map x \\<^sub>D g) \\<^sub>C \ (D.dom g) : F' y \\<^sub>C x\" proof (intro C.comp_in_homI) show "\\ (D.dom g) : F' y \\<^sub>C F y\" using g by fastforce show "\\ x (\'.map x \\<^sub>D g) : F y \\<^sub>C x\" using x g \_in_hom by auto qed qed show "\y f x. \D.ide y; \f : F' y \\<^sub>C x\\ \ \ x (\'.map x \\<^sub>D \ (C.cod f) \\<^sub>D \ y (f \\<^sub>C \'.map y)) \\<^sub>C \ (D.dom (\ (C.cod f) \\<^sub>D \ y (f \\<^sub>C \'.map y))) = f" proof - fix x y f assume y: "D.ide y" and f: "\f : F' y \\<^sub>C x\" have 1: "\\ y (f \\<^sub>C \'.map y) : y \\<^sub>D G x\" using f y \_in_hom by auto show "\ x (\'.map x \\<^sub>D \ (C.cod f) \\<^sub>D \ y (f \\<^sub>C \'.map y)) \\<^sub>C \ (D.dom (\ (C.cod f) \\<^sub>D \ y (f \\<^sub>C \'.map y))) = f" proof - have "\ x (\'.map x \\<^sub>D \ (C.cod f) \\<^sub>D \ y (f \\<^sub>C \'.map y)) \\<^sub>C \ (D.dom (\ (C.cod f) \\<^sub>D \ y (f \\<^sub>C \'.map y))) = \ x ((\'.map x \\<^sub>D \ (C.cod f)) \\<^sub>D \ y (f \\<^sub>C \'.map y)) \\<^sub>C \ (D.dom (\ (C.cod f) \\<^sub>D \ y (f \\<^sub>C \'.map y)))" using D.comp_assoc by simp also have "... = \ x (\ y (f \\<^sub>C \'.map y)) \\<^sub>C \ y" by (metis "1" C.arr_cod C.dom_cod C.ide_cod C.in_homE D.comp_ide_arr D.dom_comp D.ide_compE D.in_homE D.inverse_arrowsE \'.inverts_components \.preserves_dom \.preserves_reflects_arr category.seqI f meta_adjunction_axioms meta_adjunction_def) also have "... = f" using f y \_\ C.comp_assoc \'.inverts_components [of y] C.comp_arr_dom by fastforce finally show ?thesis by blast qed qed show "\x g y. \C.ide x; \g : y \\<^sub>D G' x\\ \ \ (C.cod (\ x (\'.map x \\<^sub>D g) \\<^sub>C \ (D.dom g))) \\<^sub>D \ y ((\ x (\'.map x \\<^sub>D g) \\<^sub>C \ (D.dom g)) \\<^sub>C \'.map y) = g" proof - fix x y g assume x: "C.ide x" and g: "\g : y \\<^sub>D G' x\" have 1: "\\ x (\'.map x \\<^sub>D g) : F y \\<^sub>C x\" using x g \_in_hom by auto show "\ (C.cod (\ x (\'.map x \\<^sub>D g) \\<^sub>C \ (D.dom g))) \\<^sub>D \ y ((\ x (\'.map x \\<^sub>D g) \\<^sub>C \ (D.dom g)) \\<^sub>C \'.map y) = g" proof - have "\ (C.cod (\ x (\'.map x \\<^sub>D g) \\<^sub>C \ (D.dom g))) \\<^sub>D \ y ((\ x (\'.map x \\<^sub>D g) \\<^sub>C \ (D.dom g)) \\<^sub>C \'.map y) = \ (C.cod (\ x (\'.map x \\<^sub>D g) \\<^sub>C \ (D.dom g))) \\<^sub>D \ y (\ x (\'.map x \\<^sub>D g) \\<^sub>C \ (D.dom g) \\<^sub>C \'.map y)" using C.comp_assoc by simp also have "... = \ x \\<^sub>D \ y (\ x (\'.map x \\<^sub>D g))" using 1 C.comp_arr_dom C.comp_arr_inv' g by fastforce also have "... = (\ x \\<^sub>D \'.map x) \\<^sub>D g" using x g \_\ D.comp_assoc by auto also have "... = g" using x g \'.inverts_components [of x] D.comp_cod_arr by fastforce finally show ?thesis by blast qed qed show "\f x x' g y' y h. \\f : x \\<^sub>C x'\; \g : y' \\<^sub>D y\; \h : F' y \\<^sub>C x\\ \ \ (C.cod (f \\<^sub>C h \\<^sub>C F' g)) \\<^sub>D \ y' ((f \\<^sub>C h \\<^sub>C F' g) \\<^sub>C \'.map y') = G' f \\<^sub>D (\ (C.cod h) \\<^sub>D \ y (h \\<^sub>C \'.map y)) \\<^sub>D g" proof - fix x y x' y' f g h assume f: "\f : x \\<^sub>C x'\" and g: "\g : y' \\<^sub>D y\" and h: "\h : F' y \\<^sub>C x\" show "\ (C.cod (f \\<^sub>C h \\<^sub>C F' g)) \\<^sub>D \ y' ((f \\<^sub>C h \\<^sub>C F' g) \\<^sub>C \'.map y') = G' f \\<^sub>D (\ (C.cod h) \\<^sub>D \ y (h \\<^sub>C \'.map y)) \\<^sub>D g" proof - have "\ (C.cod (f \\<^sub>C h \\<^sub>C F' g)) \\<^sub>D \ y' ((f \\<^sub>C h \\<^sub>C F' g) \\<^sub>C \'.map y') = \ x' \\<^sub>D \ y' ((f \\<^sub>C h \\<^sub>C F' g) \\<^sub>C \'.map y')" using f g h by fastforce also have "... = \ x' \\<^sub>D \ y' (f \\<^sub>C (h \\<^sub>C \'.map y) \\<^sub>C F g)" using g \'.naturality C.comp_assoc by auto also have "... = (\ x' \\<^sub>D G f) \\<^sub>D \ y (h \\<^sub>C \'.map y) \\<^sub>D g" using f g h \_naturality [of f x x' g y' y "h \\<^sub>C \'.map y"] D.comp_assoc by fastforce also have "... = (G' f \\<^sub>D \ x) \\<^sub>D \ y (h \\<^sub>C \'.map y) \\<^sub>D g" using f \.naturality by auto also have "... = G' f \\<^sub>D (\ (C.cod h) \\<^sub>D \ y (h \\<^sub>C \'.map y)) \\<^sub>D g" using h D.comp_assoc by auto finally show ?thesis by blast qed qed qed qed end subsection "Hom-Adjunction" text\ The bijection between hom-sets that defines an adjunction can be represented formally as a natural isomorphism of hom-functors. However, stating the definition this way is more complex than was the case for \meta_adjunction\. One reason is that we need to have a ``set category'' that is suitable as a target category for the hom-functors, and since the arrows of the categories @{term C} and @{term D} will in general have distinct types, we need a set category that simultaneously embeds both. Another reason is that we simply have to formally construct the various categories and functors required to express the definition. This is a good place to point out that I have often included more sublocales in a locale than are strictly required. The main reason for this is the fact that the locale system in Isabelle only gives one name to each entity introduced by a locale: the name that it has in the first locale in which it occurs. This means that entities that make their first appearance deeply nested in sublocales will have to be referred to by long qualified names that can be difficult to understand, or even to discover. To counteract this, I have typically introduced sublocales before the superlocales that contain them to ensure that the entities in the sublocales can be referred to by short meaningful (and predictable) names. In my opinion, though, it would be better if the locale system would make entities that occur in multiple locales accessible by \emph{all} possible qualified names, so that the most perspicuous name could be used in any particular context. \ locale hom_adjunction = C: category C + D: category D + - S: replete_set_category S + + S: set_category S setp + Cop: dual_category C + Dop: dual_category D + CopxC: product_category Cop.comp C + DopxD: product_category Dop.comp D + DopxC: product_category Dop.comp C + F: "functor" D C F + G: "functor" C D G + - HomC: hom_functor C S \C + - HomD: hom_functor D S \D + + HomC: hom_functor C S setp \C + + HomD: hom_functor D S setp \D + Fop: dual_functor Dop.comp Cop.comp F + FopxC: product_functor Dop.comp C Cop.comp C Fop.map C.map + DopxG: product_functor Dop.comp C Dop.comp D Dop.map G + Hom_FopxC: composite_functor DopxC.comp CopxC.comp S FopxC.map HomC.map + Hom_DopxG: composite_functor DopxC.comp DopxD.comp S DopxG.map HomD.map + - Hom_FopxC: set_valued_functor DopxC.comp S Hom_FopxC.map + - Hom_DopxG: set_valued_functor DopxC.comp S Hom_DopxG.map + - \: set_valued_transformation DopxC.comp S Hom_FopxC.map Hom_DopxG.map \ + - \: set_valued_transformation DopxC.comp S Hom_DopxG.map Hom_FopxC.map \ + + Hom_FopxC: set_valued_functor DopxC.comp S setp Hom_FopxC.map + + Hom_DopxG: set_valued_functor DopxC.comp S setp Hom_DopxG.map + + \: set_valued_transformation DopxC.comp S setp Hom_FopxC.map Hom_DopxG.map \ + + \: set_valued_transformation DopxC.comp S setp Hom_DopxG.map Hom_FopxC.map \ + \\: inverse_transformations DopxC.comp S Hom_FopxC.map Hom_DopxG.map \ \ for C :: "'c comp" (infixr "\\<^sub>C" 55) and D :: "'d comp" (infixr "\\<^sub>D" 55) and S :: "'s comp" (infixr "\\<^sub>S" 55) + and setp :: "'s set \ bool" and \C :: "'c * 'c \ 'c \ 's" and \D :: "'d * 'd \ 'd \ 's" and F :: "'d \ 'c" and G :: "'c \ 'd" and \ :: "'d * 'c \ 's" and \ :: "'d * 'c \ 's" begin notation C.in_hom ("\_ : _ \\<^sub>C _\") notation D.in_hom ("\_ : _ \\<^sub>D _\") abbreviation \C :: "'c * 'c \ 's \ 'c" where "\C \ HomC.\" abbreviation \D :: "'d * 'd \ 's \ 'd" where "\D \ HomD.\" end subsection "Unit/Counit Adjunction" text\ Expressed in unit/counit terms, an adjunction consists of functors \F: D \ C\ and \G: C \ D\, equipped with natural transformations \\: 1 \ GF\ and \\: FG \ 1\ satisfying certain ``triangle identities''. \ locale unit_counit_adjunction = C: category C + D: category D + F: "functor" D C F + G: "functor" C D G + GF: composite_functor D C D F G + FG: composite_functor C D C G F + FGF: composite_functor D C C F \F o G\ + GFG: composite_functor C D D G \G o F\ + \: natural_transformation D D D.map \G o F\ \ + \: natural_transformation C C \F o G\ C.map \ + F\: natural_transformation D C F \F o G o F\ \F o \\ + \G: natural_transformation C D G \G o F o G\ \\ o G\ + \F: natural_transformation D C \F o G o F\ F \\ o F\ + G\: natural_transformation C D \G o F o G\ G \G o \\ + \FoF\: vertical_composite D C F \F o G o F\ F \F o \\ \\ o F\ + G\o\G: vertical_composite C D G \G o F o G\ G \\ o G\ \G o \\ for C :: "'c comp" (infixr "\\<^sub>C" 55) and D :: "'d comp" (infixr "\\<^sub>D" 55) and F :: "'d \ 'c" and G :: "'c \ 'd" and \ :: "'d \ 'd" and \ :: "'c \ 'c" + assumes triangle_F: "\FoF\.map = F" and triangle_G: "G\o\G.map = G" begin notation C.in_hom ("\_ : _ \\<^sub>C _\") notation D.in_hom ("\_ : _ \\<^sub>D _\") end lemma unit_determines_counit: assumes "unit_counit_adjunction C D F G \ \" and "unit_counit_adjunction C D F G \ \'" shows "\ = \'" proof - (* IDEA: \' = \'FG o (FG\ o F\G) = \'\ o F\G = \FG o (\'FG o F\G) = \ *) interpret Adj: unit_counit_adjunction C D F G \ \ using assms(1) by auto interpret Adj': unit_counit_adjunction C D F G \ \' using assms(2) by auto interpret FGFG: composite_functor C D C G \F o G o F\ .. interpret FG\: natural_transformation C C \(F o G) o (F o G)\ \F o G\ \(F o G) o \\ using Adj.\.natural_transformation_axioms Adj.FG.as_nat_trans.natural_transformation_axioms horizontal_composite by fastforce interpret F\G: natural_transformation C C \F o G\ \F o G o F o G\ \F o \ o G\ using Adj.\.natural_transformation_axioms Adj.F\.natural_transformation_axioms Adj.G.as_nat_trans.natural_transformation_axioms horizontal_composite by blast interpret \'\: natural_transformation C C \F o G o F o G\ Adj.C.map \\' o \\ proof - have "natural_transformation C C ((F o G) o (F o G)) Adj.C.map (\' o \)" using Adj.\.natural_transformation_axioms Adj'.\.natural_transformation_axioms horizontal_composite Adj.C.is_functor comp_functor_identity by (metis (no_types, lifting)) thus "natural_transformation C C (F o G o F o G) Adj.C.map (\' o \)" using o_assoc by metis qed interpret \'\oF\G: vertical_composite C C \F o G\ \F o G o F o G\ Adj.C.map \F o \ o G\ \\' o \\ .. have "\' = vertical_composite.map C C (F o Adj.G\o\G.map) \'" using vcomp_ide_dom [of C C "F o G" Adj.C.map \'] Adj.triangle_G by (simp add: Adj'.\.natural_transformation_axioms) also have "... = vertical_composite.map C C (vertical_composite.map C C (F o \ o G) (F o G o \)) \'" using whisker_left Adj.F.functor_axioms Adj.G\.natural_transformation_axioms Adj.\G.natural_transformation_axioms o_assoc by (metis (no_types, lifting)) also have "... = vertical_composite.map C C (vertical_composite.map C C (F o \ o G) (\' o F o G)) \" proof - have "vertical_composite.map C C (vertical_composite.map C C (F o \ o G) (F o G o \)) \' = vertical_composite.map C C (F o \ o G) (vertical_composite.map C C (F o G o \) \')" using vcomp_assoc by (metis (no_types, lifting) Adj'.\.natural_transformation_axioms FG\.natural_transformation_axioms F\G.natural_transformation_axioms o_assoc) also have "... = vertical_composite.map C C (F o \ o G) (vertical_composite.map C C (\' o F o G) \)" using Adj'.\.natural_transformation_axioms Adj.\.natural_transformation_axioms interchange_spc [of C C "F o G" Adj.C.map \ C "F o G" Adj.C.map \'] by (metis hcomp_ide_cod hcomp_ide_dom o_assoc) also have "... = vertical_composite.map C C (vertical_composite.map C C (F o \ o G) (\' o F o G)) \" using vcomp_assoc by (metis Adj'.\F.natural_transformation_axioms Adj.G.as_nat_trans.natural_transformation_axioms Adj.\.natural_transformation_axioms F\G.natural_transformation_axioms horizontal_composite) finally show ?thesis by simp qed also have "... = vertical_composite.map C C (vertical_composite.map D C (F o \) (\' o F) o G) \" using whisker_right Adj'.\F.natural_transformation_axioms Adj.F\.natural_transformation_axioms Adj.G.functor_axioms by metis also have "... = \" using Adj'.triangle_F vcomp_ide_cod Adj.\.natural_transformation_axioms by simp finally show ?thesis by simp qed lemma counit_determines_unit: assumes "unit_counit_adjunction C D F G \ \" and "unit_counit_adjunction C D F G \' \" shows "\ = \'" proof - interpret Adj: unit_counit_adjunction C D F G \ \ using assms(1) by auto interpret Adj': unit_counit_adjunction C D F G \' \ using assms(2) by auto interpret GFGF: composite_functor D C D F \G o F o G\ .. interpret GF\: natural_transformation D D \G o F\ \(G o F) o (G o F)\ \(G o F) o \\ using Adj.\.natural_transformation_axioms Adj.GF.functor_axioms Adj.GF.as_nat_trans.natural_transformation_axioms comp_functor_identity horizontal_composite by (metis (no_types, lifting)) interpret \'GF: natural_transformation D D \G o F\ \(G o F) o (G o F)\ \\' o (G o F)\ using Adj'.\.natural_transformation_axioms Adj.GF.functor_axioms Adj.GF.as_nat_trans.natural_transformation_axioms comp_identity_functor horizontal_composite by (metis (no_types, lifting)) interpret G\F: natural_transformation D D \G o F o G o F\ \G o F\ \G o \ o F\ using Adj.\.natural_transformation_axioms Adj.F.as_nat_trans.natural_transformation_axioms Adj.G\.natural_transformation_axioms horizontal_composite by blast interpret \'\: natural_transformation D D Adj.D.map \G o F o G o F\ \\' o \\ proof - have "natural_transformation D D Adj.D.map ((G o F) o (G o F)) (\' o \)" using Adj'.\.natural_transformation_axioms Adj.D.identity_functor_axioms Adj.\.natural_transformation_axioms horizontal_composite identity_functor.is_functor by fastforce thus "natural_transformation D D Adj.D.map (G o F o G o F) (\' o \)" using o_assoc by metis qed interpret G\Fo\'\: vertical_composite D D Adj.D.map \G o F o G o F\ \G o F\ \\' o \\ \G o \ o F\ .. have "\' = vertical_composite.map D D \' (G o Adj.\FoF\.map)" using vcomp_ide_cod [of D D Adj.D.map "G o F" \'] Adj.triangle_F by (simp add: Adj'.\.natural_transformation_axioms) also have "... = vertical_composite.map D D \' (vertical_composite.map D D (G o (F o \)) (G o (\ o F)))" using whisker_left Adj.F\.natural_transformation_axioms Adj.G.functor_axioms Adj.\F.natural_transformation_axioms by fastforce also have "... = vertical_composite.map D D (vertical_composite.map D D \' (G o (F o \))) (G o \ o F)" using vcomp_assoc Adj'.\.natural_transformation_axioms GF\.natural_transformation_axioms G\F.natural_transformation_axioms o_assoc by (metis (no_types, lifting)) also have "... = vertical_composite.map D D (vertical_composite.map D D \ (\' o G o F)) (G o \ o F)" using interchange_spc [of D D Adj.D.map "G o F" \ D Adj.D.map "G o F" \'] Adj.\.natural_transformation_axioms Adj'.\.natural_transformation_axioms by (metis hcomp_ide_cod hcomp_ide_dom o_assoc) also have "... = vertical_composite.map D D \ (vertical_composite.map D D (\' o G o F) (G o \ o F))" using vcomp_assoc by (metis (no_types, lifting) Adj.\.natural_transformation_axioms G\F.natural_transformation_axioms \'GF.natural_transformation_axioms o_assoc) also have "... = vertical_composite.map D D \ (vertical_composite.map C D (\' o G) (G o \) o F)" using whisker_right Adj'.\G.natural_transformation_axioms Adj.F.functor_axioms Adj.G\.natural_transformation_axioms by fastforce also have "... = \" using Adj'.triangle_G vcomp_ide_dom Adj.GF.functor_axioms Adj.\.natural_transformation_axioms by simp finally show ?thesis by simp qed subsection "Adjunction" text\ The grand unification of everything to do with an adjunction. \ locale adjunction = C: category C + D: category D + - S: replete_set_category S + + S: set_category S setp + Cop: dual_category C + Dop: dual_category D + CopxC: product_category Cop.comp C + DopxD: product_category Dop.comp D + DopxC: product_category Dop.comp C + idDop: identity_functor Dop.comp + - HomC: hom_functor C S \C + - HomD: hom_functor D S \D + + HomC: hom_functor C S setp \C + + HomD: hom_functor D S setp \D + F: left_adjoint_functor D C F + G: right_adjoint_functor C D G + GF: composite_functor D C D F G + FG: composite_functor C D C G F + FGF: composite_functor D C C F FG.map + GFG: composite_functor C D D G GF.map + Fop: dual_functor Dop.comp Cop.comp F + FopxC: product_functor Dop.comp C Cop.comp C Fop.map C.map + DopxG: product_functor Dop.comp C Dop.comp D Dop.map G + Hom_FopxC: composite_functor DopxC.comp CopxC.comp S FopxC.map HomC.map + Hom_DopxG: composite_functor DopxC.comp DopxD.comp S DopxG.map HomD.map + - Hom_FopxC: set_valued_functor DopxC.comp S Hom_FopxC.map + - Hom_DopxG: set_valued_functor DopxC.comp S Hom_DopxG.map + + Hom_FopxC: set_valued_functor DopxC.comp S setp Hom_FopxC.map + + Hom_DopxG: set_valued_functor DopxC.comp S setp Hom_DopxG.map + \: natural_transformation D D D.map GF.map \ + \: natural_transformation C C FG.map C.map \ + F\: natural_transformation D C F \F o G o F\ \F o \\ + \G: natural_transformation C D G \G o F o G\ \\ o G\ + \F: natural_transformation D C \F o G o F\ F \\ o F\ + G\: natural_transformation C D \G o F o G\ G \G o \\ + \FoF\: vertical_composite D C F FGF.map F \F o \\ \\ o F\ + G\o\G: vertical_composite C D G GFG.map G \\ o G\ \G o \\ + \\: meta_adjunction C D F G \ \ + \\: unit_counit_adjunction C D F G \ \ + - \\: hom_adjunction C D S \C \D F G \ \ + \\: hom_adjunction C D S setp \C \D F G \ \ for C :: "'c comp" (infixr "\\<^sub>C" 55) and D :: "'d comp" (infixr "\\<^sub>D" 55) and S :: "'s comp" (infixr "\\<^sub>S" 55) + and setp :: "'s set \ bool" and \C :: "'c * 'c \ 'c \ 's" and \D :: "'d * 'd \ 'd \ 's" and F :: "'d \ 'c" and G :: "'c \ 'd" and \ :: "'d \ 'c \ 'd" and \ :: "'c \ 'd \ 'c" and \ :: "'d \ 'd" and \ :: "'c \ 'c" and \ :: "'d * 'c \ 's" and \ :: "'d * 'c \ 's" + assumes \_in_terms_of_\: "\ D.ide y; \f : F y \\<^sub>C x\ \ \ \ y f = G f \\<^sub>D \ y" and \_in_terms_of_\: "\ C.ide x; \g : y \\<^sub>D G x\ \ \ \ x g = \ x \\<^sub>C F g" and \_in_terms_of_\: "D.ide y \ \ y = \ y (F y)" and \_in_terms_of_\: "C.ide x \ \ x = \ x (G x)" and \_in_terms_of_\: "\ D.ide y; \f : F y \\<^sub>C x\ \ \ \ y f = (\\.\D (y, G x) o S.Fun (\ (y, x)) o \C (F y, x)) f" and \_in_terms_of_\: "\ C.ide x; \g : y \\<^sub>D G x\ \ \ \ x g = (\\.\C (F y, x) o S.Fun (\ (y, x)) o \D (y, G x)) g" and \_in_terms_of_\: "\ C.ide x; D.ide y \ \ \ (y, x) = S.mkArr (HomC.set (F y, x)) (HomD.set (y, G x)) (\D (y, G x) o \ y o \\.\C (F y, x))" and \_in_terms_of_\: "\ C.ide x; D.ide y \ \ \ (y, x) = S.mkArr (HomD.set (y, G x)) (HomC.set (F y, x)) (\C (F y, x) o \ x o \\.\D (y, G x))" section "Meta-Adjunctions Induce Unit/Counit Adjunctions" context meta_adjunction begin interpretation GF: composite_functor D C D F G .. interpretation FG: composite_functor C D C G F .. interpretation FGF: composite_functor D C C F FG.map .. interpretation GFG: composite_functor C D D G GF.map .. definition \o :: "'d \ 'd" where "\o y = \ y (F y)" lemma \o_in_hom: assumes "D.ide y" shows "\\o y : y \\<^sub>D G (F y)\" using assms D.ide_in_hom \o_def \_in_hom by force lemma \_in_terms_of_\o: assumes "D.ide y" and "\f : F y \\<^sub>C x\" shows "\ y f = G f \\<^sub>D \o y" proof (unfold \o_def) have 1: "\F y : F y \\<^sub>C F y\" using assms(1) D.ide_in_hom by blast hence "\ y (F y) = \ y (F y) \\<^sub>D y" by (metis assms(1) D.in_homE \_in_hom D.comp_arr_dom) thus "\ y f = G f \\<^sub>D \ y (F y)" using assms 1 D.ide_in_hom by (metis C.comp_arr_dom C.in_homE \_naturality) qed lemma \_F_char: assumes "\g : y' \\<^sub>D y\" shows "\ y' (F g) = \o y \\<^sub>D g" using assms \o_def \_in_hom [of y "F y" "F y"] D.comp_cod_arr [of "D (\ y (F y)) g" "G (F y)"] \_naturality [of "F y" "F y" "F y" g y' y "F y"] by (metis C.ide_in_hom D.arr_cod_iff_arr D.arr_dom D.cod_cod D.cod_dom D.comp_ide_arr D.comp_ide_self D.ide_cod D.in_homE F.as_nat_trans.is_natural_2 F.functor_axioms F.preserves_section_retraction \_in_hom functor.preserves_hom) interpretation \: transformation_by_components D D D.map GF.map \o proof show "\a. D.ide a \ \\o a : D.map a \\<^sub>D GF.map a\" using \o_def \_in_hom D.ide_in_hom by force fix f assume f: "D.arr f" show "\o (D.cod f) \\<^sub>D D.map f = GF.map f \\<^sub>D \o (D.dom f)" using f \_F_char [of "D.map f" "D.dom f" "D.cod f"] \_in_terms_of_\o [of "D.dom f" "F f" "F (D.cod f)"] by force qed lemma \_map_simp: assumes "D.ide y" shows "\.map y = \ y (F y)" using assms \.map_simp_ide \o_def by simp definition \o :: "'c \ 'c" where "\o x = \ x (G x)" lemma \o_in_hom: assumes "C.ide x" shows "\\o x : F (G x) \\<^sub>C x\" using assms C.ide_in_hom \o_def \_in_hom by force lemma \_in_terms_of_\o: assumes "C.ide x" and "\g : y \\<^sub>D G x\" shows "\ x g = \o x \\<^sub>C F g" proof - have "\o x \\<^sub>C F g = x \\<^sub>C \ x (G x) \\<^sub>C F g" using assms \o_def \_in_hom [of x "G x" "G x"] C.comp_cod_arr [of "\ x (G x) \\<^sub>C F g" x] by fastforce also have "... = \ x (G x \\<^sub>D G x \\<^sub>D g)" using assms \_naturality [of x x x g y "G x" "G x"] by force also have "... = \ x g" using assms D.comp_cod_arr by fastforce finally show ?thesis by simp qed lemma \_G_char: assumes "\f: x \\<^sub>C x'\" shows "\ x' (G f) = f \\<^sub>C \o x" proof (unfold \o_def) have 0: "C.ide x \ C.ide x'" using assms by auto thus "\ x' (G f) = f \\<^sub>C \ x (G x)" using 0 assms \_naturality \_in_hom [of x "G x" "G x"] G.preserves_hom \o_def \_in_terms_of_\o G.as_nat_trans.is_natural_1 C.ide_in_hom by (metis C.arrI C.in_homE) qed interpretation \: transformation_by_components C C FG.map C.map \o apply unfold_locales using \o_in_hom apply simp using \_G_char \_in_terms_of_\o by (metis C.arr_iff_in_hom C.ide_cod C.map_simp G.preserves_hom comp_apply) lemma \_map_simp: assumes "C.ide x" shows "\.map x = \ x (G x)" using assms \o_def by simp interpretation FD: composite_functor D D C D.map F .. interpretation CF: composite_functor D C C F C.map .. interpretation GC: composite_functor C C D C.map G .. interpretation DG: composite_functor C D D G D.map .. interpretation F\: natural_transformation D C F \F o G o F\ \F o \.map\ by (metis (no_types, lifting) F.as_nat_trans.natural_transformation_axioms F.functor_axioms \.natural_transformation_axioms comp_functor_identity horizontal_composite o_assoc) interpretation \F: natural_transformation D C \F o G o F\ F \\.map o F\ using \.natural_transformation_axioms F.as_nat_trans.natural_transformation_axioms horizontal_composite by fastforce interpretation \G: natural_transformation C D G \G o F o G\ \\.map o G\ using \.natural_transformation_axioms G.as_nat_trans.natural_transformation_axioms horizontal_composite by fastforce interpretation G\: natural_transformation C D \G o F o G\ G \G o \.map\ by (metis (no_types, lifting) G.as_nat_trans.natural_transformation_axioms G.functor_axioms \.natural_transformation_axioms comp_functor_identity horizontal_composite o_assoc) interpretation \FoF\: vertical_composite D C F \F o G o F\ F \F o \.map\ \\.map o F\ .. interpretation G\o\G: vertical_composite C D G \G o F o G\ G \\.map o G\ \G o \.map\ .. lemma unit_counit_F: assumes "D.ide y" shows "F y = \o (F y) \\<^sub>C F (\o y)" using assms \_in_terms_of_\o \o_def \_\ \o_in_hom F.preserves_ide C.ide_in_hom by metis lemma unit_counit_G: assumes "C.ide x" shows "G x = G (\o x) \\<^sub>D \o (G x)" using assms \_in_terms_of_\o \o_def \_\ \o_in_hom G.preserves_ide D.ide_in_hom by metis lemma induces_unit_counit_adjunction': shows "unit_counit_adjunction C D F G \.map \.map" proof show "\FoF\.map = F" using \FoF\.is_natural_transformation \FoF\.map_simp_ide unit_counit_F F.as_nat_trans.natural_transformation_axioms by (intro NaturalTransformation.eqI) auto show "G\o\G.map = G" using G\o\G.is_natural_transformation G\o\G.map_simp_ide unit_counit_G G.as_nat_trans.natural_transformation_axioms by (intro NaturalTransformation.eqI) auto qed definition \ :: "'d \ 'd" where "\ \ \.map" definition \ :: "'c \ 'c" where "\ \ \.map" theorem induces_unit_counit_adjunction: shows "unit_counit_adjunction C D F G \ \" unfolding \_def \_def using induces_unit_counit_adjunction' by simp lemma \_is_natural_transformation: shows "natural_transformation D D D.map GF.map \" unfolding \_def .. lemma \_is_natural_transformation: shows "natural_transformation C C FG.map C.map \" unfolding \_def .. text\ From the defined @{term \} and @{term \} we can recover the original @{term \} and @{term \}. \ lemma \_in_terms_of_\: assumes "D.ide y" and "\f : F y \\<^sub>C x\" shows "\ y f = G f \\<^sub>D \ y" using assms \_def by (simp add: \_in_terms_of_\o) lemma \_in_terms_of_\: assumes "C.ide x" and "\g : y \\<^sub>D G x\" shows "\ x g = \ x \\<^sub>C F g" using assms \_def by (simp add: \_in_terms_of_\o) end section "Meta-Adjunctions Induce Left and Right Adjoint Functors" context meta_adjunction begin interpretation unit_counit_adjunction C D F G \ \ using induces_unit_counit_adjunction \_def \_def by auto lemma has_terminal_arrows_from_functor: assumes x: "C.ide x" shows "terminal_arrow_from_functor D C F (G x) x (\ x)" and "\y' f. arrow_from_functor D C F y' x f \ terminal_arrow_from_functor.the_coext D C F (G x) (\ x) y' f = \ y' f" proof - interpret \x: arrow_from_functor D C F \G x\ x \\ x\ using x \.preserves_hom G.preserves_ide by unfold_locales auto have 1: "\y' f. arrow_from_functor D C F y' x f \ \x.is_coext y' f (\ y' f) \ (\g'. \x.is_coext y' f g' \ g' = \ y' f)" using x by (metis (full_types) \x.is_coext_def \_\ \_in_terms_of_\ arrow_from_functor.arrow \_in_hom \_\) interpret \x: terminal_arrow_from_functor D C F \G x\ x \\ x\ using 1 by unfold_locales blast show "terminal_arrow_from_functor D C F (G x) x (\ x)" .. show "\y' f. arrow_from_functor D C F y' x f \ \x.the_coext y' f = \ y' f" using 1 \x.the_coext_def by auto qed lemma has_left_adjoint_functor: shows "left_adjoint_functor D C F" apply unfold_locales using has_terminal_arrows_from_functor by auto lemma has_initial_arrows_to_functor: assumes y: "D.ide y" shows "initial_arrow_to_functor C D G y (F y) (\ y)" and "\x' g. arrow_to_functor C D G y x' g \ initial_arrow_to_functor.the_ext C D G (F y) (\ y) x' g = \ x' g" proof - interpret \y: arrow_to_functor C D G y \F y\ \\ y\ using y by unfold_locales auto have 1: "\x' g. arrow_to_functor C D G y x' g \ \y.is_ext x' g (\ x' g) \ (\f'. \y.is_ext x' g f' \ f' = \ x' g)" using y by (metis (full_types) \y.is_ext_def \_\ \_in_terms_of_\ arrow_to_functor.arrow \_in_hom \_\) interpret \y: initial_arrow_to_functor C D G y \F y\ \\ y\ apply unfold_locales using 1 by blast show "initial_arrow_to_functor C D G y (F y) (\ y)" .. show "\x' g. arrow_to_functor C D G y x' g \ \y.the_ext x' g = \ x' g" using 1 \y.the_ext_def by auto qed lemma has_right_adjoint_functor: shows "right_adjoint_functor C D G" apply unfold_locales using has_initial_arrows_to_functor by auto end section "Unit/Counit Adjunctions Induce Meta-Adjunctions" context unit_counit_adjunction begin definition \ :: "'d \ 'c \ 'd" where "\ y h = G h \\<^sub>D \ y" definition \ :: "'c \ 'd \ 'c" where "\ x h = \ x \\<^sub>C F h" interpretation meta_adjunction C D F G \ \ proof fix x :: 'c and y :: 'd and f :: 'c assume y: "D.ide y" and f: "\f : F y \\<^sub>C x\" show 0: "\\ y f : y \\<^sub>D G x\" using f y G.preserves_hom \.preserves_hom \_def D.ide_in_hom by auto show "\ x (\ y f) = f" proof - have "\ x (\ y f) = (\ x \\<^sub>C F (G f)) \\<^sub>C F (\ y)" using y f \_def \_def C.comp_assoc by auto also have "... = (f \\<^sub>C \ (F y)) \\<^sub>C F (\ y)" using y f \.naturality by auto also have "... = f" using y f \FoF\.map_simp_2 triangle_F C.comp_arr_dom D.ide_in_hom C.comp_assoc by fastforce finally show ?thesis by auto qed next fix x :: 'c and y :: 'd and g :: 'd assume x: "C.ide x" and g: "\g : y \\<^sub>D G x\" show "\\ x g : F y \\<^sub>C x\" using g x \_def by fastforce show "\ y (\ x g) = g" proof - have "\ y (\ x g) = (G (\ x) \\<^sub>D \ (G x)) \\<^sub>D g" using g x \_def \_def \.naturality [of g] D.comp_assoc by auto also have "... = g" using x g triangle_G D.comp_ide_arr G\o\G.map_simp_ide by auto finally show ?thesis by auto qed next fix f :: 'c and g :: 'd and h :: 'c and x :: 'c and x' :: 'c and y :: 'd and y' :: 'd assume f: "\f : x \\<^sub>C x'\" and g: "\g : y' \\<^sub>D y\" and h: "\h : F y \\<^sub>C x\" show "\ y' (f \\<^sub>C h \\<^sub>C F g) = G f \\<^sub>D \ y h \\<^sub>D g" using \_def f g h \.naturality D.comp_assoc by fastforce qed theorem induces_meta_adjunction: shows "meta_adjunction C D F G \ \" .. text\ From the defined @{term \} and @{term \} we can recover the original @{term \} and @{term \}. \ lemma \_in_terms_of_\: assumes "D.ide y" shows "\ y = \ y (F y)" using assms \_def D.comp_cod_arr by auto lemma \_in_terms_of_\: assumes "C.ide x" shows "\ x = \ x (G x)" using assms \_def C.comp_arr_dom by auto end section "Left and Right Adjoint Functors Induce Meta-Adjunctions" text\ A left adjoint functor induces a meta-adjunction, modulo the choice of a right adjoint and counit. \ context left_adjoint_functor begin definition Go :: "'c \ 'd" where "Go a = (SOME b. \e. terminal_arrow_from_functor D C F b a e)" definition \o :: "'c \ 'c" where "\o a = (SOME e. terminal_arrow_from_functor D C F (Go a) a e)" lemma Go_\o_terminal: assumes "\b e. terminal_arrow_from_functor D C F b a e" shows "terminal_arrow_from_functor D C F (Go a) a (\o a)" using assms Go_def \o_def someI_ex [of "\b. \e. terminal_arrow_from_functor D C F b a e"] someI_ex [of "\e. terminal_arrow_from_functor D C F (Go a) a e"] by simp text\ The right adjoint @{term G} to @{term F} takes each arrow @{term f} of @{term[source=true] C} to the unique @{term[source=true] D}-coextension of @{term "C f (\o (C.dom f))"} along @{term "\o (C.cod f)"}. \ definition G :: "'c \ 'd" where "G f = (if C.arr f then terminal_arrow_from_functor.the_coext D C F (Go (C.cod f)) (\o (C.cod f)) (Go (C.dom f)) (f \\<^sub>C \o (C.dom f)) else D.null)" lemma G_ide: assumes "C.ide x" shows "G x = Go x" proof - interpret terminal_arrow_from_functor D C F \Go x\ x \\o x\ using assms ex_terminal_arrow Go_\o_terminal by blast have 1: "arrow_from_functor D C F (Go x) x (\o x)" .. have "is_coext (Go x) (\o x) (Go x)" using arrow is_coext_def C.in_homE C.comp_arr_dom by auto hence "Go x = the_coext (Go x) (\o x)" using 1 the_coext_unique by blast moreover have "\o x = C x (\o (C.dom x))" using assms arrow C.comp_ide_arr C.seqI' C.ide_in_hom C.in_homE by metis ultimately show ?thesis using assms G_def C.cod_dom C.ide_in_hom C.in_homE by metis qed lemma G_is_functor: shows "functor C D G" proof fix f :: 'c assume "\C.arr f" thus "G f = D.null" using G_def by auto next fix f :: 'c assume f: "C.arr f" let ?x = "C.dom f" let ?x' = "C.cod f" interpret x\: terminal_arrow_from_functor D C F \Go ?x\ \?x\ \\o ?x\ using f ex_terminal_arrow Go_\o_terminal by simp interpret x'\: terminal_arrow_from_functor D C F \Go ?x'\ \?x'\ \\o ?x'\ using f ex_terminal_arrow Go_\o_terminal by simp have 1: "arrow_from_functor D C F (Go ?x) ?x' (C f (\o ?x))" using f x\.arrow by (unfold_locales, auto) have "G f = x'\.the_coext (Go ?x) (C f (\o ?x))" using f G_def by simp hence Gf: "\G f : Go ?x \\<^sub>D Go ?x'\ \ f \\<^sub>C \o ?x = \o ?x' \\<^sub>C F (G f)" using 1 x'\.the_coext_prop by simp show "D.arr (G f)" using Gf by auto show "D.dom (G f) = G ?x" using f Gf G_ide by auto show "D.cod (G f) = G ?x'" using f Gf G_ide by auto next fix f f' :: 'c assume ff': "C.arr (C f' f)" have f: "C.arr f" using ff' by auto let ?x = "C.dom f" let ?x' = "C.cod f" let ?x'' = "C.cod f'" interpret x\: terminal_arrow_from_functor D C F \Go ?x\ \?x\ \\o ?x\ using f ex_terminal_arrow Go_\o_terminal by simp interpret x'\: terminal_arrow_from_functor D C F \Go ?x'\ \?x'\ \\o ?x'\ using f ex_terminal_arrow Go_\o_terminal by simp interpret x''\: terminal_arrow_from_functor D C F \Go ?x''\ \?x''\ \\o ?x''\ using ff' ex_terminal_arrow Go_\o_terminal by auto have 1: "arrow_from_functor D C F (Go ?x) ?x' (f \\<^sub>C \o ?x)" using f x\.arrow by (unfold_locales, auto) have 2: "arrow_from_functor D C F (Go ?x') ?x'' (f' \\<^sub>C \o ?x')" using ff' x'\.arrow by (unfold_locales, auto) have "G f = x'\.the_coext (Go ?x) (C f (\o ?x))" using f G_def by simp hence Gf: "D.in_hom (G f) (Go ?x) (Go ?x') \ f \\<^sub>C \o ?x = \o ?x' \\<^sub>C F (G f)" using 1 x'\.the_coext_prop by simp have "G f' = x''\.the_coext (Go ?x') (f' \\<^sub>C \o ?x')" using ff' G_def by auto hence Gf': "\G f' : Go (C.cod f) \\<^sub>D Go (C.cod f')\ \ f' \\<^sub>C \o ?x' = \o ?x'' \\<^sub>C F (G f')" using 2 x''\.the_coext_prop by simp show "G (f' \\<^sub>C f) = G f' \\<^sub>D G f" proof - have "x''\.is_coext (Go ?x) ((f' \\<^sub>C f) \\<^sub>C \o ?x) (G f' \\<^sub>D G f)" proof - have 3: "\G f' \\<^sub>D G f : Go (C.dom f) \\<^sub>D Go (C.cod f')\" using 1 2 Gf Gf' by auto moreover have "(f' \\<^sub>C f) \\<^sub>C \o ?x = \o ?x'' \\<^sub>C F (G f' \\<^sub>D G f)" by (metis 3 C.comp_assoc D.in_homE Gf Gf' preserves_comp) ultimately show ?thesis using x''\.is_coext_def by auto qed moreover have "arrow_from_functor D C F (Go ?x) ?x'' ((f' \\<^sub>C f) \\<^sub>C \o ?x)" using ff' x\.arrow by unfold_locales blast ultimately show ?thesis using ff' G_def x''\.the_coext_unique C.seqE C.cod_comp C.dom_comp by auto qed qed interpretation G: "functor" C D G using G_is_functor by auto lemma G_simp: assumes "C.arr f" shows "G f = terminal_arrow_from_functor.the_coext D C F (Go (C.cod f)) (\o (C.cod f)) (Go (C.dom f)) (f \\<^sub>C \o (C.dom f))" using assms G_def by simp interpretation idC: identity_functor C .. interpretation GF: composite_functor C D C G F .. interpretation \: transformation_by_components C C GF.map C.map \o proof fix x :: 'c assume x: "C.ide x" show "\\o x : GF.map x \\<^sub>C C.map x\" proof - interpret terminal_arrow_from_functor D C F \Go x\ x \\o x\ using x Go_\o_terminal ex_terminal_arrow by simp show ?thesis using x G_ide arrow by auto qed next fix f :: 'c assume f: "C.arr f" show "\o (C.cod f) \\<^sub>C GF.map f = C.map f \\<^sub>C \o (C.dom f)" proof - let ?x = "C.dom f" let ?x' = "C.cod f" interpret x\: terminal_arrow_from_functor D C F \Go ?x\ ?x \\o ?x\ using f Go_\o_terminal ex_terminal_arrow by simp interpret x'\: terminal_arrow_from_functor D C F \Go ?x'\ ?x' \\o ?x'\ using f Go_\o_terminal ex_terminal_arrow by simp have 1: "arrow_from_functor D C F (Go ?x) ?x' (C f (\o ?x))" using f x\.arrow by unfold_locales auto have "G f = x'\.the_coext (Go ?x) (f \\<^sub>C \o ?x)" using f G_simp by blast hence "x'\.is_coext (Go ?x) (f \\<^sub>C \o ?x) (G f)" using 1 x'\.the_coext_prop x'\.is_coext_def by auto thus ?thesis using f x'\.is_coext_def by simp qed qed definition \ where "\ x h = C (\.map x) (F h)" lemma \_in_hom: assumes "C.ide x" and "\g : y \\<^sub>D G x\" shows "\\ x g : F y \\<^sub>C x\" unfolding \_def using assms \.maps_ide_in_hom by auto lemma \_natural: assumes f: "\f : x \\<^sub>C x'\" and g: "\g : y' \\<^sub>D y\" and h: "\h : y \\<^sub>D G x\" shows "f \\<^sub>C \ x h \\<^sub>C F g = \ x' ((G f \\<^sub>D h) \\<^sub>D g)" proof - have "f \\<^sub>C \ x h \\<^sub>C F g = f \\<^sub>C (\.map x \\<^sub>C F h) \\<^sub>C F g" unfolding \_def by auto also have "... = (f \\<^sub>C \.map x) \\<^sub>C F h \\<^sub>C F g" using C.comp_assoc by fastforce also have "... = (f \\<^sub>C \.map x) \\<^sub>C F (h \\<^sub>D g)" using g h by fastforce also have "... = (\.map x' \\<^sub>C F (G f)) \\<^sub>C F (h \\<^sub>D g)" using f \.naturality by auto also have "... = \.map x' \\<^sub>C F ((G f \\<^sub>D h) \\<^sub>D g)" using f g h C.comp_assoc by fastforce also have "... = \ x' ((G f \\<^sub>D h) \\<^sub>D g)" unfolding \_def by auto finally show ?thesis by auto qed lemma \_inverts_coext: assumes x: "C.ide x" and g: "\g : y \\<^sub>D G x\" shows "arrow_from_functor.is_coext D C F (G x) (\.map x) y (\ x g) g" proof - interpret x\: arrow_from_functor D C F \G x\ x \\.map x\ using x \.maps_ide_in_hom by unfold_locales auto show "x\.is_coext y (\ x g) g" using x g \_def x\.is_coext_def G_ide by blast qed lemma \_invertible: assumes y: "D.ide y" and f: "\f : F y \\<^sub>C x\" shows "\!g. \g : y \\<^sub>D G x\ \ \ x g = f" proof have x: "C.ide x" using f by auto interpret x\: terminal_arrow_from_functor D C F \Go x\ x \\o x\ using x ex_terminal_arrow Go_\o_terminal by auto have 1: "arrow_from_functor D C F y x f" using y f by (unfold_locales, auto) let ?g = "x\.the_coext y f" have "\ x ?g = f" using 1 x y \_def x\.the_coext_prop G_ide \_inverts_coext x\.is_coext_def by simp thus "\?g : y \\<^sub>D G x\ \ \ x ?g = f" using 1 x x\.the_coext_prop G_ide by simp show "\g'. \g' : y \\<^sub>D G x\ \ \ x g' = f \ g' = ?g" using 1 x y \_inverts_coext G_ide x\.the_coext_unique by force qed definition \ where "\ y f = (THE g. \g : y \\<^sub>D G (C.cod f)\ \ \ (C.cod f) g = f)" lemma \_in_hom: assumes "D.ide y" and "\f : F y \\<^sub>C x\" shows "\\ y f : y \\<^sub>D G x\" using assms \_invertible \_def theI' [of "\g. \g : y \\<^sub>D G x\ \ \ x g = f"] by auto lemma \_\: assumes "C.ide x" and "\g : y \\<^sub>D G x\" shows "\ y (\ x g) = g" proof - have "\ y (\ x g) = (THE g'. \g' : y \\<^sub>D G x\ \ \ x g' = \ x g)" proof - have "C.cod (\ x g) = x" using assms \_in_hom by auto thus ?thesis using \_def by auto qed moreover have "\!g'. \g' : y \\<^sub>D G x\ \ \ x g' = \ x g" using assms \_in_hom \_invertible D.ide_dom by blast ultimately show "\ y (\ x g) = g" using assms(2) by auto qed lemma \_\: assumes "D.ide y" and "\f : F y \\<^sub>C x\" shows "\ x (\ y f) = f" using assms \_invertible \_def theI' [of "\g. \g : y \\<^sub>D G x\ \ \ x g = f"] by auto lemma \_natural: assumes "\f : x \\<^sub>C x'\" and "\g : y' \\<^sub>D y\" and "\h : F y \\<^sub>C x\" shows "\ y' (f \\<^sub>C h \\<^sub>C F g) = (G f \\<^sub>D \ y h) \\<^sub>D g" proof - have "C.ide x' \ D.ide y \ D.in_hom (\ y h) y (G x)" using assms \_in_hom by auto thus ?thesis using assms D.comp_in_homI G.preserves_hom \_natural [of f x x' g y' y "\ y h"] \_\ \_\ by auto qed theorem induces_meta_adjunction: shows "meta_adjunction C D F G \ \" using \_in_hom \_in_hom \_\ \_\ \_natural D.comp_assoc by unfold_locales auto end text\ A right adjoint functor induces a meta-adjunction, modulo the choice of a left adjoint and unit. \ context right_adjoint_functor begin definition Fo :: "'d \ 'c" where "Fo y = (SOME x. \u. initial_arrow_to_functor C D G y x u)" definition \o :: "'d \ 'd" where "\o y = (SOME u. initial_arrow_to_functor C D G y (Fo y) u)" lemma Fo_\o_initial: assumes "\x u. initial_arrow_to_functor C D G y x u" shows "initial_arrow_to_functor C D G y (Fo y) (\o y)" using assms Fo_def \o_def someI_ex [of "\x. \u. initial_arrow_to_functor C D G y x u"] someI_ex [of "\u. initial_arrow_to_functor C D G y (Fo y) u"] by simp text\ The left adjoint @{term F} to @{term g} takes each arrow @{term g} of @{term[source=true] D} to the unique @{term[source=true] C}-extension of @{term "D (\o (D.cod g)) g"} along @{term "\o (D.dom g)"}. \ definition F :: "'d \ 'c" where "F g = (if D.arr g then initial_arrow_to_functor.the_ext C D G (Fo (D.dom g)) (\o (D.dom g)) (Fo (D.cod g)) (\o (D.cod g) \\<^sub>D g) else C.null)" lemma F_ide: assumes "D.ide y" shows "F y = Fo y" proof - interpret initial_arrow_to_functor C D G y \Fo y\ \\o y\ using assms initial_arrows_exist Fo_\o_initial by blast have 1: "arrow_to_functor C D G y (Fo y) (\o y)" .. have "is_ext (Fo y) (\o y) (Fo y)" unfolding is_ext_def using arrow D.comp_ide_arr [of "G (Fo y)" "\o y"] by force hence "Fo y = the_ext (Fo y) (\o y)" using 1 the_ext_unique by blast moreover have "\o y = D (\o (D.cod y)) y" using assms arrow D.comp_arr_ide D.comp_arr_dom by auto ultimately show ?thesis using assms F_def D.dom_cod D.in_homE D.ide_in_hom by metis qed lemma F_is_functor: shows "functor D C F" proof fix g :: 'd assume "\D.arr g" thus "F g = C.null" using F_def by auto next fix g :: 'd assume g: "D.arr g" let ?y = "D.dom g" let ?y' = "D.cod g" interpret y\: initial_arrow_to_functor C D G ?y \Fo ?y\ \\o ?y\ using g initial_arrows_exist Fo_\o_initial by simp interpret y'\: initial_arrow_to_functor C D G ?y' \Fo ?y'\ \\o ?y'\ using g initial_arrows_exist Fo_\o_initial by simp have 1: "arrow_to_functor C D G ?y (Fo ?y') (D (\o ?y') g)" using g y'\.arrow by unfold_locales auto have "F g = y\.the_ext (Fo ?y') (D (\o ?y') g)" using g F_def by simp hence Fg: "\F g : Fo ?y \\<^sub>C Fo ?y'\ \ \o ?y' \\<^sub>D g = G (F g) \\<^sub>D \o ?y" using 1 y\.the_ext_prop by simp show "C.arr (F g)" using Fg by auto show "C.dom (F g) = F ?y" using Fg g F_ide by auto show "C.cod (F g) = F ?y'" using Fg g F_ide by auto next fix g :: 'd fix g' :: 'd assume g': "D.arr (D g' g)" have g: "D.arr g" using g' by auto let ?y = "D.dom g" let ?y' = "D.cod g" let ?y'' = "D.cod g'" interpret y\: initial_arrow_to_functor C D G ?y \Fo ?y\ \\o ?y\ using g initial_arrows_exist Fo_\o_initial by simp interpret y'\: initial_arrow_to_functor C D G ?y' \Fo ?y'\ \\o ?y'\ using g initial_arrows_exist Fo_\o_initial by simp interpret y''\: initial_arrow_to_functor C D G ?y'' \Fo ?y''\ \\o ?y''\ using g' initial_arrows_exist Fo_\o_initial by auto have 1: "arrow_to_functor C D G ?y (Fo ?y') (\o ?y' \\<^sub>D g)" using g y'\.arrow by unfold_locales auto have "F g = y\.the_ext (Fo ?y') (\o ?y' \\<^sub>D g)" using g F_def by simp hence Fg: "\F g : Fo ?y \\<^sub>C Fo ?y'\ \ \o ?y' \\<^sub>D g = G (F g) \\<^sub>D \o ?y" using 1 y\.the_ext_prop by simp have 2: "arrow_to_functor C D G ?y' (Fo ?y'') (\o ?y'' \\<^sub>D g')" using g' y''\.arrow by unfold_locales auto have "F g' = y'\.the_ext (Fo ?y'') (\o ?y'' \\<^sub>D g')" using g' F_def by auto hence Fg': "\F g' : Fo ?y' \\<^sub>C Fo ?y''\ \ \o ?y'' \\<^sub>D g' = G (F g') \\<^sub>D \o ?y'" using 2 y'\.the_ext_prop by simp show "F (g' \\<^sub>D g) = F g' \\<^sub>C F g" proof - have "y\.is_ext (Fo ?y'') (\o ?y'' \\<^sub>D g' \\<^sub>D g) (F g' \\<^sub>C F g)" proof - have 3: "\F g' \\<^sub>C F g : Fo ?y \\<^sub>C Fo ?y''\" using 1 2 Fg Fg' by auto moreover have "\o ?y'' \\<^sub>D g' \\<^sub>D g = G (F g' \\<^sub>C F g) \\<^sub>D \o ?y" using Fg Fg' g g' 3 y''\.arrow by (metis C.arrI D.comp_assoc preserves_comp) ultimately show ?thesis using y\.is_ext_def by auto qed moreover have "arrow_to_functor C D G ?y (Fo ?y'') (\o ?y'' \\<^sub>D g' \\<^sub>D g)" using g g' y''\.arrow by unfold_locales auto ultimately show ?thesis using g g' F_def y\.the_ext_unique D.dom_comp D.cod_comp by auto qed qed interpretation F: "functor" D C F using F_is_functor by auto lemma F_simp: assumes "D.arr g" shows "F g = initial_arrow_to_functor.the_ext C D G (Fo (D.dom g)) (\o (D.dom g)) (Fo (D.cod g)) (\o (D.cod g) \\<^sub>D g)" using assms F_def by simp interpretation FG: composite_functor D C D F G .. interpretation \: transformation_by_components D D D.map FG.map \o proof fix y :: 'd assume y: "D.ide y" show "\\o y : D.map y \\<^sub>D FG.map y\" proof - interpret initial_arrow_to_functor C D G y \Fo y\ \\o y\ using y Fo_\o_initial initial_arrows_exist by simp show ?thesis using y F_ide arrow by auto qed next fix g :: 'd assume g: "D.arr g" show "\o (D.cod g) \\<^sub>D D.map g = FG.map g \\<^sub>D \o (D.dom g)" proof - let ?y = "D.dom g" let ?y' = "D.cod g" interpret y\: initial_arrow_to_functor C D G ?y \Fo ?y\ \\o ?y\ using g Fo_\o_initial initial_arrows_exist by simp interpret y'\: initial_arrow_to_functor C D G ?y' \Fo ?y'\ \\o ?y'\ using g Fo_\o_initial initial_arrows_exist by simp have "arrow_to_functor C D G ?y (Fo ?y') (\o ?y' \\<^sub>D g)" using g y'\.arrow by unfold_locales auto moreover have "F g = y\.the_ext (Fo ?y') (\o ?y' \\<^sub>D g)" using g F_simp by blast ultimately have "y\.is_ext (Fo ?y') (\o ?y' \\<^sub>D g) (F g)" using y\.the_ext_prop y\.is_ext_def by auto thus ?thesis using g y\.is_ext_def by simp qed qed definition \ where "\ y h = D (G h) (\.map y)" lemma \_in_hom: assumes y: "D.ide y" and f: "\f : F y \\<^sub>C x\" shows "\\ y f : y \\<^sub>D G x\" unfolding \_def using assms \.maps_ide_in_hom by auto lemma \_natural: assumes f: "\f : x \\<^sub>C x'\" and g: "\g : y' \\<^sub>D y\" and h: "\h : F y \\<^sub>C x\" shows "\ y' (f \\<^sub>C h \\<^sub>C F g) = (G f \\<^sub>D \ y h) \\<^sub>D g" proof - have "(G f \\<^sub>D \ y h) \\<^sub>D g = (G f \\<^sub>D G h \\<^sub>D \.map y) \\<^sub>D g" unfolding \_def by auto also have "... = (G f \\<^sub>D G h) \\<^sub>D \.map y \\<^sub>D g" using D.comp_assoc by fastforce also have "... = G (f \\<^sub>C h) \\<^sub>D G (F g) \\<^sub>D \.map y'" using f g h \.naturality by fastforce also have "... = (G (f \\<^sub>C h) \\<^sub>D G (F g)) \\<^sub>D \.map y'" using D.comp_assoc by fastforce also have "... = G (f \\<^sub>C h \\<^sub>C F g) \\<^sub>D \.map y'" using f g h D.comp_assoc by fastforce also have "... = \ y' (f \\<^sub>C h \\<^sub>C F g)" unfolding \_def by auto finally show ?thesis by auto qed lemma \_inverts_ext: assumes y: "D.ide y" and f: "\f : F y \\<^sub>C x\" shows "arrow_to_functor.is_ext C D G (F y) (\.map y) x (\ y f) f" proof - interpret y\: arrow_to_functor C D G y \F y\ \\.map y\ using y \.maps_ide_in_hom by unfold_locales auto show "y\.is_ext x (\ y f) f" using f y \_def y\.is_ext_def F_ide by blast qed lemma \_invertible: assumes x: "C.ide x" and g: "\g : y \\<^sub>D G x\" shows "\!f. \f : F y \\<^sub>C x\ \ \ y f = g" proof have y: "D.ide y" using g by auto interpret y\: initial_arrow_to_functor C D G y \Fo y\ \\o y\ using y initial_arrows_exist Fo_\o_initial by auto have 1: "arrow_to_functor C D G y x g" using x g by (unfold_locales, auto) let ?f = "y\.the_ext x g" have "\ y ?f = g" using \_def y\.the_ext_prop 1 F_ide x y \_inverts_ext y\.is_ext_def by fastforce moreover have "\?f : F y \\<^sub>C x\" using 1 y y\.the_ext_prop F_ide by simp ultimately show "\?f : F y \\<^sub>C x\ \ \ y ?f = g" by auto show "\f'. \f' : F y \\<^sub>C x\ \ \ y f' = g \ f' = ?f" using 1 y \_inverts_ext y\.the_ext_unique F_ide by force qed definition \ where "\ x g = (THE f. \f : F (D.dom g) \\<^sub>C x\ \ \ (D.dom g) f = g)" lemma \_in_hom: assumes "C.ide x" and "\g : y \\<^sub>D G x\" shows "C.in_hom (\ x g) (F y) x" using assms \_invertible \_def theI' [of "\f. \f : F y \\<^sub>C x\ \ \ y f = g"] by auto lemma \_\: assumes "D.ide y" and "\f : F y \\<^sub>C x\" shows "\ x (\ y f) = f" proof - have "D.dom (\ y f) = y" using assms \_in_hom by blast hence "\ x (\ y f) = (THE f'. \f' : F y \\<^sub>C x\ \ \ y f' = \ y f)" using \_def by auto moreover have "\!f'. \f' : F y \\<^sub>C x\ \ \ y f' = \ y f" using assms \_in_hom \_invertible C.ide_cod by blast ultimately show ?thesis using assms(2) by auto qed lemma \_\: assumes "C.ide x" and "\g : y \\<^sub>D G x\" shows "\ y (\ x g) = g" using assms \_invertible \_def theI' [of "\f. \f : F y \\<^sub>C x\ \ \ y f = g"] by auto theorem induces_meta_adjunction: shows "meta_adjunction C D F G \ \" using \_in_hom \_in_hom \_\ \_\ \_natural D.comp_assoc by (unfold_locales, auto) end section "Meta-Adjunctions Induce Hom-Adjunctions" text\ To obtain a hom-adjunction from a meta-adjunction, we need to exhibit hom-functors from @{term C} and @{term D} to a common set category @{term S}, so it is necessary to apply an actual concrete construction of such a category. We use the replete set category generated by the disjoint sum @{typ "('c+'d)"} of the arrow types of @{term C} and @{term D}. \ context meta_adjunction begin interpretation S: replete_setcat \undefined :: 'c+'d\ . definition inC :: "'c \ ('c+'d) setcat.arr" where "inC \ S.UP o Inl" definition inD :: "'d \ ('c+'d) setcat.arr" where "inD \ S.UP o Inr" + interpretation S: replete_setcat \undefined :: ('c+'d)\ . interpretation Cop: dual_category C .. interpretation Dop: dual_category D .. interpretation CopxC: product_category Cop.comp C .. interpretation DopxD: product_category Dop.comp D .. interpretation DopxC: product_category Dop.comp C .. - interpretation HomC: hom_functor C S.comp \\_. inC\ + interpretation HomC: hom_functor C S.comp S.setp \\_. inC\ proof show "\f. C.arr f \ inC f \ S.Univ" unfolding inC_def using S.UP_mapsto by auto + thus "\b a. \C.ide b; C.ide a\ \ inC ` C.hom b a \ S.Univ" + by blast show "\b a. \C.ide b; C.ide a\ \ inj_on inC (C.hom b a)" unfolding inC_def using S.inj_UP by (metis injD inj_Inl inj_compose inj_on_def) qed - interpretation HomD: hom_functor D S.comp \\_. inD\ + interpretation HomD: hom_functor D S.comp S.setp \\_. inD\ proof show "\f. D.arr f \ inD f \ S.Univ" unfolding inD_def using S.UP_mapsto by auto + thus "\b a. \D.ide b; D.ide a\ \ inD ` D.hom b a \ S.Univ" + by blast show "\b a. \D.ide b; D.ide a\ \ inj_on inD (D.hom b a)" unfolding inD_def using S.inj_UP by (metis injD inj_Inr inj_compose inj_on_def) qed interpretation Fop: dual_functor D C F .. interpretation FopxC: product_functor Dop.comp C Cop.comp C Fop.map C.map .. interpretation DopxG: product_functor Dop.comp C Dop.comp D Dop.map G .. interpretation Hom_FopxC: composite_functor DopxC.comp CopxC.comp S.comp FopxC.map HomC.map .. interpretation Hom_DopxG: composite_functor DopxC.comp DopxD.comp S.comp DopxG.map HomD.map .. lemma inC_\ [simp]: assumes "C.ide b" and "C.ide a" and "x \ inC ` C.hom b a" shows "inC (HomC.\ (b, a) x) = x" using assms by auto lemma \_inC [simp]: assumes "C.arr f" shows "HomC.\ (C.dom f, C.cod f) (inC f) = f" using assms HomC.\_\ by blast lemma inD_\ [simp]: assumes "D.ide b" and "D.ide a" and "x \ inD ` D.hom b a" shows "inD (HomD.\ (b, a) x) = x" using assms by auto lemma \_inD [simp]: assumes "D.arr f" shows "HomD.\ (D.dom f, D.cod f) (inD f) = f" using assms HomD.\_\ by blast lemma Hom_FopxC_simp: assumes "DopxC.arr gf" shows "Hom_FopxC.map gf = S.mkArr (HomC.set (F (D.cod (fst gf)), C.dom (snd gf))) (HomC.set (F (D.dom (fst gf)), C.cod (snd gf))) (inC \ (\h. snd gf \\<^sub>C h \\<^sub>C F (fst gf)) \ HomC.\ (F (D.cod (fst gf)), C.dom (snd gf)))" using assms HomC.map_def by simp lemma Hom_DopxG_simp: assumes "DopxC.arr gf" shows "Hom_DopxG.map gf = S.mkArr (HomD.set (D.cod (fst gf), G (C.dom (snd gf)))) (HomD.set (D.dom (fst gf), G (C.cod (snd gf)))) (inD \ (\h. G (snd gf) \\<^sub>D h \\<^sub>D fst gf) \ HomD.\ (D.cod (fst gf), G (C.dom (snd gf))))" using assms HomD.map_def by simp definition \o where "\o yx = S.mkArr (HomC.set (F (fst yx), snd yx)) (HomD.set (fst yx, G (snd yx))) (inD o \ (fst yx) o HomC.\ (F (fst yx), snd yx))" lemma \o_in_hom: assumes yx: "DopxC.ide yx" shows "\\o yx : Hom_FopxC.map yx \\<^sub>S Hom_DopxG.map yx\" proof - have "Hom_FopxC.map yx = S.mkIde (HomC.set (F (fst yx), snd yx))" using yx HomC.map_ide by auto moreover have "Hom_DopxG.map yx = S.mkIde (HomD.set (fst yx, G (snd yx)))" using yx HomD.map_ide by auto moreover have "\S.mkArr (HomC.set (F (fst yx), snd yx)) (HomD.set (fst yx, G (snd yx))) (inD \ \ (fst yx) \ HomC.\ (F (fst yx), snd yx)) : S.mkIde (HomC.set (F (fst yx), snd yx)) \\<^sub>S S.mkIde (HomD.set (fst yx, G (snd yx)))\" proof (intro S.mkArr_in_hom) show "HomC.set (F (fst yx), snd yx) \ S.Univ" using yx HomC.set_subset_Univ by simp show "HomD.set (fst yx, G (snd yx)) \ S.Univ" using yx HomD.set_subset_Univ by simp show "inD o \ (fst yx) o HomC.\ (F (fst yx), snd yx) \ HomC.set (F (fst yx), snd yx) \ HomD.set (fst yx, G (snd yx))" proof fix x assume x: "x \ HomC.set (F (fst yx), snd yx)" show "(inD o \ (fst yx) o HomC.\ (F (fst yx), snd yx)) x \ HomD.set (fst yx, G (snd yx))" using x yx HomC.\_mapsto [of "F (fst yx)" "snd yx"] \_in_hom [of "fst yx"] HomD.\_mapsto [of "fst yx" "G (snd yx)"] by auto qed qed ultimately show ?thesis using \o_def by auto qed interpretation \: transformation_by_components DopxC.comp S.comp Hom_FopxC.map Hom_DopxG.map \o proof fix yx assume yx: "DopxC.ide yx" show "\\o yx : Hom_FopxC.map yx \\<^sub>S Hom_DopxG.map yx\" using yx \o_in_hom by auto next fix gf assume gf: "DopxC.arr gf" show "S.comp (\o (DopxC.cod gf)) (Hom_FopxC.map gf) = S.comp (Hom_DopxG.map gf) (\o (DopxC.dom gf))" proof - let ?g = "fst gf" let ?f = "snd gf" let ?x = "C.dom ?f" let ?x' = "C.cod ?f" let ?y = "D.cod ?g" let ?y' = "D.dom ?g" let ?Fy = "F ?y" let ?Fy' = "F ?y'" let ?Fg = "F ?g" let ?Gx = "G ?x" let ?Gx' = "G ?x'" let ?Gf = "G ?f" have 1: "S.arr (Hom_FopxC.map gf) \ Hom_FopxC.map gf = S.mkArr (HomC.set (?Fy, ?x)) (HomC.set (?Fy', ?x')) (inC o (\h. ?f \\<^sub>C h \\<^sub>C ?Fg) o HomC.\ (?Fy, ?x))" using gf Hom_FopxC.preserves_arr Hom_FopxC_simp by blast have 2: "S.arr (\o (DopxC.cod gf)) \ \o (DopxC.cod gf) = S.mkArr (HomC.set (?Fy', ?x')) (HomD.set (?y', ?Gx')) (inD o \ ?y' o HomC.\ (?Fy', ?x'))" using gf \o_in_hom [of "DopxC.cod gf"] \o_def [of "DopxC.cod gf"] \_in_hom - S.card_of_leq by auto have 3: "S.arr (\o (DopxC.dom gf)) \ \o (DopxC.dom gf) = S.mkArr (HomC.set (?Fy, ?x)) (HomD.set (?y, ?Gx)) (inD o \ ?y o HomC.\ (?Fy, ?x))" using gf \o_in_hom [of "DopxC.dom gf"] \o_def [of "DopxC.dom gf"] \_in_hom - S.card_of_leq by auto have 4: "S.arr (Hom_DopxG.map gf) \ Hom_DopxG.map gf = S.mkArr (HomD.set (?y, ?Gx)) (HomD.set (?y', ?Gx')) (inD o (\h. ?Gf \\<^sub>D h \\<^sub>D ?g) o HomD.\ (?y, ?Gx))" using gf Hom_DopxG.preserves_arr Hom_DopxG_simp by blast have 5: "S.seq (\o (DopxC.cod gf)) (Hom_FopxC.map gf) \ S.comp (\o (DopxC.cod gf)) (Hom_FopxC.map gf) = S.mkArr (HomC.set (?Fy, ?x)) (HomD.set (?y', ?Gx')) ((inD o \ ?y' o HomC.\ (?Fy', ?x')) o (inC o (\h. ?f \\<^sub>C h \\<^sub>C ?Fg) o HomC.\ (?Fy, ?x)))" by (metis gf 1 2 DopxC.arr_iff_in_hom DopxC.ide_cod Hom_FopxC.preserves_hom S.comp_mkArr S.seqI' \o_in_hom) have 6: "S.comp (Hom_DopxG.map gf) (\o (DopxC.dom gf)) = S.mkArr (HomC.set (?Fy, ?x)) (HomD.set (?y', ?Gx')) ((inD o (\h. ?Gf \\<^sub>D h \\<^sub>D ?g) o HomD.\ (?y, ?Gx)) o (inD o \ ?y o HomC.\ (?Fy, ?x)))" by (metis 3 4 S.comp_mkArr) have 7: "restrict ((inD o \ ?y' o HomC.\ (?Fy', ?x')) o (inC o (\h. ?f \\<^sub>C h \\<^sub>C ?Fg) o HomC.\ (?Fy, ?x))) (HomC.set (?Fy, ?x)) = restrict ((inD o (\h. ?Gf \\<^sub>D h \\<^sub>D ?g) o HomD.\ (?y, ?Gx)) o (inD o \ ?y o HomC.\ (?Fy, ?x))) (HomC.set (?Fy, ?x))" proof (intro restrict_ext) show "\h. h \ HomC.set (?Fy, ?x) \ ((inD o \ ?y' o HomC.\ (?Fy', ?x')) o (inC o (\h. ?f \\<^sub>C h \\<^sub>C ?Fg) o HomC.\ (?Fy, ?x))) h = ((inD o (\h. ?Gf \\<^sub>D h \\<^sub>D ?g) o HomD.\ (?y, ?Gx)) o (inD o \ ?y o HomC.\ (?Fy, ?x))) h" proof - fix h assume h: "h \ HomC.set (?Fy, ?x)" have \h: "\HomC.\ (?Fy, ?x) h : ?Fy \\<^sub>C ?x\" using gf h HomC.\_mapsto [of ?Fy ?x] CopxC.ide_char by auto show "((inD o \ ?y' o HomC.\ (?Fy', ?x')) o (inC o (\h. ?f \\<^sub>C h \\<^sub>C ?Fg) o HomC.\ (?Fy, ?x))) h = ((inD o (\h. ?Gf \\<^sub>D h \\<^sub>D ?g) o HomD.\ (?y, ?Gx)) o (inD o \ ?y o HomC.\ (?Fy, ?x))) h" proof - have "((inD o \ ?y' o HomC.\ (?Fy', ?x')) o (inC o (\h. ?f \\<^sub>C h \\<^sub>C ?Fg) o HomC.\ (?Fy, ?x))) h = inD (\ ?y' (?f \\<^sub>C HomC.\ (?Fy, ?x) h \\<^sub>C ?Fg))" using gf \h HomC.\_mapsto HomC.\_mapsto \_in_hom \_inC [of "?f \\<^sub>C HomC.\ (?Fy, ?x) h \\<^sub>C ?Fg"] by auto also have "... = inD (D ?Gf (D (\ ?y (HomC.\ (?Fy, ?x) h)) ?g))" by (metis (no_types, lifting) C.arr_cod C.arr_dom_iff_arr C.arr_iff_in_hom C.in_homE D.arr_cod_iff_arr D.arr_iff_in_hom F.preserves_reflects_arr \_naturality \h) also have "... = ((inD o (\h. ?Gf \\<^sub>D h \\<^sub>D ?g) o HomD.\ (?y, ?Gx)) o (inD o \ ?y o HomC.\ (?Fy, ?x))) h" using gf \h \_in_hom by simp finally show ?thesis by auto qed qed qed have 8: "S.mkArr (HomC.set (?Fy, ?x)) (HomD.set (?y', ?Gx')) ((inD o \ ?y' o HomC.\ (?Fy', ?x')) o (inC o (\h. ?f \\<^sub>C h \\<^sub>C ?Fg) o HomC.\ (?Fy, ?x))) = S.mkArr (HomC.set (?Fy, ?x)) (HomD.set (?y', ?Gx')) ((inD o (\h. ?Gf \\<^sub>D h \\<^sub>D ?g) o HomD.\ (?y, ?Gx)) o (inD o \ ?y o HomC.\ (?Fy, ?x)))" using 5 7 by force show ?thesis using 5 6 8 by auto qed qed lemma \_simp: assumes YX: "DopxC.ide yx" shows "\.map yx = S.mkArr (HomC.set (F (fst yx), snd yx)) (HomD.set (fst yx, G (snd yx))) (inD o \ (fst yx) o HomC.\ (F (fst yx), snd yx))" using YX \o_def by simp abbreviation \o where "\o yx \ S.mkArr (HomD.set (fst yx, G (snd yx))) (HomC.set (F (fst yx), snd yx)) (inC o \ (snd yx) o HomD.\ (fst yx, G (snd yx)))" lemma \o_in_hom: assumes yx: "DopxC.ide yx" shows "\\o yx : Hom_DopxG.map yx \\<^sub>S Hom_FopxC.map yx\" proof - have "Hom_FopxC.map yx = S.mkIde (HomC.set (F (fst yx), snd yx))" using yx HomC.map_ide by auto moreover have "Hom_DopxG.map yx = S.mkIde (HomD.set (fst yx, G (snd yx)))" using yx HomD.map_ide by auto moreover have "\\o yx : S.mkIde (HomD.set (fst yx, G (snd yx))) \\<^sub>S S.mkIde (HomC.set (F (fst yx), snd yx))\" proof (intro S.mkArr_in_hom) show "HomC.set (F (fst yx), snd yx) \ S.Univ" using yx HomC.set_subset_Univ by simp show "HomD.set (fst yx, G (snd yx)) \ S.Univ" using yx HomD.set_subset_Univ by simp show "inC o \ (snd yx) o HomD.\ (fst yx, G (snd yx)) \ HomD.set (fst yx, G (snd yx)) \ HomC.set (F (fst yx), snd yx)" proof fix x assume x: "x \ HomD.set (fst yx, G (snd yx))" show "(inC o \ (snd yx) o HomD.\ (fst yx, G (snd yx))) x \ HomC.set (F (fst yx), snd yx)" using x yx HomD.\_mapsto [of "fst yx" "G (snd yx)"] \_in_hom [of "snd yx"] HomC.\_mapsto [of "F (fst yx)" "snd yx"] by auto qed qed ultimately show ?thesis by auto qed lemma \_inv: assumes yx: "DopxC.ide yx" shows "S.inverse_arrows (\.map yx) (\o yx)" proof - have 1: "\\.map yx : Hom_FopxC.map yx \\<^sub>S Hom_DopxG.map yx\" using yx \.preserves_hom [of yx yx yx] DopxC.ide_in_hom by blast have 2: "\\o yx : Hom_DopxG.map yx \\<^sub>S Hom_FopxC.map yx\" using yx \o_in_hom by simp have 3: "\.map yx = S.mkArr (HomC.set (F (fst yx), snd yx)) (HomD.set (fst yx, G (snd yx))) (inD o \ (fst yx) o HomC.\ (F (fst yx), snd yx))" using yx \_simp by blast have antipar: "S.antipar (\.map yx) (\o yx)" using 1 2 by blast moreover have "S.ide (S.comp (\o yx) (\.map yx))" proof - have "S.comp (\o yx) (\.map yx) = S.mkArr (HomC.set (F (fst yx), snd yx)) (HomC.set (F (fst yx), snd yx)) ((inC o \ (snd yx) o HomD.\ (fst yx, G (snd yx))) o (inD o \ (fst yx) o HomC.\ (F (fst yx), snd yx)))" using 1 2 3 antipar S.comp_mkArr by auto also have "... = S.mkArr (HomC.set (F (fst yx), snd yx)) (HomC.set (F (fst yx), snd yx)) (\x. x)" proof - have "S.mkArr (HomC.set (F (fst yx), snd yx)) (HomC.set (F (fst yx), snd yx)) (\x. x) = ..." proof show "S.arr (S.mkArr (HomC.set (F (fst yx), snd yx)) (HomC.set (F (fst yx), snd yx)) (\x. x))" - using yx HomC.set_subset_Univ S.arr_mkArr by simp + using yx HomC.set_subset_Univ by simp show "\x. x \ HomC.set (F (fst yx), snd yx) \ x = ((inC o \ (snd yx) o HomD.\ (fst yx, G (snd yx))) o (inD o \ (fst yx) o HomC.\ (F (fst yx), snd yx))) x" proof - fix x assume x: "x \ HomC.set (F (fst yx), snd yx)" have "((inC o \ (snd yx) o HomD.\ (fst yx, G (snd yx))) o (inD o \ (fst yx) o HomC.\ (F (fst yx), snd yx))) x = inC (\ (snd yx) (HomD.\ (fst yx, G (snd yx)) (inD (\ (fst yx) (HomC.\ (F (fst yx), snd yx) x)))))" by simp also have "... = inC (\ (snd yx) (\ (fst yx) (HomC.\ (F (fst yx), snd yx) x)))" using x yx HomC.\_mapsto [of "F (fst yx)" "snd yx"] \_in_hom by force also have "... = inC (HomC.\ (F (fst yx), snd yx) x)" using x yx HomC.\_mapsto [of "F (fst yx)" "snd yx"] \_\ by force also have "... = x" using x yx inC_\ by simp finally show "x = ((inC o \ (snd yx) o HomD.\ (fst yx, G (snd yx))) o (inD o \ (fst yx) o HomC.\ (F (fst yx), snd yx))) x" by auto qed qed thus ?thesis by auto qed also have "... = S.mkIde (HomC.set (F (fst yx), snd yx))" using yx S.mkIde_as_mkArr HomC.set_subset_Univ by force finally have "S.comp (\o yx) (\.map yx) = S.mkIde (HomC.set (F (fst yx), snd yx))" by auto thus ?thesis using yx HomC.set_subset_Univ S.ide_mkIde by simp qed moreover have "S.ide (S.comp (\.map yx) (\o yx))" proof - have "S.comp (\.map yx) (\o yx) = S.mkArr (HomD.set (fst yx, G (snd yx))) (HomD.set (fst yx, G (snd yx))) ((inD o \ (fst yx) o HomC.\ (F (fst yx), snd yx)) o (inC o \ (snd yx) o HomD.\ (fst yx, G (snd yx))))" using 1 2 3 S.comp_mkArr antipar by fastforce also have "... = S.mkArr (HomD.set (fst yx, G (snd yx))) (HomD.set (fst yx, G (snd yx))) (\x. x)" proof - have "S.mkArr (HomD.set (fst yx, G (snd yx))) (HomD.set (fst yx, G (snd yx))) (\x. x) = ..." proof show "S.arr (S.mkArr (HomD.set (fst yx, G (snd yx))) (HomD.set (fst yx, G (snd yx))) (\x. x))" - using yx HomD.set_subset_Univ S.arr_mkArr by simp + using yx HomD.set_subset_Univ by simp show "\x. x \ (HomD.set (fst yx, G (snd yx))) \ x = ((inD o \ (fst yx) o HomC.\ (F (fst yx), snd yx)) o (inC o \ (snd yx) o HomD.\ (fst yx, G (snd yx)))) x" proof - fix x assume x: "x \ HomD.set (fst yx, G (snd yx))" have "((inD o \ (fst yx) o HomC.\ (F (fst yx), snd yx)) o (inC o \ (snd yx) o HomD.\ (fst yx, G (snd yx)))) x = inD (\ (fst yx) (HomC.\ (F (fst yx), snd yx) (inC (\ (snd yx) (HomD.\ (fst yx, G (snd yx)) x)))))" by simp also have "... = inD (\ (fst yx) (\ (snd yx) (HomD.\ (fst yx, G (snd yx)) x)))" proof - have "\\ (snd yx) (HomD.\ (fst yx, G (snd yx)) x) : F (fst yx) \ snd yx\" using x yx HomD.\_mapsto [of "fst yx" "G (snd yx)"] \_in_hom by auto thus ?thesis by simp qed also have "... = inD (HomD.\ (fst yx, G (snd yx)) x)" using x yx HomD.\_mapsto [of "fst yx" "G (snd yx)"] \_\ by force also have "... = x" using x yx inD_\ by simp finally show "x = ((inD o \ (fst yx) o HomC.\ (F (fst yx), snd yx)) o (inC o \ (snd yx) o HomD.\ (fst yx, G (snd yx)))) x" by auto qed qed thus ?thesis by auto qed also have "... = S.mkIde (HomD.set (fst yx, G (snd yx)))" using yx S.mkIde_as_mkArr HomD.set_subset_Univ by force finally have "S.comp (\.map yx) (\o yx) = S.mkIde (HomD.set (fst yx, G (snd yx)))" by auto thus ?thesis using yx HomD.set_subset_Univ S.ide_mkIde by simp qed ultimately show ?thesis by auto qed interpretation \: natural_isomorphism DopxC.comp S.comp Hom_FopxC.map Hom_DopxG.map \.map using \_inv by unfold_locales blast interpretation \: inverse_transformation DopxC.comp S.comp Hom_FopxC.map Hom_DopxG.map \.map .. interpretation \\: inverse_transformations DopxC.comp S.comp Hom_FopxC.map Hom_DopxG.map \.map \.map using \.inverts_components by unfold_locales simp abbreviation \ where "\ \ \.map" abbreviation \ where "\ \ \.map" abbreviation HomC where "HomC \ HomC.map" abbreviation \C where "\C \ \_. inC" abbreviation HomD where "HomD \ HomD.map" abbreviation \D where "\D \ \_. inD" - theorem induces_hom_adjunction: "hom_adjunction C D S.comp \C \D F G \ \" + theorem induces_hom_adjunction: "hom_adjunction C D S.comp S.setp \C \D F G \ \" using F.is_extensional by unfold_locales auto lemma \_simp: assumes yx: "DopxC.ide yx" shows "\ yx = S.mkArr (HomD.set (fst yx, G (snd yx))) (HomC.set (F (fst yx), snd yx)) (inC o \ (snd yx) o HomD.\ (fst yx, G (snd yx)))" using assms \o_def \_inv S.inverse_unique by simp text\ The original @{term \} and @{term \} can be recovered from @{term \} and @{term \}. \ - interpretation \: set_valued_transformation DopxC.comp S.comp + interpretation \: set_valued_transformation DopxC.comp S.comp S.setp Hom_FopxC.map Hom_DopxG.map \.map .. - interpretation \: set_valued_transformation DopxC.comp S.comp + interpretation \: set_valued_transformation DopxC.comp S.comp S.setp Hom_DopxG.map Hom_FopxC.map \.map .. lemma \_in_terms_of_\': assumes y: "D.ide y" and f: "\f: F y \\<^sub>C x\" shows "\ y f = (HomD.\ (y, G x) o \.FUN (y, x) o inC) f" proof - have x: "C.ide x" using f by auto have "(HomD.\ (y, G x) o \.FUN (y, x) o inC) f = HomD.\ (y, G x) (restrict (inD o \ y o HomC.\ (F y, x)) (HomC.set (F y, x)) (inC f))" proof - have "S.arr (\ (y, x))" using x y by fastforce thus ?thesis using x y \o_def by simp qed also have "... = \ y f" using x y f HomC.\_mapsto \_in_hom HomC.\_mapsto C.ide_in_hom D.ide_in_hom by auto finally show ?thesis by auto qed lemma \_in_terms_of_\': assumes x: "C.ide x" and g: "\g : y \\<^sub>D G x\" shows "\ x g = (HomC.\ (F y, x) o \.FUN (y, x) o inD) g" proof - have y: "D.ide y" using g by auto have "(HomC.\ (F y, x) o \.FUN (y, x) o inD) g = HomC.\ (F y, x) (restrict (inC o \ x o HomD.\ (y, G x)) (HomD.set (y, G x)) (inD g))" proof - have "S.arr (\ (y, x))" using x y \.preserves_reflects_arr [of "(y, x)"] by simp thus ?thesis using x y \_simp by simp qed also have "... = \ x g" using x y g HomD.\_mapsto \_in_hom HomD.\_mapsto C.ide_in_hom D.ide_in_hom by auto finally show ?thesis by auto qed end section "Hom-Adjunctions Induce Meta-Adjunctions" context hom_adjunction begin definition \ :: "'d \ 'c \ 'd" where "\ y h = (HomD.\ (y, G (C.cod h)) o \.FUN (y, C.cod h) o \C (F y, C.cod h)) h" definition \ :: "'c \ 'd \ 'c" where "\ x h = (HomC.\ (F (D.dom h), x) o \.FUN (D.dom h, x) o \D (D.dom h, G x)) h" lemma Hom_FopxC_map_simp: assumes "DopxC.arr gf" shows "Hom_FopxC.map gf = S.mkArr (HomC.set (F (D.cod (fst gf)), C.dom (snd gf))) (HomC.set (F (D.dom (fst gf)), C.cod (snd gf))) (\C (F (D.dom (fst gf)), C.cod (snd gf)) o (\h. snd gf \\<^sub>C h \\<^sub>C F (fst gf)) o HomC.\ (F (D.cod (fst gf)), C.dom (snd gf)))" using assms HomC.map_def by simp lemma Hom_DopxG_map_simp: assumes "DopxC.arr gf" shows "Hom_DopxG.map gf = S.mkArr (HomD.set (D.cod (fst gf), G (C.dom (snd gf)))) (HomD.set (D.dom (fst gf), G (C.cod (snd gf)))) (\D (D.dom (fst gf), G (C.cod (snd gf))) o (\h. G (snd gf) \\<^sub>D h \\<^sub>D fst gf) o HomD.\ (D.cod (fst gf), G (C.dom (snd gf))))" using assms HomD.map_def by simp lemma \_Fun_mapsto: assumes "D.ide y" and "\f : F y \\<^sub>C x\" shows "\.FUN (y, x) \ HomC.set (F y, x) \ HomD.set (y, G x)" proof - have "S.arr (\ (y, x)) \ \.DOM (y, x) = HomC.set (F y, x) \ \.COD (y, x) = HomD.set (y, G x)" using assms HomC.set_map HomD.set_map by auto thus ?thesis using S.Fun_mapsto by blast qed lemma \_mapsto: assumes y: "D.ide y" shows "\ y \ C.hom (F y) x \ D.hom y (G x)" proof fix h assume h: "h \ C.hom (F y) x" hence 1: " \h : F y \\<^sub>C x\" by simp show "\ y h \ D.hom y (G x)" proof - have "\C (F y, x) h \ HomC.set (F y, x)" using y h 1 HomC.\_mapsto [of "F y" x] by fastforce hence "\.FUN (y, x) (\C (F y, x) h) \ HomD.set (y, G x)" using h y \_Fun_mapsto by auto thus ?thesis using y h 1 \_def HomC.\_mapsto HomD.\_mapsto [of y "G x"] by fastforce qed qed lemma \_simp: assumes "D.ide y" and "C.ide x" shows "S.arr (\ (y, x))" and "\ (y, x) = S.mkArr (HomC.set (F y, x)) (HomD.set (y, G x)) (\D (y, G x) o \ y o \C (F y, x))" proof - show 1: "S.arr (\ (y, x))" using assms by auto hence "\ (y, x) = S.mkArr (\.DOM (y, x)) (\.COD (y, x)) (\.FUN (y, x))" using S.mkArr_Fun by metis also have "... = S.mkArr (HomC.set (F y, x)) (HomD.set (y, G x)) (\.FUN (y, x))" using assms HomC.set_map HomD.set_map by fastforce also have "... = S.mkArr (HomC.set (F y, x)) (HomD.set (y, G x)) (\D (y, G x) o \ y o \C (F y, x))" proof (intro S.mkArr_eqI') show 2: "S.arr (S.mkArr (HomC.set (F y, x)) (HomD.set (y, G x)) (\.FUN (y, x)))" using 1 calculation by argo show "\h. h \ HomC.set (F y, x) \ \.FUN (y, x) h = (\D (y, G x) o \ y o \C (F y, x)) h" proof - fix h assume h: "h \ HomC.set (F y, x)" have "(\D (y, G x) o \ y o HomC.\ (F y, x)) h = \D (y, G x) (\D (y, G x) (\.FUN (y, x) (\C (F y, x) (\C (F y, x) h))))" proof - have "\\C (F y, x) h : F y \\<^sub>C x\" using assms h HomC.\_mapsto [of "F y" x] by auto thus ?thesis using h \_def by auto qed also have "... = \D (y, G x) (\D (y, G x) (\.FUN (y, x) h))" using assms h HomC.\_\ \_Fun_mapsto by simp also have "... = \.FUN (y, x) h" using assms h \_Fun_mapsto [of y "\C (F y, x) h"] HomC.\_mapsto HomD.\_\ [of y "G x"] C.ide_in_hom D.ide_in_hom - by (meson 2 G.preserves_ide S.arr_mkArr funcset_mem) + by blast finally show "\.FUN (y, x) h = (\D (y, G x) o \ y o \C (F y, x)) h" by auto qed qed finally show "\ (y, x) = S.mkArr (HomC.set (F y, x)) (HomD.set (y, G x)) (\D (y, G x) o \ y o \C (F y, x))" by force qed lemma \_Fun_mapsto: assumes "C.ide x" and "\g : y \\<^sub>D G x\" shows "\.FUN (y, x) \ HomD.set (y, G x) \ HomC.set (F y, x)" proof - have "S.arr (\ (y, x)) \ \.COD (y, x) = HomC.set (F y, x) \ \.DOM (y, x) = HomD.set (y, G x)" using assms HomC.set_map HomD.set_map by auto thus ?thesis using S.Fun_mapsto by fast qed lemma \_mapsto: assumes x: "C.ide x" shows "\ x \ D.hom y (G x) \ C.hom (F y) x" proof fix h assume h: "h \ D.hom y (G x)" hence 1: "\h : y \\<^sub>D G x\" by auto show "\ x h \ C.hom (F y) x" proof - have "\.FUN (y, x) (\D (y, G x) h) \ HomC.set (F y, x)" proof - have "\D (y, G x) h \ HomD.set (y, G x)" using x h 1 HomD.\_mapsto [of y "G x"] by fastforce thus ?thesis using h x \_Fun_mapsto by auto qed thus ?thesis using x h 1 \_def HomD.\_mapsto HomC.\_mapsto [of "F y" x] by fastforce qed qed lemma \_simp: assumes "D.ide y" and "C.ide x" shows "S.arr (\ (y, x))" and "\ (y, x) = S.mkArr (HomD.set (y, G x)) (HomC.set (F y, x)) (\C (F y, x) o \ x o \D (y, G x))" proof - show 1: "S.arr (\ (y, x))" using assms by auto hence "\ (y, x) = S.mkArr (\.DOM (y, x)) (\.COD (y, x)) (\.FUN (y, x))" using S.mkArr_Fun by metis also have "... = S.mkArr (HomD.set (y, G x)) (HomC.set (F y, x)) (\.FUN (y, x))" using assms HomC.set_map HomD.set_map by auto also have "... = S.mkArr (HomD.set (y, G x)) (HomC.set (F y, x)) (\C (F y, x) o \ x o \D (y, G x))" proof (intro S.mkArr_eqI') show "S.arr (S.mkArr (HomD.set (y, G x)) (HomC.set (F y, x)) (\.FUN (y, x)))" using 1 calculation by argo show "\h. h \ HomD.set (y, G x) \ \.FUN (y, x) h = (\C (F y, x) o \ x o \D (y, G x)) h" proof - fix h assume h: "h \ HomD.set (y, G x)" have "(\C (F y, x) o \ x o HomD.\ (y, G x)) h = \C (F y, x) (\C (F y, x) (\.FUN (y, x) (\D (y, G x) (\D (y, G x) h))))" proof - have "\\D (y, G x) h : y \\<^sub>D G x\" using assms h HomD.\_mapsto [of y "G x"] by auto thus ?thesis using h \_def by auto qed also have "... = \C (F y, x) (\C (F y, x) (\.FUN (y, x) h))" using assms h HomD.\_\ \_Fun_mapsto by simp also have "... = \.FUN (y, x) h" using assms h \_Fun_mapsto HomD.\_mapsto [of y "G x"] HomC.\_\ [of "F y" x] C.ide_in_hom D.ide_in_hom by blast finally show "\.FUN (y, x) h = (\C (F y, x) o \ x o HomD.\ (y, G x)) h" by auto qed qed finally show "\ (y, x) = S.mkArr (HomD.set (y, G x)) (HomC.set (F y, x)) (\C (F y, x) o \ x o \D (y, G x))" by force qed text\ The length of the next proof stems from having to use properties of composition of arrows in @{term[source=true] S} to infer properties of the composition of the corresponding functions. \ interpretation \\: meta_adjunction C D F G \ \ proof fix y :: 'd and x :: 'c and h :: 'c assume y: "D.ide y" and h: "\h : F y \\<^sub>C x\" have x: "C.ide x" using h by auto show "\\ y h : y \\<^sub>D G x\" proof - have "\.FUN (y, x) \ HomC.set (F y, x) \ HomD.set (y, G x)" using y h \_Fun_mapsto by blast thus ?thesis using x y h \_def HomD.\_mapsto [of y "G x"] HomC.\_mapsto [of "F y" x] by auto qed show "\ x (\ y h) = h" proof - have 0: "restrict (\h. h) (HomC.set (F y, x)) = restrict (\C (F y, x) o (\ x o \ y) o \C (F y, x)) (HomC.set (F y, x))" proof - have 1: "S.ide (\ (y, x) \\<^sub>S \ (y, x))" using x y \\.inv [of "(y, x)"] by auto hence 6: "S.seq (\ (y, x)) (\ (y, x))" by auto have 2: "\ (y, x) = S.mkArr (HomC.set (F y, x)) (HomD.set (y, G x)) (\D (y, G x) o \ y o \C (F y, x)) \ \ (y, x) = S.mkArr (HomD.set (y, G x)) (HomC.set (F y, x)) (\C (F y, x) o \ x o \D (y, G x))" using x y \_simp \_simp by force have 3: "S (\ (y, x)) (\ (y, x)) = S.mkArr (HomC.set (F y, x)) (HomC.set (F y, x)) (\C (F y, x) o (\ x o \ y) o \C (F y, x))" proof - have 4: "S.arr (\ (y, x) \\<^sub>S \ (y, x))" using 1 by auto hence "S (\ (y, x)) (\ (y, x)) = S.mkArr (HomC.set (F y, x)) (HomC.set (F y, x)) ((\C (F y, x) o \ x o \D (y, G x)) o (\D (y, G x) o \ y o \C (F y, x)))" using 1 2 S.ide_in_hom S.comp_mkArr by fastforce also have "... = S.mkArr (HomC.set (F y, x)) (HomC.set (F y, x)) (\C (F y, x) o (\ x o \ y) o \C (F y, x))" proof (intro S.mkArr_eqI') show "S.arr (S.mkArr (HomC.set (F y, x)) (HomC.set (F y, x)) ((\C (F y, x) o \ x o \D (y, G x)) o (\D (y, G x) o \ y o \C (F y, x))))" using 4 calculation by simp show "\h. h \ HomC.set (F y, x) \ ((\C (F y, x) o \ x o \D (y, G x)) o (\D (y, G x) o \ y o \C (F y, x))) h = (\C (F y, x) o (\ x o \ y) o \C (F y, x)) h" proof - fix h assume h: "h \ HomC.set (F y, x)" hence "\\ y (\C (F y, x) h) : y \\<^sub>D G x\" using x y h HomC.\_mapsto [of "F y" x] \_mapsto by auto thus "((\C (F y, x) o \ x o \D (y, G x)) o (\D (y, G x) o \ y o \C (F y, x))) h = (\C (F y, x) o (\ x o \ y) o \C (F y, x)) h" using x y 1 \_mapsto HomD.\_\ by simp qed qed finally show ?thesis by simp qed moreover have "\ (y, x) \\<^sub>S \ (y, x) = S.mkArr (HomC.set (F y, x)) (HomC.set (F y, x)) (\h. h)" using 1 2 6 calculation S.mkIde_as_mkArr S.arr_mkArr S.dom_mkArr S.ideD(2) by metis ultimately have 4: "S.mkArr (HomC.set (F y, x)) (HomC.set (F y, x)) (\C (F y, x) o (\ x o \ y) o \C (F y, x)) = S.mkArr (HomC.set (F y, x)) (HomC.set (F y, x)) (\h. h)" by auto have 5: "S.arr (S.mkArr (HomC.set (F y, x)) (HomC.set (F y, x)) (\C (F y, x) o (\ x o \ y) o \C (F y, x)))" using 1 3 6 by presburger hence "restrict (\C (F y, x) o (\ x o \ y) o \C (F y, x)) (HomC.set (F y, x)) = S.Fun (S.mkArr (HomC.set (F y, x)) (HomC.set (F y, x)) (\C (F y, x) o (\ x o \ y) o \C (F y, x)))" by auto also have "... = restrict (\h. h) (HomC.set (F y, x))" using 4 5 by auto finally show ?thesis by auto qed moreover have "\C (F y, x) h \ HomC.set (F y, x)" using x y h HomC.\_mapsto [of "F y" x] by auto ultimately have "\C (F y, x) h = (\C (F y, x) o (\ x o \ y) o \C (F y, x)) (\C (F y, x) h)" using x y h HomC.\_mapsto [of "F y" x] by fast hence "\C (F y, x) (\C (F y, x) h) = \C (F y, x) ((\C (F y, x) o (\ x o \ y) o \C (F y, x)) (\C (F y, x) h))" by simp hence "h = \C (F y, x) (\C (F y, x) (\ x (\ y (\C (F y, x) (\C (F y, x) h)))))" using x y h HomC.\_\ [of "F y" x] by simp also have "... = \ x (\ y h)" using x y h HomC.\_\ HomC.\_\ \_mapsto \_mapsto by (metis PiE mem_Collect_eq) finally show ?thesis by auto qed next fix x :: 'c and h :: 'd and y :: 'd assume x: "C.ide x" and h: "\h : y \\<^sub>D G x\" have y: "D.ide y" using h by auto show "\\ x h : F y \\<^sub>C x\" using x y h \_mapsto [of x y] by auto show "\ y (\ x h) = h" proof - have 0: "restrict (\h. h) (HomD.set (y, G x)) = restrict (\D (y, G x) o (\ y o \ x) o \D (y, G x)) (HomD.set (y, G x))" proof - have 1: "S.ide (S (\ (y, x)) (\ (y, x)))" using x y \\.inv by force hence 6: "S.seq (\ (y, x)) (\ (y, x))" by auto have 2: "\ (y, x) = S.mkArr (HomC.set (F y, x)) (HomD.set (y, G x)) (\D (y, G x) o \ y o \C (F y, x)) \ \ (y, x) = S.mkArr (HomD.set (y, G x)) (HomC.set (F y, x)) (\C (F y, x) o \ x o \D (y, G x))" using x h \_simp \_simp by auto have 3: "S (\ (y, x)) (\ (y, x)) = S.mkArr (HomD.set (y, G x)) (HomD.set (y, G x)) (\D (y, G x) o (\ y o \ x) o \D (y, G x))" proof - have 4: "S.seq (\ (y, x)) (\ (y, x))" using 1 by auto hence "S (\ (y, x)) (\ (y, x)) = S.mkArr (HomD.set (y, G x)) (HomD.set (y, G x)) ((\D (y, G x) o \ y o \C (F y, x)) o (\C (F y, x) o \ x o \D (y, G x)))" using 1 2 6 S.ide_in_hom S.comp_mkArr by fastforce also have "... = S.mkArr (HomD.set (y, G x)) (HomD.set (y, G x)) (\D (y, G x) o (\ y o \ x) o \D (y, G x))" proof show "S.arr (S.mkArr (HomD.set (y, G x)) (HomD.set (y, G x)) ((\D (y, G x) o \ y o \C (F y, x)) o (\C (F y, x) o \ x o \D (y, G x))))" using 4 calculation by simp show "\h. h \ HomD.set (y, G x) \ ((\D (y, G x) o \ y o \C (F y, x)) o (\C (F y, x) o \ x o \D (y, G x))) h = (\D (y, G x) o (\ y o \ x) o \D (y, G x)) h" proof - fix h assume h: "h \ HomD.set (y, G x)" hence "\\ x (\D (y, G x) h) : F y \\<^sub>C x\" using x y HomD.\_mapsto [of y "G x"] \_mapsto by auto thus "((\D (y, G x) o \ y o \C (F y, x)) o (\C (F y, x) o \ x o \D (y, G x))) h = (\D (y, G x) o (\ y o \ x) o \D (y, G x)) h" using x y HomC.\_\ by simp qed qed finally show ?thesis by auto qed moreover have "\ (y, x) \\<^sub>S \ (y, x) = S.mkArr (HomD.set (y, G x)) (HomD.set (y, G x)) (\h. h)" using 1 2 6 calculation by (metis S.arr_mkArr S.cod_mkArr S.ide_in_hom S.mkIde_as_mkArr S.in_homE) ultimately have 4: "S.mkArr (HomD.set (y, G x)) (HomD.set (y, G x)) (\D (y, G x) o (\ y o \ x) o \D (y, G x)) = S.mkArr (HomD.set (y, G x)) (HomD.set (y, G x)) (\h. h)" by auto have 5: "S.arr (S.mkArr (HomD.set (y, G x)) (HomD.set (y, G x)) (\D (y, G x) o (\ y o \ x) o \D (y, G x)))" using 1 3 by fastforce hence "restrict (\D (y, G x) o (\ y o \ x) o \D (y, G x)) (HomD.set (y, G x)) = S.Fun (S.mkArr (HomD.set (y, G x)) (HomD.set (y, G x)) (\D (y, G x) o (\ y o \ x) o \D (y, G x)))" by auto also have "... = restrict (\h. h) (HomD.set (y, G x))" using 4 5 by auto finally show ?thesis by auto qed moreover have "\D (y, G x) h \ HomD.set (y, G x)" using x y h HomD.\_mapsto [of y "G x"] by auto ultimately have "\D (y, G x) h = (\D (y, G x) o (\ y o \ x) o \D (y, G x)) (\D (y, G x) h)" by fast hence "\D (y, G x) (\D (y, G x) h) = \D (y, G x) ((\D (y, G x) o (\ y o \ x) o \D (y, G x)) (\D (y, G x) h))" by simp hence "h = \D (y, G x) (\D (y, G x) (\ y (\ x (\D (y, G x) (\D (y, G x) h)))))" using x y h HomD.\_\ by simp also have "... = \ y (\ x h)" using x y h HomD.\_\ HomD.\_\ [of "\ y (\ x h)" y "G x"] \_mapsto \_mapsto by fastforce finally show ?thesis by auto qed next fix x :: 'c and x' :: 'c and y :: 'd and y' :: 'd and f :: 'c and g :: 'd and h :: 'c assume f: "\f : x \\<^sub>C x'\" and g: "\g : y' \\<^sub>D y\" and h: "\h : F y \\<^sub>C x\" have x: "C.ide x" using f by auto have y: "D.ide y" using g by auto have x': "C.ide x'" using f by auto have y': "D.ide y'" using g by auto show "\ y' (f \\<^sub>C h \\<^sub>C F g) = G f \\<^sub>D \ y h \\<^sub>D g" proof - have 0: "restrict ((\D (y', G x') o (\h. G f \\<^sub>D h \\<^sub>D g) o \D (y, G x)) o (\D (y, G x) o \ y o \C (F y, x))) (HomC.set (F y, x)) = restrict ((\D (y', G x') o \ y' o \C (F y', x')) o (\C (F y', x') o (\h. f \\<^sub>C h \\<^sub>C F g)) o \C (F y, x)) (HomC.set (F y, x))" proof - have 1: "S.arr (\ (y, x)) \ \ (y, x) = S.mkArr (HomC.set (F y, x)) (HomD.set (y, G x)) (\D (y, G x) o \ y o \C (F y, x))" using x y \_simp [of y x] by auto have 2: "S.arr (\ (y', x')) \ \ (y', x') = S.mkArr (HomC.set (F y', x')) (HomD.set (y', G x')) (\D (y', G x') o \ y' o \C (F y', x'))" using x' y' \_simp [of y' x'] by auto have 3: "S.arr (S.mkArr (HomC.set (F y, x)) (HomD.set (y', G x')) ((\D (y', G x') o (\h. G f \\<^sub>D h \\<^sub>D g) o \D (y, G x)) o (\D (y, G x) o \ y o \C (F y, x)))) \ S.mkArr (HomC.set (F y, x)) (HomD.set (y', G x')) ((\D (y', G x') o (\h. G f \\<^sub>D h \\<^sub>D g) o \D (y, G x)) o (\D (y, G x) o \ y o \C (F y, x))) = S (S.mkArr (HomD.set (y, G x)) (HomD.set (y', G x')) (\D (y', G x') o (\h. G f \\<^sub>D h \\<^sub>D g) o \D (y, G x))) (S.mkArr (HomC.set (F y, x)) (HomD.set (y, G x)) (\D (y, G x) o \ y o \C (F y, x)))" proof - have 1: "S.seq (S.mkArr (HomD.set (y, G x)) (HomD.set (y', G x')) (\D (y', G x') o (\h. G f \\<^sub>D h \\<^sub>D g) o \D (y, G x))) (S.mkArr (HomC.set (F y, x)) (HomD.set (y, G x)) (\D (y, G x) o \ y o \C (F y, x)))" proof - have "S.arr (Hom_DopxG.map (g, f)) \ Hom_DopxG.map (g, f) = S.mkArr (HomD.set (y, G x)) (HomD.set (y', G x')) (\D (y', G x') o (\h. G f \\<^sub>D h \\<^sub>D g) o \D (y, G x))" using f g Hom_DopxG.preserves_arr Hom_DopxG_map_simp by fastforce thus ?thesis using 1 S.cod_mkArr S.dom_mkArr S.seqI by metis qed have "S.seq (S.mkArr (HomD.set (y, G x)) (HomD.set (y', G x')) (\D (y', G x') o (\h. G f \\<^sub>D h \\<^sub>D g) o \D (y, G x))) (S.mkArr (HomC.set (F y, x)) (HomD.set (y, G x)) (\D (y, G x) o \ y o \C (F y, x)))" using 1 by (intro S.seqI', auto) moreover have "S.mkArr (HomC.set (F y, x)) (HomD.set (y', G x')) ((\D (y', G x') o (\h. G f \\<^sub>D h \\<^sub>D g) o \D (y, G x)) o (\D (y, G x) o \ y o \C (F y, x))) = S (S.mkArr (HomD.set (y, G x)) (HomD.set (y', G x')) (\D (y', G x') o (\h. G f \\<^sub>D h \\<^sub>D g) o \D (y, G x))) (S.mkArr (HomC.set (F y, x)) (HomD.set (y, G x)) (\D (y, G x) o \ y o \C (F y, x)))" using 1 S.comp_mkArr by fastforce ultimately show ?thesis by auto qed moreover have 4: "S.arr (S.mkArr (HomC.set (F y, x)) (HomD.set (y', G x')) ((\D (y', G x') o \ y' o \C (F y', x')) o (\C (F y', x') o (\h. f \\<^sub>C h \\<^sub>C F g) o \C (F y, x)))) \ S.mkArr (HomC.set (F y, x)) (HomD.set (y', G x')) ((\D (y', G x') o \ y' o \C (F y', x')) o (\C (F y', x') o (\h. f \\<^sub>C h \\<^sub>C F g) o \C (F y, x))) = S (S.mkArr (HomC.set (F y', x')) (HomD.set (y', G x')) (\D (y', G x') o \ y' o \C (F y', x'))) (S.mkArr (HomC.set (F y, x)) (HomC.set (F y', x')) (\C (F y', x') o (\h. f \\<^sub>C h \\<^sub>C F g) o \C (F y, x)))" proof - have 5: "S.seq (S.mkArr (HomC.set (F y', x')) (HomD.set (y', G x')) (\D (y', G x') o \ y' o \C (F y', x'))) (S.mkArr (HomC.set (F y, x)) (HomC.set (F y', x')) (\C (F y', x') o (\h. f \\<^sub>C h \\<^sub>C F g) o \C (F y, x)))" proof - have "S.arr (Hom_FopxC.map (g, f)) \ Hom_FopxC.map (g, f) = S.mkArr (HomC.set (F y, x)) (HomC.set (F y', x')) (\C (F y', x') o (\h. f \\<^sub>C h \\<^sub>C F g) o \C (F y, x))" using f g Hom_FopxC.preserves_arr Hom_FopxC_map_simp by fastforce thus ?thesis using 2 S.cod_mkArr S.dom_mkArr S.seqI by metis qed have "S.seq (S.mkArr (HomC.set (F y', x')) (HomD.set (y', G x')) (\D (y', G x') o \ y' o \C (F y', x'))) (S.mkArr (HomC.set (F y, x)) (HomC.set (F y', x')) (\C (F y', x') o (\h. f \\<^sub>C h \\<^sub>C F g) o \C (F y, x)))" using 5 by (intro S.seqI', auto) moreover have "S.mkArr (HomC.set (F y, x)) (HomD.set (y', G x')) ((\D (y', G x') o \ y' o \C (F y', x')) o (\C (F y', x') o (\h. f \\<^sub>C h \\<^sub>C F g) o \C (F y, x))) = S (S.mkArr (HomC.set (F y', x')) (HomD.set (y', G x')) (\D (y', G x') o \ y' o \C (F y', x'))) (S.mkArr (HomC.set (F y, x)) (HomC.set (F y', x')) (\C (F y', x') o (\h. f \\<^sub>C h \\<^sub>C F g) o \C (F y, x)))" using 5 S.comp_mkArr by fastforce ultimately show ?thesis by argo qed moreover have 2: "S.mkArr (HomC.set (F y, x)) (HomD.set (y', G x')) ((\D (y', G x') o (\h. G f \\<^sub>D h \\<^sub>D g) o \D (y, G x)) o (\D (y, G x) o \ y o \C (F y, x))) = S.mkArr (HomC.set (F y, x)) (HomD.set (y', G x')) ((\D (y', G x') o \ y' o \C (F y', x')) o (\C (F y', x') o (\h. f \\<^sub>C h \\<^sub>C F g) o \C (F y, x)))" proof - have "S (Hom_DopxG.map (g, f)) (\ (y, x)) = S (\ (y', x')) (Hom_FopxC.map (g, f))" using f g \.is_natural_1 \.is_natural_2 by fastforce moreover have "Hom_DopxG.map (g, f) = S.mkArr (HomD.set (y, G x)) (HomD.set (y', G x')) (\D (y', G x') o (\h. G f \\<^sub>D h \\<^sub>D g) o \D (y, G x))" using f g Hom_DopxG_map_simp [of "(g, f)"] by fastforce moreover have "Hom_FopxC.map (g, f) = S.mkArr (HomC.set (F y, x)) (HomC.set (F y', x')) (\C (F y', x') o (\h. f \\<^sub>C h \\<^sub>C F g) o \C (F y, x))" using f g Hom_FopxC_map_simp [of "(g, f)"] by fastforce ultimately show ?thesis using 1 2 3 4 by simp qed ultimately have 6: "S.arr (S.mkArr (HomC.set (F y, x)) (HomD.set (y', G x')) ((\D (y', G x') o (\h. G f \\<^sub>D h \\<^sub>D g) o \D (y, G x)) o (\D (y, G x) o \ y o \C (F y, x))))" by fast hence "restrict ((\D (y', G x') o (\h. D (G f) (D h g)) o \D (y, G x)) o (\D (y, G x) o \ y o \C (F y, x))) (HomC.set (F y, x)) = S.Fun (S.mkArr (HomC.set (F y, x)) (HomD.set (y', G x')) ((\D (y', G x') o (\h. G f \\<^sub>D h \\<^sub>D g) o \D (y, G x)) o (\D (y, G x) o \ y o \C (F y, x))))" by simp also have "... = S.Fun (S.mkArr (HomC.set (F y, x)) (HomD.set (y', G x')) ((\D (y', G x') o \ y' o \C (F y', x')) o (\C (F y', x') o (\h. f \\<^sub>C h \\<^sub>C F g) o \C (F y, x))))" using 2 by argo also have "... = restrict ((\D (y', G x') o \ y' o \C (F y', x')) o (\C (F y', x') o (\h. f \\<^sub>C h \\<^sub>C F g) o \C (F y, x))) (HomC.set (F y, x))" using 4 S.Fun_mkArr by meson finally show ?thesis by auto qed hence 5: "((\D (y', G x') \ (\h. G f \\<^sub>D h \\<^sub>D g) \ \D (y, G x)) \ (\D (y, G x) \ \ y \ \C (F y, x))) (\C (F y, x) h) = (\D (y', G x') \ \ y' \ \C (F y', x') \ (\C (F y', x') \ (\h. f \\<^sub>C h \\<^sub>C F g)) \ \C (F y, x)) (\C (F y, x) h)" proof - have "\C (F y, x) h \ HomC.set (F y, x)" using x y h HomC.\_mapsto [of "F y" x] by auto thus ?thesis using 0 h restr_eqE [of "(\D (y', G x') \ (\h. G f \\<^sub>D h \\<^sub>D g) \ \D (y, G x)) \ (\D (y, G x) \ \ y \ \C (F y, x))" "HomC.set (F y, x)" "(\D (y', G x') \ \ y' \ \C (F y', x')) \ (\C (F y', x') \ (\h. f \\<^sub>C h \\<^sub>C F g) o \C (F y, x))"] by fast qed show ?thesis proof - have "\ y' (C f (C h (F g))) = \D (y', G x') (\D (y', G x') (\ y' (\C (F y', x') (\C (F y', x') (C f (C (\C (F y, x) (\C (F y, x) h)) (F g)))))))" proof - have "\D (y', G x') (\D (y', G x') (\ y' (\C (F y', x') (\C (F y', x') (C f (C (\C (F y, x) (\C (F y, x) h)) (F g))))))) = \D (y', G x') (\D (y', G x') (\ y' (\C (F y', x') (\C (F y', x') (C f (C h (F g)))))))" using x y h HomC.\_\ by simp also have "... = \D (y', G x') (\D (y', G x') (\ y' (C f (C h (F g)))))" using f g h HomC.\_\ [of "C f (C h (F g))"] by fastforce also have "... = \ y' (C f (C h (F g)))" proof - have "\\ y' (f \\<^sub>C h \\<^sub>C F g) : y' \\<^sub>D G x'\" using f g h y' x' \_mapsto [of y' x'] by auto thus ?thesis by simp qed finally show ?thesis by auto qed also have "... = \D (y', G x') (\D (y', G x') (G f \\<^sub>D \D (y, G x) (\D (y, G x) (\ y (\C (F y, x) (\C (F y, x) h)))) \\<^sub>D g))" using 5 by force also have "... = D (G f) (D (\ y h) g)" proof - have \yh: "\\ y h : y \\<^sub>D G x\" using x y h \_mapsto by auto have "\D (y', G x') (\D (y', G x') (G f \\<^sub>D \D (y, G x) (\D (y, G x) (\ y (\C (F y, x) (\C (F y, x) h)))) \\<^sub>D g)) = \D (y', G x') (\D (y', G x') (G f \\<^sub>D \D (y, G x) (\D (y, G x) (\ y h)) \\<^sub>D g))" using x y f g h by auto also have "... = \D (y', G x') (\D (y', G x') (G f \\<^sub>D \ y h \\<^sub>D g))" using \yh x' y' f g by simp also have "... = G f \\<^sub>D \ y h \\<^sub>D g" using \yh f g by fastforce finally show ?thesis by auto qed finally show ?thesis by auto qed qed qed theorem induces_meta_adjunction: shows "meta_adjunction C D F G \ \" .. end section "Putting it All Together" text\ Combining the above results, an interpretation of any one of the locales: \left_adjoint_functor\, \right_adjoint_functor\, \meta_adjunction\, \hom_adjunction\, and \unit_counit_adjunction\ extends to an interpretation of \adjunction\. \ context meta_adjunction begin + interpretation S: replete_setcat . interpretation F: left_adjoint_functor D C F using has_left_adjoint_functor by auto interpretation G: right_adjoint_functor C D G using has_right_adjoint_functor by auto interpretation \\: unit_counit_adjunction C D F G \ \ using induces_unit_counit_adjunction \_def \_def by auto - interpretation \\: hom_adjunction C D replete_setcat.comp \C \D F G \ \ + interpretation \\: hom_adjunction C D S.comp S.setp \C \D F G \ \ using induces_hom_adjunction by auto theorem induces_adjunction: - shows "adjunction C D replete_setcat.comp \C \D F G \ \ \ \ \ \" + shows "adjunction C D S.comp S.setp \C \D F G \ \ \ \ \ \" using \_map_simp \_map_simp \_in_terms_of_\ \_in_terms_of_\' \_in_terms_of_\ \_in_terms_of_\' \_simp \_simp \_def \_def by unfold_locales auto end context unit_counit_adjunction begin interpretation \\: meta_adjunction C D F G \ \ using induces_meta_adjunction by auto + interpretation S: replete_setcat . interpretation F: left_adjoint_functor D C F using \\.has_left_adjoint_functor by auto interpretation G: right_adjoint_functor C D G using \\.has_right_adjoint_functor by auto - interpretation \\: hom_adjunction C D replete_setcat.comp + interpretation \\: hom_adjunction C D S.comp S.setp \\.\C \\.\D F G \\.\ \\.\ using \\.induces_hom_adjunction by auto theorem induces_adjunction: - shows "adjunction C D replete_setcat.comp \\.\C \\.\D F G \ \ \ \ \\.\ \\.\" + shows "adjunction C D S.comp S.setp \\.\C \\.\D F G \ \ \ \ \\.\ \\.\" using \_in_terms_of_\ \_in_terms_of_\ \\.\_in_terms_of_\' \_def \\.\_in_terms_of_\' \\.\_simp \\.\_simp \_def by unfold_locales auto end context hom_adjunction begin interpretation \\: meta_adjunction C D F G \ \ using induces_meta_adjunction by auto interpretation F: left_adjoint_functor D C F using \\.has_left_adjoint_functor by auto interpretation G: right_adjoint_functor C D G using \\.has_right_adjoint_functor by auto interpretation \\: unit_counit_adjunction C D F G \\.\ \\.\ using \\.induces_unit_counit_adjunction \\.\_def \\.\_def by auto theorem induces_adjunction: - shows "adjunction C D S \C \D F G \ \ \\.\ \\.\ \ \" + shows "adjunction C D S setp \C \D F G \ \ \\.\ \\.\ \ \" proof fix x assume "C.ide x" thus "\\.\ x = \ x (G x)" using \\.\_map_simp \\.\_def by simp next fix y assume "D.ide y" thus "\\.\ y = \ y (F y)" using \\.\_map_simp \\.\_def by simp fix x y f assume y: "D.ide y" and f: "\f : F y \\<^sub>C x\" show "\ y f = G f \\<^sub>D \\.\ y" using y f \\.\_in_terms_of_\ \\.\_def by simp show "\ y f = (\D (y, G x) \ \.FUN (y, x) \ \C (F y, x)) f" using y f \_def by auto next fix x y g assume x: "C.ide x" and g: "\g : y \\<^sub>D G x\" show "\ x g = \\.\ x \\<^sub>C F g" using x g \\.\_in_terms_of_\ \\.\_def by simp show "\ x g = (\C (F y, x) \ \.FUN (y, x) \ \D (y, G x)) g" using x g \_def by fast next fix x y assume x: "C.ide x" and y: "D.ide y" show "\ (y, x) = S.mkArr (HomC.set (F y, x)) (HomD.set (y, G x)) (\D (y, G x) o \ y o \C (F y, x))" using x y \_simp by simp show "\ (y, x) = S.mkArr (HomD.set (y, G x)) (HomC.set (F y, x)) (\C (F y, x) o \ x o \D (y, G x))" using x y \_simp by simp qed end context left_adjoint_functor begin interpretation \\: meta_adjunction C D F G \ \ using induces_meta_adjunction by auto + interpretation S: replete_setcat . theorem induces_adjunction: - shows "adjunction C D replete_setcat.comp \\.\C \\.\D F G \ \ \\.\ \\.\ \\.\ \\.\" + shows "adjunction C D S.comp S.setp \\.\C \\.\D F G \ \ \\.\ \\.\ \\.\ \\.\" using \\.induces_adjunction by auto end context right_adjoint_functor begin interpretation \\: meta_adjunction C D F G \ \ using induces_meta_adjunction by auto + interpretation S: replete_setcat . theorem induces_adjunction: - shows "adjunction C D replete_setcat.comp \\.\C \\.\D F G \ \ \\.\ \\.\ \\.\ \\.\" + shows "adjunction C D S.comp S.setp \\.\C \\.\D F G \ \ \\.\ \\.\ \\.\ \\.\" using \\.induces_adjunction by auto end definition adjoint_functors where "adjoint_functors C D F G = (\\ \. meta_adjunction C D F G \ \)" lemma adjoint_functors_respects_naturally_isomorphic: assumes "adjoint_functors C D F G" and "naturally_isomorphic D C F' F" and "naturally_isomorphic C D G G'" shows "adjoint_functors C D F' G'" proof - obtain \ \ where \\: "meta_adjunction C D F G \ \" using assms(1) adjoint_functors_def by blast interpret \\: meta_adjunction C D F G \ \ using \\ by simp obtain \ where \: "natural_isomorphism D C F' F \" using assms(2) naturally_isomorphic_def by blast obtain \ where \: "natural_isomorphism C D G G' \" using assms(3) naturally_isomorphic_def by blast show ?thesis using adjoint_functors_def \ \ \\.respects_natural_isomorphism by blast qed lemma left_adjoint_functor_respects_naturally_isomorphic: assumes "left_adjoint_functor D C F" and "naturally_isomorphic D C F F'" shows "left_adjoint_functor D C F'" proof - interpret F: left_adjoint_functor D C F using assms(1) by simp have 1: "meta_adjunction C D F F.G F.\ F.\" using F.induces_meta_adjunction by simp interpret \\: meta_adjunction C D F F.G F.\ F.\ using 1 by simp have "adjoint_functors C D F F.G" using 1 adjoint_functors_def by blast hence 2: "adjoint_functors C D F' F.G" using assms(2) adjoint_functors_respects_naturally_isomorphic [of C D F F.G F' F.G] naturally_isomorphic_reflexive naturally_isomorphic_symmetric \\.G.functor_axioms by blast obtain \' \' where \'\': "meta_adjunction C D F' F.G \' \'" using 2 adjoint_functors_def by blast interpret \'\': meta_adjunction C D F' F.G \' \' using \'\' by simp show ?thesis using \'\'.has_left_adjoint_functor by simp qed lemma right_adjoint_functor_respects_naturally_isomorphic: assumes "right_adjoint_functor C D G" and "naturally_isomorphic C D G G'" shows "right_adjoint_functor C D G'" proof - interpret G: right_adjoint_functor C D G using assms(1) by simp have 1: "meta_adjunction C D G.F G G.\ G.\" using G.induces_meta_adjunction by simp interpret \\: meta_adjunction C D G.F G G.\ G.\ using 1 by simp have "adjoint_functors C D G.F G" using 1 adjoint_functors_def by blast hence 2: "adjoint_functors C D G.F G'" using assms(2) adjoint_functors_respects_naturally_isomorphic naturally_isomorphic_reflexive naturally_isomorphic_symmetric \\.F.functor_axioms by blast obtain \' \' where \'\': "meta_adjunction C D G.F G' \' \'" using 2 adjoint_functors_def by blast interpret \'\': meta_adjunction C D G.F G' \' \' using \'\' by simp show ?thesis using \'\'.has_right_adjoint_functor by simp qed section "Inverse Functors are Adjoints" (* TODO: This really should show that inverse functors induce an adjoint equivalence. *) lemma inverse_functors_induce_meta_adjunction: assumes "inverse_functors C D F G" shows "meta_adjunction C D F G (\x. G) (\y. F)" proof - interpret inverse_functors C D F G using assms by auto interpret meta_adjunction C D F G \\x. G\ \\y. F\ proof - have 1: "\y. B.arr y \ G (F y) = y" by (metis B.map_simp comp_apply inv) have 2: "\x. A.arr x \ F (G x) = x" by (metis A.map_simp comp_apply inv') show "meta_adjunction C D F G (\x. G) (\y. F)" proof fix y f x assume y: "B.ide y" and f: "\f : F y \\<^sub>A x\" show "\G f : y \\<^sub>B G x\" using y f 1 G.preserves_hom by (elim A.in_homE, auto) show "F (G f) = f" using f 2 by auto next fix x g y assume x: "A.ide x" and g: "\g : y \\<^sub>B G x\" show "\F g : F y \\<^sub>A x\" using x g 2 F.preserves_hom by (elim B.in_homE, auto) show "G (F g) = g" using g 1 A.map_def by blast next fix f x x' g y' y h assume f: "\f : x \\<^sub>A x'\" and g: "\g : y' \\<^sub>B y\" and h: "\h : F y \\<^sub>A x\" show "G (C f (C h (F g))) = D (G f) (D (G h) g)" using f g h 1 2 inv inv' A.map_def B.map_def by (elim A.in_homE B.in_homE, auto) qed qed show ?thesis .. qed lemma inverse_functors_are_adjoints: assumes "inverse_functors A B F G" shows "adjoint_functors A B F G" using assms inverse_functors_induce_meta_adjunction adjoint_functors_def by fast context inverse_functors begin lemma \_char: shows "meta_adjunction.\ B F (\x. G) = identity_functor.map B" proof (intro eqI) interpret meta_adjunction A B F G \\y. G\ \\x. F\ using inverse_functors_induce_meta_adjunction inverse_functors_axioms by auto - interpret adjunction A B replete_setcat.comp \C \D F G \\y. G\ \\x. F\ \ \ \ \ + interpret S: replete_setcat . + interpret adjunction A B S.comp S.setp \C \D F G \\y. G\ \\x. F\ \ \ \ \ using induces_adjunction by force show "natural_transformation B B B.map GF.map \" using \.natural_transformation_axioms by auto show "natural_transformation B B B.map GF.map B.map" by (simp add: B.as_nat_trans.natural_transformation_axioms inv) show "\b. B.ide b \ \ b = B.map b" using \_in_terms_of_\ \o_def \o_in_hom by fastforce qed lemma \_char: shows "meta_adjunction.\ A F G (\y. F) = identity_functor.map A" proof (intro eqI) interpret meta_adjunction A B F G \\y. G\ \\x. F\ using inverse_functors_induce_meta_adjunction inverse_functors_axioms by auto - interpret adjunction A B replete_setcat.comp \C \D F G \\y. G\ \\x. F\ \ \ \ \ + interpret S: replete_setcat . + interpret adjunction A B S.comp S.setp \C \D F G \\y. G\ \\x. F\ \ \ \ \ using induces_adjunction by force show "natural_transformation A A FG.map A.map \" using \.natural_transformation_axioms by auto show "natural_transformation A A FG.map A.map A.map" by (simp add: A.as_nat_trans.natural_transformation_axioms inv') show "\a. A.ide a \ \ a = A.map a" using \_in_terms_of_\ \o_def \o_in_hom by fastforce qed end section "Composition of Adjunctions" locale composite_adjunction = A: category A + B: category B + C: category C + F: "functor" B A F + G: "functor" A B G + F': "functor" C B F' + G': "functor" B C G' + FG: meta_adjunction A B F G \ \ + F'G': meta_adjunction B C F' G' \' \' for A :: "'a comp" (infixr "\\<^sub>A" 55) and B :: "'b comp" (infixr "\\<^sub>B" 55) and C :: "'c comp" (infixr "\\<^sub>C" 55) and F :: "'b \ 'a" and G :: "'a \ 'b" and F' :: "'c \ 'b" and G' :: "'b \ 'c" and \ :: "'b \ 'a \ 'b" and \ :: "'a \ 'b \ 'a" and \' :: "'c \ 'b \ 'c" and \' :: "'b \ 'c \ 'b" begin - interpretation FG: adjunction A B replete_setcat.comp + interpretation S: replete_setcat . + interpretation FG: adjunction A B S.comp S.setp FG.\C FG.\D F G \ \ FG.\ FG.\ FG.\ FG.\ using FG.induces_adjunction by simp - interpretation F'G': adjunction B C replete_setcat.comp F'G'.\C F'G'.\D F' G' \' \' + interpretation F'G': adjunction B C S.comp S.setp F'G'.\C F'G'.\D F' G' \' \' F'G'.\ F'G'.\ F'G'.\ F'G'.\ using F'G'.induces_adjunction by simp (* Notation for C.in_hom is inherited here somehow, but I don't know from where. *) lemma is_meta_adjunction: shows "meta_adjunction A C (F o F') (G' o G) (\z. \' z o \ (F' z)) (\x. \ x o \' (G x))" proof - interpret G'oG: composite_functor A B C G G' .. interpret FoF': composite_functor C B A F' F .. show ?thesis proof fix y f x assume y: "C.ide y" and f: "\f : FoF'.map y \\<^sub>A x\" show "\(\' y \ \ (F' y)) f : y \\<^sub>C G'oG.map x\" using y f FG.\_in_hom F'G'.\_in_hom by simp show "(\ x \ \' (G x)) ((\' y \ \ (F' y)) f) = f" using y f FG.\_in_hom F'G'.\_in_hom FG.\_\ F'G'.\_\ by simp next fix x g y assume x: "A.ide x" and g: "\g : y \\<^sub>C G'oG.map x\" show "\(\ x \ \' (G x)) g : FoF'.map y \\<^sub>A x\" using x g FG.\_in_hom F'G'.\_in_hom by auto show "(\' y \ \ (F' y)) ((\ x \ \' (G x)) g) = g" using x g FG.\_in_hom F'G'.\_in_hom FG.\_\ F'G'.\_\ by simp next fix f x x' g y' y h assume f: "\f : x \\<^sub>A x'\" and g: "\g : y' \\<^sub>C y\" and h: "\h : FoF'.map y \\<^sub>A x\" show "(\' y' \ \ (F' y')) (f \\<^sub>A h \\<^sub>A FoF'.map g) = G'oG.map f \\<^sub>C (\' y \ \ (F' y)) h \\<^sub>C g" using f g h FG.\_naturality [of f x x' "F' g" "F' y'" "F' y" h] F'G'.\_naturality [of "G f" "G x" "G x'" g y' y "\ (F' y) h"] FG.\_in_hom by fastforce qed qed interpretation K\H: natural_transformation C C \G' o F'\ \G' o G o F o F'\ \G' o FG.\ o F'\ proof - interpret \F': natural_transformation C B F' \(G o F) o F'\ \FG.\ o F'\ using FG.\_is_natural_transformation F'.as_nat_trans.natural_transformation_axioms horizontal_composite by fastforce interpret G'\F': natural_transformation C C \G' o F'\ \G' o (G o F o F')\ \G' o (FG.\ o F')\ using \F'.natural_transformation_axioms G'.as_nat_trans.natural_transformation_axioms horizontal_composite by blast show "natural_transformation C C (G' o F') (G' o G o F o F') (G' o FG.\ o F')" using G'\F'.natural_transformation_axioms o_assoc by metis qed interpretation G'\F'o\': vertical_composite C C C.map \G' o F'\ \G' o G o F o F'\ F'G'.\ \G' o FG.\ o F'\ .. interpretation F\G: natural_transformation A A \F o F' o G' o G\ \F o G\ \F o F'G'.\ o G\ proof - interpret F\': natural_transformation B A \F o (F' o G')\ F \F o F'G'.\\ using F'G'.\.natural_transformation_axioms F.as_nat_trans.natural_transformation_axioms horizontal_composite by fastforce interpret F\'G: natural_transformation A A \F o (F' o G') o G\ \F o G\ \F o F'G'.\ o G\ using F\'.natural_transformation_axioms G.as_nat_trans.natural_transformation_axioms horizontal_composite by blast show "natural_transformation A A (F o F' o G' o G) (F o G) (F o F'G'.\ o G)" using F\'G.natural_transformation_axioms o_assoc by metis qed interpretation \oF\'G: vertical_composite A A \F \ F' \ G' \ G\ \F o G\ A.map \F o F'G'.\ o G\ FG.\ .. interpretation meta_adjunction A C \F o F'\ \G' o G\ \\z. \' z o \ (F' z)\ \\x. \ x o \' (G x)\ using is_meta_adjunction by auto - interpretation adjunction A C replete_setcat.comp \C \D \F \ F'\ \G' \ G\ + interpretation S: replete_setcat . + interpretation adjunction A C S.comp S.setp \C \D \F \ F'\ \G' \ G\ \\z. \' z \ \ (F' z)\ \\x. \ x \ \' (G x)\ \ \ \ \ using induces_adjunction by simp lemma \_char: shows "\ = G'\F'o\'.map" proof (intro NaturalTransformation.eqI) show "natural_transformation C C C.map (G' o G o F o F') G'\F'o\'.map" .. show "natural_transformation C C C.map (G' o G o F o F') \" by (metis (no_types, lifting) \_is_natural_transformation o_assoc) fix a assume a: "C.ide a" show "\ a = G'\F'o\'.map a" unfolding \_def using a G'\F'o\'.map_def FG.\.preserves_hom [of "F' a" "F' a" "F' a"] F'G'.\_in_terms_of_\ FG.\_map_simp \_map_simp [of a] C.ide_in_hom F'G'.\_def FG.\_def by auto qed lemma \_char: shows "\ = \oF\'G.map" proof (intro NaturalTransformation.eqI) show "natural_transformation A A (F o F' o G' o G) A.map \" by (metis (no_types, lifting) \_is_natural_transformation o_assoc) show "natural_transformation A A (F \ F' \ G' \ G) A.map \oF\'G.map" .. fix a assume a: "A.ide a" show "\ a = \oF\'G.map a" proof - have "\ a = \ a (\' (G a) (G' (G a)))" using a \_in_terms_of_\ by simp also have "... = FG.\ a \\<^sub>A F (F'G'.\ (G a) \\<^sub>B F' (G' (G a)))" by (metis F'G'.\_in_terms_of_\ F'G'.\o_def F'G'.\o_in_hom F'G'.\\.\_in_terms_of_\ F'G'.\\.\_def FG.G\.natural_transformation_axioms FG.\_in_terms_of_\ a functor.preserves_ide natural_transformation_def) also have "... = \oF\'G.map a" using a B.comp_arr_dom \oF\'G.map_def by simp finally show ?thesis by blast qed qed end section "Right Adjoints are Unique up to Natural Isomorphism" text\ As an example of the use of the of the foregoing development, we show that two right adjoints to the same functor are naturally isomorphic. \ theorem two_right_adjoints_naturally_isomorphic: assumes "adjoint_functors C D F G" and "adjoint_functors C D F G'" shows "naturally_isomorphic C D G G'" proof - text\ For any object @{term x} of @{term C}, we have that \\ x \ C.hom (F (G x)) x\ is a terminal arrow from @{term F} to @{term x}, and similarly for \\' x\. We may therefore obtain the unique coextension \\ x \ D.hom (G x) (G' x)\ of \\ x\ along \\' x\. An explicit formula for \\ x\ is \D (G' (\ x)) (\' (G x))\. Similarly, we obtain \\' x = D (G (\' x)) (\ (G' x)) \ D.hom (G' x) (G x)\. We show these are the components of inverse natural transformations between @{term G} and @{term G'}. \ obtain \ \ where \\: "meta_adjunction C D F G \ \" using assms adjoint_functors_def by blast obtain \' \' where \'\': "meta_adjunction C D F G' \' \'" using assms adjoint_functors_def by blast interpret Adj: meta_adjunction C D F G \ \ using \\ by auto - interpret Adj: adjunction C D replete_setcat.comp Adj.\C Adj.\D + interpret S: replete_setcat . + interpret Adj: adjunction C D S.comp S.setp Adj.\C Adj.\D F G \ \ Adj.\ Adj.\ Adj.\ Adj.\ using Adj.induces_adjunction by auto interpret Adj': meta_adjunction C D F G' \' \' using \'\' by auto - interpret Adj': adjunction C D replete_setcat.comp Adj'.\C Adj'.\D + interpret Adj': adjunction C D S.comp S.setp Adj'.\C Adj'.\D F G' \' \' Adj'.\ Adj'.\ Adj'.\ Adj'.\ using Adj'.induces_adjunction by auto write C (infixr "\\<^sub>C" 55) write D (infixr "\\<^sub>D" 55) write Adj.C.in_hom ("\_ : _ \\<^sub>C _\") write Adj.D.in_hom ("\_ : _ \\<^sub>D _\") let ?\o = "\a. G' (Adj.\ a) \\<^sub>D Adj'.\ (G a)" interpret \: transformation_by_components C D G G' ?\o proof show "\a. Adj.C.ide a \ \G' (Adj.\ a) \\<^sub>D Adj'.\ (G a) : G a \\<^sub>D G' a\" by fastforce show "\f. Adj.C.arr f \ (G' (Adj.\ (Adj.C.cod f)) \\<^sub>D Adj'.\ (G (Adj.C.cod f))) \\<^sub>D G f = G' f \\<^sub>D G' (Adj.\ (Adj.C.dom f)) \\<^sub>D Adj'.\ (G (Adj.C.dom f))" proof - fix f assume f: "Adj.C.arr f" let ?x = "Adj.C.dom f" let ?x' = "Adj.C.cod f" have "(G' (Adj.\ (Adj.C.cod f)) \\<^sub>D Adj'.\ (G (Adj.C.cod f))) \\<^sub>D G f = G' (Adj.\ (Adj.C.cod f) \\<^sub>C F (G f)) \\<^sub>D Adj'.\ (G (Adj.C.dom f))" using f Adj'.\.naturality [of "G f"] Adj.D.comp_assoc by simp also have "... = G' (f \\<^sub>C Adj.\ (Adj.C.dom f)) \\<^sub>D Adj'.\ (G (Adj.C.dom f))" using f Adj.\.naturality by simp also have "... = G' f \\<^sub>D G' (Adj.\ (Adj.C.dom f)) \\<^sub>D Adj'.\ (G (Adj.C.dom f))" using f Adj.D.comp_assoc by simp finally show "(G' (Adj.\ (Adj.C.cod f)) \\<^sub>D Adj'.\ (G (Adj.C.cod f))) \\<^sub>D G f = G' f \\<^sub>D G' (Adj.\ (Adj.C.dom f)) \\<^sub>D Adj'.\ (G (Adj.C.dom f))" by auto qed qed interpret natural_isomorphism C D G G' \.map proof fix a assume a: "Adj.C.ide a" show "Adj.D.iso (\.map a)" proof show "Adj.D.inverse_arrows (\.map a) (\ (G' a) (Adj'.\ a))" proof text\ The proof that the two composites are identities is a modest diagram chase. This is a good example of the inference rules for the \category\, \functor\, and \natural_transformation\ locales in action. Isabelle is able to use the single hypothesis that \a\ is an identity to implicitly fill in all the details that the various quantities are in fact arrows and that the indicated composites are all well-defined, as well as to apply associativity of composition. In most cases, this is done by auto or simp without even mentioning any of the rules that are used. $$\xymatrix{ {G' a} \ar[dd]_{\eta'(G'a)} \ar[rr]^{\tau' a} \ar[dr]_{\eta(G'a)} && {G a} \ar[rr]^{\tau a} \ar[dr]_{\eta'(Ga)} && {G' a} \\ & {GFG'a} \rrtwocell\omit{\omit(2)} \ar[ur]_{G(\epsilon' a)} \ar[dr]_{\eta'(GFG'a)} && {G'FGa} \drtwocell\omit{\omit(3)} \ar[ur]_{G'(\epsilon a)} & \\ {G'FG'a} \urtwocell\omit{\omit(1)} \ar[rr]_{G'F\eta(G'a)} \ar@/_8ex/[rrrr]_{G'FG'a} && {G'FGFG'a} \dtwocell\omit{\omit(4)} \ar[ru]_{G'FG(\epsilon' a)} \ar[rr]_{G'(\epsilon(FG'a))} && {G'FG'a} \ar[uu]_{G'(\epsilon' a)} \\ &&&& }$$ \ show "Adj.D.ide (\.map a \\<^sub>D \ (G' a) (Adj'.\ a))" proof - have "\.map a \\<^sub>D \ (G' a) (Adj'.\ a) = G' a" proof - have "\.map a \\<^sub>D \ (G' a) (Adj'.\ a) = G' (Adj.\ a) \\<^sub>D (Adj'.\ (G a) \\<^sub>D G (Adj'.\ a)) \\<^sub>D Adj.\ (G' a)" using a \.map_simp_ide Adj.\_in_terms_of_\ Adj'.\_in_terms_of_\ Adj'.\.preserves_hom [of a a a] Adj.C.ide_in_hom Adj.D.comp_assoc Adj.\_def Adj.\_def by simp also have "... = G' (Adj.\ a) \\<^sub>D (G' (F (G (Adj'.\ a))) \\<^sub>D Adj'.\ (G (F (G' a)))) \\<^sub>D Adj.\ (G' a)" using a Adj'.\.naturality [of "G (Adj'.\ a)"] by auto also have "... = (G' (Adj.\ a) \\<^sub>D G' (F (G (Adj'.\ a)))) \\<^sub>D G' (F (Adj.\ (G' a))) \\<^sub>D Adj'.\ (G' a)" using a Adj'.\.naturality [of "Adj.\ (G' a)"] Adj.D.comp_assoc by auto also have "... = G' (Adj'.\ a) \\<^sub>D (G' (Adj.\ (F (G' a))) \\<^sub>D G' (F (Adj.\ (G' a)))) \\<^sub>D Adj'.\ (G' a)" proof - have "G' (Adj.\ a) \\<^sub>D G' (F (G (Adj'.\ a))) = G' (Adj'.\ a) \\<^sub>D G' (Adj.\ (F (G' a)))" proof - have "G' (Adj.\ a \\<^sub>C F (G (Adj'.\ a))) = G' (Adj'.\ a \\<^sub>C Adj.\ (F (G' a)))" using a Adj.\.naturality [of "Adj'.\ a"] by auto thus ?thesis using a by force qed thus ?thesis using Adj.D.comp_assoc by auto qed also have "... = G' (Adj'.\ a) \\<^sub>D Adj'.\ (G' a)" proof - have "G' (Adj.\ (F (G' a))) \\<^sub>D G' (F (Adj.\ (G' a))) = G' (F (G' a))" proof - have "G' (Adj.\ (F (G' a))) \\<^sub>D G' (F (Adj.\ (G' a))) = G' (Adj.\FoF\.map (G' a))" using a Adj.\FoF\.map_simp_1 by auto moreover have "Adj.\FoF\.map (G' a) = F (G' a)" using a by (simp add: Adj.\\.triangle_F) ultimately show ?thesis by auto qed thus ?thesis using a Adj.D.comp_cod_arr [of "Adj'.\ (G' a)"] by auto qed also have "... = G' a" using a Adj'.\\.triangle_G Adj'.G\o\G.map_simp_1 [of a] by auto finally show ?thesis by auto qed thus ?thesis using a by simp qed show "Adj.D.ide (\ (G' a) (Adj'.\ a) \\<^sub>D \.map a)" proof - have "\ (G' a) (Adj'.\ a) \\<^sub>D \.map a = G a" proof - have "\ (G' a) (Adj'.\ a) \\<^sub>D \.map a = G (Adj'.\ a) \\<^sub>D (Adj.\ (G' a) \\<^sub>D G' (Adj.\ a)) \\<^sub>D Adj'.\ (G a)" using a \.map_simp_ide Adj.\_in_terms_of_\ Adj'.\.preserves_hom [of a a a] Adj.C.ide_in_hom Adj.D.comp_assoc Adj.\_def by auto also have "... = G (Adj'.\ a) \\<^sub>D (G (F (G' (Adj.\ a))) \\<^sub>D Adj.\ (G' (F (G a)))) \\<^sub>D Adj'.\ (G a)" using a Adj.\.naturality [of "G' (Adj.\ a)"] by auto also have "... = (G (Adj'.\ a) \\<^sub>D G (F (G' (Adj.\ a)))) \\<^sub>D G (F (Adj'.\ (G a))) \\<^sub>D Adj.\ (G a)" using a Adj.\.naturality [of "Adj'.\ (G a)"] Adj.D.comp_assoc by auto also have "... = G (Adj.\ a) \\<^sub>D (G (Adj'.\ (F (G a))) \\<^sub>D G (F (Adj'.\ (G a)))) \\<^sub>D Adj.\ (G a)" proof - have "G (Adj'.\ a) \\<^sub>D G (F (G' (Adj.\ a))) = G (Adj.\ a) \\<^sub>D G (Adj'.\ (F (G a)))" proof - have "G (Adj'.\ a \\<^sub>C F (G' (Adj.\ a))) = G (Adj.\ a \\<^sub>C Adj'.\ (F (G a)))" using a Adj'.\.naturality [of "Adj.\ a"] by auto thus ?thesis using a by force qed thus ?thesis using Adj.D.comp_assoc by auto qed also have "... = G (Adj.\ a) \\<^sub>D Adj.\ (G a)" proof - have "G (Adj'.\ (F (G a))) \\<^sub>D G (F (Adj'.\ (G a))) = G (F (G a))" proof - have "G (Adj'.\ (F (G a))) \\<^sub>D G (F (Adj'.\ (G a))) = G (Adj'.\FoF\.map (G a))" using a Adj'.\FoF\.map_simp_1 [of "G a"] by auto moreover have "Adj'.\FoF\.map (G a) = F (G a)" using a by (simp add: Adj'.\\.triangle_F) ultimately show ?thesis by auto qed thus ?thesis using a Adj.D.comp_cod_arr by auto qed also have "... = G a" using a Adj.\\.triangle_G Adj.G\o\G.map_simp_1 [of a] by auto finally show ?thesis by auto qed thus ?thesis using a by auto qed qed qed qed have "natural_isomorphism C D G G' \.map" .. thus "naturally_isomorphic C D G G'" using naturally_isomorphic_def by blast qed end diff --git a/thys/Category3/EquivalenceOfCategories.thy b/thys/Category3/EquivalenceOfCategories.thy --- a/thys/Category3/EquivalenceOfCategories.thy +++ b/thys/Category3/EquivalenceOfCategories.thy @@ -1,913 +1,914 @@ (* Title: EquivalenceOfCategories Author: Eugene W. Stark , 2017 Maintainer: Eugene W. Stark *) chapter "Equivalence of Categories" text \ In this chapter we define the notions of equivalence and adjoint equivalence of categories and establish some properties of functors that are part of an equivalence. \ theory EquivalenceOfCategories imports Adjunction begin locale equivalence_of_categories = C: category C + D: category D + F: "functor" D C F + G: "functor" C D G + \: natural_isomorphism D D D.map "G o F" \ + \: natural_isomorphism C C "F o G" C.map \ for C :: "'c comp" (infixr "\\<^sub>C" 55) and D :: "'d comp" (infixr "\\<^sub>D" 55) and F :: "'d \ 'c" and G :: "'c \ 'd" and \ :: "'d \ 'd" and \ :: "'c \ 'c" begin notation C.in_hom ("\_ : _ \\<^sub>C _\") notation D.in_hom ("\_ : _ \\<^sub>D _\") lemma C_arr_expansion: assumes "C.arr f" shows "\ (C.cod f) \\<^sub>C F (G f) \\<^sub>C C.inv (\ (C.dom f)) = f" and "C.inv (\ (C.cod f)) \\<^sub>C f \\<^sub>C \ (C.dom f) = F (G f)" proof - have \_dom: "C.inverse_arrows (\ (C.dom f)) (C.inv (\ (C.dom f)))" using assms C.inv_is_inverse by auto have \_cod: "C.inverse_arrows (\ (C.cod f)) (C.inv (\ (C.cod f)))" using assms C.inv_is_inverse by auto have "\ (C.cod f) \\<^sub>C F (G f) \\<^sub>C C.inv (\ (C.dom f)) = (\ (C.cod f) \\<^sub>C F (G f)) \\<^sub>C C.inv (\ (C.dom f))" using C.comp_assoc by force also have 1: "... = (f \\<^sub>C \ (C.dom f)) \\<^sub>C C.inv (\ (C.dom f))" using assms \.naturality by simp also have 2: "... = f" using assms \_dom C.comp_arr_inv C.comp_arr_dom C.comp_assoc by force finally show "\ (C.cod f) \\<^sub>C F (G f) \\<^sub>C C.inv (\ (C.dom f)) = f" by blast show "C.inv (\ (C.cod f)) \\<^sub>C f \\<^sub>C \ (C.dom f) = F (G f)" using assms 1 2 \_dom \_cod C.invert_side_of_triangle C.isoI C.iso_inv_iso by metis qed lemma G_is_faithful: shows "faithful_functor C D G" proof fix f f' assume par: "C.par f f'" and eq: "G f = G f'" show "f = f'" proof - have "C.inv (\ (C.cod f)) \ C.hom (C.cod f) (F (G (C.cod f))) \ C.iso (C.inv (\ (C.cod f)))" using par by auto moreover have 1: "\ (C.dom f) \ C.hom (F (G (C.dom f))) (C.dom f) \ C.iso (\ (C.dom f))" using par by auto ultimately have 2: "f \\<^sub>C \ (C.dom f) = f' \\<^sub>C \ (C.dom f)" using par C_arr_expansion eq C.iso_is_section C.section_is_mono by (metis C_arr_expansion(1) eq) show ?thesis proof - have "C.epi (\ (C.dom f))" using 1 par C.iso_is_retraction C.retraction_is_epi by blast thus ?thesis using 2 par by auto qed qed qed lemma G_is_essentially_surjective: shows "essentially_surjective_functor C D G" proof fix b assume b: "D.ide b" have "C.ide (F b) \ D.isomorphic (G (F b)) b" proof show "C.ide (F b)" using b by simp show "D.isomorphic (G (F b)) b" proof (unfold D.isomorphic_def) have "\D.inv (\ b) : G (F b) \\<^sub>D b\ \ D.iso (D.inv (\ b))" using b by auto thus "\f. \f : G (F b) \\<^sub>D b\ \ D.iso f" by blast qed qed thus "\a. C.ide a \ D.isomorphic (G a) b" by blast qed interpretation \_inv: inverse_transformation C C \F o G\ C.map \ .. interpretation \_inv: inverse_transformation D D D.map \G o F\ \ .. interpretation GF: equivalence_of_categories D C G F \_inv.map \_inv.map .. lemma F_is_faithful: shows "faithful_functor D C F" using GF.G_is_faithful by simp lemma F_is_essentially_surjective: shows "essentially_surjective_functor D C F" using GF.G_is_essentially_surjective by simp lemma G_is_full: shows "full_functor C D G" proof fix a a' g assume a: "C.ide a" and a': "C.ide a'" assume g: "\g : G a \\<^sub>D G a'\" show "\f. \f : a \\<^sub>C a'\ \ G f = g" proof have \a: "C.inverse_arrows (\ a) (C.inv (\ a))" using a C.inv_is_inverse by auto have \a': "C.inverse_arrows (\ a') (C.inv (\ a'))" using a' C.inv_is_inverse by auto let ?f = "\ a' \\<^sub>C F g \\<^sub>C C.inv (\ a)" have f: "\?f : a \\<^sub>C a'\" using a a' g \a \a' \.preserves_hom [of a' a' a'] \_inv.preserves_hom [of a a a] by fastforce moreover have "G ?f = g" proof - interpret F: faithful_functor D C F using F_is_faithful by auto have "F (G ?f) = F g" proof - have "F (G ?f) = C.inv (\ a') \\<^sub>C ?f \\<^sub>C \ a" using f C_arr_expansion(2) [of "?f"] by auto also have "... = (C.inv (\ a') \\<^sub>C \ a') \\<^sub>C F g \\<^sub>C C.inv (\ a) \\<^sub>C \ a" using a a' f g C.comp_assoc by fastforce also have "... = F g" using a a' g \a \a' C.comp_inv_arr C.comp_arr_dom C.comp_cod_arr by auto finally show ?thesis by blast qed moreover have "D.par (G (\ a' \\<^sub>C F g \\<^sub>C C.inv (\ a))) g" using f g by fastforce ultimately show ?thesis using f g F.is_faithful by blast qed ultimately show "\?f : a \\<^sub>C a'\ \ G ?f = g" by blast qed qed end (* I'm not sure why I had to close and re-open the context here in order to * get the G_is_full fact in the interpretation GF. *) context equivalence_of_categories begin interpretation \_inv: inverse_transformation C C \F o G\ C.map \ .. interpretation \_inv: inverse_transformation D D D.map \G o F\ \ .. interpretation GF: equivalence_of_categories D C G F \_inv.map \_inv.map .. lemma F_is_full: shows "full_functor D C F" using GF.G_is_full by simp end text \ Traditionally the term "equivalence of categories" is also used for a functor that is part of an equivalence of categories. However, it seems best to use that term for a situation in which all of the structure of an equivalence is explicitly given, and to have a different term for one of the functors involved. \ locale equivalence_functor = C: category C + D: category D + "functor" C D G for C :: "'c comp" (infixr "\\<^sub>C" 55) and D :: "'d comp" (infixr "\\<^sub>D" 55) and G :: "'c \ 'd" + assumes induces_equivalence: "\F \ \. equivalence_of_categories C D F G \ \" begin notation C.in_hom ("\_ : _ \\<^sub>C _\") notation D.in_hom ("\_ : _ \\<^sub>D _\") end sublocale equivalence_of_categories \ equivalence_functor C D G using equivalence_of_categories_axioms by (unfold_locales, blast) text \ An equivalence functor is fully faithful and essentially surjective. \ sublocale equivalence_functor \ fully_faithful_functor C D G proof - obtain F \ \ where 1: "equivalence_of_categories C D F G \ \" using induces_equivalence by blast interpret equivalence_of_categories C D F G \ \ using 1 by auto show "fully_faithful_functor C D G" using G_is_full G_is_faithful fully_faithful_functor.intro by auto qed sublocale equivalence_functor \ essentially_surjective_functor C D G proof - obtain F \ \ where 1: "equivalence_of_categories C D F G \ \" using induces_equivalence by blast interpret equivalence_of_categories C D F G \ \ using 1 by auto show "essentially_surjective_functor C D G" using G_is_essentially_surjective by auto qed lemma (in inverse_functors) induce_equivalence: shows "equivalence_of_categories A B F G B.map A.map" using inv inv' A.is_extensional B.is_extensional B.comp_arr_dom B.comp_cod_arr A.comp_arr_dom A.comp_cod_arr by unfold_locales auto lemma (in invertible_functor) is_equivalence: shows "equivalence_functor A B G" using equivalence_functor_axioms.intro equivalence_functor_def equivalence_of_categories_def induce_equivalence by blast lemma (in identity_functor) is_equivalence: shows "equivalence_functor C C map" proof - interpret inverse_functors C C map map using map_def by unfold_locales auto interpret invertible_functor C C map using inverse_functors_axioms by unfold_locales blast show ?thesis using is_equivalence by blast qed text \ A special case of an equivalence functor is an endofunctor \F\ equipped with a natural isomorphism from \F\ to the identity functor. \ context endofunctor begin lemma isomorphic_to_identity_is_equivalence: assumes "natural_isomorphism A A F A.map \" shows "equivalence_functor A A F" proof - interpret \: natural_isomorphism A A F A.map \ using assms by auto interpret \': inverse_transformation A A F A.map \ .. interpret F\': natural_isomorphism A A F \F o F\ \F o \'.map\ proof - interpret F\': natural_transformation A A F \F o F\ \F o \'.map\ using \'.natural_transformation_axioms functor_axioms horizontal_composite [of A A A.map F \'.map A F F F] by simp show "natural_isomorphism A A F (F o F) (F o \'.map)" apply unfold_locales using \'.components_are_iso by fastforce qed interpret F\'o\': vertical_composite A A A.map F \F o F\ \'.map \F o \'.map\ .. interpret F\'o\': natural_isomorphism A A A.map \F o F\ F\'o\'.map using \'.natural_isomorphism_axioms F\'.natural_isomorphism_axioms natural_isomorphisms_compose by fast interpret inv_F\'o\': inverse_transformation A A A.map \F o F\ F\'o\'.map .. interpret F: equivalence_of_categories A A F F F\'o\'.map inv_F\'o\'.map .. show ?thesis .. qed end text \ An adjoint equivalence is an equivalence of categories that is also an adjunction. \ locale adjoint_equivalence = unit_counit_adjunction C D F G \ \ + \: natural_isomorphism D D D.map "G o F" \ + \: natural_isomorphism C C "F o G" C.map \ for C :: "'c comp" (infixr "\\<^sub>C" 55) and D :: "'d comp" (infixr "\\<^sub>D" 55) and F :: "'d \ 'c" and G :: "'c \ 'd" and \ :: "'d \ 'd" and \ :: "'c \ 'c" text \ An adjoint equivalence is clearly an equivalence of categories. \ sublocale adjoint_equivalence \ equivalence_of_categories .. context adjoint_equivalence begin text \ The triangle identities for an adjunction reduce to inverse relations when \\\ and \\\ are natural isomorphisms. \ lemma triangle_G': assumes "C.ide a" shows "D.inverse_arrows (\ (G a)) (G (\ a))" proof show "D.ide (G (\ a) \\<^sub>D \ (G a))" using assms triangle_G G\o\G.map_simp_ide by fastforce thus "D.ide (\ (G a) \\<^sub>D G (\ a))" using assms D.section_retraction_of_iso [of "G (\ a)" "\ (G a)"] by auto qed lemma triangle_F': assumes "D.ide b" shows "C.inverse_arrows (F (\ b)) (\ (F b))" proof show "C.ide (\ (F b) \\<^sub>C F (\ b))" using assms triangle_F \FoF\.map_simp_ide by auto thus "C.ide (F (\ b) \\<^sub>C \ (F b))" using assms C.section_retraction_of_iso [of "\ (F b)" "F (\ b)"] by auto qed text \ An adjoint equivalence can be dualized by interchanging the two functors and inverting the natural isomorphisms. This is somewhat awkward to prove, but probably useful to have done it once and for all. \ lemma dual_equivalence: assumes "adjoint_equivalence C D F G \ \" shows "adjoint_equivalence D C G F (inverse_transformation.map C C (C.map) \) (inverse_transformation.map D D (G o F) \)" proof - interpret adjoint_equivalence C D F G \ \ using assms by auto interpret \': inverse_transformation C C \F o G\ C.map \ .. interpret \': inverse_transformation D D D.map \G o F\ \ .. interpret G\': natural_transformation C D G \G o F o G\ \G o \'.map\ proof - have "natural_transformation C D G (G o (F o G)) (G o \'.map)" using G.as_nat_trans.natural_transformation_axioms \'.natural_transformation_axioms horizontal_composite by fastforce thus "natural_transformation C D G (G o F o G) (G o \'.map)" using o_assoc by metis qed interpret \'G: natural_transformation C D \G o F o G\ G \\'.map o G\ using \'.natural_transformation_axioms G.as_nat_trans.natural_transformation_axioms horizontal_composite by fastforce interpret \'F: natural_transformation D C F \F o G o F\ \\'.map o F\ using \'.natural_transformation_axioms F.as_nat_trans.natural_transformation_axioms horizontal_composite by fastforce interpret F\': natural_transformation D C \F o G o F\ F \F o \'.map\ proof - have "natural_transformation D C (F o (G o F)) F (F o \'.map)" using \'.natural_transformation_axioms F.as_nat_trans.natural_transformation_axioms horizontal_composite by fastforce thus "natural_transformation D C (F o G o F) F (F o \'.map)" using o_assoc by metis qed interpret F\'o\'F: vertical_composite D C F \(F o G) o F\ F \\'.map o F\ \F o \'.map\ .. interpret \'GoG\': vertical_composite C D G \G o F o G\ G \G o \'.map\ \\'.map o G\ .. show ?thesis proof show "\'GoG\'.map = G" proof (intro NaturalTransformation.eqI) show "natural_transformation C D G G G" using G.as_nat_trans.natural_transformation_axioms by auto show "natural_transformation C D G G \'GoG\'.map" using \'GoG\'.natural_transformation_axioms by auto show "\a. C.ide a \ \'GoG\'.map a = G a" proof - fix a assume a: "C.ide a" show "\'GoG\'.map a = G a" using a \'GoG\'.map_simp_ide triangle_G' G.preserves_ide \'.inverts_components \'.inverts_components D.inverse_unique G.preserves_inverse_arrows G\o\G.map_simp_ide D.inverse_arrows_sym triangle_G by (metis o_apply) qed qed show "F\'o\'F.map = F" proof (intro NaturalTransformation.eqI) show "natural_transformation D C F F F" using F.as_nat_trans.natural_transformation_axioms by auto show "natural_transformation D C F F F\'o\'F.map" using F\'o\'F.natural_transformation_axioms by auto show "\b. D.ide b \ F\'o\'F.map b = F b" proof - fix b assume b: "D.ide b" show "F\'o\'F.map b = F b" using b F\'o\'F.map_simp_ide \FoF\.map_simp_ide triangle_F triangle_F' \'.inverts_components \'.inverts_components F.preserves_ide C.inverse_unique F.preserves_inverse_arrows C.inverse_arrows_sym by (metis o_apply) qed qed qed qed end text \ Every fully faithful and essentially surjective functor underlies an adjoint equivalence. To prove this without repeating things that were already proved in @{theory Category3.Adjunction}, we first show that a fully faithful and essentially surjective functor is a left adjoint functor, and then we show that if the left adjoint in a unit-counit adjunction is fully faithful and essentially surjective, then the unit and counit are natural isomorphisms; hence the adjunction is in fact an adjoint equivalence. \ locale fully_faithful_and_essentially_surjective_functor = C: category C + D: category D + fully_faithful_functor C D F + essentially_surjective_functor C D F for C :: "'c comp" (infixr "\\<^sub>C" 55) and D :: "'d comp" (infixr "\\<^sub>D" 55) and F :: "'c \ 'd" begin notation C.in_hom ("\_ : _ \\<^sub>C _\") notation D.in_hom ("\_ : _ \\<^sub>D _\") lemma is_left_adjoint_functor: shows "left_adjoint_functor C D F" proof fix y assume y: "D.ide y" let ?x = "SOME x. C.ide x \ (\e. D.iso e \ \e : F x \\<^sub>D y\)" let ?e = "SOME e. D.iso e \ \e : F ?x \\<^sub>D y\" have "\x e. D.iso e \ terminal_arrow_from_functor C D F x y e" proof - have "\x. D.iso ?e \ terminal_arrow_from_functor C D F x y ?e" proof - have x: "C.ide ?x \ (\e. D.iso e \ \e : F ?x \\<^sub>D y\)" using y essentially_surjective someI_ex [of "\x. C.ide x \ (\e. D.iso e \ \e : F x \\<^sub>D y\)"] by blast hence e: "D.iso ?e \ \?e : F ?x \\<^sub>D y\" using someI_ex [of "\e. D.iso e \ \e : F ?x \\<^sub>D y\"] by blast interpret arrow_from_functor C D F ?x y ?e using x e by (unfold_locales, simp) interpret terminal_arrow_from_functor C D F ?x y ?e proof fix x' f assume 1: "arrow_from_functor C D F x' y f" interpret f: arrow_from_functor C D F x' y f using 1 by simp have f: "\f: F x' \\<^sub>D y\" by (meson f.arrow) show "\!g. is_coext x' f g" proof let ?g = "SOME g. \g : x' \\<^sub>C ?x\ \ F g = D.inv ?e \\<^sub>D f" have g: "\?g : x' \\<^sub>C ?x\ \ F ?g = D.inv ?e \\<^sub>D f" using f e x f.arrow is_full D.comp_in_homI D.inv_in_hom someI_ex [of "\g. \g : x' \\<^sub>C ?x\ \ F g = D.inv ?e \\<^sub>D f"] by auto show 1: "is_coext x' f ?g" proof - have "\?g : x' \\<^sub>C ?x\" using g by simp moreover have "?e \\<^sub>D F ?g = f" proof - have "?e \\<^sub>D F ?g = ?e \\<^sub>D D.inv ?e \\<^sub>D f" using g by simp also have "... = (?e \\<^sub>D D.inv ?e) \\<^sub>D f" using e f D.inv_in_hom by (metis D.comp_assoc) also have "... = f" proof - have "?e \\<^sub>D D.inv ?e = y" using e D.comp_arr_inv D.inv_is_inverse by auto thus ?thesis using f D.comp_cod_arr by auto qed finally show ?thesis by blast qed ultimately show ?thesis unfolding is_coext_def by simp qed show "\g'. is_coext x' f g' \ g' = ?g" proof - fix g' assume g': "is_coext x' f g'" have 2: "\g' : x' \\<^sub>C ?x\ \ ?e \\<^sub>D F g' = f" using g' is_coext_def by simp have 3: "\?g : x' \\<^sub>C ?x\ \ ?e \\<^sub>D F ?g = f" using 1 is_coext_def by simp have "F g' = F ?g" using e f 2 3 D.iso_is_section D.section_is_mono D.monoE D.arrI by (metis (no_types, lifting) D.arrI) moreover have "C.par g' ?g" using 2 3 by fastforce ultimately show "g' = ?g" using is_faithful [of g' ?g] by simp qed qed qed show ?thesis using e terminal_arrow_from_functor_axioms by auto qed thus ?thesis by auto qed thus "\x e. terminal_arrow_from_functor C D F x y e" by blast qed lemma extends_to_adjoint_equivalence: shows "\G \ \. adjoint_equivalence C D G F \ \" proof - interpret left_adjoint_functor C D F using is_left_adjoint_functor by blast interpret Adj: meta_adjunction D C F G \ \ using induces_meta_adjunction by simp - interpret Adj: adjunction D C replete_setcat.comp + interpret S: replete_setcat . + interpret Adj: adjunction D C S.comp S.setp Adj.\C Adj.\D F G \ \ Adj.\ Adj.\ Adj.\ Adj.\ using induces_adjunction by simp interpret equivalence_of_categories D C F G Adj.\ Adj.\ proof show 1: "\a. D.ide a \ D.iso (Adj.\ a)" proof - fix a assume a: "D.ide a" interpret \a: terminal_arrow_from_functor C D F \G a\ a \Adj.\ a\ using a Adj.has_terminal_arrows_from_functor [of a] by blast have "D.retraction (Adj.\ a)" proof - obtain b \ where \: "C.ide b \ D.iso \ \ \\: F b \\<^sub>D a\" using a essentially_surjective by blast interpret \: arrow_from_functor C D F b a \ using \ by (unfold_locales, simp) let ?g = "\a.the_coext b \" have 1: "\?g : b \\<^sub>C G a\ \ Adj.\ a \\<^sub>D F ?g = \" using \.arrow_from_functor_axioms \a.the_coext_prop [of b \] by simp have "a = (Adj.\ a \\<^sub>D F ?g) \\<^sub>D D.inv \" using a 1 \ D.comp_cod_arr Adj.\.preserves_hom D.invert_side_of_triangle(2) by auto also have "... = Adj.\ a \\<^sub>D F ?g \\<^sub>D D.inv \" using a 1 \ D.inv_in_hom Adj.\.preserves_hom [of a a a] D.comp_assoc by blast finally have "\f. D.ide (Adj.\ a \\<^sub>D f)" using a by metis thus ?thesis unfolding D.retraction_def by blast qed moreover have "D.mono (Adj.\ a)" proof show "D.arr (Adj.\ a)" using a by simp show "\f f'. D.seq (Adj.\ a) f \ D.seq (Adj.\ a) f' \ Adj.\ a \\<^sub>D f = Adj.\ a \\<^sub>D f' \ f = f'" proof - fix f f' assume ff': "D.seq (Adj.\ a) f \ D.seq (Adj.\ a) f' \ Adj.\ a \\<^sub>D f = Adj.\ a \\<^sub>D f'" have f: "\f : D.dom f \\<^sub>D F (G a)\" using a ff' Adj.\.preserves_hom [of a a a] by fastforce have f': "\f' : D.dom f' \\<^sub>D F (G a)\" using a ff' Adj.\.preserves_hom [of a a a] by fastforce have par: "D.par f f'" using f f' ff' D.dom_comp [of "Adj.\ a" f] by force obtain b' \ where \: "C.ide b' \ D.iso \ \ \\: F b' \\<^sub>D D.dom f\" using par essentially_surjective D.ide_dom [of f] by blast have 1: "Adj.\ a \\<^sub>D f \\<^sub>D \ = Adj.\ a \\<^sub>D f' \\<^sub>D \" using ff' \ par D.comp_assoc by metis obtain g where g: "\g : b' \\<^sub>C G a\ \ F g = f \\<^sub>D \" using a f \ is_full [of "G a" b' "f \\<^sub>D \"] by auto obtain g' where g': "\g' : b' \\<^sub>C G a\ \ F g' = f' \\<^sub>D \" using a f' par \ is_full [of "G a" b' "f' \\<^sub>D \"] by auto interpret f\: arrow_from_functor C D F b' a \Adj.\ a \\<^sub>D f \\<^sub>D \\ using a \ f Adj.\.preserves_hom by (unfold_locales, fastforce) interpret f'\: arrow_from_functor C D F b' a \Adj.\ a \\<^sub>D f' \\<^sub>D \\ using a \ f' par Adj.\.preserves_hom by (unfold_locales, fastforce) have "\a.is_coext b' (Adj.\ a \\<^sub>D f \\<^sub>D \) g" unfolding \a.is_coext_def using g 1 by auto moreover have "\a.is_coext b' (Adj.\ a \\<^sub>D f' \\<^sub>D \) g'" unfolding \a.is_coext_def using g' 1 by auto ultimately have "g = g'" using 1 f\.arrow_from_functor_axioms f'\.arrow_from_functor_axioms \a.the_coext_unique \a.the_coext_unique [of b' "Adj.\ a \\<^sub>D f' \\<^sub>D \" g'] by auto hence "f \\<^sub>D \ = f' \\<^sub>D \" using g g' is_faithful by argo thus "f = f'" using \ f f' par D.iso_is_retraction D.retraction_is_epi D.epiE [of \ f f'] by auto qed qed ultimately show "D.iso (Adj.\ a)" using D.iso_iff_mono_and_retraction by simp qed interpret \: natural_isomorphism D D \F o G\ D.map Adj.\ using 1 by (unfold_locales, auto) interpret \F: natural_isomorphism C D \F o G o F\ F \Adj.\ o F\ using \.components_are_iso by (unfold_locales, simp) show "\a. C.ide a \ C.iso (Adj.\ a)" proof - fix a assume a: "C.ide a" have "D.inverse_arrows ((Adj.\ o F) a) ((F o Adj.\) a)" using a \.components_are_iso Adj.\\.triangle_F Adj.\FoF\.map_simp_ide D.section_retraction_of_iso by simp hence "D.iso ((F o Adj.\) a)" by blast thus "C.iso (Adj.\ a)" using a reflects_iso [of "Adj.\ a"] by fastforce qed qed (* * Uggh, I should have started with "right_adjoint_functor D C G" so that the * following would come out right. Instead, another step is needed to dualize. * TODO: Maybe re-work this later. *) interpret adjoint_equivalence D C F G Adj.\ Adj.\ .. interpret \': inverse_transformation D D \F o G\ D.map Adj.\ .. interpret \': inverse_transformation C C C.map \G o F\ Adj.\ .. interpret E: adjoint_equivalence C D G F \'.map \'.map using adjoint_equivalence_axioms dual_equivalence by blast show ?thesis using E.adjoint_equivalence_axioms by auto qed lemma is_right_adjoint_functor: shows "right_adjoint_functor C D F" proof - obtain G \ \ where E: "adjoint_equivalence C D G F \ \" using extends_to_adjoint_equivalence by auto interpret E: adjoint_equivalence C D G F \ \ using E by simp interpret Adj: meta_adjunction C D G F E.\ E.\ using E.induces_meta_adjunction by simp show ?thesis using Adj.has_right_adjoint_functor by simp qed lemma is_equivalence_functor: shows "equivalence_functor C D F" proof obtain G \ \ where E: "adjoint_equivalence C D G F \ \" using extends_to_adjoint_equivalence by auto interpret E: adjoint_equivalence C D G F \ \ using E by simp have "equivalence_of_categories C D G F \ \" .. thus "\G \ \. equivalence_of_categories C D G F \ \" by blast qed sublocale equivalence_functor C D F using is_equivalence_functor by blast end context equivalence_of_categories begin text \ The following development shows that an equivalence of categories can be refined to an adjoint equivalence by replacing just the counit. \ abbreviation \' where "\' a \ \ a \\<^sub>C F (D.inv (\ (G a))) \\<^sub>C C.inv (\ (F (G a)))" interpretation \': transformation_by_components C C \F \ G\ C.map \' proof show "\a. C.ide a \ \\' a : (F \ G) a \\<^sub>C C.map a\" using \.components_are_iso \.components_are_iso by simp fix f assume f: "C.arr f" show "\' (C.cod f) \\<^sub>C (F \ G) f = C.map f \\<^sub>C \' (C.dom f)" proof - have "\' (C.cod f) \\<^sub>C (F \ G) f = \ (C.cod f) \\<^sub>C F (D.inv (\ (G (C.cod f)))) \\<^sub>C C.inv (\ (F (G (C.cod f)))) \\<^sub>C F (G f)" using f C.comp_assoc by simp also have "... = \ (C.cod f) \\<^sub>C (F (D.inv (\ (G (C.cod f)))) \\<^sub>C F (G (F (G f)))) \\<^sub>C C.inv (\ (F (G (C.dom f))))" using f \.inv_naturality [of "F (G f)"] C.comp_assoc by simp also have "... = (\ (C.cod f) \\<^sub>C F (G f)) \\<^sub>C F (D.inv (\ (G (C.dom f)))) \\<^sub>C C.inv (\ (F (G (C.dom f))))" proof - have "F (D.inv (\ (G (C.cod f)))) \\<^sub>C F (G (F (G f))) = F (G f) \\<^sub>C F (D.inv (\ (G (C.dom f))))" proof - have "F (D.inv (\ (G (C.cod f)))) \\<^sub>C F (G (F (G f))) = F (D.inv (\ (G (C.cod f))) \\<^sub>D G (F (G f)))" using f by simp also have "... = F (G f \\<^sub>D D.inv (\ (G (C.dom f))))" using f \.inv_naturality [of "G f"] by simp also have "... = F (G f) \\<^sub>C F (D.inv (\ (G (C.dom f))))" using f by simp finally show ?thesis by blast qed thus ?thesis using C.comp_assoc by simp qed also have "... = C.map f \\<^sub>C \ (C.dom f) \\<^sub>C F (D.inv (\ (G (C.dom f)))) \\<^sub>C C.inv (\ (F (G (C.dom f))))" using f \.naturality C.comp_assoc by simp finally show ?thesis by blast qed qed interpretation \': natural_isomorphism C C \F \ G\ C.map \'.map proof show "\a. C.ide a \ C.iso (\'.map a)" unfolding \'.map_def using \.components_are_iso \.components_are_iso apply simp by (intro C.isos_compose) auto qed lemma F\_inverse: assumes "D.ide b" shows "F (\ (G (F b))) = F (G (F (\ b)))" and "F (\ b) \\<^sub>C \ (F b) = \ (F (G (F b))) \\<^sub>C F (\ (G (F b)))" and "C.inverse_arrows (F (\ b)) (\' (F b))" and "F (\ b) = C.inv (\' (F b))" and "C.inv (F (\ b)) = \' (F b)" proof - let ?\' = "\a. \ a \\<^sub>C F (D.inv (\ (G a))) \\<^sub>C C.inv (\ (F (G a)))" show 1: "F (\ (G (F b))) = F (G (F (\ b)))" proof - have "F (\ (G (F b))) \\<^sub>C F (\ b) = F (G (F (\ b))) \\<^sub>C F (\ b)" proof - have "F (\ (G (F b))) \\<^sub>C F (\ b) = F (\ (G (F b)) \\<^sub>D \ b)" using assms by simp also have "... = F (G (F (\ b)) \\<^sub>D \ b)" using assms \.naturality [of "\ b"] by simp also have "... = F (G (F (\ b))) \\<^sub>C F (\ b)" using assms by simp finally show ?thesis by blast qed thus ?thesis using assms \.components_are_iso C.iso_cancel_right by simp qed show "F (\ b) \\<^sub>C \ (F b) = \ (F (G (F b))) \\<^sub>C F (\ (G (F b)))" using assms 1 \.naturality [of "F (\ b)"] by simp show 2: "C.inverse_arrows (F (\ b)) (?\' (F b))" proof show 3: "C.ide (?\' (F b) \\<^sub>C F (\ b))" proof - have "?\' (F b) \\<^sub>C F (\ b) = \ (F b) \\<^sub>C (F (D.inv (\ (G (F b)))) \\<^sub>C C.inv (\ (F (G (F b))))) \\<^sub>C F (\ b)" using C.comp_assoc by simp also have "... = \ (F b) \\<^sub>C (F (D.inv (\ (G (F b)))) \\<^sub>C F (G (F (\ b)))) \\<^sub>C C.inv (\ (F b))" using assms \.naturality [of "F (\ b)"] \.components_are_iso C.comp_assoc C.invert_opposite_sides_of_square [of "\ (F (G (F b)))" "F (G (F (\ b)))" "F (\ b)" "\ (F b)"] by simp also have "... = \ (F b) \\<^sub>C C.inv (\ (F b))" proof - have "F (D.inv (\ (G (F b)))) \\<^sub>C F (G (F (\ b))) = F (G (F b))" using assms 1 D.comp_inv_arr' \.components_are_iso by (metis D.ideD(1) D.ideD(2) F.preserves_comp F.preserves_ide G.preserves_ide \.preserves_dom D.map_simp) moreover have "F (G (F b)) \\<^sub>C C.inv (\ (F b)) = C.inv (\ (F b))" using assms D.comp_cod_arr \.components_are_iso C.inv_in_hom [of "\ (F b)"] by (metis C.comp_ide_arr C_arr_expansion(1) D.ide_char F.preserves_arr F.preserves_dom F.preserves_ide G.preserves_ide C.seqE) ultimately show ?thesis by simp qed also have "... = F b" using assms \.components_are_iso C.comp_arr_inv' by simp finally have "(\ (F b) \\<^sub>C F (D.inv (\ (G (F b)))) \\<^sub>C C.inv (\ (F (G (F b))))) \\<^sub>C F (\ b) = F b" by blast thus ?thesis using assms by simp qed show "C.ide (F (\ b) \\<^sub>C ?\' (F b))" proof - have "(F (\ b) \\<^sub>C ?\' (F b)) \\<^sub>C F (\ b) = F (G (F b)) \\<^sub>C F (\ b)" proof - have "(F (\ b) \\<^sub>C ?\' (F b)) \\<^sub>C F (\ b) = F (\ b) \\<^sub>C (\ (F b) \\<^sub>C F (D.inv (\ (G (F b)))) \\<^sub>C C.inv (\ (F (G (F b))))) \\<^sub>C F (\ b)" using C.comp_assoc by simp also have "... = F (\ b)" using assms 3 C.comp_arr_dom [of "F (\ b)" "(\ (F b) \\<^sub>C F (D.inv (\ (G (F b)))) \\<^sub>C C.inv (\ (F (G (F b))))) \\<^sub>C F (\ b)"] by auto also have "... = F (G (F b)) \\<^sub>C F (\ b)" using assms C.comp_cod_arr by simp finally show ?thesis by blast qed hence "F (\ b) \\<^sub>C ?\' (F b) = F (G (F b))" using assms C.iso_cancel_right by simp thus ?thesis using assms by simp qed qed show "C.inv (F (\ b)) = ?\' (F b)" using assms 2 C.inverse_unique by simp show "F (\ b) = C.inv (?\' (F b))" proof - have "C.inverse_arrows (?\' (F b)) (F (\ b))" using assms 2 by auto thus ?thesis using assms C.inverse_unique by simp qed qed interpretation FoGoF: composite_functor D C C F \F o G\ .. interpretation GoFoG: composite_functor C D D G \G o F\ .. interpretation natural_transformation D C F FoGoF.map \F \ \\ proof - have "F \ D.map = F" using hcomp_ide_dom F.as_nat_trans.natural_transformation_axioms by blast moreover have "F o (G o F) = FoGoF.map" by auto ultimately show "natural_transformation D C F FoGoF.map (F \ \)" using \.natural_transformation_axioms F.as_nat_trans.natural_transformation_axioms horizontal_composite [of D D D.map "G o F" \ C F F F] by simp qed interpretation natural_transformation C D G GoFoG.map \\ \ G\ using \.natural_transformation_axioms G.as_nat_trans.natural_transformation_axioms horizontal_composite [of C D G G G ] by fastforce interpretation natural_transformation D C FoGoF.map F \\'.map \ F\ using \'.natural_transformation_axioms F.as_nat_trans.natural_transformation_axioms horizontal_composite [of D C F F F] by fastforce interpretation natural_transformation C D GoFoG.map G \G \ \'.map\ proof - have "G o C.map = G" using hcomp_ide_dom G.as_nat_trans.natural_transformation_axioms by blast moreover have "G o (F o G) = GoFoG.map" by auto ultimately show "natural_transformation C D GoFoG.map G (G \ \'.map)" using G.as_nat_trans.natural_transformation_axioms \'.natural_transformation_axioms horizontal_composite [of C C "F o G" C.map \'.map D G G G] by simp qed interpretation \'F_F\: vertical_composite D C F FoGoF.map F \F \ \\ \\'.map \ F\ .. interpretation G\'_\G: vertical_composite C D G GoFoG.map G \\ o G\ \G o \'.map\ .. interpretation \\': unit_counit_adjunction C D F G \ \'.map proof show 1: "\'F_F\.map = F" proof fix g show "\'F_F\.map g = F g" proof (cases "D.arr g") show "\ D.arr g \ \'F_F\.map g = F g" using \'F_F\.is_extensional F.is_extensional by simp assume g: "D.arr g" have "\'F_F\.map g = \' (F (D.cod g)) \\<^sub>C F (\ g)" using g \'F_F\.map_def by simp also have "... = \' (F (D.cod g)) \\<^sub>C F (\ (D.cod g) \\<^sub>D g)" using g \.is_natural_2 by simp also have "... = (\' (F (D.cod g)) \\<^sub>C F (\ (D.cod g))) \\<^sub>C F g" using g C.comp_assoc by simp also have "... = F (D.cod g) \\<^sub>C F g" using g F\_inverse(3) [of "D.cod g"] by fastforce also have "... = F g" using g C.comp_cod_arr by simp finally show "\'F_F\.map g = F g" by blast qed qed show "G\'_\G.map = G" proof fix f show "G\'_\G.map f = G f" proof (cases "C.arr f") show "\ C.arr f \ G\'_\G.map f = G f" using G\'_\G.is_extensional G.is_extensional by simp assume f: "C.arr f" have "F (G\'_\G.map f) = F (G (\' (C.cod f)) \\<^sub>D \ (G f))" using f G\'_\G.map_def D.comp_assoc by simp also have "... = F (G (\' (C.cod f)) \\<^sub>D \ (G (C.cod f)) \\<^sub>D G f)" using f \.is_natural_2 [of "G f"] by simp also have "... = F (G (\' (C.cod f))) \\<^sub>C F (\ (G (C.cod f))) \\<^sub>C F (G f)" using f by simp also have "... = (F (G (\' (C.cod f))) \\<^sub>C C.inv (\' (F (G (C.cod f))))) \\<^sub>C F (G f)" using f F\_inverse(4) C.comp_assoc by simp also have "... = (C.inv (\' (C.cod f)) \\<^sub>C \' (C.cod f)) \\<^sub>C F (G f)" using f \'.inv_naturality [of "\' (C.cod f)"] by simp also have "... = F (G (C.cod f)) \\<^sub>C F (G f)" using f C.comp_inv_arr' [of "\' (C.cod f)"] \'.components_are_iso by simp also have "... = F (G f)" using f C.comp_cod_arr by simp finally have "F (G\'_\G.map f) = F (G f)" by blast moreover have "D.par (G\'_\G.map f) (G f)" using f by simp ultimately show "G\'_\G.map f = G f" using f F_is_faithful by (simp add: faithful_functor_axioms_def faithful_functor_def) qed qed qed interpretation \\': adjoint_equivalence C D F G \ \'.map .. lemma refines_to_adjoint_equivalence: shows "adjoint_equivalence C D F G \ \'.map" .. end end diff --git a/thys/Category3/HFSetCat.thy b/thys/Category3/HFSetCat.thy --- a/thys/Category3/HFSetCat.thy +++ b/thys/Category3/HFSetCat.thy @@ -1,2224 +1,2259 @@ (* Title: HFSetCat Author: Eugene W. Stark , 2020 Maintainer: Eugene W. Stark *) chapter "The Category of Hereditarily Finite Sets" theory HFSetCat imports CategoryWithFiniteLimits CartesianClosedCategory HereditarilyFinite.HF begin text\ This theory constructs a category whose objects are in bijective correspondence with the hereditarily finite sets and whose arrows correspond to the functions between such sets. We show that this category is cartesian closed and has finite limits. Note that up to this point we have not constructed any other interpretation for the @{locale cartesian_closed_category} locale, but it is important to have one to ensure that the locale assumptions are consistent. \ section "Preliminaries" text\ We begin with some preliminary definitions and facts about hereditarily finite sets, which are better targeted toward what we are trying to do here than what already exists in @{theory HereditarilyFinite.HF}. \ text\ The following defines when a hereditarily finite set \F\ represents a function from a hereditarily finite set \B\ to a hereditarily finite set \C\. Specifically, \F\ must be a relation from \B\ to \C\, whose domain is \B\, whose range is contained in \C\, and which is single-valued on its domain. \ definition hfun where "hfun B C F \ F \ B * C \ hfunction F \ hdomain F = B \ hrange F \ C" lemma hfunI [intro]: assumes "F \ A * B" and "\X. X \<^bold>\ A \ \!Y. \X, Y\ \<^bold>\ F" and "\X Y. \X, Y\ \<^bold>\ F \ Y \<^bold>\ B" shows "hfun A B F" unfolding hfun_def using assms hfunction_def hrelation_def is_hpair_def hrange_def hconverse_def hdomain_def apply (intro conjI) apply auto by fast lemma hfunE [elim]: assumes "hfun B C F" and "(\Y. Y \<^bold>\ B \ (\!Z. \Y, Z\ \<^bold>\ F) \ (\Z. \Y, Z\ \<^bold>\ F \ Z \<^bold>\ C)) \ T" shows T proof - have "\Y. Y \<^bold>\ B \ (\!Z. \Y, Z\ \<^bold>\ F) \ (\Z. \Y, Z\ \<^bold>\ F \ Z \<^bold>\ C)" proof (intro allI impI conjI) fix Y assume Y: "Y \<^bold>\ B" show "\!Z. \Y, Z\ \<^bold>\ F" proof - have "\Z. \Y, Z\ \<^bold>\ F" using assms Y hfun_def hdomain_def by auto moreover have "\Z Z'. \ \Y, Z\ \<^bold>\ F; \Y, Z'\ \<^bold>\ F \ \ Z = Z'" using assms hfun_def hfunction_def by simp ultimately show ?thesis by blast qed show "\Z. \Y, Z\ \<^bold>\ F \ Z \<^bold>\ C" using assms Y hfun_def by auto qed thus ?thesis using assms(2) by simp qed text\ The hereditarily finite set \hexp B C\ represents the collection of all functions from \B\ to \C\. \ definition hexp where "hexp B C = \F \<^bold>\ HPow (B * C). hfun B C F\" lemma hfun_in_hexp: assumes "hfun B C F" shows "F \<^bold>\ hexp B C" using assms by (simp add: hexp_def hfun_def) text\ The function \happ\ applies a function \F\ from \B\ to \C\ to an element of \B\, yielding an element of \C\. \ abbreviation happ where "happ \ app" lemma happ_mapsto: assumes "F \<^bold>\ hexp B C" and "Y \<^bold>\ B" shows "happ F Y \<^bold>\ C" and "happ F Y \<^bold>\ hrange F" proof - show "happ F Y \<^bold>\ C" using assms app_def hexp_def app_equality hdomain_def hfun_def by auto show "happ F Y \<^bold>\ hrange F" proof - have "\Y, happ F Y\ \<^bold>\ F" using assms app_def hexp_def app_equality hdomain_def hfun_def by auto thus ?thesis using hdomain_def hrange_def hconverse_def by auto qed qed lemma happ_expansion: assumes "hfun B C F" shows "F = \XY \<^bold>\ B * C. hsnd XY = happ F (hfst XY)\" proof fix XY show "XY \<^bold>\ F \ XY \<^bold>\ \XY \<^bold>\ B * C. hsnd XY = happ F (hfst XY)\" proof show "XY \<^bold>\ F \ XY \<^bold>\ \XY \<^bold>\ B * C. hsnd XY = happ F (hfst XY)\" proof - assume XY: "XY \<^bold>\ F" have "XY \<^bold>\ B * C" using assms XY hfun_def by auto moreover have "hsnd XY = happ F (hfst XY)" using assms XY hfunE app_def [of F "hfst XY"] the1_equality [of "\y. \hfst XY, y\ \<^bold>\ F"] calculation by auto ultimately show "XY \<^bold>\ \XY \<^bold>\ B * C. hsnd XY = happ F (hfst XY)\" by simp qed show "XY \<^bold>\ \XY \<^bold>\ B * C. hsnd XY = happ F (hfst XY)\ \ XY \<^bold>\ F" proof - assume XY: "XY \<^bold>\ \XY \<^bold>\ B * C. hsnd XY = happ F (hfst XY)\" show "XY \<^bold>\ F" using assms XY app_def [of F "hfst XY"] the1_equality [of "\y. \hfst XY, y\ \<^bold>\ F"] by fastforce qed qed qed text\ Function \hlam\ takes a function \F\ from \A * B\ to \C\ to a function \hlam F\ from \A\ to \hexp B C\. \ definition hlam where "hlam A B C F = \XG \<^bold>\ A * hexp B C. \YZ. YZ \<^bold>\ hsnd XG \ is_hpair YZ \ \\hfst XG, hfst YZ\, hsnd YZ\ \<^bold>\ F\" lemma hfun_hlam: assumes "hfun (A * B) C F" shows "hfun A (hexp B C) (hlam A B C F)" proof show "hlam A B C F \ A * hexp B C" using assms hlam_def by auto show "\X. X \<^bold>\ A \ \!Y. \X, Y\ \<^bold>\ hlam A B C F" proof fix X assume X: "X \<^bold>\ A" let ?G = "\YZ \<^bold>\ B * C. \\X, hfst YZ\, hsnd YZ\ \<^bold>\ F\" have 1: "?G \<^bold>\ hexp B C" using assms X hexp_def by fastforce show "\X, ?G\ \<^bold>\ hlam A B C F" using assms X 1 is_hpair_def hfun_def hlam_def by auto fix Y assume XY: "\X, Y\ \<^bold>\ hlam A B C F" show "Y = ?G" using assms X XY hlam_def hexp_def by fastforce qed show "\X Y. \X, Y\ \<^bold>\ hlam A B C F \ Y \<^bold>\ hexp B C" using assms hlam_def hexp_def by simp qed lemma happ_hlam: assumes "X \<^bold>\ A" and "hfun (A * B) C F" shows "\!G. \X, G\ \<^bold>\ hlam A B C F" and "happ (hlam A B C F) X = (THE G. \X, G\ \<^bold>\ hlam A B C F)" and "happ (hlam A B C F) X = \yz \<^bold>\ B * C. \\X, hfst yz\, hsnd yz\ \<^bold>\ F\" and "Y \<^bold>\ B \ happ (happ (hlam A B C F) X) Y = happ F \X, Y\" proof - show 1: "\!G. \X, G\ \<^bold>\ hlam A B C F" using assms(1,2) hfun_hlam hfunE by (metis (full_types)) show 2: "happ (hlam A B C F) X = (THE G. \X, G\ \<^bold>\ hlam A B C F)" using assms app_def by simp show "happ (happ (hlam A B C F) X) Y = happ F \X, Y\" proof - have 3: "\X, happ (hlam A B C F) X\ \<^bold>\ hlam A B C F" using assms(1) 1 2 theI' [of "\G. \X, G\ \<^bold>\ hlam A B C F"] by simp hence "\!Z. happ (happ (hlam A B C F) X) = Z" by simp moreover have "happ (happ (hlam A B C F) X) Y = happ F \X, Y\" using assms(1-2) 3 hlam_def is_hpair_def app_def by simp ultimately show ?thesis by simp qed show "happ (hlam A B C F) X = \YZ \<^bold>\ B * C. \\X, hfst YZ\, hsnd YZ\ \<^bold>\ F\" proof - let ?G = "\YZ \<^bold>\ B * C. \\X, hfst YZ\, hsnd YZ\ \<^bold>\ F\" have 4: "hfun B C ?G" proof show "\YZ \<^bold>\ B * C. \\X, hfst YZ\, hsnd YZ\ \<^bold>\ F\ \ B * C" using assms by auto show "\Y. Y \<^bold>\ B \ \!Z. \Y, Z\ \<^bold>\ \YZ \<^bold>\ B * C. \\X, hfst YZ\, hsnd YZ\ \<^bold>\ F\" proof - fix Y assume Y: "Y \<^bold>\ B" have XY: "\X, Y\ \<^bold>\ A * B" using assms Y by simp hence 1: "\!Z. \\X, Y\, Z\ \<^bold>\ F" using assms XY hfunE [of "A * B" C F] by metis obtain Z where Z: "\\X, Y\, Z\ \<^bold>\ F" using 1 by auto have "\Z. \Y, Z\ \<^bold>\ \YZ \<^bold>\ B * C. \\X, hfst YZ\, hsnd YZ\ \<^bold>\ F\" proof - have "\Y, Z\ \<^bold>\ B * C" using assms Y Z by blast moreover have "\\X, hfst \Y, Z\\, hsnd \Y, Z\\ \<^bold>\ F" using assms Y Z by simp ultimately show ?thesis by auto qed moreover have "\Z Z'. \\Y, Z\ \<^bold>\ \YZ \<^bold>\ B * C. \\X, hfst YZ\, hsnd YZ\ \<^bold>\ F\; \Y, Z'\ \<^bold>\ \YZ \<^bold>\ B * C. \\X, hfst YZ\, hsnd YZ\ \<^bold>\ F\\ \ Z = Z'" using assms Y by auto ultimately show "\!Z. \Y, Z\ \<^bold>\ \YZ \<^bold>\ B * C. \\X, hfst YZ\, hsnd YZ\ \<^bold>\ F\" by auto qed show "\Y Z. \Y, Z\ \<^bold>\ \YZ \<^bold>\ B * C. \\X, hfst YZ\, hsnd YZ\ \<^bold>\ F\ \ Z \<^bold>\ C" using assms by simp qed have "\X, ?G\ \<^bold>\ hlam A B C F" proof - have "\X, ?G\ \<^bold>\ A * hexp B C" using assms 4 by (simp add: hfun_in_hexp) moreover have "\YZ. YZ \<^bold>\ ?G \ is_hpair YZ \ \\X, hfst YZ\, hsnd YZ\ \<^bold>\ F" using assms 1 is_hpair_def hfun_def by auto ultimately show ?thesis using assms 1 hlam_def by simp qed thus "happ (hlam A B C F) X = ?G" using assms 2 4 app_equality hfun_def hfun_hlam by auto qed qed section "Construction of the Category" locale hfsetcat begin text\ - We construct the category of hereditarily finite sets and function simply by applying + We construct the category of hereditarily finite sets and functions simply by applying the generic ``set category'' construction, using the hereditarily finite sets as the universe, and constraining the collections of such sets that determine objects of the - category to those having cardinality less than that of the natural numbers; - \emph{i.e.}~to those that are finite. + category to those that are finite. \ - interpretation setcat \undefined :: hf\ natLeq - using Field_natLeq natLeq_Card_order - by unfold_locales auto + interpretation setcat \undefined :: hf\ finite + using finite_subset + by unfold_locales blast+ + + lemma set_ide_char: + shows "A \ set ` Collect ide \ A \ Univ \ finite A" + proof + assume A: "A \ set ` Collect ide" + show "A \ Univ \ finite A" + proof + show "A \ Univ" + using A setp_set' by auto + obtain a where a: "ide a \ A = set a" + using A by blast + have "finite (elem_of ` set a)" + using a setp_set_ide by blast + moreover have "inj_on elem_of (set a)" + proof - + have "inj_on elem_of Univ" + using bij_elem_of bij_betw_imp_inj_on by auto + moreover have "set a \ Univ" + using a setp_set' [of a] by blast + ultimately show ?thesis + using inj_on_subset by auto + qed + ultimately show "finite A" + using a A finite_imageD [of elem_of "set a"] by blast + qed + next + assume A: "A \ Univ \ finite A" + have "ide (mkIde A)" + using A ide_mkIde by simp + moreover have "set (mkIde A) = A" + using A finite_imp_setp set_mkIde by presburger + ultimately show "A \ set ` Collect ide" by blast + qed + + lemma set_ideD: + assumes "ide a" + shows "set a \ Univ" and "finite (set a)" + using assms set_ide_char by auto + + lemma ide_mkIdeI [intro]: + assumes "A \ Univ" and "finite A" + shows "ide (mkIde A)" and "set (mkIde A) = A" + using assms ide_mkIde set_mkIde by auto interpretation category_with_terminal_object comp using terminal_unity by unfold_locales auto text\ We verify that the objects of HF are indeed in bijective correspondence with the hereditarily finite sets. \ definition ide_to_hf - where "ide_to_hf a = HF (DOWN ` set a)" + where "ide_to_hf a = HF (elem_of ` set a)" definition hf_to_ide - where "hf_to_ide x = mkIde (UP ` hfset x)" + where "hf_to_ide x = mkIde (arr_of ` hfset x)" lemma ide_to_hf_mapsto: shows "ide_to_hf \ Collect ide \ UNIV" by simp lemma hf_to_ide_mapsto: shows "hf_to_ide \ UNIV \ Collect ide" proof fix x :: hf - have "finite (UP ` hfset x)" + have "finite (arr_of ` hfset x)" by simp - moreover have "UP ` hfset x \ Univ" - by (metis (mono_tags, lifting) UNIV_I bij_UP bij_betw_def imageE image_eqI subsetI) - ultimately have "ide (mkIde (UP ` hfset x))" - using ide_mkIde_finite by simp + moreover have "arr_of ` hfset x \ Univ" + by (metis (mono_tags, lifting) UNIV_I bij_arr_of bij_betw_def imageE image_eqI subsetI) + ultimately have "ide (mkIde (arr_of ` hfset x))" + using finite_imp_setp ide_mkIde by presburger thus "hf_to_ide x \ Collect ide" using hf_to_ide_def by simp qed lemma hf_to_ide_ide_to_hf: assumes "a \ Collect ide" shows "hf_to_ide (ide_to_hf a) = a" proof - - have "hf_to_ide (ide_to_hf a) = mkIde (UP ` hfset (HF (DOWN ` set a)))" + have "hf_to_ide (ide_to_hf a) = mkIde (arr_of ` hfset (HF (elem_of ` set a)))" using hf_to_ide_def ide_to_hf_def by simp also have "... = a" proof - - have "mkIde (UP ` hfset (HF (DOWN ` set a))) = mkIde (UP ` DOWN ` set a)" + have "mkIde (arr_of ` hfset (HF (elem_of ` set a))) = mkIde (arr_of ` elem_of ` set a)" proof - have "finite (set a)" - using assms set_card finite_iff_ordLess_natLeq by auto - hence "finite (DOWN ` set a)" + using assms set_ide_char by blast + hence "finite (elem_of ` set a)" by simp - hence "hfset (HF (DOWN ` set a)) = DOWN ` set a" - using hfset_HF [of "DOWN ` set a"] by simp + hence "hfset (HF (elem_of ` set a)) = elem_of ` set a" + using hfset_HF [of "elem_of ` set a"] by simp thus ?thesis by simp qed also have "... = a" proof - have "set a \ Univ" - using assms set_subset_Univ ide_char by blast - hence "\x. x \ set a \ UP (DOWN x) = x" + using assms set_ide_char by blast + hence "\x. x \ set a \ arr_of (elem_of x) = x" using assms by auto - hence "UP ` DOWN ` set a = set a" + hence "arr_of ` elem_of ` set a = set a" by force thus ?thesis using assms ide_char mkIde_set by simp qed finally show ?thesis by blast qed finally show "hf_to_ide (ide_to_hf a) = a" by blast qed lemma ide_to_hf_hf_to_ide: assumes "x \ UNIV" shows "ide_to_hf (hf_to_ide x) = x" proof - - have "HF (DOWN ` set (mkIde (UP ` hfset x))) = x" + have "HF (elem_of ` set (mkIde (arr_of ` hfset x))) = x" proof - - have "HF (DOWN ` set (mkIde (UP ` hfset x))) = HF (DOWN ` UP ` hfset x)" - proof - - have "|UP ` hfset x| A. DOWN ` UP ` A = A" - using DOWN_UP by force + have "\A. elem_of ` arr_of ` A = A" + using elem_of_arr_of by force thus ?thesis by metis qed also have "... = x" by simp finally show ?thesis by blast qed thus ?thesis using assms ide_to_hf_def hf_to_ide_def by simp qed lemma bij_betw_ide_hf_set: shows "bij_betw ide_to_hf (Collect ide) (UNIV :: hf set)" using ide_to_hf_mapsto hf_to_ide_mapsto ide_to_hf_hf_to_ide hf_to_ide_ide_to_hf by (intro bij_betwI) auto lemma ide_implies_finite_set: assumes "ide a" shows "finite (set a)" and "finite (hom unity a)" proof - show 1: "finite (set a)" - using assms ide_char set_card finite_iff_ordLess_natLeq by blast + using assms set_ide_char by blast show "finite (hom unity a)" - proof - - have "|hom unity a| =o |set a|" - using assms bij_betw_points_and_set card_of_ordIsoI by auto - thus ?thesis - using 1 by simp - qed + using assms 1 bij_betw_points_and_set finite_imageD inj_img set_def by auto qed text\ We establish the connection between the membership relation defined for hereditarily finite sets and the corresponding membership relation associated with the set category. \ - lemma UP_membI [intro]: + lemma arr_of_membI [intro]: assumes "x \<^bold>\ ide_to_hf a" - shows "UP x \ set a" + shows "arr_of x \ set a" proof - - let ?X = "inv_into (set a) DOWN x" - have "x = DOWN ?X \ ?X \ set a" + let ?X = "inv_into (set a) elem_of x" + have "x = elem_of ?X \ ?X \ set a" using assms by (simp add: f_inv_into_f ide_to_hf_def inv_into_into) thus ?thesis - by (metis (no_types, lifting) UP_DOWN elem_set_implies_incl_in incl_in_def - set_subset_Univ subsetD) + by (metis (no_types, lifting) arr_of_elem_of elem_set_implies_incl_in + elem_set_implies_set_eq_singleton incl_in_def mem_Collect_eq terminal_char2) qed - lemma DOWN_membI [intro]: + lemma elem_of_membI [intro]: assumes "ide a" and "x \ set a" - shows "DOWN x \<^bold>\ ide_to_hf a" + shows "elem_of x \<^bold>\ ide_to_hf a" proof - - have "finite (DOWN ` set a)" + have "finite (elem_of ` set a)" using assms ide_implies_finite_set [of a] by simp - hence "DOWN x \ hfset (ide_to_hf a)" - using assms ide_to_hf_def hfset_HF [of "DOWN ` set a"] by simp + hence "elem_of x \ hfset (ide_to_hf a)" + using assms ide_to_hf_def hfset_HF [of "elem_of ` set a"] by simp thus ?thesis using hmem_def by blast qed text\ We show that each hom-set \hom a b\ is in bijective correspondence with the elements of the hereditarily finite set \hfun (ide_to_hf a) (ide_to_hf b)\. \ definition arr_to_hfun where "arr_to_hfun f = \XY \<^bold>\ ide_to_hf (dom f) * ide_to_hf (cod f). - hsnd XY = DOWN (Fun f (UP (hfst XY)))\" + hsnd XY = elem_of (Fun f (arr_of (hfst XY)))\" definition hfun_to_arr where "hfun_to_arr B C F = - mkArr (UP ` hfset B) (UP ` hfset C) (\x. UP (happ F (DOWN x)))" + mkArr (arr_of ` hfset B) (arr_of ` hfset C) (\x. arr_of (happ F (elem_of x)))" lemma hfun_arr_to_hfun: assumes "arr f" shows "hfun (ide_to_hf (dom f)) (ide_to_hf (cod f)) (arr_to_hfun f)" proof show "arr_to_hfun f \ ide_to_hf (dom f) * ide_to_hf (cod f)" using assms arr_to_hfun_def by auto show "\X. X \<^bold>\ ide_to_hf (dom f) \ \!Y. \X, Y\ \<^bold>\ arr_to_hfun f" proof fix X assume X: "X \<^bold>\ ide_to_hf (dom f)" - show "\X, DOWN (Fun f (UP X))\ \<^bold>\ arr_to_hfun f" + show "\X, elem_of (Fun f (arr_of X))\ \<^bold>\ arr_to_hfun f" proof - - have "\X, DOWN (Fun f (UP X))\ \<^bold>\ \XY \<^bold>\ ide_to_hf (dom f) * ide_to_hf (cod f). - hsnd XY = DOWN (Fun f (UP (hfst XY)))\" + have "\X, elem_of (Fun f (arr_of X))\ \<^bold>\ \XY \<^bold>\ ide_to_hf (dom f) * ide_to_hf (cod f). + hsnd XY = elem_of (Fun f (arr_of (hfst XY)))\" proof - - have "hsnd \X, DOWN (Fun f (UP X))\ = - DOWN (Fun f (UP (hfst \X, DOWN (Fun f (UP X))\)))" + have "hsnd \X, elem_of (Fun f (arr_of X))\ = + elem_of (Fun f (arr_of (hfst \X, elem_of (Fun f (arr_of X))\)))" using assms X by simp - moreover have "\X, DOWN (Fun f (UP X))\ \<^bold>\ ide_to_hf (dom f) * ide_to_hf (cod f)" + moreover have "\X, elem_of (Fun f (arr_of X))\ \<^bold>\ ide_to_hf (dom f) * ide_to_hf (cod f)" proof - - have "DOWN (Fun f (UP X)) \<^bold>\ ide_to_hf (cod f)" - proof (intro DOWN_membI) + have "elem_of (Fun f (arr_of X)) \<^bold>\ ide_to_hf (cod f)" + proof (intro elem_of_membI) show "ide (cod f)" using assms ide_cod by simp - show "Fun f (UP X) \ Cod f" - using assms X Fun_mapsto UP_membI by auto + show "Fun f (arr_of X) \ Cod f" + using assms X Fun_mapsto arr_of_membI by auto qed thus ?thesis using X by simp qed ultimately show ?thesis by simp qed thus ?thesis using arr_to_hfun_def by simp qed fix Y assume XY: "\X, Y\ \<^bold>\ arr_to_hfun f" - show "Y = DOWN (Fun f (UP X))" + show "Y = elem_of (Fun f (arr_of X))" using assms X XY arr_to_hfun_def by auto qed show "\X Y. \X, Y\ \<^bold>\ arr_to_hfun f \ Y \<^bold>\ ide_to_hf (cod f)" using assms arr_to_hfun_def ide_to_hf_def \arr_to_hfun f \ ide_to_hf (dom f) * ide_to_hf (cod f)\ by blast qed lemma arr_to_hfun_in_hexp: assumes "arr f" shows "arr_to_hfun f \<^bold>\ hexp (ide_to_hf (dom f)) (ide_to_hf (cod f))" using assms arr_to_hfun_def hfun_arr_to_hfun hexp_def by auto lemma hfun_to_arr_in_hom: assumes "hfun B C F" shows "\hfun_to_arr B C F : hf_to_ide B \ hf_to_ide C\" proof - let ?f = "mkArr (UP ` hfset B) (UP ` hfset C) (\x. UP (happ F (DOWN x)))" + let ?f = "mkArr (arr_of ` hfset B) (arr_of ` hfset C) (\x. arr_of (happ F (elem_of x)))" have 0: "arr ?f" proof - - have "UP ` hfset B \ Univ \ UP ` hfset C \ Univ" - using UP_mapsto by auto - moreover have "(\x. UP (happ F (DOWN x))) \ UP ` hfset B \ UP ` hfset C" + have "arr_of ` hfset B \ Univ \ arr_of ` hfset C \ Univ" + using arr_of_mapsto by auto + moreover have "(\x. arr_of (happ F (elem_of x))) \ arr_of ` hfset B \ arr_of ` hfset C" proof fix x - assume x: "x \ UP ` hfset B" - have "happ F (DOWN x) \ hfset C" + assume x: "x \ arr_of ` hfset B" + have "happ F (elem_of x) \ hfset C" using assms x happ_mapsto hfun_in_hexp - by (metis DOWN_UP HF_hfset finite_hfset hmem_HF_iff imageE) - thus "UP (happ F (DOWN x)) \ UP ` hfset C" + by (metis elem_of_arr_of HF_hfset finite_hfset hmem_HF_iff imageE) + thus "arr_of (happ F (elem_of x)) \ arr_of ` hfset C" by simp qed ultimately show ?thesis using arr_mkArr by (meson finite_hfset finite_iff_ordLess_natLeq finite_imageI) qed show 1: "arr (hfun_to_arr B C F)" using 0 hfun_to_arr_def by simp show "dom (hfun_to_arr B C F) = hf_to_ide B" using 1 hfun_to_arr_def hf_to_ide_def dom_mkArr by auto show "cod (hfun_to_arr B C F) = hf_to_ide C" using 1 hfun_to_arr_def hf_to_ide_def cod_mkArr by auto qed text\ The comprehension notation from @{theory HereditarilyFinite.HF} interferes in an unfortunate way with the restriction notation from @{theory "HOL-Library.FuncSet"}, making it impossible to use both in the present context. \ lemma Fun_char: assumes "arr f" - shows "Fun f = restrict (\x. UP (happ (arr_to_hfun f) (DOWN x))) (Dom f)" + shows "Fun f = restrict (\x. arr_of (happ (arr_to_hfun f) (elem_of x))) (Dom f)" proof fix x - show "Fun f x = restrict (\x. UP (happ (arr_to_hfun f) (DOWN x))) (Dom f) x" + show "Fun f x = restrict (\x. arr_of (happ (arr_to_hfun f) (elem_of x))) (Dom f) x" proof (cases "x \ Dom f") show "x \ Dom f \ ?thesis" using assms Fun_mapsto Fun_def restrict_apply by simp show "x \ Dom f \ ?thesis" proof - assume x: "x \ Dom f" have 1: "hfun (ide_to_hf (dom f)) (ide_to_hf (cod f)) (arr_to_hfun f)" using assms app_def arr_to_hfun_def hfun_arr_to_hfun - the1_equality [of "\y. \DOWN x, y\ \<^bold>\ arr_to_hfun f" "DOWN (Fun f x)"] + the1_equality [of "\y. \elem_of x, y\ \<^bold>\ arr_to_hfun f" "elem_of (Fun f x)"] by simp - have 2: "\!Y. \DOWN x, Y\ \<^bold>\ arr_to_hfun f" - using assms x 1 hfunE DOWN_membI ide_dom + have 2: "\!Y. \elem_of x, Y\ \<^bold>\ arr_to_hfun f" + using assms x 1 hfunE elem_of_membI ide_dom by (metis (no_types, lifting)) - have "Fun f x = UP (DOWN (Fun f x))" + have "Fun f x = arr_of (elem_of (Fun f x))" proof - have "Fun f x \ Univ" - using assms x ide_cod Fun_mapsto [of f] set_subset_Univ by auto + using assms x ide_cod Fun_mapsto [of f] set_ide_char by blast thus ?thesis - using UP_DOWN by simp + using arr_of_elem_of by simp qed - also have "... = UP (happ (arr_to_hfun f) (DOWN x))" + also have "... = arr_of (happ (arr_to_hfun f) (elem_of x))" proof - - have "\DOWN x, DOWN (Fun f x)\ \<^bold>\ arr_to_hfun f" - using assms x 2 ide_dom arr_to_hfun_def set_subset_Univ UP_DOWN - by (metis (mono_tags, lifting) HCollectE hfst_conv hsnd_conv subsetD) - moreover have "\DOWN x, happ (arr_to_hfun f) (DOWN x)\ \<^bold>\ arr_to_hfun f" + have "\elem_of x, elem_of (Fun f x)\ \<^bold>\ arr_to_hfun f" + proof - + have "\elem_of x, elem_of (Fun f x)\ \<^bold>\ ide_to_hf (dom f) * ide_to_hf (cod f)" + using assms x ide_dom ide_cod Fun_mapsto by fast + moreover have "elem_of (Fun f x) = elem_of (Fun f (arr_of (elem_of x)))" + by (metis (no_types, lifting) arr_of_elem_of setp_set_ide assms ide_dom subsetD x) + ultimately show ?thesis + using arr_to_hfun_def by auto + qed + moreover have "\elem_of x, happ (arr_to_hfun f) (elem_of x)\ \<^bold>\ arr_to_hfun f" using assms x 1 2 app_equality hfun_def by blast ultimately show ?thesis using 2 by fastforce qed - also have "... = restrict (\x. UP (happ (arr_to_hfun f) (DOWN x))) (Dom f) x" + also have "... = restrict (\x. arr_of (happ (arr_to_hfun f) (elem_of x))) (Dom f) x" using assms x ide_dom by auto finally show ?thesis by simp qed qed qed lemma Fun_hfun_to_arr: assumes "hfun B C F" - shows "Fun (hfun_to_arr B C F) = restrict (\x. UP (happ F (DOWN x))) (UP ` hfset B)" + shows "Fun (hfun_to_arr B C F) = restrict (\x. arr_of (happ F (elem_of x))) (arr_of ` hfset B)" proof - have "arr (hfun_to_arr B C F)" using assms hfun_to_arr_in_hom by blast - hence "arr (mkArr (UP ` hfset B) (UP ` hfset C) (\x. UP (happ F (DOWN x))))" + hence "arr (mkArr (arr_of ` hfset B) (arr_of ` hfset C) (\x. arr_of (happ F (elem_of x))))" using hfun_to_arr_def by simp thus ?thesis using assms hfun_to_arr_def Fun_mkArr by simp qed - lemma UP_img_hfset_ide_to_hf: + lemma arr_of_img_hfset_ide_to_hf: assumes "ide a" - shows "UP ` hfset (ide_to_hf a) = set a" + shows "arr_of ` hfset (ide_to_hf a) = set a" proof - - have "UP ` hfset (ide_to_hf a) = UP ` hfset (HF (DOWN ` set a))" + have "arr_of ` hfset (ide_to_hf a) = arr_of ` hfset (HF (elem_of ` set a))" using ide_to_hf_def by simp - also have "... = UP ` DOWN ` set a" + also have "... = arr_of ` elem_of ` set a" using assms ide_implies_finite_set(1) ide_char by auto also have "... = set a" proof - - have "\x. x \ set a \ UP (DOWN x) = x" - using assms ide_char - by (metis (no_types, lifting) UP_DOWN set_subset_Univ subsetD) + have "\x. x \ set a \ arr_of (elem_of x) = x" + using assms ide_char arr_of_elem_of setp_set_ide by blast thus ?thesis by force qed finally show ?thesis by blast qed lemma hfun_to_arr_arr_to_hfun: assumes "arr f" shows "hfun_to_arr (ide_to_hf (dom f)) (ide_to_hf (cod f)) (arr_to_hfun f) = f" proof - have 0: "hfun_to_arr (ide_to_hf (dom f)) (ide_to_hf (cod f)) (arr_to_hfun f) = - mkArr (UP ` hfset (ide_to_hf (dom f))) (UP ` hfset (ide_to_hf (cod f))) - (\x. UP (happ (arr_to_hfun f) (DOWN x)))" + mkArr (arr_of ` hfset (ide_to_hf (dom f))) (arr_of ` hfset (ide_to_hf (cod f))) + (\x. arr_of (happ (arr_to_hfun f) (elem_of x)))" unfolding hfun_to_arr_def by blast also have "... = mkArr (Dom f) (Cod f) - (restrict (\x. UP (happ (arr_to_hfun f) (DOWN x))) (Dom f))" + (restrict (\x. arr_of (happ (arr_to_hfun f) (elem_of x))) (Dom f))" proof (intro mkArr_eqI) - show 1: "UP ` hfset (ide_to_hf (dom f)) = Dom f" - using assms UP_img_hfset_ide_to_hf ide_dom by simp - show 2: "UP ` hfset (ide_to_hf (cod f)) = Cod f" - using assms UP_img_hfset_ide_to_hf ide_cod by simp - show "arr (mkArr (UP ` hfset (ide_to_hf (dom f))) (UP ` hfset (ide_to_hf (cod f))) - (\x. UP (happ (arr_to_hfun f) (DOWN x))))" + show 1: "arr_of ` hfset (ide_to_hf (dom f)) = Dom f" + using assms arr_of_img_hfset_ide_to_hf ide_dom by simp + show 2: "arr_of ` hfset (ide_to_hf (cod f)) = Cod f" + using assms arr_of_img_hfset_ide_to_hf ide_cod by simp + show "arr (mkArr (arr_of ` hfset (ide_to_hf (dom f))) (arr_of ` hfset (ide_to_hf (cod f))) + (\x. arr_of (happ (arr_to_hfun f) (elem_of x))))" using 0 1 2 by (metis (no_types, lifting) arrI assms hfun_arr_to_hfun hfun_to_arr_in_hom) - show "\x. x \ UP ` hfset (ide_to_hf (dom f)) \ - UP (happ (arr_to_hfun f) (DOWN x)) = - restrict (\x. UP (happ (arr_to_hfun f) (DOWN x))) (Dom f) x" + show "\x. x \ arr_of ` hfset (ide_to_hf (dom f)) \ + arr_of (happ (arr_to_hfun f) (elem_of x)) = + restrict (\x. arr_of (happ (arr_to_hfun f) (elem_of x))) (Dom f) x" using assms 1 by simp qed also have "... = mkArr (Dom f) (Cod f) (Fun f)" using assms Fun_char mkArr_eqI by simp also have "... = f" using assms mkArr_Fun by blast finally show ?thesis by simp qed lemma arr_to_hfun_hfun_to_arr: assumes "hfun B C F" shows "arr_to_hfun (hfun_to_arr B C F) = F" proof - have "arr_to_hfun (hfun_to_arr B C F) = \XY \<^bold>\ ide_to_hf (dom (hfun_to_arr B C F)) * ide_to_hf (cod (hfun_to_arr B C F)). - hsnd XY = DOWN (Fun (hfun_to_arr B C F) (UP (hfst XY)))\" + hsnd XY = elem_of (Fun (hfun_to_arr B C F) (arr_of (hfst XY)))\" unfolding arr_to_hfun_def by blast also have - "... = \XY \<^bold>\ ide_to_hf (mkIde (UP ` hfset B)) * ide_to_hf (mkIde (UP ` hfset C)). - hsnd XY = DOWN (Fun (hfun_to_arr B C F) (UP (hfst XY)))\" + "... = \XY \<^bold>\ ide_to_hf (mkIde (arr_of ` hfset B)) * ide_to_hf (mkIde (arr_of ` hfset C)). + hsnd XY = elem_of (Fun (hfun_to_arr B C F) (arr_of (hfst XY)))\" using assms hfun_to_arr_in_hom [of B C F] hf_to_ide_def by (metis (no_types, lifting) in_homE) also have - "... = \XY \<^bold>\ ide_to_hf (mkIde (UP ` hfset B)) * ide_to_hf (mkIde (UP ` hfset C)). - hsnd XY = DOWN (restrict (\x. UP (happ F (DOWN x))) (UP ` hfset B) - (UP (hfst XY)))\" + "... = \XY \<^bold>\ ide_to_hf (mkIde (arr_of ` hfset B)) * ide_to_hf (mkIde (arr_of ` hfset C)). + hsnd XY = elem_of (restrict (\x. arr_of (happ F (elem_of x))) (arr_of ` hfset B) + (arr_of (hfst XY)))\" using assms Fun_hfun_to_arr by simp also have - "... = \XY \<^bold>\ ide_to_hf (mkIde (UP ` hfset B)) * ide_to_hf (mkIde (UP ` hfset C)). - hsnd XY = DOWN (UP (happ F (DOWN (UP (hfst XY)))))\" + "... = \XY \<^bold>\ ide_to_hf (mkIde (arr_of ` hfset B)) * ide_to_hf (mkIde (arr_of ` hfset C)). + hsnd XY = elem_of (arr_of (happ F (elem_of (arr_of (hfst XY)))))\" proof - have - 1: "\XY. XY \<^bold>\ ide_to_hf (mkIde (UP ` hfset B)) * ide_to_hf (mkIde (UP ` hfset C)) - \ UP (hfst XY) \ UP ` hfset B" + 1: "\XY. XY \<^bold>\ ide_to_hf (mkIde (arr_of ` hfset B)) * ide_to_hf (mkIde (arr_of ` hfset C)) + \ arr_of (hfst XY) \ arr_of ` hfset B" proof - fix XY assume - XY: "XY \<^bold>\ ide_to_hf (mkIde (UP ` hfset B)) * ide_to_hf (mkIde (UP ` hfset C))" - have "hfst XY \<^bold>\ ide_to_hf (mkIde (UP ` hfset B))" + XY: "XY \<^bold>\ ide_to_hf (mkIde (arr_of ` hfset B)) * ide_to_hf (mkIde (arr_of ` hfset C))" + have "hfst XY \<^bold>\ ide_to_hf (mkIde (arr_of ` hfset B))" using XY by auto - thus "UP (hfst XY) \ UP ` hfset B" - using assms UP_membI [of "hfst XY" "mkIde (UP ` hfset B)"] set_mkIde + thus "arr_of (hfst XY) \ arr_of ` hfset B" + using assms arr_of_membI [of "hfst XY" "mkIde (arr_of ` hfset B)"] set_mkIde by (metis (mono_tags, lifting) arrI arr_mkArr hfun_to_arr_def hfun_to_arr_in_hom) qed show ?thesis proof - have - "\XY. (XY \<^bold>\ ide_to_hf (mkIde (UP ` hfset B)) * ide_to_hf (mkIde (UP ` hfset C)) \ - hsnd XY = DOWN (restrict (\x. UP (happ F (DOWN x))) (UP ` hfset B) - (UP (hfst XY)))) + "\XY. (XY \<^bold>\ ide_to_hf (mkIde (arr_of ` hfset B)) * ide_to_hf (mkIde (arr_of ` hfset C)) \ + hsnd XY = elem_of (restrict (\x. arr_of (happ F (elem_of x))) (arr_of ` hfset B) + (arr_of (hfst XY)))) \ - (XY \<^bold>\ ide_to_hf (mkIde (UP ` hfset B)) * ide_to_hf (mkIde (UP ` hfset C)) \ - hsnd XY = DOWN (UP (happ F (DOWN (UP (hfst XY))))))" + (XY \<^bold>\ ide_to_hf (mkIde (arr_of ` hfset B)) * ide_to_hf (mkIde (arr_of ` hfset C)) \ + hsnd XY = elem_of (arr_of (happ F (elem_of (arr_of (hfst XY))))))" using 1 by auto thus ?thesis by blast qed qed also have - "... = \XY \<^bold>\ ide_to_hf (mkIde (UP ` hfset B)) * ide_to_hf (mkIde (UP ` hfset C)). + "... = \XY \<^bold>\ ide_to_hf (mkIde (arr_of ` hfset B)) * ide_to_hf (mkIde (arr_of ` hfset C)). hsnd XY = happ F (hfst XY)\" by simp also have "... = \XY \<^bold>\ B * C. hsnd XY = happ F (hfst XY)\" using assms hf_to_ide_def ide_to_hf_hf_to_ide by force also have "... = F" using assms happ_expansion by simp finally show ?thesis by simp qed lemma bij_betw_hom_hfun: assumes "ide a" and "ide b" shows "bij_betw arr_to_hfun (hom a b) {F. hfun (ide_to_hf a) (ide_to_hf b) F}" proof (intro bij_betwI) show "arr_to_hfun \ hom a b \ {F. hfun (ide_to_hf a) (ide_to_hf b) F}" using assms arr_to_hfun_in_hexp hexp_def hfun_arr_to_hfun by blast show "hfun_to_arr (ide_to_hf a) (ide_to_hf b) \ {F. hfun (ide_to_hf a) (ide_to_hf b) F} \ hom a b" using assms hfun_to_arr_in_hom by (metis (no_types, lifting) Pi_I hf_to_ide_ide_to_hf mem_Collect_eq) show "\x. x \ hom a b \ hfun_to_arr (ide_to_hf a) (ide_to_hf b) (arr_to_hfun x) = x" using assms hfun_to_arr_arr_to_hfun by blast show "\y. y \ {F. hfun (ide_to_hf a) (ide_to_hf b) F} \ arr_to_hfun (hfun_to_arr (ide_to_hf a) (ide_to_hf b) y) = y" using assms arr_to_hfun_hfun_to_arr by simp qed text\ We next relate composition of arrows in the category to the corresponding operation on hereditarily finite sets. \ definition hcomp where "hcomp G F = \XZ \<^bold>\ hdomain F * hrange G. hsnd XZ = happ G (happ F (hfst XZ))\" lemma hfun_hcomp: assumes "hfun A B F" and "hfun B C G" shows "hfun A C (hcomp G F)" proof show "hcomp G F \ A * C" using assms hcomp_def hfun_def by auto show "\X. X \<^bold>\ A \ \!Y. \X, Y\ \<^bold>\ hcomp G F" proof fix X assume X: "X \<^bold>\ A" show "\X, happ G (happ F X)\ \<^bold>\ hcomp G F" unfolding hcomp_def using assms X hfunE happ_mapsto hfun_in_hexp by (metis (mono_tags, lifting) HCollect_iff hfst_conv hfun_def hsnd_conv timesI) show "\X Y. \X \<^bold>\ A; \X, Y\ \<^bold>\ hcomp G F\ \ Y = happ G (happ F X)" unfolding hcomp_def by simp qed show "\X Y. \X, Y\ \<^bold>\ hcomp G F \ Y \<^bold>\ C" unfolding hcomp_def using assms hfunE happ_mapsto hfun_in_hexp by (metis HCollectE hfun_def hsubsetCE timesD2) qed lemma arr_to_hfun_comp: assumes "seq g f" shows "arr_to_hfun (comp g f) = hcomp (arr_to_hfun g) (arr_to_hfun f)" proof - have 1: "hdomain (arr_to_hfun f) = ide_to_hf (dom f)" using assms hfun_arr_to_hfun hfun_def by blast have "arr_to_hfun (comp g f) = \XZ \<^bold>\ ide_to_hf (dom f) * ide_to_hf (cod g). - hsnd XZ = DOWN (Fun (comp g f) (UP (hfst XZ)))\" + hsnd XZ = elem_of (Fun (comp g f) (arr_of (hfst XZ)))\" unfolding arr_to_hfun_def comp_def using assms by fastforce also have "... = \XZ \<^bold>\ hdomain (arr_to_hfun f) * hrange (arr_to_hfun g). hsnd XZ = happ (arr_to_hfun g) (happ (arr_to_hfun f) (hfst XZ))\" proof fix XZ have "hfst XZ \<^bold>\ hdomain (arr_to_hfun f) \ hsnd XZ \<^bold>\ ide_to_hf (cod g) \ - hsnd XZ = DOWN (Fun (comp g f) (UP (hfst XZ))) + hsnd XZ = elem_of (Fun (comp g f) (arr_of (hfst XZ))) \ hsnd XZ \<^bold>\ hrange (arr_to_hfun g) \ hsnd XZ = happ (arr_to_hfun g) (happ (arr_to_hfun f) (hfst XZ))" proof assume XZ: "hfst XZ \<^bold>\ hdomain (arr_to_hfun f)" - have 2: "UP (hfst XZ) \ Dom f" - using XZ 1 hfsetcat.UP_membI by auto - have 3: "UP (happ (arr_to_hfun f) (hfst XZ)) \ Dom g" + have 2: "arr_of (hfst XZ) \ Dom f" + using XZ 1 hfsetcat.arr_of_membI by auto + have 3: "arr_of (happ (arr_to_hfun f) (hfst XZ)) \ Dom g" using assms XZ 2 - by (metis (no_types, lifting) "1" happ_mapsto(1) hfsetcat.UP_membI + by (metis (no_types, lifting) "1" happ_mapsto(1) hfsetcat.arr_of_membI arr_to_hfun_in_hexp seqE) - have 4: "DOWN (Fun (comp g f) (UP (hfst XZ))) = + have 4: "elem_of (Fun (comp g f) (arr_of (hfst XZ))) = happ (arr_to_hfun g) (happ (arr_to_hfun f) (hfst XZ))" proof - - have "DOWN (Fun (comp g f) (UP (hfst XZ))) = - DOWN (restrict (Fun g o Fun f) (Dom f) (UP (hfst XZ)))" + have "elem_of (Fun (comp g f) (arr_of (hfst XZ))) = + elem_of (restrict (Fun g o Fun f) (Dom f) (arr_of (hfst XZ)))" using assms Fun_comp Fun_char by simp - also have "... = DOWN ((Fun g o Fun f) (UP (hfst XZ)))" + also have "... = elem_of ((Fun g o Fun f) (arr_of (hfst XZ)))" using XZ 2 by auto - also have "... = DOWN (Fun g (Fun f (UP (hfst XZ))))" + also have "... = elem_of (Fun g (Fun f (arr_of (hfst XZ))))" by simp also have - "... = DOWN (Fun g (restrict (\x. UP (happ (arr_to_hfun f) (DOWN x))) (Dom f) - (UP (hfst XZ))))" + "... = elem_of (Fun g (restrict (\x. arr_of (happ (arr_to_hfun f) (elem_of x))) (Dom f) + (arr_of (hfst XZ))))" proof - - have "Fun f = restrict (\x. UP (happ (arr_to_hfun f) (DOWN x))) (Dom f)" + have "Fun f = restrict (\x. arr_of (happ (arr_to_hfun f) (elem_of x))) (Dom f)" using assms Fun_char [of f] by blast thus ?thesis by simp qed - also have "... = DOWN (Fun g (UP (happ (arr_to_hfun f) (hfst XZ))))" + also have "... = elem_of (Fun g (arr_of (happ (arr_to_hfun f) (hfst XZ))))" using 2 by simp - also have "... = DOWN (restrict (\x. UP (happ (arr_to_hfun g) (DOWN x))) (Dom g) - (UP (happ (arr_to_hfun f) (hfst XZ))))" + also have "... = elem_of (restrict (\x. arr_of (happ (arr_to_hfun g) (elem_of x))) (Dom g) + (arr_of (happ (arr_to_hfun f) (hfst XZ))))" proof - - have "Fun g = restrict (\x. UP (happ (arr_to_hfun g) (DOWN x))) (Dom g)" + have "Fun g = restrict (\x. arr_of (happ (arr_to_hfun g) (elem_of x))) (Dom g)" using assms Fun_char [of g] by blast thus ?thesis by simp qed also have "... = happ (arr_to_hfun g) (happ (arr_to_hfun f) (hfst XZ))" using 3 by simp finally show ?thesis by blast qed - have 5: "DOWN (Fun (comp g f) (UP (hfst XZ))) \<^bold>\ hrange (arr_to_hfun g)" + have 5: "elem_of (Fun (comp g f) (arr_of (hfst XZ))) \<^bold>\ hrange (arr_to_hfun g)" proof - have "happ (arr_to_hfun g) (happ (arr_to_hfun f) (hfst XZ)) \<^bold>\ hrange (arr_to_hfun g)" using assms 1 3 XZ hfun_arr_to_hfun happ_mapsto arr_to_hfun_in_hexp arr_to_hfun_def by (metis (no_types, lifting) seqE) thus ?thesis using XZ 4 by simp qed show "hsnd XZ \<^bold>\ ide_to_hf (cod g) \ - hsnd XZ = DOWN (Fun (comp g f) (UP (hfst XZ))) + hsnd XZ = elem_of (Fun (comp g f) (arr_of (hfst XZ))) \ hsnd XZ \<^bold>\ hrange (arr_to_hfun g) \ hsnd XZ = happ (arr_to_hfun g) (happ (arr_to_hfun f) (hfst XZ))" using XZ 4 5 by simp show "hsnd XZ \<^bold>\ hrange (arr_to_hfun g) \ hsnd XZ = happ (arr_to_hfun g) (happ (arr_to_hfun f) (hfst XZ)) \ hsnd XZ \<^bold>\ ide_to_hf (cod g) \ - hsnd XZ = DOWN (Fun (comp g f) (UP (hfst XZ)))" + hsnd XZ = elem_of (Fun (comp g f) (arr_of (hfst XZ)))" using assms XZ 1 4 by (metis (no_types, lifting) arr_to_hfun_in_hexp happ_mapsto(1) seqE) qed thus "XZ \<^bold>\ \XZ \<^bold>\ ide_to_hf (dom f) * ide_to_hf (cod g). - hsnd XZ = DOWN (Fun (comp g f) (UP (hfst XZ)))\ + hsnd XZ = elem_of (Fun (comp g f) (arr_of (hfst XZ)))\ \ XZ \<^bold>\ \XZ \<^bold>\ hdomain (arr_to_hfun f) * hrange (arr_to_hfun g). hsnd XZ = happ (arr_to_hfun g) (happ (arr_to_hfun f) (hfst XZ))\" using 1 is_hpair_def by auto qed also have "... = hcomp (arr_to_hfun g) (arr_to_hfun f)" using assms arr_to_hfun_def hcomp_def by simp finally show ?thesis by simp qed lemma hfun_to_arr_hcomp: assumes "hfun A B F" and "hfun B C G" shows "hfun_to_arr A C (hcomp G F) = comp (hfun_to_arr B C G) (hfun_to_arr A B F)" proof - have 1: "arr_to_hfun (hfun_to_arr A C (hcomp G F)) = arr_to_hfun (comp (hfun_to_arr B C G) (hfun_to_arr A B F))" proof - have "arr_to_hfun (comp (hfun_to_arr B C G) (hfun_to_arr A B F)) = hcomp (arr_to_hfun (hfun_to_arr B C G)) (arr_to_hfun (hfun_to_arr A B F))" using assms arr_to_hfun_comp hfun_to_arr_in_hom by blast also have "... = hcomp G F" using assms by (simp add: arr_to_hfun_hfun_to_arr) also have "... = arr_to_hfun (hfun_to_arr A C (hcomp G F))" proof - have "hfun A C (hcomp G F)" using assms hfun_hcomp by simp thus ?thesis by (simp add: arr_to_hfun_hfun_to_arr) qed finally show ?thesis by simp qed show ?thesis proof - have "hfun_to_arr A C (hcomp G F) \ hom (hf_to_ide A) (hf_to_ide C)" using assms hfun_hcomp hf_to_ide_def hfun_to_arr_in_hom by auto moreover have "comp (hfun_to_arr B C G) (hfun_to_arr A B F) \ hom (hf_to_ide A) (hf_to_ide C)" using assms hfun_to_arr_in_hom hf_to_ide_def by (metis (no_types, lifting) comp_in_homI mem_Collect_eq) moreover have "inj_on arr_to_hfun (hom (hf_to_ide A) (hf_to_ide C))" proof - have "ide (hf_to_ide A) \ ide (hf_to_ide C)" using assms hf_to_ide_mapsto by auto thus ?thesis using bij_betw_hom_hfun [of "hf_to_ide A" "hf_to_ide C"] bij_betw_imp_inj_on by auto qed ultimately show ?thesis using 1 inj_on_def [of arr_to_hfun "hom (hf_to_ide A) (hf_to_ide C)"] by simp qed qed section "Binary Products" text\ The category of hereditarily finite sets has binary products, given by cartesian product of sets in the usual way. \ definition prod where "prod a b = hf_to_ide (ide_to_hf a * ide_to_hf b)" definition pr0 where "pr0 a b = (if ide a \ ide b then - mkArr (set (prod a b)) (set b) (\x. UP (hsnd (DOWN x))) + mkArr (set (prod a b)) (set b) (\x. arr_of (hsnd (elem_of x))) else null)" definition pr1 where "pr1 a b = (if ide a \ ide b then - mkArr (set (prod a b)) (set a) (\x. UP (hfst (DOWN x))) + mkArr (set (prod a b)) (set a) (\x. arr_of (hfst (elem_of x))) else null)" definition tuple where "tuple f g = mkArr (set (dom f)) (set (prod (cod f) (cod g))) - (\x. UP (hpair (DOWN (Fun f x)) (DOWN (Fun g x))))" + (\x. arr_of (hpair (elem_of (Fun f x)) (elem_of (Fun g x))))" lemma ide_prod: assumes "ide a" and "ide b" shows "ide (prod a b)" using assms prod_def hf_to_ide_mapsto ide_to_hf_mapsto by auto lemma pr1_in_hom [intro]: assumes "ide a" and "ide b" shows "\pr1 a b : prod a b \ a\" proof show 0: "arr (pr1 a b)" proof - - have "set (prod a b) \ Univ" - using assms ide_prod ide_char set_subset_Univ by blast - moreover have "set a \ Univ" - using assms ide_char set_subset_Univ by blast - moreover have "(\x. UP (hfst (DOWN x))) \ set (prod a b) \ set a" + have "set (prod a b) \ Univ \ finite (set (prod a b))" + using assms ide_implies_finite_set(1) set_ideD(1) ide_prod by presburger + moreover have "set a \ Univ \ finite (set a)" + using assms ide_char set_ide_char by blast + moreover have "(\x. arr_of (hfst (elem_of x))) \ set (prod a b) \ set a" proof (unfold prod_def) - show "(\x. UP (hfst (DOWN x))) \ set (hf_to_ide (ide_to_hf a * ide_to_hf b)) \ set a" + show "(\x. arr_of (hfst (elem_of x))) \ set (hf_to_ide (ide_to_hf a * ide_to_hf b)) \ set a" proof fix x assume x: "x \ set (hf_to_ide (ide_to_hf a * ide_to_hf b))" - have "DOWN x \ hfset (ide_to_hf a * ide_to_hf b)" + have "elem_of x \ hfset (ide_to_hf a * ide_to_hf b)" using assms ide_char x - by (metis (no_types, lifting) prod_def DOWN_membI HF_hfset UNIV_I hmem_HF_iff + by (metis (no_types, lifting) prod_def elem_of_membI HF_hfset UNIV_I hmem_HF_iff ide_prod ide_to_hf_hf_to_ide) - hence "hfst (DOWN x) \<^bold>\ ide_to_hf a" + hence "hfst (elem_of x) \<^bold>\ ide_to_hf a" by (metis HF_hfset finite_hfset hfst_conv hmem_HF_iff timesE) - thus "UP (hfst (DOWN x)) \ set a" - using UP_membI by simp + thus "arr_of (hfst (elem_of x)) \ set a" + using arr_of_membI by simp qed qed ultimately show ?thesis unfolding pr1_def - using assms arr_mkArr ide_prod set_card by presburger + using assms arr_mkArr finite_imp_setp by presburger qed show "dom (pr1 a b) = prod a b" using assms 0 ide_char ide_prod dom_mkArr by (metis (no_types, lifting) mkIde_set pr1_def) show "cod (pr1 a b) = a" using assms 0 ide_char ide_prod cod_mkArr by (metis (no_types, lifting) mkIde_set pr1_def) qed lemma pr1_simps [simp]: assumes "ide a" and "ide b" shows "arr (pr1 a b)" and "dom (pr1 a b) = prod a b" and "cod (pr1 a b) = a" using assms pr1_in_hom by blast+ lemma pr0_in_hom [intro]: assumes "ide a" and "ide b" shows "\pr0 a b : prod a b \ b\" proof show 0: "arr (pr0 a b)" proof - - have "set (prod a b) \ Univ" - using assms ide_prod ide_char set_subset_Univ by blast - moreover have "set b \ Univ" - using assms ide_char set_subset_Univ by blast - moreover have "(\x. UP (hsnd (DOWN x))) \ set (prod a b) \ set b" + have "set (prod a b) \ Univ \ finite (set (prod a b))" + using setp_set_ide assms ide_implies_finite_set(1) ide_prod by presburger + moreover have "set b \ Univ \ finite (set b)" + using assms ide_char set_ide_char by blast + moreover have "(\x. arr_of (hsnd (elem_of x))) \ set (prod a b) \ set b" proof (unfold prod_def) - show "(\x. UP (hsnd (DOWN x))) \ set (hf_to_ide (ide_to_hf a * ide_to_hf b)) \ set b" + show "(\x. arr_of (hsnd (elem_of x))) \ set (hf_to_ide (ide_to_hf a * ide_to_hf b)) \ set b" proof fix x assume x: "x \ set (hf_to_ide (ide_to_hf a * ide_to_hf b))" - have "DOWN x \ hfset (ide_to_hf a * ide_to_hf b)" + have "elem_of x \ hfset (ide_to_hf a * ide_to_hf b)" using assms ide_char x - by (metis (no_types, lifting) prod_def DOWN_membI HF_hfset UNIV_I hmem_HF_iff + by (metis (no_types, lifting) prod_def elem_of_membI HF_hfset UNIV_I hmem_HF_iff ide_prod ide_to_hf_hf_to_ide) - hence "hsnd (DOWN x) \<^bold>\ ide_to_hf b" + hence "hsnd (elem_of x) \<^bold>\ ide_to_hf b" by (metis HF_hfset finite_hfset hsnd_conv hmem_HF_iff timesE) - thus "UP (hsnd (DOWN x)) \ set b" - using UP_membI by simp + thus "arr_of (hsnd (elem_of x)) \ set b" + using arr_of_membI by simp qed qed ultimately show ?thesis unfolding pr0_def - using assms arr_mkArr ide_prod set_card by presburger + using assms arr_mkArr finite_imp_setp by presburger qed show "dom (pr0 a b) = prod a b" using assms 0 ide_char ide_prod dom_mkArr by (metis (no_types, lifting) mkIde_set pr0_def) show "cod (pr0 a b) = b" using assms 0 ide_char ide_prod cod_mkArr by (metis (no_types, lifting) mkIde_set pr0_def) qed lemma pr0_simps [simp]: assumes "ide a" and "ide b" shows "arr (pr0 a b)" and "dom (pr0 a b) = prod a b" and "cod (pr0 a b) = b" using assms pr0_in_hom by blast+ - lemma UP_tuple_DOWN_membI: + lemma arr_of_tuple_elem_of_membI: assumes "span f g" and "x \ Dom f" - shows "UP \DOWN (Fun f x), DOWN (Fun g x)\ \ set (prod (cod f) (cod g))" + shows "arr_of \elem_of (Fun f x), elem_of (Fun g x)\ \ set (prod (cod f) (cod g))" proof - have "Fun f x \ set (cod f)" using assms Fun_mapsto by blast moreover have "Fun g x \ set (cod g)" using assms Fun_mapsto by auto - ultimately have "\DOWN (Fun f x), DOWN (Fun g x)\ + ultimately have "\elem_of (Fun f x), elem_of (Fun g x)\ \<^bold>\ ide_to_hf (cod f) * ide_to_hf (cod g)" using assms ide_cod by auto moreover have "set (prod (cod f) (cod g)) \ Univ" - using assms ide_char ide_cod set_subset_Univ ide_prod by presburger + using setp_set_ide assms(1) ide_cod ide_prod by presburger ultimately show ?thesis - using prod_def UP_membI ide_to_hf_hf_to_ide by auto + using prod_def arr_of_membI ide_to_hf_hf_to_ide by auto qed lemma tuple_in_hom [intro]: assumes "span f g" shows "\tuple f g : dom f \ prod (cod f) (cod g)\" proof show 1: "arr (tuple f g)" proof - - have "Dom f \ Univ" - using assms set_subset_Univ ide_dom by blast - moreover have "set (prod (cod f) (cod g)) \ Univ" - using assms ide_char ide_cod set_subset_Univ ide_prod by presburger - moreover have "(\x. UP \DOWN (Fun f x), DOWN (Fun g x)\) + have "Dom f \ Univ \ finite (Dom f)" + using assms set_ideD(1) ide_dom ide_implies_finite_set(1) by presburger + moreover have "set (prod (cod f) (cod g)) \ Univ \ finite (set (prod (cod f) (cod g)))" + using assms set_ideD(1) ide_cod ide_prod ide_implies_finite_set(1) by presburger + moreover have "(\x. arr_of \elem_of (Fun f x), elem_of (Fun g x)\) \ Dom f \ set (prod (cod f) (cod g))" - using assms UP_tuple_DOWN_membI by simp + using assms arr_of_tuple_elem_of_membI by simp ultimately show ?thesis - using assms ide_prod tuple_def arr_mkArr set_card ide_dom ide_cod by simp + using assms ide_prod tuple_def arr_mkArr ide_dom ide_cod by simp qed show "dom (tuple f g) = dom f" using assms 1 dom_mkArr ide_dom mkIde_set tuple_def by auto show "cod (tuple f g) = prod (cod f) (cod g)" using assms 1 cod_mkArr ide_cod mkIde_set tuple_def ide_prod by auto qed lemma tuple_simps [simp]: assumes "span f g" shows "arr (tuple f g)" and "dom (tuple f g) = dom f" and "cod (tuple f g) = prod (cod f) (cod g)" using assms tuple_in_hom by blast+ lemma Fun_pr1: assumes "ide a" and "ide b" - shows "Fun (pr1 a b) = restrict (\x. UP (hfst (DOWN x))) (set (prod a b))" + shows "Fun (pr1 a b) = restrict (\x. arr_of (hfst (elem_of x))) (set (prod a b))" using assms pr1_def Fun_mkArr arr_char pr1_simps(1) by presburger lemma Fun_pr0: assumes "ide a" and "ide b" - shows "Fun (pr0 a b) = restrict (\x. UP (hsnd (DOWN x))) (set (prod a b))" + shows "Fun (pr0 a b) = restrict (\x. arr_of (hsnd (elem_of x))) (set (prod a b))" using assms pr0_def Fun_mkArr arr_char pr0_simps(1) by presburger lemma Fun_tuple: assumes "span f g" - shows "Fun (tuple f g) = restrict (\x. UP \DOWN (Fun f x), DOWN (Fun g x)\) (Dom f)" + shows "Fun (tuple f g) = restrict (\x. arr_of \elem_of (Fun f x), elem_of (Fun g x)\) (Dom f)" proof - have "arr (tuple f g)" using assms tuple_in_hom by blast thus ?thesis using assms tuple_def Fun_mkArr by simp qed lemma pr1_tuple: assumes "span f g" shows "comp (pr1 (cod f) (cod g)) (tuple f g) = f" proof (intro arr_eqI) have pr1: "\pr1 (cod f) (cod g) : prod (cod f) (cod g) \ cod f\" using assms ide_cod by blast have tuple: "\tuple f g : dom f \ prod (cod f) (cod g)\" using assms by blast show par: "par (comp (pr1 (cod f) (cod g)) (tuple f g)) f" using assms pr1_in_hom tuple_in_hom by (metis (no_types, lifting) comp_in_homI' ide_cod in_homE) show "Fun (comp (pr1 (cod f) (cod g)) (tuple f g)) = Fun f" proof - have seq: "seq (pr1 (cod f) (cod g)) (tuple f g)" using par by blast have "Fun (comp (pr1 (cod f) (cod g)) (tuple f g)) = restrict (Fun (pr1 (cod f) (cod g)) \ Fun (tuple f g)) (Dom (tuple f g))" using pr1 tuple seq Fun_comp by simp also have "... = restrict (Fun (mkArr (set (prod (cod f) (cod g))) (Cod f) - (\x. UP (hfst (DOWN x)))) \ + (\x. arr_of (hfst (elem_of x)))) \ Fun (mkArr (Dom f) (set (prod (cod f) (cod g))) - (\x. UP \DOWN (Fun f x), DOWN (Fun g x)\))) + (\x. arr_of \elem_of (Fun f x), elem_of (Fun g x)\))) (Dom (tuple f g))" unfolding pr1_def tuple_def using assms ide_cod by presburger also have "... = restrict - (restrict (\x. UP (hfst (DOWN x))) (set (prod (cod f) (cod g))) o - restrict (\x. UP \DOWN (Fun f x), DOWN (Fun g x)\) (Dom f)) + (restrict (\x. arr_of (hfst (elem_of x))) (set (prod (cod f) (cod g))) o + restrict (\x. arr_of \elem_of (Fun f x), elem_of (Fun g x)\) (Dom f)) (Dom f)" proof - - have "Fun (mkArr (set (prod (cod f) (cod g))) (Cod f) (\x. UP (hfst (DOWN x)))) = - restrict (\x. UP (hfst (DOWN x))) (set (prod (cod f) (cod g)))" + have "Fun (mkArr (set (prod (cod f) (cod g))) (Cod f) (\x. arr_of (hfst (elem_of x)))) = + restrict (\x. arr_of (hfst (elem_of x))) (set (prod (cod f) (cod g)))" using assms Fun_mkArr ide_prod pr1 by (metis (no_types, lifting) arrI ide_cod pr1_def) moreover have "Fun (mkArr (Dom f) (set (prod (cod f) (cod g))) - (\x. UP \DOWN (Fun f x), DOWN (Fun g x)\)) = - restrict (\x. UP \DOWN (Fun f x), DOWN (Fun g x)\) (Dom f)" + (\x. arr_of \elem_of (Fun f x), elem_of (Fun g x)\)) = + restrict (\x. arr_of \elem_of (Fun f x), elem_of (Fun g x)\) (Dom f)" using assms Fun_mkArr ide_prod ide_cod tuple_def tuple arrI by simp ultimately show ?thesis using assms tuple_simps(2) by simp qed also have "... = restrict - ((\x. UP (hfst (DOWN x))) o (\x. UP \DOWN (Fun f x), DOWN (Fun g x)\)) + ((\x. arr_of (hfst (elem_of x))) o (\x. arr_of \elem_of (Fun f x), elem_of (Fun g x)\)) (Dom f)" - using assms tuple tuple_def UP_tuple_DOWN_membI by auto + using assms tuple tuple_def arr_of_tuple_elem_of_membI by auto also have "... = restrict (Fun f) (Dom f)" proof fix x - have "restrict ((\x. UP (hfst (DOWN x))) o (\x. UP \DOWN (Fun f x), DOWN (Fun g x)\)) + have "restrict ((\x. arr_of (hfst (elem_of x))) o (\x. arr_of \elem_of (Fun f x), elem_of (Fun g x)\)) (Dom f) x = - restrict (\x. UP (DOWN (Fun f x))) (Dom f) x" + restrict (\x. arr_of (elem_of (Fun f x))) (Dom f) x" by simp also have "... = restrict (Fun f) (Dom f) x" proof (cases "x \ Dom f") show "x \ Dom f \ ?thesis" by simp assume x: "x \ Dom f" have "Fun f x \ Cod f" using assms x Fun_mapsto arr_char by blast moreover have "Cod f \ Univ" - using assms pr1 ide_cod set_subset_Univ by simp + using setp_set_ide assms ide_cod by blast ultimately show ?thesis - using assms UP_DOWN Fun_mapsto by auto + using assms arr_of_elem_of Fun_mapsto by auto qed - finally show "restrict ((\x. UP (hfst (DOWN x))) \ - (\x. UP \DOWN (Fun f x), DOWN (Fun g x)\)) + finally show "restrict ((\x. arr_of (hfst (elem_of x))) \ + (\x. arr_of \elem_of (Fun f x), elem_of (Fun g x)\)) (Dom f) x = restrict (Fun f) (Dom f) x" by blast qed also have "... = Fun f" using assms par Fun_mapsto Fun_mkArr mkArr_Fun by (metis (no_types, lifting)) finally show ?thesis by blast qed qed lemma pr0_tuple: assumes "span f g" shows "comp (pr0 (cod f) (cod g)) (tuple f g) = g" proof (intro arr_eqI) have pr0: "\pr0 (cod f) (cod g) : prod (cod f) (cod g) \ cod g\" using assms ide_cod by blast have tuple: "\tuple f g : dom f \ prod (cod f) (cod g)\" using assms by blast show par: "par (comp (pr0 (cod f) (cod g)) (tuple f g)) g" using assms pr0_in_hom tuple_in_hom by (metis (no_types, lifting) comp_in_homI' ide_cod in_homE) show "Fun (comp (pr0 (cod f) (cod g)) (tuple f g)) = Fun g" proof - have seq: "seq (pr0 (cod f) (cod g)) (tuple f g)" using par by blast have "Fun (comp (pr0 (cod f) (cod g)) (tuple f g)) = restrict (Fun (pr0 (cod f) (cod g)) \ Fun (tuple f g)) (Dom (tuple f g))" using pr0 tuple seq Fun_comp by simp also have "... = restrict (Fun (mkArr (set (prod (cod f) (cod g))) (Cod g) - (\x. UP (hsnd (DOWN x)))) \ + (\x. arr_of (hsnd (elem_of x)))) \ Fun (mkArr (Dom f) (set (prod (cod f) (cod g))) - (\x. UP \DOWN (Fun f x), DOWN (Fun g x)\))) + (\x. arr_of \elem_of (Fun f x), elem_of (Fun g x)\))) (Dom (tuple f g))" unfolding pr0_def tuple_def using assms ide_cod by presburger also have "... = restrict - (restrict (\x. UP (hsnd (DOWN x))) (set (prod (cod f) (cod g))) o - restrict (\x. UP \DOWN (Fun f x), DOWN (Fun g x)\) (Dom g)) + (restrict (\x. arr_of (hsnd (elem_of x))) (set (prod (cod f) (cod g))) o + restrict (\x. arr_of \elem_of (Fun f x), elem_of (Fun g x)\) (Dom g)) (Dom g)" proof - - have "Fun (mkArr (set (prod (cod f) (cod g))) (Cod g) (\x. UP (hsnd (DOWN x)))) = - restrict (\x. UP (hsnd (DOWN x))) (set (prod (cod f) (cod g)))" + have "Fun (mkArr (set (prod (cod f) (cod g))) (Cod g) (\x. arr_of (hsnd (elem_of x)))) = + restrict (\x. arr_of (hsnd (elem_of x))) (set (prod (cod f) (cod g)))" using assms Fun_mkArr ide_prod arrI by (metis (no_types, lifting) ide_cod pr0 pr0_def) moreover have "Fun (mkArr (Dom f) (set (prod (cod f) (cod g))) - (\x. UP \DOWN (Fun f x), DOWN (Fun g x)\)) = - restrict (\x. UP \DOWN (Fun f x), DOWN (Fun g x)\) (Dom f)" + (\x. arr_of \elem_of (Fun f x), elem_of (Fun g x)\)) = + restrict (\x. arr_of \elem_of (Fun f x), elem_of (Fun g x)\) (Dom f)" using assms Fun_mkArr ide_prod ide_cod tuple_def tuple arrI by simp ultimately show ?thesis using assms tuple_simps(2) by simp qed also have "... = restrict - ((\x. UP (hsnd (DOWN x))) o (\x. UP \DOWN (Fun f x), DOWN (Fun g x)\)) + ((\x. arr_of (hsnd (elem_of x))) o (\x. arr_of \elem_of (Fun f x), elem_of (Fun g x)\)) (Dom g)" - using assms tuple tuple_def UP_tuple_DOWN_membI by auto + using assms tuple tuple_def arr_of_tuple_elem_of_membI by auto also have "... = restrict (Fun g) (Dom g)" proof fix x - have "restrict ((\x. UP (hsnd (DOWN x))) - o (\x. UP \DOWN (Fun f x), DOWN (Fun g x)\)) + have "restrict ((\x. arr_of (hsnd (elem_of x))) + o (\x. arr_of \elem_of (Fun f x), elem_of (Fun g x)\)) (Dom g) x = - restrict (\x. UP (DOWN (Fun g x))) (Dom g) x" + restrict (\x. arr_of (elem_of (Fun g x))) (Dom g) x" by simp also have "... = restrict (Fun g) (Dom g) x" proof (cases "x \ Dom g") show "x \ Dom g \ ?thesis" by simp assume x: "x \ Dom g" have "Fun g x \ Cod g" using assms x Fun_mapsto arr_char by blast moreover have "Cod g \ Univ" - using assms pr0 ide_cod set_subset_Univ by simp + using assms set_ideD(1) ide_cod by blast ultimately show ?thesis - using assms UP_DOWN Fun_mapsto by auto + using assms arr_of_elem_of Fun_mapsto by auto qed - finally show "restrict ((\x. UP (hsnd (DOWN x))) \ - (\x. UP \DOWN (Fun f x), DOWN (Fun g x)\)) + finally show "restrict ((\x. arr_of (hsnd (elem_of x))) \ + (\x. arr_of \elem_of (Fun f x), elem_of (Fun g x)\)) (Dom g) x = restrict (Fun g) (Dom g) x" by blast qed also have "... = Fun g" using assms par Fun_mapsto Fun_mkArr mkArr_Fun by (metis (no_types, lifting)) finally show ?thesis by blast qed qed lemma tuple_pr: assumes "ide a" and "ide b" and "\h : dom h \ prod a b\" shows "tuple (comp (pr1 a b) h) (comp (pr0 a b) h) = h" proof (intro arr_eqI) have pr0: "\pr0 a b : prod a b \ b\" using assms pr0_in_hom ide_cod by blast have pr1: "\pr1 a b : prod a b \ a\" using assms pr1_in_hom ide_cod by blast have tuple: "\tuple (comp (pr1 a b) h) (comp (pr0 a b) h) : dom h \ prod a b\" using assms pr0 pr1 by (metis (no_types, lifting) cod_comp dom_comp pr0_simps(3) pr1_simps(3) seqI' tuple_in_hom) show par: "par (tuple (comp (pr1 a b) h) (comp (pr0 a b) h)) h" using assms tuple by (metis (no_types, lifting) in_homE) show "Fun (tuple (comp (pr1 a b) h) (comp (pr0 a b) h)) = Fun h" proof - have 1: "Fun (comp (pr1 a b) h) = - restrict (restrict (\x. UP (hfst (DOWN x))) (set (prod a b)) \ Fun h) (Dom h)" + restrict (restrict (\x. arr_of (hfst (elem_of x))) (set (prod a b)) \ Fun h) (Dom h)" using assms pr1 Fun_comp Fun_pr1 seqI' by auto have 2: "Fun (comp (pr0 a b) h) = - restrict (restrict (\x. UP (hsnd (DOWN x))) (set (prod a b)) \ Fun h) (Dom h)" + restrict (restrict (\x. arr_of (hsnd (elem_of x))) (set (prod a b)) \ Fun h) (Dom h)" using assms pr0 Fun_comp Fun_pr0 seqI' by auto have "Fun (tuple (comp (pr1 a b) h) (comp (pr0 a b) h)) = - restrict (\x. UP \DOWN (restrict - (restrict (\x. UP (hfst (DOWN x))) (set (prod a b)) \ Fun h) + restrict (\x. arr_of \elem_of (restrict + (restrict (\x. arr_of (hfst (elem_of x))) (set (prod a b)) \ Fun h) (Dom h) x), - DOWN (restrict - (restrict (\x. UP (hsnd (DOWN x))) (set (prod a b)) \ Fun h) + elem_of (restrict + (restrict (\x. arr_of (hsnd (elem_of x))) (set (prod a b)) \ Fun h) (Dom h) x)\) (Dom h)" proof - have "Dom (comp (pr1 a b) h) = Dom h" using assms pr1_in_hom by (metis (no_types, lifting) in_homE dom_comp seqI) moreover have "arr (mkArr (Dom (comp (pr1 a b) h)) (set (prod (cod (comp (pr1 a b) h)) (cod (comp (pr0 a b) h)))) - (\x. UP \DOWN (Fun (comp (pr1 a b) h) x), - DOWN (Fun (comp (pr0 a b) h) x)\))" + (\x. arr_of \elem_of (Fun (comp (pr1 a b) h) x), + elem_of (Fun (comp (pr0 a b) h) x)\))" using tuple unfolding tuple_def by blast ultimately show ?thesis using 1 2 tuple tuple_def Fun_mkArr [of "Dom (comp (pr1 a b) h)" "set (prod (cod (comp (pr1 a b) h)) (cod (comp (pr0 a b) h)))" - "\x. UP \DOWN (Fun (comp (pr1 a b) h) x), - DOWN (Fun (comp (pr0 a b) h) x)\"] + "\x. arr_of \elem_of (Fun (comp (pr1 a b) h) x), + elem_of (Fun (comp (pr0 a b) h) x)\"] by simp qed also have "... = Fun h" proof let ?f = "..." fix x show "?f x = Fun h x" proof - have "x \ Dom h \ ?f x = Fun h x" proof - assume x: "x \ Dom h" have "restrict ?f (Dom h) x = undefined" using assms x restrict_apply by auto also have "... = Fun h x" proof - have "arr h" using assms by blast thus ?thesis using assms x Fun_mapsto [of h] extensional_arb [of "Fun h" "Dom h" x] by simp qed finally show ?thesis by auto qed moreover have "x \ Dom h \ ?f x = Fun h x" proof - assume x: "x \ Dom h" have 1: "Fun h x \ set (prod a b)" proof - have "Fun h x \ Cod h" using assms x Fun_mapsto [of h] by blast moreover have "Cod h = set (prod a b)" using assms ide_prod by (metis (no_types, lifting) in_homE) ultimately show ?thesis by fast qed - have "?f x = UP \hfst (DOWN (Fun h x)), hsnd (DOWN (Fun h x))\" + have "?f x = arr_of \hfst (elem_of (Fun h x)), hsnd (elem_of (Fun h x))\" using x 1 by simp - also have "... = UP (DOWN (Fun h x))" + also have "... = arr_of (elem_of (Fun h x))" proof - - have "DOWN (Fun h x) \<^bold>\ ide_to_hf a * ide_to_hf b" + have "elem_of (Fun h x) \<^bold>\ ide_to_hf a * ide_to_hf b" using assms x 1 par - by (metis (no_types, lifting) prod_def DOWN_membI UNIV_I ide_prod + by (metis (no_types, lifting) prod_def elem_of_membI UNIV_I ide_prod ide_to_hf_hf_to_ide) thus ?thesis using x is_hpair_def by auto qed also have "... = Fun h x" - using assms 1 ide_prod set_subset_Univ UP_DOWN - by (meson subsetD) + using 1 arr_of_elem_of assms set_ideD(1) ide_prod by blast finally show ?thesis by blast qed ultimately show ?thesis by blast qed qed finally show ?thesis by blast qed qed interpretation HF': elementary_category_with_binary_products comp pr0 pr1 proof show "\a b. \ide a; ide b\ \ span (pr1 a b) (pr0 a b)" using pr0_simps(1) pr0_simps(2) pr1_simps(1) pr1_simps(2) by auto show "\a b. \ide a; ide b\ \ cod (pr0 a b) = b" using pr0_simps(1-3) by blast show "\a b. \ide a; ide b\ \ cod (pr1 a b) = a" using pr1_simps(1-3) by blast show "\f g. span f g \ \!l. comp (pr1 (cod f) (cod g)) l = f \ comp (pr0 (cod f) (cod g)) l = g" proof fix f g assume fg: "span f g" show "comp (pr1 (cod f) (cod g)) (tuple f g) = f \ comp (pr0 (cod f) (cod g)) (tuple f g) = g" using fg pr0_simps pr1_simps tuple_simps pr0_tuple pr1_tuple by presburger show "\l. \comp (pr1 (cod f) (cod g)) l = f \ comp (pr0 (cod f) (cod g)) l = g\ \ l = tuple f g " proof - fix l assume l: "comp (pr1 (cod f) (cod g)) l = f \ comp (pr0 (cod f) (cod g)) l = g" show "l = tuple f g" using fg l tuple_pr by (metis (no_types, lifting) arr_iff_in_hom ide_cod seqE pr0_simps(2)) qed qed show "\a b. \ (ide a \ ide b) \ pr0 a b = null" using pr0_def by auto show "\a b. \ (ide a \ ide b) \ pr1 a b = null" using pr1_def by auto qed text\ For reasons of economy of locale parameters, the notion \prod\ is a defined notion of the @{locale elementary_category_with_binary_products} locale. However, we need to be able to relate this notion to that of cartesian product of hereditarily finite sets, which we have already used to give a definition of \prod\. The locale assumptions for @{locale elementary_cartesian_closed_category} refer specifically to \HF'.prod\, even though in the end the notion itself does not depend on that choice. To be able to show that the locale assumptions of @{locale elementary_cartesian_closed_category} are satisfied, we need to use a choice of products that we can relate to the cartesian product of hereditarily finite sets. We therefore need to show that our previously defined \prod\ coincides (on objects) with the one defined in the @{locale elementary_category_with_binary_products} locale; \emph{i.e.}~\HF'.prod\. Note that the latter is defined for all arrows, not just identity arrows, so we need to use that for the subsequent definitions and proofs. \ lemma prod_ide_eq: assumes "ide a" and "ide b" shows "prod a b = HF'.prod a b" using assms prod_def HF'.pr_simps(2) HF'.prod_def pr0_simps(2) by presburger lemma tuple_span_eq: assumes "span f g" shows "tuple f g = HF'.tuple f g" using assms tuple_def HF'.tuple_def by (metis (no_types, lifting) HF'.tuple_eqI pr0_tuple pr1_tuple) section "Exponentials" text\ We now turn our attention to exponentials. \ definition exp where "exp b c = hf_to_ide (hexp (ide_to_hf b) (ide_to_hf c))" definition eval where "eval b c = mkArr (set (HF'.prod (exp b c) b)) (set c) - (\x. UP (happ (hfst (DOWN x)) (hsnd (DOWN x))))" + (\x. arr_of (happ (hfst (elem_of x)) (hsnd (elem_of x))))" definition \ where "\ a b c f = mkArr (set a) (set (exp b c)) - (\x. UP (happ (hlam (ide_to_hf a) (ide_to_hf b) (ide_to_hf c) + (\x. arr_of (happ (hlam (ide_to_hf a) (ide_to_hf b) (ide_to_hf c) (arr_to_hfun f)) - (DOWN x)))" + (elem_of x)))" lemma ide_exp: assumes "ide b" and "ide c" shows "ide (exp b c)" using assms exp_def hf_to_ide_mapsto ide_to_hf_mapsto by auto lemma hfset_ide_to_hf: assumes "ide a" - shows "hfset (ide_to_hf a) = DOWN ` set a" + shows "hfset (ide_to_hf a) = elem_of ` set a" using assms ide_to_hf_def ide_implies_finite_set(1) by auto lemma eval_in_hom [intro]: assumes "ide b" and "ide c" shows "in_hom (eval b c) (HF'.prod (exp b c) b) c" proof show 1: "arr (eval b c)" proof (unfold eval_def arr_mkArr, intro conjI) show "set (HF'.prod (exp b c) b) \ Univ" - using assms ide_char HF'.ide_prod ide_exp set_subset_Univ by simp + using HF'.ide_prod assms set_ideD(1) ide_exp by presburger show "set c \ Univ" - using assms ide_char set_subset_Univ by simp - show "(\x. UP (happ (hfst (DOWN x)) (hsnd (DOWN x)))) + using assms set_ideD(1) by blast + show "(\x. arr_of (happ (hfst (elem_of x)) (hsnd (elem_of x)))) \ set (HF'.prod (exp b c) b) \ set c" proof fix x assume "x \ set (HF'.prod (exp b c) b)" hence x: "x \ set (prod (exp b c) b)" using assms prod_ide_eq ide_exp by auto - show "UP (happ (hfst (DOWN x)) (hsnd (DOWN x))) \ set c" - proof (intro UP_membI) - show "happ (hfst (DOWN x)) (hsnd (DOWN x)) \<^bold>\ ide_to_hf c" + show "arr_of (happ (hfst (elem_of x)) (hsnd (elem_of x))) \ set c" + proof (intro arr_of_membI) + show "happ (hfst (elem_of x)) (hsnd (elem_of x)) \<^bold>\ ide_to_hf c" proof - - have 1: "DOWN x \<^bold>\ ide_to_hf (exp b c) * ide_to_hf b" + have 1: "elem_of x \<^bold>\ ide_to_hf (exp b c) * ide_to_hf b" proof - - have "DOWN x \<^bold>\ ide_to_hf (prod (exp b c) b)" - using assms x DOWN_membI ide_prod ide_exp by simp + have "elem_of x \<^bold>\ ide_to_hf (prod (exp b c) b)" + using assms x elem_of_membI ide_prod ide_exp by simp thus ?thesis using assms x prod_def ide_to_hf_hf_to_ide by auto qed - have "hfst (DOWN x) \<^bold>\ hexp (ide_to_hf b) (ide_to_hf c)" + have "hfst (elem_of x) \<^bold>\ hexp (ide_to_hf b) (ide_to_hf c)" using assms 1 x exp_def ide_to_hf_hf_to_ide by auto - moreover have "hsnd (DOWN x) \<^bold>\ ide_to_hf b" + moreover have "hsnd (elem_of x) \<^bold>\ ide_to_hf b" using assms 1 by auto ultimately show ?thesis - using happ_mapsto [of "hfst (DOWN x)" "ide_to_hf b" "ide_to_hf c" - "hsnd (DOWN x)"] + using happ_mapsto [of "hfst (elem_of x)" "ide_to_hf b" "ide_to_hf c" + "hsnd (elem_of x)"] by simp qed qed qed - show "|set (HF'.prod (exp b c) b)| \ hexp (ide_to_hf a) (ide_to_hf (exp b c))" using assms hfun_in_hexp hfun_hlam by (metis (no_types, lifting) prod_def HCollect_iff in_homE UNIV_I arr_to_hfun_in_hexp exp_def hexp_def ide_to_hf_hf_to_ide) lemma lam_in_hom [intro]: assumes "ide a" and "ide b" and "ide c" and "in_hom f (prod a b) c" shows "in_hom (\ a b c f) a (exp b c)" proof show 1: "arr (\ a b c f)" proof (unfold \_def arr_mkArr, intro conjI) show "set a \ Univ" - using assms(1) set_subset_Univ ide_char by blast + using assms(1) set_ideD(1) by blast show "set (exp b c) \ Univ" - using assms(2-3) set_subset_Univ ide_exp ide_char by simp - show "|set a| x. UP (happ (hlam (ide_to_hf a) (ide_to_hf b) (ide_to_hf c) (arr_to_hfun f)) - (DOWN x))) + using assms(2-3) set_ideD(1) ide_exp ide_char by blast + show "finite (elem_of ` set a)" + using assms(1) set_ideD(1) setp_set_ide by presburger + show "finite (elem_of ` set (exp b c))" + using assms(2-3) set_ideD(1) setp_set_ide ide_exp by presburger + show "(\x. arr_of (happ (hlam (ide_to_hf a) (ide_to_hf b) (ide_to_hf c) (arr_to_hfun f)) + (elem_of x))) \ set a \ set (exp b c)" proof fix x assume x: "x \ set a" - show "UP (happ (hlam (ide_to_hf a) (ide_to_hf b) (ide_to_hf c) (arr_to_hfun f)) - (DOWN x)) + show "arr_of (happ (hlam (ide_to_hf a) (ide_to_hf b) (ide_to_hf c) (arr_to_hfun f)) + (elem_of x)) \ set (exp b c)" - using assms x hlam_arr_to_hfun_in_hexp ide_to_hf_def DOWN_membI happ_mapsto - UP_membI + using assms x hlam_arr_to_hfun_in_hexp ide_to_hf_def elem_of_membI happ_mapsto + arr_of_membI by meson qed qed show "dom (\ a b c f) = a" using assms(1) 1 \_def ide_char dom_mkArr mkIde_set by auto show "cod (\ a b c f) = exp b c" using assms(2-3) 1 \_def cod_mkArr ide_exp mkIde_set by auto qed lemma lam_simps [simp]: assumes "ide a" and "ide b" and "ide c" and "in_hom f (prod a b) c" shows "arr (\ a b c f)" and "dom (\ a b c f) = a" and "cod (\ a b c f) = exp b c" using assms lam_in_hom by blast+ lemma Fun_lam: assumes "ide a" and "ide b" and "ide c" and "in_hom f (prod a b) c" shows "Fun (\ a b c f) = - restrict (\x. UP (happ (hlam (ide_to_hf a) (ide_to_hf b) (ide_to_hf c) (arr_to_hfun f)) - (DOWN x))) + restrict (\x. arr_of (happ (hlam (ide_to_hf a) (ide_to_hf b) (ide_to_hf c) (arr_to_hfun f)) + (elem_of x))) (set a)" using assms arr_char lam_simps(1) \_def Fun_mkArr by simp lemma Fun_eval: assumes "ide b" and "ide c" - shows "Fun (eval b c) = restrict (\x. UP (happ (hfst (DOWN x)) (hsnd (DOWN x)))) + shows "Fun (eval b c) = restrict (\x. arr_of (happ (hfst (elem_of x)) (hsnd (elem_of x)))) (set (HF'.prod (exp b c) b))" using assms arr_char eval_simps(1) eval_def Fun_mkArr by force lemma Fun_prod: assumes "arr f" and "arr g" and "x \ set (prod (dom f) (dom g))" - shows "Fun (HF'.prod f g) x = UP \DOWN (Fun f (UP (hfst (DOWN x)))), - DOWN (Fun g (UP (hsnd (DOWN x))))\" + shows "Fun (HF'.prod f g) x = arr_of \elem_of (Fun f (arr_of (hfst (elem_of x)))), + elem_of (Fun g (arr_of (hsnd (elem_of x))))\" proof - have 1: "span (comp f (pr1 (dom f) (dom g))) (comp g (pr0 (dom f) (dom g)))" using assms by (metis (no_types, lifting) HF'.prod_def HF'.prod_simps(1) HF'.tuple_ext not_arr_null) have 2: "Dom (comp f (pr1 (dom f) (dom g))) = set (prod (dom f) (dom g))" using assms by (metis (mono_tags, lifting) 1 dom_comp ide_dom pr0_simps(2)) have 3: "Dom (comp g (pr0 (dom f) (dom g))) = set (prod (dom f) (dom g))" using assms 1 2 by force have "Fun (HF'.prod f g) x = Fun (HF'.tuple (comp f (pr1 (dom f) (dom g))) (comp g (pr0 (dom f) (dom g)))) x" using assms(3) HF'.prod_def by simp - also have "... = restrict (\x. UP \DOWN (Fun (comp f (pr1 (dom f) (dom g))) x), - DOWN (Fun (comp g (pr0 (dom f) (dom g))) x)\) + also have "... = restrict (\x. arr_of \elem_of (Fun (comp f (pr1 (dom f) (dom g))) x), + elem_of (Fun (comp g (pr0 (dom f) (dom g))) x)\) (Dom (comp f (pr1 (dom f) (dom g)))) x" using assms 1 tuple_span_eq Fun_tuple by simp - also have "... = UP \DOWN (Fun (comp f (pr1 (dom f) (dom g))) x), - DOWN (Fun (comp g (pr0 (dom f) (dom g))) x)\" + also have "... = arr_of \elem_of (Fun (comp f (pr1 (dom f) (dom g))) x), + elem_of (Fun (comp g (pr0 (dom f) (dom g))) x)\" using assms(3) 2 by simp - also have "... = UP \DOWN (Fun f (UP (hfst (DOWN x)))), - DOWN (Fun g (UP (hsnd (DOWN x))))\" + also have "... = arr_of \elem_of (Fun f (arr_of (hfst (elem_of x)))), + elem_of (Fun g (arr_of (hsnd (elem_of x))))\" proof - - have "Fun (comp f (pr1 (dom f) (dom g))) x = Fun f (UP (hfst (DOWN x)))" + have "Fun (comp f (pr1 (dom f) (dom g))) x = Fun f (arr_of (hfst (elem_of x)))" proof - (* TODO: Figure out what is making this proof so "stiff". *) have 4: "seq f (pr1 (dom f) (dom g))" using assms 1 by blast have "Fun (comp f (pr1 (dom f) (dom g))) x = restrict (Fun f \ Fun (pr1 (dom f) (dom g))) (Dom (pr1 (dom f) (dom g))) x" using assms 1 Fun_comp [of f "pr1 (dom f) (dom g)"] by (metis (no_types, lifting)) also have "... = (Fun f \ Fun (pr1 (dom f) (dom g))) x" proof - have "x \ Dom (pr1 (dom f) (dom g))" using assms 1 2 4 by (metis (no_types, lifting) dom_comp) thus ?thesis by simp qed also have "... = Fun f (Fun (pr1 (dom f) (dom g)) x)" by simp - also have "... = Fun f (UP (hfst (DOWN x)))" + also have "... = Fun f (arr_of (hfst (elem_of x)))" using assms 1 Fun_pr1 [of "dom f" "dom g"] ide_dom by simp finally show ?thesis by blast qed moreover - have "Fun (comp g (pr0 (dom f) (dom g))) x = Fun g (UP (hsnd (DOWN x)))" + have "Fun (comp g (pr0 (dom f) (dom g))) x = Fun g (arr_of (hsnd (elem_of x)))" proof - have 4: "seq g (pr0 (dom f) (dom g))" using assms 1 by blast have "Fun (comp g (pr0 (dom f) (dom g))) x = restrict (Fun g \ Fun (pr0 (dom f) (dom g))) (Dom (pr0 (dom f) (dom g))) x" using assms 1 Fun_comp [of g "pr0 (dom f) (dom g)"] by (metis (no_types, lifting)) also have "... = (Fun g \ Fun (pr0 (dom f) (dom g))) x" proof - have "x \ Dom (pr0 (dom f) (dom g))" using assms 1 2 4 by (metis (no_types, lifting) dom_comp) thus ?thesis by simp qed also have "... = Fun g (Fun (pr0 (dom f) (dom g)) x)" by simp - also have "... = Fun g (UP (hsnd (DOWN x)))" + also have "... = Fun g (arr_of (hsnd (elem_of x)))" using assms 1 Fun_pr0 [of "dom f" "dom g"] ide_dom by simp finally show ?thesis by blast qed ultimately show ?thesis by simp qed finally show ?thesis by simp qed lemma prod_in_terms_of_tuple: assumes "arr f" and "arr g" shows "HF'.prod f g = tuple (comp f (pr1 (dom f) (dom g))) (comp g (pr0 (dom f) (dom g)))" using assms HF'.prod_def tuple_span_eq by (metis (no_types, lifting) HF'.prod_simps(1) HF'.tuple_ext not_arr_null) lemma eval_prod_lam: assumes "ide a" and "ide b" and "ide c" and "in_hom g (prod a b) c" shows "comp (eval b c) (HF'.prod (\ a b c g) b) = g" proof - have ide_dom_lam: "ide (dom (\ a b c g))" using assms lam_in_hom [of a b c g] ide_dom by blast have ide_dom_b: "ide (dom b)" using assms ide_dom ideD(1) by blast define \_pr1 where "\_pr1 = comp (\ a b c g) (pr1 (dom (\ a b c g)) (dom b))" define b_pr0 where "b_pr0 = comp b (pr0 (dom (\ a b c g)) (dom b))" have lam_pr1: "in_hom \_pr1 (prod a b) (exp b c)" proof (unfold \_pr1_def, intro comp_in_homI) show "in_hom (pr1 (dom (\ a b c g)) (dom b)) (prod a b) a" using assms ide_dom_lam ide_dom_b ideD(2) lam_simps(2) pr1_in_hom by auto show "in_hom (\ a b c g) a (exp b c)" using assms lam_in_hom by simp qed have b_pr0: "in_hom b_pr0 (prod a b) b" using assms b_pr0_def by (metis (no_types, lifting) HF'.arr_pr0_iff HF'.cod_pr0 comp_in_homI' ideD(1-3) lam_simps(2) pr0_simps(2)) have 1: "span \_pr1 b_pr0" using lam_pr1 b_pr0 by (metis (no_types, lifting) in_homE) have tuple: "in_hom (tuple \_pr1 b_pr0) (prod a b) (prod (exp b c) b)" using 1 lam_pr1 b_pr0 tuple_in_hom [of \_pr1 b_pr0] by (metis (mono_tags, lifting) in_homE) define \_pr1' where "\_pr1' = comp (\ a b c g) (pr1 a b)" define b_pr0' where "b_pr0' = pr0 a b" have lam_pr1_eq: "\_pr1 = \_pr1'" using assms \_pr1_def \_pr1'_def ideD(2) lam_simps(2) by auto have b_pr0_eq: "b_pr0 = b_pr0'" using assms b_pr0_def b_pr0'_def b_pr0 comp_ide_arr by (metis (no_types, lifting) ideD(2) in_homE lam_simps(2)) - have Fun_pr0: "Fun (pr0 a b) = restrict (\x. UP (hsnd (DOWN x))) (set (prod a b))" + have Fun_pr0: "Fun (pr0 a b) = restrict (\x. arr_of (hsnd (elem_of x))) (set (prod a b))" using assms Fun_pr0 by simp have Fun_lam_pr1: "Fun \_pr1 = restrict (Fun (\ a b c g) o - restrict (\x. UP (hfst (DOWN x))) (set (prod a b))) + restrict (\x. arr_of (hfst (elem_of x))) (set (prod a b))) (set (prod a b))" using assms 1 Fun_comp Fun_pr1 lam_pr1_eq \_pr1'_def by (metis (no_types, lifting) pr1_simps(2)) have "comp (eval b c) (HF'.prod (\ a b c g) b) = comp (eval b c) (tuple \_pr1 b_pr0)" using assms \_pr1_def b_pr0_def 1 prod_in_terms_of_tuple ideD(1) lam_simps(1) by presburger also have 5: "... = comp (eval b c) (tuple \_pr1' b_pr0')" using lam_pr1_eq b_pr0_eq by simp also have "... = g" proof (intro arr_eqI) have 2: "arr (comp (eval b c) (tuple \_pr1 b_pr0))" using assms tuple arr_char by (metis (no_types, lifting) in_homE seqI eval_simps(1-2) ide_exp prod_ide_eq) have 3: "arr g" using assms by blast have tuple': "in_hom (tuple \_pr1' b_pr0') (prod a b) (prod (exp b c) b)" using tuple lam_pr1_eq b_pr0_eq by blast have 4: "Dom g = set (prod a b)" using assms by (metis (no_types, lifting) in_homE) show par: "par (comp (eval b c) (tuple \_pr1' b_pr0')) g" using assms tuple' 2 3 5 by (metis (no_types, lifting) cod_comp dom_comp in_homE eval_simps(3)) show "Fun (comp (eval b c) (tuple \_pr1' b_pr0')) = Fun g" proof fix x have "x \ set (prod a b) \ Fun (comp (eval b c) (tuple \_pr1' b_pr0')) x = Fun g x" proof - have 5: "Fun g \ extensional (Dom g)" using assms 3 Fun_mapsto by simp moreover have "Fun (comp (eval b c) (tuple \_pr1' b_pr0')) \ extensional (Dom g)" using 5 par Fun_mapsto by (metis (no_types, lifting) Int_iff) ultimately show "x \ set (prod a b) \ Fun (comp (eval b c) (tuple \_pr1' b_pr0')) x = Fun g x" using 4 extensional_arb [of "Fun g" "Dom g" x] extensional_arb [of "Fun (comp (eval b c) (tuple \_pr1' b_pr0'))" "Dom g" x] by force qed moreover have "x \ set (prod a b) \ Fun (comp (eval b c) (tuple \_pr1' b_pr0')) x = Fun g x" proof - assume x: "x \ set (prod a b)" have 6: "Dom (tuple \_pr1' b_pr0') = set (prod a b)" using assms 4 tuple' par by (metis (no_types, lifting) in_homE) have "Fun (comp (eval b c) (tuple \_pr1' b_pr0')) x = Fun (eval b c) (Fun (tuple \_pr1' b_pr0') x)" proof - have "Fun (comp (eval b c) (tuple \_pr1' b_pr0')) x = (Fun (eval b c) \ Fun (tuple \_pr1' b_pr0')) x" using assms par x 6 Fun_comp [of "eval b c" "tuple \_pr1' b_pr0'"] by auto also have "... = Fun (eval b c) (Fun (tuple \_pr1' b_pr0') x)" by simp finally show ?thesis by blast qed - also have "... = restrict (\x. UP (happ (hfst (DOWN x)) (hsnd (DOWN x)))) + also have "... = restrict (\x. arr_of (happ (hfst (elem_of x)) (hsnd (elem_of x)))) (set (HF'.prod (exp b c) b)) (Fun (tuple \_pr1' b_pr0') x)" using assms Fun_eval by simp - also have "... = (\x. UP (happ (hfst (DOWN x)) (hsnd (DOWN x)))) + also have "... = (\x. arr_of (happ (hfst (elem_of x)) (hsnd (elem_of x)))) (Fun (tuple \_pr1' b_pr0') x)" proof - have "Fun (tuple \_pr1' b_pr0') x \ set (HF'.prod (exp b c) b)" proof - have "x \ Dom (tuple \_pr1' b_pr0')" using x 6 by blast moreover have "Cod (tuple \_pr1' b_pr0') = set (HF'.prod (exp b c) b)" by (metis (no_types, lifting) in_homE assms(2-3) ide_exp prod_ide_eq tuple') moreover have "arr (tuple \_pr1' b_pr0')" using tuple' by blast ultimately show ?thesis using tuple' Fun_mapsto [of "tuple \_pr1' b_pr0'"] by auto qed thus ?thesis using restrict_apply by simp qed - also have "... = (\x. UP (happ (hfst (DOWN x)) (hsnd (DOWN x)))) - (UP \DOWN (Fun \_pr1' x), DOWN (Fun b_pr0' x)\)" + also have "... = (\x. arr_of (happ (hfst (elem_of x)) (hsnd (elem_of x)))) + (arr_of \elem_of (Fun \_pr1' x), elem_of (Fun b_pr0' x)\)" proof - have 7: "Dom \_pr1' = set (prod a b)" using assms by (metis (no_types, lifting) 1 comp_ide_arr ideD(2) b_pr0_def lam_pr1_eq lam_simps(2) pr0_simps(2)) moreover have "span \_pr1' b_pr0'" using assms 1 b_pr0_eq lam_pr1_eq by auto moreover have "x \ Dom \_pr1'" using x 7 by simp ultimately have "Fun (tuple \_pr1' b_pr0') x = - UP \DOWN (Fun \_pr1' x), DOWN (Fun b_pr0' x)\" + arr_of \elem_of (Fun \_pr1' x), elem_of (Fun b_pr0' x)\" using assms x restrict_apply Fun_tuple by simp thus ?thesis by simp qed - also have "... = UP (happ (DOWN (Fun \_pr1' x)) (DOWN (Fun b_pr0' x)))" + also have "... = arr_of (happ (elem_of (Fun \_pr1' x)) (elem_of (Fun b_pr0' x)))" using assms by simp - also have "... = UP (happ (DOWN (UP (happ (hlam (ide_to_hf a) (ide_to_hf b) + also have "... = arr_of (happ (elem_of (arr_of (happ (hlam (ide_to_hf a) (ide_to_hf b) (ide_to_hf c) (arr_to_hfun g)) - (hfst (DOWN x))))) - (DOWN (UP (hsnd (DOWN x)))))" + (hfst (elem_of x))))) + (elem_of (arr_of (hsnd (elem_of x)))))" proof - - have "Fun b_pr0' x = UP (hsnd (DOWN x))" + have "Fun b_pr0' x = arr_of (hsnd (elem_of x))" using assms x Fun_pr0 b_pr0'_def by simp moreover have "Fun \_pr1' x = - UP (happ (hlam (ide_to_hf a) (ide_to_hf b) (ide_to_hf c) + arr_of (happ (hlam (ide_to_hf a) (ide_to_hf b) (ide_to_hf c) (arr_to_hfun g)) - (hfst (DOWN x)))" + (hfst (elem_of x)))" proof - have "Fun \_pr1' x = restrict (Fun (\ a b c g) o Fun (pr1 a b)) (Dom (pr1 a b)) x" using assms x Fun_pr1 Fun_comp lam_pr1_eq Fun_lam_pr1 pr1_simps(1-2) by presburger also have "... = Fun (\ a b c g) (Fun (pr1 a b) x)" using assms x restrict_apply Fun_lam_pr1 Fun_pr1 calculation lam_pr1_eq by auto - also have "... = restrict (\x. UP (happ (hlam (ide_to_hf a) (ide_to_hf b) + also have "... = restrict (\x. arr_of (happ (hlam (ide_to_hf a) (ide_to_hf b) (ide_to_hf c) (arr_to_hfun g)) - (DOWN x))) + (elem_of x))) (set a) (Fun (pr1 a b) x)" using assms x Fun_lam by simp - also have "... = UP (happ (hlam (ide_to_hf a) (ide_to_hf b) (ide_to_hf c) + also have "... = arr_of (happ (hlam (ide_to_hf a) (ide_to_hf b) (ide_to_hf c) (arr_to_hfun g)) - (DOWN (Fun (pr1 a b) x)))" + (elem_of (Fun (pr1 a b) x)))" proof - have "Fun (pr1 a b) x \ set a" proof - have "x \ Dom (pr1 a b)" using assms x pr1_simps(1-2) by auto moreover have "Cod (pr1 a b) = set a" using assms HF'.cod_pr1 pr1_simps(1) by auto moreover have "arr (pr1 a b)" using assms arr_char by blast ultimately show ?thesis using Fun_mapsto [of "pr1 a b"] by auto qed thus ?thesis using restrict_apply by simp qed - also have "... = UP (happ (hlam (ide_to_hf a) (ide_to_hf b) (ide_to_hf c) + also have "... = arr_of (happ (hlam (ide_to_hf a) (ide_to_hf b) (ide_to_hf c) (arr_to_hfun g)) - (hfst (DOWN x)))" + (hfst (elem_of x)))" using assms x Fun_pr1 Fun_lam [of a b c g] by simp finally show ?thesis by simp qed ultimately show ?thesis by simp qed - also have "... = UP (happ (happ (hlam (ide_to_hf a) (ide_to_hf b) (ide_to_hf c) + also have "... = arr_of (happ (happ (hlam (ide_to_hf a) (ide_to_hf b) (ide_to_hf c) (arr_to_hfun g)) - (hfst (DOWN x))) - (hsnd (DOWN x)))" + (hfst (elem_of x))) + (hsnd (elem_of x)))" by simp - also have "... = UP (happ (arr_to_hfun g) (DOWN x))" + also have "... = arr_of (happ (arr_to_hfun g) (elem_of x))" using assms x happ_hlam - by (metis (no_types, lifting) prod_def DOWN_membI HCollect_iff ide_dom + by (metis (no_types, lifting) prod_def elem_of_membI HCollect_iff ide_dom in_homE UNIV_I arr_to_hfun_in_hexp hexp_def hfst_conv hsnd_conv ide_to_hf_hf_to_ide timesE) also have "... = Fun g x" using assms x 3 4 Fun_char [of g] restrict_apply [of "Fun g" "Dom g" x] by simp finally show ?thesis by simp qed ultimately show "Fun (comp (eval b c) (tuple \_pr1' b_pr0')) x = Fun g x" by auto qed qed finally show ?thesis by simp qed lemma lam_eval_prod: assumes "ide a" and "ide b" and "ide c" and "in_hom h a (exp b c)" shows "\ a b c (comp (eval b c) (HF'.prod h b)) = h" proof (intro arr_eqI) have 0: "in_hom (comp (eval b c) (HF'.prod h b)) (prod a b) c" proof show "in_hom (HF'.prod h b) (prod a b) (HF'.prod (exp b c) b)" proof show 1: "arr (HF'.prod h b)" using assms HF'.prod_in_hom' by (metis (no_types, lifting) ideD(1) in_homE) show "dom (HF'.prod h b) = prod a b" using assms 1 by (metis (no_types, lifting) HF'.prod_simps(2) ideD(1-2) in_homE prod_ide_eq) show "cod (HF'.prod h b) = HF'.prod (exp b c) b" using assms 1 by (metis (no_types, lifting) HF'.prod_simps(3) ideD(1,3) in_homE) qed show "in_hom (eval b c) (HF'.prod (exp b c) b) c" using assms by blast qed have 1: "in_hom (\ a b c (comp (eval b c) (HF'.prod h b))) a (exp b c)" using assms 0 by blast have 2: "Fun (comp (eval b c) (HF'.prod h b)) = restrict (Fun (eval b c) \ Fun (HF'.prod h b)) (set (HF'.prod a b))" proof - have "seq (eval b c) (HF'.prod h b)" using assms 1 by (metis (no_types, lifting) 0 in_homE) moreover have "Dom (HF'.prod h b) = set (HF'.prod a b)" using assms by (metis (no_types, lifting) HF'.prod_simps(2) ideD(1-2) in_homE) ultimately show ?thesis using assms Fun_comp [of "eval b c" "HF'.prod h b"] by simp qed show par: "par (\ a b c (comp (eval b c) (HF'.prod h b))) h" using assms 1 by (metis (no_types, lifting) in_homE) show "Fun (\ a b c (comp (eval b c) (HF'.prod h b))) = Fun h" proof fix x show "Fun (\ a b c (comp (eval b c) (HF'.prod h b))) x = Fun h x" proof - have "x \ set a \ ?thesis" using assms 1 Fun_mapsto extensional_arb [of "Fun h" "set a" x] extensional_arb [of "Fun (\ a b c (comp (eval b c) (HF'.prod h b)))" "set a" x] by (metis (no_types, lifting) 0 Int_iff lam_simps(2) par) moreover have "x \ set a \ ?thesis" proof - assume x: "x \ set a" have 3: "dom (comp (eval b c) (HF'.prod h b)) = HF'.prod a b" using assms 0 in_homE prod_ide_eq by auto have 4: "cod (comp (eval b c) (HF'.prod h b)) = c" using assms 0 by blast have 5: "dom (comp (eval b c) (HF'.prod h b)) = HF'.prod a b" using assms 3 by (metis (mono_tags, lifting)) have 6: "cod (comp (eval b c) (HF'.prod h b)) = c" using assms 4 by (metis (no_types, lifting)) have 7: "arr_to_hfun (comp (eval b c) (HF'.prod h b)) = \xy \<^bold>\ ide_to_hf (HF'.prod a b) * ide_to_hf c. - hsnd xy = DOWN (Fun (comp (eval b c) (HF'.prod h b)) (UP (hfst xy)))\" + hsnd xy = elem_of (Fun (comp (eval b c) (HF'.prod h b)) (arr_of (hfst xy)))\" unfolding arr_to_hfun_def using 2 5 6 by metis have "Fun (\ a b c (comp (eval b c) (HF'.prod h b))) x = - UP (happ (hlam (ide_to_hf a) (ide_to_hf b) (ide_to_hf c) + arr_of (happ (hlam (ide_to_hf a) (ide_to_hf b) (ide_to_hf c) (arr_to_hfun (comp (eval b c) (HF'.prod h b)))) - (DOWN x))" + (elem_of x))" using assms 0 x Fun_lam by auto - also have "... = UP \yz \<^bold>\ ide_to_hf b * ide_to_hf c. - \\DOWN x, hfst yz\, hsnd yz\ + also have "... = arr_of \yz \<^bold>\ ide_to_hf b * ide_to_hf c. + \\elem_of x, hfst yz\, hsnd yz\ \<^bold>\ arr_to_hfun (comp (eval b c) (HF'.prod h b))\" proof - have "seq (eval b c) (HF'.prod h b)" using assms 0 by blast moreover have "ide_to_hf (dom (comp (eval b c) (HF'.prod h b))) = ide_to_hf a * ide_to_hf b" using assms 1 3 by (metis (no_types, lifting) prod_def UNIV_I ide_to_hf_hf_to_ide prod_ide_eq) moreover have "ide_to_hf (cod (comp (eval b c) (HF'.prod h b))) = ide_to_hf c" using assms 2 4 by auto ultimately show ?thesis - using assms 0 x happ_hlam(3) DOWN_membI + using assms 0 x happ_hlam(3) elem_of_membI hfun_arr_to_hfun [of "comp (eval b c) (HF'.prod h b)"] by simp qed - also have "... = UP \yz \<^bold>\ ide_to_hf b * ide_to_hf c. - hsnd yz = DOWN (Fun (comp (eval b c) (HF'.prod h b)) - (UP \DOWN x, hfst yz\))\" + also have "... = arr_of \yz \<^bold>\ ide_to_hf b * ide_to_hf c. + hsnd yz = elem_of (Fun (comp (eval b c) (HF'.prod h b)) + (arr_of \elem_of x, hfst yz\))\" proof - have "\yz \<^bold>\ ide_to_hf b * ide_to_hf c. - \\DOWN x, hfst yz\, hsnd yz\ + \\elem_of x, hfst yz\, hsnd yz\ \<^bold>\ arr_to_hfun (comp (eval b c) (HF'.prod h b))\ = \yz \<^bold>\ ide_to_hf b * ide_to_hf c. - hsnd yz = DOWN (Fun (comp (eval b c) (HF'.prod h b)) - (UP \DOWN x, hfst yz\))\" + hsnd yz = elem_of (Fun (comp (eval b c) (HF'.prod h b)) + (arr_of \elem_of x, hfst yz\))\" proof fix yz show "yz \<^bold>\ \yz \<^bold>\ ide_to_hf b * ide_to_hf c. - \\DOWN x, hfst yz\, hsnd yz\ + \\elem_of x, hfst yz\, hsnd yz\ \<^bold>\ arr_to_hfun (comp (eval b c) (HF'.prod h b))\ \ yz \<^bold>\ \yz \<^bold>\ ide_to_hf b * ide_to_hf c. - hsnd yz = DOWN (Fun (comp (eval b c) (HF'.prod h b)) - (UP \DOWN x, hfst yz\))\" + hsnd yz = elem_of (Fun (comp (eval b c) (HF'.prod h b)) + (arr_of \elem_of x, hfst yz\))\" proof - have "yz \<^bold>\ ide_to_hf b * ide_to_hf c \ - \\DOWN x, hfst yz\, hsnd yz\ \<^bold>\ arr_to_hfun (comp (eval b c) (HF'.prod h b)) - \ hsnd yz = DOWN (Fun (comp (eval b c) (HF'.prod h b)) - (UP \DOWN x, hfst yz\))" + \\elem_of x, hfst yz\, hsnd yz\ \<^bold>\ arr_to_hfun (comp (eval b c) (HF'.prod h b)) + \ hsnd yz = elem_of (Fun (comp (eval b c) (HF'.prod h b)) + (arr_of \elem_of x, hfst yz\))" proof - assume yz: "yz \<^bold>\ ide_to_hf b * ide_to_hf c" - have "\\DOWN x, hfst yz\, hsnd yz\ + have "\\elem_of x, hfst yz\, hsnd yz\ \<^bold>\ arr_to_hfun (comp (eval b c) (HF'.prod h b)) \ - \\DOWN x, hfst yz\, hsnd yz\ \<^bold>\ ide_to_hf (HF'.prod a b) * ide_to_hf c \ - hsnd yz = DOWN (Fun (comp (eval b c) (HF'.prod h b)) - (UP \DOWN x, hfst yz\))" + \\elem_of x, hfst yz\, hsnd yz\ \<^bold>\ ide_to_hf (HF'.prod a b) * ide_to_hf c \ + hsnd yz = elem_of (Fun (comp (eval b c) (HF'.prod h b)) + (arr_of \elem_of x, hfst yz\))" using 7 by auto - moreover have "\\DOWN x, hfst yz\, hsnd yz\ + moreover have "\\elem_of x, hfst yz\, hsnd yz\ \<^bold>\ ide_to_hf (prod a b) * ide_to_hf c" proof - - have "\DOWN x, hfst yz\ \<^bold>\ ide_to_hf (HF'.prod a b)" + have "\elem_of x, hfst yz\ \<^bold>\ ide_to_hf (HF'.prod a b)" using assms x yz - by (metis (no_types, lifting) prod_def DOWN_membI UNIV_I hfst_conv + by (metis (no_types, lifting) prod_def elem_of_membI UNIV_I hfst_conv ide_to_hf_hf_to_ide prod_ide_eq timesE times_iff) thus ?thesis using yz assms(1-2) prod_ide_eq by auto qed ultimately show ?thesis using assms(1-2) prod_ide_eq by auto qed thus ?thesis by auto qed qed thus ?thesis by simp qed - also have "... = UP \yz \<^bold>\ ide_to_hf b * ide_to_hf c. yz \<^bold>\ DOWN (Fun h x)\" + also have "... = arr_of \yz \<^bold>\ ide_to_hf b * ide_to_hf c. yz \<^bold>\ elem_of (Fun h x)\" proof - have "\yz \<^bold>\ ide_to_hf b * ide_to_hf c. - hsnd yz = DOWN (Fun (comp (eval b c) (HF'.prod h b)) - (UP \DOWN x, hfst yz\))\ = - \yz \<^bold>\ ide_to_hf b * ide_to_hf c. yz \<^bold>\ DOWN (Fun h x)\" + hsnd yz = elem_of (Fun (comp (eval b c) (HF'.prod h b)) + (arr_of \elem_of x, hfst yz\))\ = + \yz \<^bold>\ ide_to_hf b * ide_to_hf c. yz \<^bold>\ elem_of (Fun h x)\" proof - have "\yz. yz \<^bold>\ ide_to_hf b * ide_to_hf c \ - hsnd yz = DOWN (Fun (comp (eval b c) (HF'.prod h b)) - (UP \DOWN x, hfst yz\)) + hsnd yz = elem_of (Fun (comp (eval b c) (HF'.prod h b)) + (arr_of \elem_of x, hfst yz\)) \ - yz \<^bold>\ DOWN (Fun h x)" + yz \<^bold>\ elem_of (Fun h x)" proof - fix yz assume yz: "yz \<^bold>\ ide_to_hf b * ide_to_hf c" - have 7: "UP \DOWN x, hfst yz\ \ set (HF'.prod a b)" - using assms x yz UP_membI - by (metis (no_types, lifting) prod_def DOWN_membI UNIV_I hfst_conv + have 7: "arr_of \elem_of x, hfst yz\ \ set (HF'.prod a b)" + using assms x yz arr_of_membI + by (metis (no_types, lifting) prod_def elem_of_membI UNIV_I hfst_conv ide_to_hf_hf_to_ide prod_ide_eq timesE times_iff) have 8: "Fun h x \ set (exp b c)" proof - have "Fun h x \ Cod h" using assms x Fun_mapsto by blast moreover have "Cod h = set (exp b c)" using assms 0 lam_simps(3) par by auto ultimately show ?thesis by blast qed - show "hsnd yz = DOWN (Fun (comp (eval b c) (HF'.prod h b)) - (UP \DOWN x, hfst yz\)) + show "hsnd yz = elem_of (Fun (comp (eval b c) (HF'.prod h b)) + (arr_of \elem_of x, hfst yz\)) \ - yz \<^bold>\ DOWN (Fun h x)" + yz \<^bold>\ elem_of (Fun h x)" proof - - have "Fun (comp (eval b c) (HF'.prod h b)) (UP \DOWN x, hfst yz\) = - UP (happ (DOWN (Fun h x)) (hfst yz))" + have "Fun (comp (eval b c) (HF'.prod h b)) (arr_of \elem_of x, hfst yz\) = + arr_of (happ (elem_of (Fun h x)) (hfst yz))" proof - - have "Fun (comp (eval b c) (HF'.prod h b)) (UP \DOWN x, hfst yz\) = + have "Fun (comp (eval b c) (HF'.prod h b)) (arr_of \elem_of x, hfst yz\) = restrict (Fun (eval b c) \ Fun (HF'.prod h b)) (set (HF'.prod a b)) - (UP \DOWN x, hfst yz\)" + (arr_of \elem_of x, hfst yz\)" using assms x yz 2 by simp also have "... = Fun (eval b c) - (Fun (HF'.prod h b) (UP \DOWN x, hfst yz\))" + (Fun (HF'.prod h b) (arr_of \elem_of x, hfst yz\))" using 7 by simp also have "... = Fun (eval b c) - (UP \DOWN (Fun h x), - DOWN (Fun b (UP (hfst yz)))\)" + (arr_of \elem_of (Fun h x), + elem_of (Fun b (arr_of (hfst yz)))\)" proof - - have "Fun (HF'.prod h b) (UP \DOWN x, hfst yz\) = - UP \DOWN (Fun h x), DOWN (Fun b (UP (hfst yz)))\" + have "Fun (HF'.prod h b) (arr_of \elem_of x, hfst yz\) = + arr_of \elem_of (Fun h x), elem_of (Fun b (arr_of (hfst yz)))\" proof - - have "Fun (HF'.prod h b) (UP \DOWN x, hfst yz\) = - UP \DOWN (Fun h (UP (hfst (DOWN (UP \DOWN x, hfst yz\))))), - DOWN (Fun b (UP (hsnd (DOWN (UP \DOWN x, hfst yz\)))))\" + have "Fun (HF'.prod h b) (arr_of \elem_of x, hfst yz\) = + arr_of \elem_of (Fun h (arr_of (hfst (elem_of (arr_of \elem_of x, hfst yz\))))), + elem_of (Fun b (arr_of (hsnd (elem_of (arr_of \elem_of x, hfst yz\)))))\" proof - - have "UP \DOWN x, hfst yz\ \ set (prod (dom h) (dom b))" + have "arr_of \elem_of x, hfst yz\ \ set (prod (dom h) (dom b))" using assms x yz 7 by (metis (no_types, lifting) ideD(2) in_homE prod_ide_eq) thus ?thesis using assms x yz Fun_prod ideD(1) by blast qed - also have "... = UP \DOWN (Fun h (UP (DOWN x))), - DOWN (Fun b (UP (hfst yz)))\" + also have "... = arr_of \elem_of (Fun h (arr_of (elem_of x))), + elem_of (Fun b (arr_of (hfst yz)))\" using assms x yz by simp - also have "... = UP \DOWN (Fun h x), DOWN (Fun b (UP (hfst yz)))\" - using assms(1) set_subset_Univ x by force + also have "... = arr_of \elem_of (Fun h x), elem_of (Fun b (arr_of (hfst yz)))\" + using assms(1) set_ideD(1) x by force finally show ?thesis by simp qed thus ?thesis by simp qed - also have "... = Fun (eval b c) (UP \DOWN (Fun h x), hfst yz\)" - using assms x yz Fun_ide ide_char UP_membI by auto - also have "... = restrict (\x. UP (happ (hfst (DOWN x)) (hsnd (DOWN x)))) + also have "... = Fun (eval b c) (arr_of \elem_of (Fun h x), hfst yz\)" + using assms x yz Fun_ide ide_char arr_of_membI by auto + also have "... = restrict (\x. arr_of (happ (hfst (elem_of x)) (hsnd (elem_of x)))) (set (HF'.prod (exp b c) b)) - (UP \DOWN (Fun h x), hfst yz\)" + (arr_of \elem_of (Fun h x), hfst yz\)" using assms Fun_eval [of b c] by simp - also have "... = (\x. UP (happ (hfst (DOWN x)) (hsnd (DOWN x)))) - (UP \DOWN (Fun h x), hfst yz\)" + also have "... = (\x. arr_of (happ (hfst (elem_of x)) (hsnd (elem_of x)))) + (arr_of \elem_of (Fun h x), hfst yz\)" proof - - have "UP \DOWN (Fun h x), hfst yz\ + have "arr_of \elem_of (Fun h x), hfst yz\ \ set (HF'.prod (exp b c) b)" proof - have 1: "ide_to_hf (HF'.prod (exp b c) b) = - HF (DOWN ` set (HF'.prod (exp b c) b))" + HF (elem_of ` set (HF'.prod (exp b c) b))" unfolding ide_to_hf_def by blast - have "\DOWN (Fun h x), hfst yz\ - \<^bold>\ HF (DOWN ` set (HF'.prod (exp b c) b))" + have "\elem_of (Fun h x), hfst yz\ + \<^bold>\ HF (elem_of ` set (HF'.prod (exp b c) b))" using assms x yz 1 8 Fun_mapsto [of h] - by (metis (no_types, lifting) prod_def DOWN_membI UNIV_I + by (metis (no_types, lifting) prod_def elem_of_membI UNIV_I hfst_conv ide_exp ide_to_hf_hf_to_ide prod_ide_eq timesE times_iff) thus ?thesis - using assms x yz 1 UP_membI [of "\DOWN (Fun h x), hfst yz\"] + using assms x yz 1 arr_of_membI [of "\elem_of (Fun h x), hfst yz\"] by auto qed thus ?thesis by simp qed - also have "... = UP (happ (DOWN (Fun h x)) (hfst yz))" + also have "... = arr_of (happ (elem_of (Fun h x)) (hfst yz))" by simp finally show ?thesis by simp qed - hence 9: "DOWN (Fun (comp (eval b c) (HF'.prod h b)) - (UP \DOWN x, hfst yz\)) = - happ (DOWN (Fun h x)) (hfst yz)" + hence 9: "elem_of (Fun (comp (eval b c) (HF'.prod h b)) + (arr_of \elem_of x, hfst yz\)) = + happ (elem_of (Fun h x)) (hfst yz)" by simp show ?thesis proof - - have "hsnd yz = happ (DOWN (Fun h x)) (hfst yz) - \ yz \<^bold>\ DOWN (Fun h x)" + have "hsnd yz = happ (elem_of (Fun h x)) (hfst yz) + \ yz \<^bold>\ elem_of (Fun h x)" proof - have 10: "\!z. \hfst yz, z\ \<^bold>\ DOWN (Fun h x)" + have 10: "\!z. \hfst yz, z\ \<^bold>\ elem_of (Fun h x)" proof - - have "hfun (ide_to_hf b) (ide_to_hf c) (DOWN (Fun h x))" + have "hfun (ide_to_hf b) (ide_to_hf c) (elem_of (Fun h x))" using assms x 8 - by (metis (no_types, lifting) DOWN_membI HCollect_iff UNIV_I + by (metis (no_types, lifting) elem_of_membI HCollect_iff UNIV_I exp_def hexp_def ide_exp ide_to_hf_hf_to_ide) thus ?thesis using assms yz - hfunE [of "ide_to_hf b" "ide_to_hf c" "DOWN (Fun h x)"] + hfunE [of "ide_to_hf b" "ide_to_hf c" "elem_of (Fun h x)"] by (metis (no_types, lifting) hfst_conv timesE) qed - show "yz \<^bold>\ DOWN (Fun h x) - \ hsnd yz = happ (DOWN (Fun h x)) (hfst yz)" + show "yz \<^bold>\ elem_of (Fun h x) + \ hsnd yz = happ (elem_of (Fun h x)) (hfst yz)" proof - - assume yz1: "yz \<^bold>\ DOWN (Fun h x)" - show "hsnd yz = happ (DOWN (Fun h x)) (hfst yz)" + assume yz1: "yz \<^bold>\ elem_of (Fun h x)" + show "hsnd yz = happ (elem_of (Fun h x)) (hfst yz)" unfolding app_def using assms x yz yz1 10 hfun_arr_to_hfun arr_to_hfun_def the1_equality - [of "\y. \hfst yz, y\ \<^bold>\ DOWN (Fun h x)" "hsnd yz"] + [of "\y. \hfst yz, y\ \<^bold>\ elem_of (Fun h x)" "hsnd yz"] by (metis (no_types, lifting) hfst_conv hsnd_conv timesE) qed - show "hsnd yz = happ (DOWN (Fun h x)) (hfst yz) - \ yz \<^bold>\ DOWN (Fun h x)" + show "hsnd yz = happ (elem_of (Fun h x)) (hfst yz) + \ yz \<^bold>\ elem_of (Fun h x)" unfolding app_def using assms x yz 10 - theI' [of "\y. \hfst yz, y\ \<^bold>\ DOWN (Fun h x)"] + theI' [of "\y. \hfst yz, y\ \<^bold>\ elem_of (Fun h x)"] by (metis (no_types, lifting) hfst_conv hsnd_conv timesE) qed thus ?thesis using 9 by simp qed qed qed thus ?thesis by blast qed thus ?thesis by simp qed also have "... = Fun h x" proof - - have H: "Fun h x = restrict (\x. UP (happ (arr_to_hfun h) (DOWN x))) (Dom h) x" + have H: "Fun h x = restrict (\x. arr_of (happ (arr_to_hfun h) (elem_of x))) (Dom h) x" proof - have "arr h" using assms by blast thus ?thesis using assms x Fun_char by simp qed - also have "... = UP (happ (arr_to_hfun h) (DOWN x))" + also have "... = arr_of (happ (arr_to_hfun h) (elem_of x))" using assms x par by (metis (no_types, lifting) 0 lam_simps(2) restrict_apply) - also have "... = UP (THE g. \DOWN x, g\ \<^bold>\ arr_to_hfun h)" + also have "... = arr_of (THE g. \elem_of x, g\ \<^bold>\ arr_to_hfun h)" using app_def by simp - also have "... = UP \yz \<^bold>\ ide_to_hf b * ide_to_hf c. yz \<^bold>\ DOWN (Fun h x)\" + also have "... = arr_of \yz \<^bold>\ ide_to_hf b * ide_to_hf c. yz \<^bold>\ elem_of (Fun h x)\" proof - - have ex_un_g: "\!g. \DOWN x, g\ \<^bold>\ arr_to_hfun h" + have ex_un_g: "\!g. \elem_of x, g\ \<^bold>\ arr_to_hfun h" using assms x arr_to_hfun_def hfun_arr_to_hfun hfunE [of "ide_to_hf a" "ide_to_hf (exp b c)" "arr_to_hfun h"] - by (metis (no_types, lifting) DOWN_membI in_homE) + by (metis (no_types, lifting) elem_of_membI in_homE) moreover have - "\DOWN x, \yz \<^bold>\ ide_to_hf b * ide_to_hf c. yz \<^bold>\ DOWN (Fun h x)\\ + "\elem_of x, \yz \<^bold>\ ide_to_hf b * ide_to_hf c. yz \<^bold>\ elem_of (Fun h x)\\ \<^bold>\ arr_to_hfun h" proof - - have "DOWN (Fun h x) = - \yz \<^bold>\ ide_to_hf b * ide_to_hf c. yz \<^bold>\ DOWN (Fun h x)\" + have "elem_of (Fun h x) = + \yz \<^bold>\ ide_to_hf b * ide_to_hf c. yz \<^bold>\ elem_of (Fun h x)\" proof fix yz - show "yz \<^bold>\ DOWN (Fun h x) \ - yz \<^bold>\ \yz \<^bold>\ ide_to_hf b * ide_to_hf c. yz \<^bold>\ DOWN (Fun h x)\" + show "yz \<^bold>\ elem_of (Fun h x) \ + yz \<^bold>\ \yz \<^bold>\ ide_to_hf b * ide_to_hf c. yz \<^bold>\ elem_of (Fun h x)\" proof - show "yz \<^bold>\ DOWN (Fun h x) - \ yz \<^bold>\ \yz \<^bold>\ ide_to_hf b * ide_to_hf c. yz \<^bold>\ DOWN (Fun h x)\" + show "yz \<^bold>\ elem_of (Fun h x) + \ yz \<^bold>\ \yz \<^bold>\ ide_to_hf b * ide_to_hf c. yz \<^bold>\ elem_of (Fun h x)\" proof - - assume yz: "yz \<^bold>\ DOWN (Fun h x)" + assume yz: "yz \<^bold>\ elem_of (Fun h x)" have "yz \<^bold>\ ide_to_hf b * ide_to_hf c" proof - - have "DOWN (Fun h x) \<^bold>\ hexp (ide_to_hf b) (ide_to_hf c)" + have "elem_of (Fun h x) \<^bold>\ hexp (ide_to_hf b) (ide_to_hf c)" proof - have "ide (hf_to_ide (hexp (ide_to_hf b) (ide_to_hf c)))" using assms exp_def ide_exp by auto moreover have "Fun h x \ set (hf_to_ide (hexp (ide_to_hf b) (ide_to_hf c)))" proof - have "Fun h x \ Cod h" using assms x Fun_mapsto by blast moreover have "Cod h = set (hf_to_ide (hexp (ide_to_hf b) (ide_to_hf c)))" using assms 0 exp_def lam_simps(3) par by auto ultimately show ?thesis by blast qed ultimately show ?thesis - using DOWN_membI [of "hf_to_ide (hexp (ide_to_hf b) (ide_to_hf c))" + using elem_of_membI [of "hf_to_ide (hexp (ide_to_hf b) (ide_to_hf c))" "Fun h x"] by (simp add: ide_to_hf_hf_to_ide) qed thus ?thesis using assms yz hexp_def by auto qed thus ?thesis using assms x yz by blast qed - show "yz \<^bold>\ \yz \<^bold>\ ide_to_hf b * ide_to_hf c. yz \<^bold>\ DOWN (Fun h x)\ - \ yz \<^bold>\ DOWN (Fun h x)" + show "yz \<^bold>\ \yz \<^bold>\ ide_to_hf b * ide_to_hf c. yz \<^bold>\ elem_of (Fun h x)\ + \ yz \<^bold>\ elem_of (Fun h x)" using assms by simp qed qed - moreover have "UP (DOWN x) = x" - using assms x ide_char set_subset_Univ UP_DOWN - by (meson subsetD) + moreover have "arr_of (elem_of x) = x" + using arr_of_elem_of assms(1) set_ideD(1) x by blast ultimately show ?thesis using assms x arr_to_hfun_def ex_un_g by auto qed ultimately show ?thesis - using assms x theI' [of "\g. \DOWN x, g\ \<^bold>\ arr_to_hfun h"] + using assms x theI' [of "\g. \elem_of x, g\ \<^bold>\ arr_to_hfun h"] by fastforce qed finally show ?thesis using assms x by simp qed finally show ?thesis by simp qed ultimately show "Fun (\ a b c (comp (eval b c) (HF'.prod h b))) x = Fun h x" by blast qed qed qed section "The Main Results" interpretation cartesian_closed_category comp proof - interpret elementary_cartesian_closed_category comp pr0 pr1 some_terminal trm exp eval \ using ide_exp eval_in_hom lam_in_hom prod_ide_eq eval_prod_lam lam_eval_prod by unfold_locales auto show "cartesian_closed_category comp" using is_cartesian_closed_category by simp qed theorem is_cartesian_closed_category: shows "cartesian_closed_category comp" .. theorem is_category_with_finite_limits: shows "category_with_finite_limits comp" proof fix J :: "'j comp" assume J: "category J" interpret J: category J using J by simp assume finite: "finite (Collect J.arr)" have "has_products (Collect J.ide)" proof - have "Collect J.ide \ UNIV" using J.not_arr_null by blast moreover have "finite (Collect J.ide)" proof - have "Collect J.ide \ Collect J.arr" by auto thus ?thesis using finite J.ideD(1) finite_subset by blast qed ultimately show ?thesis using finite has_finite_products' by simp qed moreover have "has_products (Collect J.arr)" proof - have "Collect J.arr \ UNIV" using J.not_arr_null by blast thus ?thesis using finite has_finite_products' by simp qed ultimately show "has_limits_of_shape J" using J.category_axioms has_limits_if_has_products [of J] by simp qed end end diff --git a/thys/Category3/Limit.thy b/thys/Category3/Limit.thy --- a/thys/Category3/Limit.thy +++ b/thys/Category3/Limit.thy @@ -1,6208 +1,6271 @@ (* Title: Limit Author: Eugene W. Stark , 2016 Maintainer: Eugene W. Stark *) chapter Limit theory Limit imports FreeCategory DiscreteCategory Adjunction begin text\ This theory defines the notion of limit in terms of diagrams and cones and relates it to the concept of a representation of a functor. The diagonal functor associated with a diagram shape @{term J} is defined and it is shown that a right adjoint to the diagonal functor gives limits of shape @{term J} and that a category has limits of shape @{term J} if and only if the diagonal functor is a left adjoint functor. Products and equalizers are defined as special cases of limits, and it is shown that a category with equalizers has limits of shape @{term J} if it has products indexed by the sets of objects and arrows of @{term J}. The existence of limits in a set category is investigated, and it is shown that every set category has equalizers and that a set category @{term S} has @{term I}-indexed products if and only if the universe of @{term S} ``admits @{term I}-indexed tupling.'' The existence of limits in functor categories is also developed, showing that limits in functor categories are ``determined pointwise'' and that a functor category @{term "[A, B]"} has limits of shape @{term J} if @{term B} does. Finally, it is shown that the Yoneda functor preserves limits. This theory concerns itself only with limits; I have made no attempt to consider colimits. Although it would be possible to rework the entire development in dual form, it is possible that there is a more efficient way to dualize at least parts of it without repeating all the work. This is something that deserves further thought. \ section "Representations of Functors" text\ A representation of a contravariant functor \F: Cop \ S\, where @{term S} - is a replete set category that is the target of a hom-functor for @{term C}, consists of + is a set category that is the target of a hom-functor for @{term C}, consists of an object @{term a} of @{term C} and a natural isomorphism @{term "\: Y a \ F"}, where \Y: C \ [Cop, S]\ is the Yoneda functor. \ locale representation_of_functor = C: category C + Cop: dual_category C + - S: replete_set_category S + + S: set_category S setp + F: "functor" Cop.comp S F + - Hom: hom_functor C S \ + - Ya: yoneda_functor_fixed_object C S \ a + + Hom: hom_functor C S setp \ + + Ya: yoneda_functor_fixed_object C S setp \ a + natural_isomorphism Cop.comp S \Ya.Y a\ F \ for C :: "'c comp" (infixr "\" 55) and S :: "'s comp" (infixr "\\<^sub>S" 55) + and setp :: "'s set \ bool" and \ :: "'c * 'c \ 'c \ 's" and F :: "'c \ 's" and a :: 'c and \ :: "'c \ 's" begin abbreviation Y where "Y \ Ya.Y" abbreviation \ where "\ \ Hom.\" end text\ Two representations of the same functor are uniquely isomorphic. \ locale two_representations_one_functor = C: category C + Cop: dual_category C + - S: replete_set_category S + - F: set_valued_functor Cop.comp S F + - yoneda_functor C S \ + - Ya: yoneda_functor_fixed_object C S \ a + - Ya': yoneda_functor_fixed_object C S \ a' + - \: representation_of_functor C S \ F a \ + - \': representation_of_functor C S \ F a' \' + S: set_category S setp + + F: set_valued_functor Cop.comp S setp F + + yoneda_functor C S setp \ + + Ya: yoneda_functor_fixed_object C S setp \ a + + Ya': yoneda_functor_fixed_object C S setp \ a' + + \: representation_of_functor C S setp \ F a \ + + \': representation_of_functor C S setp \ F a' \' for C :: "'c comp" (infixr "\" 55) and S :: "'s comp" (infixr "\\<^sub>S" 55) + and setp :: "'s set \ bool" and F :: "'c \ 's" and \ :: "'c * 'c \ 'c \ 's" and a :: 'c and \ :: "'c \ 's" and a' :: 'c and \' :: "'c \ 's" begin interpretation \: inverse_transformation Cop.comp S \Y a\ F \ .. interpretation \': inverse_transformation Cop.comp S \Y a'\ F \' .. interpretation \\': vertical_composite Cop.comp S \Y a\ F \Y a'\ \ \'.map .. interpretation \'\: vertical_composite Cop.comp S \Y a'\ F \Y a\ \' \.map .. lemma are_uniquely_isomorphic: shows "\!\. \\ : a \ a'\ \ C.iso \ \ map \ = Cop_S.MkArr (Y a) (Y a') \\'.map" proof - interpret \\': natural_isomorphism Cop.comp S \Y a\ \Y a'\ \\'.map using \.natural_isomorphism_axioms \'.natural_isomorphism_axioms natural_isomorphisms_compose by blast interpret \'\: natural_isomorphism Cop.comp S \Y a'\ \Y a\ \'\.map using \'.natural_isomorphism_axioms \.natural_isomorphism_axioms natural_isomorphisms_compose by blast interpret \\'_\'\: inverse_transformations Cop.comp S \Y a\ \Y a'\ \\'.map \'\.map proof fix x assume X: "Cop.ide x" show "S.inverse_arrows (\\'.map x) (\'\.map x)" using S.inverse_arrows_compose S.inverse_arrows_sym X \'\.map_simp_ide \\'.map_simp_ide \'.inverts_components \.inverts_components by force qed have "Cop_S.inverse_arrows (Cop_S.MkArr (Y a) (Y a') \\'.map) (Cop_S.MkArr (Y a') (Y a) \'\.map)" proof - have Ya: "functor Cop.comp S (Y a)" .. have Ya': "functor Cop.comp S (Y a')" .. have \\': "natural_transformation Cop.comp S (Y a) (Y a') \\'.map" .. have \'\: "natural_transformation Cop.comp S (Y a') (Y a) \'\.map" .. show ?thesis by (metis (no_types, lifting) Cop_S.arr_MkArr Cop_S.comp_MkArr Cop_S.ide_MkIde Cop_S.inverse_arrows_def Ya'.functor_axioms Ya.functor_axioms \'\.natural_transformation_axioms \\'.natural_transformation_axioms \\'_\'\.inverse_transformations_axioms inverse_transformations_inverse(1-2) mem_Collect_eq) qed hence 3: "Cop_S.iso (Cop_S.MkArr (Y a) (Y a') \\'.map)" using Cop_S.isoI by blast hence "\f. \f : a \ a'\ \ map f = Cop_S.MkArr (Y a) (Y a') \\'.map" using Ya.ide_a Ya'.ide_a is_full Y_def Cop_S.iso_is_arr full_functor.is_full Cop_S.MkArr_in_hom \\'.natural_transformation_axioms preserves_ide by force from this obtain \ where \: "\\ : a \ a'\ \ map \ = Cop_S.MkArr (Y a) (Y a') \\'.map" by blast show ?thesis by (metis 3 C.in_homE \ is_faithful reflects_iso) qed end section "Diagrams and Cones" text\ A \emph{diagram} in a category @{term C} is a functor \D: J \ C\. We refer to the category @{term J} as the diagram \emph{shape}. Note that in the usual expositions of category theory that use set theory as their foundations, the shape @{term J} of a diagram is required to be a ``small'' category, where smallness means that the collection of objects of @{term J}, as well as each of the ``homs,'' is a set. However, in HOL there is no class of all sets, so it is not meaningful to speak of @{term J} as ``small'' in any kind of absolute sense. There is likely a meaningful notion of smallness of @{term J} \emph{relative to} @{term C} (the result below that states that a set category has @{term I}-indexed products if and only if its universe ``admits @{term I}-indexed tuples'' is suggestive of how this might be defined), but I haven't fully explored this idea at present. \ locale diagram = C: category C + J: category J + "functor" J C D for J :: "'j comp" (infixr "\\<^sub>J" 55) and C :: "'c comp" (infixr "\" 55) and D :: "'j \ 'c" begin notation J.in_hom ("\_ : _ \\<^sub>J _\") end lemma comp_diagram_functor: assumes "diagram J C D" and "functor J' J F" shows "diagram J' C (D o F)" by (meson assms(1) assms(2) diagram_def functor.axioms(1) functor_comp) text\ A \emph{cone} over a diagram \D: J \ C\ is a natural transformation from a constant functor to @{term D}. The value of the constant functor is the \emph{apex} of the cone. \ locale cone = C: category C + J: category J + D: diagram J C D + A: constant_functor J C a + natural_transformation J C A.map D \ for J :: "'j comp" (infixr "\\<^sub>J" 55) and C :: "'c comp" (infixr "\" 55) and D :: "'j \ 'c" and a :: 'c and \ :: "'j \ 'c" begin lemma ide_apex: shows "C.ide a" using A.value_is_ide by auto lemma component_in_hom: assumes "J.arr j" shows "\\ j : a \ D (J.cod j)\" using assms by auto end text\ A cone over diagram @{term D} is transformed into a cone over diagram @{term "D o F"} by pre-composing with @{term F}. \ lemma comp_cone_functor: assumes "cone J C D a \" and "functor J' J F" shows "cone J' C (D o F) a (\ o F)" proof - interpret \: cone J C D a \ using assms(1) by auto interpret F: "functor" J' J F using assms(2) by auto interpret A': constant_functor J' C a using \.A.value_is_ide by unfold_locales auto have 1: "\.A.map o F = A'.map" using \.A.map_def A'.map_def \.J.not_arr_null by auto interpret \': natural_transformation J' C A'.map \D o F\ \\ o F\ using 1 horizontal_composite F.as_nat_trans.natural_transformation_axioms \.natural_transformation_axioms by fastforce show "cone J' C (D o F) a (\ o F)" .. qed text\ A cone over diagram @{term D} can be transformed into a cone over a diagram @{term D'} by post-composing with a natural transformation from @{term D} to @{term D'}. \ lemma vcomp_transformation_cone: assumes "cone J C D a \" and "natural_transformation J C D D' \" shows "cone J C D' a (vertical_composite.map J C \ \)" by (meson assms cone.axioms(4-5) cone.intro diagram.intro natural_transformation.axioms(1-4) vertical_composite.intro vertical_composite.is_natural_transformation) context "functor" begin lemma preserves_diagrams: fixes J :: "'j comp" assumes "diagram J A D" shows "diagram J B (F o D)" by (meson assms diagram_def functor_axioms functor_comp functor_def) lemma preserves_cones: fixes J :: "'j comp" assumes "cone J A D a \" shows "cone J B (F o D) (F a) (F o \)" proof - interpret \: cone J A D a \ using assms by auto interpret Fa: constant_functor J B \F a\ using \.ide_apex by unfold_locales auto have 1: "F o \.A.map = Fa.map" proof fix f show "(F \ \.A.map) f = Fa.map f" using is_extensional Fa.is_extensional \.A.is_extensional by (cases "\.J.arr f", simp_all) qed interpret \': natural_transformation J B Fa.map \F o D\ \F o \\ using 1 horizontal_composite \.natural_transformation_axioms as_nat_trans.natural_transformation_axioms by fastforce show "cone J B (F o D) (F a) (F o \)" .. qed end context diagram begin abbreviation cone where "cone a \ \ Limit.cone J C D a \" abbreviation cones :: "'c \ ('j \ 'c) set" where "cones a \ { \. cone a \ }" text\ An arrow @{term "f \ C.hom a' a"} induces by composition a transformation from cones with apex @{term a} to cones with apex @{term a'}. This transformation is functorial in @{term f}. \ abbreviation cones_map :: "'c \ ('j \ 'c) \ ('j \ 'c)" where "cones_map f \ (\\ \ cones (C.cod f). \j. if J.arr j then \ j \ f else C.null)" lemma cones_map_mapsto: assumes "C.arr f" shows "cones_map f \ extensional (cones (C.cod f)) \ (cones (C.cod f) \ cones (C.dom f))" proof show "cones_map f \ extensional (cones (C.cod f))" by blast show "cones_map f \ cones (C.cod f) \ cones (C.dom f)" proof fix \ assume "\ \ cones (C.cod f)" hence \: "cone (C.cod f) \" by auto interpret \: cone J C D \C.cod f\ \ using \ by auto interpret B: constant_functor J C \C.dom f\ using assms by unfold_locales auto have "cone (C.dom f) (\j. if J.arr j then \ j \ f else C.null)" using assms B.value_is_ide \.is_natural_1 \.is_natural_2 apply (unfold_locales, auto) using \.is_natural_1 apply (metis C.comp_assoc) using \.is_natural_2 C.comp_arr_dom by (metis J.arr_cod_iff_arr J.cod_cod C.comp_assoc) thus "(\j. if J.arr j then \ j \ f else C.null) \ cones (C.dom f)" by auto qed qed lemma cones_map_ide: assumes "\ \ cones a" shows "cones_map a \ = \" proof - interpret \: cone J C D a \ using assms by auto show ?thesis proof fix j show "cones_map a \ j = \ j" using assms \.A.value_is_ide \.preserves_hom C.comp_arr_dom \.is_extensional by (cases "J.arr j", auto) qed qed lemma cones_map_comp: assumes "C.seq f g" shows "cones_map (f \ g) = restrict (cones_map g o cones_map f) (cones (C.cod f))" proof (intro restr_eqI) show "cones (C.cod (f \ g)) = cones (C.cod f)" using assms by simp show "\\. \ \ cones (C.cod (f \ g)) \ (\j. if J.arr j then \ j \ f \ g else C.null) = (cones_map g o cones_map f) \" proof - fix \ assume \: "\ \ cones (C.cod (f \ g))" show "(\j. if J.arr j then \ j \ f \ g else C.null) = (cones_map g o cones_map f) \" proof - have "((cones_map g) o (cones_map f)) \ = cones_map g (cones_map f \)" by force also have "... = (\j. if J.arr j then (\j. if J.arr j then \ j \ f else C.null) j \ g else C.null)" proof fix j have "cone (C.dom f) (cones_map f \)" using assms \ cones_map_mapsto by (elim C.seqE, force) thus "cones_map g (cones_map f \) j = (if J.arr j then C (if J.arr j then \ j \ f else C.null) g else C.null)" using \ assms by auto qed also have "... = (\j. if J.arr j then \ j \ f \ g else C.null)" using C.comp_assoc by fastforce finally show ?thesis by auto qed qed qed end text\ Changing the apex of a cone by pre-composing with an arrow @{term f} commutes with changing the diagram of a cone by post-composing with a natural transformation. \ lemma cones_map_vcomp: assumes "diagram J C D" and "diagram J C D'" and "natural_transformation J C D D' \" and "cone J C D a \" and f: "partial_magma.in_hom C f a' a" shows "diagram.cones_map J C D' f (vertical_composite.map J C \ \) = vertical_composite.map J C (diagram.cones_map J C D f \) \" proof - interpret D: diagram J C D using assms(1) by auto interpret D': diagram J C D' using assms(2) by auto interpret \: natural_transformation J C D D' \ using assms(3) by auto interpret \: cone J C D a \ using assms(4) by auto interpret \o\: vertical_composite J C \.A.map D D' \ \ .. interpret \o\: cone J C D' a \o\.map .. interpret \f: cone J C D a' \D.cones_map f \\ using f \.cone_axioms D.cones_map_mapsto by blast interpret \o\f: vertical_composite J C \f.A.map D D' \D.cones_map f \\ \ .. interpret \o\_f: cone J C D' a' \D'.cones_map f \o\.map\ using f \o\.cone_axioms D'.cones_map_mapsto [of f] by blast write C (infixr "\" 55) show "D'.cones_map f \o\.map = \o\f.map" proof (intro NaturalTransformation.eqI) show "natural_transformation J C \f.A.map D' (D'.cones_map f \o\.map)" .. show "natural_transformation J C \f.A.map D' \o\f.map" .. show "\j. D.J.ide j \ D'.cones_map f \o\.map j = \o\f.map j" proof - fix j assume j: "D.J.ide j" have "D'.cones_map f \o\.map j = \o\.map j \ f" using f \o\.cone_axioms \o\.map_simp_2 \o\.is_extensional by auto also have "... = (\ j \ \ (D.J.dom j)) \ f" using j \o\.map_simp_2 by simp also have "... = \ j \ \ (D.J.dom j) \ f" using D.C.comp_assoc by simp also have "... = \o\f.map j" using j f \.cone_axioms \o\f.map_simp_2 by auto finally show "D'.cones_map f \o\.map j = \o\f.map j" by auto qed qed qed text\ Given a diagram @{term D}, we can construct a contravariant set-valued functor, which takes each object @{term a} of @{term C} to the set of cones over @{term D} with apex @{term a}, and takes each arrow @{term f} of @{term C} to the function on cones over @{term D} induced by pre-composition with @{term f}. For this, we need to introduce a set category @{term S} whose universe is large enough to contain all the cones over @{term D}, and we need to have an explicit correspondence between cones and elements of the universe of @{term S}. A replete set category @{term S} equipped with an injective mapping @{term_type "\ :: ('j => 'c) => 's"} serves this purpose. \ locale cones_functor = C: category C + Cop: dual_category C + J: category J + D: diagram J C D + S: replete_concrete_set_category S UNIV \ for J :: "'j comp" (infixr "\\<^sub>J" 55) and C :: "'c comp" (infixr "\" 55) and D :: "'j \ 'c" and S :: "'s comp" (infixr "\\<^sub>S" 55) and \ :: "('j \ 'c) \ 's" begin notation S.in_hom ("\_ : _ \\<^sub>S _\") - abbreviation \ where "\ \ S.\" + abbreviation \ where "\ \ S.DN" definition map :: "'c \ 's" where "map = (\f. if C.arr f then S.mkArr (\ ` D.cones (C.cod f)) (\ ` D.cones (C.dom f)) (\ o D.cones_map f o \) else S.null)" lemma map_simp [simp]: assumes "C.arr f" shows "map f = S.mkArr (\ ` D.cones (C.cod f)) (\ ` D.cones (C.dom f)) (\ o D.cones_map f o \)" using assms map_def by auto lemma arr_map: assumes "C.arr f" shows "S.arr (map f)" proof - have "\ o D.cones_map f o \ \ \ ` D.cones (C.cod f) \ \ ` D.cones (C.dom f)" using assms D.cones_map_mapsto by force - thus ?thesis using assms S.\_mapsto S.arr_mkArr by auto + thus ?thesis using assms S.UP_mapsto by auto qed lemma map_ide: assumes "C.ide a" shows "map a = S.mkIde (\ ` D.cones a)" proof - have "map a = S.mkArr (\ ` D.cones a) (\ ` D.cones a) (\ o D.cones_map a o \)" using assms map_simp by force also have "... = S.mkArr (\ ` D.cones a) (\ ` D.cones a) (\x. x)" - using S.\_mapsto D.cones_map_ide S.arr_mkArr by force + using S.UP_mapsto D.cones_map_ide by force also have "... = S.mkIde (\ ` D.cones a)" - using assms S.mkIde_as_mkArr S.\_mapsto by blast + using assms S.mkIde_as_mkArr S.UP_mapsto by blast finally show ?thesis by auto qed lemma map_preserves_dom: assumes "Cop.arr f" shows "map (Cop.dom f) = S.dom (map f)" using assms arr_map map_ide by auto lemma map_preserves_cod: assumes "Cop.arr f" shows "map (Cop.cod f) = S.cod (map f)" using assms arr_map map_ide by auto lemma map_preserves_comp: assumes "Cop.seq g f" shows "map (g \\<^sup>o\<^sup>p f) = map g \\<^sub>S map f" proof - have "map (g \\<^sup>o\<^sup>p f) = S.mkArr (\ ` D.cones (C.cod f)) (\ ` D.cones (C.dom g)) ((\ o D.cones_map g o \) o (\ o D.cones_map f o \))" proof - have 1: "S.arr (map (g \\<^sup>o\<^sup>p f))" using assms arr_map [of "C f g"] by simp have "map (g \\<^sup>o\<^sup>p f) = S.mkArr (\ ` D.cones (C.cod f)) (\ ` D.cones (C.dom g)) (\ o D.cones_map (C f g) o \)" using assms map_simp [of "C f g"] by simp also have "... = S.mkArr (\ ` D.cones (C.cod f)) (\ ` D.cones (C.dom g)) ((\ o D.cones_map g o \) o (\ o D.cones_map f o \))" using assms 1 calculation D.cones_map_mapsto D.cones_map_comp by auto finally show ?thesis by blast qed also have "... = map g \\<^sub>S map f" using assms arr_map [of f] arr_map [of g] map_simp S.comp_mkArr by auto finally show ?thesis by auto qed lemma is_functor: shows "functor Cop.comp S map" apply (unfold_locales) using map_def arr_map map_preserves_dom map_preserves_cod map_preserves_comp by auto end sublocale cones_functor \ "functor" Cop.comp S map using is_functor by auto - sublocale cones_functor \ set_valued_functor Cop.comp S map .. + sublocale cones_functor \ set_valued_functor Cop.comp S \\A. A \ S.Univ\ map .. section Limits subsection "Limit Cones" text\ A \emph{limit cone} for a diagram @{term D} is a cone @{term \} over @{term D} with the universal property that any other cone @{term \'} over the diagram @{term D} factors uniquely through @{term \}. \ locale limit_cone = C: category C + J: category J + D: diagram J C D + cone J C D a \ for J :: "'j comp" (infixr "\\<^sub>J" 55) and C :: "'c comp" (infixr "\" 55) and D :: "'j \ 'c" and a :: 'c and \ :: "'j \ 'c" + assumes is_universal: "cone J C D a' \' \ \!f. \f : a' \ a\ \ D.cones_map f \ = \'" begin definition induced_arrow :: "'c \ ('j \ 'c) \ 'c" where "induced_arrow a' \' = (THE f. \f : a' \ a\ \ D.cones_map f \ = \')" lemma induced_arrowI: assumes \': "\' \ D.cones a'" shows "\induced_arrow a' \' : a' \ a\" and "D.cones_map (induced_arrow a' \') \ = \'" proof - have "\!f. \f : a' \ a\ \ D.cones_map f \ = \'" using assms \' is_universal by simp hence 1: "\induced_arrow a' \' : a' \ a\ \ D.cones_map (induced_arrow a' \') \ = \'" using theI' [of "\f. \f : a' \ a\ \ D.cones_map f \ = \'"] induced_arrow_def by presburger show "\induced_arrow a' \' : a' \ a\" using 1 by simp show "D.cones_map (induced_arrow a' \') \ = \'" using 1 by simp qed lemma cones_map_induced_arrow: shows "induced_arrow a' \ D.cones a' \ C.hom a' a" and "\\'. \' \ D.cones a' \ D.cones_map (induced_arrow a' \') \ = \'" using induced_arrowI by auto lemma induced_arrow_cones_map: assumes "C.ide a'" shows "(\f. D.cones_map f \) \ C.hom a' a \ D.cones a'" and "\f. \f : a' \ a\ \ induced_arrow a' (D.cones_map f \) = f" proof - have a': "C.ide a'" using assms by (simp add: cone.ide_apex) have cone_\: "cone J C D a \" .. show "(\f. D.cones_map f \) \ C.hom a' a \ D.cones a'" using cone_\ D.cones_map_mapsto by blast fix f assume f: "\f : a' \ a\" show "induced_arrow a' (D.cones_map f \) = f" proof - have "D.cones_map f \ \ D.cones a'" using f cone_\ D.cones_map_mapsto by blast hence "\!f'. \f' : a' \ a\ \ D.cones_map f' \ = D.cones_map f \" using assms is_universal by auto thus ?thesis using f induced_arrow_def the1_equality [of "\f'. \f' : a' \ a\ \ D.cones_map f' \ = D.cones_map f \"] by presburger qed qed text\ For a limit cone @{term \} with apex @{term a}, for each object @{term a'} the hom-set @{term "C.hom a' a"} is in bijective correspondence with the set of cones with apex @{term a'}. \ lemma bij_betw_hom_and_cones: assumes "C.ide a'" shows "bij_betw (\f. D.cones_map f \) (C.hom a' a) (D.cones a')" proof (intro bij_betwI) show "(\f. D.cones_map f \) \ C.hom a' a \ D.cones a'" using assms induced_arrow_cones_map by blast show "induced_arrow a' \ D.cones a' \ C.hom a' a" using assms cones_map_induced_arrow by blast show "\f. f \ C.hom a' a \ induced_arrow a' (D.cones_map f \) = f" using assms induced_arrow_cones_map by blast show "\\'. \' \ D.cones a' \ D.cones_map (induced_arrow a' \') \ = \'" using assms cones_map_induced_arrow by blast qed lemma induced_arrow_eqI: assumes "D.cone a' \'" and "\f : a' \ a\" and "D.cones_map f \ = \'" shows "induced_arrow a' \' = f" using assms is_universal induced_arrow_def the1_equality [of "\f. f \ C.hom a' a \ D.cones_map f \ = \'" f] by simp lemma induced_arrow_self: shows "induced_arrow a \ = a" proof - have "\a : a \ a\ \ D.cones_map a \ = \" using ide_apex cone_axioms D.cones_map_ide by force thus ?thesis using induced_arrow_eqI cone_axioms by auto qed end context diagram begin abbreviation limit_cone where "limit_cone a \ \ Limit.limit_cone J C D a \" text\ A diagram @{term D} has object @{term a} as a limit if @{term a} is the apex of some limit cone over @{term D}. \ abbreviation has_as_limit :: "'c \ bool" where "has_as_limit a \ (\\. limit_cone a \)" abbreviation has_limit where "has_limit \ (\a \. limit_cone a \)" definition some_limit :: 'c where "some_limit = (SOME a. \\. limit_cone a \)" definition some_limit_cone :: "'j \ 'c" where "some_limit_cone = (SOME \. limit_cone some_limit \)" lemma limit_cone_some_limit_cone: assumes has_limit shows "limit_cone some_limit some_limit_cone" proof - have "\a. has_as_limit a" using assms by simp hence "has_as_limit some_limit" using some_limit_def someI_ex [of "\a. \\. limit_cone a \"] by simp thus "limit_cone some_limit some_limit_cone" using assms some_limit_cone_def someI_ex [of "\\. limit_cone some_limit \"] by simp qed lemma ex_limitE: assumes "\a. has_as_limit a" obtains a \ where "limit_cone a \" using assms someI_ex by blast end subsection "Limits by Representation" text\ A limit for a diagram D can also be given by a representation \(a, \)\ of the cones functor. \ locale representation_of_cones_functor = C: category C + Cop: dual_category C + J: category J + D: diagram J C D + S: replete_concrete_set_category S UNIV \ + Cones: cones_functor J C D S \ + - Hom: hom_functor C S \ + - representation_of_functor C S \ Cones.map a \ + Hom: hom_functor C S \\A. A \ S.Univ\ \ + + representation_of_functor C S S.setp \ Cones.map a \ for J :: "'j comp" (infixr "\\<^sub>J" 55) and C :: "'c comp" (infixr "\" 55) and D :: "'j \ 'c" and S :: "'s comp" (infixr "\\<^sub>S" 55) and \ :: "'c * 'c \ 'c \ 's" and \ :: "('j \ 'c) \ 's" and a :: 'c and \ :: "'c \ 's" subsection "Putting it all Together" text\ A ``limit situation'' combines and connects the ways of presenting a limit. \ locale limit_situation = C: category C + Cop: dual_category C + J: category J + D: diagram J C D + S: replete_concrete_set_category S UNIV \ + Cones: cones_functor J C D S \ + - Hom: hom_functor C S \ + - \: representation_of_functor C S \ Cones.map a \ + + Hom: hom_functor C S S.setp \ + + \: representation_of_functor C S S.setp \ Cones.map a \ + \: limit_cone J C D a \ for J :: "'j comp" (infixr "\\<^sub>J" 55) and C :: "'c comp" (infixr "\" 55) and D :: "'j \ 'c" and S :: "'s comp" (infixr "\\<^sub>S" 55) and \ :: "'c * 'c \ 'c \ 's" and \ :: "('j \ 'c) \ 's" and a :: 'c and \ :: "'c \ 's" and \ :: "'j \ 'c" + - assumes \_in_terms_of_\: "\ = S.\ (S.Fun (\ a) (\ (a, a) a))" + assumes \_in_terms_of_\: "\ = S.DN (S.Fun (\ a) (\ (a, a) a))" and \_in_terms_of_\: "Cop.ide a' \ \ a' = S.mkArr (Hom.set (a', a)) (\ ` D.cones a') (\x. \ (D.cones_map (Hom.\ (a', a) x) \))" text (in limit_situation) \ The assumption @{prop \_in_terms_of_\} states that the universal cone @{term \} is obtained by applying the function @{term "S.Fun (\ a)"} to the identity @{term a} of @{term[source=true] C} (after taking into account the necessary coercions). \ text (in limit_situation) \ The assumption @{prop \_in_terms_of_\} states that the component of @{term \} at @{term a'} is the arrow of @{term[source=true] S} corresponding to the function that takes an arrow @{term "f \ C.hom a' a"} and produces the cone with vertex @{term a'} obtained by transforming the universal cone @{term \} by @{term f}. \ subsection "Limit Cones Induce Limit Situations" text\ To obtain a limit situation from a limit cone, we need to introduce a set category that is large enough to contain the hom-sets of @{term C} as well as the cones over @{term D}. We use the category of all @{typ "('c + ('j \ 'c))"}-sets for this. \ context limit_cone begin interpretation Cop: dual_category C .. interpretation CopxC: product_category Cop.comp C .. interpretation S: replete_setcat \undefined :: 'c + ('j \ 'c)\ . notation S.comp (infixr "\\<^sub>S" 55) interpretation Sr: replete_concrete_set_category S.comp UNIV \S.UP o Inr\ apply unfold_locales using S.UP_mapsto apply auto[1] using S.inj_UP inj_Inr inj_compose by metis interpretation Cones: cones_functor J C D S.comp \S.UP o Inr\ .. - interpretation Hom: hom_functor C S.comp \\_. S.UP o Inl\ + interpretation Hom: hom_functor C S.comp S.setp \\_. S.UP o Inl\ apply (unfold_locales) using S.UP_mapsto - apply auto[1] + apply auto[1] using S.inj_UP injD inj_onI inj_Inl inj_compose - by (metis (no_types, lifting)) - - interpretation Y: yoneda_functor C S.comp \\_. S.UP o Inl\ .. - interpretation Ya: yoneda_functor_fixed_object C S.comp \\_. S.UP o Inl\ a + apply (metis (no_types, lifting)) + using S.UP_mapsto + by auto + + interpretation Y: yoneda_functor C S.comp S.setp \\_. S.UP o Inl\ .. + interpretation Ya: yoneda_functor_fixed_object C S.comp S.setp \\_. S.UP o Inl\ a apply (unfold_locales) using ide_apex by auto abbreviation inl :: "'c \ 'c + ('j \ 'c)" where "inl \ Inl" abbreviation inr :: "('j \ 'c) \ 'c + ('j \ 'c)" where "inr \ Inr" abbreviation \ where "\ \ S.UP o inr" abbreviation \ where "\ \ Cones.\" abbreviation \ where "\ \ \_. S.UP o inl" abbreviation \ where "\ \ Hom.\" abbreviation Y where "Y \ Y.Y" lemma Ya_ide: assumes a': "C.ide a'" shows "Y a a' = S.mkIde (Hom.set (a', a))" using assms ide_apex Y.Y_simp Hom.map_ide by simp lemma Ya_arr: assumes g: "C.arr g" shows "Y a g = S.mkArr (Hom.set (C.cod g, a)) (Hom.set (C.dom g, a)) (\ (C.dom g, a) o Cop.comp g o \ (C.cod g, a))" using ide_apex g Y.Y_ide_arr [of a g "C.dom g" "C.cod g"] by auto lemma cone_\ [simp]: shows "\ \ D.cones a" using cone_axioms by simp text\ For each object @{term a'} of @{term[source=true] C} we have a function mapping @{term "C.hom a' a"} to the set of cones over @{term D} with apex @{term a'}, which takes @{term "f \ C.hom a' a"} to \\f\, where \\f\ is the cone obtained by composing @{term \} with @{term f} (after accounting for coercions to and from the universe of @{term S}). The corresponding arrows of @{term S} are the components of a natural isomorphism from @{term "Y a"} to \Cones\. \ definition \o :: "'c \ ('c + ('j \ 'c)) setcat.arr" where "\o a' = S.mkArr (Hom.set (a', a)) (\ ` D.cones a') (\x. \ (D.cones_map (\ (a', a) x) \))" lemma \o_in_hom: assumes a': "C.ide a'" shows "\\o a' : S.mkIde (Hom.set (a', a)) \\<^sub>S S.mkIde (\ ` D.cones a')\" proof - have " \S.mkArr (Hom.set (a', a)) (\ ` D.cones a') (\x. \ (D.cones_map (\ (a', a) x) \)) : S.mkIde (Hom.set (a', a)) \\<^sub>S S.mkIde (\ ` D.cones a')\" proof - have "(\x. \ (D.cones_map (\ (a', a) x) \)) \ Hom.set (a', a) \ \ ` D.cones a'" proof fix x assume x: "x \ Hom.set (a', a)" hence "\\ (a', a) x : a' \ a\" using ide_apex a' Hom.\_mapsto by auto hence "D.cones_map (\ (a', a) x) \ \ D.cones a'" using ide_apex a' x D.cones_map_mapsto cone_\ by force thus "\ (D.cones_map (\ (a', a) x) \) \ \ ` D.cones a'" by simp qed moreover have "Hom.set (a', a) \ S.Univ" using ide_apex a' Hom.set_subset_Univ by auto moreover have "\ ` D.cones a' \ S.Univ" using S.UP_mapsto by auto ultimately show ?thesis using S.mkArr_in_hom by simp qed thus ?thesis using \o_def [of a'] by auto qed interpretation \: transformation_by_components Cop.comp S.comp \Y a\ Cones.map \o proof fix a' assume A': "Cop.ide a'" show "\\o a' : Y a a' \\<^sub>S Cones.map a'\" using A' Ya_ide \o_in_hom Cones.map_ide by auto next fix g assume g: "Cop.arr g" show "\o (Cop.cod g) \\<^sub>S Y a g = Cones.map g \\<^sub>S \o (Cop.dom g)" proof - let ?A = "Hom.set (C.cod g, a)" let ?B = "Hom.set (C.dom g, a)" let ?B' = "\ ` D.cones (C.cod g)" let ?C = "\ ` D.cones (C.dom g)" let ?F = "\ (C.dom g, a) o Cop.comp g o \ (C.cod g, a)" let ?F' = "\ o D.cones_map g o \" let ?G = "\x. \ (D.cones_map (\ (C.dom g, a) x) \)" let ?G' = "\x. \ (D.cones_map (\ (C.cod g, a) x) \)" have "S.arr (Y a g) \ Y a g = S.mkArr ?A ?B ?F" using ide_apex g Ya.preserves_arr Ya_arr by fastforce moreover have "S.arr (\o (Cop.cod g))" using g \o_in_hom [of "Cop.cod g"] by auto moreover have "\o (Cop.cod g) = S.mkArr ?B ?C ?G" using g \o_def [of "C.dom g"] by auto moreover have "S.seq (\o (Cop.cod g)) (Y a g)" using ide_apex g \o_in_hom [of "Cop.cod g"] by auto ultimately have 1: "S.seq (\o (Cop.cod g)) (Y a g) \ \o (Cop.cod g) \\<^sub>S Y a g = S.mkArr ?A ?C (?G o ?F)" using S.comp_mkArr [of ?A ?B ?F ?C ?G] by argo have "Cones.map g = S.mkArr (\ ` D.cones (C.cod g)) (\ ` D.cones (C.dom g)) ?F'" using g Cones.map_simp by fastforce moreover have "\o (Cop.dom g) = S.mkArr ?A ?B' ?G'" using g \o_def by fastforce moreover have "S.seq (Cones.map g) (\o (Cop.dom g))" using g Cones.preserves_hom [of g "C.cod g" "C.dom g"] \o_in_hom [of "Cop.dom g"] by force ultimately have 2: "S.seq (Cones.map g) (\o (Cop.dom g)) \ Cones.map g \\<^sub>S \o (Cop.dom g) = S.mkArr ?A ?C (?F' o ?G')" using S.seqI' [of "\o (Cop.dom g)" "Cones.map g"] S.comp_mkArr by auto have "\o (Cop.cod g) \\<^sub>S Y a g = S.mkArr ?A ?C (?G o ?F)" using 1 by auto also have "... = S.mkArr ?A ?C (?F' o ?G')" proof (intro S.mkArr_eqI') show "S.arr (S.mkArr ?A ?C (?G o ?F))" using 1 by force show "\x. x \ ?A \ (?G o ?F) x = (?F' o ?G') x" proof - fix x assume x: "x \ ?A" hence 1: "\\ (C.cod g, a) x : C.cod g \ a\" using ide_apex g Hom.\_mapsto [of "C.cod g" a] by auto have "(?G o ?F) x = \ (D.cones_map (\ (C.dom g, a) (\ (C.dom g, a) (\ (C.cod g, a) x \ g))) \)" proof - (* Why is it so balky with this proof? *) have "(?G o ?F) x = ?G (?F x)" by simp also have "... = \ (D.cones_map (\ (C.dom g, a) (\ (C.dom g, a) (\ (C.cod g, a) x \ g))) \)" by (metis Cop.comp_def comp_apply) finally show ?thesis by auto qed also have "... = \ (D.cones_map (\ (C.cod g, a) x \ g) \)" proof - have "\\ (C.cod g, a) x \ g : C.dom g \ a\" using g 1 by auto thus ?thesis using Hom.\_\ by presburger qed also have "... = \ (D.cones_map g (D.cones_map (\ (C.cod g, a) x) \))" using g x 1 cone_\ D.cones_map_comp [of "\ (C.cod g, a) x" g] by fastforce also have "... = \ (D.cones_map g (\ (\ (D.cones_map (\ (C.cod g, a) x) \))))" - using 1 cone_\ D.cones_map_mapsto Sr.\_\ by auto + using 1 cone_\ D.cones_map_mapsto Sr.DN_UP by auto also have "... = (?F' o ?G') x" by simp finally show "(?G o ?F) x = (?F' o ?G') x" by auto qed qed also have "... = Cones.map g \\<^sub>S \o (Cop.dom g)" using 2 by auto finally show ?thesis by auto qed qed - interpretation \: set_valued_transformation Cop.comp S.comp \Y a\ Cones.map \.map .. + interpretation \: set_valued_transformation Cop.comp S.comp S.setp + \Y a\ Cones.map \.map .. interpretation \: natural_isomorphism Cop.comp S.comp \Y a\ Cones.map \.map proof fix a' assume a': "Cop.ide a'" show "S.iso (\.map a')" proof - let ?F = "\x. \ (D.cones_map (\ (a', a) x) \)" have bij: "bij_betw ?F (Hom.set (a', a)) (\ ` D.cones a')" proof - have "\x x'. \ x \ Hom.set (a', a); x' \ Hom.set (a', a); \ (D.cones_map (\ (a', a) x) \) = \ (D.cones_map (\ (a', a) x') \) \ \ x = x'" proof - fix x x' assume x: "x \ Hom.set (a', a)" and x': "x' \ Hom.set (a', a)" and xx': "\ (D.cones_map (\ (a', a) x) \) = \ (D.cones_map (\ (a', a) x') \)" have \x: "\\ (a', a) x : a' \ a\" using x ide_apex a' Hom.\_mapsto by auto have \x': "\\ (a', a) x' : a' \ a\" using x' ide_apex a' Hom.\_mapsto by auto have 1: "\!f. \f : a' \ a\ \ \ (D.cones_map f \) = \ (D.cones_map (\ (a', a) x) \)" proof - have "D.cones_map (\ (a', a) x) \ \ D.cones a'" using \x a' cone_\ D.cones_map_mapsto by force hence 2: "\!f. \f : a' \ a\ \ D.cones_map f \ = D.cones_map (\ (a', a) x) \" using a' is_universal by simp show "\!f. \f : a' \ a\ \ \ (D.cones_map f \) = \ (D.cones_map (\ (a', a) x) \)" proof - have "\f. \ (D.cones_map f \) = \ (D.cones_map (\ (a', a) x) \) \ D.cones_map f \ = D.cones_map (\ (a', a) x) \" proof - fix f :: 'c have "D.cones_map f \ = D.cones_map (\ (a', a) x) \ \ \ (D.cones_map f \) = \ (D.cones_map (\ (a', a) x) \)" by simp thus "(\ (D.cones_map f \) = \ (D.cones_map (\ (a', a) x) \)) = (D.cones_map f \ = D.cones_map (\ (a', a) x) \)" - by (meson Sr.inj_\ injD) + by (meson Sr.inj_UP injD) qed thus ?thesis using 2 by auto qed qed have 2: "\!x''. x'' \ Hom.set (a', a) \ \ (D.cones_map (\ (a', a) x'') \) = \ (D.cones_map (\ (a', a) x) \)" proof - from 1 obtain f'' where f'': "\f'' : a' \ a\ \ \ (D.cones_map f'' \) = \ (D.cones_map (\ (a', a) x) \)" by blast have "\ (a', a) f'' \ Hom.set (a', a) \ \ (D.cones_map (\ (a', a) (\ (a', a) f'')) \) = \ (D.cones_map (\ (a', a) x) \)" proof show "\ (a', a) f'' \ Hom.set (a', a)" using f'' Hom.set_def by auto show "\ (D.cones_map (\ (a', a) (\ (a', a) f'')) \) = \ (D.cones_map (\ (a', a) x) \)" using f'' Hom.\_\ by presburger qed moreover have "\x''. x'' \ Hom.set (a', a) \ \ (D.cones_map (\ (a', a) x'') \) = \ (D.cones_map (\ (a', a) x) \) \ x'' = \ (a', a) f''" proof - fix x'' assume x'': "x'' \ Hom.set (a', a) \ \ (D.cones_map (\ (a', a) x'') \) = \ (D.cones_map (\ (a', a) x) \)" hence "\\ (a', a) x'' : a' \ a\ \ \ (D.cones_map (\ (a', a) x'') \) = \ (D.cones_map (\ (a', a) x) \)" using ide_apex a' Hom.set_def Hom.\_mapsto [of a' a] by auto hence "\ (a', a) (\ (a', a) x'') = \ (a', a) f''" using 1 f'' by auto thus "x'' = \ (a', a) f''" using ide_apex a' x'' Hom.\_\ by simp qed ultimately show ?thesis using ex1I [of "\x'. x' \ Hom.set (a', a) \ \ (D.cones_map (\ (a', a) x') \) = \ (D.cones_map (\ (a', a) x) \)" "\ (a', a) f''"] by simp qed thus "x = x'" using x x' xx' by auto qed hence "inj_on ?F (Hom.set (a', a))" using inj_onI [of "Hom.set (a', a)" ?F] by auto moreover have "?F ` Hom.set (a', a) = \ ` D.cones a'" proof show "?F ` Hom.set (a', a) \ \ ` D.cones a'" proof fix X' assume X': "X' \ ?F ` Hom.set (a', a)" from this obtain x' where x': "x' \ Hom.set (a', a) \ ?F x' = X'" by blast show "X' \ \ ` D.cones a'" proof - have "X' = \ (D.cones_map (\ (a', a) x') \)" using x' by blast hence "X' = \ (D.cones_map (\ (a', a) x') \)" using x' by force moreover have "\\ (a', a) x' : a' \ a\" using ide_apex a' x' Hom.set_def Hom.\_\ by auto ultimately show ?thesis using x' cone_\ D.cones_map_mapsto by force qed qed show "\ ` D.cones a' \ ?F ` Hom.set (a', a)" proof fix X' assume X': "X' \ \ ` D.cones a'" hence "\ X' \ \ ` \ ` D.cones a'" by simp - with Sr.\_\ have "\ X' \ D.cones a'" + with Sr.DN_UP have "\ X' \ D.cones a'" by auto hence "\!f. \f : a' \ a\ \ D.cones_map f \ = \ X'" using a' is_universal by simp from this obtain f where "\f : a' \ a\ \ D.cones_map f \ = \ X'" by auto hence f: "\f : a' \ a\ \ \ (D.cones_map f \) = X'" - using X' Sr.\_\ by auto + using X' Sr.UP_DN by auto have "X' = ?F (\ (a', a) f)" using f Hom.\_\ by presburger thus "X' \ ?F ` Hom.set (a', a)" using f Hom.set_def by force qed qed ultimately show ?thesis using bij_betw_def [of ?F "Hom.set (a', a)" "\ ` D.cones a'"] inj_on_def by auto qed let ?f = "S.mkArr (Hom.set (a', a)) (\ ` D.cones a') ?F" have iso: "S.iso ?f" proof - have "?F \ Hom.set (a', a) \ \ ` D.cones a'" using bij bij_betw_imp_funcset by fast hence 1: "S.arr ?f" - using ide_apex a' Hom.set_subset_Univ S.\_mapsto S.arr_mkArr by auto - thus ?thesis using bij S.iso_char S.arr_mkArr S.set_mkIde by fastforce + using ide_apex a' Hom.set_subset_Univ S.UP_mapsto by auto + thus ?thesis using bij S.iso_char S.set_mkIde by fastforce qed moreover have "?f = \.map a'" using a' \o_def by force finally show ?thesis by auto qed qed - interpretation R: representation_of_functor C S.comp \ Cones.map a \.map .. + interpretation R: representation_of_functor C S.comp S.setp \ Cones.map a \.map .. lemma \_in_terms_of_\: shows "\ = \ (\.FUN a (\ (a, a) a))" proof - have "\.FUN a (\ (a, a) a) = (\x \ Hom.set (a, a). \ (D.cones_map (\ (a, a) x) \)) (\ (a, a) a)" using ide_apex S.Fun_mkArr \.map_simp_ide \o_def \.preserves_reflects_arr [of a] by simp also have "... = \ (D.cones_map a \)" proof - have "(\x \ Hom.set (a, a). \ (D.cones_map (\ (a, a) x) \)) (\ (a, a) a) = \ (D.cones_map (\ (a, a) (\ (a, a) a)) \)" proof - have "\ (a, a) a \ Hom.set (a, a)" using ide_apex Hom.\_mapsto by fastforce thus ?thesis using restrict_apply' [of "\ (a, a) a" "Hom.set (a, a)"] by blast qed also have "... = \ (D.cones_map a \)" proof - have "\ (a, a) (\ (a, a) a) = a" using ide_apex Hom.\_\ [of a a a] by fastforce thus ?thesis by metis qed finally show ?thesis by auto qed finally have "\.FUN a (\ (a, a) a) = \ (D.cones_map a \)" by auto also have "... = \ \" using ide_apex D.cones_map_ide [of \ a] cone_\ by simp finally have "\.FUN a (\ (a, a) a) = \ \" by blast hence "\ (\.FUN a (\ (a, a) a)) = \ (\ \)" by simp - thus ?thesis using cone_\ Sr.\_\ by simp + thus ?thesis using cone_\ Sr.DN_UP by simp qed abbreviation Hom where "Hom \ Hom.map" abbreviation \ where "\ \ \.map" lemma induces_limit_situation: shows "limit_situation J C D S.comp \ \ a \ \" using \_in_terms_of_\ \o_def by unfold_locales auto no_notation S.comp (infixr "\\<^sub>S" 55) end sublocale limit_cone \ limit_situation J C D replete_setcat.comp \ \ a \ \ using induces_limit_situation by auto subsection "Representations of the Cones Functor Induce Limit Situations" context representation_of_cones_functor begin - interpretation \: set_valued_transformation Cop.comp S \Y a\ Cones.map \ .. + interpretation \: set_valued_transformation Cop.comp S S.setp \Y a\ Cones.map \ .. interpretation \: inverse_transformation Cop.comp S \Y a\ Cones.map \ .. - interpretation \: set_valued_transformation Cop.comp S Cones.map \Y a\ \.map .. + interpretation \: set_valued_transformation Cop.comp S S.setp + Cones.map \Y a\ \.map .. abbreviation \ where "\ \ Cones.\" abbreviation \ where "\ \ \ (S.Fun (\ a) (\ (a, a) a))" lemma Cones_SET_eq_\_img_cones: assumes "C.ide a'" shows "Cones.SET a' = \ ` D.cones a'" proof - - have "\ ` D.cones a' \ S.Univ" using S.\_mapsto by auto + have "\ ` D.cones a' \ S.Univ" using S.UP_mapsto by auto thus ?thesis using assms Cones.map_ide S.set_mkIde by auto qed lemma \\: shows "\ \ = S.Fun (\ a) (\ (a, a) a)" proof - have "S.Fun (\ a) (\ (a, a) a) \ Cones.SET a" using Ya.ide_a Hom.\_mapsto S.Fun_mapsto [of "\ a"] Hom.set_map by fastforce thus ?thesis using Ya.ide_a Cones_SET_eq_\_img_cones by auto qed interpretation \: cone J C D a \ proof - have "\ \ \ \ ` D.cones a" using Ya.ide_a \\ S.Fun_mapsto [of "\ a"] Hom.\_mapsto Hom.set_map Cones_SET_eq_\_img_cones by fastforce thus "D.cone a \" - by (metis (no_types, lifting) S.\_\ UNIV_I f_inv_into_f inv_into_into mem_Collect_eq) + by (metis (no_types, lifting) S.DN_UP UNIV_I f_inv_into_f inv_into_into mem_Collect_eq) qed lemma cone_\: shows "D.cone a \" .. lemma \_FUN_simp: assumes a': "C.ide a'" and x: "x \ Hom.set (a', a)" shows "\.FUN a' x = Cones.FUN (\ (a', a) x) (\ \)" proof - have \x: "\\ (a', a) x : a' \ a\" using Ya.ide_a a' x Hom.\_mapsto by blast have \a: "\ (a, a) a \ Hom.set (a, a)" using Ya.ide_a Hom.\_mapsto by fastforce have "\.FUN a' x = (\.FUN a' o Ya.FUN (\ (a', a) x)) (\ (a, a) a)" proof - have "\ (a', a) (a \ \ (a', a) x) = x" using Ya.ide_a a' x \x Hom.\_\ C.comp_cod_arr by fastforce moreover have "S.arr (S.mkArr (Hom.set (a, a)) (Hom.set (a', a)) (\ (a', a) \ Cop.comp (\ (a', a) x) \ \ (a, a)))" - by (metis (no_types) Cop.hom_char Ya.Y_ide_arr(2) Ya.preserves_reflects_arr - \.ide_apex \x Cop.in_homE) + by (metis C.arrI Cop.arr_char Ya.Y_ide_arr(2) Ya.preserves_arr \.ide_apex \x) ultimately show ?thesis using Ya.ide_a a' x Ya.Y_ide_arr \x \a C.ide_in_hom by auto qed also have "... = (Cones.FUN (\ (a', a) x) o \.FUN a) (\ (a, a) a)" proof - have "(\.FUN a' o Ya.FUN (\ (a', a) x)) (\ (a, a) a) = S.Fun (\ a' \\<^sub>S Y a (\ (a', a) x)) (\ (a, a) a)" using \x a' \a Ya.ide_a Ya.map_simp Hom.set_map by (elim C.in_homE, auto) also have "... = S.Fun (S (Cones.map (\ (a', a) x)) (\ a)) (\ (a, a) a)" using \x is_natural_1 [of "\ (a', a) x"] is_natural_2 [of "\ (a', a) x"] by auto also have "... = (Cones.FUN (\ (a', a) x) o \.FUN a) (\ (a, a) a)" proof - have "S.seq (Cones.map (\ (a', a) x)) (\ a)" using Ya.ide_a \x Cones.map_preserves_dom [of "\ (a', a) x"] apply (intro S.seqI) apply auto[2] by fastforce thus ?thesis using Ya.ide_a \a Hom.set_map by auto qed finally show ?thesis by simp qed also have "... = Cones.FUN (\ (a', a) x) (\ \)" using \\ by simp finally show ?thesis by auto qed lemma \_is_universal: assumes "D.cone a' \'" shows "\\ (a', a) (\.FUN a' (\ \')) : a' \ a\" and "D.cones_map (\ (a', a) (\.FUN a' (\ \'))) \ = \'" and "\ \f' : a' \ a\; D.cones_map f' \ = \' \ \ f' = \ (a', a) (\.FUN a' (\ \'))" proof - interpret \': cone J C D a' \' using assms by auto have a': "C.ide a'" using \'.ide_apex by simp have \\': "\ \' \ Cones.SET a'" using assms a' Cones_SET_eq_\_img_cones by auto let ?f = "\ (a', a) (\.FUN a' (\ \'))" have A: "\.FUN a' (\ \') \ Hom.set (a', a)" proof - have "\.FUN a' \ Cones.SET a' \ Ya.SET a'" using a' \.preserves_hom [of a' a' a'] S.Fun_mapsto [of "\.map a'"] by fastforce thus ?thesis using a' \\' Ya.ide_a Hom.set_map by auto qed show f: "\?f : a' \ a\" using A a' Ya.ide_a Hom.\_mapsto [of a' a] by auto have E: "\f. \f : a' \ a\ \ Cones.FUN f (\ \) = \.FUN a' (\ (a', a) f)" proof - fix f assume f: "\f : a' \ a\" have "\ (a', a) f \ Hom.set (a', a)" using a' Ya.ide_a f Hom.\_mapsto by auto thus "Cones.FUN f (\ \) = \.FUN a' (\ (a', a) f)" using a' f \_FUN_simp by simp qed have I: "\.FUN a' (\.FUN a' (\ \')) = \ \'" proof - have "\.FUN a' (\.FUN a' (\ \')) = compose (\.DOM a') (\.FUN a') (\.FUN a') (\ \')" using a' \\' Cones.map_ide \.preserves_hom [of a' a' a'] by force also have "... = (\x \ \.DOM a'. x) (\ \')" using a' \.inverts_components S.inverse_arrows_char by force also have "... = \ \'" using a' \\' Cones.map_ide \.preserves_hom [of a' a' a'] by force finally show ?thesis by auto qed show f\: "D.cones_map ?f \ = \'" proof - have "D.cones_map ?f \ = (\ o Cones.FUN ?f o \) \" using f Cones.preserves_arr [of ?f] cone_\ by (cases "D.cone a \", auto) also have "... = \'" using f Ya.ide_a a' A E I by auto finally show ?thesis by auto qed show "\ \f' : a' \ a\; D.cones_map f' \ = \' \ \ f' = ?f" proof - assume f': "\f' : a' \ a\" and f'\: "D.cones_map f' \ = \'" show "f' = ?f" proof - have 1: "\ (a', a) f' \ Hom.set (a', a) \ \ (a', a) ?f \ Hom.set (a', a)" using Ya.ide_a a' f f' Hom.\_mapsto by auto have "S.iso (\ a')" using \'.ide_apex components_are_iso by auto hence 2: "S.arr (\ a') \ bij_betw (\.FUN a') (Hom.set (a', a)) (Cones.SET a')" using Ya.ide_a a' S.iso_char Hom.set_map by auto have "\.FUN a' (\ (a', a) f') = \.FUN a' (\ (a', a) ?f)" proof - have "\.FUN a' (\ (a', a) ?f) = \ \'" using A I Hom.\_\ Ya.ide_a a' by simp also have "... = Cones.FUN f' (\ \)" using f f' A E cone_\ Cones.preserves_arr f\ f'\ by (elim C.in_homE, auto) also have "... = \.FUN a' (\ (a', a) f')" using f' E by simp finally show ?thesis by argo qed moreover have "inj_on (\.FUN a') (Hom.set (a', a))" using 2 bij_betw_imp_inj_on by blast ultimately have 3: "\ (a', a) f' = \ (a', a) ?f" using 1 inj_on_def [of "\.FUN a'" "Hom.set (a', a)"] by blast show ?thesis proof - have "f' = \ (a', a) (\ (a', a) f')" using Ya.ide_a a' f' Hom.\_\ by simp also have "... = \ (a', a) (\.FUN a' (\ \'))" using Ya.ide_a a' Hom.\_\ A 3 by simp finally show ?thesis by blast qed qed qed qed interpretation \: limit_cone J C D a \ proof show "\a' \'. D.cone a' \' \ \!f. \f : a' \ a\ \ D.cones_map f \ = \'" proof - fix a' \' assume 1: "D.cone a' \'" show "\!f. \f : a' \ a\ \ D.cones_map f \ = \'" proof show "\\ (a', a) (\.FUN a' (\ \')) : a' \ a\ \ D.cones_map (\ (a', a) (\.FUN a' (\ \'))) \ = \'" using 1 \_is_universal by blast show "\f. \f : a' \ a\ \ D.cones_map f \ = \' \ f = \ (a', a) (\.FUN a' (\ \'))" using 1 \_is_universal by blast qed qed qed lemma \_is_limit_cone: shows "D.limit_cone a \" .. lemma induces_limit_situation: shows "limit_situation J C D S \ \ a \ \" proof show "\ = \" by simp fix a' assume a': "Cop.ide a'" let ?F = "\x. \ (D.cones_map (\ (a', a) x) \)" show "\ a' = S.mkArr (Hom.set (a', a)) (\ ` D.cones a') ?F" proof - have 1: "\\ a' : S.mkIde (Hom.set (a', a)) \\<^sub>S S.mkIde (\ ` D.cones a')\" using a' Cones.map_ide Ya.ide_a by auto moreover have "\.DOM a' = Hom.set (a', a)" using 1 Hom.set_subset_Univ a' Ya.ide_a Hom.set_map by simp moreover have "\.COD a' = \ ` D.cones a'" using a' Cones_SET_eq_\_img_cones by fastforce ultimately have 2: "\ a' = S.mkArr (Hom.set (a', a)) (\ ` D.cones a') (\.FUN a')" using S.mkArr_Fun [of "\ a'"] by fastforce also have "... = S.mkArr (Hom.set (a', a)) (\ ` D.cones a') ?F" proof show "S.arr (S.mkArr (Hom.set (a', a)) (\ ` D.cones a') (\.FUN a'))" using 1 2 by auto show "\x. x \ Hom.set (a', a) \ \.FUN a' x = ?F x" proof - fix x assume x: "x \ Hom.set (a', a)" hence \x: "\\ (a', a) x : a' \ a\" using a' Ya.ide_a Hom.\_mapsto by auto show "\.FUN a' x = ?F x" proof - have "\.FUN a' x = Cones.FUN (\ (a', a) x) (\ \)" using a' x \_FUN_simp by simp also have "... = restrict (\ o D.cones_map (\ (a', a) x) o \) (\ ` D.cones a) (\ \)" using \x Cones.map_simp Cones.preserves_arr [of "\ (a', a) x"] S.Fun_mkArr by (elim C.in_homE, auto) also have "... = ?F x" using cone_\ by simp ultimately show ?thesis by simp qed qed qed finally show "\ a' = S.mkArr (Hom.set (a', a)) (\ ` D.cones a') ?F" by auto qed qed end sublocale representation_of_cones_functor \ limit_situation J C D S \ \ a \ \ using induces_limit_situation by auto section "Categories with Limits" context category begin text\ A category @{term[source=true] C} has limits of shape @{term J} if every diagram of shape @{term J} admits a limit cone. \ definition has_limits_of_shape where "has_limits_of_shape J \ \D. diagram J C D \ (\a \. limit_cone J C D a \)" text\ A category has limits at a type @{typ 'j} if it has limits of shape @{term J} for every category @{term J} whose arrows are of type @{typ 'j}. \ definition has_limits where "has_limits (_ :: 'j) \ \J :: 'j comp. category J \ has_limits_of_shape J" text\ Whether a category has limits of shape \J\ truly depends only on the ``shape'' (\emph{i.e.}~isomorphism class) of \J\ and not on details of its construction. \ lemma has_limits_preserved_by_isomorphism: assumes "has_limits_of_shape J" and "isomorphic_categories J J'" shows "has_limits_of_shape J'" proof - interpret J: category J using assms(2) isomorphic_categories_def isomorphic_categories_axioms_def by auto interpret J': category J' using assms(2) isomorphic_categories_def isomorphic_categories_axioms_def by auto from assms(2) obtain \ \ where IF: "inverse_functors J' J \ \" using isomorphic_categories_def isomorphic_categories_axioms_def inverse_functors_sym by blast interpret IF: inverse_functors J' J \ \ using IF by auto have \\: "\ o \ = J.map" using IF.inv by metis have \\: "\ o \ = J'.map" using IF.inv' by metis have "\D'. diagram J' C D' \ \a \. limit_cone J' C D' a \" proof - fix D' assume D': "diagram J' C D'" interpret D': diagram J' C D' using D' by auto interpret D: composite_functor J J' C \ D' .. interpret D: diagram J C \D' o \\ .. have D: "diagram J C (D' o \)" .. from assms(1) obtain a \ where \: "D.limit_cone a \" using D has_limits_of_shape_def by blast interpret \: limit_cone J C \D' o \\ a \ using \ by auto interpret A': constant_functor J' C a using \.ide_apex by (unfold_locales, auto) have \o\: "cone J' C (D' o \ o \) a (\ o \)" using comp_cone_functor IF.G.functor_axioms \.cone_axioms by fastforce hence \o\: "cone J' C D' a (\ o \)" using \\ by (metis D'.functor_axioms Fun.comp_assoc comp_functor_identity) interpret \o\: cone J' C D' a \\ o \\ using \o\ by auto interpret \o\: limit_cone J' C D' a \\ o \\ proof fix a' \' assume \': "D'.cone a' \'" interpret \': cone J' C D' a' \' using \' by auto have \'o\: "cone J C (D' o \) a' (\' o \)" using \' comp_cone_functor IF.F.functor_axioms by fastforce interpret \'o\: cone J C \D' o \\ a' \\' o \\ using \'o\ by auto have "cone J C (D' o \) a' (\' o \)" .. hence 1: "\!f. \f : a' \ a\ \ D.cones_map f \ = \' o \" using \.is_universal by simp show "\!f. \f : a' \ a\ \ D'.cones_map f (\ o \) = \'" proof let ?f = "THE f. \f : a' \ a\ \ D.cones_map f \ = \' o \" have f: "\?f : a' \ a\ \ D.cones_map ?f \ = \' o \" using 1 theI' [of "\f. \f : a' \ a\ \ D.cones_map f \ = \' o \"] by blast have f_in_hom: "\?f : a' \ a\" using f by blast have "D'.cones_map ?f (\ o \) = \'" proof fix j' have "\J'.arr j' \ D'.cones_map ?f (\ o \) j' = \' j'" proof - assume j': "\J'.arr j'" have "D'.cones_map ?f (\ o \) j' = null" using j' f_in_hom \o\ by fastforce thus ?thesis using j' \'.is_extensional by simp qed moreover have "J'.arr j' \ D'.cones_map ?f (\ o \) j' = \' j'" proof - assume j': "J'.arr j'" have "D'.cones_map ?f (\ o \) j' = \ (\ j') \ ?f" using j' f \o\ by fastforce also have "... = D.cones_map ?f \ (\ j')" using j' f_in_hom \ \.cone_\ by fastforce also have "... = \' j'" using j' f \ \\ Fun.comp_def J'.map_simp by metis finally show "D'.cones_map ?f (\ o \) j' = \' j'" by auto qed ultimately show "D'.cones_map ?f (\ o \) j' = \' j'" by blast qed thus "\?f : a' \ a\ \ D'.cones_map ?f (\ o \) = \'" using f by auto fix f' assume f': "\f' : a' \ a\ \ D'.cones_map f' (\ o \) = \'" have "D.cones_map f' \ = \' o \" proof fix j have "\J.arr j \ D.cones_map f' \ j = (\' o \) j" using f' \ \'o\.is_extensional \.cone_\ mem_Collect_eq restrict_apply by auto moreover have "J.arr j \ D.cones_map f' \ j = (\' o \) j" proof - assume j: "J.arr j" have "D.cones_map f' \ j = C (\ j) f'" using j f' \.cone_\ by auto also have "... = C ((\ o \) (\ j)) f'" using j f' \\ by (metis comp_apply J.map_simp) also have "... = D'.cones_map f' (\ o \) (\ j)" using j f' \o\ by fastforce also have "... = (\' o \) j" using j f' by auto finally show "D.cones_map f' \ j = (\' o \) j" by auto qed ultimately show "D.cones_map f' \ j = (\' o \) j" by blast qed hence "\f' : a' \ a\ \ D.cones_map f' \ = \' o \" using f' by auto moreover have "\P x x'. (\!x. P x) \ P x \ P x' \ x = x'" by auto ultimately show "f' = ?f" using 1 f by blast qed qed have "limit_cone J' C D' a (\ o \)" .. thus "\a \. limit_cone J' C D' a \" by blast qed thus ?thesis using has_limits_of_shape_def by auto qed end subsection "Diagonal Functors" text\ The existence of limits can also be expressed in terms of adjunctions: a category @{term C} has limits of shape @{term J} if the diagonal functor taking each object @{term a} in @{term C} to the constant-@{term a} diagram and each arrow \f \ C.hom a a'\ to the constant-@{term f} natural transformation between diagrams is a left adjoint functor. \ locale diagonal_functor = C: category C + J: category J + J_C: functor_category J C for J :: "'j comp" (infixr "\\<^sub>J" 55) and C :: "'c comp" (infixr "\" 55) begin notation J.in_hom ("\_ : _ \\<^sub>J _\") notation J_C.comp (infixr "\\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>]" 55) notation J_C.in_hom ("\_ : _ \\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] _\") definition map :: "'c \ ('j, 'c) J_C.arr" where "map f = (if C.arr f then J_C.MkArr (constant_functor.map J C (C.dom f)) (constant_functor.map J C (C.cod f)) (constant_transformation.map J C f) else J_C.null)" lemma is_functor: shows "functor C J_C.comp map" proof fix f show "\ C.arr f \ local.map f = J_C.null" using map_def by simp assume f: "C.arr f" interpret Dom_f: constant_functor J C \C.dom f\ using f by (unfold_locales, auto) interpret Cod_f: constant_functor J C \C.cod f\ using f by (unfold_locales, auto) interpret Fun_f: constant_transformation J C f using f by (unfold_locales, auto) show 1: "J_C.arr (map f)" using f map_def by (simp add: Fun_f.natural_transformation_axioms) show "J_C.dom (map f) = map (C.dom f)" proof - have "constant_transformation J C (C.dom f)" using f by unfold_locales auto hence "constant_transformation.map J C (C.dom f) = Dom_f.map" using Dom_f.map_def constant_transformation.map_def [of J C "C.dom f"] by auto thus ?thesis using f 1 by (simp add: map_def J_C.dom_char) qed show "J_C.cod (map f) = map (C.cod f)" proof - have "constant_transformation J C (C.cod f)" using f by unfold_locales auto hence "constant_transformation.map J C (C.cod f) = Cod_f.map" using Cod_f.map_def constant_transformation.map_def [of J C "C.cod f"] by auto thus ?thesis using f 1 by (simp add: map_def J_C.cod_char) qed next fix f g assume g: "C.seq g f" have f: "C.arr f" using g by auto interpret Dom_f: constant_functor J C \C.dom f\ using f by unfold_locales auto interpret Cod_f: constant_functor J C \C.cod f\ using f by unfold_locales auto interpret Fun_f: constant_transformation J C f using f by unfold_locales auto interpret Cod_g: constant_functor J C \C.cod g\ using g by unfold_locales auto interpret Fun_g: constant_transformation J C g using g by unfold_locales auto interpret Fun_g: natural_transformation J C Cod_f.map Cod_g.map Fun_g.map apply unfold_locales using f g C.seqE [of g f] C.comp_arr_dom C.comp_cod_arr Fun_g.is_extensional by auto interpret Fun_fg: vertical_composite J C Dom_f.map Cod_f.map Cod_g.map Fun_f.map Fun_g.map .. have 1: "J_C.arr (map f)" using f map_def by (simp add: Fun_f.natural_transformation_axioms) show "map (g \ f) = map g \\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map f" proof - have "map (C g f) = J_C.MkArr Dom_f.map Cod_g.map (constant_transformation.map J C (C g f))" using f g map_def by simp also have "... = J_C.MkArr Dom_f.map Cod_g.map (\j. if J.arr j then C g f else C.null)" proof - have "constant_transformation J C (g \ f)" using g by unfold_locales auto thus ?thesis using constant_transformation.map_def by metis qed also have "... = J_C.comp (J_C.MkArr Cod_f.map Cod_g.map Fun_g.map) (J_C.MkArr Dom_f.map Cod_f.map Fun_f.map)" proof - have "J_C.MkArr Cod_f.map Cod_g.map Fun_g.map \\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] J_C.MkArr Dom_f.map Cod_f.map Fun_f.map = J_C.MkArr Dom_f.map Cod_g.map Fun_fg.map" using J_C.comp_char J_C.comp_MkArr Fun_f.natural_transformation_axioms Fun_g.natural_transformation_axioms by blast also have "... = J_C.MkArr Dom_f.map Cod_g.map (\j. if J.arr j then g \ f else C.null)" using Fun_fg.is_extensional Fun_fg.map_simp_2 by auto finally show ?thesis by auto qed also have "... = map g \\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map f" using f g map_def by fastforce finally show ?thesis by auto qed qed sublocale "functor" C J_C.comp map using is_functor by auto text\ The objects of \[J, C]\ correspond bijectively to diagrams of shape @{term J} in @{term C}. \ lemma ide_determines_diagram: assumes "J_C.ide d" shows "diagram J C (J_C.Map d)" and "J_C.MkIde (J_C.Map d) = d" proof - interpret \: natural_transformation J C \J_C.Map d\ \J_C.Map d\ \J_C.Map d\ using assms J_C.ide_char J_C.arr_MkArr by fastforce interpret D: "functor" J C \J_C.Map d\ .. show "diagram J C (J_C.Map d)" .. show "J_C.MkIde (J_C.Map d) = d" using assms J_C.ide_char by (metis J_C.ideD(1) J_C.MkArr_Map) qed lemma diagram_determines_ide: assumes "diagram J C D" shows "J_C.ide (J_C.MkIde D)" and "J_C.Map (J_C.MkIde D) = D" proof - interpret D: diagram J C D using assms by auto show "J_C.ide (J_C.MkIde D)" using J_C.ide_char using D.functor_axioms J_C.ide_MkIde by auto thus "J_C.Map (J_C.MkIde D) = D" using J_C.in_homE by simp qed lemma bij_betw_ide_diagram: shows "bij_betw J_C.Map (Collect J_C.ide) (Collect (diagram J C))" proof (intro bij_betwI) show "J_C.Map \ Collect J_C.ide \ Collect (diagram J C)" using ide_determines_diagram by blast show "J_C.MkIde \ Collect (diagram J C) \ Collect J_C.ide" using diagram_determines_ide by blast show "\d. d \ Collect J_C.ide \ J_C.MkIde (J_C.Map d) = d" using ide_determines_diagram by blast show "\D. D \ Collect (diagram J C) \ J_C.Map (J_C.MkIde D) = D" using diagram_determines_ide by blast qed text\ Arrows from from the diagonal functor correspond bijectively to cones. \ lemma arrow_determines_cone: assumes "J_C.ide d" and "arrow_from_functor C J_C.comp map a d x" shows "cone J C (J_C.Map d) a (J_C.Map x)" and "J_C.MkArr (constant_functor.map J C a) (J_C.Map d) (J_C.Map x) = x" proof - interpret D: diagram J C \J_C.Map d\ using assms ide_determines_diagram by auto interpret x: arrow_from_functor C J_C.comp map a d x using assms by auto interpret A: constant_functor J C a using x.arrow by (unfold_locales, auto) interpret \: constant_transformation J C a using x.arrow by (unfold_locales, auto) have Dom_x: "J_C.Dom x = A.map" using J_C.in_hom_char map_def x.arrow by force have Cod_x: "J_C.Cod x = J_C.Map d" using x.arrow by auto interpret \: natural_transformation J C A.map \J_C.Map d\ \J_C.Map x\ using x.arrow J_C.arr_char [of x] Dom_x Cod_x by force show "D.cone a (J_C.Map x)" .. show "J_C.MkArr A.map (J_C.Map d) (J_C.Map x) = x" using x.arrow Dom_x Cod_x \.natural_transformation_axioms by (intro J_C.arr_eqI, auto) qed lemma cone_determines_arrow: assumes "J_C.ide d" and "cone J C (J_C.Map d) a \" shows "arrow_from_functor C J_C.comp map a d (J_C.MkArr (constant_functor.map J C a) (J_C.Map d) \)" and "J_C.Map (J_C.MkArr (constant_functor.map J C a) (J_C.Map d) \) = \" proof - interpret \: cone J C \J_C.Map d\ a \ using assms(2) by auto let ?x = "J_C.MkArr \.A.map (J_C.Map d) \" interpret x: arrow_from_functor C J_C.comp map a d ?x proof have "\J_C.MkArr \.A.map (J_C.Map d) \ : J_C.MkIde \.A.map \\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] J_C.MkIde (J_C.Map d)\" using \.natural_transformation_axioms by auto moreover have "J_C.MkIde \.A.map = map a" using \.A.value_is_ide map_def \.A.map_def C.ide_char by (metis (no_types, lifting) J_C.dom_MkArr preserves_arr preserves_dom) moreover have "J_C.MkIde (J_C.Map d) = d" using assms ide_determines_diagram(2) by simp ultimately show "C.ide a \ \J_C.MkArr \.A.map (J_C.Map d) \ : map a \\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] d\" using \.A.value_is_ide by simp qed show "arrow_from_functor C J_C.comp map a d ?x" .. show "J_C.Map (J_C.MkArr (constant_functor.map J C a) (J_C.Map d) \) = \" by (simp add: \.natural_transformation_axioms) qed text\ Transforming a cone by composing at the apex with an arrow @{term g} corresponds, via the preceding bijections, to composition in \[J, C]\ with the image of @{term g} under the diagonal functor. \ lemma cones_map_is_composition: assumes "\g : a' \ a\" and "cone J C D a \" shows "J_C.MkArr (constant_functor.map J C a') D (diagram.cones_map J C D g \) = J_C.MkArr (constant_functor.map J C a) D \ \\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map g" proof - interpret A: constant_transformation J C a using assms(1) by (unfold_locales, auto) interpret \: cone J C D a \ using assms(2) by auto have cone_\: "cone J C D a \" .. interpret A': constant_transformation J C a' using assms(1) by (unfold_locales, auto) let ?\' = "\.D.cones_map g \" interpret \': cone J C D a' ?\' using assms(1) cone_\ \.D.cones_map_mapsto by blast let ?x = "J_C.MkArr \.A.map D \" let ?x' = "J_C.MkArr \'.A.map D ?\'" show "?x' = J_C.comp ?x (map g)" proof (intro J_C.arr_eqI) have x: "J_C.arr ?x" using \.natural_transformation_axioms J_C.arr_char [of ?x] by simp show x': "J_C.arr ?x'" using \'.natural_transformation_axioms J_C.arr_char [of ?x'] by simp have 3: "\?x : map a \\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] J_C.MkIde D\" using \.D.diagram_axioms arrow_from_functor.arrow cone_\ cone_determines_arrow(1) diagram_determines_ide(1) by fastforce have 4: "\?x' : map a' \\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] J_C.MkIde D\" by (metis (no_types, lifting) J_C.Dom.simps(1) J_C.Dom_cod J_C.Map_cod J_C.cod_MkArr \'.cone_axioms arrow_from_functor.arrow category.ide_cod cone_determines_arrow(1) functor_def is_functor x) have seq_xg: "J_C.seq ?x (map g)" using assms(1) 3 preserves_hom [of g] by (intro J_C.seqI', auto) show 2: "J_C.seq ?x (map g)" using seq_xg J_C.seqI' by blast show "J_C.Dom ?x' = J_C.Dom (?x \\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map g)" proof - have "J_C.Dom ?x' = J_C.Dom (J_C.dom ?x')" using x' J_C.Dom_dom by simp also have "... = J_C.Dom (map a')" using 4 by force also have "... = J_C.Dom (J_C.dom (?x \\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map g))" using assms(1) 2 by auto also have "... = J_C.Dom (?x \\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map g)" using seq_xg J_C.Dom_dom J_C.seqI' by blast finally show ?thesis by auto qed show "J_C.Cod ?x' = J_C.Cod (?x \\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map g)" proof - have "J_C.Cod ?x' = J_C.Cod (J_C.cod ?x')" using x' J_C.Cod_cod by simp also have "... = J_C.Cod (J_C.MkIde D)" using 4 by force also have "... = J_C.Cod (J_C.cod (?x \\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map g))" using 2 3 J_C.cod_comp J_C.in_homE by metis also have "... = J_C.Cod (?x \\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map g)" using seq_xg J_C.Cod_cod J_C.seqI' by blast finally show ?thesis by auto qed show "J_C.Map ?x' = J_C.Map (?x \\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map g)" proof - interpret g: constant_transformation J C g apply unfold_locales using assms(1) by auto interpret \og: vertical_composite J C A'.map \.A.map D g.map \ using assms(1) C.comp_arr_dom C.comp_cod_arr A'.is_extensional g.is_extensional apply (unfold_locales, auto) by (elim J.seqE, auto) have "J_C.Map (?x \\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map g) = \og.map" using assms(1) 2 J_C.comp_char map_def by auto also have "... = J_C.Map ?x'" using x' \og.map_def J_C.arr_char [of ?x'] natural_transformation.is_extensional assms(1) cone_\ \og.map_simp_2 by fastforce finally show ?thesis by auto qed qed qed text\ Coextension along an arrow from a functor is equivalent to a transformation of cones. \ lemma coextension_iff_cones_map: assumes x: "arrow_from_functor C J_C.comp map a d x" and g: "\g : a' \ a\" and x': "\x' : map a' \\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] d\" shows "arrow_from_functor.is_coext C J_C.comp map a x a' x' g \ J_C.Map x' = diagram.cones_map J C (J_C.Map d) g (J_C.Map x)" proof - interpret x: arrow_from_functor C J_C.comp map a d x using assms by auto interpret A': constant_functor J C a' using assms(2) by (unfold_locales, auto) have x': "arrow_from_functor C J_C.comp map a' d x'" using A'.value_is_ide assms(3) by (unfold_locales, blast) have d: "J_C.ide d" using J_C.ide_cod x.arrow by blast let ?D = "J_C.Map d" let ?\ = "J_C.Map x" let ?\' = "J_C.Map x'" interpret D: diagram J C ?D using ide_determines_diagram J_C.ide_cod x.arrow by blast interpret \: cone J C ?D a ?\ using assms(1) d arrow_determines_cone by simp interpret \: constant_transformation J C g using g \.ide_apex by (unfold_locales, auto) interpret \og: vertical_composite J C A'.map \.A.map ?D \.map ?\ using g C.comp_arr_dom C.comp_cod_arr \.is_extensional by (unfold_locales, auto) show ?thesis proof assume 0: "x.is_coext a' x' g" show "?\' = D.cones_map g ?\" proof - have 1: "x' = x \\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map g" using 0 x.is_coext_def by blast hence "?\' = J_C.Map x'" using 0 x.is_coext_def by fast moreover have "... = D.cones_map g ?\" proof - have "J_C.MkArr A'.map (J_C.Map d) (D.cones_map g (J_C.Map x)) = x'" proof - have "J_C.MkArr A'.map (J_C.Map d) (D.cones_map g (J_C.Map x)) = x \\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map g" using d g cones_map_is_composition arrow_determines_cone(2) \.cone_axioms x.arrow_from_functor_axioms by auto thus ?thesis by (metis 1) qed moreover have "J_C.arr (J_C.MkArr A'.map (J_C.Map d) (D.cones_map g (J_C.Map x)))" using 1 d g cones_map_is_composition preserves_arr arrow_determines_cone(2) \.cone_axioms x.arrow_from_functor_axioms assms(3) by auto ultimately show ?thesis by auto qed ultimately show ?thesis by blast qed next assume X': "?\' = D.cones_map g ?\" show "x.is_coext a' x' g" proof - have 4: "J_C.seq x (map g)" using g x.arrow mem_Collect_eq preserves_arr preserves_cod by (elim C.in_homE, auto) hence 1: "x \\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map g = J_C.MkArr (J_C.Dom (map g)) (J_C.Cod x) (vertical_composite.map J C (J_C.Map (map g)) ?\)" using J_C.comp_char [of x "map g"] by simp have 2: "vertical_composite.map J C (J_C.Map (map g)) ?\ = \og.map" by (simp add: map_def \.value_is_arr \.natural_transformation_axioms) have 3: "... = D.cones_map g ?\" using g \og.map_simp_2 \.cone_axioms \og.is_extensional by auto have "J_C.MkArr A'.map ?D ?\' = J_C.comp x (map g)" proof - have f1: "A'.map = J_C.Dom (map g)" using \.natural_transformation_axioms map_def g by auto have "J_C.Map d = J_C.Cod x" using x.arrow by auto thus ?thesis using f1 X' 1 2 3 by argo qed moreover have "J_C.MkArr A'.map ?D ?\' = x'" using d x' arrow_determines_cone by blast ultimately show ?thesis using g x.is_coext_def by simp qed qed qed end locale right_adjoint_to_diagonal_functor = C: category C + J: category J + J_C: functor_category J C + \: diagonal_functor J C + "functor" J_C.comp C G + Adj: meta_adjunction J_C.comp C \.map G \ \ for J :: "'j comp" (infixr "\\<^sub>J" 55) and C :: "'c comp" (infixr "\" 55) and G :: "('j, 'c) functor_category.arr \ 'c" and \ :: "'c \ ('j, 'c) functor_category.arr \ 'c" and \ :: "('j, 'c) functor_category.arr \ 'c \ ('j, 'c) functor_category.arr" + assumes adjoint: "adjoint_functors J_C.comp C \.map G" begin - interpretation Adj: adjunction J_C.comp C replete_setcat.comp Adj.\C Adj.\D \.map G + interpretation S: replete_setcat . + interpretation Adj: adjunction J_C.comp C S.comp S.setp Adj.\C Adj.\D \.map G \ \ Adj.\ Adj.\ Adj.\ Adj.\ using Adj.induces_adjunction by simp text\ A right adjoint @{term G} to a diagonal functor maps each object @{term d} of \[J, C]\ (corresponding to a diagram @{term D} of shape @{term J} in @{term C} to an object of @{term C}. This object is the limit object, and the component at @{term d} of the counit of the adjunction determines the limit cone. \ lemma gives_limit_cones: assumes "diagram J C D" shows "limit_cone J C D (G (J_C.MkIde D)) (J_C.Map (Adj.\ (J_C.MkIde D)))" proof - interpret D: diagram J C D using assms by auto let ?d = "J_C.MkIde D" let ?a = "G ?d" let ?x = "Adj.\ ?d" let ?\ = "J_C.Map ?x" have "diagram J C D" .. hence 1: "J_C.ide ?d" using \.diagram_determines_ide by auto hence 2: "J_C.Map (J_C.MkIde D) = D" using assms 1 J_C.in_homE \.diagram_determines_ide(2) by simp interpret x: terminal_arrow_from_functor C J_C.comp \.map ?a ?d ?x apply unfold_locales apply (metis (no_types, lifting) "1" preserves_ide Adj.\_in_terms_of_\ Adj.\o_def Adj.\o_in_hom) by (metis 1 Adj.has_terminal_arrows_from_functor(1) terminal_arrow_from_functor.is_terminal) have 3: "arrow_from_functor C J_C.comp \.map ?a ?d ?x" .. interpret \: cone J C D ?a ?\ using 1 2 3 \.arrow_determines_cone [of ?d] by auto have cone_\: "D.cone ?a ?\" .. interpret \: limit_cone J C D ?a ?\ proof fix a' \' assume cone_\': "D.cone a' \'" interpret \': cone J C D a' \' using cone_\' by auto let ?x' = "J_C.MkArr \'.A.map D \'" interpret x': arrow_from_functor C J_C.comp \.map a' ?d ?x' using 1 2 by (metis \.cone_determines_arrow(1) cone_\') have "arrow_from_functor C J_C.comp \.map a' ?d ?x'" .. hence 4: "\!g. x.is_coext a' ?x' g" using x.is_terminal by simp have 5: "\g. \g : a' \\<^sub>C ?a\ \ x.is_coext a' ?x' g \ D.cones_map g ?\ = \'" using \.coextension_iff_cones_map x'.arrow x.arrow_from_functor_axioms by auto have 6: "\g. x.is_coext a' ?x' g \ \g : a' \\<^sub>C ?a\" using x.is_coext_def by simp show "\!g. \g : a' \\<^sub>C ?a\ \ D.cones_map g ?\ = \'" proof - have "\g. \g : a' \\<^sub>C ?a\ \ D.cones_map g ?\ = \'" using 4 5 6 by meson thus ?thesis using 4 5 6 by blast qed qed show "D.limit_cone ?a ?\" .. qed corollary gives_limits: assumes "diagram J C D" shows "diagram.has_as_limit J C D (G (J_C.MkIde D))" using assms gives_limit_cones by fastforce end lemma (in category) has_limits_iff_left_adjoint_diagonal: assumes "category J" shows "has_limits_of_shape J \ left_adjoint_functor C (functor_category.comp J C) (diagonal_functor.map J C)" proof - interpret J: category J using assms by auto interpret J_C: functor_category J C .. interpret \: diagonal_functor J C .. show ?thesis proof assume A: "left_adjoint_functor C J_C.comp \.map" interpret \: left_adjoint_functor C J_C.comp \.map using A by auto interpret Adj: meta_adjunction J_C.comp C \.map \.G \.\ \.\ using \.induces_meta_adjunction by auto have 1: "adjoint_functors J_C.comp C \.map \.G" using adjoint_functors_def \.induces_meta_adjunction by blast interpret G: right_adjoint_to_diagonal_functor J C \.G \.\ \.\ using 1 by unfold_locales auto show "has_limits_of_shape J" using A G.gives_limits has_limits_of_shape_def by blast next text\ If @{term "has_limits J"}, then every diagram @{term D} from @{term J} to @{term[source=true] C} has a limit cone. This means that, for every object @{term d} of the functor category \[J, C]\, there exists an object @{term a} of @{term C} and a terminal arrow from \\ a\ to @{term d} in \[J, C]\. The terminal arrow is given by the limit cone. \ assume A: "has_limits_of_shape J" show "left_adjoint_functor C J_C.comp \.map" proof fix d assume D: "J_C.ide d" interpret D: diagram J C \J_C.Map d\ using D \.ide_determines_diagram by auto let ?D = "J_C.Map d" have "diagram J C (J_C.Map d)" .. from this obtain a \ where limit: "limit_cone J C ?D a \" using A has_limits_of_shape_def by blast interpret A: constant_functor J C a using limit by (simp add: Limit.cone_def limit_cone_def) interpret \: limit_cone J C ?D a \ using limit by simp have cone_\: "cone J C ?D a \" .. let ?x = "J_C.MkArr A.map ?D \" interpret x: arrow_from_functor C J_C.comp \.map a d ?x using D cone_\ \.cone_determines_arrow by auto have "terminal_arrow_from_functor C J_C.comp \.map a d ?x" proof show "\a' x'. arrow_from_functor C J_C.comp \.map a' d x' \ \!g. x.is_coext a' x' g" proof - fix a' x' assume x': "arrow_from_functor C J_C.comp \.map a' d x'" interpret x': arrow_from_functor C J_C.comp \.map a' d x' using x' by auto interpret A': constant_functor J C a' by (unfold_locales, simp add: x'.arrow) let ?\' = "J_C.Map x'" interpret \': cone J C ?D a' ?\' using D x' \.arrow_determines_cone by auto have cone_\': "cone J C ?D a' ?\'" .. let ?g = "\.induced_arrow a' ?\'" show "\!g. x.is_coext a' x' g" proof show "x.is_coext a' x' ?g" proof (unfold x.is_coext_def) have 1: "\?g : a' \ a\ \ D.cones_map ?g \ = ?\'" using \.induced_arrow_def \.is_universal cone_\' theI' [of "\f. \f : a' \ a\ \ D.cones_map f \ = ?\'"] by presburger hence 2: "x' = ?x \\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] \.map ?g" proof - have "x' = J_C.MkArr A'.map ?D ?\'" using D \.arrow_determines_cone(2) x'.arrow_from_functor_axioms by auto thus ?thesis using 1 cone_\ \.cones_map_is_composition [of ?g a' a ?D \] by simp qed show "\?g : a' \ a\ \ x' = ?x \\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] \.map ?g" using 1 2 by auto qed next fix g assume X: "x.is_coext a' x' g" show "g = ?g" proof - have "\g : a' \ a\ \ D.cones_map g \ = ?\'" proof show G: "\g : a' \ a\" using X x.is_coext_def by blast show "D.cones_map g \ = ?\'" proof - have "?\' = J_C.Map (?x \\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] \.map g)" using X x.is_coext_def [of a' x' g] by fast also have "... = D.cones_map g \" proof - interpret map_g: constant_transformation J C g using G by (unfold_locales, auto) interpret \': vertical_composite J C map_g.F.map A.map \\.\.Ya.Cop_S.Map d\ map_g.map \ proof (intro_locales) have "map_g.G.map = A.map" using G by blast thus "natural_transformation_axioms J (\) map_g.F.map A.map map_g.map" using map_g.natural_transformation_axioms by (simp add: natural_transformation_def) qed have "J_C.Map (?x \\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] \.map g) = vertical_composite.map J C map_g.map \" proof - have "J_C.seq ?x (\.map g)" using G x.arrow by auto thus ?thesis using G \.map_def J_C.Map_comp' [of ?x "\.map g"] by auto qed also have "... = D.cones_map g \" using G cone_\ \'.map_def map_g.map_def \.is_natural_2 \'.map_simp_2 by auto finally show ?thesis by blast qed finally show ?thesis by auto qed qed thus ?thesis using cone_\' \.is_universal \.induced_arrow_def theI_unique [of "\g. \g : a' \ a\ \ D.cones_map g \ = ?\'" g] by presburger qed qed qed qed thus "\a x. terminal_arrow_from_functor C J_C.comp \.map a d x" by auto qed qed qed section "Right Adjoint Functors Preserve Limits" context right_adjoint_functor begin lemma preserves_limits: fixes J :: "'j comp" assumes "diagram J C E" and "diagram.has_as_limit J C E a" shows "diagram.has_as_limit J D (G o E) (G a)" proof - text\ From the assumption that @{term E} has a limit, obtain a limit cone @{term \}. \ interpret J: category J using assms(1) diagram_def by auto interpret E: diagram J C E using assms(1) by auto from assms(2) obtain \ where \: "limit_cone J C E a \" by auto interpret \: limit_cone J C E a \ using \ by auto have a: "C.ide a" using \.ide_apex by auto text\ Form the @{term E}-image \GE\ of the diagram @{term E}. \ interpret GE: composite_functor J C D E G .. interpret GE: diagram J D GE.map .. text\Let \G\\ be the @{term G}-image of the cone @{term \}, and note that it is a cone over \GE\.\ let ?G\ = "G o \" interpret G\: cone J D GE.map \G a\ ?G\ using \.cone_axioms preserves_cones by blast text\ Claim that \G\\ is a limit cone for diagram \GE\. \ interpret G\: limit_cone J D GE.map \G a\ ?G\ proof text \ Let @{term \} be an arbitrary cone over \GE\. \ fix b \ assume \: "GE.cone b \" interpret \: cone J D GE.map b \ using \ by auto interpret Fb: constant_functor J C \F b\ apply unfold_locales by (meson F_is_functor \.ide_apex functor.preserves_ide) interpret Adj: meta_adjunction C D F G \ \ using induces_meta_adjunction by auto - interpret Adj: adjunction C D replete_setcat.comp + interpret S: replete_setcat . + interpret Adj: adjunction C D S.comp S.setp Adj.\C Adj.\D F G \ \ Adj.\ Adj.\ Adj.\ Adj.\ using Adj.induces_adjunction by simp text\ For each arrow @{term j} of @{term J}, let @{term "\' j"} be defined to be the adjunct of @{term "\ j"}. We claim that @{term \'} is a cone over @{term E}. \ let ?\' = "\j. if J.arr j then Adj.\ (C.cod (E j)) \\<^sub>C F (\ j) else C.null" have cone_\': "E.cone (F b) ?\'" proof show "\j. \J.arr j \ ?\' j = C.null" by simp fix j assume j: "J.arr j" show "C.dom (?\' j) = Fb.map (J.dom j)" using j \_in_hom by simp show "C.cod (?\' j) = E (J.cod j)" using j \_in_hom by simp show "E j \\<^sub>C ?\' (J.dom j) = ?\' j" proof - have "E j \\<^sub>C ?\' (J.dom j) = (E j \\<^sub>C Adj.\ (E (J.dom j))) \\<^sub>C F (\ (J.dom j))" using j C.comp_assoc by simp also have "... = Adj.\ (E (J.cod j)) \\<^sub>C F (\ j)" proof - have "(E j \\<^sub>C Adj.\ (E (J.dom j))) \\<^sub>C F (\ (J.dom j)) = (Adj.\ (C.cod (E j)) \\<^sub>C Adj.FG.map (E j)) \\<^sub>C F (\ (J.dom j))" using j Adj.\.naturality [of "E j"] by fastforce also have "... = Adj.\ (C.cod (E j)) \\<^sub>C Adj.FG.map (E j) \\<^sub>C F (\ (J.dom j))" using C.comp_assoc by simp also have "... = Adj.\ (E (J.cod j)) \\<^sub>C F (\ j)" proof - have "Adj.FG.map (E j) \\<^sub>C F (\ (J.dom j)) = F (GE.map j \\<^sub>D \ (J.dom j))" using j by simp hence "Adj.FG.map (E j) \\<^sub>C F (\ (J.dom j)) = F (\ j)" using j \.is_natural_1 by metis thus ?thesis using j by simp qed finally show ?thesis by auto qed also have "... = ?\' j" using j by simp finally show ?thesis by auto qed show "?\' (J.cod j) \\<^sub>C Fb.map j = ?\' j" proof - have "?\' (J.cod j) \\<^sub>C Fb.map j = Adj.\ (E (J.cod j)) \\<^sub>C F (\ (J.cod j))" using j Fb.value_is_ide Adj.\.preserves_hom C.comp_arr_dom [of "F (\ (J.cod j))"] C.comp_assoc by simp also have "... = Adj.\ (E (J.cod j)) \\<^sub>C F (\ j)" using j \.is_natural_1 \.is_natural_2 Adj.\.naturality J.arr_cod_iff_arr by (metis J.cod_cod \.A.map_simp) also have "... = ?\' j" using j by simp finally show ?thesis by auto qed qed text\ Using the universal property of the limit cone @{term \}, obtain the unique arrow @{term f} that transforms @{term \} into @{term \'}. \ from this \.is_universal [of "F b" ?\'] obtain f where f: "\f : F b \\<^sub>C a\ \ E.cones_map f \ = ?\'" by auto text\ Let @{term g} be the adjunct of @{term f}, and show that @{term g} transforms @{term G\} into @{term \}. \ let ?g = "G f \\<^sub>D Adj.\ b" have 1: "\?g : b \\<^sub>D G a\" using f \.ide_apex by fastforce moreover have "GE.cones_map ?g ?G\ = \" proof fix j have "\J.arr j \ GE.cones_map ?g ?G\ j = \ j" using 1 G\.cone_axioms \.is_extensional by auto moreover have "J.arr j \ GE.cones_map ?g ?G\ j = \ j" proof - fix j assume j: "J.arr j" have "GE.cones_map ?g ?G\ j = G (\ j) \\<^sub>D ?g" using j 1 G\.cone_axioms mem_Collect_eq restrict_apply by auto also have "... = G (\ j \\<^sub>C f) \\<^sub>D Adj.\ b" using j f \.preserves_hom [of j "J.dom j" "J.cod j"] D.comp_assoc by fastforce also have "... = G (E.cones_map f \ j) \\<^sub>D Adj.\ b" proof - have "\ j \\<^sub>C f = Adj.\ (C.cod (E j)) \\<^sub>C F (\ j)" proof - have "\ j \\<^sub>C f = E.cones_map f \ j" proof - have "E.cone (C.cod f) \" using f \.cone_axioms by blast thus ?thesis using \.is_extensional by simp qed also have "... = Adj.\ (C.cod (E j)) \\<^sub>C F (\ j)" using j f by simp finally show ?thesis by blast qed thus ?thesis using f mem_Collect_eq restrict_apply Adj.F.is_extensional by simp qed also have "... = (G (Adj.\ (C.cod (E j))) \\<^sub>D Adj.\ (D.cod (GE.map j))) \\<^sub>D \ j" using j f Adj.\.naturality [of "\ j"] D.comp_assoc by auto also have "... = D.cod (\ j) \\<^sub>D \ j" using j Adj.\\.triangle_G Adj.\_in_terms_of_\ Adj.\o_def Adj.\_in_terms_of_\ Adj.\o_def Adj.unit_counit_G by fastforce also have "... = \ j" using j D.comp_cod_arr by simp finally show "GE.cones_map ?g ?G\ j = \ j" by metis qed ultimately show "GE.cones_map ?g ?G\ j = \ j" by auto qed ultimately have "\?g : b \\<^sub>D G a\ \ GE.cones_map ?g ?G\ = \" by auto text\ It remains to be shown that @{term g} is the unique such arrow. Given any @{term g'} that transforms @{term G\} into @{term \}, its adjunct transforms @{term \} into @{term \'}. The adjunct of @{term g'} is therefore equal to @{term f}, which implies @{term g'} = @{term g}. \ moreover have "\g'. \g' : b \\<^sub>D G a\ \ GE.cones_map g' ?G\ = \ \ g' = ?g" proof - fix g' assume g': "\g' : b \\<^sub>D G a\ \ GE.cones_map g' ?G\ = \" have 1: "\\ a g' : F b \\<^sub>C a\" using g' a \_in_hom by simp have 2: "E.cones_map (\ a g') \ = ?\'" proof fix j have "\J.arr j \ E.cones_map (\ a g') \ j = ?\' j" using 1 \.cone_axioms by auto moreover have "J.arr j \ E.cones_map (\ a g') \ j = ?\' j" proof - fix j assume j: "J.arr j" have "E.cones_map (\ a g') \ j = \ j \\<^sub>C \ a g'" using 1 \.cone_axioms \.is_extensional by auto also have "... = (\ j \\<^sub>C Adj.\ a) \\<^sub>C F g'" using j a g' Adj.\_in_terms_of_\ C.comp_assoc Adj.\_def by auto also have "... = (Adj.\ (C.cod (E j)) \\<^sub>C F (G (\ j))) \\<^sub>C F g'" using j a g' Adj.\.naturality [of "\ j"] by simp also have "... = Adj.\ (C.cod (E j)) \\<^sub>C F (\ j)" using j a g' G\.cone_axioms C.comp_assoc by auto finally show "E.cones_map (\ a g') \ j = ?\' j" by (simp add: j) qed ultimately show "E.cones_map (\ a g') \ j = ?\' j" by auto qed have "\ a g' = f" proof - have "\!f. \f : F b \\<^sub>C a\ \ E.cones_map f \ = ?\'" using cone_\' \.is_universal by simp moreover have "\\ a g' : F b \\<^sub>C a\ \ E.cones_map (\ a g') \ = ?\'" using 1 2 by simp ultimately show ?thesis using ex1E [of "\f. \f : F b \\<^sub>C a\ \ E.cones_map f \ = ?\'" "\ a g' = f"] using 1 2 Adj.\.is_extensional C.comp_null(2) C.ex_un_null \.cone_axioms f mem_Collect_eq restrict_apply by blast qed hence "\ b (\ a g') = \ b f" by auto hence "g' = \ b f" using \.ide_apex g' by (simp add: \_\) moreover have "?g = \ b f" using f Adj.\_in_terms_of_\ \.ide_apex Adj.\_def by auto ultimately show "g' = ?g" by argo qed ultimately show "\!g. \g : b \\<^sub>D G a\ \ GE.cones_map g ?G\ = \" by blast qed have "GE.limit_cone (G a) ?G\" .. thus ?thesis by auto qed end section "Special Kinds of Limits" subsection "Terminal Objects" text\ An object of a category @{term C} is a terminal object if and only if it is a limit of the empty diagram in @{term C}. \ locale empty_diagram = diagram J C D for J :: "'j comp" (infixr "\\<^sub>J" 55) and C :: "'c comp" (infixr "\" 55) and D :: "'j \ 'c" + assumes is_empty: "\J.arr j" begin lemma has_as_limit_iff_terminal: shows "has_as_limit a \ C.terminal a" proof assume a: "has_as_limit a" show "C.terminal a" proof have "\\. limit_cone a \" using a by auto from this obtain \ where \: "limit_cone a \" by blast interpret \: limit_cone J C D a \ using \ by auto have cone_\: "cone a \" .. show "C.ide a" using \.ide_apex by auto have 1: "\ = (\j. C.null)" using is_empty \.is_extensional by auto show "\a'. C.ide a' \ \!f. \f : a' \ a\" proof - fix a' assume a': "C.ide a'" interpret A': constant_functor J C a' apply unfold_locales using a' by auto let ?\' = "\j. C.null" have cone_\': "cone a' ?\'" using a' is_empty apply unfold_locales by auto hence "\!f. \f : a' \ a\ \ cones_map f \ = ?\'" using \.is_universal by force moreover have "\f. \f : a' \ a\ \ cones_map f \ = ?\'" using 1 cone_\ by auto ultimately show "\!f. \f : a' \ a\" by blast qed qed next assume a: "C.terminal a" show "has_as_limit a" proof - let ?\ = "\j. C.null" have "C.ide a" using a C.terminal_def by simp interpret A: constant_functor J C a apply unfold_locales using \C.ide a\ by simp interpret \: cone J C D a ?\ using \C.ide a\ is_empty by (unfold_locales, auto) have cone_\: "cone a ?\" .. have 1: "\a' \'. cone a' \' \ \' = (\j. C.null)" proof - fix a' \' assume \': "cone a' \'" interpret \': cone J C D a' \' using \' by auto show "\' = (\j. C.null)" using is_empty \'.is_extensional by metis qed have "limit_cone a ?\" proof fix a' \' assume \': "cone a' \'" have 2: "\' = (\j. C.null)" using 1 \' by simp interpret \': cone J C D a' \' using \' by auto have "\!f. \f : a' \ a\" using a C.terminal_def \'.ide_apex by simp moreover have "\f. \f : a' \ a\ \ cones_map f ?\ = \'" using 1 2 cones_map_mapsto cone_\ \'.cone_axioms mem_Collect_eq by blast ultimately show "\!f. \f : a' \ a\ \ cones_map f (\j. C.null) = \'" by blast qed thus ?thesis by auto qed qed end subsection "Products" text\ A \emph{product} in a category @{term C} is a limit of a discrete diagram in @{term C}. \ locale discrete_diagram = J: category J + diagram J C D for J :: "'j comp" (infixr "\\<^sub>J" 55) and C :: "'c comp" (infixr "\" 55) and D :: "'j \ 'c" + assumes is_discrete: "J.arr = J.ide" begin abbreviation mkCone where "mkCone F \ (\j. if J.arr j then F j else C.null)" lemma cone_mkCone: assumes "C.ide a" and "\j. J.arr j \ \F j : a \ D j\" shows "cone a (mkCone F)" proof - interpret A: constant_functor J C a using assms(1) by unfold_locales auto show "cone a (mkCone F)" using assms(2) is_discrete apply unfold_locales apply auto apply (metis C.in_homE C.comp_cod_arr) using C.comp_arr_ide by fastforce qed lemma mkCone_cone: assumes "cone a \" shows "mkCone \ = \" proof - interpret \: cone J C D a \ using assms by auto show "mkCone \ = \" using \.is_extensional by auto qed end text\ The following locale defines a discrete diagram in a category @{term C}, given an index set @{term I} and a function @{term D} mapping @{term I} to objects of @{term C}. Here we obtain the diagram shape @{term J} using a discrete category construction that allows us to directly identify the objects of @{term J} with the elements of @{term I}, however this construction can only be applied in case the set @{term I} is not the universe of its element type. \ locale discrete_diagram_from_map = J: discrete_category I null + C: category C for I :: "'i set" and C :: "'c comp" (infixr "\" 55) and D :: "'i \ 'c" and null :: 'i + assumes maps_to_ide: "i \ I \ C.ide (D i)" begin definition map where "map j \ if J.arr j then D j else C.null" end sublocale discrete_diagram_from_map \ discrete_diagram J.comp C map using map_def maps_to_ide J.arr_char J.Null_not_in_Obj J.null_char by unfold_locales auto locale product_cone = J: category J + C: category C + D: discrete_diagram J C D + limit_cone J C D a \ for J :: "'j comp" (infixr "\\<^sub>J" 55) and C :: "'c comp" (infixr "\" 55) and D :: "'j \ 'c" and a :: 'c and \ :: "'j \ 'c" begin lemma is_cone: shows "D.cone a \" .. text\ The following versions of @{prop is_universal} and @{prop induced_arrowI} from the \limit_cone\ locale are specialized to the case in which the underlying diagram is a product diagram. \ lemma is_universal': assumes "C.ide b" and "\j. J.arr j \ \F j: b \ D j\" shows "\!f. \f : b \ a\ \ (\j. J.arr j \ \ j \ f = F j)" proof - let ?\ = "D.mkCone F" interpret B: constant_functor J C b using assms(1) by unfold_locales auto have cone_\: "D.cone b ?\" using assms D.is_discrete D.cone_mkCone by blast interpret \: cone J C D b ?\ using cone_\ by auto have "\!f. \f : b \ a\ \ D.cones_map f \ = ?\" using cone_\ is_universal by force moreover have "\f. \f : b \ a\ \ D.cones_map f \ = ?\ \ (\j. J.arr j \ \ j \ f = F j)" proof - fix f assume f: "\f : b \ a\" show "D.cones_map f \ = ?\ \ (\j. J.arr j \ \ j \ f = F j)" proof assume 1: "D.cones_map f \ = ?\" show "\j. J.arr j \ \ j \ f = F j" proof - have "\j. J.arr j \ \ j \ f = F j" proof - fix j assume j: "J.arr j" have "\ j \ f = D.cones_map f \ j" using j f cone_axioms by force also have "... = F j" using j 1 by simp finally show "\ j \ f = F j" by auto qed thus ?thesis by auto qed next assume 1: "\j. J.arr j \ \ j \ f = F j" show "D.cones_map f \ = ?\" using 1 f is_cone \.is_extensional D.is_discrete is_cone cone_\ by auto qed qed ultimately show ?thesis by blast qed abbreviation induced_arrow' :: "'c \ ('j \ 'c) \ 'c" where "induced_arrow' b F \ induced_arrow b (D.mkCone F)" lemma induced_arrowI': assumes "C.ide b" and "\j. J.arr j \ \F j : b \ D j\" shows "\j. J.arr j \ \ j \ induced_arrow' b F = F j" proof - interpret B: constant_functor J C b using assms(1) by unfold_locales auto interpret \: cone J C D b \D.mkCone F\ using assms D.cone_mkCone by blast have cone_\: "D.cone b (D.mkCone F)" .. hence 1: "D.cones_map (induced_arrow' b F) \ = D.mkCone F" using induced_arrowI by blast fix j assume j: "J.arr j" have "\ j \ induced_arrow' b F = D.cones_map (induced_arrow' b F) \ j" using induced_arrowI(1) cone_\ is_cone is_extensional by force also have "... = F j" using j 1 by auto finally show "\ j \ induced_arrow' b F = F j" by auto qed end context discrete_diagram begin lemma product_coneI: assumes "limit_cone a \" shows "product_cone J C D a \" by (meson assms discrete_diagram_axioms functor_axioms functor_def product_cone.intro) end context category begin definition has_as_product where "has_as_product J D a \ (\\. product_cone J C D a \)" lemma product_is_ide: assumes "has_as_product J D a" shows "ide a" proof - obtain \ where \: "product_cone J C D a \" using assms has_as_product_def by blast interpret \: product_cone J C D a \ using \ by auto show ?thesis using \.ide_apex by auto qed text\ A category has @{term I}-indexed products for an @{typ 'i}-set @{term I} if every @{term I}-indexed discrete diagram has a product. In order to reap the benefits of being able to directly identify the elements of a set I with the objects of discrete category it generates (thereby avoiding the use of coercion maps), it is necessary to assume that @{term "I \ UNIV"}. If we want to assert that a category has products indexed by the universe of some type @{typ 'i}, we have to pass to a larger type, such as @{typ "'i option"}. \ definition has_products where "has_products (I :: 'i set) \ I \ UNIV \ (\J D. discrete_diagram J C D \ Collect (partial_magma.arr J) = I \ (\a. has_as_product J D a))" lemma ex_productE: assumes "\a. has_as_product J D a" obtains a \ where "product_cone J C D a \" using assms has_as_product_def someI_ex [of "\a. has_as_product J D a"] by metis lemma has_products_if_has_limits: assumes "has_limits (undefined :: 'j)" and "I \ (UNIV :: 'j set)" shows "has_products I" proof (unfold has_products_def, intro conjI allI impI) show "I \ UNIV" by fact fix J D assume D: "discrete_diagram J C D \ Collect (partial_magma.arr J) = I" interpret D: discrete_diagram J C D using D by simp have 1: "\a. D.has_as_limit a" using assms D D.diagram_axioms D.J.category_axioms by (simp add: has_limits_of_shape_def has_limits_def) show "\a. has_as_product J D a" using 1 has_as_product_def D.product_coneI by blast qed lemma has_finite_products_if_has_finite_limits: assumes "\J :: 'j comp. (finite (Collect (partial_magma.arr J))) \ has_limits_of_shape J" and "finite (I :: 'j set)" and "I \ UNIV" shows "has_products I" proof (unfold has_products_def, intro conjI allI impI) show "I \ UNIV" by fact fix J D assume D: "discrete_diagram J C D \ Collect (partial_magma.arr J) = I" interpret D: discrete_diagram J C D using D by simp have 1: "\a. D.has_as_limit a" using assms D has_limits_of_shape_def D.diagram_axioms by auto show "\a. has_as_product J D a" using 1 has_as_product_def D.product_coneI by blast qed lemma has_products_preserved_by_bijection: assumes "has_products I" and "bij_betw \ I I'" and "I' \ UNIV" shows "has_products I'" proof (unfold has_products_def, intro conjI allI impI) show "I' \ UNIV" by fact show "\J' D'. discrete_diagram J' C D' \ Collect (partial_magma.arr J') = I' \ \a. has_as_product J' D' a" proof - fix J' D' assume 1: "discrete_diagram J' C D' \ Collect (partial_magma.arr J') = I'" interpret J': category J' using 1 by (simp add: discrete_diagram_def) interpret D': discrete_diagram J' C D' using 1 by simp interpret J: discrete_category I \SOME x. x \ I\ using assms has_products_def [of I] someI_ex [of "\x. x \ I"] by unfold_locales auto have 2: "Collect J.arr = I \ Collect J'.arr = I'" using 1 by auto have \: "bij_betw \ (Collect J.arr) (Collect J'.arr)" using 2 assms(2) by simp let ?\ = "\j. if J.arr j then \ j else J'.null" let ?\' = "\j'. if J'.arr j' then the_inv_into I \ j' else J.null" interpret \: "functor" J.comp J' ?\ proof - have "\ ` I = I'" using \ 2 bij_betw_def [of \ I I'] by simp hence "\j. J.arr j \ J'.arr (?\ j)" using 1 D'.is_discrete by auto thus "functor J.comp J' ?\" using D'.is_discrete J.is_discrete J.seqE by unfold_locales auto qed interpret \': "functor" J' J.comp ?\' proof - have "the_inv_into I \ ` I' = I" using assms(2) \ bij_betw_the_inv_into bij_betw_imp_surj_on by metis hence "\j'. J'.arr j' \ J.arr (?\' j')" using 2 D'.is_discrete J.is_discrete by auto thus "functor J' J.comp ?\'" using D'.is_discrete J.is_discrete J'.seqE by unfold_locales auto qed let ?D = "\i. D' (\ i)" interpret D: discrete_diagram_from_map I C ?D \SOME j. j \ I\ using assms 1 D'.is_discrete bij_betw_imp_surj_on \.preserves_ide by unfold_locales auto obtain a where a: "has_as_product J.comp D.map a" using assms D.discrete_diagram_axioms has_products_def [of I] by auto obtain \ where \: "product_cone J.comp C D.map a \" using a has_as_product_def by blast interpret \: product_cone J.comp C D.map a \ using \ by simp let ?\' = "\ o ?\'" interpret A: constant_functor J' C a using \.ide_apex by unfold_locales simp interpret \': natural_transformation J' C A.map D' ?\' proof - have "\.A.map \ ?\' = A.map" using \ A.map_def \'.preserves_arr \.A.is_extensional J.not_arr_null by auto moreover have "D.map \ ?\' = D'" proof fix j' have "J'.arr j' \ (D.map \ ?\') j' = D' j'" proof - assume 2: "J'.arr j'" have 3: "inj_on \ I" using assms(2) bij_betw_imp_inj_on by auto have "\ ` I = I'" by (metis (no_types) assms(2) bij_betw_imp_surj_on) hence "\ ` I = Collect J'.arr" using 1 by force thus ?thesis using 2 3 D.map_def \'.preserves_arr f_the_inv_into_f by fastforce qed moreover have "\ J'.arr j' \ (D.map \ ?\') j' = D' j'" using D.is_extensional D'.is_extensional by (simp add: J.Null_not_in_Obj J.null_char) ultimately show "(D.map \ ?\') j' = D' j'" by blast qed ultimately show "natural_transformation J' C A.map D' ?\'" using \.natural_transformation_axioms \'.as_nat_trans.natural_transformation_axioms horizontal_composite [of J' J.comp ?\' ?\' ?\' C \.A.map D.map \] by simp qed interpret \': cone J' C D' a ?\' .. interpret \': product_cone J' C D' a ?\' proof fix a' \' assume \': "D'.cone a' \'" interpret \': cone J' C D' a' \' using \' by simp show "\!f. \f : a' \ a\ \ D'.cones_map f (\ \ ?\') = \'" proof - let ?\ = "\' o ?\" interpret A': constant_functor J.comp C a' using \'.ide_apex by unfold_locales simp interpret \: natural_transformation J.comp C A'.map D.map ?\ proof - have "\'.A.map \ ?\ = A'.map" using \ \.preserves_arr A'.map_def \'.A.is_extensional by auto moreover have "D' \ ?\ = D.map" using \ D.map_def D'.is_extensional by auto ultimately show "natural_transformation J.comp C A'.map D.map ?\" using \'.natural_transformation_axioms \.as_nat_trans.natural_transformation_axioms horizontal_composite [of J.comp J' ?\ ?\ ?\ C \'.A.map D' \'] by simp qed interpret \: cone J.comp C D.map a' ?\ .. have *: "\!f. \f : a' \ a\ \ D.cones_map f \ = ?\" using \.is_universal \.cone_axioms by simp show "\!f. \f : a' \ a\ \ D'.cones_map f ?\' = \'" proof - have "\f. \f : a' \ a\ \ D'.cones_map f ?\' = \'" proof - obtain f where f: "\f : a' \ a\ \ D.cones_map f \ = ?\" using * by blast have "D'.cones_map f ?\' = \'" proof fix j' show "D'.cones_map f ?\' j' = \' j'" proof (cases "J'.arr j'") assume j': "\ J'.arr j'" show "D'.cones_map f ?\' j' = \' j'" using f j' \'.is_extensional \'.cone_axioms by auto next assume j': "J'.arr j'" show "D'.cones_map f ?\' j' = \' j'" proof - have "D'.cones_map f ?\' j' = \ (the_inv_into I \ j') \ f" using f j' \'.cone_axioms by auto also have "... = D.cones_map f \ (the_inv_into I \ j')" proof - have "arr f \ dom f = a' \ cod f = a" using f by blast thus ?thesis using \'.preserves_arr \.cone_\ j' by auto qed also have "... = (\' \ ?\) (the_inv_into I \ j')" using f by simp also have "... = \' j'" using assms(2) j' 2 bij_betw_def [of \ I I'] bij_betw_imp_inj_on \'.preserves_arr f_the_inv_into_f by fastforce finally show ?thesis by simp qed qed qed thus ?thesis using f by blast qed moreover have "\f f'. \ \f : a' \ a\; D'.cones_map f ?\' = \'; \f' : a' \ a\; D'.cones_map f' ?\' = \' \ \ f = f'" proof - fix f f' assume f: "\f : a' \ a\" and f': "\f' : a' \ a\" and f\': "D'.cones_map f ?\' = \'" and f'\': "D'.cones_map f' ?\' = \'" have "D.cones_map f \ = \' \ ?\ \ D.cones_map f' \ = \' o ?\" proof (intro conjI) show "D.cones_map f \ = \' \ ?\" proof fix j have "\ J.arr j \ D.cones_map f \ j = (\' \ ?\) j" using f f\' \.cone_axioms \.is_extensional by auto moreover have "J.arr j \ D.cones_map f \ j = (\' \ ?\) j" proof - assume j: "J.arr j" have 1: "j = the_inv_into I \ (\ j)" using assms(2) j \ the_inv_into_f_f bij_betw_imp_inj_on J.arr_char by metis have "D.cones_map f \ j = D.cones_map f \ (the_inv_into I \ (\ j))" using 1 by simp also have "... = (\' \ ?\) j" using f j f\' 1 \.cone_axioms \'.cone_axioms \.preserves_arr by auto finally show "D.cones_map f \ j = (\' \ ?\) j" by blast qed ultimately show "D.cones_map f \ j = (\' \ ?\) j" by blast qed show "D.cones_map f' \ = \' \ ?\" proof fix j have "\ J.arr j \ D.cones_map f' \ j = (\' \ ?\) j" using f' f\' \.cone_axioms \.is_extensional by auto moreover have "J.arr j \ D.cones_map f' \ j = (\' \ ?\) j" proof - assume j: "J.arr j" have 1: "j = the_inv_into I \ (\ j)" using assms(2) j \ the_inv_into_f_f bij_betw_imp_inj_on J.arr_char by metis have "D.cones_map f' \ j = D.cones_map f' \ (the_inv_into I \ (\ j))" using 1 by simp also have "... = (\' \ ?\) j" using f' j f'\' 1 \.cone_axioms \'.cone_axioms \.preserves_arr by auto finally show "D.cones_map f' \ j = (\' \ ?\) j" by blast qed ultimately show "D.cones_map f' \ j = (\' \ ?\) j" by blast qed qed thus "f = f'" using f f' * by auto qed ultimately show ?thesis by blast qed qed qed have "has_as_product J' D' a" using has_as_product_def \'.product_cone_axioms by auto thus "\a. has_as_product J' D' a" by blast qed qed lemma ide_is_unary_product: assumes "ide a" shows "\m n :: nat. m \ n \ has_as_product (discrete_category.comp {m :: nat} (n :: nat)) (\i. if i = m then a else null) a" proof - fix m n :: nat assume neq: "m \ n" have "{m :: nat} \ UNIV" proof - have "finite {m :: nat}" by simp moreover have "\finite (UNIV :: nat set)" by simp ultimately show ?thesis by fastforce qed interpret J: discrete_category "{m :: nat}" \n :: nat\ using neq \{m :: nat} \ UNIV\ by unfold_locales auto let ?D = "\i. if i = m then a else null" interpret D: discrete_diagram J.comp C ?D apply unfold_locales using assms J.null_char neq apply auto by metis interpret A: constant_functor J.comp C a using assms by unfold_locales auto show "has_as_product J.comp ?D a" proof (unfold has_as_product_def) let ?\ = "\i :: nat. if i = m then a else null" interpret \: natural_transformation J.comp C A.map ?D ?\ using assms J.arr_char J.dom_char J.cod_char by unfold_locales auto interpret \: cone J.comp C ?D a ?\ .. interpret \: product_cone J.comp C ?D a ?\ proof fix a' \' assume \': "D.cone a' \'" interpret \': cone J.comp C ?D a' \' using \' by auto show "\!f. \f : a' \ a\ \ D.cones_map f ?\ = \'" proof show "\\' m : a' \ a\ \ D.cones_map (\' m) ?\ = \'" proof show 1: "\\' m : a' \ a\" using \'.preserves_hom \'.A.map_def J.arr_char J.dom_char J.cod_char by auto show "D.cones_map (\' m) ?\ = \'" proof fix j show "D.cones_map (\' m) (\i. if i = m then a else null) j = \' j" using J.arr_char J.dom_char J.cod_char J.ide_char \.cone_axioms comp_cod_arr apply (cases "j = m") apply simp using \'.is_extensional by simp qed qed show "\f. \f : a' \ a\ \ D.cones_map f ?\ = \' \ f = \' m" proof - fix f assume f: "\f : a' \ a\ \ D.cones_map f ?\ = \'" show "f = \' m" using assms f \'.preserves_hom J.arr_char J.dom_char J.cod_char \.cone_axioms comp_cod_arr by auto qed qed qed show "\\. product_cone J.comp C (\i. if i = m then a else null) a \" using \.product_cone_axioms by blast qed qed lemma has_unary_products: assumes "card I = 1" and "I \ UNIV" shows "has_products I" proof - obtain i where i: "I = {i}" using assms card_1_singletonE by blast obtain n where n: "n \ I" using assms by auto have "has_products {1 :: nat}" proof (unfold has_products_def, intro conjI) show "{1 :: nat} \ UNIV" by auto show "\(J :: nat comp) D. discrete_diagram J (\) D \ Collect (partial_magma.arr J) = {1} \ (\a. has_as_product J D a)" proof fix J :: "nat comp" show "\D. discrete_diagram J (\) D \ Collect (partial_magma.arr J) = {1} \ (\a. has_as_product J D a)" proof (intro allI impI) fix D assume D: "discrete_diagram J (\) D \ Collect (partial_magma.arr J) = {1}" interpret D: discrete_diagram J C D using D by auto interpret J: discrete_category \{1 :: nat}\ D.J.null by (metis D D.J.not_arr_null discrete_category.intro mem_Collect_eq) have 1: "has_as_product J.comp D (D 1)" proof - have "has_as_product J.comp (\i. if i = 1 then D 1 else null) (D 1)" using ide_is_unary_product D.preserves_ide D.is_discrete D J.null_char by (metis J.Null_not_in_Obj insertCI mem_Collect_eq) moreover have "D = (\i. if i = 1 then D 1 else null)" proof fix j show "D j = (if j = 1 then D 1 else null)" using D D.is_extensional by auto qed ultimately show ?thesis by simp qed moreover have "J = J.comp" proof - have "\j j'. J j j' = J.comp j j'" proof - fix j j' show "J j j' = J.comp j j'" using J.comp_char D.is_discrete D by (metis D.J.category_axioms D.J.ext D.J.ide_def J.null_char category.seqE mem_Collect_eq) qed thus ?thesis by auto qed ultimately show "\a. has_as_product J D a" by auto qed qed qed moreover have "bij_betw (\k. if k = 1 then i else n) {1 :: nat} I" using i by (intro bij_betwI') auto ultimately show "has_products I" using assms has_products_preserved_by_bijection by blast qed end subsection "Equalizers" text\ An \emph{equalizer} in a category @{term C} is a limit of a parallel pair of arrows in @{term C}. \ locale parallel_pair_diagram = J: parallel_pair + C: category C for C :: "'c comp" (infixr "\" 55) and f0 :: 'c and f1 :: 'c + assumes is_parallel: "C.par f0 f1" begin no_notation J.comp (infixr "\" 55) notation J.comp (infixr "\\<^sub>J" 55) definition map where "map \ (\j. if j = J.Zero then C.dom f0 else if j = J.One then C.cod f0 else if j = J.j0 then f0 else if j = J.j1 then f1 else C.null)" lemma map_simp: shows "map J.Zero = C.dom f0" and "map J.One = C.cod f0" and "map J.j0 = f0" and "map J.j1 = f1" proof - show "map J.Zero = C.dom f0" using map_def by metis show "map J.One = C.cod f0" using map_def J.Zero_not_eq_One by metis show "map J.j0 = f0" using map_def J.Zero_not_eq_j0 J.One_not_eq_j0 by metis show "map J.j1 = f1" using map_def J.Zero_not_eq_j1 J.One_not_eq_j1 J.j0_not_eq_j1 by metis qed end sublocale parallel_pair_diagram \ diagram J.comp C map apply unfold_locales apply (simp add: J.arr_char map_def) using map_def is_parallel J.arr_char J.cod_simp J.dom_simp apply auto[2] proof - show 1: "\j. J.arr j \ C.cod (map j) = map (J.cod j)" using is_parallel map_simp(1-4) J.arr_char by auto next fix j j' assume jj': "J.seq j' j" show "map (j' \\<^sub>J j) = map j' \ map j" proof - have 1: "(j = J.Zero \ j' \ J.One) \ (j \ J.Zero \ j' = J.One)" using jj' J.seq_char by blast moreover have "j = J.Zero \ j' \ J.One \ ?thesis" by (metis (no_types, lifting) C.arr_dom_iff_arr C.comp_arr_dom C.dom_dom J.comp_simp(1) jj' map_simp(1,3-4) J.arr_char is_parallel) moreover have "j \ J.Zero \ j' = J.One \ ?thesis" by (metis (no_types, lifting) C.comp_arr_dom C.comp_cod_arr C.seqE J.comp_char jj' map_simp(2-4) J.arr_char is_parallel) ultimately show ?thesis by blast qed qed context parallel_pair_diagram begin definition mkCone where "mkCone e \ \j. if J.arr j then if j = J.Zero then e else f0 \ e else C.null" abbreviation is_equalized_by where "is_equalized_by e \ C.seq f0 e \ f0 \ e = f1 \ e" abbreviation has_as_equalizer where "has_as_equalizer e \ limit_cone (C.dom e) (mkCone e)" lemma cone_mkCone: assumes "is_equalized_by e" shows "cone (C.dom e) (mkCone e)" proof - interpret E: constant_functor J.comp C \C.dom e\ apply unfold_locales using assms by auto show "cone (C.dom e) (mkCone e)" using assms mkCone_def apply unfold_locales apply auto[2] using C.dom_comp C.seqE C.cod_comp J.Zero_not_eq_One J.arr_char' J.cod_char map_def apply (metis (no_types, lifting) C.not_arr_null parallel_pair.cod_simp(1) preserves_arr) proof - show "\j. J.arr j \ map j \ mkCone e (J.dom j) = mkCone e j" proof - fix j assume j: "J.arr j" have 1: "j = J.Zero \ map j \ mkCone e (J.dom j) = mkCone e j" using assms j mkCone_def C.cod_comp by (metis (no_types, lifting) C.comp_cod_arr J.arr_char J.dom_char map_def J.dom_simp(2)) show "map j \ mkCone e (J.dom j) = mkCone e j" by (metis (no_types, lifting) 1 C.comp_cod_arr C.seqE assms j map_simp(1) mkCone_def J.dom_simp(1)) qed next show "\j. J.arr j \ mkCone e (J.cod j) \ E.map j = mkCone e j" by (metis (no_types, lifting) C.comp_arr_dom C.comp_assoc C.seqE J.arr_cod_iff_arr J.seqI J.seq_char assms E.map_simp mkCone_def J.cod_simp(1) J.dom_simp(1)) qed qed lemma is_equalized_by_cone: assumes "cone a \" shows "is_equalized_by (\ (J.Zero))" proof - interpret \: cone J.comp C map a \ using assms by auto show ?thesis by (metis (no_types, lifting) J.arr_char J.cod_char cone_def \.component_in_hom \.is_natural_1 \.naturality assms C.in_homE constant_functor.map_simp J.dom_simp(3-4) map_simp(3-4)) qed lemma mkCone_cone: assumes "cone a \" shows "mkCone (\ J.Zero) = \" proof - interpret \: cone J.comp C map a \ using assms by auto have 1: "is_equalized_by (\ J.Zero)" using assms is_equalized_by_cone by blast show ?thesis proof fix j have "j = J.Zero \ mkCone (\ J.Zero) j = \ j" using mkCone_def \.is_extensional by simp moreover have "j = J.One \ j = J.j0 \ j = J.j1 \ mkCone (\ J.Zero) j = \ j" using J.arr_char J.cod_char J.dom_char J.seq_char mkCone_def \.is_natural_1 \.is_natural_2 \.A.map_simp map_def by (metis (no_types, lifting) J.Zero_not_eq_j0 J.dom_simp(2)) ultimately have "J.arr j \ mkCone (\ J.Zero) j = \ j" using J.arr_char by auto thus "mkCone (\ J.Zero) j = \ j" using mkCone_def \.is_extensional by fastforce qed qed end locale equalizer_cone = J: parallel_pair + C: category C + D: parallel_pair_diagram C f0 f1 + limit_cone J.comp C D.map "C.dom e" "D.mkCone e" for C :: "'c comp" (infixr "\" 55) and f0 :: 'c and f1 :: 'c and e :: 'c begin lemma equalizes: shows "D.is_equalized_by e" proof show "C.seq f0 e" proof (intro C.seqI) show "C.arr e" using ide_apex C.arr_dom_iff_arr by fastforce show "C.arr f0" using D.map_simp D.preserves_arr J.arr_char by metis show "C.dom f0 = C.cod e" using J.arr_char J.ide_char D.mkCone_def D.map_simp preserves_cod [of J.Zero] by auto qed show "f0 \ e = f1 \ e" using D.map_simp D.mkCone_def J.arr_char naturality [of J.j0] naturality [of J.j1] by force qed lemma is_universal': assumes "D.is_equalized_by e'" shows "\!h. \h : C.dom e' \ C.dom e\ \ e \ h = e'" proof - have "D.cone (C.dom e') (D.mkCone e')" using assms D.cone_mkCone by blast moreover have 0: "D.cone (C.dom e) (D.mkCone e)" .. ultimately have 1: "\!h. \h : C.dom e' \ C.dom e\ \ D.cones_map h (D.mkCone e) = D.mkCone e'" using is_universal [of "C.dom e'" "D.mkCone e'"] by auto have 2: "\h. \h : C.dom e' \ C.dom e\ \ D.cones_map h (D.mkCone e) = D.mkCone e' \ e \ h = e'" proof - fix h assume h: "\h : C.dom e' \ C.dom e\" show "D.cones_map h (D.mkCone e) = D.mkCone e' \ e \ h = e'" proof assume 3: "D.cones_map h (D.mkCone e) = D.mkCone e'" show "e \ h = e'" proof - have "e' = D.mkCone e' J.Zero" using D.mkCone_def J.arr_char by simp also have "... = D.cones_map h (D.mkCone e) J.Zero" using 3 by simp also have "... = e \ h" using 0 h D.mkCone_def J.arr_char by auto finally show ?thesis by auto qed next assume e': "e \ h = e'" show "D.cones_map h (D.mkCone e) = D.mkCone e'" proof fix j have "\J.arr j \ D.cones_map h (D.mkCone e) j = D.mkCone e' j" using h cone_axioms D.mkCone_def by auto moreover have "j = J.Zero \ D.cones_map h (D.mkCone e) j = D.mkCone e' j" using h e' cone_\ D.mkCone_def J.arr_char [of J.Zero] by force moreover have "J.arr j \ j \ J.Zero \ D.cones_map h (D.mkCone e) j = D.mkCone e' j" using C.comp_assoc D.mkCone_def cone_\ e' h by auto ultimately show "D.cones_map h (D.mkCone e) j = D.mkCone e' j" by blast qed qed qed thus ?thesis using 1 by blast qed lemma induced_arrowI': assumes "D.is_equalized_by e'" shows "\induced_arrow (C.dom e') (D.mkCone e') : C.dom e' \ C.dom e\" and "e \ induced_arrow (C.dom e') (D.mkCone e') = e'" proof - interpret A': constant_functor J.comp C \C.dom e'\ using assms by (unfold_locales, auto) have cone: "D.cone (C.dom e') (D.mkCone e')" using assms D.cone_mkCone [of e'] by blast have "e \ induced_arrow (C.dom e') (D.mkCone e') = D.cones_map (induced_arrow (C.dom e') (D.mkCone e')) (D.mkCone e) J.Zero" using cone induced_arrowI(1) D.mkCone_def J.arr_char cone_\ by force also have "... = e'" proof - have "D.cones_map (induced_arrow (C.dom e') (D.mkCone e')) (D.mkCone e) = D.mkCone e'" using cone induced_arrowI by blast thus ?thesis using J.arr_char D.mkCone_def by simp qed finally have 1: "e \ induced_arrow (C.dom e') (D.mkCone e') = e'" by auto show "\induced_arrow (C.dom e') (D.mkCone e') : C.dom e' \ C.dom e\" using 1 cone induced_arrowI by simp show "e \ induced_arrow (C.dom e') (D.mkCone e') = e'" using 1 cone induced_arrowI by simp qed end context category begin definition has_as_equalizer where "has_as_equalizer f0 f1 e \ par f0 f1 \ parallel_pair_diagram.has_as_equalizer C f0 f1 e" definition has_equalizers where "has_equalizers = (\f0 f1. par f0 f1 \ (\e. has_as_equalizer f0 f1 e))" end section "Limits by Products and Equalizers" text\ A category with equalizers has limits of shape @{term J} if it has products indexed by the set of arrows of @{term J} and the set of objects of @{term J}. The proof is patterned after \cite{MacLane}, Theorem 2, page 109: \begin{quotation} \noindent ``The limit of \F: J \ C\ is the equalizer \e\ of \f, g: \\<^sub>i F\<^sub>i \ \\<^sub>u F\<^sub>c\<^sub>o\<^sub>d \<^sub>u (u \ arr J, i \ J)\ where \p\<^sub>u f = p\<^sub>c\<^sub>o\<^sub>d \<^sub>u\, \p\<^sub>u g = F\<^sub>u o p\<^sub>d\<^sub>o\<^sub>m \<^sub>u\; the limiting cone \\\ is \\\<^sub>j = p\<^sub>j e\, for \j \ J\.'' \end{quotation} \ locale category_with_equalizers = category C for C :: "'c comp" (infixr "\" 55) + assumes has_equalizers: "has_equalizers" begin lemma has_limits_if_has_products: fixes J :: "'j comp" (infixr "\\<^sub>J" 55) assumes "category J" and "has_products (Collect (partial_magma.ide J))" and "has_products (Collect (partial_magma.arr J))" shows "has_limits_of_shape J" proof (unfold has_limits_of_shape_def) interpret J: category J using assms(1) by auto have "\D. diagram J C D \ (\a \. limit_cone J C D a \)" proof - fix D assume D: "diagram J C D" interpret D: diagram J C D using D by auto text\ First, construct the two required products and their cones. \ interpret Obj: discrete_category \Collect J.ide\ J.null using J.not_arr_null J.ideD(1) mem_Collect_eq by (unfold_locales, blast) interpret \o: discrete_diagram_from_map \Collect J.ide\ C D J.null using D.preserves_ide by (unfold_locales, auto) have "\p. has_as_product Obj.comp \o.map p" using assms(2) \o.diagram_axioms has_products_def Obj.arr_char by (metis (no_types, lifting) Collect_cong \o.discrete_diagram_axioms mem_Collect_eq) from this obtain \o \o where \o: "product_cone Obj.comp C \o.map \o \o" using ex_productE [of Obj.comp \o.map] by auto interpret \o: product_cone Obj.comp C \o.map \o \o using \o by auto have \o_in_hom: "\j. Obj.arr j \ \\o j : \o \ D j\" using \o.preserves_dom \o.preserves_cod \o.map_def by auto interpret Arr: discrete_category \Collect J.arr\ J.null using J.not_arr_null by (unfold_locales, blast) interpret \a: discrete_diagram_from_map \Collect J.arr\ C \D o J.cod\ J.null by (unfold_locales, auto) have "\p. has_as_product Arr.comp \a.map p" using assms(3) has_products_def [of "Collect J.arr"] \a.discrete_diagram_axioms by blast from this obtain \a \a where \a: "product_cone Arr.comp C \a.map \a \a" using ex_productE [of Arr.comp \a.map] by auto interpret \a: product_cone Arr.comp C \a.map \a \a using \a by auto have \a_in_hom: "\j. Arr.arr j \ \\a j : \a \ D (J.cod j)\" using \a.preserves_cod \a.preserves_dom \a.map_def by auto text\ Next, construct a parallel pair of arrows \f, g: \o \ \a\ that expresses the commutativity constraints imposed by the diagram. \ interpret \o: constant_functor Arr.comp C \o using \o.ide_apex by (unfold_locales, auto) let ?\ = "\j. if Arr.arr j then \o (J.cod j) else null" interpret \: cone Arr.comp C \a.map \o ?\ using \o.ide_apex \o_in_hom \a.map_def \o.map_def \o.is_discrete \o.is_natural_2 comp_cod_arr by (unfold_locales, auto) let ?f = "\a.induced_arrow \o ?\" have f_in_hom: "\?f : \o \ \a\" using \.cone_axioms \a.induced_arrowI by blast have f_map: "\a.cones_map ?f \a = ?\" using \.cone_axioms \a.induced_arrowI by blast have ff: "\j. J.arr j \ \a j \ ?f = \o (J.cod j)" using \.component_in_hom \a.induced_arrowI' \o.ide_apex by auto let ?\' = "\j. if Arr.arr j then D j \ \o (J.dom j) else null" interpret \': cone Arr.comp C \a.map \o ?\' using \o.ide_apex \o_in_hom \o.map_def \a.map_def comp_arr_dom comp_cod_arr by (unfold_locales, auto) let ?g = "\a.induced_arrow \o ?\'" have g_in_hom: "\?g : \o \ \a\" using \'.cone_axioms \a.induced_arrowI by blast have g_map: "\a.cones_map ?g \a = ?\'" using \'.cone_axioms \a.induced_arrowI by blast have gg: "\j. J.arr j \ \a j \ ?g = D j \ \o (J.dom j)" using \'.component_in_hom \a.induced_arrowI' \o.ide_apex by force interpret PP: parallel_pair_diagram C ?f ?g using f_in_hom g_in_hom by (elim in_homE, unfold_locales, auto) from PP.is_parallel obtain e where equ: "PP.has_as_equalizer e" using has_equalizers has_equalizers_def has_as_equalizer_def by blast interpret EQU: limit_cone PP.J.comp C PP.map \dom e\ \PP.mkCone e\ using equ by auto interpret EQU: equalizer_cone C ?f ?g e .. text\ An arrow @{term h} with @{term "cod h = \o"} equalizes @{term f} and @{term g} if and only if it satisfies the commutativity condition required for a cone over @{term D}. \ have E: "\h. \h : dom h \ \o\ \ ?f \ h = ?g \ h \ (\j. J.arr j \ ?\ j \ h = ?\' j \ h)" proof fix h assume h: "\h : dom h \ \o\" show "?f \ h = ?g \ h \ \j. J.arr j \ ?\ j \ h = ?\' j \ h" proof - assume E: "?f \ h = ?g \ h" have "\j. J.arr j \ ?\ j \ h = ?\' j \ h" proof - fix j assume j: "J.arr j" have "?\ j \ h = \a.cones_map ?f \a j \ h" using j f_map by fastforce also have "... = \a j \ ?f \ h" using j f_in_hom \a.map_def \a.cone_\ comp_assoc by auto also have "... = \a j \ ?g \ h" using j E by simp also have "... = \a.cones_map ?g \a j \ h" using j g_in_hom \a.map_def \a.cone_\ comp_assoc by auto also have "... = ?\' j \ h" using j g_map by force finally show "?\ j \ h = ?\' j \ h" by auto qed thus "\j. J.arr j \ ?\ j \ h = ?\' j \ h" by blast qed show "\j. J.arr j \ ?\ j \ h = ?\' j \ h \ ?f \ h = ?g \ h" proof - assume 1: "\j. J.arr j \ ?\ j \ h = ?\' j \ h" have 2: "\j. j \ Collect J.arr \ \a j \ ?f \ h = \a j \ ?g \ h" proof - fix j assume j: "j \ Collect J.arr" have "\a j \ ?f \ h = (\a j \ ?f) \ h" using comp_assoc by simp also have "... = ?\ j \ h" using ff j by force also have "... = ?\' j \ h" using 1 j by auto also have "... = (\a j \ ?g) \ h" using gg j by force also have "... = \a j \ ?g \ h" using comp_assoc by simp finally show "\a j \ ?f \ h = \a j \ ?g \ h" by auto qed show "C ?f h = C ?g h" proof - have "\j. Arr.arr j \ \\a j \ ?f \ h : dom h \ \a.map j\" using f_in_hom h \a_in_hom by (elim in_homE, auto) hence 3: "\!k. \k : dom h \ \a\ \ (\j. Arr.arr j \ \a j \ k = \a j \ ?f \ h)" using h \a \a.is_universal' [of "dom h" "\j. \a j \ ?f \ h"] \a.map_def ide_dom [of h] by blast have 4: "\P x x'. \!k. P k x \ P x x \ P x' x \ x' = x" by auto let ?P = "\ k x. \k : dom h \ \a\ \ (\j. j \ Collect J.arr \ \a j \ k = \a j \ x)" have "?P (?g \ h) (?g \ h)" using g_in_hom h by force moreover have "?P (?f \ h) (?g \ h)" using 2 f_in_hom g_in_hom h by force ultimately show ?thesis using 3 4 [of ?P "?f \ h" "?g \ h"] by auto qed qed qed have E': "\e. \e : dom e \ \o\ \ ?f \ e = ?g \ e \ (\j. J.arr j \ (D (J.cod j) \ \o (J.cod j) \ e) \ dom e = D j \ \o (J.dom j) \ e)" proof - have 1: "\e j. \e : dom e \ \o\ \ J.arr j \ ?\ j \ e = (D (J.cod j) \ \o (J.cod j) \ e) \ dom e" proof - fix e j assume e: "\e : dom e \ \o\" assume j: "J.arr j" have "\\o (J.cod j) \ e : dom e \ D (J.cod j)\" using e j \o_in_hom by auto thus "?\ j \ e = (D (J.cod j) \ \o (J.cod j) \ e) \ dom e" using j comp_arr_dom comp_cod_arr by (elim in_homE, auto) qed have 2: "\e j. \e : dom e \ \o\ \ J.arr j \ ?\' j \ e = D j \ \o (J.dom j) \ e" by (simp add: comp_assoc) show "\e. \e : dom e \ \o\ \ ?f \ e = ?g \ e \ (\j. J.arr j \ (D (J.cod j) \ \o (J.cod j) \ e) \ dom e = D j \ \o (J.dom j) \ e)" using 1 2 E by presburger qed text\ The composites of @{term e} with the projections from the product @{term \o} determine a limit cone @{term \} for @{term D}. The component of @{term \} at an object @{term j} of @{term[source=true] J} is the composite @{term "C (\o j) e"}. However, we need to extend @{term \} to all arrows @{term j} of @{term[source=true] J}, so the correct definition is @{term "\ j = C (D j) (C (\o (J.dom j)) e)"}. \ have e_in_hom: "\e : dom e \ \o\" using EQU.equalizes f_in_hom in_homI by (metis (no_types, lifting) seqE in_homE) have e_map: "C ?f e = C ?g e" using EQU.equalizes f_in_hom in_homI by fastforce interpret domE: constant_functor J C \dom e\ using e_in_hom by (unfold_locales, auto) let ?\ = "\j. if J.arr j then D j \ \o (J.dom j) \ e else null" have \: "\j. J.arr j \ \?\ j : dom e \ D (J.cod j)\" proof - fix j assume j: "J.arr j" show "\?\ j : dom e \ D (J.cod j)\" using j e_in_hom \o_in_hom [of "J.dom j"] by auto qed interpret \: cone J C D \dom e\ ?\ using \ comp_cod_arr e_in_hom e_map E' apply unfold_locales apply auto by (metis D.as_nat_trans.is_natural_1 comp_assoc) text\ If @{term \} is any cone over @{term D} then @{term \} restricts to a cone over @{term \o} for which the induced arrow to @{term \o} equalizes @{term f} and @{term g}. \ have R: "\a \. cone J C D a \ \ cone Obj.comp C \o.map a (\o.mkCone \) \ ?f \ \o.induced_arrow a (\o.mkCone \) = ?g \ \o.induced_arrow a (\o.mkCone \)" proof - fix a \ assume cone_\: "cone J C D a \" interpret \: cone J C D a \ using cone_\ by auto interpret A: constant_functor Obj.comp C a using \.ide_apex by (unfold_locales, auto) interpret \o: cone Obj.comp C \o.map a \\o.mkCone \\ using A.value_is_ide \o.map_def comp_cod_arr comp_arr_dom by (unfold_locales, auto) let ?e = "\o.induced_arrow a (\o.mkCone \)" have mkCone_\: "\o.mkCone \ \ \o.cones a" using \o.cone_axioms by auto have e: "\?e : a \ \o\" using mkCone_\ \o.induced_arrowI by simp have ee: "\j. J.ide j \ \o j \ ?e = \ j" proof - fix j assume j: "J.ide j" have "\o j \ ?e = \o.cones_map ?e \o j" using j e \o.cone_axioms by force also have "... = \o.mkCone \ j" using j mkCone_\ \o.induced_arrowI [of "\o.mkCone \" a] by fastforce also have "... = \ j" using j by simp finally show "\o j \ ?e = \ j" by auto qed have "\j. J.arr j \ (D (J.cod j) \ \o (J.cod j) \ ?e) \ dom ?e = D j \ \o (J.dom j) \ ?e" proof - fix j assume j: "J.arr j" have 1: "\\o (J.cod j) : \o \ D (J.cod j)\" using j \o_in_hom by simp have 2: "(D (J.cod j) \ \o (J.cod j) \ ?e) \ dom ?e = D (J.cod j) \ \o (J.cod j) \ ?e" proof - have "seq (D (J.cod j)) (\o (J.cod j))" using j 1 by auto moreover have "seq (\o (J.cod j)) ?e" using j e by fastforce ultimately show ?thesis using comp_arr_dom by auto qed also have 3: "... = \o (J.cod j) \ ?e" using j e 1 comp_cod_arr by (elim in_homE, auto) also have "... = D j \ \o (J.dom j) \ ?e" using j e ee 2 3 \.naturality \.A.map_simp \.ide_apex comp_cod_arr by auto finally show "(D (J.cod j) \ \o (J.cod j) \ ?e) \ dom ?e = D j \ \o (J.dom j) \ ?e" by auto qed hence "C ?f ?e = C ?g ?e" using E' \o.induced_arrowI \o.cone_axioms mem_Collect_eq by blast thus "cone Obj.comp C \o.map a (\o.mkCone \) \ C ?f ?e = C ?g ?e" using \o.cone_axioms by auto qed text\ Finally, show that @{term \} is a limit cone. \ interpret \: limit_cone J C D \dom e\ ?\ proof fix a \ assume cone_\: "cone J C D a \" interpret \: cone J C D a \ using cone_\ by auto interpret A: constant_functor Obj.comp C a using \.ide_apex by unfold_locales auto have cone_\o: "cone Obj.comp C \o.map a (\o.mkCone \)" using A.value_is_ide \o.map_def D.preserves_ide comp_cod_arr comp_arr_dom \.preserves_hom by unfold_locales auto show "\!h. \h : a \ dom e\ \ D.cones_map h ?\ = \" proof let ?e' = "\o.induced_arrow a (\o.mkCone \)" have e'_in_hom: "\?e' : a \ \o\" using cone_\ R \o.induced_arrowI by auto have e'_map: "?f \ ?e' = ?g \ ?e' \ \o.cones_map ?e' \o = \o.mkCone \" using cone_\ R \o.induced_arrowI [of "\o.mkCone \" a] by auto have equ: "PP.is_equalized_by ?e'" using e'_map e'_in_hom f_in_hom seqI' by blast let ?h = "EQU.induced_arrow a (PP.mkCone ?e')" have h_in_hom: "\?h : a \ dom e\" using EQU.induced_arrowI PP.cone_mkCone [of ?e'] e'_in_hom equ by fastforce have h_map: "PP.cones_map ?h (PP.mkCone e) = PP.mkCone ?e'" using EQU.induced_arrowI [of "PP.mkCone ?e'" a] PP.cone_mkCone [of ?e'] e'_in_hom equ by fastforce have 3: "D.cones_map ?h ?\ = \" proof fix j have "\J.arr j \ D.cones_map ?h ?\ j = \ j" using h_in_hom \.cone_axioms cone_\ \.is_extensional by force moreover have "J.arr j \ D.cones_map ?h ?\ j = \ j" proof - fix j assume j: "J.arr j" have 1: "\\o (J.dom j) \ e : dom e \ D (J.dom j)\" using j e_in_hom \o_in_hom [of "J.dom j"] by auto have "D.cones_map ?h ?\ j = ?\ j \ ?h" using h_in_hom j \.cone_axioms by auto also have "... = D j \ (\o (J.dom j) \ e) \ ?h" using j comp_assoc by simp also have "... = D j \ \ (J.dom j)" proof - have "(\o (J.dom j) \ e) \ ?h = \ (J.dom j)" proof - have "(\o (J.dom j) \ e) \ ?h = \o (J.dom j) \ e \ ?h" using j 1 e_in_hom h_in_hom \o arrI comp_assoc by auto also have "... = \o (J.dom j) \ ?e'" using equ e'_in_hom EQU.induced_arrowI' [of ?e'] by auto also have "... = \o.cones_map ?e' \o (J.dom j)" using j e'_in_hom \o.cone_axioms by auto also have "... = \ (J.dom j)" using j e'_map by simp finally show ?thesis by auto qed thus ?thesis by simp qed also have "... = \ j" using j \.is_natural_1 by simp finally show "D.cones_map ?h ?\ j = \ j" by auto qed ultimately show "D.cones_map ?h ?\ j = \ j" by auto qed show "\?h : a \ dom e\ \ D.cones_map ?h ?\ = \" using h_in_hom 3 by simp show "\h'. \h' : a \ dom e\ \ D.cones_map h' ?\ = \ \ h' = ?h" proof - fix h' assume h': "\h' : a \ dom e\ \ D.cones_map h' ?\ = \" have h'_in_hom: "\h' : a \ dom e\" using h' by simp have h'_map: "D.cones_map h' ?\ = \" using h' by simp show "h' = ?h" proof - have 1: "\e \ h' : a \ \o\ \ ?f \ e \ h' = ?g \ e \ h' \ \o.cones_map (C e h') \o = \o.mkCone \" proof - have 2: "\e \ h' : a \ \o\" using h'_in_hom e_in_hom by auto moreover have "?f \ e \ h' = ?g \ e \ h'" by (metis (no_types, lifting) EQU.equalizes comp_assoc) moreover have "\o.cones_map (e \ h') \o = \o.mkCone \" proof have "\o.cones_map (e \ h') \o = \o.cones_map h' (\o.cones_map e \o)" using \o.cone_axioms e_in_hom h'_in_hom \o.cones_map_comp [of e h'] by fastforce fix j have "\Obj.arr j \ \o.cones_map (e \ h') \o j = \o.mkCone \ j" using 2 e_in_hom h'_in_hom \o.cone_axioms by auto moreover have "Obj.arr j \ \o.cones_map (e \ h') \o j = \o.mkCone \ j" proof - assume j: "Obj.arr j" have "\o.cones_map (e \ h') \o j = \o j \ e \ h'" using 2 j \o.cone_axioms by auto also have "... = (\o j \ e) \ h'" using comp_assoc by auto also have "... = \o.mkCone ?\ j \ h'" using j e_in_hom \o_in_hom comp_ide_arr [of "D j" "\o j \ e"] by fastforce also have "... = \o.mkCone \ j" using j h' \.cone_axioms mem_Collect_eq by auto finally show "\o.cones_map (e \ h') \o j = \o.mkCone \ j" by auto qed ultimately show "\o.cones_map (e \ h') \o j = \o.mkCone \ j" by auto qed ultimately show ?thesis by auto qed have "\e \ h' : a \ \o\" using 1 by simp moreover have "e \ h' = ?e'" using 1 cone_\o e'_in_hom e'_map \o.is_universal \o by blast ultimately show "h' = ?h" using 1 h'_in_hom h'_map EQU.is_universal' [of "e \ h'"] EQU.induced_arrowI' [of ?e'] equ by (elim in_homE) auto qed qed qed qed have "limit_cone J C D (dom e) ?\" .. thus "\a \. limit_cone J C D a \" by auto qed thus "\D. diagram J C D \ (\a \. limit_cone J C D a \)" by blast qed end section "Limits in a Set Category" text\ In this section, we consider the special case of limits in a set category. \ locale diagram_in_set_category = J: category J + - S: set_category S \ + + S: set_category S is_set + diagram J S D for J :: "'j comp" (infixr "\\<^sub>J" 55) and S :: "'s comp" (infixr "\" 55) - and \ :: "'a rel" + and is_set :: "'s set \ bool" and D :: "'j \ 's" begin notation S.in_hom ("\_ : _ \ _\") text\ An object @{term a} of a set category @{term[source=true] S} is a limit of a diagram in @{term[source=true] S} if and only if there is a bijection between the set @{term "S.hom S.unity a"} of points of @{term a} and the set of cones over the diagram that have apex @{term S.unity}. \ lemma limits_are_sets_of_cones: shows "has_as_limit a \ S.ide a \ (\\. bij_betw \ (S.hom S.unity a) (cones S.unity))" proof text\ If \has_limit a\, then by the universal property of the limit cone, composition in @{term[source=true] S} yields a bijection between @{term "S.hom S.unity a"} and @{term "cones S.unity"}. \ assume a: "has_as_limit a" hence "S.ide a" using limit_cone_def cone.ide_apex by metis from a obtain \ where \: "limit_cone a \" by auto interpret \: limit_cone J S D a \ using \ by auto have "bij_betw (\f. cones_map f \) (S.hom S.unity a) (cones S.unity)" using \.bij_betw_hom_and_cones S.ide_unity by simp thus "S.ide a \ (\\. bij_betw \ (S.hom S.unity a) (cones S.unity))" using \S.ide a\ by blast next text\ Conversely, an arbitrary bijection @{term \} between @{term "S.hom S.unity a"} and cones unity extends pointwise to a natural bijection @{term "\ a'"} between @{term "S.hom a' a"} and @{term "cones a'"}, showing that @{term a} is a limit. In more detail, the hypotheses give us a correspondence between points of @{term a} and cones with apex @{term "S.unity"}. We extend this to a correspondence between functions to @{term a} and general cones, with each arrow from @{term a'} to @{term a} determining a cone with apex @{term a'}. If @{term "f \ hom a' a"} then composition with @{term f} takes each point @{term y} of @{term a'} to the point @{term "S f y"} of @{term a}. To this we may apply the given bijection @{term \} to obtain @{term "\ (S f y) \ cones S.unity"}. The component @{term "\ (S f y) j"} at @{term j} of this cone is a point of @{term "S.cod (D j)"}. Thus, @{term "f \ hom a' a"} determines a cone @{term \f} with apex @{term a'} whose component at @{term j} is the unique arrow @{term "\f j"} of @{term[source=true] S} such that @{term "\f j \ hom a' (cod (D j))"} and @{term "S (\f j) y = \ (S f y) j"} for all points @{term y} of @{term a'}. The cone @{term \a} corresponding to @{term "a \ S.hom a a"} is then a limit cone. \ assume a: "S.ide a \ (\\. bij_betw \ (S.hom S.unity a) (cones S.unity))" hence ide_a: "S.ide a" by auto show "has_as_limit a" proof - from a obtain \ where \: "bij_betw \ (S.hom S.unity a) (cones S.unity)" by blast have X: "\f j y. \ \f : S.dom f \ a\; J.arr j; \y : S.unity \ S.dom f\ \ \ \\ (f \ y) j : S.unity \ S.cod (D j)\" proof - fix f j y assume f: "\f : S.dom f \ a\" and j: "J.arr j" and y: "\y : S.unity \ S.dom f\" interpret \: cone J S D S.unity \\ (S f y)\ using f y \ bij_betw_imp_funcset funcset_mem by blast show "\\ (f \ y) j : S.unity \ S.cod (D j)\" using j by auto qed text\ We want to define the component @{term "\j \ S.hom (S.dom f) (S.cod (D j))"} at @{term j} of a cone by specifying how it acts by composition on points @{term "y \ S.hom S.unity (S.dom f)"}. We can do this because @{term[source=true] S} is a set category. \ let ?P = "\f j \j. \\j : S.dom f \ S.cod (D j)\ \ (\y. \y : S.unity \ S.dom f\ \ \j \ y = \ (f \ y) j)" let ?\ = "\f j. if J.arr j then (THE \j. ?P f j \j) else S.null" have \: "\f j. \ \f : S.dom f \ a\; J.arr j \ \ ?P f j (?\ f j)" proof - fix b f j assume f: "\f : S.dom f \ a\" and j: "J.arr j" interpret B: constant_functor J S \S.dom f\ using f by (unfold_locales) auto have "(\y. \ (f \ y) j) \ S.hom S.unity (S.dom f) \ S.hom S.unity (S.cod (D j))" using f j X Pi_I' by simp hence "\!\j. ?P f j \j" using f j S.fun_complete' by (elim S.in_homE) auto thus "?P f j (?\ f j)" using j theI' [of "?P f j"] by simp qed text\ The arrows @{term "\ f j"} are in fact the components of a cone with apex @{term "S.dom f"}. \ have cone: "\f. \f : S.dom f \ a\ \ cone (S.dom f) (?\ f)" proof - fix f assume f: "\f : S.dom f \ a\" interpret B: constant_functor J S \S.dom f\ using f by unfold_locales auto show "cone (S.dom f) (?\ f)" proof show "\j. \J.arr j \ ?\ f j = S.null" by simp fix j assume j: "J.arr j" have 0: "\?\ f j : S.dom f \ S.cod (D j)\" using f j \ by simp show "S.dom (?\ f j) = B.map (J.dom j)" using f j \ by auto show "S.cod (?\ f j) = D (J.cod j)" using f j \ by auto have par2: "S.par (?\ f (J.cod j) \ B.map j) (?\ f j)" using f j 0 \ [of f "J.cod j"] by (elim S.in_homE, auto) have nat: "\y. \y : S.unity \ S.dom f\ \ (D j \ ?\ f (J.dom j)) \ y = ?\ f j \ y \ (?\ f (J.cod j) \ B.map j) \ y = ?\ f j \ y" proof - fix y assume y: "\y : S.unity \ S.dom f\" show "(D j \ ?\ f (J.dom j)) \ y = ?\ f j \ y \ (?\ f (J.cod j) \ B.map j) \ y = ?\ f j \ y" proof have 1: "\ (f \ y) \ cones S.unity" using f y \ bij_betw_imp_funcset PiE S.seqI S.cod_comp S.dom_comp mem_Collect_eq by fastforce interpret \: cone J S D S.unity \\ (f \ y)\ using 1 by simp show "(D j \ ?\ f (J.dom j)) \ y = ?\ f j \ y" using J.arr_dom S.comp_assoc \ \.is_natural_1 f j y by presburger have "(?\ f (J.cod j) \ B.map j) \ y = ?\ f (J.cod j) \ y" using j B.map_simp par2 B.value_is_ide S.comp_arr_ide by (metis (no_types, lifting)) also have "... = \ (f \ y) (J.cod j)" using f y \ \.is_extensional by simp also have "... = \ (f \ y) j" using j \.is_natural_2 by (metis J.arr_cod \.A.map_simp J.cod_cod) also have "... = ?\ f j \ y" using f y \ \.is_extensional by simp finally show "(?\ f (J.cod j) \ B.map j) \ y = ?\ f j \ y" by auto qed qed show "D j \ ?\ f (J.dom j) = ?\ f j" proof - have "S.par (D j \ ?\ f (J.dom j)) (?\ f j)" using f j 0 \ [of f "J.dom j"] by (elim S.in_homE, auto) thus ?thesis using nat 0 apply (intro S.arr_eqI' [of "D j \ ?\ f (J.dom j)" "?\ f j"]) apply force by auto qed show "?\ f (J.cod j) \ B.map j = ?\ f j" using par2 nat 0 f j \ apply (intro S.arr_eqI' [of "?\ f (J.cod j) \ B.map j" "?\ f j"]) apply force by (metis (no_types, lifting) S.in_homE) qed qed interpret \a: cone J S D a \?\ a\ using a cone [of a] by fastforce text\ Finally, show that \\ a\ is a limit cone. \ interpret \a: limit_cone J S D a \?\ a\ proof fix a' \' assume cone_\': "cone a' \'" interpret \': cone J S D a' \' using cone_\' by auto show "\!f. \f : a' \ a\ \ cones_map f (?\ a) = \'" proof let ?\ = "inv_into (S.hom S.unity a) \" have \: "?\ \ cones S.unity \ S.hom S.unity a" using \ bij_betw_inv_into bij_betwE by blast let ?P = "\f. \f : a' \ a\ \ (\y. y \ S.hom S.unity a' \ f \ y = ?\ (cones_map y \'))" have 1: "\!f. ?P f" proof - have "(\y. ?\ (cones_map y \')) \ S.hom S.unity a' \ S.hom S.unity a" proof fix x assume "x \ S.hom S.unity a'" hence "\x : S.unity \ a'\" by simp hence "cones_map x \ cones a' \ cones S.unity" using cones_map_mapsto [of x] by (elim S.in_homE) auto hence "cones_map x \' \ cones S.unity" using cone_\' by blast thus "?\ (cones_map x \') \ S.hom S.unity a" using \ by auto qed thus ?thesis using S.fun_complete' a \'.ide_apex by simp qed let ?f = "THE f. ?P f" have f: "?P ?f" using 1 theI' [of ?P] by simp have f_in_hom: "\?f : a' \ a\" using f by simp have f_map: "cones_map ?f (?\ a) = \'" proof - have 1: "cone a' (cones_map ?f (?\ a))" proof - have "cones_map ?f \ cones a \ cones a'" using f_in_hom cones_map_mapsto [of ?f] by (elim S.in_homE) auto hence "cones_map ?f (?\ a) \ cones a'" using \a.cone_axioms by blast thus ?thesis by simp qed interpret f\a: cone J S D a' \cones_map ?f (?\ a)\ using 1 by simp show ?thesis proof fix j have "\J.arr j \ cones_map ?f (?\ a) j = \' j" using 1 \'.is_extensional f\a.is_extensional by presburger moreover have "J.arr j \ cones_map ?f (?\ a) j = \' j" proof - assume j: "J.arr j" show "cones_map ?f (?\ a) j = \' j" proof (intro S.arr_eqI' [of "cones_map ?f (?\ a) j" "\' j"]) show par: "S.par (cones_map ?f (?\ a) j) (\' j)" using j \'.preserves_cod \'.preserves_dom \'.preserves_reflects_arr f\a.preserves_cod f\a.preserves_dom f\a.preserves_reflects_arr by presburger fix y assume "\y : S.unity \ S.dom (cones_map ?f (?\ a) j)\" hence y: "\y : S.unity \ a'\" using j f\a.preserves_dom by simp have 1: "\?\ a j : a \ D (J.cod j)\" using j \a.preserves_hom by force have 2: "\?f \ y : S.unity \ a\" using f_in_hom y by blast have "cones_map ?f (?\ a) j \ y = (?\ a j \ ?f) \ y" proof - have "S.cod ?f = a" using f_in_hom by blast thus ?thesis using j \a.cone_axioms by simp qed also have "... = ?\ a j \ ?f \ y" using 1 j y f_in_hom S.comp_assoc S.seqI' by blast also have "... = \ (a \ ?f \ y) j" using 1 2 ide_a f j y \ [of a] by (simp add: S.ide_in_hom) also have "... = \ (?f \ y) j" using a 2 y S.comp_cod_arr by (elim S.in_homE, auto) also have "... = \ (?\ (cones_map y \')) j" using j y f by simp also have "... = cones_map y \' j" proof - have "cones_map y \' \ cones S.unity" using cone_\' y cones_map_mapsto by force hence "\ (?\ (cones_map y \')) = cones_map y \'" using \ bij_betw_inv_into_right [of \] by simp thus ?thesis by auto qed also have "... = \' j \ y" using cone_\' j y by auto finally show "cones_map ?f (?\ a) j \ y = \' j \ y" by auto qed qed ultimately show "cones_map ?f (?\ a) j = \' j" by blast qed qed show "\?f : a' \ a\ \ cones_map ?f (?\ a) = \'" using f_in_hom f_map by simp show "\f'. \f' : a' \ a\ \ cones_map f' (?\ a) = \' \ f' = ?f" proof - fix f' assume f': "\f' : a' \ a\ \ cones_map f' (?\ a) = \'" have f'_in_hom: "\f' : a' \ a\" using f' by simp have f'_map: "cones_map f' (?\ a) = \'" using f' by simp show "f' = ?f" proof (intro S.arr_eqI' [of f' ?f]) show "S.par f' ?f" using f_in_hom f'_in_hom by (elim S.in_homE, auto) show "\y'. \y' : S.unity \ S.dom f'\ \ f' \ y' = ?f \ y'" proof - fix y' assume y': "\y' : S.unity \ S.dom f'\" have 0: "\ (f' \ y') = cones_map y' \'" proof fix j have 1: "\f' \ y' : S.unity \ a\" using f'_in_hom y' by auto hence 2: "\ (f' \ y') \ cones S.unity" using \ bij_betw_imp_funcset [of \ "S.hom S.unity a" "cones S.unity"] by auto interpret \'': cone J S D S.unity \\ (f' \ y')\ using 2 by auto have "\J.arr j \ \ (f' \ y') j = cones_map y' \' j" using f' y' cone_\' \''.is_extensional mem_Collect_eq restrict_apply by (elim S.in_homE, auto) moreover have "J.arr j \ \ (f' \ y') j = cones_map y' \' j" proof - assume j: "J.arr j" have 3: "\?\ a j : a \ D (J.cod j)\" using j \a.preserves_hom by force have "\ (f' \ y') j = \ (a \ f' \ y') j" using a f' y' j S.comp_cod_arr by (elim S.in_homE, auto) also have "... = ?\ a j \ f' \ y'" using 1 3 \ [of a] a f' y' j by fastforce also have "... = (?\ a j \ f') \ y'" using S.comp_assoc by simp also have "... = cones_map f' (?\ a) j \ y'" using f' y' j \a.cone_axioms by auto also have "... = \' j \ y'" using f' by blast also have "... = cones_map y' \' j" using y' j cone_\' f' mem_Collect_eq restrict_apply by force finally show "\ (f' \ y') j = cones_map y' \' j" by auto qed ultimately show "\ (f' \ y') j = cones_map y' \' j" by auto qed hence "f' \ y' = ?\ (cones_map y' \')" using \ f'_in_hom y' S.comp_in_homI bij_betw_inv_into_left [of \ "S.hom S.unity a" "cones S.unity" "f' \ y'"] by (elim S.in_homE, auto) moreover have "?f \ y' = ?\ (cones_map y' \')" using \ 0 1 f f_in_hom f'_in_hom y' S.comp_in_homI bij_betw_inv_into_left [of \ "S.hom S.unity a" "cones S.unity" "?f \ y'"] by (elim S.in_homE, auto) ultimately show "f' \ y' = ?f \ y'" by auto qed qed qed qed qed have "limit_cone a (?\ a)" .. thus ?thesis by auto qed qed end locale diagram_in_replete_set_category = J: category J + S: replete_set_category S + diagram J S D for J :: "'j comp" (infixr "\\<^sub>J" 55) and S :: "'s comp" (infixr "\" 55) and D :: "'j \ 's" begin - sublocale diagram_in_set_category J S \cardSuc (cmax (card_of (UNIV :: 's set)) natLeq)\ D + sublocale diagram_in_set_category J S S.setp D .. end context set_category begin text\ A set category has an equalizer for any parallel pair of arrows. \ lemma has_equalizers: shows "has_equalizers" proof (unfold has_equalizers_def) have "\f0 f1. par f0 f1 \ \e. has_as_equalizer f0 f1 e" proof - fix f0 f1 assume par: "par f0 f1" interpret J: parallel_pair . interpret PP: parallel_pair_diagram S f0 f1 using par by unfold_locales auto - interpret PP: diagram_in_set_category J.comp S \ PP.map .. + interpret PP: diagram_in_set_category J.comp S setp PP.map .. text\ Let @{term a} be the object corresponding to the set of all images of equalizing points of @{term "dom f0"}, and let @{term e} be the inclusion of @{term a} in @{term "dom f0"}. \ let ?a = "mkIde (img ` {e. e \ hom unity (dom f0) \ f0 \ e = f1 \ e})" have 0: "{e. e \ hom unity (dom f0) \ f0 \ e = f1 \ e} \ hom unity (dom f0)" by auto hence 1: "img ` {e. e \ hom unity (dom f0) \ f0 \ e = f1 \ e} \ Univ" using img_point_in_Univ by auto - have 2: "|img ` {e. e \ hom unity (dom f0) \ f0 \ e = f1 \ e}| " + have 2: "setp (img ` {e. e \ hom unity (dom f0) \ f0 \ e = f1 \ e})" proof - - have "|hom unity (dom f0)| =o |Dom f0|" - using par bij_betw_points_and_set [of "dom f0"] - by (simp add: card_of_ordIsoI) - moreover have "|Dom f0| " - using par set_card by simp - ultimately have "|hom unity (dom f0)| " - using ordIso_ordLess_trans by blast - moreover have - "|{e. e \ hom unity (dom f0) \ f0 \ e = f1 \ e}| \o |hom unity (dom f0)|" - using 0 by simp - ultimately have "|{e. e \ hom unity (dom f0) \ f0 \ e = f1 \ e}| " - using ordLeq_ordLess_trans by blast - thus ?thesis - using card_of_image ordLeq_ordLess_trans by blast + have "setp (img ` hom unity (dom f0))" + using ide_dom par setp_img_points by blast + moreover have "img ` {e. e \ hom unity (dom f0) \ f0 \ e = f1 \ e} \ + img ` hom unity (dom f0)" + by blast + ultimately show ?thesis + by (meson setp_respects_subset) qed have ide_a: "ide ?a" using 1 2 ide_mkIde by auto have set_a: "set ?a = img ` {e. e \ hom unity (dom f0) \ f0 \ e = f1 \ e}" using 1 2 set_mkIde by simp have incl_in_a: "incl_in ?a (dom f0)" proof - have "ide (dom f0)" using PP.is_parallel by simp moreover have "set ?a \ set (dom f0)" using img_point_elem_set set_a by fastforce ultimately show ?thesis using incl_in_def \ide ?a\ by simp qed text\ Then @{term "set a"} is in bijective correspondence with @{term "PP.cones unity"}. \ let ?\ = "\t. PP.mkCone (mkPoint (dom f0) t)" let ?\ = "\\. img (\ (J.Zero))" have bij: "bij_betw ?\ (set ?a) (PP.cones unity)" proof (intro bij_betwI) show "?\ \ set ?a \ PP.cones unity" proof fix t assume t: "t \ set ?a" hence 1: "t \ img ` {e. e \ hom unity (dom f0) \ f0 \ e = f1 \ e}" using set_a by blast then have 2: "mkPoint (dom f0) t \ hom unity (dom f0)" using mkPoint_in_hom imageE mem_Collect_eq mkPoint_img(2) by auto with 1 have 3: "mkPoint (dom f0) t \ {e. e \ hom unity (dom f0) \ f0 \ e = f1 \ e}" using mkPoint_img(2) by auto then have "PP.is_equalized_by (mkPoint (dom f0) t)" using CollectD par by fastforce thus "PP.mkCone (mkPoint (dom f0) t) \ PP.cones unity" using 2 PP.cone_mkCone [of "mkPoint (dom f0) t"] by auto qed show "?\ \ PP.cones unity \ set ?a" proof fix \ assume \: "\ \ PP.cones unity" interpret \: cone J.comp S PP.map unity \ using \ by auto have "\ (J.Zero) \ hom unity (dom f0) \ f0 \ \ (J.Zero) = f1 \ \ (J.Zero)" using \ PP.map_def PP.is_equalized_by_cone J.arr_char by auto hence "img (\ (J.Zero)) \ set ?a" using set_a by simp thus "?\ \ \ set ?a" by blast qed show "\t. t \ set ?a \ ?\ (?\ t) = t" using set_a J.arr_char PP.mkCone_def imageE mem_Collect_eq mkPoint_img(2) by auto show "\\. \ \ PP.cones unity \ ?\ (?\ \) = \" proof - fix \ assume \: "\ \ PP.cones unity" interpret \: cone J.comp S PP.map unity \ using \ by auto have 1: "\ (J.Zero) \ hom unity (dom f0) \ f0 \ \ (J.Zero) = f1 \ \ (J.Zero)" using \ PP.map_def PP.is_equalized_by_cone J.arr_char by auto hence "img (\ (J.Zero)) \ set ?a" using set_a by simp hence "img (\ (J.Zero)) \ set (dom f0)" using incl_in_a incl_in_def by auto hence "mkPoint (dom f0) (img (\ J.Zero)) = \ J.Zero" using 1 mkPoint_img(2) by blast hence "?\ (?\ \) = PP.mkCone (\ J.Zero)" by simp also have "... = \" using \ PP.mkCone_cone by simp finally show "?\ (?\ \) = \" by auto qed qed text\ It follows that @{term a} is a limit of \PP\, and that the limit cone gives an equalizer of @{term f0} and @{term f1}. \ have "PP.has_as_limit ?a" proof - have "\\. bij_betw \ (hom unity ?a) (set ?a)" using bij_betw_points_and_set ide_a by auto from this obtain \ where \: "bij_betw \ (hom unity ?a) (set ?a)" by blast have "bij_betw (?\ o \) (hom unity ?a) (PP.cones unity)" using bij \ bij_betw_comp_iff by blast hence "\\. bij_betw \ (hom unity ?a) (PP.cones unity)" by auto thus ?thesis using ide_a PP.limits_are_sets_of_cones by simp qed from this obtain \ where \: "limit_cone J.comp S PP.map ?a \" by auto have "PP.has_as_equalizer (\ J.Zero)" proof - interpret \: limit_cone J.comp S PP.map ?a \ using \ by auto have "PP.mkCone (\ (J.Zero)) = \" using \ PP.mkCone_cone \.cone_axioms by simp moreover have "dom (\ (J.Zero)) = ?a" using J.ide_char \.preserves_hom \.A.map_def by simp ultimately show ?thesis using \ by simp qed thus "\e. has_as_equalizer f0 f1 e" using par has_as_equalizer_def by auto qed thus "\f0 f1. par f0 f1 \ (\e. has_as_equalizer f0 f1 e)" by auto qed end sublocale set_category \ category_with_equalizers S apply unfold_locales using has_equalizers by auto context set_category begin text\ The aim of the next results is to characterize the conditions under which a set category has products. In a traditional development of category theory, one shows that the category \textbf{Set} of \emph{all} sets has all small (\emph{i.e.}~set-indexed) products. In the present context we do not have a category of \emph{all} sets, but rather only a category of all sets with elements at a particular type. Clearly, we cannot expect such a category to have products indexed by arbitrarily large sets. The existence of @{term I}-indexed products in a set category @{term[source=true] S} implies that the universe \S.Univ\ of @{term[source=true] S} must be large enough to admit the formation of @{term I}-tuples of its elements. Conversely, for a set category @{term[source=true] S} the ability to form @{term I}-tuples in @{term Univ} implies that @{term[source=true] S} has @{term I}-indexed products. Below we make this precise by defining the notion of when a set category @{term[source=true] S} ``admits @{term I}-indexed tupling'' and we show that @{term[source=true] S} has @{term I}-indexed products if and only if it admits @{term I}-indexed tupling. The definition of ``@{term[source=true] S} admits @{term I}-indexed tupling'' says that there is an injective map, from the space of extensional functions from @{term I} to @{term Univ}, to @{term Univ}. However for a convenient statement and proof of the desired result, the definition of extensional function from theory @{theory "HOL-Library.FuncSet"} needs to be modified. The theory @{theory "HOL-Library.FuncSet"} uses the definite, but arbitrarily chosen value @{term undefined} as the value to be assumed by an extensional function outside of its domain. In the context of the \set_category\, though, it is more natural to use \S.unity\, which is guaranteed to be an element of the universe of @{term[source=true] S}, for this purpose. Doing things that way makes it simpler to establish a bijective correspondence between cones over @{term D} with apex @{term unity} and the set of extensional functions @{term d} that map each arrow @{term j} of @{term J} to an element @{term "d j"} of @{term "set (D j)"}. Possibly it makes sense to go back and make this change in \set_category\, but that would mean completely abandoning @{theory "HOL-Library.FuncSet"} and essentially introducing a duplicate version for use with \set_category\. As a compromise, what I have done here is to locally redefine the few notions from @{theory "HOL-Library.FuncSet"} that I need in order to prove the next set of results. The redefined notions are primed to avoid confusion with the original versions. \ definition extensional' where "extensional' A \ {f. \x. x \ A \ f x = unity}" abbreviation PiE' where "PiE' A B \ Pi A B \ extensional' A" abbreviation restrict' where "restrict' f A \ \x. if x \ A then f x else unity" lemma extensional'I [intro]: assumes "\x. x \ A \ f x = unity" shows "f \ extensional' A" using assms extensional'_def by auto lemma extensional'_arb: assumes "f \ extensional' A" and "x \ A" shows "f x = unity" using assms extensional'_def by fast lemma extensional'_monotone: assumes "A \ B" shows "extensional' A \ extensional' B" using assms extensional'_arb by fastforce lemma PiE'_mono: "(\x. x \ A \ B x \ C x) \ PiE' A B \ PiE' A C" by auto end - locale discrete_diagram_in_replete_set_category = - S: replete_set_category S + + locale discrete_diagram_in_set_category = + S: set_category S \ + discrete_diagram J S D + - diagram_in_replete_set_category J S D + diagram_in_set_category J S \ D for J :: "'j comp" (infixr "\\<^sub>J" 55) and S :: "'s comp" (infixr "\" 55) + and \ :: "'s set \ bool" and D :: "'j \ 's" begin text\ For @{term D} a discrete diagram in a set category, there is a bijective correspondence between cones over @{term D} with apex unity and the set of extensional functions @{term d} that map each arrow @{term j} of @{term[source=true] J} to an element of @{term "S.set (D j)"}. \ abbreviation I where "I \ Collect J.arr" definition funToCone where "funToCone F \ \j. if J.arr j then S.mkPoint (D j) (F j) else S.null" definition coneToFun where "coneToFun \ \ \j. if J.arr j then S.img (\ j) else S.unity" lemma funToCone_mapsto: shows "funToCone \ S.PiE' I (S.set o D) \ cones S.unity" proof fix F assume F: "F \ S.PiE' I (S.set o D)" interpret U: constant_functor J S S.unity apply unfold_locales using S.ide_unity by auto - have 1: "S.ide (S.mkIde S.Univ)" - using S.ide_mkIde by simp have "cone S.unity (funToCone F)" proof show "\j. \J.arr j \ funToCone F j = S.null" using funToCone_def by simp fix j assume j: "J.arr j" have "funToCone F j = S.mkPoint (D j) (F j)" using j funToCone_def by simp moreover have "... \ S.hom S.unity (D j)" using F j is_discrete S.img_mkPoint(1) [of "D j"] by force ultimately have 2: "funToCone F j \ S.hom S.unity (D j)" by auto show 3: "S.dom (funToCone F j) = U.map (J.dom j)" using 2 j U.map_simp by auto show 4: "S.cod (funToCone F j) = D (J.cod j)" using 2 j is_discrete by auto show "D j \ funToCone F (J.dom j) = funToCone F j" using 2 j is_discrete S.comp_cod_arr by auto show "funToCone F (J.cod j) \ (U.map j) = funToCone F j" using 3 j is_discrete U.map_simp S.arr_dom_iff_arr S.comp_arr_dom U.preserves_arr by (metis J.ide_char) qed thus "funToCone F \ cones S.unity" by auto qed lemma coneToFun_mapsto: shows "coneToFun \ cones S.unity \ S.PiE' I (S.set o D)" proof fix \ assume \: "\ \ cones S.unity" interpret \: cone J S D S.unity \ using \ by auto show "coneToFun \ \ S.PiE' I (S.set o D)" proof show "coneToFun \ \ Pi I (S.set o D)" using S.mkPoint_img(1) coneToFun_def is_discrete \.component_in_hom by (simp add: S.img_point_elem_set restrict_apply') show "coneToFun \ \ S.extensional' I" by (metis S.extensional'I coneToFun_def mem_Collect_eq) qed qed lemma funToCone_coneToFun: assumes "\ \ cones S.unity" shows "funToCone (coneToFun \) = \" proof interpret \: cone J S D S.unity \ using assms by auto fix j have "\J.arr j \ funToCone (coneToFun \) j = \ j" using funToCone_def \.is_extensional by simp moreover have "J.arr j \ funToCone (coneToFun \) j = \ j" using funToCone_def coneToFun_def S.mkPoint_img(2) is_discrete \.component_in_hom by auto ultimately show "funToCone (coneToFun \) j = \ j" by blast qed lemma coneToFun_funToCone: assumes "F \ S.PiE' I (S.set o D)" shows "coneToFun (funToCone F) = F" proof fix i have "i \ I \ coneToFun (funToCone F) i = F i" using assms coneToFun_def S.extensional'_arb [of F I i] by auto moreover have "i \ I \ coneToFun (funToCone F) i = F i" proof - assume i: "i \ I" have "coneToFun (funToCone F) i = S.img (funToCone F i)" using i coneToFun_def by simp also have "... = S.img (S.mkPoint (D i) (F i))" using i funToCone_def by auto also have "... = F i" using assms i is_discrete S.img_mkPoint(2) by force finally show "coneToFun (funToCone F) i = F i" by auto qed ultimately show "coneToFun (funToCone F) i = F i" by auto qed lemma bij_coneToFun: shows "bij_betw coneToFun (cones S.unity) (S.PiE' I (S.set o D))" using coneToFun_mapsto funToCone_mapsto funToCone_coneToFun coneToFun_funToCone bij_betwI by blast lemma bij_funToCone: shows "bij_betw funToCone (S.PiE' I (S.set o D)) (cones S.unity)" using coneToFun_mapsto funToCone_mapsto funToCone_coneToFun coneToFun_funToCone bij_betwI by blast end - context replete_set_category + context set_category begin text\ A set category admits @{term I}-indexed tupling if there is an injective map that takes each extensional function from @{term I} to @{term Univ} to an element of @{term Univ}. \ definition admits_tupling where "admits_tupling I \ \\. \ \ PiE' I (\_. Univ) \ Univ \ inj_on \ (PiE' I (\_. Univ))" lemma admits_tupling_monotone: assumes "admits_tupling I" and "I' \ I" shows "admits_tupling I'" proof - from assms(1) obtain \ where \: "\ \ PiE' I (\_. Univ) \ Univ \ inj_on \ (PiE' I (\_. Univ))" using admits_tupling_def by metis have "\ \ PiE' I' (\_. Univ) \ Univ" proof fix f assume f: "f \ PiE' I' (\_. Univ)" have "f \ PiE' I (\_. Univ)" using assms(2) f extensional'_def [of I'] terminal_unity extensional'_monotone by auto thus "\ f \ Univ" using \ by auto qed moreover have "inj_on \ (PiE' I' (\_. Univ))" proof - have 1: "\F A A'. inj_on F A \ A' \ A \ inj_on F A'" using subset_inj_on by blast moreover have "PiE' I' (\_. Univ) \ PiE' I (\_. Univ)" using assms(2) extensional'_def [of I'] terminal_unity by auto ultimately show ?thesis using \ assms(2) by blast qed ultimately show ?thesis using admits_tupling_def by metis qed + lemma admits_tupling_respects_bij: + assumes "admits_tupling I" and "bij_betw \ I I'" + shows "admits_tupling I'" + proof - + obtain \ where \: "\ \ (I \ Univ) \ extensional' I \ Univ \ + inj_on \ ((I \ Univ) \ extensional' I)" + using assms(1) admits_tupling_def by metis + have inv: "bij_betw (inv_into I \) I' I" + using assms(2) bij_betw_inv_into by blast + let ?C = "\f x. if x \ I then f (\ x) else unity" + let ?\' = "\f. \ (?C f)" + have 1: "\f. f \ (I' \ Univ) \ extensional' I' \ ?C f \ (I \ Univ) \ extensional' I" + using assms bij_betw_apply by fastforce + have "?\' \ (I' \ Univ) \ extensional' I' \ Univ \ + inj_on ?\' ((I' \ Univ) \ extensional' I')" + proof + show "(\f. \ (?C f)) \ (I' \ Univ) \ extensional' I' \ Univ" + using 1 \ by blast + show "inj_on ?\' ((I' \ Univ) \ extensional' I')" + proof + fix f g + assume f: "f \ (I' \ Univ) \ extensional' I'" + assume g: "g \ (I' \ Univ) \ extensional' I'" + assume eq: "?\' f = ?\' g" + have f': "?C f \ (I \ Univ) \ extensional' I" + using f 1 by simp + have g': "?C g \ (I \ Univ) \ extensional' I" + using g 1 by simp + have 2: "?C f = ?C g" + using f' g' \ eq by (simp add: inj_on_def) + show "f = g" + proof + fix x + show "f x = g x" + proof (cases "x \ I'") + show "x \ I' \ ?thesis" + using f g + by (metis (no_types, opaque_lifting) "2" assms(2) bij_betw_apply + bij_betw_inv_into_right inv) + show "x \ I' \ ?thesis" + using f g by (metis IntD2 extensional'_arb) + qed + qed + qed + qed + thus ?thesis + using admits_tupling_def by blast + qed + + end + + context replete_set_category + begin + lemma has_products_iff_admits_tupling: fixes I :: "'i set" shows "has_products I \ I \ UNIV \ admits_tupling I" proof text\ If @{term[source=true] S} has @{term I}-indexed products, then for every @{term I}-indexed discrete diagram @{term D} in @{term[source=true] S} there is an object @{term \D} of @{term[source=true] S} whose points are in bijective correspondence with the set of cones over @{term D} with apex @{term unity}. In particular this is true for the diagram @{term D} that assigns to each element of @{term I} the ``universal object'' @{term "mkIde Univ"}. \ assume has_products: "has_products I" have I: "I \ UNIV" using has_products has_products_def by auto interpret J: discrete_category I \SOME x. x \ I\ using I someI_ex [of "\x. x \ I"] by (unfold_locales, auto) let ?D = "\i. mkIde Univ" interpret D: discrete_diagram_from_map I S ?D \SOME j. j \ I\ using J.not_arr_null J.arr_char ide_mkIde by (unfold_locales, auto) - interpret D: discrete_diagram_in_replete_set_category J.comp S D.map .. + interpret D: discrete_diagram_in_set_category J.comp S \\A. A \ Univ\ D.map .. have "discrete_diagram J.comp S D.map" .. from this obtain \D \ where \: "product_cone J.comp S D.map \D \" using has_products has_products_def [of I] ex_productE [of "J.comp" D.map] D.diagram_axioms by blast interpret \: product_cone J.comp S D.map \D \ using \ by auto have "D.has_as_limit \D" using \.limit_cone_axioms by auto hence \D: "ide \D \ (\\. bij_betw \ (hom unity \D) (D.cones unity))" using D.limits_are_sets_of_cones by simp from this obtain \ where \: "bij_betw \ (hom unity \D) (D.cones unity)" by blast have \': "inv_into (hom unity \D) \ \ D.cones unity \ hom unity \D \ inj_on (inv_into (hom unity \D) \) (D.cones unity)" using \ bij_betw_inv_into bij_betw_imp_inj_on bij_betw_imp_funcset by blast let ?\ = "img o (inv_into (hom unity \D) \) o D.funToCone" have 1: "D.funToCone \ PiE' I (set o D.map) \ D.cones unity" using D.funToCone_mapsto extensional'_def [of I] by auto have 2: "inv_into (hom unity \D) \ \ D.cones unity \ hom unity \D" using \' by auto have 3: "img \ hom unity \D \ Univ" using img_point_in_Univ by blast have 4: "inj_on D.funToCone (PiE' I (set o D.map))" proof - have "D.I = I" by auto thus ?thesis using D.bij_funToCone bij_betw_imp_inj_on by auto qed have 5: "inj_on (inv_into (hom unity \D) \) (D.cones unity)" using \' by auto have 6: "inj_on img (hom unity \D)" using \D bij_betw_points_and_set bij_betw_imp_inj_on [of img "hom unity \D" "set \D"] by simp have "?\ \ PiE' I (set o D.map) \ Univ" using 1 2 3 by force moreover have "inj_on ?\ (PiE' I (set o D.map))" proof - have 7: "\A B C D F G H. F \ A \ B \ G \ B \ C \ H \ C \ D \ inj_on F A \ inj_on G B \ inj_on H C \ inj_on (H o G o F) A" by (simp add: Pi_iff inj_on_def) show ?thesis using 1 2 3 4 5 6 7 [of D.funToCone "PiE' I (set o D.map)" "D.cones unity" "inv_into (hom unity \D) \" "hom unity \D" img Univ] by fastforce qed moreover have "PiE' I (set o D.map) = PiE' I (\x. Univ)" proof - have "\i. i \ I \ (set o D.map) i = Univ" - using J.arr_char D.map_def set_mkIde by simp + using J.arr_char D.map_def by simp thus ?thesis by blast qed ultimately have "?\ \ (PiE' I (\x. Univ)) \ Univ \ inj_on ?\ (PiE' I (\x. Univ))" by auto thus "I \ UNIV \ admits_tupling I" using I admits_tupling_def by auto next assume ex_\: "I \ UNIV \ admits_tupling I" show "has_products I" proof (unfold has_products_def) from ex_\ obtain \ where \: "\ \ (PiE' I (\x. Univ)) \ Univ \ inj_on \ (PiE' I (\x. Univ))" using admits_tupling_def by metis text\ Given an @{term I}-indexed discrete diagram @{term D}, obtain the object @{term \D} of @{term[source=true] S} corresponding to the set @{term "\ ` PiE I D"} of all @{term "\ d"} where \d \ d \ J \\<^sub>E Univ\ and @{term "d i \ D i"} for all @{term "i \ I"}. The elements of @{term \D} are in bijective correspondence with the set of cones over @{term D}, hence @{term \D} is a limit of @{term D}. \ have "\J D. discrete_diagram J S D \ Collect (partial_magma.arr J) = I \ \\D. has_as_product J D \D" proof fix J :: "'i comp" and D assume D: "discrete_diagram J S D \ Collect (partial_magma.arr J) = I" interpret J: category J using D discrete_diagram.axioms(1) by blast interpret D: discrete_diagram J S D using D by simp - interpret D: discrete_diagram_in_replete_set_category J S D .. + interpret D: discrete_diagram_in_set_category J S \\A. A \ Univ\ D .. let ?\D = "mkIde (\ ` PiE' I (set o D))" have 0: "ide ?\D" proof - have "set o D \ I \ Pow Univ" - using Pow_iff incl_in_def o_apply elem_set_implies_incl_in - set_subset_Univ subsetI - by (metis (mono_tags, lifting) Pi_I') + using Pow_iff incl_in_def o_apply elem_set_implies_incl_in subsetI Pi_I' + setp_set_ide + by (metis (mono_tags, lifting)) hence "\ ` PiE' I (set o D) \ Univ" using \ by blast thus ?thesis using \ ide_mkIde by simp qed hence set_\D: "\ ` PiE' I (set o D) = set ?\D" using 0 ide_in_hom arr_mkIde set_mkIde by auto text\ The elements of @{term \D} are all values of the form @{term "\ d"}, where @{term d} satisfies @{term "d i \ set (D i)"} for all @{term "i \ I"}. Such @{term d} correspond bijectively to cones. Since @{term \} is injective, the values @{term "\ d"} correspond bijectively to cones. \ let ?\ = "mkPoint ?\D o \ o D.coneToFun" let ?\' = "D.funToCone o inv_into (PiE' I (set o D)) \ o img" have 1: "\ \ PiE' I (set o D) \ set ?\D \ inj_on \ (PiE' I (set o D))" proof - have "PiE' I (set o D) \ PiE' I (\x. Univ)" - by (metis PiE'_mono comp_apply elem_set_implies_incl_in incl_in_def - set_subset_Univ subsetI) + using setp_set_ide elem_set_implies_incl_in elem_set_implies_set_eq_singleton + incl_in_def PiE'_mono comp_apply subsetI + by (metis (no_types, lifting)) thus ?thesis using \ subset_inj_on set_\D Pi_I' imageI by fastforce qed have 2: "inv_into (PiE' I (set o D)) \ \ set ?\D \ PiE' I (set o D)" proof fix y assume y: "y \ set ?\D" have "y \ \ ` (PiE' I (set o D))" using y set_\D by auto thus "inv_into (PiE' I (set o D)) \ y \ PiE' I (set o D)" using inv_into_into [of y \ "PiE' I (set o D)"] by simp qed have 3: "\x. x \ set ?\D \ \ (inv_into (PiE' I (set o D)) \ x) = x" using set_\D by (simp add: f_inv_into_f) have 4: "\d. d \ PiE' I (set o D) \ inv_into (PiE' I (set o D)) \ (\ d) = d" using 1 by auto have 5: "D.I = I" using D by auto have "bij_betw ?\ (D.cones unity) (hom unity ?\D)" proof (intro bij_betwI) show "?\ \ D.cones unity \ hom unity ?\D" proof fix \ assume \: "\ \ D.cones unity" show "?\ \ \ hom unity ?\D" using \ 0 1 5 D.coneToFun_mapsto mkPoint_in_hom [of ?\D] by (simp, blast) qed show "?\' \ hom unity ?\D \ D.cones unity" proof fix x assume x: "x \ hom unity ?\D" hence "img x \ set ?\D" using img_point_elem_set by blast hence "inv_into (PiE' I (set o D)) \ (img x) \ Pi I (set \ D) \ extensional' I" using 2 by blast thus "?\' x \ D.cones unity" using 5 D.funToCone_mapsto by auto qed show "\x. x \ hom unity ?\D \ ?\ (?\' x) = x" proof - fix x assume x: "x \ hom unity ?\D" show "?\ (?\' x) = x" proof - have "D.coneToFun (D.funToCone (inv_into (PiE' I (set o D)) \ (img x))) = inv_into (PiE' I (set o D)) \ (img x)" using x 1 5 img_point_elem_set set_\D D.coneToFun_funToCone by force hence "\ (D.coneToFun (D.funToCone (inv_into (PiE' I (set o D)) \ (img x)))) = img x" using x 3 img_point_elem_set set_\D by force thus ?thesis using x 0 mkPoint_img by auto qed qed show "\\. \ \ D.cones unity \ ?\' (?\ \) = \" proof - fix \ assume \: "\ \ D.cones unity" show "?\' (?\ \) = \" proof - have "img (mkPoint ?\D (\ (D.coneToFun \))) = \ (D.coneToFun \)" using \ 0 1 5 D.coneToFun_mapsto img_mkPoint(2) by blast hence "inv_into (PiE' I (set o D)) \ (img (mkPoint ?\D (\ (D.coneToFun \)))) = D.coneToFun \" - using \ D.coneToFun_mapsto 4 5 by (metis PiE) + using \ D.coneToFun_mapsto 4 5 + by (metis (no_types, lifting) PiE) hence "D.funToCone (inv_into (PiE' I (set o D)) \ (img (mkPoint ?\D (\ (D.coneToFun \))))) = \" using \ D.funToCone_coneToFun by auto thus ?thesis by auto qed qed qed hence "bij_betw (inv_into (D.cones unity) ?\) (hom unity ?\D) (D.cones unity)" using bij_betw_inv_into by blast hence "\\. bij_betw \ (hom unity ?\D) (D.cones unity)" by blast hence "D.has_as_limit ?\D" using \ide ?\D\ D.limits_are_sets_of_cones by simp from this obtain \ where \: "limit_cone J S D ?\D \" by blast interpret \: limit_cone J S D ?\D \ using \ by auto interpret P: product_cone J S D ?\D \ using \ D.product_coneI by blast have "product_cone J S D ?\D \" .. thus "has_as_product J D ?\D" using has_as_product_def by auto qed thus "I \ UNIV \ (\J D. discrete_diagram J S D \ Collect (partial_magma.arr J) = I \ (\\D. has_as_product J D \D))" using ex_\ by blast qed qed + end + + context replete_set_category + begin + text\ Characterization of the completeness properties enjoyed by a set category: A set category @{term[source=true] S} has all limits at a type @{typ 'j}, if and only if @{term[source=true] S} admits @{term I}-indexed tupling for all @{typ 'j}-sets @{term I} such that @{term "I \ UNIV"}. \ theorem has_limits_iff_admits_tupling: shows "has_limits (undefined :: 'j) \ (\I :: 'j set. I \ UNIV \ admits_tupling I)" proof assume has_limits: "has_limits (undefined :: 'j)" show "\I :: 'j set. I \ UNIV \ admits_tupling I" using has_limits has_products_if_has_limits has_products_iff_admits_tupling by blast next assume admits_tupling: "\I :: 'j set. I \ UNIV \ admits_tupling I" show "has_limits (undefined :: 'j)" using has_limits_def admits_tupling has_products_iff_admits_tupling by (metis category.axioms(1) category.ideD(1) has_limits_if_has_products iso_tuple_UNIV_I mem_Collect_eq partial_magma.not_arr_null) qed end section "Limits in Functor Categories" text\ In this section, we consider the special case of limits in functor categories, with the objective of showing that limits in a functor category \[A, B]\ are given pointwise, and that \[A, B]\ has all limits that @{term B} has. \ locale parametrized_diagram = J: category J + A: category A + B: category B + JxA: product_category J A + binary_functor J A B D for J :: "'j comp" (infixr "\\<^sub>J" 55) and A :: "'a comp" (infixr "\\<^sub>A" 55) and B :: "'b comp" (infixr "\\<^sub>B" 55) and D :: "'j * 'a \ 'b" begin (* Notation for A.in_hom and B.in_hom is being inherited, but from where? *) notation J.in_hom ("\_ : _ \\<^sub>J _\") notation JxA.comp (infixr "\\<^sub>J\<^sub>x\<^sub>A" 55) notation JxA.in_hom ("\_ : _ \\<^sub>J\<^sub>x\<^sub>A _\") text\ A choice of limit cone for each diagram \D (-, a)\, where @{term a} is an object of @{term[source=true] A}, extends to a functor \L: A \ B\, where the action of @{term L} on arrows of @{term[source=true] A} is determined by universality. \ abbreviation L where "L \ \l \. \a. if A.arr a then limit_cone.induced_arrow J B (\j. D (j, A.cod a)) (l (A.cod a)) (\ (A.cod a)) (l (A.dom a)) (vertical_composite.map J B (\ (A.dom a)) (\j. D (j, a))) else B.null" abbreviation P where "P \ \l \. \a f. \f : l (A.dom a) \\<^sub>B l (A.cod a)\ \ diagram.cones_map J B (\j. D (j, A.cod a)) f (\ (A.cod a)) = vertical_composite.map J B (\ (A.dom a)) (\j. D (j, a))" lemma L_arr: assumes "\a. A.ide a \ limit_cone J B (\j. D (j, a)) (l a) (\ a)" shows "\a. A.arr a \ (\!f. P l \ a f) \ P l \ a (L l \ a)" proof fix a assume a: "A.arr a" interpret \_dom_a: limit_cone J B \\j. D (j, A.dom a)\ \l (A.dom a)\ \\ (A.dom a)\ using a assms by auto interpret \_cod_a: limit_cone J B \\j. D (j, A.cod a)\ \l (A.cod a)\ \\ (A.cod a)\ using a assms by auto interpret Da: natural_transformation J B \\j. D (j, A.dom a)\ \\j. D (j, A.cod a)\ \\j. D (j, a)\ using a fixing_arr_gives_natural_transformation_2 by simp interpret Dao\_dom_a: vertical_composite J B \_dom_a.A.map \\j. D (j, A.dom a)\ \\j. D (j, A.cod a)\ \\ (A.dom a)\ \\j. D (j, a)\ .. interpret Dao\_dom_a: cone J B \\j. D (j, A.cod a)\ \l (A.dom a)\ Dao\_dom_a.map .. show "P l \ a (L l \ a)" using a Dao\_dom_a.cone_axioms \_cod_a.induced_arrowI [of Dao\_dom_a.map "l (A.dom a)"] by auto show "\!f. P l \ a f" using \_cod_a.is_universal Dao\_dom_a.cone_axioms by blast qed lemma L_ide: assumes "\a. A.ide a \ limit_cone J B (\j. D (j, a)) (l a) (\ a)" shows "\a. A.ide a \ L l \ a = l a" proof - let ?L = "L l \" let ?P = "P l \" fix a assume a: "A.ide a" interpret \a: limit_cone J B \\j. D (j, a)\ \l a\ \\ a\ using a assms by auto have Pa: "?P a = (\f. f \ B.hom (l a) (l a) \ diagram.cones_map J B (\j. D (j, a)) f (\ a) = \ a)" using a vcomp_ide_dom \a.natural_transformation_axioms by simp have "?P a (?L a)" using assms a L_arr [of l \ a] by fastforce moreover have "?P a (l a)" proof - have "?P a (l a) \ l a \ B.hom (l a) (l a) \ \a.D.cones_map (l a) (\ a) = \ a" using Pa by meson thus ?thesis using a \a.ide_apex \a.cone_axioms \a.D.cones_map_ide [of "\ a" "l a"] by force qed moreover have "\!f. ?P a f" using a Pa \a.is_universal \a.cone_axioms by force ultimately show "?L a = l a" by blast qed lemma chosen_limits_induce_functor: assumes "\a. A.ide a \ limit_cone J B (\j. D (j, a)) (l a) (\ a)" shows "functor A B (L l \)" proof - let ?L = "L l \" let ?P = "\a. \f. \f : l (A.dom a) \\<^sub>B l (A.cod a)\ \ diagram.cones_map J B (\j. D (j, A.cod a)) f (\ (A.cod a)) = vertical_composite.map J B (\ (A.dom a)) (\j. D (j, a))" interpret L: "functor" A B ?L apply unfold_locales using assms L_arr [of l] L_ide apply auto[4] proof - fix a' a assume 1: "A.arr (A a' a)" have a: "A.arr a" using 1 by auto have a': "\a' : A.cod a \\<^sub>A A.cod a'\" using 1 by auto have a'a: "A.seq a' a" using 1 by auto interpret \_dom_a: limit_cone J B \\j. D (j, A.dom a)\ \l (A.dom a)\ \\ (A.dom a)\ using a assms by auto interpret \_cod_a: limit_cone J B \\j. D (j, A.cod a)\ \l (A.cod a)\ \\ (A.cod a)\ using a'a assms by auto interpret \_dom_a'a: limit_cone J B \\j. D (j, A.dom (a' \\<^sub>A a))\ \l (A.dom (a' \\<^sub>A a))\ \\ (A.dom (a' \\<^sub>A a))\ using a'a assms by auto interpret \_cod_a'a: limit_cone J B \\j. D (j, A.cod (a' \\<^sub>A a))\ \l (A.cod (a' \\<^sub>A a))\ \\ (A.cod (a' \\<^sub>A a))\ using a'a assms by auto interpret Da: natural_transformation J B \\j. D (j, A.dom a)\ \\j. D (j, A.cod a)\ \\j. D (j, a)\ using a fixing_arr_gives_natural_transformation_2 by simp interpret Da': natural_transformation J B \\j. D (j, A.cod a)\ \\j. D (j, A.cod (a' \\<^sub>A a))\ \\j. D (j, a')\ using a a'a fixing_arr_gives_natural_transformation_2 by fastforce interpret Da'o\_cod_a: vertical_composite J B \_cod_a.A.map \\j. D (j, A.cod a)\ \\j. D (j, A.cod (a' \\<^sub>A a))\ \\ (A.cod a)\ \\j. D (j, a')\.. interpret Da'o\_cod_a: cone J B \\j. D (j, A.cod (a' \\<^sub>A a))\ \l (A.cod a)\ Da'o\_cod_a.map .. interpret Da'a: natural_transformation J B \\j. D (j, A.dom (a' \\<^sub>A a))\ \\j. D (j, A.cod (a' \\<^sub>A a))\ \\j. D (j, a' \\<^sub>A a)\ using a'a fixing_arr_gives_natural_transformation_2 [of "a' \\<^sub>A a"] by auto interpret Da'ao\_dom_a'a: vertical_composite J B \_dom_a'a.A.map \\j. D (j, A.dom (a' \\<^sub>A a))\ \\j. D (j, A.cod (a' \\<^sub>A a))\ \\ (A.dom (a' \\<^sub>A a))\ \\j. D (j, a' \\<^sub>A a)\ .. interpret Da'ao\_dom_a'a: cone J B \\j. D (j, A.cod (a' \\<^sub>A a))\ \l (A.dom (a' \\<^sub>A a))\ Da'ao\_dom_a'a.map .. show "?L (a' \\<^sub>A a) = ?L a' \\<^sub>B ?L a" proof - have "?P (a' \\<^sub>A a) (?L (a' \\<^sub>A a))" using assms a'a L_arr [of l \ "a' \\<^sub>A a"] by fastforce moreover have "?P (a' \\<^sub>A a) (?L a' \\<^sub>B ?L a)" proof have La: "\?L a : l (A.dom a) \\<^sub>B l (A.cod a)\" using assms a L_arr by fast moreover have La': "\?L a' : l (A.cod a) \\<^sub>B l (A.cod a')\" using assms a a' L_arr [of l \ a'] by auto ultimately have seq: "B.seq (?L a') (?L a)" by (elim B.in_homE, auto) thus La'_La: "\?L a' \\<^sub>B ?L a : l (A.dom (a' \\<^sub>A a)) \\<^sub>B l (A.cod (a' \\<^sub>A a))\" using a a' 1 La La' by (intro B.comp_in_homI, auto) show "\_cod_a'a.D.cones_map (?L a' \\<^sub>B ?L a) (\ (A.cod (a' \\<^sub>A a))) = Da'ao\_dom_a'a.map" proof - have "\_cod_a'a.D.cones_map (?L a' \\<^sub>B ?L a) (\ (A.cod (a' \\<^sub>A a))) = (\_cod_a'a.D.cones_map (?L a) o \_cod_a'a.D.cones_map (?L a')) (\ (A.cod a'))" proof - have "\_cod_a'a.D.cones_map (?L a' \\<^sub>B ?L a) (\ (A.cod (a' \\<^sub>A a))) = restrict (\_cod_a'a.D.cones_map (?L a) \ \_cod_a'a.D.cones_map (?L a')) (\_cod_a'a.D.cones (B.cod (?L a'))) (\ (A.cod (a' \\<^sub>A a)))" using seq \_cod_a'a.cone_axioms \_cod_a'a.D.cones_map_comp [of "?L a'" "?L a"] by argo also have "... = (\_cod_a'a.D.cones_map (?L a) o \_cod_a'a.D.cones_map (?L a')) (\ (A.cod a'))" proof - have "\ (A.cod a') \ \_cod_a'a.D.cones (l (A.cod a'))" using \_cod_a'a.cone_axioms a'a by simp moreover have "B.cod (?L a') = l (A.cod a')" using assms a' L_arr [of l] by auto ultimately show ?thesis using a' a'a by simp qed finally show ?thesis by blast qed also have "... = \_cod_a'a.D.cones_map (?L a) (\_cod_a'a.D.cones_map (?L a') (\ (A.cod a')))" by simp also have "... = \_cod_a'a.D.cones_map (?L a) Da'o\_cod_a.map" proof - have "?P a' (?L a')" using assms a' L_arr [of l \ a'] by fast moreover have "?P a' = (\f. f \ B.hom (l (A.cod a)) (l (A.cod a')) \ \_cod_a'a.D.cones_map f (\ (A.cod a')) = Da'o\_cod_a.map)" using a'a by force ultimately show ?thesis using a'a by force qed also have "... = vertical_composite.map J B (\_cod_a.D.cones_map (?L a) (\ (A.cod a))) (\j. D (j, a'))" using assms \_cod_a.D.diagram_axioms \_cod_a'a.D.diagram_axioms Da'.natural_transformation_axioms \_cod_a.cone_axioms La cones_map_vcomp [of J B "\j. D (j, A.cod a)" "\j. D (j, A.cod (a' \\<^sub>A a))" "\j. D (j, a')" "l (A.cod a)" "\ (A.cod a)" "?L a" "l (A.dom a)"] by blast also have "... = vertical_composite.map J B (vertical_composite.map J B (\ (A.dom a)) (\j. D (j, a))) (\j. D (j, a'))" using assms a L_arr by presburger also have "... = vertical_composite.map J B (\ (A.dom a)) (vertical_composite.map J B (\j. D (j, a)) (\j. D (j, a')))" using a'a Da.natural_transformation_axioms Da'.natural_transformation_axioms \_dom_a.natural_transformation_axioms vcomp_assoc by auto also have "... = vertical_composite.map J B (\ (A.dom (a' \\<^sub>A a))) (\j. D (j, a' \\<^sub>A a))" using a'a preserves_comp_2 by simp finally show ?thesis by auto qed qed moreover have "\!f. ?P (a' \\<^sub>A a) f" using \_cod_a'a.is_universal [of "l (A.dom (a' \\<^sub>A a))" "vertical_composite.map J B (\ (A.dom (a' \\<^sub>A a))) (\j. D (j, a' \\<^sub>A a))"] Da'ao\_dom_a'a.cone_axioms by fast ultimately show ?thesis by blast qed qed show ?thesis .. qed end locale diagram_in_functor_category = A: category A + B: category B + A_B: functor_category A B + diagram J A_B.comp D for A :: "'a comp" (infixr "\\<^sub>A" 55) and B :: "'b comp" (infixr "\\<^sub>B" 55) and J :: "'j comp" (infixr "\\<^sub>J" 55) and D :: "'j \ ('a, 'b) functor_category.arr" begin interpretation JxA: product_category J A .. interpretation A_BxA: product_category A_B.comp A .. interpretation E: evaluation_functor A B .. interpretation Curry: currying J A B .. notation JxA.comp (infixr "\\<^sub>J\<^sub>x\<^sub>A" 55) notation JxA.in_hom ("\_ : _ \\<^sub>J\<^sub>x\<^sub>A _\") text\ Evaluation of a functor or natural transformation from @{term[source=true] J} to \[A, B]\ at an arrow @{term a} of @{term[source=true] A}. \ abbreviation at where "at a \ \ \j. Curry.uncurry \ (j, a)" lemma at_simp: assumes "A.arr a" and "J.arr j" and "A_B.arr (\ j)" shows "at a \ j = A_B.Map (\ j) a" using assms Curry.uncurry_def E.map_simp by simp lemma functor_at_ide_is_functor: assumes "functor J A_B.comp F" and "A.ide a" shows "functor J B (at a F)" proof - interpret uncurry_F: "functor" JxA.comp B \Curry.uncurry F\ using assms(1) Curry.uncurry_preserves_functors by simp interpret uncurry_F: binary_functor J A B \Curry.uncurry F\ .. show ?thesis using assms(2) uncurry_F.fixing_ide_gives_functor_2 by simp qed lemma functor_at_arr_is_transformation: assumes "functor J A_B.comp F" and "A.arr a" shows "natural_transformation J B (at (A.dom a) F) (at (A.cod a) F) (at a F)" proof - interpret uncurry_F: "functor" JxA.comp B \Curry.uncurry F\ using assms(1) Curry.uncurry_preserves_functors by simp interpret uncurry_F: binary_functor J A B \Curry.uncurry F\ .. show ?thesis using assms(2) uncurry_F.fixing_arr_gives_natural_transformation_2 by simp qed lemma transformation_at_ide_is_transformation: assumes "natural_transformation J A_B.comp F G \" and "A.ide a" shows "natural_transformation J B (at a F) (at a G) (at a \)" proof - interpret \: natural_transformation J A_B.comp F G \ using assms(1) by auto interpret uncurry_F: "functor" JxA.comp B \Curry.uncurry F\ using Curry.uncurry_preserves_functors \.F.functor_axioms by simp interpret uncurry_f: binary_functor J A B \Curry.uncurry F\ .. interpret uncurry_G: "functor" JxA.comp B \Curry.uncurry G\ using Curry.uncurry_preserves_functors \.G.functor_axioms by simp interpret uncurry_G: binary_functor J A B \Curry.uncurry G\ .. interpret uncurry_\: natural_transformation JxA.comp B \Curry.uncurry F\ \Curry.uncurry G\ \Curry.uncurry \\ using Curry.uncurry_preserves_transformations \.natural_transformation_axioms by simp interpret uncurry_\: binary_functor_transformation J A B \Curry.uncurry F\ \Curry.uncurry G\ \Curry.uncurry \\ .. show ?thesis using assms(2) uncurry_\.fixing_ide_gives_natural_transformation_2 by simp qed lemma constant_at_ide_is_constant: assumes "cone x \" and a: "A.ide a" shows "at a (constant_functor.map J A_B.comp x) = constant_functor.map J B (A_B.Map x a)" proof - interpret \: cone J A_B.comp D x \ using assms(1) by auto have x: "A_B.ide x" using \.ide_apex by auto interpret Fun_x: "functor" A B \A_B.Map x\ using x A_B.ide_char by simp interpret Da: "functor" J B \at a D\ using a functor_at_ide_is_functor functor_axioms by blast interpret Da: diagram J B \at a D\ .. interpret Xa: constant_functor J B \A_B.Map x a\ using a Fun_x.preserves_ide by unfold_locales simp show "at a \.A.map = Xa.map" using a x Curry.uncurry_def E.map_def Xa.is_extensional by auto qed lemma at_ide_is_diagram: assumes a: "A.ide a" shows "diagram J B (at a D)" proof - interpret Da: "functor" J B "at a D" using a functor_at_ide_is_functor functor_axioms by simp show ?thesis .. qed lemma cone_at_ide_is_cone: assumes "cone x \" and a: "A.ide a" shows "diagram.cone J B (at a D) (A_B.Map x a) (at a \)" proof - interpret \: cone J A_B.comp D x \ using assms(1) by auto have x: "A_B.ide x" using \.ide_apex by auto interpret Fun_x: "functor" A B \A_B.Map x\ using x A_B.ide_char by simp interpret Da: diagram J B \at a D\ using a at_ide_is_diagram by auto interpret Xa: constant_functor J B \A_B.Map x a\ using a by (unfold_locales, simp) interpret \a: natural_transformation J B Xa.map \at a D\ \at a \\ using assms(1) x a transformation_at_ide_is_transformation \.natural_transformation_axioms constant_at_ide_is_constant by fastforce interpret \a: cone J B \at a D\ \A_B.Map x a\ \at a \\ .. show cone_\a: "Da.cone (A_B.Map x a) (at a \)" .. qed lemma at_preserves_comp: assumes "A.seq a' a" shows "at (A a' a) D = vertical_composite.map J B (at a D) (at a' D)" proof - interpret Da: natural_transformation J B \at (A.dom a) D\ \at (A.cod a) D\ \at a D\ using assms functor_at_arr_is_transformation functor_axioms by blast interpret Da': natural_transformation J B \at (A.cod a) D\ \at (A.cod a') D\ \at a' D\ using assms functor_at_arr_is_transformation [of D a'] functor_axioms by fastforce interpret Da'oDa: vertical_composite J B \at (A.dom a) D\ \at (A.cod a) D\ \at (A.cod a') D\ \at a D\ \at a' D\ .. interpret Da'a: natural_transformation J B \at (A.dom a) D\ \at (A.cod a') D\ \at (a' \\<^sub>A a) D\ using assms functor_at_arr_is_transformation [of D "a' \\<^sub>A a"] functor_axioms by simp show "at (a' \\<^sub>A a) D = Da'oDa.map" proof (intro NaturalTransformation.eqI) show "natural_transformation J B (at (A.dom a) D) (at (A.cod a') D) Da'oDa.map" .. show "natural_transformation J B (at (A.dom a) D) (at (A.cod a') D) (at (a' \\<^sub>A a) D)" .. show "\j. J.ide j \ at (a' \\<^sub>A a) D j = Da'oDa.map j" proof - fix j assume j: "J.ide j" interpret Dj: "functor" A B \A_B.Map (D j)\ using j preserves_ide A_B.ide_char by simp show "at (a' \\<^sub>A a) D j = Da'oDa.map j" using assms j Dj.preserves_comp at_simp Da'oDa.map_simp_ide by auto qed qed qed lemma cones_map_pointwise: assumes "cone x \" and "cone x' \'" and f: "f \ A_B.hom x' x" shows "cones_map f \ = \' \ (\a. A.ide a \ diagram.cones_map J B (at a D) (A_B.Map f a) (at a \) = at a \')" proof interpret \: cone J A_B.comp D x \ using assms(1) by auto interpret \': cone J A_B.comp D x' \' using assms(2) by auto have x: "A_B.ide x" using \.ide_apex by auto have x': "A_B.ide x'" using \'.ide_apex by auto interpret \f: cone J A_B.comp D x' \cones_map f \\ using x' f assms(1) cones_map_mapsto by blast interpret Fun_x: "functor" A B \A_B.Map x\ using x A_B.ide_char by simp interpret Fun_x': "functor" A B \A_B.Map x'\ using x' A_B.ide_char by simp show "cones_map f \ = \' \ (\a. A.ide a \ diagram.cones_map J B (at a D) (A_B.Map f a) (at a \) = at a \')" proof - assume \': "cones_map f \ = \'" have "\a. A.ide a \ diagram.cones_map J B (at a D) (A_B.Map f a) (at a \) = at a \'" proof - fix a assume a: "A.ide a" interpret Da: diagram J B \at a D\ using a at_ide_is_diagram by auto interpret \a: cone J B \at a D\ \A_B.Map x a\ \at a \\ using a assms(1) cone_at_ide_is_cone by simp interpret \'a: cone J B \at a D\ \A_B.Map x' a\ \at a \'\ using a assms(2) cone_at_ide_is_cone by simp have 1: "\A_B.Map f a : A_B.Map x' a \\<^sub>B A_B.Map x a\" using f a A_B.arr_char A_B.Map_cod A_B.Map_dom mem_Collect_eq natural_transformation.preserves_hom A.ide_in_hom by (metis (no_types, lifting) A_B.in_homE) interpret \fa: cone J B \at a D\ \A_B.Map x' a\ \Da.cones_map (A_B.Map f a) (at a \)\ using 1 \a.cone_axioms Da.cones_map_mapsto by force show "Da.cones_map (A_B.Map f a) (at a \) = at a \'" proof fix j have "\J.arr j \ Da.cones_map (A_B.Map f a) (at a \) j = at a \' j" using \'a.is_extensional \fa.is_extensional [of j] by simp moreover have "J.arr j \ Da.cones_map (A_B.Map f a) (at a \) j = at a \' j" using a f 1 \.cone_axioms \a.cone_axioms at_simp apply simp apply (elim A_B.in_homE B.in_homE, auto) using \' \.A.map_simp A_B.Map_comp [of "\ j" f a a] by auto ultimately show "Da.cones_map (A_B.Map f a) (at a \) j = at a \' j" by blast qed qed thus "\a. A.ide a \ diagram.cones_map J B (at a D) (A_B.Map f a) (at a \) = at a \'" by simp qed show "\a. A.ide a \ diagram.cones_map J B (at a D) (A_B.Map f a) (at a \) = at a \' \ cones_map f \ = \'" proof - assume A: "\a. A.ide a \ diagram.cones_map J B (at a D) (A_B.Map f a) (at a \) = at a \'" show "cones_map f \ = \'" proof (intro NaturalTransformation.eqI) show "natural_transformation J A_B.comp \'.A.map D (cones_map f \)" .. show "natural_transformation J A_B.comp \'.A.map D \'" .. show "\j. J.ide j \ cones_map f \ j = \' j" proof (intro A_B.arr_eqI) fix j assume j: "J.ide j" show 1: "A_B.arr (cones_map f \ j)" using j \f.preserves_reflects_arr by simp show "A_B.arr (\' j)" using j by auto have Dom_\f_j: "A_B.Dom (cones_map f \ j) = A_B.Map x'" using x' j 1 A_B.Map_dom \'.A.map_simp \f.preserves_dom J.ide_in_hom by (metis (no_types, lifting) J.ideD(2) \f.preserves_reflects_arr) also have Dom_\'_j: "... = A_B.Dom (\' j)" using x' j A_B.Map_dom [of "\' j"] \'.preserves_hom \'.A.map_simp by simp finally show "A_B.Dom (cones_map f \ j) = A_B.Dom (\' j)" by auto have Cod_\f_j: "A_B.Cod (cones_map f \ j) = A_B.Map (D (J.cod j))" using j A_B.Map_cod A_B.cod_char J.ide_in_hom \f.preserves_hom by (metis (no_types, lifting) "1" J.ideD(1) \f.preserves_cod) also have Cod_\'_j: "... = A_B.Cod (\' j)" using j A_B.Map_cod [of "\' j"] \'.preserves_hom by simp finally show "A_B.Cod (cones_map f \ j) = A_B.Cod (\' j)" by auto show "A_B.Map (cones_map f \ j) = A_B.Map (\' j)" proof (intro NaturalTransformation.eqI) interpret \fj: natural_transformation A B \A_B.Map x'\ \A_B.Map (D (J.cod j))\ \A_B.Map (cones_map f \ j)\ using j \f.preserves_reflects_arr A_B.arr_char [of "cones_map f \ j"] Dom_\f_j Cod_\f_j by simp show "natural_transformation A B (A_B.Map x') (A_B.Map (D (J.cod j))) (A_B.Map (cones_map f \ j))" .. interpret \'j: natural_transformation A B \A_B.Map x'\ \A_B.Map (D (J.cod j))\ \A_B.Map (\' j)\ using j A_B.arr_char [of "\' j"] Dom_\'_j Cod_\'_j by simp show "natural_transformation A B (A_B.Map x') (A_B.Map (D (J.cod j))) (A_B.Map (\' j))" .. show "\a. A.ide a \ A_B.Map (cones_map f \ j) a = A_B.Map (\' j) a" proof - fix a assume a: "A.ide a" interpret Da: diagram J B \at a D\ using a at_ide_is_diagram by auto have cone_\a: "Da.cone (A_B.Map x a) (at a \)" using a assms(1) cone_at_ide_is_cone by simp interpret \a: cone J B \at a D\ \A_B.Map x a\ \at a \\ using cone_\a by auto interpret Fun_f: natural_transformation A B \A_B.Dom f\ \A_B.Cod f\ \A_B.Map f\ using f A_B.arr_char by fast have fa: "A_B.Map f a \ B.hom (A_B.Map x' a) (A_B.Map x a)" using a f Fun_f.preserves_hom A.ide_in_hom by auto have "A_B.Map (cones_map f \ j) a = Da.cones_map (A_B.Map f a) (at a \) j" proof - have "A_B.Map (cones_map f \ j) a = A_B.Map (A_B.comp (\ j) f) a" using assms(1) f \.is_extensional by auto also have "... = B (A_B.Map (\ j) a) (A_B.Map f a)" using f j a \.preserves_hom A.ide_in_hom J.ide_in_hom A_B.Map_comp \.A.map_simp by (metis (no_types, lifting) A.comp_ide_self A.ideD(1) A_B.seqI' J.ideD(1) mem_Collect_eq) also have "... = Da.cones_map (A_B.Map f a) (at a \) j" using j a cone_\a fa Curry.uncurry_def E.map_simp by auto finally show ?thesis by auto qed also have "... = at a \' j" using j a A by simp also have "... = A_B.Map (\' j) a" using j Curry.uncurry_def E.map_simp \'j.is_extensional by simp finally show "A_B.Map (cones_map f \ j) a = A_B.Map (\' j) a" by auto qed qed qed qed qed qed text\ If @{term \} is a cone with apex @{term a} over @{term D}, then @{term \} is a limit cone if, for each object @{term x} of @{term X}, the cone obtained by evaluating @{term \} at @{term x} is a limit cone with apex @{term "A_B.Map a x"} for the diagram in @{term C} obtained by evaluating @{term D} at @{term x}. \ lemma cone_is_limit_if_pointwise_limit: assumes cone_\: "cone x \" and "\a. A.ide a \ diagram.limit_cone J B (at a D) (A_B.Map x a) (at a \)" shows "limit_cone x \" proof - interpret \: cone J A_B.comp D x \ using assms by auto have x: "A_B.ide x" using \.ide_apex by auto show "limit_cone x \" proof fix x' \' assume cone_\': "cone x' \'" interpret \': cone J A_B.comp D x' \' using cone_\' by auto have x': "A_B.ide x'" using \'.ide_apex by auto text\ The universality of the limit cone \at a \\ yields, for each object \a\ of \A\, a unique arrow \fa\ that transforms \at a \\ to \at a \'\. \ have EU: "\a. A.ide a \ \!fa. fa \ B.hom (A_B.Map x' a) (A_B.Map x a) \ diagram.cones_map J B (at a D) fa (at a \) = at a \'" proof - fix a assume a: "A.ide a" interpret Da: diagram J B \at a D\ using a at_ide_is_diagram by auto interpret \a: limit_cone J B \at a D\ \A_B.Map x a\ \at a \\ using assms(2) a by auto interpret \'a: cone J B \at a D\ \A_B.Map x' a\ \at a \'\ using a cone_\' cone_at_ide_is_cone by auto have "Da.cone (A_B.Map x' a) (at a \')" .. thus "\!fa. fa \ B.hom (A_B.Map x' a) (A_B.Map x a) \ Da.cones_map fa (at a \) = at a \'" using \a.is_universal by simp qed text\ Our objective is to show the existence of a unique arrow \f\ that transforms \\\ into \\'\. We obtain \f\ by bundling the arrows \fa\ of \C\ and proving that this yields a natural transformation from \X\ to \C\, hence an arrow of \[X, C]\. \ show "\!f. \f : x' \\<^sub>[\<^sub>A\<^sub>,\<^sub>B\<^sub>] x\ \ cones_map f \ = \'" proof let ?P = "\a fa. \fa : A_B.Map x' a \\<^sub>B A_B.Map x a\ \ diagram.cones_map J B (at a D) fa (at a \) = at a \'" have AaPa: "\a. A.ide a \ ?P a (THE fa. ?P a fa)" proof - fix a assume a: "A.ide a" have "\!fa. ?P a fa" using a EU by simp thus "?P a (THE fa. ?P a fa)" using a theI' [of "?P a"] by fastforce qed have AaPa_in_hom: "\a. A.ide a \ \THE fa. ?P a fa : A_B.Map x' a \\<^sub>B A_B.Map x a\" using AaPa by blast have AaPa_map: "\a. A.ide a \ diagram.cones_map J B (at a D) (THE fa. ?P a fa) (at a \) = at a \'" using AaPa by blast let ?Fun_f = "\a. if A.ide a then (THE fa. ?P a fa) else B.null" interpret Fun_x: "functor" A B \\a. A_B.Map x a\ using x A_B.ide_char by simp interpret Fun_x': "functor" A B \\a. A_B.Map x' a\ using x' A_B.ide_char by simp text\ The arrows \Fun_f a\ are the components of a natural transformation. It is more work to verify the naturality than it seems like it ought to be. \ interpret \: transformation_by_components A B \\a. A_B.Map x' a\ \\a. A_B.Map x a\ ?Fun_f proof fix a assume a: "A.ide a" show "\?Fun_f a : A_B.Map x' a \\<^sub>B A_B.Map x a\" using a AaPa by simp next fix a assume a: "A.arr a" text\ \newcommand\xdom{\mathop{\rm dom}} \newcommand\xcod{\mathop{\rm cod}} $$\xymatrix{ {x_{\xdom a}} \drtwocell\omit{\omit(A)} \ar[d]_{\chi_{\xdom a}} \ar[r]^{x_a} & {x_{\xcod a}} \ar[d]^{\chi_{\xcod a}} \\ {D_{\xdom a}} \ar[r]^{D_a} & {D_{\xcod a}} \\ {x'_{\xdom a}} \urtwocell\omit{\omit(B)} \ar@/^5em/[uu]^{f_{\xdom a}}_{\hspace{1em}(C)} \ar[u]^{\chi'_{\xdom a}} \ar[r]_{x'_a} & {x'_{\xcod a}} \ar[u]_{x'_{\xcod a}} \ar@/_5em/[uu]_{f_{\xcod a}} }$$ \ let ?x_dom_a = "A_B.Map x (A.dom a)" let ?x_cod_a = "A_B.Map x (A.cod a)" let ?x_a = "A_B.Map x a" have x_a: "\?x_a : ?x_dom_a \\<^sub>B ?x_cod_a\" using a x A_B.ide_char by auto let ?x'_dom_a = "A_B.Map x' (A.dom a)" let ?x'_cod_a = "A_B.Map x' (A.cod a)" let ?x'_a = "A_B.Map x' a" have x'_a: "\?x'_a : ?x'_dom_a \\<^sub>B ?x'_cod_a\" using a x' A_B.ide_char by auto let ?f_dom_a = "?Fun_f (A.dom a)" let ?f_cod_a = "?Fun_f (A.cod a)" have f_dom_a: "\?f_dom_a : ?x'_dom_a \\<^sub>B ?x_dom_a\" using a AaPa by simp have f_cod_a: "\?f_cod_a : ?x'_cod_a \\<^sub>B ?x_cod_a\" using a AaPa by simp interpret D_dom_a: diagram J B \at (A.dom a) D\ using a at_ide_is_diagram by simp interpret D_cod_a: diagram J B \at (A.cod a) D\ using a at_ide_is_diagram by simp interpret Da: natural_transformation J B \at (A.dom a) D\ \at (A.cod a) D\ \at a D\ using a functor_axioms functor_at_arr_is_transformation by simp interpret \_dom_a: limit_cone J B \at (A.dom a) D\ \A_B.Map x (A.dom a)\ \at (A.dom a) \\ using assms(2) a by auto interpret \_cod_a: limit_cone J B \at (A.cod a) D\ \A_B.Map x (A.cod a)\ \at (A.cod a) \\ using assms(2) a by auto interpret \'_dom_a: cone J B \at (A.dom a) D\ \A_B.Map x' (A.dom a)\ \at (A.dom a) \'\ using a cone_\' cone_at_ide_is_cone by auto interpret \'_cod_a: cone J B \at (A.cod a) D\ \A_B.Map x' (A.cod a)\ \at (A.cod a) \'\ using a cone_\' cone_at_ide_is_cone by auto text\ Now construct cones with apexes \x_dom_a\ and \x'_dom_a\ over @{term "at (A.cod a) D"} by forming the vertical composites of @{term "at (A.dom a) \"} and @{term "at (A.cod a) \'"} with the natural transformation @{term "at a D"}. \ interpret Dao\_dom_a: vertical_composite J B \_dom_a.A.map \at (A.dom a) D\ \at (A.cod a) D\ \at (A.dom a) \\ \at a D\ .. interpret Dao\_dom_a: cone J B \at (A.cod a) D\ ?x_dom_a Dao\_dom_a.map using \_dom_a.cone_axioms Da.natural_transformation_axioms vcomp_transformation_cone by metis interpret Dao\'_dom_a: vertical_composite J B \'_dom_a.A.map \at (A.dom a) D\ \at (A.cod a) D\ \at (A.dom a) \'\ \at a D\ .. interpret Dao\'_dom_a: cone J B \at (A.cod a) D\ ?x'_dom_a Dao\'_dom_a.map using \'_dom_a.cone_axioms Da.natural_transformation_axioms vcomp_transformation_cone by metis have Dao\_dom_a: "D_cod_a.cone ?x_dom_a Dao\_dom_a.map" .. have Dao\'_dom_a: "D_cod_a.cone ?x'_dom_a Dao\'_dom_a.map" .. text\ These cones are also obtained by transforming the cones @{term "at (A.cod a) \"} and @{term "at (A.cod a) \'"} by \x_a\ and \x'_a\, respectively. \ have A: "Dao\_dom_a.map = D_cod_a.cones_map ?x_a (at (A.cod a) \)" proof fix j have "\J.arr j \ Dao\_dom_a.map j = D_cod_a.cones_map ?x_a (at (A.cod a) \) j" using Dao\_dom_a.is_extensional \_cod_a.cone_axioms x_a by force moreover have "J.arr j \ Dao\_dom_a.map j = D_cod_a.cones_map ?x_a (at (A.cod a) \) j" proof - assume j: "J.arr j" have "Dao\_dom_a.map j = at a D j \\<^sub>B at (A.dom a) \ (J.dom j)" using j Dao\_dom_a.map_simp_2 by simp also have "... = A_B.Map (D j) a \\<^sub>B A_B.Map (\ (J.dom j)) (A.dom a)" using a j at_simp by simp also have "... = A_B.Map (A_B.comp (D j) (\ (J.dom j))) a" using a j A_B.Map_comp by (metis (no_types, lifting) A.comp_arr_dom \.is_natural_1 \.preserves_reflects_arr) also have "... = A_B.Map (A_B.comp (\ (J.cod j)) (\.A.map j)) a" using a j \.naturality by simp also have "... = A_B.Map (\ (J.cod j)) (A.cod a) \\<^sub>B A_B.Map x a" using a j x A_B.Map_comp by (metis (no_types, lifting) A.comp_cod_arr \.A.map_simp \.is_natural_2 \.preserves_reflects_arr) also have "... = at (A.cod a) \ (J.cod j) \\<^sub>B A_B.Map x a" using a j at_simp by simp also have "... = at (A.cod a) \ j \\<^sub>B A_B.Map x a" using a j \_cod_a.is_natural_2 \_cod_a.A.map_simp by (metis J.arr_cod_iff_arr J.cod_cod) also have "... = D_cod_a.cones_map ?x_a (at (A.cod a) \) j" using a j x \_cod_a.cone_axioms preserves_cod by simp finally show ?thesis by blast qed ultimately show "Dao\_dom_a.map j = D_cod_a.cones_map ?x_a (at (A.cod a) \) j" by blast qed have B: "Dao\'_dom_a.map = D_cod_a.cones_map ?x'_a (at (A.cod a) \')" proof fix j have "\J.arr j \ Dao\'_dom_a.map j = D_cod_a.cones_map ?x'_a (at (A.cod a) \') j" using Dao\'_dom_a.is_extensional \'_cod_a.cone_axioms x'_a by force moreover have "J.arr j \ Dao\'_dom_a.map j = D_cod_a.cones_map ?x'_a (at (A.cod a) \') j" proof - assume j: "J.arr j" have "Dao\'_dom_a.map j = at a D j \\<^sub>B at (A.dom a) \' (J.dom j)" using j Dao\'_dom_a.map_simp_2 by simp also have "... = A_B.Map (D j) a \\<^sub>B A_B.Map (\' (J.dom j)) (A.dom a)" using a j at_simp by simp also have "... = A_B.Map (A_B.comp (D j) (\' (J.dom j))) a" using a j A_B.Map_comp by (metis (no_types, lifting) A.comp_arr_dom \'.is_natural_1 \'.preserves_reflects_arr) also have "... = A_B.Map (A_B.comp (\' (J.cod j)) (\'.A.map j)) a" using a j \'.naturality by simp also have "... = A_B.Map (\' (J.cod j)) (A.cod a) \\<^sub>B A_B.Map x' a" using a j x' A_B.Map_comp by (metis (no_types, lifting) A.comp_cod_arr \'.A.map_simp \'.is_natural_2 \'.preserves_reflects_arr) also have "... = at (A.cod a) \' (J.cod j) \\<^sub>B A_B.Map x' a" using a j at_simp by simp also have "... = at (A.cod a) \' j \\<^sub>B A_B.Map x' a" using a j \'_cod_a.is_natural_2 \'_cod_a.A.map_simp by (metis J.arr_cod_iff_arr J.cod_cod) also have "... = D_cod_a.cones_map ?x'_a (at (A.cod a) \') j" using a j x' \'_cod_a.cone_axioms preserves_cod by simp finally show ?thesis by blast qed ultimately show "Dao\'_dom_a.map j = D_cod_a.cones_map ?x'_a (at (A.cod a) \') j" by blast qed text\ Next, we show that \f_dom_a\, which is the unique arrow that transforms \\_dom_a\ into \\'_dom_a\, is also the unique arrow that transforms \Dao\_dom_a\ into \Dao\'_dom_a\. \ have C: "D_cod_a.cones_map ?f_dom_a Dao\_dom_a.map = Dao\'_dom_a.map" proof (intro NaturalTransformation.eqI) show "natural_transformation J B \'_dom_a.A.map (at (A.cod a) D) Dao\'_dom_a.map" .. show "natural_transformation J B \'_dom_a.A.map (at (A.cod a) D) (D_cod_a.cones_map ?f_dom_a Dao\_dom_a.map)" proof - interpret \: cone J B \at (A.cod a) D\ ?x'_dom_a \D_cod_a.cones_map ?f_dom_a Dao\_dom_a.map\ proof - have "\b b' f. \ f \ B.hom b' b; D_cod_a.cone b Dao\_dom_a.map \ \ D_cod_a.cone b' (D_cod_a.cones_map f Dao\_dom_a.map)" using D_cod_a.cones_map_mapsto by blast moreover have "D_cod_a.cone ?x_dom_a Dao\_dom_a.map" .. ultimately show "D_cod_a.cone ?x'_dom_a (D_cod_a.cones_map ?f_dom_a Dao\_dom_a.map)" using f_dom_a by simp qed show ?thesis .. qed show "\j. J.ide j \ D_cod_a.cones_map ?f_dom_a Dao\_dom_a.map j = Dao\'_dom_a.map j" proof - fix j assume j: "J.ide j" have "D_cod_a.cones_map ?f_dom_a Dao\_dom_a.map j = Dao\_dom_a.map j \\<^sub>B ?f_dom_a" using j f_dom_a Dao\_dom_a.cone_axioms by (elim B.in_homE, auto) also have "... = (at a D j \\<^sub>B at (A.dom a) \ j) \\<^sub>B ?f_dom_a" using j Dao\_dom_a.map_simp_ide by simp also have "... = at a D j \\<^sub>B at (A.dom a) \ j \\<^sub>B ?f_dom_a" using B.comp_assoc by simp also have "... = at a D j \\<^sub>B D_dom_a.cones_map ?f_dom_a (at (A.dom a) \) j" using j \_dom_a.cone_axioms f_dom_a by (elim B.in_homE, auto) also have "... = at a D j \\<^sub>B at (A.dom a) \' j" using a AaPa A.ide_dom by presburger also have "... = Dao\'_dom_a.map j" using j Dao\'_dom_a.map_simp_ide by simp finally show "D_cod_a.cones_map ?f_dom_a Dao\_dom_a.map j = Dao\'_dom_a.map j" by auto qed qed text\ Naturality amounts to showing that \C f_cod_a x'_a = C x_a f_dom_a\. To do this, we show that both arrows transform @{term "at (A.cod a) \"} into \Dao\'_cod_a\, thus they are equal by the universality of @{term "at (A.cod a) \"}. \ have "\!fa. \fa : ?x'_dom_a \\<^sub>B ?x_cod_a\ \ D_cod_a.cones_map fa (at (A.cod a) \) = Dao\'_dom_a.map" using Dao\'_dom_a.cone_axioms a \_cod_a.is_universal [of ?x'_dom_a Dao\'_dom_a.map] by fast moreover have "?f_cod_a \\<^sub>B ?x'_a \ B.hom ?x'_dom_a ?x_cod_a \ D_cod_a.cones_map (?f_cod_a \\<^sub>B ?x'_a) (at (A.cod a) \) = Dao\'_dom_a.map" proof show "?f_cod_a \\<^sub>B ?x'_a \ B.hom ?x'_dom_a ?x_cod_a" using f_cod_a x'_a by blast show "D_cod_a.cones_map (?f_cod_a \\<^sub>B ?x'_a) (at (A.cod a) \) = Dao\'_dom_a.map" proof - have "D_cod_a.cones_map (?f_cod_a \\<^sub>B ?x'_a) (at (A.cod a) \) = restrict (D_cod_a.cones_map ?x'_a o D_cod_a.cones_map ?f_cod_a) (D_cod_a.cones (?x_cod_a)) (at (A.cod a) \)" using x'_a D_cod_a.cones_map_comp [of ?f_cod_a ?x'_a] f_cod_a by (elim B.in_homE, auto) also have "... = D_cod_a.cones_map ?x'_a (D_cod_a.cones_map ?f_cod_a (at (A.cod a) \))" using \_cod_a.cone_axioms by simp also have "... = Dao\'_dom_a.map" using a B AaPa_map A.ide_cod by presburger finally show ?thesis by auto qed qed moreover have "?x_a \\<^sub>B ?f_dom_a \ B.hom ?x'_dom_a ?x_cod_a \ D_cod_a.cones_map (?x_a \\<^sub>B ?f_dom_a) (at (A.cod a) \) = Dao\'_dom_a.map" proof show "?x_a \\<^sub>B ?f_dom_a \ B.hom ?x'_dom_a ?x_cod_a" using f_dom_a x_a by blast show "D_cod_a.cones_map (?x_a \\<^sub>B ?f_dom_a) (at (A.cod a) \) = Dao\'_dom_a.map" proof - have "D_cod_a.cones (B.cod (A_B.Map x a)) = D_cod_a.cones (A_B.Map x (A.cod a))" using a x by simp moreover have "B.seq ?x_a ?f_dom_a" using f_dom_a x_a by (elim B.in_homE, auto) ultimately have "D_cod_a.cones_map (?x_a \\<^sub>B ?f_dom_a) (at (A.cod a) \) = restrict (D_cod_a.cones_map ?f_dom_a o D_cod_a.cones_map ?x_a) (D_cod_a.cones (?x_cod_a)) (at (A.cod a) \)" using D_cod_a.cones_map_comp [of ?x_a ?f_dom_a] x_a by argo also have "... = D_cod_a.cones_map ?f_dom_a (D_cod_a.cones_map ?x_a (at (A.cod a) \))" using \_cod_a.cone_axioms by simp also have "... = Dao\'_dom_a.map" using A C a AaPa by argo finally show ?thesis by blast qed qed ultimately show "?f_cod_a \\<^sub>B ?x'_a = ?x_a \\<^sub>B ?f_dom_a" using a \_cod_a.is_universal by blast qed text\ The arrow from @{term x'} to @{term x} in \[A, B]\ determined by the natural transformation \\\ transforms @{term \} into @{term \'}. Moreover, it is the unique such arrow, since the components of \\\ are each determined by universality. \ let ?f = "A_B.MkArr (\a. A_B.Map x' a) (\a. A_B.Map x a) \.map" have f_in_hom: "?f \ A_B.hom x' x" proof - have arr_f: "A_B.arr ?f" using x' x A_B.arr_MkArr \.natural_transformation_axioms by simp moreover have "A_B.MkIde (\a. A_B.Map x a) = x" using x A_B.ide_char A_B.MkArr_Map A_B.in_homE A_B.ide_in_hom by metis moreover have "A_B.MkIde (\a. A_B.Map x' a) = x'" using x' A_B.ide_char A_B.MkArr_Map A_B.in_homE A_B.ide_in_hom by metis ultimately show ?thesis using A_B.dom_char A_B.cod_char by auto qed have Fun_f: "\a. A.ide a \ A_B.Map ?f a = (THE fa. ?P a fa)" using f_in_hom \.map_simp_ide by fastforce have cones_map_f: "cones_map ?f \ = \'" using AaPa Fun_f at_ide_is_diagram assms(2) x x' cone_\ cone_\' f_in_hom Fun_f cones_map_pointwise by presburger show "\?f : x' \\<^sub>[\<^sub>A\<^sub>,\<^sub>B\<^sub>] x\ \ cones_map ?f \ = \'" using f_in_hom cones_map_f by auto show "\f'. \f' : x' \\<^sub>[\<^sub>A\<^sub>,\<^sub>B\<^sub>] x\ \ cones_map f' \ = \' \ f' = ?f" proof - fix f' assume f': "\f' : x' \\<^sub>[\<^sub>A\<^sub>,\<^sub>B\<^sub>] x\ \ cones_map f' \ = \'" have 0: "\a. A.ide a \ diagram.cones_map J B (at a D) (A_B.Map f' a) (at a \) = at a \'" using f' cone_\ cone_\' cones_map_pointwise by blast have "f' = A_B.MkArr (A_B.Dom f') (A_B.Cod f') (A_B.Map f')" using f' A_B.MkArr_Map by auto also have "... = ?f" proof (intro A_B.MkArr_eqI) show "A_B.arr (A_B.MkArr (A_B.Dom f') (A_B.Cod f') (A_B.Map f'))" using f' calculation by blast show 1: "A_B.Dom f' = A_B.Map x'" using f' A_B.Map_dom by auto show 2: "A_B.Cod f' = A_B.Map x" using f' A_B.Map_cod by auto show "A_B.Map f' = \.map" proof (intro NaturalTransformation.eqI) show "natural_transformation A B (A_B.Map x') (A_B.Map x) \.map" .. show "natural_transformation A B (A_B.Map x') (A_B.Map x) (A_B.Map f')" using f' 1 2 A_B.arr_char [of f'] by auto show "\a. A.ide a \ A_B.Map f' a = \.map a" proof - fix a assume a: "A.ide a" interpret Da: diagram J B \at a D\ using a at_ide_is_diagram by auto interpret Fun_f': natural_transformation A B \A_B.Dom f'\ \A_B.Cod f'\ \A_B.Map f'\ using f' A_B.arr_char by fast have "A_B.Map f' a \ B.hom (A_B.Map x' a) (A_B.Map x a)" using a f' Fun_f'.preserves_hom A.ide_in_hom by auto hence "?P a (A_B.Map f' a)" using a 0 [of a] by simp moreover have "?P a (\.map a)" using a \.map_simp_ide Fun_f AaPa by presburger ultimately show "A_B.Map f' a = \.map a" using a EU by blast qed qed qed finally show "f' = ?f" by auto qed qed qed qed end context functor_category begin text\ A functor category \[A, B]\ has limits of shape @{term[source=true] J} whenever @{term B} has limits of shape @{term[source=true] J}. \ lemma has_limits_of_shape_if_target_does: assumes "category (J :: 'j comp)" and "B.has_limits_of_shape J" shows "has_limits_of_shape J" proof (unfold has_limits_of_shape_def) have "\D. diagram J comp D \ (\x \. limit_cone J comp D x \)" proof - fix D assume D: "diagram J comp D" interpret J: category J using assms(1) by auto interpret JxA: product_category J A .. interpret D: diagram J comp D using D by auto interpret D: diagram_in_functor_category A B J D .. interpret Curry: currying J A B .. text\ Given diagram @{term D} in \[A, B]\, choose for each object \a\ of \A\ a limit cone \(la, \a)\ for \at a D\ in \B\. \ let ?l = "\a. diagram.some_limit J B (D.at a D)" let ?\ = "\a. diagram.some_limit_cone J B (D.at a D)" have l\: "\a. A.ide a \ diagram.limit_cone J B (D.at a D) (?l a) (?\ a)" using B.has_limits_of_shape_def D.at_ide_is_diagram assms(2) diagram.limit_cone_some_limit_cone by blast text\ The choice of limit cones induces a limit functor from \A\ to \B\. \ interpret uncurry_D: diagram JxA.comp B "Curry.uncurry D" proof - interpret "functor" JxA.comp B \Curry.uncurry D\ using D.functor_axioms Curry.uncurry_preserves_functors by simp interpret binary_functor J A B \Curry.uncurry D\ .. show "diagram JxA.comp B (Curry.uncurry D)" .. qed interpret uncurry_D: parametrized_diagram J A B \Curry.uncurry D\ .. let ?L = "uncurry_D.L ?l ?\" let ?P = "uncurry_D.P ?l ?\" interpret L: "functor" A B ?L using l\ uncurry_D.chosen_limits_induce_functor [of ?l ?\] by simp have L_ide: "\a. A.ide a \ ?L a = ?l a" using uncurry_D.L_ide [of ?l ?\] l\ by blast have L_arr: "\a. A.arr a \ (\!f. ?P a f) \ ?P a (?L a)" using uncurry_D.L_arr [of ?l ?\] l\ by blast text\ The functor \L\ extends to a functor \L'\ from \JxA\ to \B\ that is constant on \J\. \ let ?L' = "\ja. if JxA.arr ja then ?L (snd ja) else B.null" let ?P' = "\ja. ?P (snd ja)" interpret L': "functor" JxA.comp B ?L' apply unfold_locales using L.preserves_arr L.preserves_dom L.preserves_cod apply auto[4] using L.preserves_comp JxA.comp_char by (elim JxA.seqE, auto) have "\ja. JxA.arr ja \ (\!f. ?P' ja f) \ ?P' ja (?L' ja)" proof - fix ja assume ja: "JxA.arr ja" have "A.arr (snd ja)" using ja by blast thus "(\!f. ?P' ja f) \ ?P' ja (?L' ja)" using ja L_arr by presburger qed hence L'_arr: "\ja. JxA.arr ja \ ?P' ja (?L' ja)" by blast have L'_ide: "\ja. \ J.arr (fst ja); A.ide (snd ja) \ \ ?L' ja = ?l (snd ja)" using L_ide l\ by force have L'_arr_map: "\ja. JxA.arr ja \ uncurry_D.P ?l ?\ (snd ja) (uncurry_D.L ?l ?\ (snd ja))" using L'_arr by presburger text\ The map that takes an object \(j, a)\ of \JxA\ to the component \\ a j\ of the limit cone \\ a\ is a natural transformation from \L\ to uncurry \D\. \ let ?\' = "\ja. ?\ (snd ja) (fst ja)" interpret \': transformation_by_components JxA.comp B ?L' \Curry.uncurry D\ ?\' proof fix ja assume ja: "JxA.ide ja" let ?j = "fst ja" let ?a = "snd ja" interpret \a: limit_cone J B \D.at ?a D\ \?l ?a\ \?\ ?a\ using ja l\ by blast show "\?\' ja : ?L' ja \\<^sub>B Curry.uncurry D ja\" using ja L'_ide [of ja] by force next fix ja assume ja: "JxA.arr ja" let ?j = "fst ja" let ?a = "snd ja" have j: "J.arr ?j" using ja by simp have a: "A.arr ?a" using ja by simp interpret D_dom_a: diagram J B \D.at (A.dom ?a) D\ using a D.at_ide_is_diagram by auto interpret D_cod_a: diagram J B \D.at (A.cod ?a) D\ using a D.at_ide_is_diagram by auto interpret Da: natural_transformation J B \D.at (A.dom ?a) D\ \D.at (A.cod ?a) D\ \D.at ?a D\ using a D.functor_axioms D.functor_at_arr_is_transformation by simp interpret \_dom_a: limit_cone J B \D.at (A.dom ?a) D\ \?l (A.dom ?a)\ \?\ (A.dom ?a)\ using a l\ by simp interpret \_cod_a: limit_cone J B \D.at (A.cod ?a) D\ \?l (A.cod ?a)\ \?\ (A.cod ?a)\ using a l\ by simp interpret Dao\_dom_a: vertical_composite J B \_dom_a.A.map \D.at (A.dom ?a) D\ \D.at (A.cod ?a) D\ \?\ (A.dom ?a)\ \D.at ?a D\ .. interpret Dao\_dom_a: cone J B \D.at (A.cod ?a) D\ \?l (A.dom ?a)\ Dao\_dom_a.map .. show "?\' (JxA.cod ja) \\<^sub>B ?L' ja = B (Curry.uncurry D ja) (?\' (JxA.dom ja))" proof - have "?\' (JxA.cod ja) \\<^sub>B ?L' ja = ?\ (A.cod ?a) (J.cod ?j) \\<^sub>B ?L' ja" using ja by fastforce also have "... = D_cod_a.cones_map (?L' ja) (?\ (A.cod ?a)) (J.cod ?j)" using ja L'_arr_map [of ja] \_cod_a.cone_axioms by auto also have "... = Dao\_dom_a.map (J.cod ?j)" using ja \_cod_a.induced_arrowI Dao\_dom_a.cone_axioms L'_arr by presburger also have "... = D.at ?a D (J.cod ?j) \\<^sub>B D_dom_a.some_limit_cone (J.cod ?j)" using ja Dao\_dom_a.map_simp_ide by fastforce also have "... = D.at ?a D (J.cod ?j) \\<^sub>B D.at (A.dom ?a) D ?j \\<^sub>B ?\' (JxA.dom ja)" using ja \_dom_a.naturality \_dom_a.ide_apex apply simp by (metis B.comp_arr_ide \_dom_a.preserves_reflects_arr) also have "... = (D.at ?a D (J.cod ?j) \\<^sub>B D.at (A.dom ?a) D ?j) \\<^sub>B ?\' (JxA.dom ja)" using j ja B.comp_assoc by presburger also have "... = B (D.at ?a D ?j) (?\' (JxA.dom ja))" using a j ja Map_comp A.comp_arr_dom D.as_nat_trans.is_natural_2 by simp also have "... = Curry.uncurry D ja \\<^sub>B ?\' (JxA.dom ja)" using Curry.uncurry_def by simp finally show ?thesis by auto qed qed text\ Since \\'\ is constant on \J\, \curry \'\ is a cone over \D\. \ interpret constL: constant_functor J comp \MkIde ?L\ using L.as_nat_trans.natural_transformation_axioms MkArr_in_hom ide_in_hom L.functor_axioms by unfold_locales blast (* TODO: This seems a little too involved. *) have curry_L': "constL.map = Curry.curry ?L' ?L' ?L'" proof fix j have "\J.arr j \ constL.map j = Curry.curry ?L' ?L' ?L' j" using Curry.curry_def constL.is_extensional by simp moreover have "J.arr j \ constL.map j = Curry.curry ?L' ?L' ?L' j" using Curry.curry_def constL.value_is_ide in_homE ide_in_hom by auto ultimately show "constL.map j = Curry.curry ?L' ?L' ?L' j" by blast qed hence uncurry_constL: "Curry.uncurry constL.map = ?L'" using L'.as_nat_trans.natural_transformation_axioms Curry.uncurry_curry by simp interpret curry_\': natural_transformation J comp constL.map D \Curry.curry ?L' (Curry.uncurry D) \'.map\ proof - have "Curry.curry (Curry.uncurry D) (Curry.uncurry D) (Curry.uncurry D) = D" using Curry.curry_uncurry D.functor_axioms D.as_nat_trans.natural_transformation_axioms by blast thus "natural_transformation J comp constL.map D (Curry.curry ?L' (Curry.uncurry D) \'.map)" using Curry.curry_preserves_transformations curry_L' \'.natural_transformation_axioms by force qed interpret curry_\': cone J comp D \MkIde ?L\ \Curry.curry ?L' (Curry.uncurry D) \'.map\ .. text\ The value of \curry_\'\ at each object \a\ of \A\ is the limit cone \\ a\, hence \curry_\'\ is a limit cone. \ have 1: "\a. A.ide a \ D.at a (Curry.curry ?L' (Curry.uncurry D) \'.map) = ?\ a" proof - fix a assume a: "A.ide a" have "D.at a (Curry.curry ?L' (Curry.uncurry D) \'.map) = (\j. Curry.uncurry (Curry.curry ?L' (Curry.uncurry D) \'.map) (j, a))" using a by simp moreover have "... = (\j. \'.map (j, a))" using a Curry.uncurry_curry \'.natural_transformation_axioms by simp moreover have "... = ?\ a" proof (intro NaturalTransformation.eqI) interpret \a: limit_cone J B \D.at a D\ \?l a\ \?\ a\ using a l\ by simp interpret \': binary_functor_transformation J A B ?L' \Curry.uncurry D\ \'.map .. show "natural_transformation J B \a.A.map (D.at a D) (?\ a)" .. show "natural_transformation J B \a.A.map (D.at a D) (\j. \'.map (j, a))" proof - have "\a.A.map = (\j. ?L' (j, a))" using a \a.A.map_def L'_ide by auto thus ?thesis using a \'.fixing_ide_gives_natural_transformation_2 by simp qed fix j assume j: "J.ide j" show "\'.map (j, a) = ?\ a j" using a j \'.map_simp_ide by simp qed ultimately show "D.at a (Curry.curry ?L' (Curry.uncurry D) \'.map) = ?\ a" by simp qed hence 2: "\a. A.ide a \ diagram.limit_cone J B (D.at a D) (?l a) (D.at a (Curry.curry ?L' (Curry.uncurry D) \'.map))" using l\ by simp hence "limit_cone J comp D (MkIde ?L) (Curry.curry ?L' (Curry.uncurry D) \'.map)" using 1 2 L.functor_axioms L_ide curry_\'.cone_axioms curry_L' D.cone_is_limit_if_pointwise_limit by simp thus "\x \. limit_cone J comp D x \" by blast qed thus "\D. diagram J comp D \ (\x \. limit_cone J comp D x \)" by blast qed lemma has_limits_if_target_does: assumes "B.has_limits (undefined :: 'j)" shows "has_limits (undefined :: 'j)" using assms B.has_limits_def has_limits_def has_limits_of_shape_if_target_does by fast end section "The Yoneda Functor Preserves Limits" text\ In this section, we show that the Yoneda functor from \C\ to \[Cop, S]\ preserves limits. \ context yoneda_functor begin lemma preserves_limits: fixes J :: "'j comp" assumes "diagram J C D" and "diagram.has_as_limit J C D a" shows "diagram.has_as_limit J Cop_S.comp (map o D) (map a)" proof - text\ The basic idea of the proof is as follows: If \\\ is a limit cone in \C\, then for every object \a'\ of \Cop\ the evaluation of \Y o \\ at \a'\ is a limit cone in \S\. By the results on limits in functor categories, this implies that \Y o \\ is a limit cone in \[Cop, S]\. \ interpret J: category J using assms(1) diagram_def by auto interpret D: diagram J C D using assms(1) by auto from assms(2) obtain \ where \: "D.limit_cone a \" by blast interpret \: limit_cone J C D a \ using \ by auto have a: "C.ide a" using \.ide_apex by auto interpret YoD: diagram J Cop_S.comp \map o D\ using D.diagram_axioms functor_axioms preserves_diagrams [of J D] by simp interpret YoD: diagram_in_functor_category Cop.comp S J \map o D\ .. interpret Yo\: cone J Cop_S.comp \map o D\ \map a\ \map o \\ using \.cone_axioms preserves_cones by blast have "\a'. C.ide a' \ limit_cone J S (YoD.at a' (map o D)) (Cop_S.Map (map a) a') (YoD.at a' (map o \))" proof - fix a' assume a': "C.ide a'" interpret A': constant_functor J C a' using a' by (unfold_locales, auto) interpret YoD_a': diagram J S \YoD.at a' (map o D)\ using a' YoD.at_ide_is_diagram by simp interpret Yo\_a': cone J S \YoD.at a' (map o D)\ \Cop_S.Map (map a) a'\ \YoD.at a' (map o \)\ using a' YoD.cone_at_ide_is_cone Yo\.cone_axioms by fastforce have eval_at_ide: "\j. J.ide j \ YoD.at a' (map \ D) j = Hom.map (a', D j)" proof - fix j assume j: "J.ide j" have "YoD.at a' (map \ D) j = Cop_S.Map (map (D j)) a'" using a' j YoD.at_simp YoD.preserves_arr [of j] by auto also have "... = Y (D j) a'" using Y_def by simp also have "... = Hom.map (a', D j)" using a' j D.preserves_arr by simp finally show "YoD.at a' (map \ D) j = Hom.map (a', D j)" by auto qed have eval_at_arr: "\j. J.arr j \ YoD.at a' (map \ \) j = Hom.map (a', \ j)" proof - fix j assume j: "J.arr j" have "YoD.at a' (map \ \) j = Cop_S.Map ((map o \) j) a'" using a' j YoD.at_simp [of a' j "map o \"] preserves_arr by fastforce also have "... = Y (\ j) a'" using Y_def by simp also have "... = Hom.map (a', \ j)" using a' j by simp finally show "YoD.at a' (map \ \) j = Hom.map (a', \ j)" by auto qed have Fun_map_a_a': "Cop_S.Map (map a) a' = Hom.map (a', a)" using a a' map_simp preserves_arr [of a] by simp show "limit_cone J S (YoD.at a' (map o D)) (Cop_S.Map (map a) a') (YoD.at a' (map o \))" proof fix x \ assume \: "YoD_a'.cone x \" interpret \: cone J S \YoD.at a' (map o D)\ x \ using \ by auto have x: "S.ide x" using \.ide_apex by simp text\ For each object \j\ of \J\, the component \\ j\ is an arrow in \S.hom x (Hom.map (a', D j))\. Each element \e \ S.set x\ therefore determines an arrow \\ (a', D j) (S.Fun (\ j) e) \ C.hom a' (D j)\. These arrows are the components of a cone \\ e\ over @{term D} with apex @{term a'}. \ have \j: "\j. J.ide j \ \\ j : x \\<^sub>S Hom.map (a', D j)\" using eval_at_ide \.preserves_hom J.ide_in_hom by force have \: "\e. e \ S.set x \ transformation_by_components J C A'.map D (\j. \ (a', D j) (S.Fun (\ j) e))" proof - fix e assume e: "e \ S.set x" show "transformation_by_components J C A'.map D (\j. \ (a', D j) (S.Fun (\ j) e))" proof fix j assume j: "J.ide j" show "\\ (a', D j) (S.Fun (\ j) e) : A'.map j \ D j\" using e j S.Fun_mapsto [of "\ j"] A'.preserves_ide Hom.set_map eval_at_ide Hom.\_mapsto [of "A'.map j" "D j"] by force next fix j assume j: "J.arr j" show "\ (a', D (J.cod j)) (S.Fun (\ (J.cod j)) e) \ A'.map j = D j \ \ (a', D (J.dom j)) (S.Fun (\ (J.dom j)) e)" proof - have "\ (a', D (J.cod j)) (S.Fun (\ (J.cod j)) e) \ A'.map j = \ (a', D (J.cod j)) (S.Fun (\ (J.cod j)) e) \ a'" using A'.map_simp j by simp also have "... = \ (a', D (J.cod j)) (S.Fun (\ (J.cod j)) e)" proof - have "\ (a', D (J.cod j)) (S.Fun (\ (J.cod j)) e) \ C.hom a' (D (J.cod j))" using a' e j Hom.\_mapsto [of "A'.map j" "D (J.cod j)"] A'.map_simp S.Fun_mapsto [of "\ (J.cod j)"] Hom.set_map eval_at_ide by auto thus ?thesis using C.comp_arr_dom by fastforce qed also have "... = \ (a', D (J.cod j)) (S.Fun (Y (D j) a') (S.Fun (\ (J.dom j)) e))" proof - have "S.Fun (Y (D j) a') (S.Fun (\ (J.dom j)) e) = (S.Fun (Y (D j) a') o S.Fun (\ (J.dom j))) e" by simp also have "... = S.Fun (Y (D j) a' \\<^sub>S \ (J.dom j)) e" using a' e j Y_arr_ide(1) S.in_homE \j eval_at_ide S.Fun_comp by force also have "... = S.Fun (\ (J.cod j)) e" using a' j x \.is_natural_2 \.A.map_simp S.comp_arr_dom J.arr_cod_iff_arr J.cod_cod YoD.preserves_arr \.is_natural_1 YoD.at_simp by auto finally have "S.Fun (Y (D j) a') (S.Fun (\ (J.dom j)) e) = S.Fun (\ (J.cod j)) e" by auto thus ?thesis by simp qed also have "... = D j \ \ (a', D (J.dom j)) (S.Fun (\ (J.dom j)) e)" proof - have "S.Fun (Y (D j) a') (S.Fun (\ (J.dom j)) e) = \ (a', D (J.cod j)) (D j \ \ (a', D (J.dom j)) (S.Fun (\ (J.dom j)) e))" proof - have "S.Fun (\ (J.dom j)) e \ Hom.set (a', D (J.dom j))" using a' e j \j S.Fun_mapsto [of "\ (J.dom j)"] Hom.set_map YoD.at_simp eval_at_ide by auto moreover have "C.arr (\ (a', D (J.dom j)) (S.Fun (\ (J.dom j)) e)) \ C.dom (\ (a', D (J.dom j)) (S.Fun (\ (J.dom j)) e)) = a'" using a' e j \j S.Fun_mapsto [of "\ (J.dom j)"] Hom.set_map eval_at_ide Hom.\_mapsto [of a' "D (J.dom j)"] by auto ultimately show ?thesis using a' e j Hom.Fun_map C.comp_arr_dom by force qed moreover have "D j \ \ (a', D (J.dom j)) (S.Fun (\ (J.dom j)) e) \ C.hom a' (D (J.cod j))" proof - have "\ (a', D (J.dom j)) (S.Fun (\ (J.dom j)) e) \ C.hom a' (D (J.dom j))" using a' e j Hom.\_mapsto [of a' "D (J.dom j)"] eval_at_ide S.Fun_mapsto [of "\ (J.dom j)"] Hom.set_map by auto thus ?thesis using j D.preserves_hom by blast qed ultimately show ?thesis using a' j Hom.\_\ by simp qed finally show ?thesis by auto qed qed qed let ?\ = "\e. transformation_by_components.map J C A'.map (\j. \ (a', D j) (S.Fun (\ j) e))" have cone_\e: "\e. e \ S.set x \ D.cone a' (?\ e)" proof - fix e assume e: "e \ S.set x" interpret \e: transformation_by_components J C A'.map D \\j. \ (a', D j) (S.Fun (\ j) e)\ using e \ by blast show "D.cone a' (?\ e)" .. qed text\ Since \\ e\ is a cone for each element \e\ of \S.set x\, by the universal property of the limit cone \\\ there is a unique arrow \fe \ C.hom a' a\ that transforms \\\ to \\ e\. \ have ex_fe: "\e. e \ S.set x \ \!fe. \fe : a' \ a\ \ D.cones_map fe \ = ?\ e" using cone_\e \.is_universal by simp text\ The map taking \e \ S.set x\ to \fe \ C.hom a' a\ determines an arrow \f \ S.hom x (Hom (a', a))\ that transforms the cone obtained by evaluating \Y o \\ at \a'\ to the cone \\\. \ let ?f = "S.mkArr (S.set x) (Hom.set (a', a)) (\e. \ (a', a) (\.induced_arrow a' (?\ e)))" have 0: "(\e. \ (a', a) (\.induced_arrow a' (?\ e))) \ S.set x \ Hom.set (a', a)" proof fix e assume e: "e \ S.set x" interpret \e: cone J C D a' \?\ e\ using e cone_\e by simp have "\.induced_arrow a' (?\ e) \ C.hom a' a" using a a' e ex_fe \.induced_arrowI \e.cone_axioms by simp thus "\ (a', a) (\.induced_arrow a' (?\ e)) \ Hom.set (a', a)" using a a' Hom.\_mapsto by auto qed have f: "\?f : x \\<^sub>S Hom.map (a', a)\" proof - have "(\e. \ (a', a) (\.induced_arrow a' (?\ e))) \ S.set x \ Hom.set (a', a)" proof fix e assume e: "e \ S.set x" interpret \e: cone J C D a' \?\ e\ using e cone_\e by simp have "\.induced_arrow a' (?\ e) \ C.hom a' a" using a a' e ex_fe \.induced_arrowI \e.cone_axioms by simp thus "\ (a', a) (\.induced_arrow a' (?\ e)) \ Hom.set (a', a)" using a a' Hom.\_mapsto by auto qed - thus ?thesis + moreover have "setp (Hom.set (a', a))" + using a a' Hom.small_homs + by (metis Fun_map_a_a' Hom.map_ide S.arr_mkIde S.ideD(1) Yo\_a'.ide_apex) + ultimately show ?thesis using a a' x \.ide_apex S.mkArr_in_hom [of "S.set x" "Hom.set (a', a)"] Hom.set_subset_Univ S.mkIde_set by simp qed have "YoD_a'.cones_map ?f (YoD.at a' (map o \)) = \" proof (intro NaturalTransformation.eqI) show "natural_transformation J S \.A.map (YoD.at a' (map o D)) \" using \.natural_transformation_axioms by auto have 1: "S.cod ?f = Cop_S.Map (map a) a'" using f Fun_map_a_a' by force interpret YoD_a'of: cone J S \YoD.at a' (map o D)\ x \YoD_a'.cones_map ?f (YoD.at a' (map o \))\ proof - have "YoD_a'.cone (S.cod ?f) (YoD.at a' (map o \))" using a a' f Yo\_a'.cone_axioms preserves_arr [of a] by auto hence "YoD_a'.cone (S.dom ?f) (YoD_a'.cones_map ?f (YoD.at a' (map o \)))" using f YoD_a'.cones_map_mapsto S.arrI by blast thus "cone J S (YoD.at a' (map o D)) x (YoD_a'.cones_map ?f (YoD.at a' (map o \)))" using f by auto qed show "natural_transformation J S \.A.map (YoD.at a' (map o D)) (YoD_a'.cones_map ?f (YoD.at a' (map o \)))" .. fix j assume j: "J.ide j" have "YoD_a'.cones_map ?f (YoD.at a' (map o \)) j = YoD.at a' (map o \) j \\<^sub>S ?f" using f j Fun_map_a_a' Yo\_a'.cone_axioms by fastforce also have "... = \ j" proof (intro S.arr_eqI) show "S.par (YoD.at a' (map o \) j \\<^sub>S ?f) (\ j)" using 1 f j x YoD_a'.preserves_hom by fastforce show "S.Fun (YoD.at a' (map o \) j \\<^sub>S ?f) = S.Fun (\ j)" proof fix e have "e \ S.set x \ S.Fun (YoD.at a' (map o \) j \\<^sub>S ?f) e = S.Fun (\ j) e" using 1 f j x S.Fun_mapsto [of "\ j"] \.A.map_simp extensional_arb [of "S.Fun (\ j)"] by auto moreover have "e \ S.set x \ S.Fun (YoD.at a' (map o \) j \\<^sub>S ?f) e = S.Fun (\ j) e" proof - assume e: "e \ S.set x" interpret \e: transformation_by_components J C A'.map D \\j. \ (a', D j) (S.Fun (\ j) e)\ using e \ by blast interpret \e: cone J C D a' \?\ e\ using e cone_\e by simp have induced_arrow: "\.induced_arrow a' (?\ e) \ C.hom a' a" using a a' e ex_fe \.induced_arrowI \e.cone_axioms by simp have "S.Fun (YoD.at a' (map o \) j \\<^sub>S ?f) e = restrict (S.Fun (YoD.at a' (map o \) j) o S.Fun ?f) (S.set x) e" using 1 e f j S.Fun_comp YoD_a'.preserves_hom by force also have "... = (\ (a', D j) o C (\ j) o \ (a', a)) (S.Fun ?f e)" using j a' f e Hom.map_simp_2 S.Fun_mkArr Hom.preserves_arr [of "(a', \ j)"] - eval_at_arr S.arr_mkArr + eval_at_arr by (elim S.in_homE, auto) also have "... = (\ (a', D j) o C (\ j) o \ (a', a)) (\ (a', a) (\.induced_arrow a' (?\ e)))" using e f S.Fun_mkArr by fastforce also have "... = \ (a', D j) (D.cones_map (\.induced_arrow a' (?\ e)) \ j)" using a a' e j 0 Hom.\_\ induced_arrow \.cone_axioms by auto also have "... = \ (a', D j) (?\ e j)" using \.induced_arrowI \e.cone_axioms by fastforce also have "... = \ (a', D j) (\ (a', D j) (S.Fun (\ j) e))" using j \e.map_def [of j] by simp also have "... = S.Fun (\ j) e" proof - have "S.Fun (\ j) e \ Hom.set (a', D j)" using a' e j S.Fun_mapsto [of "\ j"] eval_at_ide Hom.set_map by auto thus ?thesis using a' j Hom.\_\ C.ide_in_hom J.ide_in_hom by blast qed finally show "S.Fun (YoD.at a' (map o \) j \\<^sub>S ?f) e = S.Fun (\ j) e" by auto qed ultimately show "S.Fun (YoD.at a' (map o \) j \\<^sub>S ?f) e = S.Fun (\ j) e" by auto qed qed finally show "YoD_a'.cones_map ?f (YoD.at a' (map o \)) j = \ j" by auto qed hence ff: "?f \ S.hom x (Hom.map (a', a)) \ YoD_a'.cones_map ?f (YoD.at a' (map o \)) = \" using f by auto text\ Any other arrow \f' \ S.hom x (Hom.map (a', a))\ that transforms the cone obtained by evaluating \Y o \\ at @{term a'} to the cone @{term \}, must equal \f\, showing that \f\ is unique. \ moreover have "\f'. \f' : x \\<^sub>S Hom.map (a', a)\ \ YoD_a'.cones_map f' (YoD.at a' (map o \)) = \ \ f' = ?f" proof - fix f' assume f': "\f' : x \\<^sub>S Hom.map (a', a)\ \ YoD_a'.cones_map f' (YoD.at a' (map o \)) = \" show "f' = ?f" proof (intro S.arr_eqI) show par: "S.par f' ?f" using f f' by (elim S.in_homE, auto) show "S.Fun f' = S.Fun ?f" proof fix e have "e \ S.set x \ S.Fun f' e = S.Fun ?f e" using f f' S.Fun_in_terms_of_comp by fastforce moreover have "e \ S.set x \ S.Fun f' e = S.Fun ?f e" proof - assume e: "e \ S.set x" have fe: "S.Fun ?f e \ Hom.set (a', a)" - using e f S.arr_mkArr par by auto + using e f par by auto have f'e: "S.Fun f' e \ Hom.set (a', a)" using a a' e f' S.Fun_mapsto Hom.set_map by fastforce have 1: "\\ (a', a) (S.Fun f' e) : a' \ a\" using a a' e f' f'e S.Fun_mapsto Hom.\_mapsto Hom.set_map by blast have 2: "\\ (a', a) (S.Fun ?f e) : a' \ a\" using a a' e f' fe S.Fun_mapsto Hom.\_mapsto Hom.set_map by blast interpret \ofe: cone J C D a' \D.cones_map (\ (a', a) (S.Fun ?f e)) \\ proof - have "D.cones_map (\ (a', a) (S.Fun ?f e)) \ D.cones a \ D.cones a'" using 2 D.cones_map_mapsto [of "\ (a', a) (S.Fun ?f e)"] by (elim C.in_homE, auto) thus "cone J C D a' (D.cones_map (\ (a', a) (S.Fun ?f e)) \)" using \.cone_axioms by blast qed have A: "\h j. h \ C.hom a' a \ J.arr j \ S.Fun (YoD.at a' (map o \) j) (\ (a', a) h) = \ (a', D (J.cod j)) (\ j \ h)" proof - fix h j assume j: "J.arr j" assume h: "h \ C.hom a' a" have "S.Fun (YoD.at a' (map o \) j) (\ (a', a) h) = (\ (a', D (J.cod j)) \ C (\ j) \ \ (a', a)) (\ (a', a) h)" proof - have "S.Fun (YoD.at a' (map o \) j) = restrict (\ (a', D (J.cod j)) \ C (\ j) \ \ (a', a)) (Hom.set (a', a))" proof - have "S.Fun (YoD.at a' (map o \) j) = S.Fun (Y (\ j) a')" using a' j YoD.at_simp Y_def Yo\.preserves_reflects_arr [of j] by simp also have "... = restrict (\ (a', D (J.cod j)) \ C (\ j) \ \ (a', a)) (Hom.set (a', a))" using a' j \.preserves_hom [of j "J.dom j" "J.cod j"] Y_arr_ide [of a' "\ j" a "D (J.cod j)"] \.A.map_simp S.Fun_mkArr by fastforce finally show ?thesis by blast qed thus ?thesis using a a' h Hom.\_mapsto by auto qed also have "... = \ (a', D (J.cod j)) (\ j \ h)" using a a' h Hom.\_\ by simp finally show "S.Fun (YoD.at a' (map o \) j) (\ (a', a) h) = \ (a', D (J.cod j)) (\ j \ h)" by auto qed have "D.cones_map (\ (a', a) (S.Fun f' e)) \ = D.cones_map (\ (a', a) (S.Fun ?f e)) \" proof fix j have "\J.arr j \ D.cones_map (\ (a', a) (S.Fun f' e)) \ j = D.cones_map (\ (a', a) (S.Fun ?f e)) \ j" using 1 2 \.cone_axioms by (elim C.in_homE, auto) moreover have "J.arr j \ D.cones_map (\ (a', a) (S.Fun f' e)) \ j = D.cones_map (\ (a', a) (S.Fun ?f e)) \ j" proof - assume j: "J.arr j" have "D.cones_map (\ (a', a) (S.Fun f' e)) \ j = \ j \ \ (a', a) (S.Fun f' e)" using j 1 \.cone_axioms by auto also have "... = \ (a', D (J.cod j)) (S.Fun (\ j) e)" proof - have "\ (a', D (J.cod j)) (S.Fun (YoD.at a' (map o \) j) (S.Fun f' e)) = \ (a', D (J.cod j)) (\ (a', D (J.cod j)) (\ j \ \ (a', a) (S.Fun f' e)))" using j a a' f'e A Hom.\_\ Hom.\_mapsto by force moreover have "\ j \ \ (a', a) (S.Fun f' e) \ C.hom a' (D (J.cod j))" using a a' j f'e Hom.\_mapsto \.preserves_hom [of j "J.dom j" "J.cod j"] \.A.map_simp by auto moreover have "S.Fun (YoD.at a' (map o \) j) (S.Fun f' e) = S.Fun (\ j) e" using Fun_map_a_a' a a' j f' e x Yo\_a'.A.map_simp eval_at_ide Yo\_a'.cone_axioms by auto ultimately show ?thesis using a a' Hom.\_\ by auto qed also have "... = \ j \ \ (a', a) (S.Fun ?f e)" proof - have "S.Fun (YoD.at a' (map o \) j) (S.Fun ?f e) = \ (a', D (J.cod j)) (\ j \ \ (a', a) (S.Fun ?f e))" using j a a' fe A [of "\ (a', a) (S.Fun ?f e)" j] Hom.\_\ Hom.\_mapsto by auto hence "\ (a', D (J.cod j)) (S.Fun (YoD.at a' (map o \) j) (S.Fun ?f e)) = \ (a', D (J.cod j)) (\ (a', D (J.cod j)) (\ j \ \ (a', a) (S.Fun ?f e)))" by simp moreover have "\ j \ \ (a', a) (S.Fun ?f e) \ C.hom a' (D (J.cod j))" using a a' j fe Hom.\_mapsto \.preserves_hom [of j "J.dom j" "J.cod j"] \.A.map_simp by auto moreover have "S.Fun (YoD.at a' (map o \) j) (S.Fun ?f e) = S.Fun (\ j) e" proof - have "S.Fun (YoD.at a' (map o \) j) (S.Fun ?f e) = (S.Fun (YoD.at a' (map o \) j) o S.Fun ?f) e" by simp also have "... = S.Fun (YoD.at a' (map o \) j \\<^sub>S ?f) e" using Fun_map_a_a' a a' j f e x Yo\_a'.A.map_simp eval_at_ide by auto also have "... = S.Fun (\ j) e" proof - have "YoD.at a' (map o \) j \\<^sub>S ?f = YoD_a'.cones_map ?f (YoD.at a' (map o \)) j" using j f Yo\_a'.cone_axioms Fun_map_a_a' by auto thus ?thesis using j ff by argo qed finally show ?thesis by auto qed ultimately show ?thesis using a a' Hom.\_\ by auto qed also have "... = D.cones_map (\ (a', a) (S.Fun ?f e)) \ j" using j 2 \.cone_axioms by force finally show "D.cones_map (\ (a', a) (S.Fun f' e)) \ j = D.cones_map (\ (a', a) (S.Fun ?f e)) \ j" by auto qed ultimately show "D.cones_map (\ (a', a) (S.Fun f' e)) \ j = D.cones_map (\ (a', a) (S.Fun ?f e)) \ j" by auto qed hence "\ (a', a) (S.Fun f' e) = \ (a', a) (S.Fun ?f e)" using 1 2 \ofe.cone_axioms \.cone_axioms \.is_universal by blast hence "\ (a', a) (\ (a', a) (S.Fun f' e)) = \ (a', a) (\ (a', a) (S.Fun ?f e))" by simp thus "S.Fun f' e = S.Fun ?f e" using a a' fe f'e Hom.\_\ by force qed ultimately show "S.Fun f' e = S.Fun ?f e" by auto qed qed qed ultimately have "\!f. \f : x \\<^sub>S Hom.map (a', a)\ \ YoD_a'.cones_map f (YoD.at a' (map o \)) = \" using ex1I [of "\f. S.in_hom x (Hom.map (a', a)) f \ YoD_a'.cones_map f (YoD.at a' (map o \)) = \"] by blast thus "\!f. \f : x \\<^sub>S Cop_S.Map (map a) a'\ \ YoD_a'.cones_map f (YoD.at a' (map o \)) = \" using a a' Y_def by simp qed qed thus "YoD.has_as_limit (map a)" using YoD.cone_is_limit_if_pointwise_limit Yo\.cone_axioms by auto qed end end diff --git a/thys/Category3/ROOT b/thys/Category3/ROOT --- a/thys/Category3/ROOT +++ b/thys/Category3/ROOT @@ -1,12 +1,11 @@ chapter AFP session "Category3" (AFP) = "HOL-Library" + options [timeout = 3000, names_unique = false] sessions - "HOL-Cardinals" - "HereditarilyFinite" + "HereditarilyFinite" "ZFC_in_HOL" theories - Limit Subcategory EquivalenceOfCategories HFSetCat + EquivalenceOfCategories HFSetCat ZFC_SetCat document_files "root.bib" "root.tex" diff --git a/thys/Category3/SetCat.thy b/thys/Category3/SetCat.thy --- a/thys/Category3/SetCat.thy +++ b/thys/Category3/SetCat.thy @@ -1,874 +1,1347 @@ (* Title: SetCat Author: Eugene W. Stark , 2016 Maintainer: Eugene W. Stark *) chapter SetCat theory SetCat imports SetCategory ConcreteCategory begin text\ This theory proves the consistency of the \set_category\ locale by giving a particular concrete construction of an interpretation for it. Applying the general construction given by @{locale concrete_category}, we define arrows to be terms \MkArr A B F\, where \A\ and \B\ are sets and \F\ is an extensional function that maps \A\ to \B\. \ text\ This locale uses an extra dummy parameter just to fix the element type for sets. Without this, a type is used for each interpretation, which makes it impossible to construct set categories whose element types are related to the context. + An additional parameter, @{term Setp}, allows some control over which subsets of + the element type are assumed to correspond to objects of the category. \ locale setcat = fixes dummy :: 'e - and \ :: "'a rel" - assumes cardinal: "Card_order \ \ infinite (Field \)" + and Setp :: "'e set \ bool" + assumes Setp_singleton: "Setp {x}" + and Setp_respects_subset: "A' \ A \ Setp A \ Setp A'" + and union_preserves_Setp: "\ Setp A; Setp B \ \ Setp (A \ B)" begin - lemma finite_imp_card_less: - assumes "finite A" - shows "|A| " - proof - - have "finite (Field |A| )" - using assms by simp - thus ?thesis - using cardinal card_of_Well_order card_order_on_def finite_ordLess_infinite - by blast - qed + lemma finite_imp_Setp: "finite A \ Setp A" + using Setp_singleton + by (metis finite_induct insert_is_Un Setp_respects_subset singleton_insert_inj_eq + union_preserves_Setp) type_synonym 'b arr = "('b set, 'b \ 'b) concrete_category.arr" - interpretation concrete_category \{A :: 'e set. |A| }\ \\A B. extensional A \ (A \ B)\ - \\A. \x \ A. x\ \\C B A g f. compose A g f\ + interpretation S: concrete_category \Collect Setp\ \\A B. extensional A \ (A \ B)\ + \\A. \x \ A. x\ \\C B A g f. compose A g f\ using compose_Id Id_compose apply unfold_locales apply auto[3] apply blast by (metis IntD2 compose_assoc) abbreviation comp :: "'e setcat.arr comp" (infixr "\" 55) - where "comp \ COMP" - notation in_hom ("\_ : _ \ _\") + where "comp \ S.COMP" + notation S.in_hom ("\_ : _ \ _\") + + lemma is_category: + shows "category comp" + using S.category_axioms by simp lemma MkArr_expansion: - assumes "arr f" - shows "f = MkArr (Dom f) (Cod f) (\x \ Dom f. Map f x)" - proof (intro arr_eqI) - show "arr f" by fact - show "arr (MkArr (Dom f) (Cod f) (\x \ Dom f. Map f x))" - using assms arr_char - by (metis (mono_tags, lifting) Int_iff MkArr_Map extensional_restrict) - show "Dom f = Dom (MkArr (Dom f) (Cod f) (\x \ Dom f. Map f x))" + assumes "S.arr f" + shows "f = S.MkArr (S.Dom f) (S.Cod f) (\x \ S.Dom f. S.Map f x)" + proof (intro S.arr_eqI) + show "S.arr f" by fact + show "S.arr (S.MkArr (S.Dom f) (S.Cod f) (\x \ S.Dom f. S.Map f x))" + using assms S.arr_char + by (metis (mono_tags, lifting) Int_iff S.MkArr_Map extensional_restrict) + show "S.Dom f = S.Dom (S.MkArr (S.Dom f) (S.Cod f) (\x \ S.Dom f. S.Map f x))" by simp - show "Cod f = Cod (MkArr (Dom f) (Cod f) (\x \ Dom f. Map f x))" + show "S.Cod f = S.Cod (S.MkArr (S.Dom f) (S.Cod f) (\x \ S.Dom f. S.Map f x))" by simp - show "Map f = Map (MkArr (Dom f) (Cod f) (\x \ Dom f. Map f x))" - using assms arr_char - by (metis (mono_tags, lifting) Int_iff MkArr_Map extensional_restrict) + show "S.Map f = S.Map (S.MkArr (S.Dom f) (S.Cod f) (\x \ S.Dom f. S.Map f x))" + using assms S.arr_char + by (metis (mono_tags, lifting) Int_iff S.MkArr_Map extensional_restrict) qed lemma arr_char: - shows "arr f \ f \ Null \ |Dom f| \ |Cod f| \ - Map f \ extensional (Dom f) \ (Dom f \ Cod f)" - using arr_char by auto + shows "S.arr f \ f \ S.Null \ Setp (S.Dom f) \ Setp (S.Cod f) \ + S.Map f \ extensional (S.Dom f) \ (S.Dom f \ S.Cod f)" + using S.arr_char by auto lemma terminal_char: - shows "terminal a \ (\x. a = MkIde {x})" + shows "S.terminal a \ (\x. a = S.MkIde {x})" proof - show "\x. a = MkIde {x} \ terminal a" + show "\x. a = S.MkIde {x} \ S.terminal a" proof - - assume a: "\x. a = MkIde {x}" - from this obtain x where x: "a = MkIde {x}" by blast - have "terminal (MkIde {x})" + assume a: "\x. a = S.MkIde {x}" + from this obtain x where x: "a = S.MkIde {x}" by blast + have "S.terminal (S.MkIde {x})" proof - show 1: "ide (MkIde {x})" - using finite_imp_card_less ide_MkIde by auto - show "\a. ide a \ \!f. \f : a \ MkIde {x}\" + show 1: "S.ide (S.MkIde {x})" + using finite_imp_Setp S.ide_MkIde by auto + show "\a. S.ide a \ \!f. \f : a \ S.MkIde {x}\" proof fix a :: "'e setcat.arr" - assume a: "ide a" - show "\MkArr (Dom a) {x} (\_\Dom a. x) : a \ MkIde {x}\" + assume a: "S.ide a" + show "\S.MkArr (S.Dom a) {x} (\_\S.Dom a. x) : a \ S.MkIde {x}\" proof - show 2: "arr (MkArr (Dom a) {x} (\_ \ Dom a. x))" - using a 1 arr_MkArr [of "Dom a" "{x}"] ide_char by force - show "dom (MkArr (Dom a) {x} (\_ \ Dom a. x)) = a" - using a 2 dom_MkArr by force - show "cod (MkArr (Dom a) {x} (\_\Dom a. x)) = MkIde {x}" - using 2 cod_MkArr by blast + show 2: "S.arr (S.MkArr (S.Dom a) {x} (\_ \ S.Dom a. x))" + using a 1 S.arr_MkArr [of "S.Dom a" "{x}"] S.ide_char by force + show "S.dom (S.MkArr (S.Dom a) {x} (\_ \ S.Dom a. x)) = a" + using a 2 S.dom_MkArr by force + show "S.cod (S.MkArr (S.Dom a) {x} (\_\S.Dom a. x)) = S.MkIde {x}" + using 2 S.cod_MkArr by blast qed fix f :: "'e setcat.arr" - assume f: "\f : a \ MkIde {x}\" - show "f = MkArr (Dom a) {x} (\_ \ Dom a. x)" + assume f: "\f : a \ S.MkIde {x}\" + show "f = S.MkArr (S.Dom a) {x} (\_ \ S.Dom a. x)" proof - - have 1: "Dom f = Dom a \ Cod f = {x}" - using a f by (metis (mono_tags, lifting) Dom.simps(1) in_hom_char) - moreover have "Map f = (\_ \ Dom a. x)" + have 1: "S.Dom f = S.Dom a \ S.Cod f = {x}" + using a f by (metis (mono_tags, lifting) S.Dom.simps(1) S.in_hom_char) + moreover have "S.Map f = (\_ \ S.Dom a. x)" proof fix z - have "z \ Dom a \ Map f z = (\_ \ Dom a. x) z" + have "z \ S.Dom a \ S.Map f z = (\_ \ S.Dom a. x) z" using f 1 MkArr_expansion - by (metis (mono_tags, lifting) Map.simps(1) in_homE restrict_apply) - moreover have "z \ Dom a \ Map f z = (\_ \ Dom a. x) z" + by (metis (mono_tags, lifting) S.Map.simps(1) S.in_homE restrict_apply) + moreover have "z \ S.Dom a \ S.Map f z = (\_ \ S.Dom a. x) z" using f 1 arr_char [of f] by fastforce - ultimately show "Map f z = (\_ \ Dom a. x) z" by auto + ultimately show "S.Map f z = (\_ \ S.Dom a. x) z" by auto qed ultimately show ?thesis using f MkArr_expansion [of f] by fastforce qed qed qed - thus "terminal a" using x by simp + thus "S.terminal a" using x by simp qed - show "terminal a \ \x. a = MkIde {x}" + show "S.terminal a \ \x. a = S.MkIde {x}" proof - - assume a: "terminal a" - hence ide_a: "ide a" using terminal_def by auto - have 1: "\!x. x \ Dom a" + assume a: "S.terminal a" + hence ide_a: "S.ide a" using S.terminal_def by auto + have 1: "\!x. x \ S.Dom a" proof - - have "Dom a = {} \ \terminal a" + have "S.Dom a = {} \ \S.terminal a" proof - - assume "Dom a = {}" - hence 1: "a = MkIde {}" - using MkIde_Dom' \ide a\ by fastforce - have "\f. f \ hom (MkIde {undefined}) (MkIde ({} :: 'e set)) - \ Map f \ {undefined} \ {}" + assume "S.Dom a = {}" + hence 1: "a = S.MkIde {}" + using S.MkIde_Dom' \S.ide a\ by fastforce + have "\f. f \ S.hom (S.MkIde {undefined}) (S.MkIde ({} :: 'e set)) + \ S.Map f \ {undefined} \ {}" proof - fix f - assume f: "f \ hom (MkIde {undefined}) (MkIde ({} :: 'e set))" - show "Map f \ {undefined} \ {}" - using f MkArr_expansion arr_char [of f] in_hom_char by auto + assume f: "f \ S.hom (S.MkIde {undefined}) (S.MkIde ({} :: 'e set))" + show "S.Map f \ {undefined} \ {}" + using f MkArr_expansion arr_char [of f] S.in_hom_char by auto qed - hence "hom (MkIde {undefined}) a = {}" using 1 by auto - moreover have "ide (MkIde {undefined})" - using finite_imp_card_less - by (metis (mono_tags, lifting) finite.intros(1-2) ide_MkIde mem_Collect_eq) - ultimately show "\terminal a" by blast + hence "S.hom (S.MkIde {undefined}) a = {}" using 1 by auto + moreover have "S.ide (S.MkIde {undefined})" + using finite_imp_Setp + by (metis (mono_tags, lifting) finite.intros(1-2) S.ide_MkIde mem_Collect_eq) + ultimately show "\S.terminal a" by blast qed - moreover have "\x x'. x \ Dom a \ x' \ Dom a \ x \ x' \ \terminal a" + moreover have "\x x'. x \ S.Dom a \ x' \ S.Dom a \ x \ x' \ \S.terminal a" proof - fix x x' - assume 1: "x \ Dom a \ x' \ Dom a \ x \ x'" - let ?f = "MkArr {undefined} (Dom a) (\_ \ {undefined}. x)" - let ?f' = "MkArr {undefined} (Dom a) (\_ \ {undefined}. x')" - have "\?f : MkIde {undefined} \ a\" - proof - show 2: "arr ?f" - proof (intro arr_MkArr) - show "{undefined} \ {A. |A| }" - by (simp add: finite_imp_card_less) - show "Dom a \ {A. |A| }" - using ide_a ide_char by blast - show "(\_ \ {undefined}. x) \ extensional {undefined} \ ({undefined} \ Dom a)" - using 1 by blast + assume 1: "x \ S.Dom a \ x' \ S.Dom a \ x \ x'" + let ?f = "S.MkArr {undefined} (S.Dom a) (\_ \ {undefined}. x)" + let ?f' = "S.MkArr {undefined} (S.Dom a) (\_ \ {undefined}. x')" + have "S.par ?f ?f'" + proof - + have "\?f : S.MkIde {undefined} \ a\" + proof + show 2: "S.arr ?f" + proof (intro S.arr_MkArr) + show "{undefined} \ {A. Setp A}" + by (simp add: finite_imp_Setp) + show "S.Dom a \ {A. Setp A}" + using ide_a S.ide_char by blast + show "(\_ \ {undefined}. x) \ extensional {undefined} \ ({undefined} \ S.Dom a)" + using 1 by blast + qed + show "S.dom ?f = S.MkIde {undefined}" + using 2 S.dom_MkArr by auto + show "S.cod ?f = a" + using 2 S.cod_MkArr ide_a by force qed - show "dom ?f = MkIde {undefined}" - using 2 dom_MkArr by auto - show "cod ?f = a" - using 2 cod_MkArr ide_a by force - qed - moreover have "\?f' : MkIde {undefined} \ a\" - proof - show 2: "arr ?f'" - using 1 ide_a ide_char arr_MkArr [of "{undefined}" "Dom a"] - finite_imp_card_less - proof (intro arr_MkArr) - show "{undefined} \ {A. |A| }" - by (simp add: finite_imp_card_less) - show "Dom a \ {A. |A| }" - using ide_a ide_char by blast - show "(\_ \ {undefined}. x') \ extensional {undefined} \ ({undefined} \ Dom a)" - using 1 by blast + moreover have "\?f' : S.MkIde {undefined} \ a\" + proof + show 2: "S.arr ?f'" + proof (intro S.arr_MkArr) + show "{undefined} \ {A. Setp A}" + by (simp add: finite_imp_Setp) + show "S.Dom a \ {A. Setp A}" + using ide_a S.ide_char by blast + show "(\_ \ {undefined}. x') \ extensional {undefined} \ ({undefined} \ S.Dom a)" + using 1 by blast + qed + show "S.dom ?f' = S.MkIde {undefined}" + using 2 S.dom_MkArr by auto + show "S.cod ?f' = a" + using 2 S.cod_MkArr ide_a by force qed - show "dom ?f' = MkIde {undefined}" - using 2 dom_MkArr by auto - show "cod ?f' = a" - using 2 cod_MkArr ide_a by force + ultimately show ?thesis + by (metis (no_types, lifting) S.in_homE) qed moreover have "?f \ ?f'" - using 1 by (metis arr.inject restrict_apply' singletonI) - ultimately show "\terminal a" - using terminal_arr_unique - by (metis (mono_tags, lifting) in_homE) + using 1 by (metis S.arr.inject restrict_apply' singletonI) + ultimately show "\S.terminal a" + using S.cod_MkArr ide_a S.terminal_arr_unique [of ?f ?f'] by auto qed ultimately show ?thesis using a by auto qed - hence "Dom a = {THE x. x \ Dom a}" - using theI [of "\x. x \ Dom a"] by auto - hence "a = MkIde {THE x. x \ Dom a}" - using a terminal_def by (metis (mono_tags, lifting) MkIde_Dom') - thus "\x. a = MkIde {x}" + hence "S.Dom a = {THE x. x \ S.Dom a}" + using theI [of "\x. x \ S.Dom a"] by auto + hence "a = S.MkIde {THE x. x \ S.Dom a}" + using a S.terminal_def by (metis (mono_tags, lifting) S.MkIde_Dom') + thus "\x. a = S.MkIde {x}" by auto qed qed definition IMG :: "'e setcat.arr \ 'e setcat.arr" - where "IMG f = MkIde (Map f ` Dom f)" + where "IMG f = S.MkIde (S.Map f ` S.Dom f)" - interpretation set_category_data comp IMG .. + interpretation S: set_category_data comp IMG + .. lemma terminal_unity: - shows "terminal unity" - using terminal_char unity_def someI_ex [of terminal] + shows "S.terminal S.unity" + using terminal_char S.unity_def someI_ex [of S.terminal] by (metis (mono_tags, lifting)) text\ - The inverse maps @{term UP} and @{term DOWN} are used to pass back and forth between + The inverse maps @{term arr_of} and @{term elem_of} are used to pass back and forth between the inhabitants of type @{typ 'a} and the corresponding terminal objects. These are exported so that a client of the theory can relate the concrete element type @{typ 'a} to the otherwise abstract arrow type. \ - definition UP :: "'e \ 'e setcat.arr" - where "UP x \ MkIde {x}" + definition arr_of :: "'e \ 'e setcat.arr" + where "arr_of x \ S.MkIde {x}" - definition DOWN :: "'e setcat.arr \ 'e" - where "DOWN t \ the_elem (Dom t)" + definition elem_of :: "'e setcat.arr \ 'e" + where "elem_of t \ the_elem (S.Dom t)" abbreviation U - where "U \ DOWN unity" + where "U \ elem_of S.unity" - lemma UP_mapsto: - shows "UP \ UNIV \ Univ" - using terminal_char UP_def by fast + lemma arr_of_mapsto: + shows "arr_of \ UNIV \ S.Univ" + using terminal_char arr_of_def by fast - lemma DOWN_mapsto: - shows "DOWN \ Univ \ UNIV" + lemma elem_of_mapsto: + shows "elem_of \ Univ \ UNIV" by auto - lemma DOWN_UP [simp]: - shows "DOWN (UP x) = x" - by (simp add: DOWN_def UP_def) + lemma elem_of_arr_of [simp]: + shows "elem_of (arr_of x) = x" + by (simp add: elem_of_def arr_of_def) - lemma UP_DOWN [simp]: - assumes "t \ Univ" - shows "UP (DOWN t) = t" - using assms terminal_char UP_def DOWN_def - by (metis (mono_tags, lifting) mem_Collect_eq DOWN_UP) + lemma arr_of_elem_of [simp]: + assumes "t \ S.Univ" + shows "arr_of (elem_of t) = t" + using assms terminal_char arr_of_def elem_of_def + by (metis (mono_tags, lifting) mem_Collect_eq elem_of_arr_of) - lemma inj_UP: - shows "inj UP" - by (metis DOWN_UP injI) + lemma inj_arr_of: + shows "inj arr_of" + by (metis elem_of_arr_of injI) - lemma bij_UP: - shows "bij_betw UP UNIV Univ" + lemma bij_arr_of: + shows "bij_betw arr_of UNIV S.Univ" proof (intro bij_betwI) - interpret category comp using is_category by auto - show DOWN_UP: "\x :: 'e. DOWN (UP x) = x" by simp - show UP_DOWN: "\t. t \ Univ \ UP (DOWN t) = t" by simp - show "UP \ UNIV \ Univ" using UP_mapsto by auto - show "DOWN \ Collect terminal \ UNIV" by auto + show "\x :: 'e. elem_of (arr_of x) = x" by simp + show "\t. t \ S.Univ \ arr_of (elem_of t) = t" by simp + show "arr_of \ UNIV \ S.Univ" using arr_of_mapsto by auto + show "elem_of \ Collect S.terminal \ UNIV" by auto qed + lemma bij_elem_of: + shows "bij_betw elem_of S.Univ UNIV" + proof (intro bij_betwI) + show "\t. t \ S.Univ \ arr_of (elem_of t) = t" by simp + show "\x. x \ UNIV \ elem_of (arr_of x) = x" by simp + show "elem_of \ S.Univ \ UNIV" using elem_of_mapsto by auto + show "arr_of \ UNIV \ S.Univ" using arr_of_mapsto by auto + qed + + lemma elem_of_img_arr_of_img [simp]: + shows "elem_of ` arr_of ` A = A" + by force + + lemma arr_of_img_elem_of_img [simp]: + assumes "A \ S.Univ" + shows "arr_of ` elem_of ` A = A" + using assms by force + lemma Dom_terminal: - assumes "terminal t" - shows "Dom t = {DOWN t}" - using assms UP_def - by (metis (mono_tags, lifting) Dom.simps(1) DOWN_def terminal_char the_elem_eq) + assumes "S.terminal t" + shows "S.Dom t = {elem_of t}" + using assms arr_of_def + by (metis (mono_tags, lifting) S.Dom.simps(1) elem_of_def terminal_char the_elem_eq) text\ The image of a point @{term "p \ hom unity a"} is a terminal object, which is given - by the formula @{term "(UP o Fun p o DOWN) unity"}. + by the formula @{term "(arr_of o Fun p o elem_of) unity"}. \ lemma IMG_point: - assumes "\p : unity \ a\" - shows "IMG \ hom unity a \ Univ" - and "IMG p = (UP o Map p o DOWN) unity" + assumes "\p : S.unity \ a\" + shows "IMG \ S.hom S.unity a \ S.Univ" + and "IMG p = (arr_of o S.Map p o elem_of) S.unity" proof - - show "IMG \ hom unity a \ Univ" + show "IMG \ S.hom S.unity a \ S.Univ" proof fix f - assume f: "f \ hom unity a" - have "terminal (MkIde (Map f ` Dom unity))" + assume f: "f \ S.hom S.unity a" + have "S.terminal (S.MkIde (S.Map f ` S.Dom S.unity))" proof - - obtain u :: 'e where u: "unity = MkIde {u}" + obtain u :: 'e where u: "S.unity = S.MkIde {u}" using terminal_unity terminal_char by (metis (mono_tags, lifting)) - have "Map f ` Dom unity = {Map f u}" + have "S.Map f ` S.Dom S.unity = {S.Map f u}" using u by simp thus ?thesis using terminal_char by auto qed - hence "MkIde (Map f ` Dom unity) \ Univ" by simp - moreover have "MkIde (Map f ` Dom unity) = IMG f" - using f IMG_def in_hom_char + hence "S.MkIde (S.Map f ` S.Dom S.unity) \ S.Univ" by simp + moreover have "S.MkIde (S.Map f ` S.Dom S.unity) = IMG f" + using f IMG_def S.in_hom_char by (metis (mono_tags, lifting) mem_Collect_eq) - ultimately show "IMG f \ Univ" by auto + ultimately show "IMG f \ S.Univ" by auto qed - have "IMG p = MkIde (Map p ` Dom p)" using IMG_def by blast - also have "... = MkIde (Map p ` {U})" - using assms in_hom_char terminal_unity Dom_terminal + have "IMG p = S.MkIde (S.Map p ` S.Dom p)" using IMG_def by blast + also have "... = S.MkIde (S.Map p ` {U})" + using assms S.in_hom_char terminal_unity Dom_terminal by (metis (mono_tags, lifting)) - also have "... = (UP o Map p o DOWN) unity" by (simp add: UP_def) - finally show "IMG p = (UP o Map p o DOWN) unity" using assms by auto + also have "... = (arr_of o S.Map p o elem_of) S.unity" by (simp add: arr_of_def) + finally show "IMG p = (arr_of o S.Map p o elem_of) S.unity" using assms by auto qed text\ The function @{term IMG} is injective on @{term "hom unity a"} and its inverse takes a terminal object @{term t} to the arrow in @{term "hom unity a"} corresponding to the constant-@{term t} function. \ abbreviation MkElem :: "'e setcat.arr => 'e setcat.arr => 'e setcat.arr" - where "MkElem t a \ MkArr {U} (Dom a) (\_ \ {U}. DOWN t)" + where "MkElem t a \ S.MkArr {U} (S.Dom a) (\_ \ {U}. elem_of t)" lemma MkElem_in_hom: - assumes "arr f" and "x \ Dom f" - shows "\MkElem (UP x) (dom f) : unity \ dom f\" + assumes "S.arr f" and "x \ S.Dom f" + shows "\MkElem (arr_of x) (S.dom f) : S.unity \ S.dom f\" proof - - have "(\_ \ {U}. DOWN (UP x)) \ {U} \ Dom (dom f)" - using assms dom_char [of f] by simp - moreover have "MkIde {U} = unity" + have "(\_ \ {U}. elem_of (arr_of x)) \ {U} \ S.Dom (S.dom f)" + using assms S.dom_char [of f] by simp + moreover have "S.MkIde {U} = S.unity" using terminal_char terminal_unity - by (metis (mono_tags, lifting) DOWN_UP UP_def) - moreover have "MkIde (Dom (dom f)) = dom f" - using assms dom_char MkIde_Dom' ide_dom by blast + by (metis (mono_tags, lifting) elem_of_arr_of arr_of_def) + moreover have "S.MkIde (S.Dom (S.dom f)) = S.dom f" + using assms S.dom_char S.MkIde_Dom' S.ide_dom by blast ultimately show ?thesis - using assms MkArr_in_hom [of "{U}" "Dom (dom f)" "\_ \ {U}. DOWN (UP x)"] - by (metis (no_types, lifting) Dom.simps(1) Dom_in_Obj IntI arr_dom ideD(1) - restrict_extensional terminal_def terminal_unity) + using assms S.MkArr_in_hom [of "{U}" "S.Dom (S.dom f)" "\_ \ {U}. elem_of (arr_of x)"] + by (metis (no_types, lifting) S.Dom.simps(1) S.Dom_in_Obj IntI S.arr_dom S.ideD(1) + restrict_extensional S.terminal_def terminal_unity) qed lemma MkElem_IMG: - assumes "p \ hom unity a" + assumes "p \ S.hom S.unity a" shows "MkElem (IMG p) a = p" proof - - have 0: "IMG p = UP (Map p U)" + have 0: "IMG p = arr_of (S.Map p U)" using assms IMG_point(2) by auto - have 1: "Dom p = {U}" + have 1: "S.Dom p = {U}" using assms terminal_unity Dom_terminal - by (metis (mono_tags, lifting) in_hom_char mem_Collect_eq) - moreover have "Cod p = Dom a" + by (metis (mono_tags, lifting) S.in_hom_char mem_Collect_eq) + moreover have "S.Cod p = S.Dom a" using assms - by (metis (mono_tags, lifting) in_hom_char mem_Collect_eq) - moreover have "Map p = (\_ \ {U}. DOWN (IMG p))" + by (metis (mono_tags, lifting) S.in_hom_char mem_Collect_eq) + moreover have "S.Map p = (\_ \ {U}. elem_of (IMG p))" proof fix e - show "Map p e = (\_ \ {U}. DOWN (IMG p)) e" + show "S.Map p e = (\_ \ {U}. elem_of (IMG p)) e" proof - - have "Map p e = (\x \ Dom p. Map p x) e" + have "S.Map p e = (\x \ S.Dom p. S.Map p x) e" using assms MkArr_expansion [of p] - by (metis (mono_tags, lifting) CollectD Map.simps(1) in_homE) - also have "... = (\_ \ {U}. DOWN (IMG p)) e" + by (metis (mono_tags, lifting) CollectD S.Map.simps(1) S.in_homE) + also have "... = (\_ \ {U}. elem_of (IMG p)) e" using assms 0 1 by simp finally show ?thesis by blast qed qed ultimately show "MkElem (IMG p) a = p" - using assms MkArr_Map CollectD - by (metis (mono_tags, lifting) in_homE mem_Collect_eq) + using assms S.MkArr_Map CollectD + by (metis (mono_tags, lifting) S.in_homE mem_Collect_eq) qed lemma inj_IMG: - assumes "ide a" - shows "inj_on IMG (hom unity a)" + assumes "S.ide a" + shows "inj_on IMG (S.hom S.unity a)" proof (intro inj_onI) fix x y - assume x: "x \ hom unity a" - assume y: "y \ hom unity a" + assume x: "x \ S.hom S.unity a" + assume y: "y \ S.hom S.unity a" assume eq: "IMG x = IMG y" show "x = y" - proof (intro arr_eqI) - show "arr x" using x by blast - show "arr y" using y by blast - show "Dom x = Dom y" - using x y in_hom_char by (metis (mono_tags, lifting) CollectD) - show "Cod x = Cod y" - using x y in_hom_char by (metis (mono_tags, lifting) CollectD) - show "Map x = Map y" + proof (intro S.arr_eqI) + show "S.arr x" using x by blast + show "S.arr y" using y by blast + show "S.Dom x = S.Dom y" + using x y S.in_hom_char by (metis (mono_tags, lifting) CollectD) + show "S.Cod x = S.Cod y" + using x y S.in_hom_char by (metis (mono_tags, lifting) CollectD) + show "S.Map x = S.Map y" proof - - have "\a. y \ hom unity a \ MkArr {U} (Dom a) (\e\{U}. DOWN (IMG x)) = y" + have "\a. y \ S.hom S.unity a \ S.MkArr {U} (S.Dom a) (\e\{U}. elem_of (IMG x)) = y" using MkElem_IMG eq by presburger hence "y = x" using MkElem_IMG x y by blast thus ?thesis by meson qed qed qed lemma set_char: - assumes "ide a" - shows "set a = UP ` Dom a" + assumes "S.ide a" + shows "S.set a = arr_of ` S.Dom a" proof - show "set a \ UP ` Dom a" + show "S.set a \ arr_of ` S.Dom a" proof fix t - assume "t \ set a" - from this obtain p where p: "\p : unity \ a\ \ t = IMG p" - using set_def by blast - have "t = (UP o Map p o DOWN) unity" + assume "t \ S.set a" + from this obtain p where p: "\p : S.unity \ a\ \ t = IMG p" + using S.set_def by blast + have "t = (arr_of o S.Map p o elem_of) S.unity" using p IMG_point(2) by blast - moreover have "(Map p o DOWN) unity \ Dom a" - using p arr_char in_hom_char Dom_terminal terminal_unity + moreover have "(S.Map p o elem_of) S.unity \ S.Dom a" + using p arr_char S.in_hom_char Dom_terminal terminal_unity by (metis (mono_tags, lifting) IntD2 Pi_split_insert_domain o_apply) - ultimately show "t \ UP ` Dom a" by simp + ultimately show "t \ arr_of ` S.Dom a" by simp qed - show "UP ` Dom a \ set a" + show "arr_of ` S.Dom a \ S.set a" proof fix t - assume "t \ UP ` Dom a" - from this obtain x where x: "x \ Dom a \ t = UP x" by blast - let ?p = "MkElem (UP x) a" - have p: "?p \ hom unity a" - using assms x MkElem_in_hom [of "dom a"] ideD(1-2) by force + assume "t \ arr_of ` S.Dom a" + from this obtain x where x: "x \ S.Dom a \ t = arr_of x" by blast + let ?p = "MkElem (arr_of x) a" + have p: "?p \ S.hom S.unity a" + using assms x MkElem_in_hom [of "S.dom a"] S.ideD(1-2) by force moreover have "IMG ?p = t" - using p x DOWN_UP IMG_def UP_def - by (metis (no_types, lifting) Dom.simps(1) Map.simps(1) image_empty + using p x elem_of_arr_of IMG_def arr_of_def + by (metis (no_types, lifting) S.Dom.simps(1) S.Map.simps(1) image_empty image_insert image_restrict_eq) - ultimately show "t \ set a" using set_def by blast + ultimately show "t \ S.set a" using S.set_def by blast qed qed lemma Map_via_comp: - assumes "arr f" - shows "Map f = (\x \ Dom f. Map (f \ MkElem (UP x) (dom f)) U)" + assumes "S.arr f" + shows "S.Map f = (\x \ S.Dom f. S.Map (f \ MkElem (arr_of x) (S.dom f)) U)" proof fix x - have "x \ Dom f \ Map f x = (\x \ Dom f. Map (f \ MkElem (UP x) (dom f)) U) x" + have "x \ S.Dom f \ S.Map f x = (\x \ S.Dom f. S.Map (f \ MkElem (arr_of x) (S.dom f)) U) x" using assms arr_char [of f] IntD1 extensional_arb restrict_apply by fastforce moreover have - "x \ Dom f \ Map f x = (\x \ Dom f. Map (f \ MkElem (UP x) (dom f)) U) x" + "x \ S.Dom f \ S.Map f x = (\x \ S.Dom f. S.Map (f \ MkElem (arr_of x) (S.dom f)) U) x" proof - - assume x: "x \ Dom f" - let ?X = "MkElem (UP x) (dom f)" - have "\?X : unity \ dom f\" + assume x: "x \ S.Dom f" + let ?X = "MkElem (arr_of x) (S.dom f)" + have "\?X : S.unity \ S.dom f\" using assms x MkElem_in_hom by auto - moreover have "Dom ?X = {U} \ Map ?X = (\_ \ {U}. x)" + moreover have "S.Dom ?X = {U} \ S.Map ?X = (\_ \ {U}. x)" using x by simp ultimately have - "Map (f \ MkElem (UP x) (dom f)) = compose {U} (Map f) (\_ \ {U}. x)" - using assms x Map_comp [of "MkElem (UP x) (dom f)" f] - by (metis (mono_tags, lifting) Cod.simps(1) Dom_dom arr_iff_in_hom seqE seqI') + "S.Map (f \ MkElem (arr_of x) (S.dom f)) = compose {U} (S.Map f) (\_ \ {U}. x)" + using assms x S.Map_comp [of "MkElem (arr_of x) (S.dom f)" f] + by (metis (mono_tags, lifting) S.Cod.simps(1) S.Dom_dom S.arr_iff_in_hom S.seqE S.seqI') thus ?thesis using x by (simp add: compose_eq restrict_apply' singletonI) qed - ultimately show "Map f x = (\x \ Dom f. Map (f \ MkElem (UP x) (dom f)) U) x" + ultimately show "S.Map f x = (\x \ S.Dom f. S.Map (f \ MkElem (arr_of x) (S.dom f)) U) x" by auto qed lemma arr_eqI': - assumes "par f f'" and "\t. \t : unity \ dom f\ \ f \ t = f' \ t" + assumes "S.par f f'" and "\t. \t : S.unity \ S.dom f\ \ f \ t = f' \ t" shows "f = f'" - proof (intro arr_eqI) - show "arr f" using assms by simp - show "arr f'" using assms by simp - show "Dom f = Dom f'" - using assms by (metis (mono_tags, lifting) Dom_dom) - show "Cod f = Cod f'" - using assms by (metis (mono_tags, lifting) Cod_cod) - show "Map f = Map f'" + proof (intro S.arr_eqI) + show "S.arr f" using assms by simp + show "S.arr f'" using assms by simp + show "S.Dom f = S.Dom f'" + using assms by (metis (mono_tags, lifting) S.Dom_dom) + show "S.Cod f = S.Cod f'" + using assms by (metis (mono_tags, lifting) S.Cod_cod) + show "S.Map f = S.Map f'" proof - have 1: "\x. x \ Dom f \ \MkElem (UP x) (dom f) : unity \ dom f\" + have 1: "\x. x \ S.Dom f \ \MkElem (arr_of x) (S.dom f) : S.unity \ S.dom f\" using MkElem_in_hom by (metis (mono_tags, lifting) assms(1)) fix x - show "Map f x = Map f' x" - using assms 1 \Dom f = Dom f'\ by (simp add: Map_via_comp) + show "S.Map f x = S.Map f' x" + using assms 1 \S.Dom f = S.Dom f'\ by (simp add: Map_via_comp) qed qed - text \ - We need to show that the cardinality constraint on the sets that determine objects - implies a corresponding constraint on the sets of global elements of those objects. - \ - - lemma card_points_less: - assumes "ide a" - shows "|hom unity a| " + lemma Setp_elem_of_img: + assumes "A \ S.set ` Collect S.ide" + shows "Setp (elem_of ` A)" proof - - have "bij_betw (\f. Map f U) (hom unity a) (Dom a)" - proof (intro bij_betwI') - show "\x. x \ hom unity a \ Map x (DOWN unity) \ Dom a" - using arr_char Dom_terminal terminal_unity in_hom_char by auto - show "\x y. \x \ hom unity a; y \ hom unity a\ \ Map x U = Map y U \ x = y" - proof - - fix x y - assume x: "x \ hom unity a" and y: "y \ hom unity a" - have 1: "Map x \ extensional {U} \ Map y \ extensional {U}" - using x y in_hom_char Dom_terminal terminal_unity - by (metis (mono_tags, lifting) Map_via_comp mem_Collect_eq restrict_extensional) - show "Map x U = Map y U \ x = y" - proof - show "x = y \ Map x U = Map y U" - by simp - show "Map x U = Map y U \ x = y" - proof - - assume 2: "Map x U = Map y U" - have "Map x = Map y" - proof - fix z - show "Map x z = Map y z" - using 1 2 extensional_arb [of "Map x"] extensional_arb [of "Map y"] - by (cases "z = U") auto - qed - thus "x = y" - using x y 1 in_hom_char - by (intro arr_eqI) auto - qed - qed - qed - show "\y. y \ Dom a \ \x \ hom unity a. y = Map x (DOWN unity)" - proof - - fix y - assume y: "y \ Dom a" - let ?x = "MkArr {DOWN unity} (Dom a) (\_ \ {U}. y)" - have "arr ?x" - proof (intro arr_MkArr) - show "{U} \ {A. |A| }" - by (metis (mono_tags, lifting) Dom_terminal ide_char terminal_def terminal_unity) - show "Dom a \ {A. |A| }" - using assms ide_char by blast - show "(\_ \ {U}. y) \ extensional {U} \ ({U} \ Dom a)" - using assms y by blast - qed - hence "?x \ hom unity a" - using UP_DOWN UP_def assms cod_MkArr dom_char in_homI terminal_unity by simp - moreover have "y = Map ?x (DOWN unity)" - by simp - ultimately show "\x \ hom unity a. y = Map x (DOWN unity)" - by auto - qed + obtain a where a: "S.ide a \ S.set a = A" + using assms by blast + have "elem_of ` S.set a = S.Dom a" + proof - + have "S.set a = arr_of ` S.Dom a" + using a set_char by blast + moreover have "elem_of ` arr_of ` S.Dom a = S.Dom a" + using elem_of_arr_of by force + ultimately show ?thesis by presburger qed - hence "|hom unity a| =o |Dom a|" - using card_of_ordIsoI by auto - moreover have "|Dom a| " - using assms ide_char by auto - ultimately show "|hom unity a| " - using ordIso_ordLess_trans by auto + thus ?thesis + using a S.ide_char by auto + qed + + lemma set_MkIde_elem_of_img: + assumes "A \ S.Univ" and "S.ide (S.MkIde (elem_of ` A))" + shows "S.set (S.MkIde (elem_of ` A)) = A" + proof - + have "S.Dom (S.MkIde (elem_of ` A)) = elem_of ` A" + by simp + moreover have "arr_of ` elem_of ` A = A" + using assms arr_of_elem_of by force + ultimately show ?thesis + using assms Setp_elem_of_img set_char S.ide_MkIde by auto + qed + + (* + * We need this result, which characterizes what sets of terminals correspond + * to sets. + *) + lemma set_img_Collect_ide_iff: + shows "A \ S.set ` Collect S.ide \ A \ S.Univ \ Setp (elem_of ` A)" + proof + show "A \ S.set ` Collect S.ide \ A \ S.Univ \ Setp (elem_of ` A)" + using set_char arr_of_mapsto Setp_elem_of_img by auto + assume A: "A \ S.Univ \ Setp (elem_of ` A)" + have "S.ide (S.MkIde (elem_of ` A))" + using A S.ide_MkIde by blast + moreover have "S.set (S.MkIde (elem_of ` A)) = A" + using A calculation set_MkIde_elem_of_img by presburger + ultimately show "A \ S.set ` Collect S.ide" by blast qed text\ The main result, which establishes the consistency of the \set_category\ locale and provides us with a way of obtaining ``set categories'' at arbitrary types. \ theorem is_set_category: - shows "set_category comp \" + shows "set_category comp (\A. A \ S.Univ \ Setp (elem_of ` A))" proof - show "\img :: 'e setcat.arr \ 'e setcat.arr. set_category_given_img comp img \" + show "\img :: 'e setcat.arr \ 'e setcat.arr. + set_category_given_img comp img (\A. A \ S.Univ \ Setp (elem_of ` A))" proof - show "set_category_given_img (comp :: 'e setcat.arr comp) IMG \" + show "set_category_given_img comp IMG (\A. A \ S.Univ \ Setp (elem_of ` A))" proof - show "Card_order \ \ infinite (Field \ )" - using cardinal by simp - show "Univ \ {}" using terminal_char by blast + show "S.Univ \ {}" using terminal_char by blast fix a :: "'e setcat.arr" - assume a: "ide a" - show "IMG \ hom unity a \ Univ" using IMG_point terminal_unity by blast - show "|hom unity a| " using a card_points_less by simp - show "inj_on IMG (hom unity a)" using a inj_IMG terminal_unity by blast + assume a: "S.ide a" + show "S.set a \ S.Univ \ Setp (elem_of ` S.set a)" + using a set_img_Collect_ide_iff by auto + show "inj_on IMG (S.hom S.unity a)" using a inj_IMG terminal_unity by blast next fix t :: "'e setcat.arr" - assume t: "terminal t" - show "t \ IMG ` hom unity t" + assume t: "S.terminal t" + show "t \ IMG ` S.hom S.unity t" proof - - have "t \ set t" + have "t \ S.set t" using t set_char [of t] - by (metis (mono_tags, lifting) Dom.simps(1) image_insert insertI1 UP_def - terminal_char terminal_def) + by (metis (mono_tags, lifting) S.Dom.simps(1) image_insert insertI1 arr_of_def + terminal_char S.terminal_def) thus ?thesis - using t set_def [of t] by simp + using t S.set_def [of t] by simp qed next - fix A :: "'e setcat.arr set" - assume A: "A \ Univ" and 0: "|A| " - show "\a. ide a \ set a = A" - proof - let ?a = "MkArr (DOWN ` A) (DOWN ` A) (\x \ (DOWN ` A). x)" - show "ide ?a \ set ?a = A" - proof - have "|DOWN ` A| " - using 0 card_of_image ordLeq_ordLess_trans by blast - thus 1: "ide ?a" - using ide_char [of ?a] by simp - show "set ?a = A" - proof - - have 2: "\x. x \ A \ x = UP (DOWN x)" - using A UP_DOWN by force - hence "UP ` DOWN ` A = A" - using A UP_DOWN by auto - thus ?thesis - using 1 A set_char [of ?a] by simp - qed - qed - qed + show "\A. A \ S.Univ \ Setp (elem_of ` A) \ \a. S.ide a \ S.set a = A" + using set_img_Collect_ide_iff by fast next fix a b :: "'e setcat.arr" - assume a: "ide a" and b: "ide b" and ab: "set a = set b" + assume a: "S.ide a" and b: "S.ide b" and ab: "S.set a = S.set b" show "a = b" - using a b ab set_char inj_UP inj_image_eq_iff dom_char in_homE ide_in_hom + using a b ab set_char inj_arr_of inj_image_eq_iff S.dom_char S.in_homE S.ide_in_hom by (metis (mono_tags, lifting)) next fix f f' :: "'e setcat.arr" - assume par: "par f f'" and ff': "\x. \x : unity \ dom f\ \ f \ x = f' \ x" + assume par: "S.par f f'" and ff': "\x. \x : S.unity \ S.dom f\ \ f \ x = f' \ x" show "f = f'" using par ff' arr_eqI' by blast next fix a b :: "'e setcat.arr" and F :: "'e setcat.arr \ 'e setcat.arr" - assume a: "ide a" and b: "ide b" and F: "F \ hom unity a \ hom unity b" - show "\f. \f : a \ b\ \ (\x. \x : unity \ dom f\ \ f \ x = F x)" + assume a: "S.ide a" and b: "S.ide b" and F: "F \ S.hom S.unity a \ S.hom S.unity b" + show "\f. \f : a \ b\ \ (\x. \x : S.unity \ S.dom f\ \ f \ x = F x)" proof - let ?f = "MkArr (Dom a) (Dom b) (\x \ Dom a. Map (F (MkElem (UP x) a)) U)" + let ?f = "S.MkArr (S.Dom a) (S.Dom b) (\x \ S.Dom a. S.Map (F (MkElem (arr_of x) a)) U)" have 1: "\?f : a \ b\" proof - - have "(\x \ Dom a. Map (F (MkElem (UP x) a)) U) - \ extensional (Dom a) \ (Dom a \ Dom b)" + have "(\x \ S.Dom a. S.Map (F (MkElem (arr_of x) a)) U) + \ extensional (S.Dom a) \ (S.Dom a \ S.Dom b)" proof - show "(\x \ Dom a. Map (F (MkElem (UP x) a)) U) \ extensional (Dom a)" + show "(\x \ S.Dom a. S.Map (F (MkElem (arr_of x) a)) U) \ extensional (S.Dom a)" using a F by simp - show "(\x \ Dom a. Map (F (MkElem (UP x) a)) U) \ Dom a \ Dom b" + show "(\x \ S.Dom a. S.Map (F (MkElem (arr_of x) a)) U) \ S.Dom a \ S.Dom b" proof fix x - assume x: "x \ Dom a" - have "MkElem (UP x) a \ hom unity a" - using x a MkElem_in_hom [of a x] ide_char ideD(1-2) by force - hence 1: "F (MkElem (UP x) a) \ hom unity b" + assume x: "x \ S.Dom a" + have "MkElem (arr_of x) a \ S.hom S.unity a" + using x a MkElem_in_hom [of a x] S.ide_char S.ideD(1-2) by force + hence 1: "F (MkElem (arr_of x) a) \ S.hom S.unity b" using F by auto - moreover have "Dom (F (MkElem (UP x) a)) = {U}" + moreover have "S.Dom (F (MkElem (arr_of x) a)) = {U}" using 1 MkElem_IMG - by (metis (mono_tags, lifting) Dom.simps(1)) - moreover have "Cod (F (MkElem (UP x) a)) = Dom b" - using 1 by (metis (mono_tags, lifting) CollectD in_hom_char) - ultimately have "Map (F (MkElem (UP x) a)) \ {U} \ Dom b" - using arr_char [of "F (MkElem (UP x) a)"] by blast - thus "Map (F (MkElem (UP x) a)) U \ Dom b" by blast + by (metis (mono_tags, lifting) S.Dom.simps(1)) + moreover have "S.Cod (F (MkElem (arr_of x) a)) = S.Dom b" + using 1 by (metis (mono_tags, lifting) CollectD S.in_hom_char) + ultimately have "S.Map (F (MkElem (arr_of x) a)) \ {U} \ S.Dom b" + using arr_char [of "F (MkElem (arr_of x) a)"] by blast + thus "S.Map (F (MkElem (arr_of x) a)) U \ S.Dom b" by blast qed qed - hence "\?f : MkIde (Dom a) \ MkIde (Dom b)\" - using a b MkArr_in_hom ide_char by blast + hence "\?f : S.MkIde (S.Dom a) \ S.MkIde (S.Dom b)\" + using a b S.MkArr_in_hom S.ide_char by blast thus ?thesis using a b by simp qed - moreover have "\x. \x : unity \ dom ?f\ \ ?f \ x = F x" + moreover have "\x. \x : S.unity \ S.dom ?f\ \ ?f \ x = F x" proof - fix x - assume x: "\x : unity \ dom ?f\" + assume x: "\x : S.unity \ S.dom ?f\" have 2: "x = MkElem (IMG x) a" using a x 1 MkElem_IMG [of x a] - by (metis (mono_tags, lifting) in_homE mem_Collect_eq) - moreover have 5: "Dom x = {U} \ Cod x = Dom a \ - Map x = (\_ \ {U}. DOWN (IMG x))" + by (metis (mono_tags, lifting) S.in_homE mem_Collect_eq) + moreover have 5: "S.Dom x = {U} \ S.Cod x = S.Dom a \ + S.Map x = (\_ \ {U}. elem_of (IMG x))" using x 2 - by (metis (no_types, lifting) Cod.simps(1) Dom.simps(1) Map.simps(1)) - moreover have "Cod ?f = Dom b" using 1 by simp + by (metis (no_types, lifting) S.Cod.simps(1) S.Dom.simps(1) S.Map.simps(1)) + moreover have "S.Cod ?f = S.Dom b" using 1 by simp ultimately have 3: "?f \ x = - MkArr {U} (Dom b) (compose {U} (Map ?f) (\_ \ {U}. DOWN (IMG x)))" - by (metis (no_types, lifting) "1" Map.simps(1) comp_MkArr in_homE x) - have 4: "compose {U} (Map ?f) (\_ \ {U}. DOWN (IMG x)) = Map (F x)" + S.MkArr {U} (S.Dom b) (compose {U} (S.Map ?f) (\_ \ {U}. elem_of (IMG x)))" + by (metis (no_types, lifting) "1" S.Map.simps(1) S.comp_MkArr S.in_homE x) + have 4: "compose {U} (S.Map ?f) (\_ \ {U}. elem_of (IMG x)) = S.Map (F x)" proof fix y have "y \ {U} \ - compose {U} (Map ?f) (\_ \ {U}. DOWN (IMG x)) y = Map (F x) y" + compose {U} (S.Map ?f) (\_ \ {U}. elem_of (IMG x)) y = S.Map (F x) y" proof - assume y: "y \ {U}" - have "compose {U} (Map ?f) (\_ \ {U}. DOWN (IMG x)) y = undefined" + have "compose {U} (S.Map ?f) (\_ \ {U}. elem_of (IMG x)) y = undefined" using y compose_def extensional_arb by simp - also have "... = Map (F x) y" + also have "... = S.Map (F x) y" proof - - have 5: "F x \ hom unity b" using x F 1 by fastforce - hence "Dom (F x) = {U}" - by (metis (mono_tags, lifting) "2" CollectD Dom.simps(1) in_hom_char x) + have 5: "F x \ S.hom S.unity b" using x F 1 by fastforce + hence "S.Dom (F x) = {U}" + by (metis (mono_tags, lifting) "2" CollectD S.Dom.simps(1) S.in_hom_char x) thus ?thesis - using x y F 5 arr_char [of "F x"] extensional_arb [of "Map (F x)" "{U}" y] - by (metis (mono_tags, lifting) CollectD Int_iff in_hom_char) + using x y F 5 arr_char [of "F x"] extensional_arb [of "S.Map (F x)" "{U}" y] + by (metis (mono_tags, lifting) CollectD Int_iff S.in_hom_char) qed ultimately show ?thesis by auto qed moreover have "y \ {U} \ - compose {U} (Map ?f) (\_ \ {U}. DOWN (IMG x)) y = Map (F x) y" + compose {U} (S.Map ?f) (\_ \ {U}. elem_of (IMG x)) y = S.Map (F x) y" proof - assume y: "y \ {U}" - have "compose {U} (Map ?f) (\_ \ {U}. DOWN (IMG x)) y = - Map ?f (DOWN (IMG x))" + have "compose {U} (S.Map ?f) (\_ \ {U}. elem_of (IMG x)) y = + S.Map ?f (elem_of (IMG x))" using y by (simp add: compose_eq restrict_apply') - also have "... = (\x. Map (F (MkElem (UP x) a)) U) (DOWN (IMG x))" + also have "... = (\x. S.Map (F (MkElem (arr_of x) a)) U) (elem_of (IMG x))" proof - - have "DOWN (IMG x) \ Dom a" - using x y a 5 arr_char in_homE restrict_apply by force + have "elem_of (IMG x) \ S.Dom a" + using x y a 5 arr_char S.in_homE restrict_apply by force thus ?thesis using restrict_apply by simp qed - also have "... = Map (F x) y" + also have "... = S.Map (F x) y" using x y 1 2 MkElem_IMG [of x a] by simp finally show - "compose {U} (Map ?f) (\_ \ {U}. DOWN (IMG x)) y = Map (F x) y" + "compose {U} (S.Map ?f) (\_ \ {U}. elem_of (IMG x)) y = S.Map (F x) y" by auto qed ultimately show - "compose {U} (Map ?f) (\_ \ {U}. DOWN (IMG x)) y = Map (F x) y" + "compose {U} (S.Map ?f) (\_ \ {U}. elem_of (IMG x)) y = S.Map (F x) y" by auto qed show "?f \ x = F x" - proof (intro arr_eqI) - have 5: "?f \ x \ hom unity b" using 1 x by blast - have 6: "F x \ hom unity b" + proof (intro S.arr_eqI) + have 5: "?f \ x \ S.hom S.unity b" using 1 x by blast + have 6: "F x \ S.hom S.unity b" using x F 1 - by (metis (mono_tags, lifting) PiE in_homE mem_Collect_eq) - show "arr (comp ?f x)" using 5 by blast - show "arr (F x)" using 6 by blast - show "Dom (comp ?f x) = Dom (F x)" - using 5 6 by (metis (mono_tags, lifting) CollectD in_hom_char) - show "Cod (comp ?f x) = Cod (F x)" - using 5 6 by (metis (mono_tags, lifting) CollectD in_hom_char) - show "Map (comp ?f x) = Map (F x)" + by (metis (mono_tags, lifting) PiE S.in_homE mem_Collect_eq) + show "S.arr (comp ?f x)" using 5 by blast + show "S.arr (F x)" using 6 by blast + show "S.Dom (comp ?f x) = S.Dom (F x)" + using 5 6 by (metis (mono_tags, lifting) CollectD S.in_hom_char) + show "S.Cod (comp ?f x) = S.Cod (F x)" + using 5 6 by (metis (mono_tags, lifting) CollectD S.in_hom_char) + show "S.Map (comp ?f x) = S.Map (F x)" using 3 4 by simp qed qed - thus "\?f : a \ b\ \ (\x. \x : unity \ dom ?f\ \ comp ?f x = F x)" + thus "\?f : a \ b\ \ (\x. \x : S.unity \ S.dom ?f\ \ comp ?f x = F x)" using 1 by blast qed + next + show "\A. A \ S.Univ \ Setp (elem_of ` A) \ A \ S.Univ" + by simp + show "\A' A. \A' \ A; A \ S.Univ \ Setp (elem_of ` A)\ \ A' \ S.Univ \ Setp (elem_of ` A')" + by (meson image_mono setcat.Setp_respects_subset setcat_axioms subset_trans) + show "\A B. \A \ S.Univ \ Setp (elem_of ` A); B \ S.Univ \ Setp (elem_of ` B)\ + \ A \ B \ S.Univ \ Setp (elem_of ` (A \ B))" + by (simp add: image_Un union_preserves_Setp) qed qed qed text\ \SetCat\ can be viewed as a concrete set category over its own element type - @{typ 'a}, using @{term UP} as the required injection from @{typ 'a} to the universe + @{typ 'a}, using @{term arr_of} as the required injection from @{typ 'a} to the universe of \SetCat\. \ corollary is_concrete_set_category: - shows "concrete_set_category comp \ UNIV UP" + shows "concrete_set_category comp (\A. A \ S.Univ \ Setp (elem_of ` A)) UNIV arr_of" proof - - interpret S: set_category comp \ using is_set_category by auto + interpret S': set_category comp \\A. A \ S.Univ \ Setp (elem_of ` A)\ + using is_set_category by auto show ?thesis proof - show 1: "UP \ UNIV \ S.Univ" - using UP_def terminal_char by force - show "inj_on UP UNIV" - using inj_UP by blast + show 1: "arr_of \ UNIV \ S'.Univ" + using arr_of_def terminal_char by force + show "inj_on arr_of UNIV" + using inj_arr_of by blast qed qed text\ As a consequence of the categoricity of the \set_category\ axioms, if @{term S} interprets \set_category\, and if @{term \} is a bijection between the universe of @{term S} and the elements of type @{typ 'a}, then @{term S} is isomorphic to the category \setcat\ of @{typ 'a} sets and functions between them constructed here. \ corollary set_category_iso_SetCat: fixes S :: "'s comp" and \ :: "'s \ 'e" - assumes "set_category S \" - and "bij_betw \ (Collect (category.terminal S)) UNIV" - shows "\\. invertible_functor S (comp :: 'e setcat.arr comp) \ - \ (\m. set_category.incl S \ m \ set_category.incl comp \ (\ m))" + assumes "set_category S \" + and "bij_betw \ (set_category_data.Univ S) UNIV" + and "\A. \ A \ A \ set_category_data.Univ S \ (arr_of \ \) ` A \ S.set ` Collect S.ide" + shows "\\. invertible_functor S comp \ + \ (\m. set_category.incl S \ m + \ set_category.incl comp (\A. A \ S.set ` Collect S.ide) (\ m))" proof - - interpret S: set_category S using assms by auto - let ?\ = "inv_into S.Univ \" - have "bij_betw (UP o \) S.Univ (Collect terminal)" + interpret S': set_category S \ using assms by auto + let ?\ = "inv_into S'.Univ \" + have "bij_betw (arr_of o \) S'.Univ S.Univ" proof (intro bij_betwI) - show "UP o \ \ S.Univ \ Collect terminal" - using assms(2) UP_mapsto by auto - show "?\ o DOWN \ Collect terminal \ S.Univ" + show "arr_of o \ \ S'.Univ \ Collect S.terminal" + using assms(2) arr_of_mapsto by auto + show "?\ o elem_of \ S.Univ \ S'.Univ" proof fix x :: "'e setcat.arr" - assume x: "x \ Univ" - show "(inv_into S.Univ \ \ DOWN) x \ S.Univ" + assume x: "x \ S.Univ" + show "(inv_into S'.Univ \ \ elem_of) x \ S'.Univ" using x assms(2) bij_betw_def comp_apply inv_into_into by (metis UNIV_I) qed fix t - assume "t \ S.Univ" - thus "(?\ o DOWN) ((UP o \) t) = t" + assume "t \ S'.Univ" + thus "(?\ o elem_of) ((arr_of o \) t) = t" using assms(2) bij_betw_inv_into_left - by (metis comp_apply DOWN_UP) + by (metis comp_apply elem_of_arr_of) next fix t' :: "'e setcat.arr" - assume "t' \ Collect terminal" - thus "(UP o \) ((?\ o DOWN) t') = t'" + assume "t' \ S.Univ" + thus "(arr_of o \) ((?\ o elem_of) t') = t'" using assms(2) by (simp add: bij_betw_def f_inv_into_f) qed thus ?thesis - using assms(1) set_category_is_categorical [of S \ comp "UP o \"] is_set_category - by auto + using assms is_set_category set_img_Collect_ide_iff + set_category_is_categorical + [of S \ comp "\A. A \ S.set ` Collect S.ide" "arr_of o \"] + by simp + qed + + sublocale category comp + using is_category by blast + sublocale set_category comp \\A. A \ Collect S.terminal \ Setp (elem_of ` A)\ + using is_set_category set_img_Collect_ide_iff by simp + sublocale concrete_set_category comp \\A. A \ Collect S.terminal \ Setp (elem_of ` A)\ + UNIV arr_of + using is_concrete_set_category by simp + + end + + text\ + Here we discard the temporary interpretations \S\, leaving only the exported + definitions and facts. +\ + + context setcat + begin + + text\ + We establish mappings to pass back and forth between objects and arrows of the category + and sets and functions on the underlying elements. +\ + + definition set_of_ide :: "'e setcat.arr \ 'e set" + where "set_of_ide a \ elem_of ` set a" + + definition ide_of_set :: "'e set \ 'e setcat.arr" + where "ide_of_set A \ mkIde (arr_of ` A)" + + lemma bij_betw_ide_set: + shows "set_of_ide \ Collect ide \ Collect Setp" + and "ide_of_set \ Collect Setp \ Collect ide" + and [simp]: "ide a \ ide_of_set (set_of_ide a) = a" + and [simp]: "Setp A \ set_of_ide (ide_of_set A) = A" + and "bij_betw set_of_ide (Collect ide) (Collect Setp)" + and "bij_betw ide_of_set (Collect Setp) (Collect ide)" + proof - + show 1: "set_of_ide \ Collect ide \ Collect Setp" + using setp_set_ide set_of_ide_def by auto + show 2: "ide_of_set \ Collect Setp \ Collect ide" + proof + fix A :: "'e set" + assume A: "A \ Collect Setp" + have "arr_of ` A \ Univ" + using A arr_of_mapsto by auto + moreover have "Setp (elem_of ` arr_of ` A)" + using A elem_of_arr_of Setp_respects_subset by simp + ultimately have "ide (mkIde (arr_of ` A))" + using ide_mkIde by blast + thus "ide_of_set A \ Collect ide" + using ide_of_set_def by simp + qed + show 3: "\a. ide a \ ide_of_set (set_of_ide a) = a" + using arr_of_img_elem_of_img ide_of_set_def mkIde_set set_of_ide_def setp_set_ide + by presburger + show 4: "\A. Setp A \ set_of_ide (ide_of_set A) = A" + proof - + fix A :: "'e set" + assume A: "Setp A" + have "elem_of ` set (mkIde (arr_of ` A)) = elem_of ` arr_of ` A" + proof - + have "arr_of ` A \ Univ" + using A arr_of_mapsto by blast + moreover have "Setp (elem_of ` arr_of ` A)" + using A by simp + ultimately have "set (mkIde (arr_of ` A)) = arr_of ` A" + using A set_mkIde by blast + thus ?thesis by auto + qed + also have "... = A" + using A elem_of_arr_of by force + finally show "set_of_ide (ide_of_set A) = A" + using ide_of_set_def set_of_ide_def by simp + qed + show "bij_betw set_of_ide (Collect ide) (Collect Setp)" + using 1 2 3 4 + by (intro bij_betwI) blast+ + show "bij_betw ide_of_set (Collect Setp) (Collect ide)" + using 1 2 3 4 + by (intro bij_betwI) blast+ + qed + + definition fun_of_arr :: "'e setcat.arr \ 'e \ 'e" + where "fun_of_arr f \ restrict (elem_of o Fun f o arr_of) (elem_of `Dom f)" + + definition arr_of_fun :: "'e set \ 'e set \ ('e \ 'e) \ 'e setcat.arr" + where "arr_of_fun A B F \ mkArr (arr_of ` A) (arr_of ` B) (arr_of o F o elem_of)" + + lemma bij_betw_hom_fun: + shows "fun_of_arr \ hom a b \ extensional (set_of_ide a) \ (set_of_ide a \ set_of_ide b)" + and "\Setp A; Setp B\ \ arr_of_fun A B \ (A \ B) \ hom (ide_of_set A) (ide_of_set B)" + and "f \ hom a b \ arr_of_fun (set_of_ide a) (set_of_ide b) (fun_of_arr f) = f" + and "\Setp A; Setp B; F \ A \ B; F \ extensional A\ \ fun_of_arr (arr_of_fun A B F) = F" + and "\ide a; ide b\ \ bij_betw fun_of_arr (hom a b) + (extensional (set_of_ide a) \ (set_of_ide a \ set_of_ide b))" + and "\Setp A; Setp B\ \ + bij_betw (arr_of_fun A B) + (extensional A \ (A \ B)) (hom (ide_of_set A) (ide_of_set B))" + proof - + show 1: "\a b. fun_of_arr \ + hom a b \ extensional (set_of_ide a) \ (set_of_ide a \ set_of_ide b)" + proof + fix a b f + assume f: "f \ hom a b" + have Dom: "Dom f = arr_of ` set_of_ide a" + using f set_of_ide_def + by (metis (mono_tags, lifting) arr_dom arr_mkIde bij_betw_ide_set(3) + ide_dom ide_of_set_def in_homE mem_Collect_eq set_mkIde) + have Cod: "Cod f = arr_of ` set_of_ide b" + using f set_of_ide_def + by (metis (mono_tags, lifting) arr_cod arr_mkIde bij_betw_ide_set(3) + ide_cod ide_of_set_def in_homE mem_Collect_eq set_mkIde) + have "fun_of_arr f \ set_of_ide a \ set_of_ide b" + proof + fix x + assume x: "x \ set_of_ide a" + have 1: "arr_of x \ Dom f" + using x Dom by blast + hence "Fun f (arr_of x) \ Cod f" + using f Fun_mapsto Cod by blast + hence "elem_of (Fun f (arr_of x)) \ set_of_ide b" + using Cod by auto + hence "restrict (elem_of o Fun f o arr_of) (elem_of ` Dom f) x \ set_of_ide b" + using 1 by force + thus "fun_of_arr f x \ set_of_ide b" + unfolding fun_of_arr_def by auto + qed + moreover have "fun_of_arr f \ extensional (set_of_ide a)" + unfolding fun_of_arr_def + using set_of_ide_def f by blast + ultimately show "fun_of_arr f \ extensional (set_of_ide a) \ (set_of_ide a \ set_of_ide b)" + by blast + qed + show 2: "\A B. \Setp A; Setp B\ \ + arr_of_fun A B \ (A \ B) \ hom (ide_of_set A) (ide_of_set B)" + proof + fix x and A B :: "'e set" + assume A: "Setp A" and B: "Setp B" + assume x: "x \ A \ B" + show "arr_of_fun A B x \ hom (ide_of_set A) (ide_of_set B)" + proof + show "\arr_of_fun A B x : ide_of_set A \ ide_of_set B\" + proof + show 1: "arr (arr_of_fun A B x)" + proof - + have "arr_of ` A \ Univ \ Setp (elem_of ` arr_of ` A)" + using A arr_of_mapsto elem_of_arr_of + by (metis (no_types, lifting) PiE arr_mkIde bij_betw_ide_set(2) + ide_implies_incl ide_of_set_def incl_def mem_Collect_eq) + moreover have "arr_of ` B \ Univ \ Setp (elem_of ` arr_of ` B)" + using B arr_of_mapsto elem_of_arr_of + by (metis (no_types, lifting) PiE arr_mkIde bij_betw_ide_set(2) + ide_implies_incl ide_of_set_def incl_def mem_Collect_eq) + moreover have "arr_of \ x \ elem_of \ arr_of ` A \ arr_of ` B" + using x by auto + ultimately show ?thesis + unfolding arr_of_fun_def by blast + qed + show "dom (arr_of_fun A B x) = ide_of_set A" + using 1 dom_mkArr ide_of_set_def arr_of_fun_def by simp + show "cod (arr_of_fun A B x) = ide_of_set B" + using 1 cod_mkArr ide_of_set_def arr_of_fun_def by simp + qed + qed + qed + show 3: "\a b f. f \ hom a b \ arr_of_fun (set_of_ide a) (set_of_ide b) (fun_of_arr f) = f" + proof - + fix a b f + assume f: "f \ hom a b" + have 1: "Dom f = set a \ Cod f = set b" + using f by blast + have Dom: "Dom f \ Univ \ Setp (elem_of ` Dom f)" + using setp_set_ide f ide_dom by blast + have Cod: "Cod f \ Univ \ Setp (elem_of ` Cod f)" + using setp_set_ide f ide_cod by blast + have "mkArr (set a) (set b) + (arr_of \ restrict (elem_of \ Fun f \ arr_of) (elem_of ` Dom f) \ elem_of) = + mkArr (Dom f) (Cod f) + (arr_of \ restrict (elem_of \ Fun f \ arr_of) (elem_of ` Dom f) \ elem_of)" + using 1 by simp + also have "... = mkArr (Dom f) (Cod f) (Fun f)" + proof (intro mkArr_eqI') + show 2: "\x. x \ Dom f \ + (arr_of \ + restrict (elem_of \ Fun f \ arr_of) (elem_of ` Dom f) \ + elem_of) x = + Fun f x" + proof - + fix x + assume x: "x \ Dom f" + hence 1: "elem_of x \ elem_of ` Dom f" + by blast + have "(arr_of \ restrict (elem_of \ Fun f \ arr_of) (elem_of ` Dom f) \ elem_of) x = + arr_of (restrict (elem_of o Fun f o arr_of) (elem_of ` Dom f) (elem_of x))" + by auto + also have "... = arr_of ((elem_of o Fun f o arr_of) (elem_of x))" + using 1 by auto + also have "... = arr_of (elem_of (Fun f (arr_of (elem_of x))))" + by auto + also have "... = arr_of (elem_of (Fun f x))" + using x arr_of_elem_of \Dom f \ Univ \ Setp (elem_of ` Dom f)\ by auto + also have "... = Fun f x" + using x f Dom Cod Fun_mapsto arr_of_elem_of [of "Fun f x"] by blast + finally show "(arr_of \ + restrict (elem_of \ Fun f \ arr_of) (elem_of ` Dom f) \ + elem_of) x = + Fun f x" + by blast + qed + have "arr_of \ restrict (elem_of \ Fun f \ arr_of) (elem_of ` Dom f) \ elem_of + \ Dom f \ Cod f" + proof + fix x + assume x: "x \ Dom f" + have "(arr_of \ restrict (elem_of \ Fun f \ arr_of) (elem_of ` Dom f) \ elem_of) x = + Fun f x" + using 2 x by blast + moreover have "... \ Cod f" + using f x Fun_mapsto by blast + ultimately show "(arr_of \ + restrict (elem_of \ Fun f \ arr_of) (elem_of ` Dom f) \ + elem_of) x + \ Cod f" + by argo + qed + thus "arr (mkArr (Dom f) (Cod f) + (arr_of \ + restrict (elem_of \ Fun f \ arr_of) (elem_of ` Dom f) \ + elem_of))" + using Dom Cod by blast + qed + finally have "mkArr (set a) (set b) + (arr_of \ + restrict (elem_of \ Fun f \ arr_of) (elem_of ` Dom f) \ + elem_of) = f" + using f mkArr_Fun + by (metis (no_types, lifting) in_homE mem_Collect_eq) + thus "arr_of_fun (set_of_ide a) (set_of_ide b) (fun_of_arr f) = f" + using 1 f + by (metis (no_types, lifting) Cod Dom arr_of_img_elem_of_img arr_of_fun_def + fun_of_arr_def set_of_ide_def) + qed + show 4: "\A B F. \Setp A; Setp B; F \ A \ B; F \ extensional A\ + \ fun_of_arr (arr_of_fun A B F) = F" + proof + fix F and A B :: "'e set" + assume A: "Setp A" and B: "Setp B" + assume F: "F \ A \ B" and ext: "F \ extensional A" + have 4: "arr (mkArr (arr_of ` A) (arr_of ` B) (arr_of \ F \ elem_of))" + proof - + have "arr_of ` A \ Univ \ Setp (elem_of ` arr_of ` A)" + using A + by (metis (no_types, lifting) PiE arr_mkIde bij_betw_ide_set(2) ide_implies_incl + ide_of_set_def incl_def mem_Collect_eq) + moreover have "arr_of ` B \ Univ \ Setp (elem_of ` arr_of ` B)" + using B + by (metis (no_types, lifting) PiE arr_mkIde bij_betw_ide_set(2) ide_implies_incl + ide_of_set_def incl_def mem_Collect_eq) + moreover have "arr_of \ F \ elem_of \ arr_of ` A \ arr_of ` B" + using F by auto + ultimately show ?thesis by blast + qed + show "\x. fun_of_arr (arr_of_fun A B F) x = F x" + proof - + fix x + have "fun_of_arr (arr_of_fun A B F) x = + restrict (elem_of \ + Fun (mkArr (arr_of ` A) (arr_of ` B) (arr_of \ F \ elem_of)) \ + arr_of) A x" + proof - + have "elem_of ` Dom (mkArr (arr_of ` A) (arr_of ` B) (arr_of \ F \ elem_of)) = A" + using A 4 elem_of_arr_of dom_mkArr set_of_ide_def bij_betw_ide_set(4) ide_of_set_def + by auto + thus ?thesis + using arr_of_fun_def fun_of_arr_def by auto + qed + also have "... = F x" + proof (cases "x \ A") + show "x \ A \ ?thesis" + using ext extensional_arb by fastforce + assume x: "x \ A" + show "restrict + (elem_of \ + Fun (mkArr (arr_of ` A) (arr_of ` B) (arr_of \ F \ elem_of)) \ + arr_of) A x = + F x" + using x 4 Fun_mkArr by simp + qed + finally show "fun_of_arr (arr_of_fun A B F) x = F x" + by blast + qed + qed + show "\Setp A; Setp B\ \ + bij_betw (arr_of_fun A B) (extensional A \ (A \ B)) + (hom (ide_of_set A) (ide_of_set B))" + proof - + assume A: "Setp A" and B: "Setp B" + have "ide (ide_of_set A) \ ide (ide_of_set B)" + using A B bij_betw_ide_set(2) by auto + have "set_of_ide (ide_of_set A) = A \ set_of_ide (ide_of_set B) = B" + using A B by simp + show ?thesis + using A B 1 [of "ide_of_set A" "ide_of_set B"] 2 3 4 + apply (intro bij_betwI) + apply auto + apply blast + by fastforce + qed + show "\ide a; ide b\ \ bij_betw fun_of_arr (hom a b) + (extensional (set_of_ide a) \ (set_of_ide a \ set_of_ide b))" + proof (intro bij_betwI) + assume a: "ide a" and b: "ide b" + show "fun_of_arr \ hom a b \ extensional (set_of_ide a) \ (set_of_ide a \ set_of_ide b)" + using 1 by blast + show "arr_of_fun (set_of_ide a) (set_of_ide b) \ + extensional (set_of_ide a) \ (set_of_ide a \ set_of_ide b) \ hom a b" + using a b 2 [of "set_of_ide a" "set_of_ide b"] setp_set_ide set_of_ide_def + bij_betw_ide_set(3) + by auto + show "\f. f \ hom a b \ arr_of_fun (set_of_ide a) (set_of_ide b) (fun_of_arr f) = f" + using a b 3 by blast + show "\F. F \ extensional (set_of_ide a) \ (set_of_ide a \ set_of_ide b) \ + fun_of_arr (arr_of_fun (set_of_ide a) (set_of_ide b) F) = F" + using a b 4 [of "set_of_ide a" "set_of_ide b"] + by (metis (no_types, lifting) IntE set_of_ide_def setp_set_ide) + qed + qed + + lemma fun_of_arr_ide: + assumes "ide a" + shows "fun_of_arr a = restrict id (elem_of ` Dom a)" + proof + fix x + show "fun_of_arr a x = restrict id (elem_of ` Dom a) x" + proof (cases "x \ elem_of ` Dom a") + show "x \ elem_of ` Dom a \ ?thesis" + using fun_of_arr_def extensional_arb by auto + assume x: "x \ elem_of ` Dom a" + have "fun_of_arr a x = restrict (elem_of \ Fun a \ arr_of) (elem_of ` Dom a) x" + using x fun_of_arr_def by simp + also have "... = elem_of (Fun a (arr_of x))" + using x by auto + also have "... = elem_of ((\x\set a. x) (arr_of x))" + using assms x Fun_ide by auto + also have "... = elem_of (arr_of x)" + proof - + have "x \ elem_of ` set a" + using assms x ideD(2) by force + hence "arr_of x \ set a" + by (metis (mono_tags, lifting) arr_of_img_elem_of_img assms image_eqI setp_set_ide) + thus ?thesis by simp + qed + also have "... = x" + by simp + also have "... = restrict id (elem_of ` Dom a) x" + using x by auto + finally show ?thesis by blast + qed + qed + + lemma arr_of_fun_id: + assumes "Setp A" + shows "arr_of_fun A A (restrict id A) = ide_of_set A" + proof - + have A: "arr_of ` A \ Univ \ Setp (elem_of ` arr_of ` A)" + using assms elem_of_arr_of + by (metis (no_types, lifting) PiE arr_mkIde bij_betw_ide_set(2) ide_implies_incl + ide_of_set_def incl_def mem_Collect_eq) + have "arr_of_fun A A (restrict id A) = + mkArr (arr_of ` A) (arr_of ` A) (arr_of \ restrict id A \ elem_of)" + unfolding arr_of_fun_def by simp + also have "... = mkArr (arr_of ` A) (arr_of ` A) (\x. x)" + using A arr_mkArr + by (intro mkArr_eqI') auto + also have "... = ide_of_set A" + using A ide_of_set_def mkIde_as_mkArr by simp + finally show ?thesis by blast + qed + + lemma fun_of_arr_comp: + assumes "f \ hom a b" and "g \ hom b c" + shows "fun_of_arr (comp g f) = restrict (fun_of_arr g \ fun_of_arr f) (set_of_ide a)" + proof - + have 1: "seq g f" + using assms by blast + have "fun_of_arr (comp g f) = + restrict (elem_of \ Fun (comp g f) \ arr_of) (elem_of ` Dom (comp g f))" + unfolding fun_of_arr_def by blast + also have "... = restrict (elem_of \ Fun (comp g f) \ arr_of) (elem_of ` Dom f)" + using assms dom_comp seqI' by auto + also have "... = restrict (elem_of \ restrict (Fun g \ Fun f) (Dom f) \ arr_of) + (elem_of ` Dom f)" + using 1 Fun_comp by auto + also have "... = restrict (restrict (elem_of o Fun g o arr_of) (elem_of ` Dom g) \ + restrict (elem_of o Fun f o arr_of) (elem_of ` Dom f)) + (elem_of ` Dom f)" + proof + fix x + show "restrict (elem_of \ restrict (Fun g \ Fun f) (Dom f) \ arr_of) (elem_of ` Dom f) x = + restrict (restrict (elem_of \ Fun g \ arr_of) (elem_of ` Dom g) \ + restrict (elem_of \ Fun f \ arr_of) (elem_of ` Dom f)) + (elem_of ` Dom f) x" + proof (cases "x \ elem_of ` Dom f") + show "x \ elem_of ` Dom f \ ?thesis" + by auto + assume x: "x \ elem_of ` Dom f" + have 2: "arr_of x \ Dom f" + proof - + have "arr_of x \ arr_of ` elem_of ` Dom f" + using x by simp + thus ?thesis + by (metis (no_types, lifting) 1 arr_of_img_elem_of_img ide_dom seqE setp_set_ide) + qed + have 3: "Dom g = Cod f" + using assms by fastforce + have "restrict (elem_of \ restrict (Fun g \ Fun f) (Dom f) \ arr_of) (elem_of ` Dom f) x = + elem_of (Fun g (Fun f (arr_of x)))" + using x 2 by simp + also have "... = restrict + (restrict (elem_of \ Fun g \ arr_of) (elem_of ` Dom g) \ + restrict (elem_of \ Fun f \ arr_of) (elem_of ` Dom f)) + (elem_of ` Dom f) x" + proof - + have "restrict (restrict (elem_of \ Fun g \ arr_of) (elem_of ` Dom g) \ + restrict (elem_of \ Fun f \ arr_of) (elem_of ` Dom f)) + (elem_of ` Dom f) x = + elem_of (Fun g (Fun f (arr_of x)))" + proof - + have "restrict (restrict (elem_of \ Fun g \ arr_of) (elem_of ` Dom g) \ + restrict (elem_of \ Fun f \ arr_of) (elem_of ` Dom f)) + (elem_of ` Dom f) x = + (restrict (elem_of \ Fun g \ arr_of) (elem_of ` Dom g) o + restrict (elem_of \ Fun f \ arr_of) (elem_of ` Dom f)) x" + using 2 by force + also have "... = restrict (elem_of \ Fun g \ arr_of) (elem_of ` Dom g) + (restrict (elem_of \ Fun f \ arr_of) (elem_of ` Dom f) x)" + by simp + also have "... = restrict (elem_of \ Fun g \ arr_of) (elem_of ` Dom g) + (elem_of (Fun f (arr_of x)))" + using 2 by force + also have "... = (elem_of o Fun g o arr_of) (elem_of (Fun f (arr_of x)))" + proof - + have "elem_of (Fun f (arr_of x)) \ elem_of ` Dom g" + using assms 2 3 Fun_mapsto [of f] by blast + thus ?thesis by simp + qed + also have "... = elem_of (Fun g (arr_of (elem_of (Fun f (arr_of x)))))" + by simp + also have "... = elem_of (Fun g (Fun f (arr_of x)))" + proof - + have "Fun f (arr_of x) \ Univ" + using assms 2 setp_set_ide ide_cod Fun_mapsto by blast + thus ?thesis + using 2 by simp + qed + finally show ?thesis by blast + qed + thus ?thesis by simp + qed + finally show ?thesis by blast + qed + qed + also have "... = restrict (fun_of_arr g o fun_of_arr f) (elem_of ` Dom f)" + unfolding fun_of_arr_def by blast + finally show ?thesis + unfolding set_of_ide_def + using assms by blast + qed + + lemma arr_of_fun_comp: + assumes "Setp A" and "Setp B" and "Setp C" + and "F \ extensional A \ (A \ B)" and "G \ extensional B \ (B \ C)" + shows "arr_of_fun A C (G o F) = comp (arr_of_fun B C G) (arr_of_fun A B F)" + proof - + have A: "arr_of ` A \ Univ \ Setp (elem_of ` arr_of ` A)" + using assms elem_of_arr_of + by (metis (no_types, lifting) Pi_iff arr_mkIde bij_betw_ide_set(2) + ide_implies_incl ide_of_set_def incl_def mem_Collect_eq) + have B: "arr_of ` B \ Univ \ Setp (elem_of ` arr_of ` B)" + using assms elem_of_arr_of + by (metis (no_types, lifting) Pi_iff arr_mkIde bij_betw_ide_set(2) + ide_implies_incl ide_of_set_def incl_def mem_Collect_eq) + have C: "arr_of ` C \ Univ \ Setp (elem_of ` arr_of ` C)" + using assms elem_of_arr_of + by (metis (no_types, lifting) Pi_iff arr_mkIde bij_betw_ide_set(2) + ide_implies_incl ide_of_set_def incl_def mem_Collect_eq) + have "arr_of_fun A C (G o F) = mkArr (arr_of ` A) (arr_of ` C) (arr_of \ (G \ F) \ elem_of)" + unfolding arr_of_fun_def by simp + also have "... = mkArr (arr_of ` A) (arr_of ` C) + ((arr_of \ G \ elem_of) o (arr_of o F \ elem_of))" + proof (intro mkArr_eqI') + show "arr (mkArr (arr_of ` A) (arr_of ` C) (arr_of \ (G \ F) \ elem_of))" + proof - + have "arr_of \ (G \ F) \ elem_of \ arr_of ` A \ arr_of ` C" + using assms by force + thus ?thesis + using A B C by blast + qed + show "\x. x \ arr_of ` A \ + (arr_of \ (G \ F) \ elem_of) x = + ((arr_of \ G \ elem_of) o (arr_of o F \ elem_of)) x" + by simp + qed + also have "... = comp (mkArr (arr_of ` B) (arr_of ` C) (arr_of \ G \ elem_of)) + (mkArr (arr_of ` A) (arr_of ` B) (arr_of o F o elem_of))" + proof - + have "arr (mkArr (arr_of ` B) (arr_of ` C) (arr_of \ G \ elem_of))" + using assms B C elem_of_arr_of by fastforce + moreover have "arr (mkArr (arr_of ` A) (arr_of ` B) (arr_of o F o elem_of))" + using assms A B elem_of_arr_of by fastforce + ultimately show ?thesis + using comp_mkArr by auto + qed + also have "... = comp (arr_of_fun B C G) (arr_of_fun A B F)" + using arr_of_fun_def by presburger + finally show ?thesis by blast qed end - sublocale setcat \ set_category comp \ - using is_set_category by simp - sublocale setcat \ concrete_set_category comp \ UNIV UP - using is_concrete_set_category by simp - text\ - By using a large enough cardinal, we can effectively eliminate the cardinality constraint - on the sets that determine objects and thereby obtain a set category that is replete. - This is the normal use case, which we want to streamline as much as possible, + When there is no restriction on the sets that determine objects, the resulting set category + is replete. This is the normal use case, which we want to streamline as much as possible, so it is useful to introduce a special locale for this purpose. \ locale replete_setcat = fixes dummy :: 'e begin - interpretation SC: setcat dummy - \cardSuc (cmax (card_of (UNIV :: 'e setcat.arr set)) natLeq)\ + interpretation SC: setcat dummy \\_. True\ proof - show "Card_order (cardSuc (cmax (card_of (UNIV :: 'e setcat.arr set)) natLeq)) \ - infinite (Field (cardSuc (cmax (card_of (UNIV :: 'e setcat.arr set)) natLeq)))" - by (metis Card_order_cmax Field_natLeq cardSuc_Card_order cardSuc_finite - card_of_Card_order finite_cmax infinite_UNIV_char_0 natLeq_Card_order) + show True by blast qed - text\ - We don't want to expose the concrete details of the construction used to obtain - the interpretation \SC\; instead, we want any facts proved about it to be derived - solely from the assumptions of the @{locale set_category} locales. - So we create another level of definitions here. - \ - definition comp where "comp \ SC.comp" - definition UP - where "UP \ SC.UP" - - definition DOWN - where "DOWN \ SC.DOWN" - - sublocale set_category comp \cardSuc (cmax (card_of (UNIV :: 'e setcat.arr set)) natLeq)\ - using SC.is_set_category comp_def by simp + definition arr_of + where "arr_of \ SC.arr_of" - sublocale concrete_set_category comp - \cardSuc (cmax (card_of (UNIV :: 'e setcat.arr set)) natLeq)\ UNIV UP - using SC.is_concrete_set_category comp_def UP_def by simp - - sublocale replete_set_category comp .. + definition elem_of + where "elem_of \ SC.elem_of" - lemma UP_mapsto: - shows "UP \ UNIV \ Univ" - using SC.UP_mapsto - by (simp add: UP_def comp_def) + sublocale replete_set_category comp + unfolding comp_def + using SC.set_img_Collect_ide_iff SC.set_category_given_img_axioms + by unfold_locales auto - lemma DOWN_mapsto: - shows "DOWN \ Univ \ UNIV" + sublocale concrete_set_category comp setp UNIV arr_of + using SC.is_concrete_set_category comp_def SC.set_img_Collect_ide_iff arr_of_def by auto - lemma DOWN_UP [simp]: - shows "DOWN (UP x) = x" - by (simp add: DOWN_def UP_def) + lemma bij_arr_of: + shows "bij_betw arr_of UNIV Univ" + using SC.bij_arr_of comp_def arr_of_def by presburger - lemma UP_DOWN [simp]: - assumes "t \ Univ" - shows "UP (DOWN t) = t" - using assms DOWN_def UP_def - by (simp add: DOWN_def UP_def comp_def) - - lemma inj_UP: - shows "inj UP" - by (metis DOWN_UP injI) - - lemma bij_UP: - shows "bij_betw UP UNIV Univ" - by (metis SC.bij_UP UP_def comp_def) + lemma bij_elem_of: + shows "bij_betw elem_of Univ UNIV" + using SC.bij_elem_of comp_def elem_of_def by auto end end diff --git a/thys/Category3/SetCategory.thy b/thys/Category3/SetCategory.thy --- a/thys/Category3/SetCategory.thy +++ b/thys/Category3/SetCategory.thy @@ -1,2475 +1,2676 @@ (* Title: SetCategory Author: Eugene W. Stark , 2016 Maintainer: Eugene W. Stark *) chapter SetCategory theory SetCategory -imports Category Functor "HOL-Cardinals.Cardinals" +imports Category Functor Subcategory begin text\ This theory defines a locale \set_category\ that axiomatizes the notion ``category of @{typ 'a}-sets and functions between them'' in the context of HOL. A primary reason for doing this is to make it possible to prove results (such as the Yoneda Lemma) that use such categories without having to commit to a particular element type @{typ 'a} and without having the results depend on the concrete details of a particular construction. The axiomatization given here is categorical, in the sense that if categories @{term S} and @{term S'} each interpret the \set_category\ locale, then a bijection between the sets of terminal objects of @{term S} and @{term S'} extends to an isomorphism of @{term S} and @{term S'} as categories. The axiomatization is based on the following idea: if, for some type @{typ 'a}, category @{term S} is the category of all @{typ 'a}-sets and functions between them, then the elements of type @{typ 'a} are in bijective correspondence with the terminal objects of category @{term S}. In addition, if @{term unity} is an arbitrarily chosen terminal object of @{term S}, then for each object @{term a}, the hom-set @{term "hom unity a"} (\emph{i.e.} the set of ``points'' or ``global elements'' of @{term a}) is in bijective correspondence with a subset of the terminal objects of @{term S}. By making a specific, but arbitrary, choice of such a correspondence, we can then associate with each object @{term a} of @{term S} a set @{term "set a"} that consists of all terminal objects @{term t} that correspond to some point @{term x} of @{term a}. Each arrow @{term f} then induces a function \Fun f \ set (dom f) \ set (cod f)\, defined on terminal objects of @{term S} by passing to points of @{term "dom f"}, composing with @{term f}, then passing back from points of @{term "cod f"} to terminal objects. Once we can associate a set with each object of @{term S} and a function with each arrow, we can force @{term S} to be isomorphic to the category of @{typ 'a}-sets by imposing suitable extensionality and completeness axioms. \ section "Some Lemmas about Restriction" text\ \sloppypar The development of the \set_category\ locale makes heavy use of the theory @{theory "HOL-Library.FuncSet"}. However, in some cases, I found that that theory did not provide results about restriction in the form that was most useful to me. I used the following additional results in various places. \ (* This variant of FuncSet.restrict_ext is sometimes useful. *) lemma restr_eqI: assumes "A = A'" and "\x. x \ A \ F x = F' x" shows "restrict F A = restrict F' A'" using assms by force (* This rule avoided a long proof in at least one instance where FuncSet.restrict_apply did not work. *) lemma restr_eqE [elim]: assumes "restrict F A = restrict F' A" and "x \ A" shows "F x = F' x" using assms restrict_def by metis (* This seems more useful than compose_eq in FuncSet. *) lemma compose_eq' [simp]: shows "compose A G F = restrict (G o F) A" unfolding compose_def restrict_def by auto section "Set Categories" text\ We first define the locale \set_category_data\, which sets out the basic data and definitions for the \set_category\ locale, without imposing any conditions other than that @{term S} is a category and that @{term img} is a function defined on the arrow type of @{term S}. The function @{term img} should be thought of as a mapping that takes a point @{term "x \ hom unity a"} to a corresponding terminal object @{term "img x"}. Eventually, assumptions will be introduced so - that this is in fact the case. + that this is in fact the case. The set of terminal objects of the category will + serve as abstract ``elements'' of sets; we will refer to the set of \emph{all} + terminal objects as the \emph{universe}. \ locale set_category_data = category S for S :: "'s comp" (infixr "\" 55) and img :: "'s \ 's" begin notation in_hom ("\_ : _ \ _\") text\ Call the set of all terminal objects of S the ``universe''. \ abbreviation Univ :: "'s set" where "Univ \ Collect terminal" text\ Choose an arbitrary element of the universe and call it @{term unity}. \ definition unity :: 's where "unity = (SOME t. terminal t)" text\ Each object @{term a} determines a subset @{term "set a"} of the universe, consisting of all those terminal objects @{term t} such that @{term "t = img x"} for some @{term "x \ hom unity a"}. \ definition set :: "'s \ 's set" where "set a = img ` hom unity a" end text\ Next, we define a locale \set_category_given_img\ that augments the \set_category_data\ locale with assumptions that serve to define the notion of a set category with a chosen correspondence between points and terminal objects. The assumptions require that the universe be nonempty (so that the definition of @{term unity} makes sense), that the map @{term img} is a locally injective map taking points to terminal objects, that each terminal object @{term t} belongs to @{term "set t"}, that two objects of @{term S} are equal if they determine the same set, that two parallel arrows of @{term S} are equal if they determine the same - function, that there is an object corresponding to each subset of the universe - whose cardinality is less than a specified cardinal \\\, and that for any objects - @{term a} and @{term b} and function @{term "F \ hom unity a \ hom unity b"} - there is an arrow @{term "f \ hom a b"} whose action under the composition - of @{term S} coincides with the function @{term F}. + function, and that for any objects @{term a} and @{term b} and function + @{term "F \ hom unity a \ hom unity b"} there is an arrow @{term "f \ hom a b"} + whose action under the composition of @{term S} coincides with the function @{term F}. - The cardinal \\\, which is given as a parameter to the locale, has been introduced - because most of the familiar properties of a set category do not depend on - there being an object corresponding to \emph{every} subset of the universe, - and we would like to consider such situations; for example, the situation in - which there only \emph{finite} subsets determine objects. + The parameter @{term setp} is a predicate that determines which subsets of the + universe are to be regarded as defining objects of the category. This parameter + has been introduced because most of the characteristic properties of a category + of sets and functions do not depend on there being an object corresponding to + \emph{every} subset of the universe, and we intend to consider in particular the + cases in which only finite subsets or only ``small'' subsets of the universe + determine objects. Accordingly, we assume that there is an object corresponding + to each subset of the universe that satisfies @{term setp}. + It is also necessary to assume some basic regularity properties of the predicate + @{term setp}; namely, that it holds for all subsets of the universe corresponding + to objects of @{term S}, and that it respects subset and union. \ locale set_category_given_img = set_category_data S img for S :: "'s comp" (infixr "\" 55) and img :: "'s \ 's" - and \ :: "'t rel" + - assumes nonempty_Univ: "Univ \ {}" - and img_mapsto: "ide a \ img \ hom unity a \ Univ" - and card_points: "ide a \ |hom unity a| " + and setp :: "'s set \ bool" + + assumes setp_imp_subset_Univ: "setp A \ A \ Univ" + and setp_set_ide: "ide a \ setp (set a)" + and setp_respects_subset: "A' \ A \ setp A \ setp A'" + and setp_respects_union: "\ setp A; setp B \ \ setp (A \ B)" + and nonempty_Univ: "Univ \ {}" and inj_img: "ide a \ inj_on img (hom unity a)" and stable_img: "terminal t \ t \ img ` hom unity t" and extensional_set: "\ ide a; ide b; set a = set b \ \ a = b" and extensional_arr: "\ par f f'; \x. \x : unity \ dom f\ \ f \ x = f' \ x \ \ f = f'" - and set_complete: "\ A \ Univ; |A| \ \ \a. ide a \ set a = A" - and fun_complete1: "\ ide a; ide b; F \ hom unity a \ hom unity b \ - \ \f. \f : a \ b\ \ (\x. \x : unity \ dom f\ \ f \ x = F x)" - and cardinal: "Card_order \ \ infinite (Field \)" + and set_complete: "setp A \ \a. ide a \ set a = A" + and fun_complete_ax: "\ ide a; ide b; F \ hom unity a \ hom unity b \ + \ \f. \f : a \ b\ \ (\x. \x : unity \ dom f\ \ f \ x = F x)" begin - text\ - The inverse of the map @{term set} is a map @{term mkIde} that takes each subset - of the universe of cardinality less than \\\ to an identity of @{term[source=true] S}. -\ - definition mkIde :: "'s set \ 's" - where "mkIde A = (if A \ Univ \ |A| then inv_into (Collect ide) set A else null)" + lemma setp_singleton: + assumes "terminal a" + shows "setp {a}" + using assms + by (metis setp_set_ide Set.set_insert Un_upper1 insert_is_Un set_def + setp_respects_subset stable_img terminal_def) + + lemma setp_empty: + shows "setp {}" + using setp_singleton setp_respects_subset nonempty_Univ by blast + + lemma finite_imp_setp: + assumes "A \ Univ" and "finite A" + shows "setp A" + using setp_empty setp_singleton setp_respects_union + by (metis assms(1-2) finite_subset_induct insert_is_Un mem_Collect_eq) text\ Each arrow @{term "f \ hom a b"} determines a function @{term "Fun f \ Univ \ Univ"}, by passing from @{term Univ} to @{term "hom a unity"}, composing with @{term f}, then passing back to @{term Univ}. \ definition Fun :: "'s \ 's \ 's" where "Fun f = restrict (img o S f o inv_into (hom unity (dom f)) img) (set (dom f))" lemma comp_arr_point: assumes "arr f" and "\x : unity \ dom f\" shows "f \ x = inv_into (hom unity (cod f)) img (Fun f (img x))" proof - have "\f \ x : unity \ cod f\" using assms by blast thus ?thesis using assms Fun_def inj_img set_def by simp qed text\ Parallel arrows that determine the same function are equal. \ lemma arr_eqI: assumes "par f f'" and "Fun f = Fun f'" shows "f = f'" using assms comp_arr_point extensional_arr by metis lemma terminal_unity: shows "terminal unity" using unity_def nonempty_Univ by (simp add: someI_ex) lemma ide_unity [simp]: shows "ide unity" using terminal_unity terminal_def by blast - lemma set_subset_Univ [simp]: + lemma setp_set' [simp]: assumes "ide a" - shows "set a \ Univ" - using assms set_def img_mapsto by auto + shows "setp (set a)" + using assms setp_set_ide by auto lemma inj_on_set: shows "inj_on set (Collect ide)" using extensional_set by (intro inj_onI, auto) text\ - The mapping @{term mkIde}, which takes subsets of the universe to identities, - and @{term set}, which takes identities to subsets of the universe, are inverses. + The inverse of the map @{term set} is a map @{term mkIde} that takes each subset + of the universe to an identity of @{term[source=true] S}. \ - - lemma set_card: - assumes "ide a" - shows "|set a| " - using assms card_points set_def - by (metis card_of_image ordLeq_ordLess_trans) + definition mkIde :: "'s set \ 's" + where "mkIde A = (if setp A then inv_into (Collect ide) set A else null)" lemma mkIde_set [simp]: assumes "ide a" shows "mkIde (set a) = a" - using assms mkIde_def inj_on_set inv_into_f_f set_card - by (simp add: ordLess_imp_ordLeq) + by (simp add: assms inj_on_set mkIde_def) - lemma set_mkIde: - assumes "A \ Univ" and "|A| " + lemma set_mkIde [simp]: + assumes "setp A" shows "set (mkIde A) = A" using assms mkIde_def set_complete someI_ex [of "\a. a \ Collect ide \ set a = A"] - by (metis set_category_given_img.mkIde_set set_category_given_img_axioms) + mkIde_set + by metis - lemma ide_mkIde: - assumes "A \ Univ" and "|A| " + lemma ide_mkIde [simp]: + assumes "setp A" shows "ide (mkIde A)" using assms mkIde_def mkIde_set set_complete by metis - text\ - Because we have assumed the cardinal \\\ to be infinite, there is an object corresponding - to every finite subset of the universe. - \ - - lemma ide_mkIde_finite: - assumes "A \ Univ" and "finite A" - shows "ide (mkIde A)" - by (meson assms card_of_Field_ordIso cardinal finite_ordLess_infinite2 ide_mkIde - ordLess_ordIso_trans) - - lemma arr_mkIde: - shows "arr (mkIde A) \ A \ Univ \ |A| " - using ide_mkIde mkIde_def not_arr_null by force + lemma arr_mkIde [iff]: + shows "arr (mkIde A) \ setp A" + using not_arr_null mkIde_def ide_mkIde by auto - lemma dom_mkIde: - assumes "A \ Univ" and "|A| " + lemma dom_mkIde [simp]: + assumes "setp A" shows "dom (mkIde A) = mkIde A" using assms ide_mkIde by simp - lemma cod_mkIde: - assumes "A \ Univ" and "|A| " + lemma cod_mkIde [simp]: + assumes "setp A" shows "cod (mkIde A) = mkIde A" using assms ide_mkIde by simp text\ Each arrow @{term f} determines an extensional function from @{term "set (dom f)"} to @{term "set (cod f)"}. \ lemma Fun_mapsto: assumes "arr f" shows "Fun f \ extensional (set (dom f)) \ (set (dom f) \ set (cod f))" proof show "Fun f \ extensional (set (dom f))" using Fun_def by fastforce show "Fun f \ set (dom f) \ set (cod f)" proof fix t assume t: "t \ set (dom f)" have "Fun f t = img (f \ inv_into (hom unity (dom f)) img t)" using assms t Fun_def comp_def by simp moreover have "... \ set (cod f)" using assms t set_def inv_into_into [of t img "hom unity (dom f)"] by blast ultimately show "Fun f t \ set (cod f)" by auto qed qed text\ Identities of @{term[source=true] S} correspond to restrictions of the identity function. \ lemma Fun_ide: assumes "ide a" shows "Fun a = restrict (\x. x) (set a)" using assms Fun_def inj_img set_def comp_cod_arr by fastforce - lemma Fun_mkIde: - assumes "A \ Univ" and "|A| " + lemma Fun_mkIde [simp]: + assumes "setp A" shows "Fun (mkIde A) = restrict (\x. x) A" using assms ide_mkIde set_mkIde Fun_ide by simp text\ Composition in @{term S} corresponds to extensional function composition. \ lemma Fun_comp [simp]: assumes "seq g f" shows "Fun (g \ f) = restrict (Fun g o Fun f) (set (dom f))" proof - have "restrict (img o S (g \ f) o (inv_into (hom unity (dom (g \ f))) img)) (set (dom (g \ f))) = restrict (Fun g o Fun f) (set (dom f))" proof - let ?img' = "\a. \t. inv_into (hom unity a) img t" have 1: "set (dom (g \ f)) = set (dom f)" using assms by auto moreover have "\t. t \ set (dom (g \ f)) \ (img o S (g \ f) o ?img' (dom (g \ f))) t = (Fun g o Fun f) t" proof - fix t assume "t \ set (dom (g \ f))" hence t: "t \ set (dom f)" by (simp add: 1) have "(img o S (g \ f) o ?img' (dom (g \ f))) t = img (g \ f \ ?img' (dom f) t)" using assms dom_comp comp_assoc by simp also have "... = img (g \ ?img' (dom g) (Fun f t))" proof - have "\a x. x \ hom unity a \ ?img' a (img x) = x" using assms inj_img ide_cod inv_into_f_eq by (metis arrI in_homE mem_Collect_eq) thus ?thesis using assms t Fun_def set_def comp_arr_point by auto qed also have "... = Fun g (Fun f t)" proof - have "Fun f t \ img ` hom unity (cod f)" using assms t Fun_mapsto set_def by fast - thus ?thesis using assms by (auto simp add: set_def Fun_def) + thus ?thesis + using assms by (auto simp add: set_def Fun_def) qed finally show "(img o S (g \ f) o ?img' (dom (g \ f))) t = (Fun g o Fun f) t" by auto qed ultimately show ?thesis by auto qed thus ?thesis using Fun_def by auto qed text\ The constructor @{term mkArr} is used to obtain an arrow given subsets @{term A} and @{term B} of the universe and a function @{term "F \ A \ B"}. \ definition mkArr :: "'s set \ 's set \ ('s \ 's) \ 's" - where "mkArr A B F = (if A \ Univ \ |A| \ B \ Univ \ |B| \ F \ A \ B + where "mkArr A B F = (if setp A \ setp B \ F \ A \ B then (THE f. f \ hom (mkIde A) (mkIde B) \ Fun f = restrict F A) else null)" text\ Each function @{term "F \ set a \ set b"} determines a unique arrow @{term "f \ hom a b"}, such that @{term "Fun f"} is the restriction of @{term F} to @{term "set a"}. \ lemma fun_complete: assumes "ide a" and "ide b" and "F \ set a \ set b" shows "\!f. \f : a \ b\ \ Fun f = restrict F (set a)" proof - let ?P = "\f. \f : a \ b\ \ Fun f = restrict F (set a)" show "\!f. ?P f" proof have "\f. ?P f" proof - let ?F' = "\x. inv_into (hom unity b) img (F (img x))" have "?F' \ hom unity a \ hom unity b" proof fix x assume x: "x \ hom unity a" have "F (img x) \ set b" using assms(3) x set_def by auto thus "inv_into (hom unity b) img (F (img x)) \ hom unity b" - using assms inj_img set_def by auto + using assms setp_set_ide inj_img set_def by auto qed hence "\f. \f : a \ b\ \ (\x. \x : unity \ a\ \ f \ x = ?F' x)" - using assms fun_complete1 [of a b] by force + using assms fun_complete_ax [of a b] by force from this obtain f where f: "\f : a \ b\ \ (\x. \x : unity \ a\ \ f \ x = ?F' x)" by blast let ?img' = "\a. \t. inv_into (hom unity a) img t" have "Fun f = restrict F (set a)" proof (unfold Fun_def, intro restr_eqI) show "set (dom f) = set a" using f by auto show "\t. t \ set (dom f) \ (img \ S f \ ?img' (dom f)) t = F t" proof - fix t assume t: "t \ set (dom f)" have "(img \ S f \ ?img' (dom f)) t = img (f \ ?img' (dom f) t)" by simp also have "... = img (?F' (?img' (dom f) t))" by (metis f in_homE inv_into_into set_def mem_Collect_eq t) also have "... = img (?img' (cod f) (F t))" using f t set_def inj_img by auto also have "... = F t" proof - have "F t \ set (cod f)" using assms f t by auto thus ?thesis using f t set_def inj_img by auto qed finally show "(img \ S f \ ?img' (dom f)) t = F t" by auto qed qed thus ?thesis using f by blast qed thus F: "?P (SOME f. ?P f)" using someI_ex [of ?P] by fast show "\f'. ?P f' \ f' = (SOME f. ?P f)" using F arr_eqI by (metis (no_types, lifting) in_homE) qed qed lemma mkArr_in_hom: - assumes "A \ Univ" and "|A| " and "B \ Univ" and "|B| " and "F \ A \ B" + assumes "setp A" and "setp B" and "F \ A \ B" shows "\mkArr A B F : mkIde A \ mkIde B\" using assms mkArr_def fun_complete [of "mkIde A" "mkIde B" F] ide_mkIde set_mkIde theI' [of "\f. f \ hom (mkIde A) (mkIde B) \ Fun f = restrict F A"] + setp_imp_subset_Univ by simp text\ The ``only if'' direction of the next lemma can be achieved only if there exists a non-arrow element of type @{typ 's}, which can be used as the value of @{term "mkArr A B F"} in cases where @{term "F \ A \ B"}. Nevertheless, it is essential to have this, because without the ``only if'' direction, we can't derive any useful consequences from an assumption of the form @{term "arr (mkArr A B F)"}; instead we have to obtain @{term "F \ A \ B"} some other way. This is is usually highly inconvenient and it makes the theory very weak and almost unusable in practice. The observation that having a non-arrow value of type @{typ 's} solves this problem is ultimately what led me to incorporate @{term null} first into the definition of the \set_category\ locale and then, ultimately, into the definition of the \category\ locale. I believe this idea is critical to the usability of the entire development. \ - (* TODO: This gets used as an introduction rule, but the conjunction on the right-hand side - is not very convenient. *) - lemma arr_mkArr: - shows "arr (mkArr A B F) \ - A \ Univ \ |A| \ B \ Univ \ |B| \ F \ A \ B" + lemma arr_mkArr [iff]: + shows "arr (mkArr A B F) \ setp A \ setp B \ F \ A \ B" proof - show "arr (mkArr A B F) \ - A \ Univ \ |A| \ B \ Univ \ |B| \ F \ A \ B" - using mkArr_def domains_null codomains_null has_domain_iff_arr has_codomain_iff_arr - not_arr_null - by (intro conjI) metis+ - show "A \ Univ \ |A| \ B \ Univ \ |B| \ F \ A \ B - \ arr (mkArr A B F)" + show "arr (mkArr A B F) \ setp A \ setp B \ F \ A \ B" + using mkArr_def not_arr_null ex_un_null someI_ex [of "\f. \arr f"] setp_imp_subset_Univ + by metis + show "setp A \ setp B \ F \ A \ B \ arr (mkArr A B F)" using mkArr_in_hom by auto qed + lemma arr_mkArrI [intro]: + assumes "setp A" and "setp B" and "F \ A \ B" + shows "arr (mkArr A B F)" + using assms arr_mkArr by blast + lemma Fun_mkArr': assumes "arr (mkArr A B F)" shows "\mkArr A B F : mkIde A \ mkIde B\" and "Fun (mkArr A B F) = restrict F A" proof - - have 1: "A \ Univ \ |A| \ B \ Univ \ |B| \ F \ A \ B" - using assms arr_mkArr by simp + have 1: "setp A \ setp B \ F \ A \ B" using assms by fast have 2: "mkArr A B F \ hom (mkIde A) (mkIde B) \ Fun (mkArr A B F) = restrict F (set (mkIde A))" proof - have "\!f. f \ hom (mkIde A) (mkIde B) \ Fun f = restrict F (set (mkIde A))" using 1 fun_complete [of "mkIde A" "mkIde B" F] ide_mkIde set_mkIde by simp thus ?thesis using 1 mkArr_def theI' set_mkIde by simp qed show "\mkArr A B F : mkIde A \ mkIde B\" using 1 2 by auto show "Fun (mkArr A B F) = restrict F A" using 1 2 set_mkIde by auto qed lemma mkArr_Fun: assumes "arr f" shows "mkArr (set (dom f)) (set (cod f)) (Fun f) = f" proof - - have 1: "set (dom f) \ Univ \ |set (dom f)| \ - set (cod f) \ Univ \ |set (cod f)| \ - ide (dom f) \ ide (cod f) \ + have 1: "setp (set (dom f)) \ setp (set (cod f)) \ ide (dom f) \ ide (cod f) \ Fun f \ extensional (set (dom f)) \ (set (dom f) \ set (cod f))" - using assms Fun_mapsto set_def set_card set_subset_Univ by auto + using Fun_mapsto assms ide_cod ide_dom setp_set' by presburger hence "\!f'. f' \ hom (dom f) (cod f) \ Fun f' = restrict (Fun f) (set (dom f))" using fun_complete by force moreover have "f \ hom (dom f) (cod f) \ Fun f = restrict (Fun f) (set (dom f))" using assms 1 extensional_restrict by force ultimately have "f = (THE f'. f' \ hom (dom f) (cod f) \ Fun f' = restrict (Fun f) (set (dom f)))" using theI' [of "\f'. f' \ hom (dom f) (cod f) \ Fun f' = restrict (Fun f) (set (dom f))"] by blast also have "... = mkArr (set (dom f)) (set (cod f)) (Fun f)" using assms 1 mkArr_def mkIde_set by simp finally show ?thesis by auto qed lemma dom_mkArr [simp]: assumes "arr (mkArr A B F)" shows "dom (mkArr A B F) = mkIde A" using assms Fun_mkArr' by auto lemma cod_mkArr [simp]: assumes "arr (mkArr A B F)" shows "cod (mkArr A B F) = mkIde B" using assms Fun_mkArr' by auto lemma Fun_mkArr [simp]: assumes "arr (mkArr A B F)" shows "Fun (mkArr A B F) = restrict F A" using assms Fun_mkArr' by auto text\ The following provides the basic technique for showing that arrows constructed using @{term mkArr} are equal. \ lemma mkArr_eqI [intro]: assumes "arr (mkArr A B F)" and "A = A'" and "B = B'" and "\x. x \ A \ F x = F' x" shows "mkArr A B F = mkArr A' B' F'" - using assms arr_mkArr Fun_mkArr + using assms Fun_mkArr by (intro arr_eqI, auto simp add: Pi_iff) text\ This version avoids trivial proof obligations when the domain and codomain sets are identical from the context. \ lemma mkArr_eqI' [intro]: assumes "arr (mkArr A B F)" and "\x. x \ A \ F x = F' x" shows "mkArr A B F = mkArr A B F'" - using assms mkArr_eqI arr_mkArr by simp + using assms mkArr_eqI by simp lemma mkArr_restrict_eq: assumes "arr (mkArr A B F)" shows "mkArr A B (restrict F A) = mkArr A B F" - using assms arr_mkArr by (intro mkArr_eqI', auto) + using assms by (intro mkArr_eqI', auto) lemma mkArr_restrict_eq': assumes "arr (mkArr A B (restrict F A))" shows "mkArr A B (restrict F A) = mkArr A B F" using assms by (intro mkArr_eqI', auto) - lemma mkIde_as_mkArr: - assumes "A \ Univ" and "|A| " + lemma mkIde_as_mkArr [simp]: + assumes "setp A" shows "mkArr A A (\x. x) = mkIde A" - using assms arr_mkIde arr_mkArr dom_mkIde cod_mkIde Fun_mkIde + using assms arr_mkIde dom_mkIde cod_mkIde Fun_mkIde by (intro arr_eqI, auto) lemma comp_mkArr: assumes "arr (mkArr A B F)" and "arr (mkArr B C G)" shows "mkArr B C G \ mkArr A B F = mkArr A C (G \ F)" proof (intro arr_eqI) have 1: "seq (mkArr B C G) (mkArr A B F)" using assms by force - have 2: "G o F \ A \ C" using assms arr_mkArr by auto + have 2: "G o F \ A \ C" using assms by auto show "par (mkArr B C G \ mkArr A B F) (mkArr A C (G \ F))" - using assms 1 2 arr_mkArr + using assms 1 2 by (intro conjI) simp_all show "Fun (mkArr B C G \ mkArr A B F) = Fun (mkArr A C (G \ F))" - using 1 2 arr_mkArr set_mkIde by fastforce + using 1 2 by fastforce qed text\ The locale assumption @{prop stable_img} forces @{term "t \ set t"} in case @{term t} is a terminal object. This is very convenient, as it results in the characterization of terminal objects as identities @{term t} for which @{term "set t = {t}"}. However, it is not absolutely necessary to have this. The following weaker characterization of terminal objects can be proved without the @{prop stable_img} assumption. \ lemma terminal_char1: shows "terminal t \ ide t \ (\!x. x \ set t)" proof - have "terminal t \ ide t \ (\!x. x \ set t)" proof - assume t: "terminal t" have "ide t" using t terminal_def by auto moreover have "\!x. x \ set t" proof - have "\!x. x \ hom unity t" using t terminal_unity terminal_def by auto thus ?thesis using set_def by auto qed ultimately show "ide t \ (\!x. x \ set t)" by auto qed moreover have "ide t \ (\!x. x \ set t) \ terminal t" proof - assume t: "ide t \ (\!x. x \ set t)" from this obtain t' where "set t = {t'}" by blast - hence t': "set t = {t'} \ {t'} \ Univ \ t = mkIde {t'}" - using t set_subset_Univ mkIde_set by metis + hence t': "set t = {t'} \ setp {t'} \ t = mkIde {t'}" + using t setp_set_ide mkIde_set by metis show "terminal t" proof show "ide t" using t by simp show "\a. ide a \ \!f. \f : a \ t\" proof - fix a assume a: "ide a" show "\!f. \f : a \ t\" proof show 1: "\mkArr (set a) {t'} (\x. t') : a \ t\" - by (metis Pi_I a mkIde_set set_card mkArr_in_hom set_subset_Univ - singletonD t t') + using a t t' mkArr_in_hom + by (metis Pi_I' mkIde_set setp_set_ide singletonD) show "\f. \f : a \ t\ \ f = mkArr (set a) {t'} (\x. t')" proof - fix f assume f: "\f : a \ t\" show "f = mkArr (set a) {t'} (\x. t')" proof (intro arr_eqI) show 1: "par f (mkArr (set a) {t'} (\x. t'))" using 1 f in_homE by metis show "Fun f = Fun (mkArr (set a) {t'} (\x. t'))" proof - have "Fun (mkArr (set a) {t'} (\x. t')) = (\x\set a. t')" using 1 Fun_mkArr by simp also have "... = Fun f" proof - have "\x. x \ set a \ Fun f x = t'" using f t' Fun_def mkArr_Fun arr_mkArr by (metis PiE in_homE singletonD) moreover have "\x. x \ set a \ Fun f x = undefined" using f Fun_def by auto ultimately show ?thesis by auto qed finally show ?thesis by force qed qed qed qed qed qed qed ultimately show ?thesis by blast qed text\ As stated above, in the presence of the @{prop stable_img} assumption we have the following stronger characterization of terminal objects. \ lemma terminal_char2: shows "terminal t \ ide t \ set t = {t}" proof assume t: "terminal t" show "ide t \ set t = {t}" proof show "ide t" using t terminal_char1 by auto show "set t = {t}" proof - have "\!x. x \ hom unity t" using t terminal_def terminal_unity by force moreover have "t \ img ` hom unity t" using t stable_img set_def by simp ultimately show ?thesis using set_def by auto qed qed next assume "ide t \ set t = {t}" thus "terminal t" using terminal_char1 by force qed end text\ At last, we define the \set_category\ locale by existentially quantifying out the choice of a particular @{term img} map. We need to know that such a map exists, but it does not matter which one we choose. \ locale set_category = category S for S :: "'s comp" (infixr "\" 55) - and \ :: "'t rel" + - assumes ex_img: "\img. set_category_given_img S img \" + and setp :: "'s set \ bool" + + assumes ex_img: "\img. set_category_given_img S img setp" begin notation in_hom ("\_ : _ \ _\") definition some_img - where "some_img = (SOME img. set_category_given_img S img \)" + where "some_img = (SOME img. set_category_given_img S img setp)" + sublocale set_category_given_img S some_img setp + proof - + have "\img. set_category_given_img S img setp" using ex_img by auto + thus "set_category_given_img S some_img setp" + using someI_ex [of "\img. set_category_given_img S img setp"] some_img_def + by metis + qed + end - sublocale set_category \ set_category_given_img S some_img \ - proof - - have "\img. set_category_given_img S img \" using ex_img by auto - thus "set_category_given_img S some_img \" - using someI_ex [of "\img. set_category_given_img S img \"] some_img_def - by metis - qed - - text\ - For a set category, if the cardinal \\\ is large enough, then it imposes no constraint - on what subsets of the universe determine objects. In this case, we call the set category - \emph{replete} and we can eliminate the cardinality assumptions from various facts. - \ + text\We call a set category \emph{replete} if there is an object corresponding to + every subset of the universe. +\ locale replete_set_category = - set_category S \cardSuc (cmax (card_of (UNIV :: 's set)) natLeq)\ + category S + + set_category S \\A. A \ Collect terminal\ for S :: "'s comp" (infixr "\" 55) begin - lemma card_of_leq: - assumes "A \ Univ" - shows "|A| o cmax (card_of (UNIV :: 's set)) natLeq" - using assms card_of_Card_order natLeq_Card_order ordLeq_cmax1 - ordLeq_transitive - by (metis card_of_UNIV) - thus ?thesis - by (simp add: Card_order_cmax natLeq_Card_order) - qed - - lemma set_mkIde [simp]: - assumes "A \ Univ" - shows "set (mkIde A) = A" - using assms card_of_leq set_mkIde by simp - - lemma ide_mkIde [simp]: - assumes "A \ Univ" - shows "ide (mkIde A)" - using assms card_of_leq ide_mkIde by simp - - lemma arr_mkIde [iff]: - shows "arr (mkIde A) \ A \ Univ" - using card_of_leq arr_mkIde by auto - - lemma dom_mkIde [simp]: - assumes "A \ Univ" - shows "dom (mkIde A) = mkIde A" - using assms ide_mkIde by simp - - lemma cod_mkIde [simp]: - assumes "A \ Univ" - shows "cod (mkIde A) = mkIde A" - using assms ide_mkIde by simp - - lemma Fun_mkIde [simp]: - assumes "A \ Univ" - shows "Fun (mkIde A) = restrict (\x. x) A" - using assms set_mkIde ide_mkIde Fun_ide by simp - - lemma mkArr_in_hom [intro]: - assumes "A \ Univ" and "B \ Univ" and "F \ A \ B" - shows "\mkArr A B F : mkIde A \ mkIde B\" - using assms card_of_leq arr_mkArr by auto - - lemma arr_mkArr: - shows "arr (mkArr A B F) \ A \ Univ \ B \ Univ \ F \ A \ B" - using card_of_leq arr_mkArr by auto - - lemma mkIde_as_mkArr: - assumes "A \ Univ" - shows "mkArr A A (\x. x) = mkIde A" - using assms card_of_leq set_mkIde arr_mkIde arr_mkArr dom_mkIde cod_mkIde Fun_mkIde - by (intro arr_eqI, auto) + abbreviation setp + where "setp \ \A. A \ Univ" end context set_category begin text\ The arbitrary choice of @{term img} induces a system of arrows corresponding to inclusions of subsets. \ definition incl :: "'s \ bool" where "incl f = (arr f \ set (dom f) \ set (cod f) \ f = mkArr (set (dom f)) (set (cod f)) (\x. x))" lemma Fun_incl: assumes "incl f" shows "Fun f = (\x \ set (dom f). x)" using assms incl_def by (metis Fun_mkArr) lemma ex_incl_iff_subset: assumes "ide a" and "ide b" shows "(\f. \f : a \ b\ \ incl f) \ set a \ set b" proof show "\f. \f : a \ b\ \ incl f \ set a \ set b" using incl_def by auto show "set a \ set b \ \f. \f : a \ b\ \ incl f" proof assume 1: "set a \ set b" show "\mkArr (set a) (set b) (\x. x) : a \ b\ \ incl (mkArr (set a) (set b) (\x. x))" proof show "\mkArr (set a) (set b) (\x. x) : a \ b\" - by (metis 1 assms image_ident image_subset_iff_funcset mkIde_set set_card - mkArr_in_hom set_subset_Univ) + by (metis 1 assms image_ident image_subset_iff_funcset mkIde_set + mkArr_in_hom setp_set_ide) thus "incl (mkArr (set a) (set b) (\x. x))" using 1 incl_def by force qed qed qed end section "Categoricity" text\ In this section we show that the \set_category\ locale completely characterizes the structure of its interpretations as categories, in the sense that for any two - interpretations @{term S} and @{term S'} for the same cardinal \\\, - a bijection between the universe of @{term S} and the universe of @{term S'} extends + interpretations @{term S} and @{term S'}, a @{term setp}-respecting bijection + between the universe of @{term S} and the universe of @{term S'} extends to an isomorphism of @{term S} and @{term S'}. \ locale two_set_categories_bij_betw_Univ = - S: set_category S \ + - S': set_category S' \ + S: set_category S setp + + S': set_category S' setp' for S :: "'s comp" (infixr "\" 55) + and setp :: "'s set \ bool" and S' :: "'t comp" (infixr "\\" 55) - and \ :: "'u rel" + and setp' :: "'t set \ bool" and \ :: "'s \ 't" + assumes bij_\: "bij_betw \ S.Univ S'.Univ" + and \_respects_setp: "A \ S.Univ \ setp' (\ ` A) \ setp A" begin notation S.in_hom ("\_ : _ \ _\") notation S'.in_hom ("\_ : _ \'' _\") abbreviation \ where "\ \ inv_into S.Univ \" lemma \_\: assumes "t \ S.Univ" shows "\ (\ t) = t" using assms bij_\ bij_betw_inv_into_left by metis lemma \_\: assumes "t' \ S'.Univ" shows "\ (\ t') = t'" using assms bij_\ bij_betw_inv_into_right by metis lemma \_img_\_img: assumes "A \ S.Univ" shows "\ ` \ ` A = A" using assms bij_\ by (simp add: bij_betw_def) lemma \_img_\_img: assumes "A' \ S'.Univ" shows "\ ` \ ` A' = A'" using assms bij_\ by (simp add: bij_betw_def image_inv_into_cancel) text\ We define the object map @{term \o} of a functor from @{term[source=true] S} to @{term[source=true] S'}. \ definition \o where "\o = (\a \ Collect S.ide. S'.mkIde (\ ` S.set a))" lemma set_\o: assumes "S.ide a" shows "S'.set (\o a) = \ ` S.set a" - proof - - from assms have "S.set a \ S.Univ \ |S.set a| " - by (simp add: S.set_card) - moreover have "|\ ` S.set a| " - by (meson calculation card_of_image ordLeq_ordLess_trans) - ultimately show ?thesis - using S'.set_mkIde \o_def assms bij_\ bij_betw_def image_mono mem_Collect_eq restrict_def - by (metis (no_types, lifting)) - qed + by (simp add: S.setp_imp_subset_Univ \o_def \_respects_setp assms) lemma \o_preserves_ide: assumes "S.ide a" shows "S'.ide (\o a)" - using assms S'.ide_mkIde S.set_subset_Univ bij_\ bij_betw_def image_mono restrict_apply' + using assms S'.ide_mkIde bij_\ bij_betw_def image_mono restrict_apply' S.setp_set' + \_respects_setp S.setp_imp_subset_Univ unfolding \o_def - by (metis (no_types, lifting) S.set_card card_of_image mem_Collect_eq ordLeq_ordLess_trans) + by simp text\ The map @{term \a} assigns to each arrow @{term f} of @{term[source=true] S} the function on the universe of @{term[source=true] S'} that is the same as the function induced by @{term f} on the universe of @{term[source=true] S}, up to the bijection @{term \} between the two universes. \ definition \a where "\a = (\f. \x' \ \ ` S.set (S.dom f). \ (S.Fun f (\ x')))" lemma \a_mapsto: assumes "S.arr f" shows "\a f \ S'.set (\o (S.dom f)) \ S'.set (\o (S.cod f))" proof - have "\a f \ \ ` S.set (S.dom f) \ \ ` S.set (S.cod f)" proof fix x assume x: "x \ \ ` S.set (S.dom f)" have "\ x \ S.set (S.dom f)" - using assms x \_img_\_img [of "S.set (S.dom f)"] S.set_subset_Univ by auto + using assms x \_img_\_img [of "S.set (S.dom f)"] S.setp_imp_subset_Univ by auto hence "S.Fun f (\ x) \ S.set (S.cod f)" using assms S.Fun_mapsto by auto hence "\ (S.Fun f (\ x)) \ \ ` S.set (S.cod f)" by simp thus "\a f x \ \ ` S.set (S.cod f)" using x \a_def by auto qed thus ?thesis using assms set_\o \o_preserves_ide by auto qed text\ The map @{term \a} takes composition of arrows to extensional composition of functions. \ lemma \a_comp: assumes gf: "S.seq g f" shows "\a (g \ f) = restrict (\a g o \a f) (S'.set (\o (S.dom f)))" proof - have "\a (g \ f) = (\x' \ \ ` S.set (S.dom f). \ (S.Fun (S g f) (\ x')))" using gf \a_def by auto also have "... = (\x' \ \ ` S.set (S.dom f). \ (restrict (S.Fun g o S.Fun f) (S.set (S.dom f)) (\ x')))" using gf set_\o S.Fun_comp by simp also have "... = restrict (\a g o \a f) (S'.set (\o (S.dom f)))" proof - have "\x'. x' \ \ ` S.set (S.dom f) \ \ (restrict (S.Fun g o S.Fun f) (S.set (S.dom f)) (\ x')) = \a g (\a f x')" proof - fix x' assume X': "x' \ \ ` S.set (S.dom f)" hence 1: "\ x' \ S.set (S.dom f)" - using gf \_img_\_img [of "S.set (S.dom f)"] S.set_subset_Univ S.ide_dom by blast + using gf \_img_\_img S.setp_imp_subset_Univ S.ide_dom S.setp_set_ide + by blast hence "\ (restrict (S.Fun g o S.Fun f) (S.set (S.dom f)) (\ x')) = \ (S.Fun g (S.Fun f (\ x')))" using restrict_apply by auto also have "... = \ (S.Fun g (\ (\ (S.Fun f (\ x')))))" proof - have "S.Fun f (\ x') \ S.set (S.cod f)" using gf 1 S.Fun_mapsto by fast hence "\ (\ (S.Fun f (\ x'))) = S.Fun f (\ x')" - using assms bij_\ S.set_subset_Univ bij_betw_def inv_into_f_f subsetCE S.ide_cod + using assms bij_\ S.setp_imp_subset_Univ bij_betw_def inv_into_f_f subsetCE + S.ide_cod S.setp_set_ide by (metis S.seqE) thus ?thesis by auto qed also have "... = \a g (\a f x')" proof - have "\a f x' \ \ ` S.set (S.cod f)" using gf S.ide_dom S.ide_cod X' \a_mapsto [of f] set_\o [of "S.dom f"] set_\o [of "S.cod f"] by blast thus ?thesis using gf X' \a_def by auto qed finally show "\ (restrict (S.Fun g o S.Fun f) (S.set (S.dom f)) (\ x')) = \a g (\a f x')" by auto qed thus ?thesis using assms set_\o by fastforce qed finally show ?thesis by auto qed text\ Finally, we use @{term \o} and @{term \a} to define a functor @{term \}. \ definition \ where "\ f = (if S.arr f then S'.mkArr (S'.set (\o (S.dom f))) (S'.set (\o (S.cod f))) (\a f) else S'.null)" lemma \_in_hom: assumes "S.arr f" shows "\ f \ S'.hom (\o (S.dom f)) (\o (S.cod f))" proof - have "\\ f : S'.dom (\ f) \' S'.cod (\ f)\" - using assms \_def [of f] \a_mapsto [of f] \o_preserves_ide S'.set_card S'.arr_mkArr + using assms \_def \a_mapsto \o_preserves_ide by (intro S'.in_homI) auto thus ?thesis - using assms \_def \a_mapsto \o_preserves_ide S'.set_card S'.arr_mkArr S'.mkIde_set - by auto + using assms \_def \a_mapsto \o_preserves_ide by auto qed lemma \_ide [simp]: assumes "S.ide a" shows "\ a = \o a" proof - have "\ a = S'.mkArr (S'.set (\o a)) (S'.set (\o a)) (\x'. x')" proof - have "\\ a : \o a \' \o a\" using assms \_in_hom S.ide_in_hom by fastforce moreover have "\a a = restrict (\x'. x') (S'.set (\o a))" proof - have "\a a = (\x' \ \ ` S.set a. \ (S.Fun a (\ x')))" using assms \a_def restrict_apply by auto also have "... = (\x' \ S'.set (\o a). \ (\ x'))" proof - - have "S.Fun a = (\x \ S.set a. x)" using assms S.Fun_ide by simp + have "S.Fun a = (\x \ S.set a. x)" + using assms S.Fun_ide by auto moreover have "\x'. x' \ \ ` S.set a \ \ x' \ S.set a" - using assms bij_\ S.set_subset_Univ image_iff by (metis \_img_\_img) + using assms bij_\ S.setp_imp_subset_Univ image_iff S.setp_set_ide + by (metis \_img_\_img) ultimately show ?thesis using assms set_\o by auto qed also have "... = restrict (\x'. x') (S'.set (\o a))" - using assms S'.set_subset_Univ \o_preserves_ide \_\ + using assms S'.setp_imp_subset_Univ S'.setp_set_ide \o_preserves_ide \_\ by (meson restr_eqI subsetCE) ultimately show ?thesis by auto qed ultimately show ?thesis using assms \_def \o_preserves_ide S'.mkArr_restrict_eq' by (metis S'.arrI S.ide_char) qed thus ?thesis - using assms S'.mkIde_as_mkArr \o_preserves_ide \_in_hom S'.set_card S'.mkIde_set + using assms S'.mkIde_as_mkArr \o_preserves_ide \_in_hom S'.mkIde_set by simp qed lemma set_dom_\: assumes "S.arr f" shows "S'.set (S'.dom (\ f)) = \ ` (S.set (S.dom f))" using assms S.ide_dom \_in_hom \_ide set_\o by fastforce lemma \_comp: assumes "S.seq g f" shows "\ (g \ f) = \ g \\ \ f" proof - have "\ (g \ f) = S'.mkArr (S'.set (\o (S.dom f))) (S'.set (\o (S.cod g))) (\a (S g f))" using \_def assms by auto also have "... = S'.mkArr (S'.set (\o (S.dom f))) (S'.set (\o (S.cod g))) (restrict (\a g o \a f) (S'.set (\o (S.dom f))))" using assms \a_comp set_\o by force also have "... = S'.mkArr (S'.set (\o (S.dom f))) (S'.set (\o (S.cod g))) (\a g o \a f)" by (metis S'.mkArr_restrict_eq' \_in_hom assms calculation S'.in_homE mem_Collect_eq) also have "... = S' (S'.mkArr (S'.set (\o (S.dom g))) (S'.set (\o (S.cod g))) (\a g)) (S'.mkArr (S'.set (\o (S.dom f))) (S'.set (\o (S.cod f))) (\a f))" - by (metis S'.comp_mkArr S.seqE \_def \_in_hom assms S'.in_homE mem_Collect_eq) + proof - + have "S'.arr (S'.mkArr (S'.set (\o (S.dom f))) (S'.set (\o (S.cod f))) (\a f))" + using assms \a_mapsto set_\o S.ide_dom S.ide_cod \o_preserves_ide + S'.arr_mkArr S'.setp_imp_subset_Univ S'.setp_set_ide S.seqE + by metis + moreover have "S'.arr (S'.mkArr (S'.set (\o (S.dom g))) (S'.set (\o (S.cod g))) + (\a g))" + using assms \a_mapsto set_\o S.ide_dom S.ide_cod \o_preserves_ide S'.arr_mkArr + S'.setp_imp_subset_Univ S'.setp_set_ide S.seqE + by metis + ultimately show ?thesis using assms S'.comp_mkArr by auto + qed also have "... = \ g \\ \ f" using assms \_def by force finally show ?thesis by fast qed interpretation \: "functor" S S' \ apply unfold_locales using \_def apply simp using \_in_hom \_comp by auto lemma \_is_functor: shows "functor S S' \" .. lemma Fun_\: assumes "S.arr f" and "x \ S.set (S.dom f)" shows "S'.Fun (\ f) (\ x) = \a f (\ x)" using assms \_def \.preserves_arr set_\o by auto lemma \_acts_elementwise: assumes "S.ide a" shows "S'.set (\ a) = \ ` S.set a" proof have 0: "S'.set (\ a) = \ ` S.set a" using assms \_ide set_\o by simp have 1: "\x. x \ S.set a \ \ x = \ x" proof - fix x assume x: "x \ S.set a" - have 1: "S.terminal x" using assms x S.set_subset_Univ by blast + have 1: "S.terminal x" using assms x S.setp_imp_subset_Univ S.setp_set_ide by blast hence 2: "S'.terminal (\ x)" by (metis CollectD CollectI bij_\ bij_betw_def image_iff) have "\ x = \o x" using assms x 1 \_ide S.terminal_def by auto also have "... = \ x" proof - have "\o x = S'.mkIde (\ ` S.set x)" using assms 1 x \o_def S.terminal_def by auto moreover have "S'.mkIde (\ ` S.set x) = \ x" using assms x 1 2 S.terminal_char2 S'.terminal_char2 S'.mkIde_set bij_\ - by (metis image_empty image_insert) + by (metis (no_types, lifting) empty_is_image image_insert) ultimately show ?thesis by auto qed finally show "\ x = \ x" by auto qed show "S'.set (\ a) \ \ ` S.set a" using 0 1 by force show "\ ` S.set a \ S'.set (\ a)" using 0 1 by force qed lemma \_preserves_incl: assumes "S.incl m" shows "S'.incl (\ m)" proof - have 1: "S.arr m \ S.set (S.dom m) \ S.set (S.cod m) \ m = S.mkArr (S.set (S.dom m)) (S.set (S.cod m)) (\x. x)" using assms S.incl_def by blast have "S'.arr (\ m)" using 1 by auto moreover have 2: "S'.set (S'.dom (\ m)) \ S'.set (S'.cod (\ m))" - using 1 \.preserves_dom \.preserves_cod \_acts_elementwise - by (metis (full_types) S.ide_cod S.ide_dom image_mono) + using 1 \.preserves_dom \.preserves_cod \_acts_elementwise by auto moreover have "\ m = S'.mkArr (S'.set (S'.dom (\ m))) (S'.set (S'.cod (\ m))) (\x'. x')" proof - have "\ m = S'.mkArr (S'.set (\o (S.dom m))) (S'.set (\o (S.cod m))) (\a m)" using 1 \_def by simp also have "... = S'.mkArr (S'.set (S'.dom (\ m))) (S'.set (S'.cod (\ m))) (\a m)" using 1 \_ide by auto finally have 3: "\ m = S'.mkArr (S'.set (S'.dom (\ m))) (S'.set (S'.cod (\ m))) (\a m)" by auto also have "... = S'.mkArr (S'.set (S'.dom (\ m))) (S'.set (S'.cod (\ m))) (\x'. x')" proof - have 4: "S.Fun m = restrict (\x. x) (S.set (S.dom m))" using assms S.incl_def by (metis (full_types) S.Fun_mkArr) hence "\a m = restrict (\x'. x') (\ ` (S.set (S.dom m)))" proof - + have 5: "\x'. x' \ \ ` S.set (S.dom m) \ \ (\ x') = x'" + by (meson 1 S.ide_dom S.setp_imp_subset_Univ S.setp_set' f_inv_into_f + image_mono subset_eq) have "\a m = restrict (\x'. \ (S.Fun m (\ x'))) (\ ` S.set (S.dom m))" using \a_def by simp also have "... = restrict (\x'. x') (\ ` S.set (S.dom m))" proof - have "\x. x \ \ ` (S.set (S.dom m)) \ \ (S.Fun m (\ x)) = x" proof - fix x assume x: "x \ \ ` (S.set (S.dom m))" hence "\ x \ S.set (S.dom m)" - using 1 S.ide_dom S.set_subset_Univ \_img_\_img image_eqI by metis - moreover have "\x'. x' \ \ ` S.set (S.dom m) \ \ (\ x') = x'" - by (metis 1 S'.set_subset_Univ S.ide_dom \o_preserves_ide \_\ set_\o subsetD) - ultimately show "\ (S.Fun m (\ x)) = x" using 1 4 x by simp + using 1 S.ide_dom S.setp_imp_subset_Univ S.setp_set_ide \_img_\_img image_eqI + by metis + thus "\ (S.Fun m (\ x)) = x" using 1 4 5 x by simp qed thus ?thesis by auto qed finally show ?thesis by auto qed hence "\a m = restrict (\x'. x') (S'.set (S'.dom (\ m)))" using 1 set_dom_\ by auto thus ?thesis using 2 3 \S'.arr (\ m)\ S'.mkArr_restrict_eq S'.ide_cod S'.ide_dom S'.incl_def by (metis S'.arr_mkArr image_restrict_eq image_subset_iff_funcset) qed finally show ?thesis by auto qed ultimately show ?thesis using S'.incl_def by blast qed + lemma \_respects_sets: + assumes "A' \ S'.Univ" + shows "setp (\ ` A') \ setp' A'" + using assms \_respects_setp \_img_\_img bij_\ + by (metis \_img_\_img bij_betw_def image_mono order_refl) + text\ Interchange the role of @{term \} and @{term \} to obtain a functor \\\ from @{term[source=true] S'} to @{term[source=true] S}. \ - interpretation INV: two_set_categories_bij_betw_Univ S' S \ \ - apply unfold_locales by (simp add: bij_\ bij_betw_inv_into) + interpretation INV: two_set_categories_bij_betw_Univ S' setp' S setp \ + using \_respects_sets bij_\ bij_betw_inv_into + by unfold_locales auto abbreviation \o where "\o \ INV.\o" abbreviation \a where "\a \ INV.\a" abbreviation \ where "\ \ INV.\" interpretation \: "functor" S' S \ using INV.\_is_functor by auto text\ The functors @{term \} and @{term \} are inverses. \ lemma Fun_\: assumes "S'.arr f'" and "x' \ S'.set (S'.dom f')" shows "S.Fun (\ f') (\ x') = \a f' (\ x')" using assms INV.Fun_\ by blast lemma \o_\o: assumes "S.ide a" shows "\o (\o a) = a" using assms \o_def INV.\o_def \_img_\_img \o_preserves_ide set_\o S.mkIde_set - by force + by (simp add: S.setp_imp_subset_Univ) lemma \\: assumes "S.arr f" shows "\ (\ f) = f" proof (intro S.arr_eqI) show par: "S.par (\ (\ f)) f" using assms \o_preserves_ide \o_\o by auto show "S.Fun (\ (\ f)) = S.Fun f" proof - have "S.arr (\ (\ f))" using assms by auto moreover have "\ (\ f) = S.mkArr (S.set (S.dom f)) (S.set (S.cod f)) (\a (\ f))" using assms INV.\_def \_in_hom \o_\o by auto moreover have "\a (\ f) = (\x \ S.set (S.dom f). \ (S'.Fun (\ f) (\ x)))" proof - have "\a (\ f) = (\x \ \ ` S'.set (S'.dom (\ f)). \ (S'.Fun (\ f) (\ x)))" proof - have "\x. x \ \ ` S'.set (S'.dom (\ f)) \ INV.\ x = \ x" - using assms S.ide_dom S.set_subset_Univ \.preserves_reflects_arr par bij_\ + using assms S.ide_dom S.setp_imp_subset_Univ \.preserves_reflects_arr par bij_\ inv_into_inv_into_eq subsetCE INV.set_dom_\ - by metis + by (metis (no_types) S.setp_set') thus ?thesis using INV.\a_def by auto qed moreover have "\ ` S'.set (S'.dom (\ f)) = S.set (S.dom f)" using assms by (metis par \.preserves_reflects_arr INV.set_dom_\) ultimately show ?thesis by auto qed ultimately have 1: "S.Fun (\ (\ f)) = (\x \ S.set (S.dom f). \ (S'.Fun (\ f) (\ x)))" using S'.Fun_mkArr by simp show ?thesis proof fix x have "x \ S.set (S.dom f) \ S.Fun (\ (\ f)) x = S.Fun f x" using 1 assms extensional_def S.Fun_mapsto S.Fun_def by auto moreover have "x \ S.set (S.dom f) \ S.Fun (\ (\ f)) x = S.Fun f x" proof - assume x: "x \ S.set (S.dom f)" have "S.Fun (\ (\ f)) x = \ (\ (S.Fun f (\ (\ x))))" using assms x 1 Fun_\ bij_\ \a_def by auto also have "... = S.Fun f x" proof - have 2: "\x. x \ S.Univ \ \ (\ x) = x" using bij_\ bij_betw_inv_into_left by fast have "S.Fun f (\ (\ x)) = S.Fun f x" - using assms x 2 - by (metis S.ide_dom S.set_subset_Univ subsetCE) + using assms x 2 S.ide_dom S.setp_imp_subset_Univ + by (metis S.setp_set' subsetD) moreover have "S.Fun f x \ S.Univ" - using x assms S.Fun_mapsto S.set_subset_Univ S.ide_cod by blast + using x assms S.Fun_mapsto S.setp_imp_subset_Univ S.setp_set' S.ide_cod + by blast ultimately show ?thesis using 2 by auto qed finally show ?thesis by auto qed ultimately show "S.Fun (\ (\ f)) x = S.Fun f x" by auto qed qed qed lemma \o_\o: assumes "S'.ide a'" shows "\o (\o a') = a'" using assms \o_def INV.\o_def \_img_\_img INV.\o_preserves_ide \_\ INV.set_\o - S'.mkIde_set + S'.mkIde_set S'.setp_imp_subset_Univ by force lemma \\: assumes "S'.arr f'" shows "\ (\ f') = f'" proof (intro S'.arr_eqI) show par: "S'.par (\ (\ f')) f'" using assms \.preserves_ide \.preserves_ide \_ide INV.\_ide \o_\o by auto show "S'.Fun (\ (\ f')) = S'.Fun f'" proof - have "S'.arr (\ (\ f'))" using assms by blast moreover have "\ (\ f') = S'.mkArr (S'.set (S'.dom f')) (S'.set (S'.cod f')) (\a (\ f'))" using assms \_def INV.\_in_hom \o_\o by simp moreover have "\a (\ f') = (\x' \ S'.set (S'.dom f'). \ (S.Fun (\ f') (\ x')))" unfolding \a_def using assms par \.preserves_arr set_dom_\ by metis ultimately have 1: "S'.Fun (\ (\ f')) = (\x' \ S'.set (S'.dom f'). \ (S.Fun (\ f') (\ x')))" using S'.Fun_mkArr by simp show ?thesis proof fix x' have "x' \ S'.set (S'.dom f') \ S'.Fun (\ (\ f')) x' = S'.Fun f' x'" using 1 assms S'.Fun_mapsto extensional_def by (simp add: S'.Fun_def) moreover have "x' \ S'.set (S'.dom f') \ S'.Fun (\ (\ f')) x' = S'.Fun f' x'" proof - assume x': "x' \ S'.set (S'.dom f')" have "S'.Fun (\ (\ f')) x' = \ (S.Fun (\ f') (\ x'))" using x' 1 by auto also have "... = \ (\a f' (\ x'))" - using Fun_\ x' assms S'.set_subset_Univ bij_\ by metis + using Fun_\ x' assms S'.setp_imp_subset_Univ bij_\ by metis also have "... = \ (\ (S'.Fun f' (\ (\ x'))))" proof - have "\ (\a f' (\ x')) = \ (\ (S'.Fun f' x'))" proof - have "x' \ S'.Univ" - by (meson S'.ide_dom S'.set_subset_Univ assms subsetCE x') + by (meson S'.ide_dom S'.setp_imp_subset_Univ S'.setp_set_ide assms subsetCE x') thus ?thesis by (simp add: INV.\a_def INV.\_\ x') qed also have "... = \ (\ (S'.Fun f' (\ (\ x'))))" - using assms x' \_\ S'.set_subset_Univ S'.ide_dom by (metis subsetCE) + using assms x' \_\ S'.setp_imp_subset_Univ S'.setp_set_ide S'.ide_dom + by (metis subsetCE) finally show ?thesis by auto qed also have "... = S'.Fun f' x'" proof - have 2: "\x'. x' \ S'.Univ \ \ (\ x') = x'" using bij_\ bij_betw_inv_into_right by fast have "S'.Fun f' (\ (\ x')) = S'.Fun f' x'" - using assms x' 2 S'.set_subset_Univ S'.ide_dom by (metis subsetCE) + using assms x' 2 S'.setp_imp_subset_Univ S'.setp_set_ide S'.ide_dom + by (metis subsetCE) moreover have "S'.Fun f' x' \ S'.Univ" - using x' assms S'.Fun_mapsto S'.set_subset_Univ S'.ide_cod by blast + using x' assms S'.Fun_mapsto S'.setp_imp_subset_Univ S'.setp_set_ide S'.ide_cod + by blast ultimately show ?thesis using 2 by auto qed finally show ?thesis by auto qed ultimately show "S'.Fun (\ (\ f')) x' = S'.Fun f' x'" by auto qed qed qed lemma inverse_functors_\_\: shows "inverse_functors S S' \ \" proof - interpret \\: composite_functor S S' S \ \ .. have inv: "\ o \ = S.map" using \\ S.map_def \\.is_extensional by auto interpret \\: composite_functor S' S S' \ \ .. have inv': "\ o \ = S'.map" using \\ S'.map_def \\.is_extensional by auto show ?thesis using inv inv' by (unfold_locales, auto) qed lemma are_isomorphic: shows "\\. invertible_functor S S' \ \ (\m. S.incl m \ S'.incl (\ m))" proof - interpret inverse_functors S S' \ \ using inverse_functors_\_\ by auto have 1: "inverse_functors S S' \ \" .. interpret invertible_functor S S' \ apply unfold_locales using 1 by auto have "invertible_functor S S' \" .. thus ?thesis using \_preserves_incl by auto qed end text\ The main result: @{locale set_category} is categorical, in the following (logical) sense: - If \S\ and \S'\ are two ``set categories'' for the same cardinal \\\, - and if the sets of terminal objects of \S\ and \S'\ are in bijective correspondence, - then \S\ and \S'\ are isomorphic as categories, via a functor that preserves inclusion maps, - hence the inclusion relation between sets. - \ + If \S\ and \S'\ are two "set categories", and if the sets of terminal objects of \S\ and \S'\ + are in correspondence via a @{term setp}-preserving bijection, then \S\ and \S'\ are + isomorphic as categories, via a functor that preserves inclusion maps, hence also the + inclusion relation between sets. +\ theorem set_category_is_categorical: - assumes "set_category S \" and "set_category S' \" + assumes "set_category S setp" and "set_category S' setp'" and "bij_betw \ (set_category_data.Univ S) (set_category_data.Univ S')" + and "\A. A \ set_category_data.Univ S \ setp' (\ ` A) \ setp A" shows "\\. invertible_functor S S' \ \ - (\m. set_category.incl S \ m \ set_category.incl S' \ (\ m))" + (\m. set_category.incl S setp m \ set_category.incl S' setp' (\ m))" proof - - interpret S: set_category S using assms(1) by auto - interpret S': set_category S' using assms(2) by auto - interpret two_set_categories_bij_betw_Univ S S' \ \ - apply (unfold_locales) using assms(3) by auto + interpret S: set_category S setp using assms(1) by auto + interpret S': set_category S' setp' using assms(2) by auto + interpret two_set_categories_bij_betw_Univ S setp S' setp' \ + apply (unfold_locales) using assms(3-4) by auto show ?thesis using are_isomorphic by auto qed section "Further Properties of Set Categories" text\ In this section we further develop the consequences of the \set_category\ axioms, and establish characterizations of a number of standard category-theoretic notions for a \set_category\. \ context set_category begin abbreviation Dom where "Dom f \ set (dom f)" abbreviation Cod where "Cod f \ set (cod f)" subsection "Initial Object" text\ The object corresponding to the empty set is an initial object. \ definition empty where "empty = mkIde {}" lemma initial_empty: shows "initial empty" proof show 0: "ide empty" - using empty_def ide_mkIde - by (simp add: ide_mkIde_finite) + using empty_def by (simp add: setp_empty) show "\b. ide b \ \!f. \f : empty \ b\" proof - fix b assume b: "ide b" show "\!f. \f : empty \ b\" proof show 1: "\mkArr {} (set b) (\x. x) : empty \ b\" - using 0 b empty_def mkArr_in_hom mkIde_set set_subset_Univ arr_mkIde - by (metis (no_types, lifting) Pi_I empty_iff ide_def mkIde_def) + using 0 b empty_def mkArr_in_hom mkIde_set setp_imp_subset_Univ arr_mkIde + by (metis Pi_I empty_iff ide_def mkIde_def) show "\f. \f : empty \ b\ \ f = mkArr {} (set b) (\x. x)" by (metis "1" arr_mkArr empty_iff in_homE empty_def mkArr_Fun mkArr_eqI set_mkIde) qed qed qed subsection "Identity Arrows" text\ Identity arrows correspond to restrictions of the identity function. \ lemma ide_char: assumes "arr f" shows "ide f \ Dom f = Cod f \ Fun f = (\x \ Dom f. x)" using assms mkIde_as_mkArr mkArr_Fun Fun_ide in_homE ide_cod mkArr_Fun mkIde_set by (metis ide_char) lemma ideI: assumes "arr f" and "Dom f = Cod f" and "\x. x \ Dom f \ Fun f x = x" shows "ide f" proof - have "Fun f = (\x \ Dom f. x)" using assms Fun_def by auto thus ?thesis using assms ide_char by blast qed subsection "Inclusions" lemma ide_implies_incl: assumes "ide a" shows "incl a" - by (simp add: assms set_card incl_def mkIde_as_mkArr) + by (simp add: assms incl_def) definition incl_in :: "'s \ 's \ bool" where "incl_in a b = (ide a \ ide b \ set a \ set b)" abbreviation incl_of where "incl_of a b \ mkArr (set a) (set b) (\x. x)" lemma elem_set_implies_set_eq_singleton: assumes "a \ set b" shows "set a = {a}" proof - have "ide b" using assms set_def by auto - thus ?thesis using assms set_subset_Univ terminal_char2 - by (metis mem_Collect_eq subsetCE) + thus ?thesis using assms setp_imp_subset_Univ terminal_char2 + by (metis setp_set' insert_subset mem_Collect_eq mk_disjoint_insert) qed lemma elem_set_implies_incl_in: assumes "a \ set b" shows "incl_in a b" proof - have b: "ide b" using assms set_def by auto - hence "set b \ Univ" by simp + hence "setp (set b)" by simp hence "a \ Univ \ set a \ set b" - using assms elem_set_implies_set_eq_singleton by auto + using setp_imp_subset_Univ assms elem_set_implies_set_eq_singleton by auto hence "ide a \ set a \ set b" using b terminal_char1 by simp thus ?thesis using b incl_in_def by simp qed lemma incl_incl_of [simp]: assumes "incl_in a b" shows "incl (incl_of a b)" and "\incl_of a b : a \ b\" proof - show "\incl_of a b : a \ b\" - using assms incl_in_def mkArr_in_hom mkIde_set set_subset_Univ - by (metis image_ident image_subset_iff_funcset set_card) + using assms incl_in_def mkArr_in_hom + by (metis image_ident image_subset_iff_funcset mkIde_set setp_set_ide) thus "incl (incl_of a b)" using assms incl_def incl_in_def by fastforce qed text\ There is at most one inclusion between any pair of objects. \ lemma incls_coherent: assumes "par f f'" and "incl f" and "incl f'" shows "f = f'" using assms incl_def fun_complete by auto text\ The set of inclusions is closed under composition. \ lemma incl_comp [simp]: assumes "incl f" and "incl g" and "cod f = dom g" shows "incl (g \ f)" proof - have 1: "seq g f" using assms incl_def by auto moreover have 2: "Dom (g \ f) \ Cod (g \ f)" using assms 1 incl_def by auto moreover have "g \ f = mkArr (Dom f) (Cod g) (restrict (\x. x) (Dom f))" proof (intro arr_eqI) have 3: "arr (mkArr (Dom f) (Cod g) (\x\Dom f. x))" by (metis 1 2 cod_comp dom_comp ide_cod ide_dom incl_def incl_in_def incl_incl_of(1) mkArr_restrict_eq) show 4: "par (g \ f) (mkArr (Dom f) (Cod g) (\x\Dom f. x))" using assms 1 3 mkIde_set by auto show "Fun (g \ f) = Fun (mkArr (Dom f) (Cod g) (\x\Dom f. x))" using assms 3 4 Fun_comp Fun_mkArr - by (metis Fun_ide comp_cod_arr ide_cod mkArr_restrict_eq' incl_def) + by (metis comp_cod_arr dom_cod ide_cod ide_implies_incl incl_def mkArr_restrict_eq') qed - ultimately show ?thesis using incl_def arr_mkArr set_mkIde by force + ultimately show ?thesis using incl_def by force qed subsection "Image Factorization" text\ The image of an arrow is the object that corresponds to the set-theoretic image of the domain set under the function induced by the arrow. \ abbreviation Img where "Img f \ Fun f ` Dom f" definition img where "img f = mkIde (Img f)" lemma ide_img [simp]: assumes "arr f" shows "ide (img f)" proof - have "Fun f ` Dom f \ Cod f" using assms Fun_mapsto by blast - moreover have "Cod f \ Univ \ |Cod f| " - using assms by (simp add: set_card) - ultimately have "Fun f ` Dom f \ Univ \ |Fun f ` Dom f| " - by (meson assms card_of_image ide_dom ordLeq_ordLess_trans set_card subset_eq) - thus ?thesis using img_def ide_mkIde by simp + moreover have "setp (Cod f)" using assms by simp + ultimately show ?thesis using img_def setp_respects_subset by auto qed lemma set_img [simp]: assumes "arr f" shows "set (img f) = Img f" proof - - have "Fun f ` set (dom f) \ set (cod f) \ set (cod f) \ Univ" + have 1: "Img f \ Cod f \ setp (set (cod f))" using assms Fun_mapsto by auto - hence "Fun f ` set (dom f) \ Univ \ |Fun f ` Dom f| " - by (metis assms ide_def ide_img img_def mkIde_def) - thus ?thesis using assms img_def set_mkIde by auto + hence "Fun f ` set (dom f) \ Univ" + using setp_imp_subset_Univ by blast + thus ?thesis + using assms 1 img_def set_mkIde setp_respects_subset by auto qed lemma img_point_in_Univ: assumes "\x : unity \ a\" shows "img x \ Univ" proof - have "set (img x) = {Fun x unity}" using assms terminal_char2 terminal_unity by auto thus "img x \ Univ" using assms terminal_char1 by auto qed lemma incl_in_img_cod: assumes "arr f" shows "incl_in (img f) (cod f)" proof (unfold img_def) - have 1: "Img f \ Cod f \ Cod f \ Univ \ |Cod f| " - using assms Fun_mapsto - by (metis arr_mkArr image_subset_iff_funcset mkArr_Fun) + have 1: "Img f \ Cod f \ setp (Cod f)" + using assms Fun_mapsto by auto hence 2: "ide (mkIde (Img f))" - using assms ide_img img_def by auto + using setp_respects_subset by auto + moreover have "ide (cod f)" using assms by auto moreover have "set (mkIde (Img f)) \ Cod f" - using 1 2 - by (metis ideD(1) arr_mkIde set_mkIde) + using 1 2 using setp_respects_subset by force ultimately show "incl_in (mkIde (Img f)) (cod f)" using assms incl_in_def ide_cod by blast qed lemma img_point_elem_set: assumes "\x : unity \ a\" shows "img x \ set a" by (metis assms img_point_in_Univ in_homE incl_in_img_cod insert_subset mem_Collect_eq incl_in_def terminal_char2) text\ The corestriction of an arrow @{term f} is the arrow @{term "corestr f \ hom (dom f) (img f)"} that induces the same function on the universe as @{term f}. \ definition corestr where "corestr f = mkArr (Dom f) (Img f) (Fun f)" lemma corestr_in_hom: assumes "arr f" shows "\corestr f : dom f \ img f\" - proof - - have "Fun f \ Dom f \ Fun f ` Dom f \ Dom f \ Univ" - using assms by auto - moreover have "Fun f ` Dom f \ Univ" - by (metis assms ide_img set_img set_subset_Univ) - moreover have "|Fun f ` Dom f| " - using assms by (metis ide_img set_card set_img) - ultimately have "mkArr (Dom f) (Fun f ` Dom f) (Fun f) \ hom (dom f) (img f)" - using assms img_def mkArr_in_hom [of "Dom f" "Fun f ` Dom f" "Fun f"] mkIde_set - by (simp add: set_card) - thus ?thesis using corestr_def by fastforce - qed + by (metis assms corestr_def equalityD2 ide_dom ide_img image_subset_iff_funcset + mkIde_set set_img mkArr_in_hom setp_set_ide) text\ Every arrow factors as a corestriction followed by an inclusion. \ lemma img_fact: assumes "arr f" shows "S (incl_of (img f) (cod f)) (corestr f) = f" proof (intro arr_eqI) have 1: "\corestr f : dom f \ img f\" using assms corestr_in_hom by blast moreover have 2: "\incl_of (img f) (cod f) : img f \ cod f\" using assms incl_in_img_cod incl_incl_of by fast ultimately show P: "par (incl_of (img f) (cod f) \ corestr f) f" using assms in_homE by blast show "Fun (incl_of (img f) (cod f) \ corestr f) = Fun f" - proof - - have "Fun (incl_of (img f) (cod f) \ corestr f) - = restrict (Fun (incl_of (img f) (cod f)) o Fun (corestr f)) (Dom f)" - using Fun_comp 1 2 P by auto - also have "... = Fun f" - by (metis 1 2 Fun_comp Fun_ide P comp_cod_arr corestr_def ide_img in_homE - mkArr_Fun Fun_mkArr) - finally show ?thesis by auto - qed + by (metis (no_types, lifting) 1 2 Fun_comp Fun_ide Fun_mkArr P comp_cod_arr + corestr_def ide_img in_homE mkArr_Fun) qed lemma Fun_corestr: assumes "arr f" shows "Fun (corestr f) = Fun f" - proof - - have 1: "f = incl_of (img f) (cod f) \ corestr f" - using assms img_fact by auto - hence 2: "Fun f = restrict (Fun (incl_of (img f) (cod f)) o Fun (corestr f)) (Dom f)" - using assms by (metis Fun_comp dom_comp) - also have "... = restrict (Fun (corestr f)) (Dom f)" - using assms by (metis 1 2 Fun_mkArr seqE mkArr_Fun corestr_def) - also have "... = Fun (corestr f)" - using assms 1 by (metis Fun_def dom_comp extensional_restrict restrict_extensional) - finally show ?thesis by auto - qed + by (metis Fun_mkArr arrI assms corestr_def corestr_in_hom mkArr_Fun) subsection "Points and Terminal Objects" text\ To each element @{term t} of @{term "set a"} is associated a point @{term "mkPoint a t \ hom unity a"}. The function induced by such a point is the constant-@{term t} function on the set @{term "{unity}"}. \ definition mkPoint where "mkPoint a t \ mkArr {unity} (set a) (\_. t)" lemma mkPoint_in_hom: assumes "ide a" and "t \ set a" shows "\mkPoint a t : unity \ a\" - using assms mkArr_in_hom mkIde_set set_subset_Univ terminal_char2 terminal_unity - mkPoint_def set_card - by (metis Pi_I) + using assms mkArr_in_hom + by (metis Pi_I mkIde_set setp_set_ide terminal_char2 terminal_unity mkPoint_def) lemma Fun_mkPoint: assumes "ide a" and "t \ set a" shows "Fun (mkPoint a t) = (\_ \ {unity}. t)" - using assms mkPoint_def mkPoint_in_hom Fun_mkArr by force + using assms mkPoint_def terminal_unity mkPoint_in_hom by fastforce text\ For each object @{term a} the function @{term "mkPoint a"} has as its inverse the restriction of the function @{term img} to @{term "hom unity a"} \ lemma mkPoint_img: shows "img \ hom unity a \ set a" and "\x. \x : unity \ a\ \ mkPoint a (img x) = x" proof - show "img \ hom unity a \ set a" using img_point_elem_set by simp show "\x. \x : unity \ a\ \ mkPoint a (img x) = x" proof - fix x assume x: "\x : unity \ a\" show "mkPoint a (img x) = x" proof (intro arr_eqI) have 0: "img x \ set a" using x img_point_elem_set by metis hence 1: "mkPoint a (img x) \ hom unity a" using x mkPoint_in_hom by force thus 2: "par (mkPoint a (img x)) x" using x by fastforce have "Fun (mkPoint a (img x)) = (\_ \ {unity}. img x)" using 1 mkPoint_def by auto also have "... = Fun x" - proof - fix z - have "z \ unity \ (\_ \ {unity}. img x) z = Fun x z" - using x Fun_mapsto Fun_def restrict_apply singletonD terminal_char2 terminal_unity - by auto - moreover have "(\_ \ {unity}. img x) unity = Fun x unity" - by (metis (no_types, lifting) 0 image_empty image_insert in_homE insertI1 - restrict_apply elem_set_implies_set_eq_singleton set_img terminal_char2 - terminal_unity the_elem_eq x) - ultimately show "(\_ \ {unity}. img x) z = Fun x z" by auto - qed + by (metis 0 Fun_corestr calculation elem_set_implies_set_eq_singleton + ide_cod ide_unity in_homE mem_Collect_eq Fun_mkPoint corestr_in_hom + img_point_in_Univ mkPoint_in_hom singletonI terminalE x) finally show "Fun (mkPoint a (img x)) = Fun x" by auto qed qed qed lemma img_mkPoint: assumes "ide a" shows "mkPoint a \ set a \ hom unity a" and "\t. t \ set a \ img (mkPoint a t) = t" proof - show "mkPoint a \ set a \ hom unity a" using assms(1) mkPoint_in_hom by simp show "\t. t \ set a \ img (mkPoint a t) = t" proof - fix t assume t: "t \ set a" show "img (mkPoint a t) = t" proof - have 1: "arr (mkPoint a t)" using assms t mkPoint_in_hom by auto have "Fun (mkPoint a t) ` {unity} = {t}" using 1 mkPoint_def by simp thus ?thesis by (metis in_homE img_def mkIde_set mkPoint_in_hom elem_set_implies_incl_in elem_set_implies_set_eq_singleton incl_in_def t terminal_char2 terminal_unity) qed qed qed text\ For each object @{term a} the elements of @{term "hom unity a"} are therefore in bijective correspondence with @{term "set a"}. \ lemma bij_betw_points_and_set: assumes "ide a" shows "bij_betw img (hom unity a) (set a)" proof (intro bij_betwI) show "img \ hom unity a \ set a" using assms mkPoint_img by auto show "mkPoint a \ set a \ hom unity a" using assms img_mkPoint by auto show "\x. x \ hom unity a \ mkPoint a (img x) = x" using assms mkPoint_img by auto show "\t. t \ set a \ img (mkPoint a t) = t" using assms img_mkPoint by auto qed + lemma setp_img_points: + assumes "ide a" + shows "setp (img ` hom unity a)" + using assms + by (metis image_subset_iff_funcset mkPoint_img(1) setp_respects_subset setp_set_ide) + text\ The function on the universe induced by an arrow @{term f} agrees, under the bijection between @{term "hom unity (dom f)"} and @{term "Dom f"}, with the action of @{term f} by composition on @{term "hom unity (dom f)"}. \ lemma Fun_point: assumes "\x : unity \ a\" shows "Fun x = (\_ \ {unity}. img x)" using assms mkPoint_img img_mkPoint Fun_mkPoint [of a "img x"] img_point_elem_set by auto lemma comp_arr_mkPoint: assumes "arr f" and "t \ Dom f" shows "f \ mkPoint (dom f) t = mkPoint (cod f) (Fun f t)" proof (intro arr_eqI) have 0: "seq f (mkPoint (dom f) t)" using assms mkPoint_in_hom [of "dom f" t] by auto have 1: "\f \ mkPoint (dom f) t : unity \ cod f\" using assms mkPoint_in_hom [of "dom f" t] by auto show "par (f \ mkPoint (dom f) t) (mkPoint (cod f) (Fun f t))" proof - have "\mkPoint (cod f) (Fun f t) : unity \ cod f\" using assms Fun_mapsto mkPoint_in_hom [of "cod f" "Fun f t"] by auto thus ?thesis using 1 by fastforce qed show "Fun (f \ mkPoint (dom f) t) = Fun (mkPoint (cod f) (Fun f t))" proof - have "Fun (f \ mkPoint (dom f) t) = restrict (Fun f o Fun (mkPoint (dom f) t)) {unity}" using assms 0 1 Fun_comp terminal_char2 terminal_unity by auto also have "... = (\_ \ {unity}. Fun f t)" using assms Fun_mkPoint by auto also have "... = Fun (mkPoint (cod f) (Fun f t))" using assms Fun_mkPoint [of "cod f" "Fun f t"] Fun_mapsto by fastforce finally show ?thesis by auto qed qed lemma comp_arr_point: assumes "arr f" and "\x : unity \ dom f\" shows "f \ x = mkPoint (cod f) (Fun f (img x))" by (metis assms comp_arr_mkPoint img_point_elem_set mkPoint_img(2)) text\ This agreement allows us to express @{term "Fun f"} in terms of composition. \ lemma Fun_in_terms_of_comp: assumes "arr f" shows "Fun f = restrict (img o S f o mkPoint (dom f)) (Dom f)" proof fix t have "t \ Dom f \ Fun f t = restrict (img o S f o mkPoint (dom f)) (Dom f) t" using assms by (simp add: Fun_def) moreover have "t \ Dom f \ Fun f t = restrict (img o S f o mkPoint (dom f)) (Dom f) t" proof - assume t: "t \ Dom f" have 1: "f \ mkPoint (dom f) t = mkPoint (cod f) (Fun f t)" using assms t comp_arr_mkPoint by simp hence "img (f \ mkPoint (dom f) t) = img (mkPoint (cod f) (Fun f t))" by simp thus ?thesis proof - have "Fun f t \ Cod f" using assms t Fun_mapsto by auto thus ?thesis using assms t 1 img_mkPoint by auto qed qed ultimately show "Fun f t = restrict (img o S f o mkPoint (dom f)) (Dom f) t" by auto qed text\ We therefore obtain a rule for proving parallel arrows equal by showing that they have the same action by composition on points. \ (* * TODO: Find out why attempts to use this as the main rule for a proof loop * unless the specific instance is given. *) lemma arr_eqI': assumes "par f f'" and "\x. \x : unity \ dom f\ \ f \ x = f' \ x" shows "f = f'" using assms Fun_in_terms_of_comp mkPoint_in_hom by (intro arr_eqI, auto) text\ An arrow can therefore be specified by giving its action by composition on points. In many situations, this is more natural than specifying it as a function on the universe. \ definition mkArr' where "mkArr' a b F = mkArr (set a) (set b) (img o F o mkPoint a)" lemma mkArr'_in_hom: assumes "ide a" and "ide b" and "F \ hom unity a \ hom unity b" shows "\mkArr' a b F : a \ b\" proof - have "img o F o mkPoint a \ set a \ set b" - proof - fix t - assume t: "t \ set a" - thus "(img o F o mkPoint a) t \ set b" - using assms mkPoint_in_hom img_point_elem_set [of "F (mkPoint a t)" b] - by auto - qed + using assms(1,3) img_mkPoint(1) mkPoint_img(1) by fastforce thus ?thesis - using assms mkArr'_def mkArr_in_hom [of "set a" "set b"] set_card mkIde_set by simp + using assms mkArr'_def mkArr_in_hom [of "set a" "set b"] mkIde_set by simp qed lemma comp_point_mkArr': assumes "ide a" and "ide b" and "F \ hom unity a \ hom unity b" shows "\x. \x : unity \ a\ \ mkArr' a b F \ x = F x" proof - fix x assume x: "\x : unity \ a\" have "Fun (mkArr' a b F) (img x) = img (F x)" unfolding mkArr'_def - using assms x Fun_mkArr arr_mkArr img_point_elem_set mkPoint_img mkPoint_in_hom - by (simp add: set_card Pi_iff) + using assms x Fun_mkArr img_point_elem_set mkPoint_img mkPoint_in_hom + by (simp add: Pi_iff) hence "mkArr' a b F \ x = mkPoint b (img (F x))" using assms x mkArr'_in_hom [of a b F] comp_arr_point by auto thus "mkArr' a b F \ x = F x" using assms x mkPoint_img(2) by auto qed text\ A third characterization of terminal objects is as those objects whose set of points is a singleton. \ lemma terminal_char3: assumes "\!x. \x : unity \ a\" shows "terminal a" proof - have a: "ide a" using assms ide_cod mem_Collect_eq by blast hence "bij_betw img (hom unity a) (set a)" using assms bij_betw_points_and_set by auto hence "img ` (hom unity a) = set a" by (simp add: bij_betw_def) moreover have "hom unity a = {THE x. x \ hom unity a}" using assms theI' [of "\x. x \ hom unity a"] by auto ultimately have "set a = {img (THE x. x \ hom unity a)}" by (metis image_empty image_insert) thus ?thesis using a terminal_char1 by simp qed text\ The following is an alternative formulation of functional completeness, which says that any function on points uniquely determines an arrow. \ lemma fun_complete': assumes "ide a" and "ide b" and "F \ hom unity a \ hom unity b" shows "\!f. \f : a \ b\ \ (\x. \x : unity \ a\ \ f \ x = F x)" proof have 1: "\mkArr' a b F : a \ b\" using assms mkArr'_in_hom by auto moreover have 2: "\x. \x : unity \ a\ \ mkArr' a b F \ x = F x" using assms comp_point_mkArr' by auto ultimately show "\mkArr' a b F : a \ b\ \ (\x. \x : unity \ a\ \ mkArr' a b F \ x = F x)" by blast fix f assume f: "\f : a \ b\ \ (\x. \x : unity \ a\ \ f \ x = F x)" show "f = mkArr' a b F" using f 1 2 by (intro arr_eqI' [of f "mkArr' a b F"], fastforce, auto) qed subsection "The `Determines Same Function' Relation on Arrows" text\ An important part of understanding the structure of a category of sets and functions is to characterize when it is that two arrows ``determine the same function''. The following result provides one answer to this: two arrows with a common domain determine the same function if and only if they can be rendered equal by composing with a cospan of inclusions. \ lemma eq_Fun_iff_incl_joinable: assumes "span f f'" shows "Fun f = Fun f' \ (\m m'. incl m \ incl m' \ seq m f \ seq m' f' \ m \ f = m' \ f')" proof assume ff': "Fun f = Fun f'" let ?b = "mkIde (Cod f \ Cod f')" let ?m = "incl_of (cod f) ?b" let ?m' = "incl_of (cod f') ?b" - have 1: "|Cod f \ Cod f'| " - using assms set_card cardinal card_order_infinite_isLimOrd by simp have incl_m: "incl ?m" - using assms 1 incl_incl_of [of "cod f" ?b] incl_in_def ide_mkIde set_mkIde by simp + using assms incl_incl_of [of "cod f" ?b] incl_in_def setp_respects_union by simp have incl_m': "incl ?m'" - using assms 1 incl_incl_of [of "cod f'" ?b] incl_in_def ide_mkIde set_mkIde by simp + using assms incl_incl_of [of "cod f'" ?b] incl_in_def setp_respects_union by simp have m: "?m = mkArr (Cod f) (Cod f \ Cod f') (\x. x)" - by (simp add: assms 1 set_mkIde) + using setp_respects_union by (simp add: assms) have m': "?m' = mkArr (Cod f') (Cod f \ Cod f') (\x. x)" - by (simp add: assms 1 set_mkIde) + using setp_respects_union by (simp add: assms) have seq: "seq ?m f \ seq ?m' f'" - using assms m m' - by (metis ide_cod incl_m incl_m' mkIde_set seqI incl_def dom_mkArr) + using assms m m' using setp_respects_union by simp have "?m \ f = ?m' \ f'" - by (metis assms comp_mkArr ff' incl_m incl_m' mkArr_Fun incl_def) + by (metis assms comp_mkArr ff' incl_def incl_m incl_m' mkArr_Fun) hence "incl ?m \ incl ?m' \ seq ?m f \ seq ?m' f' \ ?m \ f = ?m' \ f'" using seq \incl ?m\ \incl ?m'\ by simp thus "\m m'. incl m \ incl m' \ seq m f \ seq m' f' \ m \ f = m' \ f'" by auto next assume ff': "\m m'. incl m \ incl m' \ seq m f \ seq m' f' \ m \ f = m' \ f'" show "Fun f = Fun f'" using ff' by (metis Fun_comp Fun_ide comp_cod_arr ide_cod seqE Fun_incl) qed text\ Another answer to the same question: two arrows with a common domain determine the same function if and only if their corestrictions are equal. \ lemma eq_Fun_iff_eq_corestr: assumes "span f f'" shows "Fun f = Fun f' \ corestr f = corestr f'" using assms corestr_def Fun_corestr by metis subsection "Retractions, Sections, and Isomorphisms" text\ An arrow is a retraction if and only if its image coincides with its codomain. \ lemma retraction_if_Img_eq_Cod: assumes "arr g" and "Img g = Cod g" shows "retraction g" and "ide (g \ mkArr (Cod g) (Dom g) (inv_into (Dom g) (Fun g)))" proof - let ?F = "inv_into (Dom g) (Fun g)" let ?f = "mkArr (Cod g) (Dom g) ?F" have f: "arr ?f" - by (simp add: assms inv_into_into set_card arr_mkArr) + by (simp add: assms inv_into_into) show "ide (g \ ?f)" proof - have "g = mkArr (Dom g) (Cod g) (Fun g)" using assms mkArr_Fun by auto hence "g \ ?f = mkArr (Cod g) (Cod g) (Fun g o ?F)" using assms(1) f comp_mkArr by metis moreover have "mkArr (Cod g) (Cod g) (\y. y) = ..." proof (intro mkArr_eqI') show "arr (mkArr (Cod g) (Cod g) (\y. y))" - using assms arr_cod_iff_arr set_card arr_mkArr by auto + using assms arr_cod_iff_arr by auto show "\y. y \ Cod g \ y = (Fun g o ?F) y" using assms by (simp add: f_inv_into_f) qed ultimately show ?thesis - using assms f mkIde_as_mkArr arr_mkArr mkIde_set by auto + using assms f mkIde_as_mkArr by auto qed thus "retraction g" by auto qed lemma retraction_char: shows "retraction g \ arr g \ Img g = Cod g" proof assume G: "retraction g" show "arr g \ Img g = Cod g" proof show "arr g" using G by blast show "Img g = Cod g" proof - from G obtain f where f: "ide (g \ f)" by blast have "Fun g ` Fun f ` Cod g = Cod g" proof - have "restrict (Fun g o Fun f) (Cod g) = restrict (\x. x) (Cod g)" using f Fun_comp Fun_ide ide_compE by metis thus ?thesis by (metis image_comp image_ident image_restrict_eq) qed moreover have "Fun f ` Cod g \ Dom g" using f Fun_mapsto arr_mkArr mkArr_Fun funcset_image by (metis seqE ide_compE ide_compE) moreover have "Img g \ Cod g" using f Fun_mapsto by blast ultimately show ?thesis by blast qed qed next assume "arr g \ Img g = Cod g" thus "retraction g" using retraction_if_Img_eq_Cod by blast qed text\ Every corestriction is a retraction. \ lemma retraction_corestr: assumes "arr f" shows "retraction (corestr f)" using assms retraction_char Fun_corestr corestr_in_hom in_homE set_img by metis text\ An arrow is a section if and only if it induces an injective function on its domain, except in the special case that it has an empty domain set and a nonempty codomain set. \ lemma section_if_inj: assumes "arr f" and "inj_on (Fun f) (Dom f)" and "Dom f = {} \ Cod f = {}" shows "section f" and "ide (mkArr (Cod f) (Dom f) (\y. if y \ Img f then SOME x. x \ Dom f \ Fun f x = y else SOME x. x \ Dom f) \ f)" proof - let ?P= "\y. \x. x \ Dom f \ Fun f x = y" let ?G = "\y. if y \ Img f then SOME x. ?P y x else SOME x. x \ Dom f" let ?g = "mkArr (Cod f) (Dom f) ?G" have g: "arr ?g" proof - - have 1: "Cod f \ Univ" using assms by simp - have 2: "Dom f \ Univ" using assms by simp - have 3: "|Cod f| \ |Dom f| " - using assms set_card by simp - have 4: "?G \ Cod f \ Dom f" + have 1: "setp (Cod f)" using assms by simp + have 2: "setp (Dom f)" using assms by simp + have 3: "?G \ Cod f \ Dom f" proof fix y assume Y: "y \ Cod f" show "?G y \ Dom f" proof (cases "y \ Img f") assume "y \ Img f" hence "(\x. ?P y x) \ ?G y = (SOME x. ?P y x)" using Y by auto hence "?P y (?G y)" using someI_ex [of "?P y"] by argo thus "?G y \ Dom f" by auto next assume "y \ Img f" hence "(\x. x \ Dom f) \ ?G y = (SOME x. x \ Dom f)" using assms Y by auto thus "?G y \ Dom f" using someI_ex [of "\x. x \ Dom f"] by argo qed qed - show ?thesis using 1 2 3 4 arr_mkArr by simp + show ?thesis using 1 2 3 by simp qed show "ide (?g \ f)" proof - have "f = mkArr (Dom f) (Cod f) (Fun f)" using assms mkArr_Fun by auto hence "?g \ f = mkArr (Dom f) (Dom f) (?G o Fun f)" using assms(1) g comp_mkArr [of "Dom f" "Cod f" "Fun f" "Dom f" ?G] by argo moreover have "mkArr (Dom f) (Dom f) (\x. x) = ..." proof (intro mkArr_eqI') show "arr (mkArr (Dom f) (Dom f) (\x. x))" - using assms set_card arr_mkArr by auto + using assms by auto show "\x. x \ Dom f \ x = (?G o Fun f) x" proof - fix x assume x: "x \ Dom f" have "Fun f x \ Img f" using x by blast hence *: "(\x'. ?P (Fun f x) x') \ ?G (Fun f x) = (SOME x'. ?P (Fun f x) x')" by auto then have "?P (Fun f x) (?G (Fun f x))" using someI_ex [of "?P (Fun f x)"] by argo with * have "x = ?G (Fun f x)" using assms x inj_on_def [of "Fun f" "Dom f"] by simp thus "x = (?G o Fun f) x" by simp qed qed ultimately show ?thesis - using assms set_card mkIde_as_mkArr mkIde_set by auto + using assms mkIde_as_mkArr mkIde_set by auto qed thus "section f" by auto qed lemma section_char: shows "section f \ arr f \ (Dom f = {} \ Cod f = {}) \ inj_on (Fun f) (Dom f)" proof assume f: "section f" from f obtain g where g: "ide (g \ f)" using section_def by blast show "arr f \ (Dom f = {} \ Cod f = {}) \ inj_on (Fun f) (Dom f)" proof - have "arr f" using f by blast moreover have "Dom f = {} \ Cod f = {}" proof - have "Cod f \ {} \ Dom f \ {}" proof assume "Cod f \ {}" from this obtain y where "y \ Cod f" by blast hence "Fun g y \ Dom f" using g Fun_mapsto by (metis seqE ide_compE image_eqI retractionI retraction_char) thus "Dom f \ {}" by blast qed thus ?thesis by auto qed moreover have "inj_on (Fun f) (Dom f)" by (metis Fun_comp Fun_ide g ide_compE inj_on_id2 inj_on_imageI2 inj_on_restrict_eq) ultimately show ?thesis by auto qed next assume F: "arr f \ (Dom f = {} \ Cod f = {}) \ inj_on (Fun f) (Dom f)" thus "section f" using section_if_inj by auto qed text\ Section-retraction pairs can also be characterized by an inverse relationship between the functions they induce. \ lemma section_retraction_char: shows "ide (g \ f) \ antipar f g \ compose (Dom f) (Fun g) (Fun f) = (\x \ Dom f. x)" by (metis Fun_comp cod_comp compose_eq' dom_comp ide_char ide_compE seqE) text\ Antiparallel arrows @{term f} and @{term g} are inverses if the functions they induce are inverses. \ lemma inverse_arrows_char: shows "inverse_arrows f g \ antipar f g \ compose (Dom f) (Fun g) (Fun f) = (\x \ Dom f. x) \ compose (Dom g) (Fun f) (Fun g) = (\y \ Dom g. y)" using section_retraction_char by blast text\ An arrow is an isomorphism if and only if the function it induces is a bijection. \ lemma iso_char: shows "iso f \ arr f \ bij_betw (Fun f) (Dom f) (Cod f)" by (metis bij_betw_def image_empty iso_iff_section_and_retraction retraction_char section_char) text\ The inverse of an isomorphism is constructed by inverting the induced function. \ lemma inv_char: assumes "iso f" shows "inv f = mkArr (Cod f) (Dom f) (inv_into (Dom f) (Fun f))" proof - let ?g = "mkArr (Cod f) (Dom f) (inv_into (Dom f) (Fun f))" have "ide (f \ ?g)" using assms iso_is_retraction retraction_char retraction_if_Img_eq_Cod by simp moreover have "ide (?g \ f)" proof - let ?g' = "mkArr (Cod f) (Dom f) (\y. if y \ Img f then SOME x. x \ Dom f \ Fun f x = y else SOME x. x \ Dom f)" have 1: "ide (?g' \ f)" using assms iso_is_section section_char section_if_inj by simp moreover have "?g' = ?g" proof show "arr ?g'" using 1 ide_compE by blast show "\y. y \ Cod f \ (if y \ Img f then SOME x. x \ Dom f \ Fun f x = y else SOME x. x \ Dom f) = inv_into (Dom f) (Fun f) y" proof - fix y assume "y \ Cod f" hence "y \ Img f" using assms iso_is_retraction retraction_char by metis thus "(if y \ Img f then SOME x. x \ Dom f \ Fun f x = y else SOME x. x \ Dom f) = inv_into (Dom f) (Fun f) y" using inv_into_def by metis qed qed ultimately show ?thesis by auto qed ultimately have "inverse_arrows f ?g" by auto thus ?thesis using inverse_unique by blast qed lemma Fun_inv: assumes "iso f" shows "Fun (inv f) = restrict (inv_into (Dom f) (Fun f)) (Cod f)" using assms inv_in_hom inv_char iso_inv_iso iso_is_arr Fun_mkArr by metis subsection "Monomorphisms and Epimorphisms" text\ An arrow is a monomorphism if and only if the function it induces is injective. \ lemma mono_char: shows "mono f \ arr f \ inj_on (Fun f) (Dom f)" proof assume f: "mono f" hence "arr f" using mono_def by auto moreover have "inj_on (Fun f) (Dom f)" proof (intro inj_onI) have 0: "inj_on (S f) (hom unity (dom f))" proof - have "hom unity (dom f) \ {g. seq f g}" using f mono_def arrI by auto hence "\A. hom unity (dom f) \ A \ inj_on (S f) A" using f mono_def by auto thus ?thesis by (meson subset_inj_on) qed fix x x' assume x: "x \ Dom f" and x': "x' \ Dom f" and xx': "Fun f x = Fun f x'" have "mkPoint (dom f) x \ hom unity (dom f) \ mkPoint (dom f) x' \ hom unity (dom f)" using x x' \arr f\ mkPoint_in_hom by simp moreover have "f \ mkPoint (dom f) x = f \ mkPoint (dom f) x'" using \arr f\ x x' xx' comp_arr_mkPoint by simp ultimately have "mkPoint (dom f) x = mkPoint (dom f) x'" using 0 inj_onD [of "S f" "hom unity (dom f)" "mkPoint (dom f) x"] by simp thus "x = x'" using \arr f\ x x' img_mkPoint(2) img_mkPoint(2) ide_dom by metis qed ultimately show "arr f \ inj_on (Fun f) (Dom f)" by auto next assume f: "arr f \ inj_on (Fun f) (Dom f)" show "mono f" proof show "arr f" using f by auto show "\g g'. seq f g \ seq f g' \ f \ g = f \ g' \ g = g'" proof - fix g g' assume gg': "seq f g \ seq f g' \ f \ g = f \ g'" show "g = g'" proof (intro arr_eqI) show par: "par g g'" using gg' dom_comp by (metis seqE) show "Fun g = Fun g'" proof fix x have "x \ Dom g \ Fun g x = Fun g' x" using gg' by (simp add: par Fun_def) moreover have "x \ Dom g \ Fun g x = Fun g' x" proof - assume x: "x \ Dom g" have "Fun f (Fun g x) = Fun (f \ g) x" using gg' x Fun_comp [of f g] by auto also have "... = Fun f (Fun g' x)" using par f gg' x monoE by simp finally have "Fun f (Fun g x) = Fun f (Fun g' x)" by auto moreover have "Fun g x \ Dom f \ Fun g' x \ Dom f" using par gg' x Fun_mapsto by fastforce ultimately show "Fun g x = Fun g' x" using f gg' inj_onD [of "Fun f" "Dom f" "Fun g x" "Fun g' x"] by simp qed ultimately show "Fun g x = Fun g' x" by auto qed qed qed qed qed text\ Inclusions are monomorphisms. \ lemma mono_imp_incl: assumes "incl f" shows "mono f" using assms incl_def Fun_incl mono_char by auto text\ A monomorphism is a section, except in case it has an empty domain set and a nonempty codomain set. \ lemma mono_imp_section: assumes "mono f" and "Dom f = {} \ Cod f = {}" shows "section f" using assms mono_char section_char by auto text\ An arrow is an epimorphism if and only if either its image coincides with its codomain, or else the universe has only a single element (in which case all arrows are epimorphisms). \ lemma epi_char: shows "epi f \ arr f \ (Img f = Cod f \ (\t t'. t \ Univ \ t' \ Univ \ t = t'))" proof assume epi: "epi f" show "arr f \ (Img f = Cod f \ (\t t'. t \ Univ \ t' \ Univ \ t = t'))" proof - have f: "arr f" using epi epi_implies_arr by auto moreover have "\(\t t'. t \ Univ \ t' \ Univ \ t = t') \ Img f = Cod f" proof - assume "\(\t t'. t \ Univ \ t' \ Univ \ t = t')" from this obtain tt and ff where B: "tt \ Univ \ ff \ Univ \ tt \ ff" by blast show "Img f = Cod f" proof show "Img f \ Cod f" using f Fun_mapsto by auto show "Cod f \ Img f" proof let ?g = "mkArr (Cod f) {ff, tt} (\y. tt)" let ?g' = "mkArr (Cod f) {ff, tt} (\y. if \x. x \ Dom f \ Fun f x = y then tt else ff)" let ?b = "mkIde {ff, tt}" have b: "ide ?b" - using B ide_mkIde_finite by simp + by (metis B finite.emptyI finite_imp_setp finite_insert ide_mkIde + insert_subset setp_imp_subset_Univ setp_singleton mem_Collect_eq) have g: "\?g : cod f \ ?b\ \ Fun ?g = (\y \ Cod f. tt)" - proof - - have "arr ?g" - proof - - have "arr (mkIde {ff, tt})" - using b ideD(1) by presburger - thus ?thesis - by (simp add: f set_card arr_mkIde arr_mkArr) - qed - thus ?thesis - using f b B in_homI [of ?g] mkIde_set by simp - qed + using f B in_homI [of ?g "cod f" "mkIde {ff, tt}"] finite_imp_setp by simp have g': "?g' \ hom (cod f) ?b \ Fun ?g' = (\y \ Cod f. if \x. x \ Dom f \ Fun f x = y then tt else ff)" - proof - - have "arr ?g" - proof - - have "arr (mkIde {ff, tt})" - using b ideD(1) by presburger - thus ?thesis - by (simp add: f set_card arr_mkIde arr_mkArr) - qed - thus ?thesis - using f b B in_homI [of ?g'] arr_mkArr mkIde_set by simp - qed + using f B in_homI [of ?g'] finite_imp_setp by simp have "?g \ f = ?g' \ f" proof (intro arr_eqI) show "par (?g \ f) (?g' \ f)" - using f g g' arr_mkArr by auto + using f g g' by auto show "Fun (?g \ f) = Fun (?g' \ f)" using f g g' Fun_comp comp_mkArr by fastforce qed hence gg': "?g = ?g'" by (metis (no_types, lifting) epiE epi f g in_homE seqI) fix y assume y: "y \ Cod f" have "Fun ?g' y = tt" using gg' g y by simp hence "(if \x. x \ Dom f \ Fun f x = y then tt else ff) = tt" using g' y by simp hence "\x. x \ Dom f \ Fun f x = y" using B by argo thus "y \ Img f" by blast qed qed qed ultimately show "arr f \ (Img f = Cod f \ (\t t'. t \ Univ \ t' \ Univ \ t = t'))" by fast qed next show "arr f \ (Img f = Cod f \ (\t t'. t \ Univ \ t' \ Univ \ t = t')) \ epi f" proof - have "arr f \ Img f = Cod f \ epi f" using retraction_char retraction_is_epi by presburger moreover have "arr f \ (\t t'. t \ Univ \ t' \ Univ \ t = t') \ epi f" proof - assume f: "arr f \ (\t t'. t \ Univ \ t' \ Univ \ t = t')" have "\f f'. par f f' \ f = f'" proof - fix f f' assume ff': "par f f'" show "f = f'" proof (intro arr_eqI) show "par f f'" using ff' by simp have "\t t'. t \ Cod f \ t' \ Cod f \ t = t'" - using f ff' set_subset_Univ ide_cod subsetD by blast + using f ff' setp_imp_subset_Univ setp_set_ide ide_cod subsetD by blast thus "Fun f = Fun f'" using ff' Fun_mapsto [of f] Fun_mapsto [of f'] extensional_arb [of "Fun f" "Dom f"] extensional_arb [of "Fun f'" "Dom f"] by fastforce qed qed moreover have "\g g'. par (g \ f) (g' \ f) \ par g g'" by force ultimately show "epi f" using f by (intro epiI; metis) qed ultimately show "arr f \ (Img f = Cod f \ (\t t'. t \ Univ \ t' \ Univ \ t = t')) \ epi f" by auto qed qed text\ An epimorphism is a retraction, except in the case of a degenerate universe with only a single element. \ lemma epi_imp_retraction: assumes "epi f" and "\t t'. t \ Univ \ t' \ Univ \ t \ t'" shows "retraction f" using assms epi_char retraction_char by auto text\ Retraction/inclusion factorization is unique (not just up to isomorphism -- remember that the notion of inclusion is not categorical but depends on the arbitrarily chosen @{term img}). \ lemma unique_retr_incl_fact: assumes "seq m e" and "seq m' e'" and "m \ e = m' \ e'" and "incl m" and "incl m'" and "retraction e" and "retraction e'" shows "m = m'" and "e = e'" proof - have 1: "cod m = cod m' \ dom e = dom e'" using assms(1-3) by (metis dom_comp cod_comp) hence 2: "span e e'" using assms(1-2) by blast hence 3: "Fun e = Fun e'" using assms eq_Fun_iff_incl_joinable by meson hence "img e = img e'" using assms 1 img_def by auto moreover have "img e = cod e \ img e' = cod e'" using assms(6-7) retraction_char img_def mkIde_set by simp ultimately have "par e e'" using 2 by simp thus "e = e'" using 3 arr_eqI by blast hence "par m m'" using assms(1) assms(2) 1 by fastforce thus "m = m'" using assms(4) assms(5) incls_coherent by blast qed end section "Concrete Set Categories" text\ The \set_category\ locale is useful for stating results that depend on a category of @{typ 'a}-sets and functions, without having to commit to a particular element type @{typ 'a}. However, in applications we often need to work with a category of sets and functions that is guaranteed to contain sets corresponding to the subsets of some extrinsically given type @{typ 'a}. A \emph{concrete set category} is a set category \S\ that is equipped with an injective function @{term \} from type @{typ 'a} to \S.Univ\. The following locale serves to facilitate some of the technical aspects of passing back and forth between elements of type @{typ 'a} and the elements of \S.Univ\. \ - locale concrete_set_category = set_category S \ + locale concrete_set_category = set_category S setp for S :: "'s comp" (infixr "\\<^sub>S" 55) - and \ :: "'t rel" + and setp :: "'s set \ bool" and U :: "'a set" and \ :: "'a \ 's" + - assumes \_mapsto: "\ \ U \ Univ" - and inj_\: "inj_on \ U" + assumes UP_mapsto: "\ \ U \ Univ" + and inj_UP: "inj_on \ U" begin - abbreviation \ - where "\ \ inv_into U \" + abbreviation UP + where "UP \ \" - lemma \_mapsto: - shows "\ \ \ ` U \ U" + abbreviation DN + where "DN \ inv_into U UP" + + lemma DN_mapsto: + shows "DN \ UP ` U \ U" by (simp add: inv_into_into) - lemma \_\ [simp]: + lemma DN_UP [simp]: assumes "x \ U" - shows "\ (\ x) = x" - using assms inj_\ inv_into_f_f by simp + shows "DN (UP x) = x" + using assms inj_UP inv_into_f_f by simp - lemma \_\ [simp]: - assumes "t \ \ ` U" - shows "\ (\ t) = t" - using assms o_def inj_\ by auto + lemma UP_DN [simp]: + assumes "t \ UP ` U" + shows "UP (DN t) = t" + using assms o_def inj_UP by auto + + lemma bij_UP: + shows "bij_betw UP U (UP ` U)" + using inj_UP inj_on_imp_bij_betw by blast + + lemma bij_DN: + shows "bij_betw DN (UP ` U) U" + using bij_UP bij_betw_inv_into by blast end locale replete_concrete_set_category = replete_set_category S + - concrete_set_category S \cardSuc (cmax (card_of (UNIV :: 's set)) natLeq)\ U \ + concrete_set_category S \\A. A \ Univ\ U UP for S :: "'s comp" (infixr "\\<^sub>S" 55) and U :: "'a set" - and \ :: "'a \ 's" + and UP :: "'a \ 's" + + section "Sub-Set Categories" + + text\ + In this section, we show that a full subcategory of a set category, obtained + by imposing suitable further restrictions on the subsets of the universe + that correspond to objects, is again a set category. +\ + + locale sub_set_category = + S: set_category + + fixes ssetp :: "'a set \ bool" + assumes ssetp_singleton: "\t. t \ S.Univ \ ssetp {t}" + and subset_closed: "\B A. \B \ A; ssetp A\ \ ssetp B" + and union_closed: "\A B. \ssetp A; ssetp B\ \ ssetp (A \ B)" + and containment: "\A. ssetp A \ setp A" + begin + + sublocale full_subcategory S \\a. S.ide a \ ssetp (S.set a)\ + by unfold_locales auto + + lemma is_full_subcategory: + shows "full_subcategory S (\a. S.ide a \ ssetp (S.set a))" + .. + + lemma ide_char: + shows "ide a \ S.ide a \ ssetp (S.set a)" + using ide_char arr_char by fastforce + + lemma terminal_unity: + shows "terminal S.unity" + proof + show "ide S.unity" + using S.terminal_unity S.terminal_def [of S.unity] S.terminal_char2 ide_char + ssetp_singleton + by force + thus "\a. ide a \ \!f. in_hom f a S.unity" + using S.terminal_unity S.terminal_def ide_char ide_char' in_hom_char + by (metis (no_types, lifting)) + qed + + lemma terminal_char: + shows "terminal t \ S.terminal t" + proof + fix t + assume t: "S.terminal t" + have "ide t" + using t ssetp_singleton ide_char S.terminal_char2 by force + thus "terminal t" + using t in_hom_char ide_char arr_char S.terminal_def terminalI by auto + next + assume t: "terminal t" + have 1: "S.ide t" + using t ide_char terminal_def by simp + moreover have "card (S.set t) = 1" + proof - + have "card (S.set t) = card (S.hom S.unity t)" + using S.set_def S.inj_img + by (metis 1 S.bij_betw_points_and_set bij_betw_same_card) + also have "... = card (hom S.unity t)" + using t in_hom_char terminal_def terminal_unity by auto + also have "... = 1" + proof - + have "\!f. f \ hom S.unity t" + using t terminal_def terminal_unity by force + moreover have "\A. card A = 1 \ (\!a. a \ A)" + apply (intro iffI) + apply (metis card_1_singletonE empty_iff insert_iff) + using card_1_singleton_iff by auto + ultimately show ?thesis by auto + qed + finally show ?thesis by blast + qed + ultimately show "S.terminal t" + using 1 S.terminal_char1 card_1_singleton_iff + by (metis One_nat_def singleton_iff) + qed + + sublocale set_category comp ssetp + proof + text\ + Here things are simpler if we define \img\ appropriately so that we have + \set = T.set\ after accounting for the definition \unity \ SOME t. terminal t\, + which is different from that of S.unity. +\ + have 1: "terminal (SOME t. terminal t)" + using terminal_unity someI_ex [of terminal] by blast + obtain i where i: "\i : S.unity \ SOME t. terminal t\" + using terminal_unity someI_ex [of terminal] in_hom_char terminal_def + by auto + obtain i' where i': "\i' : (SOME t. terminal t) \ S.unity\" + using terminal_unity someI_ex [of S.terminal] S.terminal_def + by (metis (no_types, lifting) 1 terminal_def) + have ii': "inverse_arrows i i'" + proof + have "i' \ i = S.unity" + using i i' terminal_unity + by (metis (no_types, lifting) S.comp_in_homI' S.ide_in_hom S.ide_unity S.in_homE + S.terminalE S.terminal_unity in_hom_char) + thus 2: "ide (comp i' i)" + by (metis (no_types, lifting) cod_comp comp_simp i i' ide_char' in_homE seqI') + have "i \ i' = (SOME t. terminal t)" + using 1 + by (metis (no_types, lifting) 2 comp_simp i' ide_compE in_homE inverse_arrowsE + iso_iff_mono_and_retraction point_is_mono retractionI section_retraction_of_iso(2)) + thus "ide (comp i i')" + using comp_char + by (metis (no_types, lifting) 2 ide_char' dom_comp i' ide_compE in_homE seq_char) + qed + interpret set_category_data comp \\x. S.some_img (x \ i)\ .. + have i_in_hom: "\i : S.unity \ unity\" + using i unity_def by simp + have i'_in_hom: "\i' : unity \ S.unity\" + using i' unity_def by simp + have "\a. ide a \ set a = S.set a" + proof - + fix a + assume a: "ide a" + have "set a = (\x. S.some_img (x \ i)) ` hom unity a" + using set_def by simp + also have "... = S.some_img ` S.hom S.unity a" + proof + show "(\x. S.some_img (x \ i)) ` hom unity a \ S.some_img ` S.hom S.unity a" + using i in_hom_char i_in_hom by auto + show "S.some_img ` S.hom S.unity a \ (\x. S.some_img (x \ i)) ` hom unity a" + proof + fix b + assume b: "b \ S.some_img ` S.hom S.unity a" + obtain x where x: "x \ S.hom S.unity a \ S.some_img x = b" + using b by blast + have "x \ i' \ hom unity a" + using x in_hom_char S.comp_in_homI a i' ideD(1) unity_def by force + moreover have "S.some_img ((x \ i') \ i) = b" + by (metis (no_types, lifting) i ii' x S.comp_assoc calculation comp_simp + ide_compE in_homE inverse_arrowsE mem_Collect_eq S.comp_arr_ide seqI' + seq_char S.ide_unity unity_def) + ultimately show "b \ (\x. S.some_img (x \ i)) ` hom unity a" by blast + qed + qed + also have "... = S.set a" + using S.set_def by auto + finally show "set a = S.set a" by blast + qed + interpret T: set_category_given_img comp \\x. S.some_img (x \ i)\ ssetp + proof + show "Collect terminal \ {}" + using terminal_unity by blast + show "\A' A. \A' \ A; ssetp A\ \ ssetp A'" + using subset_closed by blast + show "\A B. \ssetp A; ssetp B\ \ ssetp (A \ B)" + using union_closed by simp + show "\A. ssetp A \ A \ Univ" + using S.setp_imp_subset_Univ containment terminal_char by presburger + show "\a b. \ide a; ide b; set a = set b\ \ a = b" + using ide_char \\a. ide a \ set a = S.set a\ S.extensional_set by auto + show "\a. ide a \ ssetp (set a)" + using \\a. ide a \ set a = S.set a\ ide_char by force + show "\A. ssetp A \ \a. ide a \ set a = A" + using S.set_complete \\a. ide a \ set a = S.set a\ containment ide_char by blast + show "\t. terminal t \ t \ (\x. S.some_img (x \ i)) ` hom unity t" + using S.set_def S.stable_img \\a. ide a \ set a = S.set a\ set_def + terminal_char terminal_def + by force + show "\a. ide a \ inj_on (\x. S.some_img (x \ i)) (hom unity a)" + proof - + fix a + assume a: "ide a" + show "inj_on (\x. S.some_img (x \ i)) (hom unity a)" + proof + fix x y + assume x: "x \ hom unity a" and y: "y \ hom unity a" + and eq: "S.some_img (x \ i) = S.some_img (y \ i)" + have "x \ i = y \ i" + proof - + have "x \ i \ S.hom S.unity a \ y \ i \ S.hom S.unity a" + using in_hom_char \\i : S.unity \ unity\\ x y by blast + thus ?thesis + using a eq ide_char S.inj_img [of a] inj_on_def [of S.some_img] by simp + qed + have "x = (x \ i) \ i'" + by (metis (no_types, lifting) S.comp_arr_ide S.comp_assoc S.inverse_arrowsE + S.match_4 i i' ii' inclusion_preserves_inverse mem_Collect_eq seqI' + seq_char unity_def x) + also have "... = (y \ i) \ i'" + using \x \ i = y \ i\ by simp + also have "... = y" + by (metis (no_types, lifting) S.comp_arr_ide S.comp_assoc S.inverse_arrowsE + S.match_4 i i' ii' inclusion_preserves_inverse mem_Collect_eq seqI' + seq_char unity_def y) + finally show "x = y" by simp + qed + qed + show "\f f'. \par f f'; \x. in_hom x unity (dom f) \ comp f x = comp f' x\ + \ f = f'" + proof - + fix f f' + assume par: "par f f'" + assume eq: "\x. in_hom x unity (dom f) \ comp f x = comp f' x" + have "S.par f f'" + using par arr_char dom_char cod_char by auto + moreover have "\x. S.in_hom x S.unity (S.dom f) \ f \ x = f' \ x" + proof - + fix x + assume x: "S.in_hom x S.unity (S.dom f)" + have "S.in_hom (x \ i') unity (S.dom f)" + using i'_in_hom in_hom_char x by blast + hence 1: "in_hom (x \ i') unity (dom f)" + using arr_dom dom_simp i in_hom_char par unity_def by force + hence "comp f (x \ i') = comp f' (x \ i')" + using eq by blast + hence "(f \ (x \ i')) \ i = (f' \ (x \ i')) \ i" + using comp_char + by (metis (no_types, lifting) 1 comp_simp in_homE seqI par) + thus "f \ x = f' \ x" + by (metis (no_types, lifting) S.comp_arr_dom S.comp_assoc S.comp_inv_arr + S.in_homE i_in_hom ii' in_hom_char inclusion_preserves_inverse x) + qed + ultimately show "f = f'" + using S.extensional_arr by blast + qed + show "\a b F. \ide a; ide b; F \ hom unity a \ hom unity b\ + \ \f. in_hom f a b \ + (\x. in_hom x unity (dom f) \ comp f x = F x)" + proof - + fix a b F + assume a: "ide a" and b: "ide b" and F: "F \ hom unity a \ hom unity b" + have "S.ide a" + using a ide_char by blast + have "S.ide b" + using b ide_char by blast + have 1: "(\x. F (x \ i') \ i) \ S.hom S.unity a \ S.hom S.unity b" + proof + fix x + assume x: "x \ S.hom S.unity a" + have "x \ i' \ S.hom unity a" + using i'_in_hom in_hom_char x by blast + hence "x \ i' \ hom unity a" + using a in_hom_char + by (metis (no_types, lifting) ideD(1) i'_in_hom in_hom_char mem_Collect_eq) + hence "F (x \ i') \ hom unity b" + using a b F by blast + hence "F (x \ i') \ S.hom unity b" + using b in_hom_char by blast + thus "F (x \ i') \ i \ S.hom S.unity b" + using i in_hom_char unity_def by auto + qed + obtain f where f: "S.in_hom f a b \ (\x. S.in_hom x S.unity (S.dom f) + \ f \ x = (\x. F (x \ i') \ i) x)" + using 1 S.fun_complete_ax \S.ide a\ \S.ide b\ by presburger + show "\f. in_hom f a b \ (\x. in_hom x unity (dom f) \ comp f x = F x)" + proof - + have "in_hom f a b" + using f in_hom_char ideD(1) a b by presburger + moreover have "\x. in_hom x unity (dom f) \ comp f x = F x" + proof (intro allI impI) + fix x + assume x: "in_hom x unity (dom f)" + have xi: "S.in_hom (x \ i) S.unity (S.dom f)" + using f x i in_hom_char dom_char + by (metis (no_types, lifting) in_homE unity_def calculation S.comp_in_homI) + hence 1: "f \ (x \ i) = F ((x \ i) \ i') \ i" + using f by blast + hence "((f \ x) \ i) \ i' = (F x \ i) \ i'" + by (metis (no_types, lifting) xi S.comp_assoc S.inverse_arrowsE + S.seqI' i' ii' in_hom_char inclusion_preserves_inverse S.comp_arr_ide) + hence "f \ x = F x" + by (metis (no_types, lifting) xi 1 S.invert_side_of_triangle(2) S.match_2 S.match_3 + S.seqI arr_char calculation S.in_homE S.inverse_unique S.isoI + ii' in_homE inclusion_preserves_inverse) + thus "comp f x = F x" + using comp_char + by (metis (no_types, lifting) comp_simp in_homE seqI calculation x) + qed + ultimately show ?thesis by blast + qed + qed + qed + show "\img. set_category_given_img comp img ssetp" + using T.set_category_given_img_axioms by blast + qed + + lemma is_set_category: + shows "set_category comp ssetp" + .. + + end end diff --git a/thys/Category3/Subcategory.thy b/thys/Category3/Subcategory.thy --- a/thys/Category3/Subcategory.thy +++ b/thys/Category3/Subcategory.thy @@ -1,427 +1,431 @@ (* Title: Subcategory Author: Eugene W. Stark , 2017 Maintainer: Eugene W. Stark *) chapter "Subcategory" text\ In this chapter we give a construction of the subcategory of a category defined by a predicate on arrows subject to closure conditions. The arrows of the subcategory are directly identified with the arrows of the ambient category. We also define the related notions of full subcategory and inclusion functor. \ theory Subcategory imports Functor begin locale subcategory = C: category C for C :: "'a comp" (infixr "\\<^sub>C" 55) and Arr :: "'a \ bool" + assumes inclusion: "Arr f \ C.arr f" and dom_closed: "Arr f \ Arr (C.dom f)" and cod_closed: "Arr f \ Arr (C.cod f)" and comp_closed: "\ Arr f; Arr g; C.cod f = C.dom g \ \ Arr (g \\<^sub>C f)" begin no_notation C.in_hom ("\_ : _ \ _\") notation C.in_hom ("\_ : _ \\<^sub>C _\") definition comp (infixr "\" 55) where "g \ f = (if Arr f \ Arr g \ C.cod f = C.dom g then g \\<^sub>C f else C.null)" interpretation partial_magma comp proof show "\!n. \f. n \ f = n \ f \ n = n" proof show 1: "\f. C.null \ f = C.null \ f \ C.null = C.null" by (metis C.comp_null(1) C.ex_un_null comp_def) show "\n. \f. n \ f = n \ f \ n = n \ n = C.null" using 1 C.ex_un_null by metis qed qed lemma null_char [simp]: shows "null = C.null" proof - have "\f. C.null \ f = C.null \ f \ C.null = C.null" by (metis C.comp_null(1) C.ex_un_null comp_def) thus ?thesis using ex_un_null by (metis comp_null(2)) qed lemma ideI: assumes "Arr a" and "C.ide a" shows "ide a" unfolding ide_def using assms null_char C.ide_def comp_def by auto lemma Arr_iff_dom_in_domain: shows "Arr f \ C.dom f \ domains f" proof show "C.dom f \ domains f \ Arr f" using domains_def comp_def ide_def by fastforce show "Arr f \ C.dom f \ domains f" proof - assume f: "Arr f" have "ide (C.dom f)" using f inclusion C.dom_in_domains C.has_domain_iff_arr C.domains_def dom_closed ideI by auto moreover have "f \ C.dom f \ null" using f comp_def dom_closed null_char inclusion C.comp_arr_dom by force ultimately show ?thesis using domains_def by simp qed qed lemma Arr_iff_cod_in_codomain: shows "Arr f \ C.cod f \ codomains f" proof show "C.cod f \ codomains f \ Arr f" using codomains_def comp_def ide_def by fastforce show "Arr f \ C.cod f \ codomains f" proof - assume f: "Arr f" have "ide (C.cod f)" using f inclusion C.cod_in_codomains C.has_codomain_iff_arr C.codomains_def cod_closed ideI by auto moreover have "C.cod f \ f \ null" using f comp_def cod_closed null_char inclusion C.comp_cod_arr by force ultimately show ?thesis using codomains_def by simp qed qed lemma arr_char: shows "arr f \ Arr f" proof show "Arr f \ arr f" using arr_def comp_def Arr_iff_dom_in_domain Arr_iff_cod_in_codomain by auto show "arr f \ Arr f" proof - assume f: "arr f" obtain a where a: "a \ domains f \ a \ codomains f" using f arr_def by auto have "f \ a \ C.null \ a \ f \ C.null" using a domains_def codomains_def null_char by auto thus "Arr f" using comp_def by metis qed qed lemma arrI [intro]: assumes "Arr f" shows "arr f" using assms arr_char by simp lemma arrE [elim]: assumes "arr f" shows "Arr f" using assms arr_char by simp interpretation category comp using comp_def null_char inclusion comp_closed dom_closed cod_closed apply unfold_locales apply auto[1] apply (metis Arr_iff_dom_in_domain Arr_iff_cod_in_codomain arr_char arr_def emptyE) proof - fix f g h assume gf: "seq g f" and hg: "seq h g" show 1: "seq (h \ g) f" using gf hg inclusion comp_closed comp_def by auto show "(h \ g) \ f = h \ g \ f" using gf hg 1 C.not_arr_null inclusion comp_def arr_char by (metis (full_types) C.cod_comp C.comp_assoc) next fix f g h assume hg: "seq h g" and hgf: "seq (h \ g) f" show "seq g f" using hg hgf comp_def null_char inclusion arr_char comp_closed by (metis (full_types) C.dom_comp) next fix f g h assume hgf: "seq h (g \ f)" and gf: "seq g f" show "seq h g" using hgf gf comp_def null_char arr_char comp_closed by (metis C.seqE C.ext C.match_2) qed theorem is_category: shows "category comp" .. notation in_hom ("\_ : _ \ _\") lemma dom_simp: assumes "arr f" shows "dom f = C.dom f" proof - have "ide (C.dom f)" using assms ideI by (meson C.ide_dom arr_char dom_closed inclusion) moreover have "f \ C.dom f \ null" using assms inclusion comp_def null_char dom_closed not_arr_null C.comp_arr_dom arr_char by auto ultimately show ?thesis using dom_eqI ext by blast qed lemma dom_char: shows "dom f = (if arr f then C.dom f else C.null)" using dom_simp dom_def arr_def arr_char by auto lemma cod_simp: assumes "arr f" shows "cod f = C.cod f" proof - have "ide (C.cod f)" using assms ideI by (meson C.ide_cod arr_char cod_closed inclusion) moreover have "C.cod f \ f \ null" using assms inclusion comp_def null_char cod_closed not_arr_null C.comp_cod_arr arr_char by auto ultimately show ?thesis using cod_eqI ext by blast qed lemma cod_char: shows "cod f = (if arr f then C.cod f else C.null)" using cod_simp cod_def arr_def by auto lemma in_hom_char: shows "\f : a \ b\ \ arr a \ arr b \ arr f \ \f : a \\<^sub>C b\" using inclusion arr_char cod_closed dom_closed by (metis C.arr_iff_in_hom C.in_homE arr_iff_in_hom cod_simp dom_simp in_homE) lemma ide_char: shows "ide a \ arr a \ C.ide a" using ide_in_hom C.ide_in_hom in_hom_char by simp lemma seq_char: shows "seq g f \ arr f \ arr g \ C.seq g f" proof show "arr f \ arr g \ C.seq g f \ seq g f" using arr_char dom_char cod_char by (intro seqI, auto) show "seq g f \ arr f \ arr g \ C.seq g f" apply (elim seqE, auto) using inclusion arr_char dom_simp cod_simp by auto qed lemma hom_char: shows "hom a b = C.hom a b \ Collect Arr" proof show "hom a b \ C.hom a b \ Collect Arr" using in_hom_char by auto show "C.hom a b \ Collect Arr \ hom a b" using arr_char dom_char cod_char by force qed lemma comp_char: shows "g \ f = (if arr f \ arr g \ C.seq g f then g \\<^sub>C f else C.null)" using arr_char comp_def comp_closed C.ext by force lemma comp_simp: assumes "seq g f" shows "g \ f = g \\<^sub>C f" using assms comp_char seq_char by metis lemma inclusion_preserves_inverse: assumes "inverse_arrows f g" shows "C.inverse_arrows f g" using assms ide_char comp_simp arr_char by (intro C.inverse_arrowsI, auto) lemma iso_char: shows "iso f \ C.iso f \ arr f \ arr (C.inv f)" proof assume f: "iso f" show "C.iso f \ arr f \ arr (C.inv f)" proof - have 1: "inverse_arrows f (inv f)" using f inv_is_inverse by auto have 2: "C.inverse_arrows f (inv f)" using 1 inclusion_preserves_inverse by blast moreover have "arr (inv f)" using 1 iso_is_arr by blast moreover have "inv f = C.inv f" using 1 2 C.inv_is_inverse C.inverse_arrow_unique by blast ultimately show ?thesis using f 2 iso_is_arr by auto qed next assume f: "C.iso f \ arr f \ arr (C.inv f)" show "iso f" proof have 1: "C.inverse_arrows f (C.inv f)" using f C.inv_is_inverse by blast show "inverse_arrows f (C.inv f)" proof have 2: "C.inv f \ f = C.inv f \\<^sub>C f \ f \ C.inv f = f \\<^sub>C C.inv f" using f 1 comp_char by fastforce have 3: "antipar f (C.inv f)" using f C.seqE seqI dom_simp cod_simp by simp show "ide (C.inv f \ f)" using 1 2 3 ide_char apply (elim C.inverse_arrowsE) by simp show "ide (f \ C.inv f)" using 1 2 3 ide_char apply (elim C.inverse_arrowsE) by simp qed qed qed lemma inv_char: assumes "iso f" shows "inv f = C.inv f" proof - have "C.inverse_arrows f (inv f)" proof have 1: "inverse_arrows f (inv f)" using assms iso_char inv_is_inverse by blast show "C.ide (inv f \\<^sub>C f)" proof - have "inv f \ f = inv f \\<^sub>C f" using assms 1 inv_in_hom iso_char [of f] comp_char [of "inv f" f] seq_char by auto thus ?thesis using 1 ide_char arr_char by force qed show "C.ide (f \\<^sub>C inv f)" proof - have "f \ inv f = f \\<^sub>C inv f" using assms 1 inv_in_hom iso_char [of f] comp_char [of f "inv f"] seq_char by auto thus ?thesis using 1 ide_char arr_char by force qed qed thus ?thesis using C.inverse_arrow_unique C.inv_is_inverse by blast qed end sublocale subcategory \ category comp using is_category by auto section "Full Subcategory" locale full_subcategory = C: category C for C :: "'a comp" and Ide :: "'a \ bool" + assumes inclusion: "Ide f \ C.ide f" sublocale full_subcategory \ subcategory C "\f. C.arr f \ Ide (C.dom f) \ Ide (C.cod f)" by (unfold_locales; simp) context full_subcategory begin + lemma in_hom_char: + shows "\f : a \ b\ \ arr a \ arr b \ \f : a \\<^sub>C b\" + using arr_char in_hom_char by auto + text \ Isomorphisms in a full subcategory are inherited from the ambient category. \ lemma iso_char: shows "iso f \ arr f \ C.iso f" proof assume f: "iso f" obtain g where g: "inverse_arrows f g" using f by blast show "arr f \ C.iso f" proof - have "C.inverse_arrows f g" using g apply (elim inverse_arrowsE, intro C.inverse_arrowsI) using comp_simp ide_char arr_char by auto thus ?thesis using f iso_is_arr by blast qed next assume f: "arr f \ C.iso f" obtain g where g: "C.inverse_arrows f g" using f by blast have "inverse_arrows f g" proof show "ide (comp g f)" using f g by (metis (no_types, lifting) C.seqE C.ide_compE C.inverse_arrowsE arr_char dom_simp ide_dom comp_def) show "ide (comp f g)" using f g C.inverse_arrows_sym [of f g] by (metis (no_types, lifting) C.seqE C.ide_compE C.inverse_arrowsE arr_char dom_simp ide_dom comp_def) qed thus "iso f" by auto qed end section "Inclusion Functor" text \ If \S\ is a subcategory of \C\, then there is an inclusion functor from \S\ to \C\. Inclusion functors are faithful embeddings. \ locale inclusion_functor = C: category C + S: subcategory C Arr for C :: "'a comp" and Arr :: "'a \ bool" begin interpretation "functor" S.comp C S.map using S.map_def S.arr_char S.inclusion S.dom_char S.cod_char S.dom_closed S.cod_closed S.comp_closed S.arr_char S.comp_char apply unfold_locales apply auto[4] by (elim S.seqE, auto) lemma is_functor: shows "functor S.comp C S.map" .. interpretation faithful_functor S.comp C S.map apply unfold_locales by simp lemma is_faithful_functor: shows "faithful_functor S.comp C S.map" .. interpretation embedding_functor S.comp C S.map apply unfold_locales by simp lemma is_embedding_functor: shows "embedding_functor S.comp C S.map" .. end sublocale inclusion_functor \ faithful_functor S.comp C S.map using is_faithful_functor by auto sublocale inclusion_functor \ embedding_functor S.comp C S.map using is_embedding_functor by auto text \ The inclusion of a full subcategory is a special case. Such functors are fully faithful. \ locale full_inclusion_functor = C: category C + S: full_subcategory C Ide for C :: "'a comp" and Ide :: "'a \ bool" sublocale full_inclusion_functor \ inclusion_functor C "\f. C.arr f \ Ide (C.dom f) \ Ide (C.cod f)" .. context full_inclusion_functor begin interpretation full_functor S.comp C S.map apply unfold_locales using S.ide_in_hom by (metis (no_types, lifting) C.in_homE S.arr_char S.in_hom_char S.map_simp) lemma is_full_functor: shows "full_functor S.comp C S.map" .. end sublocale full_inclusion_functor \ full_functor S.comp C S.map using is_full_functor by auto sublocale full_inclusion_functor \ fully_faithful_functor S.comp C S.map .. end diff --git a/thys/Category3/Yoneda.thy b/thys/Category3/Yoneda.thy --- a/thys/Category3/Yoneda.thy +++ b/thys/Category3/Yoneda.thy @@ -1,1088 +1,1080 @@ (* Title: Yoneda Author: Eugene W. Stark , 2016 Maintainer: Eugene W. Stark *) chapter Yoneda theory Yoneda imports DualCategory SetCat FunctorCategory begin text\ This theory defines the notion of a ``hom-functor'' and gives a proof of the Yoneda Lemma. In traditional developments of category theory based on set theories such as ZFC, hom-functors are normally defined to be functors into the large category \textbf{Set} whose objects are of \emph{all} sets and whose arrows are functions between sets. However, in HOL there does not exist a single ``type of all sets'', so the notion of the category of \emph{all} sets and functions does not make sense. To work around this, we consider a more general setting consisting of a category @{term C} together with a set category @{term S} and a function @{term "\ :: 'c * 'c \ 'c \ 's"} such that whenever @{term b} and @{term a} are objects of C then @{term "\ (b, a)"} maps \C.hom b a\ injectively to \S.Univ\. We show that these data induce a binary functor \Hom\ from \Cop\C\ to @{term S} in such a way that @{term \} is rendered natural in @{term "(b, a)"}. The Yoneda lemma is then proved for the Yoneda functor determined by \Hom\. \ section "Hom-Functors" text\ A hom-functor for a category @{term C} allows us to regard the hom-sets of @{term C} as objects of a category @{term S} of sets and functions. Any description of a hom-functor for @{term C} must therefore specify the category @{term S} and provide some sort of correspondence between arrows of @{term C} and elements of objects of @{term S}. If we are to think of each hom-set \C.hom b a\ of \C\ as corresponding to an object \Hom (b, a)\ of @{term S} then at a minimum it ought to be the case that the correspondence between arrows and elements is bijective between \C.hom b a\ and \Hom (b, a)\. The \hom_functor\ locale defined below captures this idea by assuming a set category @{term S} and a function @{term \} taking arrows of @{term C} to elements of \S.Univ\, such that @{term \} is injective on each set \C.hom b a\. We show that these data induce a functor \Hom\ from \Cop\C\ to \S\ in such a way that @{term \} becomes a natural bijection between \C.hom b a\ and \Hom (b, a)\. \ locale hom_functor = C: category C + - Cop: dual_category C + - CopxC: product_category Cop.comp C + - S: replete_set_category S + S: set_category S setp for C :: "'c comp" (infixr "\" 55) and S :: "'s comp" (infixr "\\<^sub>S" 55) + and setp :: "'s set \ bool" and \ :: "'c * 'c \ 'c \ 's" + assumes maps_arr_to_Univ: "C.arr f \ \ (C.dom f, C.cod f) f \ S.Univ" and local_inj: "\ C.ide b; C.ide a \ \ inj_on (\ (b, a)) (C.hom b a)" + and small_homs: "\ C.ide b; C.ide a \ \ setp (\ (b, a) ` C.hom b a)" begin + sublocale Cop: dual_category C .. + sublocale CopxC: product_category Cop.comp C .. + notation S.in_hom ("\_ : _ \\<^sub>S _\") notation CopxC.comp (infixr "\" 55) notation CopxC.in_hom ("\_ : _ \ _\") definition set where "set ba \ \ (fst ba, snd ba) ` C.hom (fst ba) (snd ba)" lemma set_subset_Univ: assumes "C.ide b" and "C.ide a" shows "set (b, a) \ S.Univ" using assms set_def maps_arr_to_Univ CopxC.ide_char by auto definition \ :: "'c * 'c \ 's \ 'c" where "\ ba = inv_into (C.hom (fst ba) (snd ba)) (\ ba)" lemma \_mapsto: assumes "C.ide b" and "C.ide a" shows "\ (b, a) \ C.hom b a \ set (b, a)" using assms set_def maps_arr_to_Univ by auto lemma \_mapsto: assumes "C.ide b" and "C.ide a" shows "\ (b, a) \ set (b, a) \ C.hom b a" using assms set_def \_def local_inj by auto lemma \_\ [simp]: assumes "\f : b \ a\" shows "\ (b, a) (\ (b, a) f) = f" using assms local_inj [of b a] \_def by fastforce lemma \_\ [simp]: assumes "C.ide b" and "C.ide a" and "x \ set (b, a)" shows "\ (b, a) (\ (b, a) x) = x" using assms set_def local_inj \_def by auto lemma \_img_set: assumes "C.ide b" and "C.ide a" shows "\ (b, a) ` set (b, a) = C.hom b a" using assms \_def set_def local_inj by auto text\ A hom-functor maps each arrow @{term "(g, f)"} of @{term "CopxC"} to the arrow of the set category @{term[source=true] S} corresponding to the function that takes an arrow @{term h} of @{term C} to the arrow @{term "C f (C h g)"} of @{term C} obtained by precomposing with @{term g} and postcomposing with @{term f}. \ definition map where "map gf = (if CopxC.arr gf then S.mkArr (set (CopxC.dom gf)) (set (CopxC.cod gf)) (\ (CopxC.cod gf) o (\h. snd gf \ h \ fst gf) o \ (CopxC.dom gf)) else S.null)" lemma arr_map: assumes "CopxC.arr gf" shows "S.arr (map gf)" proof - have "\ (CopxC.cod gf) o (\h. snd gf \ h \ fst gf) o \ (CopxC.dom gf) \ set (CopxC.dom gf) \ set (CopxC.cod gf)" using assms \_mapsto [of "fst (CopxC.cod gf)" "snd (CopxC.cod gf)"] \_mapsto [of "fst (CopxC.dom gf)" "snd (CopxC.dom gf)"] by fastforce thus ?thesis - using assms map_def S.arr_mkArr set_subset_Univ S.card_of_leq by simp + using assms map_def set_subset_Univ small_homs + by (simp add: set_def) qed lemma map_ide [simp]: assumes "C.ide b" and "C.ide a" shows "map (b, a) = S.mkIde (set (b, a))" proof - have "map (b, a) = S.mkArr (set (b, a)) (set (b, a)) (\ (b, a) o (\h. a \ h \ b) o \ (b, a))" using assms map_def by auto also have "... = S.mkArr (set (b, a)) (set (b, a)) (\h. h)" proof - have "S.mkArr (set (b, a)) (set (b, a)) (\h. h) = ..." - using assms S.arr_mkArr set_subset_Univ set_def C.comp_arr_dom C.comp_cod_arr - S.card_of_leq S.arr_mkIde + using assms set_subset_Univ set_def C.comp_arr_dom C.comp_cod_arr + S.arr_mkIde small_homs by (intro S.mkArr_eqI', simp, fastforce) thus ?thesis by auto qed also have "... = S.mkIde (set (b, a))" - using assms S.mkIde_as_mkArr set_subset_Univ by simp + using assms S.mkIde_as_mkArr set_subset_Univ small_homs + by (simp add: set_def) finally show ?thesis by auto qed lemma set_map: assumes "C.ide a" and "C.ide b" shows "S.set (map (b, a)) = set (b, a)" - using assms map_ide S.set_mkIde set_subset_Univ by simp + using assms map_ide set_subset_Univ small_homs set_def by simp text\ The definition does in fact yield a functor. \ - interpretation "functor" CopxC.comp S map + sublocale "functor" CopxC.comp S map proof fix gf assume "\CopxC.arr gf" thus "map gf = S.null" using map_def by auto next fix gf assume gf: "CopxC.arr gf" thus arr: "S.arr (map gf)" using gf arr_map by blast show "S.dom (map gf) = map (CopxC.dom gf)" using arr gf local.map_def map_ide by auto show "S.cod (map gf) = map (CopxC.cod gf)" - using gf set_subset_Univ \_mapsto map_def set_def S.card_of_leq S.arr_mkIde S.arr_mkArr - arr map_ide + using gf set_subset_Univ \_mapsto map_def set_def S.arr_mkIde arr map_ide by auto next fix gf gf' assume gf': "CopxC.seq gf' gf" hence seq: "C.arr (fst gf) \ C.arr (snd gf) \ C.dom (snd gf') = C.cod (snd gf) \ C.arr (fst gf') \ C.arr (snd gf') \ C.dom (fst gf) = C.cod (fst gf')" by (elim CopxC.seqE C.seqE, auto) have 0: "S.arr (map (CopxC.comp gf' gf))" using gf' arr_map by blast have 1: "map (gf' \ gf) = S.mkArr (set (CopxC.dom gf)) (set (CopxC.cod gf')) (\ (CopxC.cod gf') o (\h. snd (gf' \ gf) \ h \ fst (gf' \ gf)) o \ (CopxC.dom gf))" using gf' map_def using CopxC.cod_comp CopxC.dom_comp by auto also have "... = S.mkArr (set (CopxC.dom gf)) (set (CopxC.cod gf')) (\ (CopxC.cod gf') \ (\h. snd gf' \ h \ fst gf') \ \ (CopxC.dom gf') \ (\ (CopxC.cod gf) \ (\h. snd gf \ h \ fst gf) \ \ (CopxC.dom gf)))" proof (intro S.mkArr_eqI') show "S.arr (S.mkArr (set (CopxC.dom gf)) (set (CopxC.cod gf')) (\ (CopxC.cod gf') \ (\h. snd (gf' \ gf) \ h \ fst (gf' \ gf)) \ \ (CopxC.dom gf)))" using 0 1 by simp show "\x. x \ set (CopxC.dom gf) \ (\ (CopxC.cod gf') \ (\h. snd (gf' \ gf) \ h \ fst (gf' \ gf)) \ \ (CopxC.dom gf)) x = (\ (CopxC.cod gf') \ (\h. snd gf' \ h \ fst gf') \ \ (CopxC.dom gf') \ (\ (CopxC.cod gf) \ (\h. snd gf \ h \ fst gf) \ \ (CopxC.dom gf))) x" proof - fix x assume "x \ set (CopxC.dom gf)" hence x: "x \ set (C.cod (fst gf), C.dom (snd gf))" by (simp add: seq) show "(\ (CopxC.cod gf') \ (\h. snd (gf' \ gf) \ h \ fst (gf' \ gf)) \ \ (CopxC.dom gf)) x = (\ (CopxC.cod gf') \ (\h. snd gf' \ h \ fst gf') \ \ (CopxC.dom gf') \ (\ (CopxC.cod gf) \ (\h. snd gf \ h \ fst gf) \ \ (CopxC.dom gf))) x" proof - have "(\ (CopxC.cod gf') o (\h. snd (gf' \ gf) \ h \ fst (gf' \ gf)) o \ (CopxC.dom gf)) x = \ (CopxC.cod gf') (snd (gf' \ gf) \ \ (CopxC.dom gf) x \ fst (gf' \ gf))" by simp also have "... = \ (CopxC.cod gf') (((\h. snd gf' \ h \ fst gf') \ \ (CopxC.dom gf') \ (\ (CopxC.dom gf') \ (\h. snd gf \ h \ fst gf))) (\ (CopxC.dom gf) x))" proof - have "C.ide (C.cod (fst gf)) \ C.ide (C.dom (snd gf))" using gf' by (elim CopxC.seqE, auto) hence "\\ (C.cod (fst gf), C.dom (snd gf)) x : C.cod (fst gf) \ C.dom (snd gf)\" using x \_mapsto by auto hence "\snd gf \ \ (C.cod (fst gf), C.dom (snd gf)) x \ fst gf : C.cod (fst gf') \ C.dom (snd gf')\" using x seq by auto thus ?thesis using seq \_\ C.comp_assoc by auto qed also have "... = (\ (CopxC.cod gf') \ (\h. snd gf' \ h \ fst gf') \ \ (CopxC.dom gf') \ (\ (CopxC.dom gf') \ (\h. snd gf \ h \ fst gf) \ \ (CopxC.dom gf))) x" by auto finally show ?thesis using seq by simp qed qed qed also have "... = map gf' \\<^sub>S map gf" using seq gf' map_def arr_map [of gf] arr_map [of gf'] S.comp_mkArr by auto finally show "map (gf' \ gf) = map gf' \\<^sub>S map gf" using seq gf' by auto qed - interpretation binary_functor Cop.comp C S map .. + lemma is_functor: + shows "functor CopxC.comp S map" .. + + sublocale binary_functor Cop.comp C S map .. lemma is_binary_functor: shows "binary_functor Cop.comp C S map" .. - end - - sublocale hom_functor \ binary_functor Cop.comp C S map - using is_binary_functor by auto - - context hom_functor - begin - text\ The map @{term \} determines a bijection between @{term "C.hom b a"} and @{term "set (b, a)"} which is natural in @{term "(b, a)"}. \ lemma \_local_bij: assumes "C.ide b" and "C.ide a" shows "bij_betw (\ (b, a)) (C.hom b a) (set (b, a))" using assms local_inj inj_on_imp_bij_betw set_def by auto lemma \_natural: assumes "C.arr g" and "C.arr f" and "h \ C.hom (C.cod g) (C.dom f)" shows "\ (C.dom g, C.cod f) (f \ h \ g) = S.Fun (map (g, f)) (\ (C.cod g, C.dom f) h)" proof - let ?\h = "\ (C.cod g, C.dom f) h" have \h: "?\h \ set (C.cod g, C.dom f)" using assms \_mapsto set_def by simp have gf: "CopxC.arr (g, f)" using assms by simp have "map (g, f) = S.mkArr (set (C.cod g, C.dom f)) (set (C.dom g, C.cod f)) (\ (C.dom g, C.cod f) \ (\h. f \ h \ g) \ \ (C.cod g, C.dom f))" using assms map_def by simp moreover have "S.arr (map (g, f))" using gf by simp ultimately have "S.Fun (map (g, f)) = restrict (\ (C.dom g, C.cod f) \ (\h. f \ h \ g) \ \ (C.cod g, C.dom f)) (set (C.cod g, C.dom f))" using S.Fun_mkArr by simp hence "S.Fun (map (g, f)) ?\h = (\ (C.dom g, C.cod f) \ (\h. f \ h \ g) \ \ (C.cod g, C.dom f)) ?\h" using \h by simp also have "... = \ (C.dom g, C.cod f) (f \ h \ g)" using assms(3) by simp finally show ?thesis by auto qed lemma Dom_map: assumes "C.arr g" and "C.arr f" shows "S.Dom (map (g, f)) = set (C.cod g, C.dom f)" - using assms map_def preserves_arr S.set_mkIde S.arr_mkArr by auto + using assms map_def preserves_arr by auto lemma Cod_map: assumes "C.arr g" and "C.arr f" shows "S.Cod (map (g, f)) = set (C.dom g, C.cod f)" - using assms map_def preserves_arr S.set_mkIde S.arr_mkArr by auto + using assms map_def preserves_arr by auto lemma Fun_map: assumes "C.arr g" and "C.arr f" shows "S.Fun (map (g, f)) = restrict (\ (C.dom g, C.cod f) o (\h. f \ h \ g) o \ (C.cod g, C.dom f)) (set (C.cod g, C.dom f))" using assms map_def preserves_arr by force lemma map_simp_1: assumes "C.arr g" and "C.ide a" shows "map (g, a) = S.mkArr (set (C.cod g, a)) (set (C.dom g, a)) (\ (C.dom g, a) o Cop.comp g o \ (C.cod g, a))" proof - have 1: "map (g, a) = S.mkArr (set (C.cod g, a)) (set (C.dom g, a)) (\ (C.dom g, a) o (\h. a \ h \ g) o \ (C.cod g, a))" using assms map_def by force also have "... = S.mkArr (set (C.cod g, a)) (set (C.dom g, a)) (\ (C.dom g, a) o Cop.comp g o \ (C.cod g, a))" using assms 1 preserves_arr [of "(g, a)"] set_def C.in_homI C.comp_cod_arr by (intro S.mkArr_eqI) auto finally show ?thesis by blast qed lemma map_simp_2: assumes "C.ide b" and "C.arr f" shows "map (b, f) = S.mkArr (set (b, C.dom f)) (set (b, C.cod f)) (\ (b, C.cod f) o C f o \ (b, C.dom f))" proof - have 1: "map (b, f) = S.mkArr (set (b, C.dom f)) (set (b, C.cod f)) (\ (b, C.cod f) o (\h. f \ h \ b) o \ (b, C.dom f))" using assms map_def by force also have "... = S.mkArr (set (b, C.dom f)) (set (b, C.cod f)) (\ (b, C.cod f) o C f o \ (b, C.dom f))" using assms 1 preserves_arr [of "(b, f)"] set_def C.in_homI C.comp_arr_dom by (intro S.mkArr_eqI) auto finally show ?thesis by blast qed end text\ Every category @{term C} has a hom-functor: take @{term S} to be the replete set category generated by the arrow type \'a\ of @{term C} and take @{term "\ (b, a)"} to be the map \S.UP :: 'a \ 'a SC.arr\. \ context category begin interpretation Cop: dual_category C .. interpretation CopxC: product_category Cop.comp C .. interpretation S: replete_setcat \undefined :: 'a\ . lemma has_hom_functor: - shows "hom_functor C (S.comp :: 'a setcat.arr comp) (\_. S.UP)" + shows "hom_functor C S.comp S.setp (\_. S.UP)" proof show "\f. arr f \ S.UP f \ S.Univ" using S.UP_mapsto by auto + thus "\b a. \ide b; ide a\ \ S.UP ` hom b a \ S.Univ" + by force show "\b a. \ide b; ide a\ \ inj_on S.UP (hom b a)" by (meson S.inj_UP injD inj_onI) qed end text\ The locales \set_valued_functor\ and \set_valued_transformation\ provide some abbreviations that are convenient when working with functors and natural transformations into a set category. \ locale set_valued_functor = C: category C + - S: replete_set_category S + + S: set_category S setp + "functor" C S F for C :: "'c comp" and S :: "'s comp" + and setp :: "'s set \ bool" and F :: "'c \ 's" begin abbreviation SET :: "'c \ 's set" where "SET a \ S.set (F a)" abbreviation DOM :: "'c \ 's set" where "DOM f \ S.Dom (F f)" abbreviation COD :: "'c \ 's set" where "COD f \ S.Cod (F f)" abbreviation FUN :: "'c \ 's \ 's" where "FUN f \ S.Fun (F f)" end locale set_valued_transformation = C: category C + - S: replete_set_category S + - F: set_valued_functor C S F + - G: set_valued_functor C S G + + S: set_category S setp + + F: set_valued_functor C S setp F + + G: set_valued_functor C S setp G + natural_transformation C S F G \ for C :: "'c comp" and S :: "'s comp" + and setp :: "'s set \ bool" and F :: "'c \ 's" and G :: "'c \ 's" and \ :: "'c \ 's" begin abbreviation DOM :: "'c \ 's set" where "DOM f \ S.Dom (\ f)" abbreviation COD :: "'c \ 's set" where "COD f \ S.Cod (\ f)" abbreviation FUN :: "'c \ 's \ 's" where "FUN f \ S.Fun (\ f)" end section "Yoneda Functors" text\ A Yoneda functor is the functor from @{term C} to \[Cop, S]\ obtained by ``currying'' a hom-functor in its first argument. \ locale yoneda_functor = C: category C + Cop: dual_category C + CopxC: product_category Cop.comp C + - S: replete_set_category S + - Hom: hom_functor C S \ + - Cop_S: functor_category Cop.comp S + - curried_functor' Cop.comp C S Hom.map + S: set_category S setp + + Hom: hom_functor C S setp \ for C :: "'c comp" (infixr "\" 55) and S :: "'s comp" (infixr "\\<^sub>S" 55) + and setp :: "'s set \ bool" and \ :: "'c * 'c \ 'c \ 's" begin + sublocale Cop_S: functor_category Cop.comp S .. + sublocale curried_functor' Cop.comp C S Hom.map .. + notation Cop_S.in_hom ("\_ : _ \\<^sub>[\<^sub>C\<^sub>o\<^sub>p\<^sub>,\<^sub>S\<^sub>] _\") abbreviation \ where "\ \ Hom.\" text\ An arrow of the functor category \[Cop, S]\ consists of a natural transformation bundled together with its domain and codomain functors. However, when considering a Yoneda functor from @{term[source=true] C} to \[Cop, S]\ we generally are only interested in the mapping @{term Y} that takes each arrow @{term f} of @{term[source=true] C} to the corresponding natural transformation @{term "Y f"}. The domain and codomain functors are then the identity transformations @{term "Y (C.dom f)"} and @{term "Y (C.cod f)"}. \ definition Y where "Y f \ Cop_S.Map (map f)" lemma Y_simp [simp]: assumes "C.arr f" shows "Y f = (\g. Hom.map (g, f))" using assms preserves_arr Y_def by simp lemma Y_ide_is_functor: assumes "C.ide a" shows "functor Cop.comp S (Y a)" using assms Y_def Hom.fixing_ide_gives_functor_2 by force lemma Y_arr_is_transformation: assumes "C.arr f" shows "natural_transformation Cop.comp S (Y (C.dom f)) (Y (C.cod f)) (Y f)" using assms Y_def [of f] map_def Hom.fixing_arr_gives_natural_transformation_2 preserves_dom preserves_cod by fastforce lemma Y_ide_arr [simp]: assumes a: "C.ide a" and "\g : b' \ b\" shows "\Y a g : Hom.map (b, a) \\<^sub>S Hom.map (b', a)\" and "Y a g = S.mkArr (Hom.set (b, a)) (Hom.set (b', a)) (\ (b', a) o Cop.comp g o \ (b, a))" using assms Hom.map_simp_1 by (fastforce, auto) lemma Y_arr_ide [simp]: assumes "C.ide b" and "\f : a \ a'\" shows "\Y f b : Hom.map (b, a) \\<^sub>S Hom.map (b, a')\" and "Y f b = S.mkArr (Hom.set (b, a)) (Hom.set (b, a')) (\ (b, a') o C f o \ (b, a))" using assms apply fastforce using assms Hom.map_simp_2 by auto end locale yoneda_functor_fixed_object = - yoneda_functor C S \ - for C :: "'c comp" (infixr "\" 55) - and S :: "'s comp" (infixr "\\<^sub>S" 55) - and \ :: "'c * 'c \ 'c \ 's" - and a :: 'c + + yoneda_functor + + fixes a assumes ide_a: "C.ide a" + begin - sublocale yoneda_functor_fixed_object \ "functor" Cop.comp S "(Y a)" - using ide_a Y_ide_is_functor by auto - sublocale yoneda_functor_fixed_object \ set_valued_functor Cop.comp S "(Y a)" .. + sublocale "functor" Cop.comp S \Y a\ + using ide_a Y_ide_is_functor by auto + sublocale set_valued_functor Cop.comp S setp \Y a\ .. + + end text\ The Yoneda lemma states that, given a category @{term C} and a functor @{term F} from @{term Cop} to a set category @{term S}, for each object @{term a} of @{term C}, the set of natural transformations from the contravariant functor @{term "Y a"} to @{term F} is in bijective correspondence with the set \F.SET a\ of elements of @{term "F a"}. Explicitly, if @{term e} is an arbitrary element of the set \F.SET a\, then the functions \\x. F.FUN (\ (b, a) x) e\ are the components of a natural transformation from @{term "Y a"} to @{term F}. Conversely, if @{term \} is a natural transformation from @{term "Y a"} to @{term F}, then the component @{term "\ b"} of @{term \} at an arbitrary object @{term b} is completely determined by the single arrow \\.FUN a (\ (a, a) a)))\, which is the the element of \F.SET a\ that corresponds to the image of the identity @{term a} under the function \\.FUN a\. Then @{term "\ b"} is the arrow from @{term "Y a b"} to @{term "F b"} corresponding to the function \\x. (F.FUN (\ (b, a) x) (\.FUN a (\ (a, a) a)))\ from \S.set (Y a b)\ to \F.SET b\. The above expressions look somewhat more complicated than the usual versions due to the need to account for the coercions @{term \} and @{term \}. \ locale yoneda_lemma = - C: category C + - Cop: dual_category C + - S: replete_set_category S + - F: set_valued_functor Cop.comp S F + - yoneda_functor_fixed_object C S \ a + yoneda_functor_fixed_object C S setp \ a + + F: set_valued_functor Cop.comp S setp F for C :: "'c comp" (infixr "\" 55) and S :: "'s comp" (infixr "\\<^sub>S" 55) + and setp :: "'s set \ bool" and \ :: "'c * 'c \ 'c \ 's" and F :: "'c \ 's" and a :: 'c begin text\ The mapping that evaluates the component @{term "\ a"} at @{term a} of a natural transformation @{term \} from @{term Y} to @{term F} on the element @{term "\ (a, a) a"} of @{term "SET a"}, yielding an element of @{term "F.SET a"}. \ definition \ :: "('c \ 's) \ 's" where "\ \ = S.Fun (\ a) (\ (a, a) a)" text\ The mapping that takes an element @{term e} of @{term "F.SET a"} and produces a map on objects of @{term[source=true] C} whose value at @{term b} is the arrow of @{term[source=true] S} corresponding to the function @{term "(\x. F.FUN (\ (b, a) x) e) \ Hom.set (b, a) \ F.SET b"}. \ - definition \o :: "'s \ 'c \ 's" - where "\o e b = S.mkArr (Hom.set (b, a)) (F.SET b) (\x. F.FUN (\ (b, a) x) e)" + definition \\<^sub>o :: "'s \ 'c \ 's" + where "\\<^sub>o e b = S.mkArr (Hom.set (b, a)) (F.SET b) (\x. F.FUN (\ (b, a) x) e)" - lemma \o_e_ide: + lemma \\<^sub>o_in_hom: assumes e: "e \ S.set (F a)" and b: "C.ide b" - shows "\\o e b : Y a b \\<^sub>S F b\" - and "\o e b = S.mkArr (Hom.set (b, a)) (F.SET b) (\x. F.FUN (\ (b, a) x) e)" + shows "\\\<^sub>o e b : Y a b \\<^sub>S F b\" proof - - show "\o e b = S.mkArr (Hom.set (b, a)) (F.SET b) (\x. F.FUN (\ (b, a) x) e)" - using \o_def by auto - moreover have "(\x. F.FUN (\ (b, a) x) e) \ Hom.set (b, a) \ F.SET b" + have "(\x. F.FUN (\ (b, a) x) e) \ Hom.set (b, a) \ F.SET b" proof fix x assume x: "x \ Hom.set (b, a)" thus "F.FUN (\ (b, a) x) e \ F.SET b" using assms e ide_a Hom.\_mapsto S.Fun_mapsto [of "F (\ (b, a) x)"] by force qed - ultimately show "\\o e b : Y a b \\<^sub>S F b\" + thus ?thesis using ide_a b S.mkArr_in_hom [of "Hom.set (b, a)" "F.SET b"] Hom.set_subset_Univ - S.mkIde_set - by force + S.mkIde_set \\<^sub>o_def + by (metis C.ideD(1) Cop.ide_char F.preserves_ide Hom.set_map S.setp_set_ide + preserves_ide Y_simp) qed text\ - For each @{term "e \ F.SET a"}, the mapping @{term "\o e"} gives the components + For each @{term "e \ F.SET a"}, the mapping @{term "\\<^sub>o e"} gives the components of a natural transformation @{term \} from @{term "Y a"} to @{term F}. \ - lemma \o_e_induces_transformation: + lemma \\<^sub>o_induces_transformation: assumes e: "e \ S.set (F a)" - shows "transformation_by_components Cop.comp S (Y a) F (\o e)" + shows "transformation_by_components Cop.comp S (Y a) F (\\<^sub>o e)" proof fix b :: 'c assume b: "Cop.ide b" - show "\\o e b : Y a b \\<^sub>S F b\" - using ide_a b e \o_e_ide by simp + show "\\\<^sub>o e b : Y a b \\<^sub>S F b\" + using ide_a b e \\<^sub>o_in_hom by simp next fix g :: 'c assume g: "Cop.arr g" let ?b = "Cop.dom g" let ?b' = "Cop.cod g" - show "\o e (Cop.cod g) \\<^sub>S Y a g = F g \\<^sub>S \o e (Cop.dom g)" + show "\\<^sub>o e (Cop.cod g) \\<^sub>S Y a g = F g \\<^sub>S \\<^sub>o e (Cop.dom g)" proof - - have 1: "\o e (Cop.cod g) \\<^sub>S Y a g + have 1: "\\<^sub>o e (Cop.cod g) \\<^sub>S Y a g = S.mkArr (Hom.set (?b, a)) (F.SET ?b') ((\x. F.FUN (\ (?b', a) x) e) o (\ (?b', a) o Cop.comp g o \ (?b, a)))" proof - have "S.arr (S.mkArr (Hom.set (Cop.cod g, a)) (F.SET (Cop.cod g)) (\s. F.FUN (\ (Cop.cod g, a) s) e)) \ S.dom (S.mkArr (Hom.set (Cop.cod g, a)) (F.SET (Cop.cod g)) (\s. F.FUN (\ (Cop.cod g, a) s) e)) = Y a (Cop.cod g) \ S.cod (S.mkArr (Hom.set (Cop.cod g, a)) (F.SET (Cop.cod g)) (\s. F.FUN (\ (Cop.cod g, a) s) e)) = F (Cop.cod g)" - using Cop.cod_char \o_e_ide [of e ?b'] \o_e_ide [of e ?b'] e g + using Cop.cod_char \\<^sub>o_def \\<^sub>o_in_hom [of e ?b'] e g by (metis Cop.ide_char Cop.ide_cod S.in_homE) moreover have "Y a g = S.mkArr (Hom.set (Cop.dom g, a)) (Hom.set (Cop.cod g, a)) (\ (Cop.cod g, a) \ Cop.comp g \ \ (Cop.dom g, a))" using Y_ide_arr [of a g ?b' ?b] ide_a g by auto ultimately show ?thesis - using ide_a e g Y_ide_arr Cop.cod_char \o_e_ide + using ide_a e g Y_ide_arr Cop.cod_char \\<^sub>o_def S.comp_mkArr [of "Hom.set (?b, a)" "Hom.set (?b', a)" "\ (?b', a) o Cop.comp g o \ (?b, a)" "F.SET ?b'" "\x. F.FUN (\ (?b', a) x) e"] - by (metis C.ide_dom Cop.arr_char preserves_arr) + by (metis preserves_arr) qed also have "... = S.mkArr (Hom.set (?b, a)) (F.SET ?b') (F.FUN g o (\x. F.FUN (\ (?b, a) x) e))" proof (intro S.mkArr_eqI') have "(\x. F.FUN (\ (?b', a) x) e) o (\ (?b', a) o Cop.comp g o \ (?b, a)) \ Hom.set (?b, a) \ F.SET ?b'" proof - - have "S.arr (S (\o e ?b') (Y a g))" - using ide_a e g \o_e_ide [of e ?b'] Y_ide_arr(1) [of a "C.dom g" "C.cod g" g] + have "S.arr (S (\\<^sub>o e ?b') (Y a g))" + using ide_a e g \\<^sub>o_in_hom [of e ?b'] Y_ide_arr(1) [of a "C.dom g" "C.cod g" g] Cop.ide_char Cop.ide_cod by blast - thus ?thesis using 1 S.arr_mkArr by simp + thus ?thesis using 1 by simp qed thus "S.arr (S.mkArr (Hom.set (?b, a)) (F.SET ?b') ((\x. F.FUN (\ (?b', a) x) e) o (\ (?b', a) o Cop.comp g o \ (?b, a))))" - using ide_a e g Hom.set_subset_Univ S.arr_mkArr by force + using ide_a e g Hom.set_subset_Univ S.arr_mkArr + by (metis C.dom_dom C.ide_dom Cop.arr_char Cop.cod_char Cop.comp_cod_arr + Cop.comp_in_homI' Cop.seqE S.seqI' \\<^sub>o_in_hom calculation preserves_hom) show "\x. x \ Hom.set (?b, a) \ ((\x. F.FUN (\ (?b', a) x) e) o (\ (?b', a) o Cop.comp g o \ (?b, a))) x = (F.FUN g o (\x. F.FUN (\ (?b, a) x) e)) x" proof - fix x assume x: "x \ Hom.set (?b, a)" have "((\x. (F.FUN o \ (?b', a)) x e) o (\ (?b', a) o Cop.comp g o \ (?b, a))) x = F.FUN (\ (?b', a) (\ (?b', a) (C (\ (?b, a) x) g))) e" by simp also have "... = (F.FUN g o (F.FUN o \ (?b, a)) x) e" proof - have 1: "\\ (Cop.dom g, a) x : Cop.dom g \ a\" using ide_a x g Hom.\_mapsto [of ?b a] by auto moreover have "S.seq (F g) (F (\ (C.cod g, a) x))" using 1 g by (intro S.seqI', auto) moreover have "\ (C.dom g, a) (\ (C.dom g, a) (C (\ (C.cod g, a) x) g)) = C (\ (C.cod g, a) x) g" using g 1 Hom.\_\ [of "C (\ (?b, a) x) g" ?b' a] by fastforce ultimately show ?thesis using assms F.preserves_comp by fastforce qed also have "... = (F.FUN g o (\x. F.FUN (\ (?b, a) x) e)) x" by fastforce finally show "((\x. F.FUN (\ (?b', a) x) e) o (\ (?b', a) o Cop.comp g o \ (?b, a))) x = (F.FUN g o (\x. F.FUN (\ (?b, a) x) e)) x" by simp qed qed - also have "... = F g \\<^sub>S \o e (Cop.dom g)" + also have "... = F g \\<^sub>S \\<^sub>o e (Cop.dom g)" proof - have "S.arr (F g) \ F g = S.mkArr (F.SET ?b) (F.SET ?b') (F.FUN g)" using g S.mkArr_Fun [of "F g"] by simp moreover have - "S.arr (\o e ?b) \ - \o e ?b = S.mkArr (Hom.set (?b, a)) (F.SET ?b) (\x. F.FUN (\ (?b, a) x) e)" - using e g \o_e_ide + "S.arr (\\<^sub>o e ?b) \ + \\<^sub>o e ?b = S.mkArr (Hom.set (?b, a)) (F.SET ?b) (\x. F.FUN (\ (?b, a) x) e)" + using e g \\<^sub>o_def \\<^sub>o_in_hom by (metis C.ide_cod Cop.arr_char Cop.dom_char S.in_homE) ultimately show ?thesis using S.comp_mkArr [of "Hom.set (?b, a)" "F.SET ?b" "\x. F.FUN (\ (?b, a) x) e" "F.SET ?b'" "F.FUN g"] by metis qed finally show ?thesis by blast qed qed - abbreviation \ :: "'s \ 'c \ 's" - where "\ e \ transformation_by_components.map Cop.comp S (Y a) (\o e)" + definition \ :: "'s \ 'c \ 's" + where "\ e \ transformation_by_components.map Cop.comp S (Y a) (\\<^sub>o e)" end locale yoneda_lemma_fixed_e = - yoneda_lemma C S \ F a - for C :: "'c comp" (infixr "\" 55) - and S :: "'s comp" (infixr "\\<^sub>S" 55) - and \ :: "'c * 'c \ 'c \ 's" - and F :: "'c \ 's" - and a :: 'c - and e :: 's + + yoneda_lemma + + fixes e assumes E: "e \ F.SET a" begin - interpretation \e: transformation_by_components Cop.comp S \Y a\ F \\o e\ - using E \o_e_induces_transformation by auto + interpretation \e: transformation_by_components Cop.comp S \Y a\ F \\\<^sub>o e\ + using E \\<^sub>o_induces_transformation by auto + sublocale \e: natural_transformation Cop.comp S \Y a\ F \\ e\ + unfolding \_def .. lemma natural_transformation_\e: shows "natural_transformation Cop.comp S (Y a) F (\ e)" .. lemma \e_ide: assumes "Cop.ide b" shows "S.arr (\ e b)" and "\ e b = S.mkArr (Hom.set (b, a)) (F.SET b) (\x. F.FUN (\ (b, a) x) e)" using assms apply fastforce - using assms \o_def by auto + using assms \\<^sub>o_def \_def by auto end locale yoneda_lemma_fixed_\ = - yoneda_lemma C S \ F a + - \: set_valued_transformation Cop.comp S "Y a" F \ - for C :: "'c comp" (infixr "\" 55) - and S :: "'s comp" (infixr "\\<^sub>S" 55) - and \ :: "'c * 'c \ 'c \ 's" - and F :: "'c \ 's" - and a :: 'c - and \ :: "'c \ 's" + yoneda_lemma + + \: natural_transformation Cop.comp S \Y a\ F \ + for \ begin + sublocale \: set_valued_transformation Cop.comp S setp \Y a\ F \ .. + text\ The key lemma: The component @{term "\ b"} of @{term \} at an arbitrary object @{term b} is completely determined by the single element @{term "\.FUN a (\ (a, a) a) \ F.SET a"}. \ lemma \_ide: assumes b: "Cop.ide b" shows "\ b = S.mkArr (Hom.set (b, a)) (F.SET b) (\x. (F.FUN (\ (b, a) x) (\.FUN a (\ (a, a) a))))" proof - let ?\a = "\ (a, a) a" have \a: "\ (a, a) a \ Hom.set (a, a)" using ide_a Hom.\_mapsto [of a a] by fastforce have 1: "\ b = S.mkArr (Hom.set (b, a)) (F.SET b) (\.FUN b)" using ide_a b S.mkArr_Fun [of "\ b"] Hom.set_map by auto also have "... = S.mkArr (Hom.set (b, a)) (F.SET b) (\x. (F.FUN (\ (b, a) x) (\.FUN a ?\a)))" proof (intro S.mkArr_eqI') show "S.arr (S.mkArr (Hom.set (b, a)) (F.SET b) (\.FUN b))" using ide_a b 1 S.mkArr_Fun [of "\ b"] Hom.set_map by auto show "\x. x \ Hom.set (b, a) \ \.FUN b x = (F.FUN (\ (b, a) x) (\.FUN a ?\a))" proof - fix x assume x: "x \ Hom.set (b, a)" let ?\x = "\ (b, a) x" have \x: "\?\x : b \ a\" using ide_a b x Hom.\_mapsto [of b a] by auto show "\.FUN b x = (F.FUN (\ (b, a) x) (\.FUN a ?\a))" proof - have "\.FUN b x = S.Fun (\ b \\<^sub>S Y a ?\x) ?\a" proof - have "\.FUN b x = \.FUN b ((\ (b, a) o Cop.comp ?\x) a)" using ide_a b x \x Hom.\_\ by (metis C.comp_cod_arr C.in_homE C.ide_dom Cop.comp_def comp_apply) also have "\.FUN b ((\ (b, a) o Cop.comp ?\x) a) = (\.FUN b o (\ (b, a) o Cop.comp ?\x o \ (a, a))) ?\a" using ide_a b C.ide_in_hom by simp also have "... = S.Fun (\ b \\<^sub>S Y a ?\x) ?\a" proof - have "S.arr (Y a ?\x)" using ide_a \x preserves_arr by (elim C.in_homE, auto) moreover have "Y a ?\x = S.mkArr (Hom.set (a, a)) (SET b) (\ (b, a) \ Cop.comp ?\x \ \ (a, a))" using ide_a b \x preserves_hom Y_ide_arr Hom.set_map C.arrI by auto moreover have "S.arr (\ b) \ \ b = S.mkArr (SET b) (F.SET b) (\.FUN b)" using ide_a b S.mkArr_Fun [of "\ b"] by simp ultimately have "S.seq (\ b) (Y a ?\x) \ \ b \\<^sub>S Y a ?\x = S.mkArr (Hom.set (a, a)) (F.SET b) (\.FUN b o (\ (b, a) \ Cop.comp ?\x \ \ (a, a)))" using 1 S.comp_mkArr S.seqI by (metis S.cod_mkArr S.dom_mkArr) thus ?thesis using ide_a b x Hom.\_mapsto S.Fun_mkArr by force qed finally show ?thesis by auto qed also have "... = S.Fun (F ?\x \\<^sub>S \ a) ?\a" using ide_a b \x \.naturality [of ?\x] by force also have "... = F.FUN ?\x (\.FUN a ?\a)" proof - have "restrict (S.Fun (F ?\x \\<^sub>S \ a)) (Hom.set (a, a)) = restrict (F.FUN (\ (b, a) x) o \.FUN a) (Hom.set (a, a))" proof - have "S.arr (F ?\x \\<^sub>S \ a) \ F ?\x \\<^sub>S \ a = S.mkArr (Hom.set (a, a)) (F.SET b) (F.FUN ?\x o \.FUN a)" proof show 1: "S.seq (F ?\x) (\ a)" using \x ide_a \.preserves_cod F.preserves_dom by (elim C.in_homE, auto) have "\ a = S.mkArr (Hom.set (a, a)) (F.SET a) (\.FUN a)" using ide_a 1 S.mkArr_Fun [of "\ a"] Hom.set_map by auto moreover have "F ?\x = S.mkArr (F.SET a) (F.SET b) (F.FUN ?\x)" using x \x 1 S.mkArr_Fun [of "F ?\x"] by fastforce ultimately show "F ?\x \\<^sub>S \ a = S.mkArr (Hom.set (a, a)) (F.SET b) (F.FUN ?\x o \.FUN a)" using 1 S.comp_mkArr [of "Hom.set (a, a)" "F.SET a" "\.FUN a" "F.SET b" "F.FUN ?\x"] by (elim S.seqE, auto) qed thus ?thesis by force qed thus "S.Fun (F (\ (b, a) x) \\<^sub>S \ a) ?\a = F.FUN ?\x (\.FUN a ?\a)" using ide_a \a restr_eqE [of "S.Fun (F ?\x \\<^sub>S \ a)" "Hom.set (a, a)" "F.FUN ?\x o \.FUN a"] by simp qed finally show ?thesis by simp qed qed qed finally show ?thesis by auto qed text\ Consequently, if @{term \'} is any natural transformation from @{term "Y a"} to @{term F} that agrees with @{term \} at @{term a}, then @{term "\' = \"}. \ lemma eqI: assumes "natural_transformation Cop.comp S (Y a) F \'" and "\' a = \ a" shows "\' = \" proof (intro NaturalTransformation.eqI) interpret \': natural_transformation Cop.comp S \Y a\ F \' using assms by auto - interpret T': yoneda_lemma_fixed_\ C S \ F a \' .. + interpret T': yoneda_lemma_fixed_\ C S setp \ F a \' .. show "natural_transformation Cop.comp S (Y a) F \" .. show "natural_transformation Cop.comp S (Y a) F \'" .. show "\b. Cop.ide b \ \' b = \ b" using assms(2) \_ide T'.\_ide by simp qed end context yoneda_lemma begin text\ One half of the Yoneda lemma: The mapping @{term \} is an injection, with left inverse @{term \}, from the set @{term "F.SET a"} to the set of natural transformations from @{term "Y a"} to @{term F}. \ lemma \_is_injection: assumes "e \ F.SET a" shows "natural_transformation Cop.comp S (Y a) F (\ e)" and "\ (\ e) = e" proof - - interpret yoneda_lemma_fixed_e C S \ F a e + interpret yoneda_lemma_fixed_e C S setp \ F a e using assms by (unfold_locales, auto) - interpret \e: natural_transformation Cop.comp S \Y a\ F \\ e\ - using natural_transformation_\e by auto show "natural_transformation Cop.comp S (Y a) F (\ e)" .. show "\ (\ e) = e" unfolding \_def using assms \e_ide S.Fun_mkArr Hom.\_mapsto Hom.\_\ ide_a F.preserves_ide S.Fun_ide restrict_apply C.ide_in_hom by (auto simp add: Pi_iff) qed - lemma \\_in_Fa: + lemma \\_mapsto: assumes "natural_transformation Cop.comp S (Y a) F \" shows "\ \ \ F.SET a" proof - - interpret \: natural_transformation Cop.comp S \Y a\ F \ using assms by auto - interpret yoneda_lemma_fixed_\ C S \ F a \ .. + interpret \: natural_transformation Cop.comp S \Y a\ F \ + using assms by auto + interpret yoneda_lemma_fixed_\ C S setp \ F a \ .. show ?thesis proof (unfold \_def) have "S.arr (\ a) \ S.Dom (\ a) = Hom.set (a, a) \ S.Cod (\ a) = F.SET a" using ide_a Hom.set_map by auto hence "\.FUN a \ Hom.set (a, a) \ F.SET a" using S.Fun_mapsto by blast thus "\.FUN a (\ (a, a) a) \ F.SET a" using ide_a Hom.\_mapsto by fastforce qed qed text\ The other half of the Yoneda lemma: The mapping @{term \} is a surjection, with right inverse @{term \}, taking natural transformations from @{term "Y a"} to @{term F} to elements of @{term "F.SET a"}. \ lemma \_is_surjection: assumes "natural_transformation Cop.comp S (Y a) F \" - shows "\ \ \ F.SET a" and "\ (\ \) = \" + shows "\ (\ \) = \" proof - - interpret natural_transformation Cop.comp S \Y a\ F \ using assms by auto - interpret yoneda_lemma_fixed_\ C S \ F a \ .. - show 1: "\ \ \ F.SET a" using assms \\_in_Fa by auto - interpret yoneda_lemma_fixed_e C S \ F a \\ \\ - using 1 by (unfold_locales, auto) - interpret \e: natural_transformation Cop.comp S \Y a\ F \\ (\ \)\ - using natural_transformation_\e by auto + interpret natural_transformation Cop.comp S \Y a\ F \ + using assms by auto + interpret yoneda_lemma_fixed_\ C S setp \ F a \ .. + interpret yoneda_lemma_fixed_e C S setp \ F a \\ \\ + using assms \\_mapsto by unfold_locales auto show "\ (\ \) = \" proof (intro eqI) show "natural_transformation Cop.comp S (Y a) F (\ (\ \))" .. show "\ (\ \) a = \ a" using ide_a \_ide [of a] \e_ide \_def by simp qed qed text\ The main result. \ theorem yoneda_lemma: shows "bij_betw \ (F.SET a) {\. natural_transformation Cop.comp S (Y a) F \}" - using \_is_injection \_is_surjection by (intro bij_betwI, auto) + using \\_mapsto \_is_injection \_is_surjection + by (intro bij_betwI, auto) end text\ We now consider the special case in which @{term F} is the contravariant functor @{term "Y a'"}. Then for any @{term e} in \Hom.set (a, a')\ we have @{term "\ e = Y (\ (a, a') e)"}, and @{term \} is a bijection from \Hom.set (a, a')\ to the set of natural transformations from @{term "Y a"} to @{term "Y a'"}. It then follows that that the Yoneda functor @{term Y} is a fully faithful functor from @{term C} to the functor category \[Cop, S]\. \ locale yoneda_lemma_for_hom = - C: category C + - Cop: dual_category C + - S: replete_set_category S + - yoneda_functor_fixed_object C S \ a + - Ya': yoneda_functor_fixed_object C S \ a' + - yoneda_lemma C S \ "Y a'" a + yoneda_functor_fixed_object C S setp \ a + + Ya': yoneda_functor_fixed_object C S setp \ a' + + yoneda_lemma C S setp \ "Y a'" a for C :: "'c comp" (infixr "\" 55) and S :: "'s comp" (infixr "\\<^sub>S" 55) + and setp :: "'s set \ bool" and \ :: "'c * 'c \ 'c \ 's" - and F :: "'c \ 's" and a :: 'c and a' :: 'c + assumes ide_a': "C.ide a'" begin text\ In case @{term F} is the functor @{term "Y a'"}, for any @{term "e \ Hom.set (a, a')"} the induced natural transformation @{term "\ e"} from @{term "Y a"} to @{term "Y a'"} is just @{term "Y (\ (a, a') e)"}. \ - lemma \_equals_Yo\: + lemma app_\_equals: assumes e: "e \ Hom.set (a, a')" shows "\ e = Y (\ (a, a') e)" proof - let ?\e = "\ (a, a') e" have \e: "\?\e : a \ a'\" using ide_a ide_a' e Hom.\_mapsto [of a a'] by auto interpret Ye: natural_transformation Cop.comp S \Y a\ \Y a'\ \Y ?\e\ using Y_arr_is_transformation [of ?\e] \e by (elim C.in_homE, auto) - interpret yoneda_lemma_fixed_e C S \ \Y a'\ a e - using ide_a ide_a' e S.set_mkIde Hom.set_map + interpret yoneda_lemma_fixed_e C S setp \ \Y a'\ a e + using ide_a ide_a' e Hom.set_map by (unfold_locales, simp_all) - interpret \e: natural_transformation Cop.comp S \Y a\ \Y a'\ \\ e\ - using natural_transformation_\e by auto - interpret yoneda_lemma_fixed_\ C S \ \Y a'\ a \\ e\ .. + interpret yoneda_lemma_fixed_\ C S setp \ \Y a'\ a \\ e\ .. have "natural_transformation Cop.comp S (Y a) (Y a') (Y ?\e)" .. moreover have "natural_transformation Cop.comp S (Y a) (Y a') (\ e)" .. moreover have "\ e a = Y ?\e a" proof - have 1: "S.arr (\ e a)" using ide_a e \e.preserves_reflects_arr by simp have 2: "\ e a = S.mkArr (Hom.set (a, a)) (Ya'.SET a) (\x. Ya'.FUN (\ (a, a) x) e)" - using ide_a \o_def \e_ide by simp + using ide_a \\<^sub>o_def \e_ide by simp also have "... = S.mkArr (Hom.set (a, a)) (Hom.set (a, a')) (\ (a, a') o C ?\e o \ (a, a))" proof (intro S.mkArr_eqI) show "S.arr (S.mkArr (Hom.set (a, a)) (Ya'.SET a) (\x. Ya'.FUN (\ (a, a) x) e))" using ide_a e 1 2 by simp show "Hom.set (a, a) = Hom.set (a, a)" .. show 3: "Ya'.SET a = Hom.set (a, a')" using ide_a ide_a' Y_simp Hom.set_map by simp show "\x. x \ Hom.set (a, a) \ Ya'.FUN (\ (a, a) x) e = (\ (a, a') o C ?\e o \ (a, a)) x" proof - fix x assume x: "x \ Hom.set (a, a)" have \x: "\\ (a, a) x : a \ a\" using ide_a x Hom.\_mapsto [of a a] by auto have "S.arr (Y a' (\ (a, a) x)) \ Y a' (\ (a, a) x) = S.mkArr (Hom.set (a, a')) (Hom.set (a, a')) (\ (a, a') \ Cop.comp (\ (a, a) x) \ \ (a, a'))" using Y_ide_arr ide_a ide_a' \x by blast hence "Ya'.FUN (\ (a, a) x) e = (\ (a, a') \ Cop.comp (\ (a, a) x) \ \ (a, a')) e" using e 3 S.Fun_mkArr Ya'.preserves_reflects_arr [of "\ (a, a) x"] by simp also have "... = (\ (a, a') o C ?\e o \ (a, a)) x" by simp finally show "Ya'.FUN (\ (a, a) x) e = (\ (a, a') o C ?\e o \ (a, a)) x" by auto qed qed also have "... = Y ?\e a" using ide_a ide_a' Y_arr_ide \e by simp finally show "\ e a = Y ?\e a" by auto qed ultimately show ?thesis using eqI by auto qed - lemma Y_injective_on_homs: - assumes "\f : a \ a'\" and "\f' : a \ a'\" and "map f = map f'" - shows "f = f'" - proof - - have "f = \ (a, a') (\ (a, a') f)" - using assms ide_a Hom.\_\ by simp - also have "... = \ (a, a') (\ (\ (\ (a, a') f)))" - using ide_a ide_a' assms(1) \_is_injection Hom.\_mapsto Hom.set_map - by (elim C.in_homE, simp add: Pi_iff) - also have "... = \ (a, a') (\ (Y (\ (a, a') (\ (a, a') f))))" - using assms Hom.\_mapsto [of a a'] \_equals_Yo\ [of "\ (a, a') f"] by force - also have "... = \ (a, a') (\ (\ (\ (a, a') f')))" - using assms Hom.\_mapsto [of a a'] ide_a Hom.\_\ Y_def - \_equals_Yo\ [of "\ (a, a') f'"] - by fastforce - also have "... = \ (a, a') (\ (a, a') f')" - using ide_a ide_a' assms(2) \_is_injection Hom.\_mapsto Hom.set_map - by (elim C.in_homE, simp add: Pi_iff) - also have "... = f'" - using assms ide_a Hom.\_\ by simp - finally show "f = f'" by auto + lemma is_injective_on_homs: + shows "inj_on map (C.hom a a')" + proof (intro inj_onI) + fix f f' + assume f: "f \ C.hom a a'" and f': "f' \ C.hom a a'" + assume eq: "map f = map f'" + show "f = f'" + proof - + have "f = \ (a, a') (\ (a, a') f)" + using f ide_a Hom.\_\ by simp + also have "... = \ (a, a') (\ (\ (\ (a, a') f)))" + using ide_a ide_a' f \_is_injection Hom.\_mapsto Hom.set_map + by (metis C.ideD(1) PiE Y_simp mem_Collect_eq) + also have "... = \ (a, a') (\ (Y (\ (a, a') (\ (a, a') f))))" + using f Hom.\_mapsto app_\_equals [of "\ (a, a') f"] by force + also have "... = \ (a, a') (\ (\ (\ (a, a') f')))" + using f f' eq Hom.\_mapsto [of a a'] ide_a Hom.\_\ Y_def + app_\_equals [of "\ (a, a') f'"] + by fastforce + also have "... = \ (a, a') (\ (a, a') f')" + using ide_a ide_a' f' \_is_injection Hom.\_mapsto Hom.set_map + by (metis C.ideD(1) PiE Y_simp mem_Collect_eq) + also have "... = f'" + using f' ide_a Hom.\_\ by simp + finally show "f = f'" by auto + qed qed - lemma Y_surjective_on_homs: - assumes \: "natural_transformation Cop.comp S (Y a) (Y a') \" - shows "Y (\ (a, a') (\ \)) = \" - using ide_a ide_a' \ \_is_surjection \_equals_Yo\ \\_in_Fa Hom.set_map by simp - end context yoneda_functor begin lemma is_faithful_functor: shows "faithful_functor C Cop_S.comp map" proof fix f :: 'c and f' :: 'c assume par: "C.par f f'" and ff': "map f = map f'" show "f = f'" proof - - interpret Ya': yoneda_functor_fixed_object C S \ \C.cod f\ + interpret Ya': yoneda_functor_fixed_object C S setp \ \C.cod f\ using par by (unfold_locales, auto) - interpret yoneda_lemma_for_hom C S \ \Y (C.cod f)\ \C.dom f\ \C.cod f\ + interpret yoneda_lemma_for_hom C S setp \ \C.dom f\ \C.cod f\ using par by (unfold_locales, auto) - show "f = f'" using par ff' Y_injective_on_homs [of f f'] by fastforce + show "f = f'" + using par ff' is_injective_on_homs inj_on_def [of map "C.hom (C.dom f) (C.cod f)"] + by force qed qed lemma is_full_functor: shows "full_functor C Cop_S.comp map" proof fix a :: 'c and a' :: 'c and t assume a: "C.ide a" and a': "C.ide a'" assume t: "\t : map a \\<^sub>[\<^sub>C\<^sub>o\<^sub>p\<^sub>,\<^sub>S\<^sub>] map a'\" show "\e. \e : a \ a'\ \ map e = t" proof - interpret Ya': yoneda_functor_fixed_object C S \ a' + interpret Ya': yoneda_functor_fixed_object C S setp \ a' using a' by (unfold_locales, auto) - interpret yoneda_lemma_for_hom C S \ \Y a'\ a a' + interpret yoneda_lemma_for_hom C S setp \ a a' using a a' by (unfold_locales, auto) have NT: "natural_transformation Cop.comp S (Y a) (Y a') (Cop_S.Map t)" using t a' Y_def Cop_S.Map_dom Cop_S.Map_cod Cop_S.dom_char Cop_S.cod_char Cop_S.in_homE Cop_S.arrE by metis hence 1: "\ (Cop_S.Map t) \ Hom.set (a, a')" - using \\_in_Fa ide_a ide_a' Hom.set_map by simp + using \\_mapsto ide_a ide_a' Hom.set_map by simp moreover have "map (\ (a, a') (\ (Cop_S.Map t))) = t" proof (intro Cop_S.arr_eqI) have 2: "\map (\ (a, a') (\ (Cop_S.Map t))) : map a \\<^sub>[\<^sub>C\<^sub>o\<^sub>p\<^sub>,\<^sub>S\<^sub>] map a'\" using 1 ide_a ide_a' Hom.\_mapsto [of a a'] by blast show "Cop_S.arr t" using t by blast show "Cop_S.arr (map (\ (a, a') (\ (Cop_S.Map t))))" using 2 by blast show 3: "Cop_S.Map (map (\ (a, a') (\ (Cop_S.Map t)))) = Cop_S.Map t" - using NT Y_surjective_on_homs Y_def by simp + using NT 1 Y_def \_is_surjection app_\_equals \\_mapsto by metis show 4: "Cop_S.Dom (map (\ (a, a') (\ (Cop_S.Map t)))) = Cop_S.Dom t" using t 2 functor_axioms Cop_S.Map_dom by (metis Cop_S.in_homE) show "Cop_S.Cod (map (\ (a, a') (\ (Cop_S.Map t)))) = Cop_S.Cod t" using 2 3 4 t Cop_S.Map_cod by (metis Cop_S.in_homE) qed ultimately show "\\ (a, a') (\ (Cop_S.Map t)) : a \ a'\ \ map (\ (a, a') (\ (Cop_S.Map t))) = t" using ide_a ide_a' Hom.\_mapsto by auto qed qed end sublocale yoneda_functor \ faithful_functor C Cop_S.comp map using is_faithful_functor by auto sublocale yoneda_functor \ full_functor C Cop_S.comp map using is_full_functor by auto sublocale yoneda_functor \ fully_faithful_functor C Cop_S.comp map .. end diff --git a/thys/Category3/ZFC_SetCat.thy b/thys/Category3/ZFC_SetCat.thy new file mode 100644 --- /dev/null +++ b/thys/Category3/ZFC_SetCat.thy @@ -0,0 +1,1012 @@ +(* Title: ZFC_SetCat + Author: Eugene W. Stark , 2022 + Maintainer: Eugene W. Stark +*) + +chapter "ZFC SetCat" + +text\ + In the statement and proof of the Yoneda Lemma given in theory \Yoneda\, + we sidestepped the issue, of not having a category of ``all'' sets, by axiomatizing + the notion of a ``set category'', showing that for every category we could obtain + a hom-functor into a set category at a higher type, and then proving the Yoneda + lemma for that particular hom-functor. This is perhaps the best we can do within HOL, + because HOL does not provide any type that contains a universe of sets with the closure + properties usually associated with a category \Set\ of sets and functions between them. + However, a significant aspect of category theory involves considering ``all'' + algebraic structures of a particular kind as the objects of a ``large'' category having + nice closure or completeness properties. Being able to consider a category of sets that + is ``small-complete'', or a cartesian closed category of sets and functions that includes + some infinite sets as objects, are basic examples of this kind of situation. + + The purpose of this section is to demonstrate that, although it cannot be done in + pure HOL, if we are willing to accept the existence of a type \V\ whose inhabitants + correspond to sets satisfying the axioms of ZFC, then it is possible to construct, + for example, the ``large'' category of sets and functions as it is usually understood + in category theory. Moreover, assuming the existence of such a type is essentially + all we have to do; all the category theory we have developed so far still applies. + Specifically, what we do in this section is to use theory \ZFC_in_HOL\, which provides + an axiomatization of a set-theoretic universe \V\, to construct a ``set category'' + \ZFC_SetCat\, whose objects correspond to \V\-sets, whose arrows correspond to functions + between \V\-sets, and which has the small-completeness property traditionally ascribed + to the category of all small sets and functions between them. +\ + +theory ZFC_SetCat +imports ZFC_in_HOL.ZFC_Cardinals Limit +begin + + text\ + The following locale constructs the category of classes and functions between them + and shows that it is small complete. The category is obtained simply as the replete + set category at type \V\. This is not yet the category of sets we want, because it + contains objects corresponding to ``large'' \V\-sets. + \ + + locale ZFC_class_cat + begin + + sublocale replete_setcat \undefined :: V\ . + + lemma admits_small_V_tupling: + assumes "small (I :: V set)" + shows "admits_tupling I" + proof (unfold admits_tupling_def) + let ?\ = "\f. UP (VLambda (ZFC_in_HOL.set I) (DN o f))" + have "?\ \ (I \ Univ) \ extensional' I \ Univ" + using UP_mapsto by force + moreover have "inj_on ?\ ((I \ Univ) \ extensional' I)" + proof + fix f g + assume f: "f \ (I \ Univ) \ extensional' I" + assume g: "g \ (I \ Univ) \ extensional' I" + assume eq: "UP (VLambda (ZFC_in_HOL.set I) (DN \ f)) = + UP (VLambda (ZFC_in_HOL.set I) (DN \ g))" + have 1: "VLambda (ZFC_in_HOL.set I) (DN \ f) = VLambda (ZFC_in_HOL.set I) (DN \ g)" + using f g eq + by (meson injD inj_UP) + show "f = g" + proof + fix x + have "x \ I \ f x = g x" + using f g extensional'_def [of I] by auto + moreover have "x \ I \ f x = g x" + proof - + assume x: "x \ I" + hence "(DN o f) x = (DN o g) x" + using assms 1 x elts_of_set VLambda_eq_D2 by fastforce + thus "f x = g x" + using f g x comp_apply UP_DN + by (metis IntD1 PiE bij_arr_of bij_betw_imp_surj_on) + qed + ultimately show "f x = g x" by blast + qed + qed + ultimately show "\\. \ \ (I \ Univ) \ extensional' I \ Univ \ + inj_on \ ((I \ Univ) \ extensional' I)" + by blast + qed + + corollary admits_small_tupling: + assumes "small I" + shows "admits_tupling I" + proof - + obtain \ where \: "inj_on \ I \ \ ` I \ range elts" + using assms small_def by metis + have "admits_tupling (\ ` I)" + using \ admits_small_V_tupling by fastforce + moreover have inv: "bij_betw (inv_into I \) (\ ` I) I" + by (simp add: \ bij_betw_inv_into inj_on_imp_bij_betw) + ultimately show ?thesis + using admits_tupling_respects_bij by blast + qed + + lemma has_small_products: + assumes "small (I :: 'i set)" and "I \ UNIV" + shows "has_products I" + proof - + have 1: "\I :: V set. small I \ has_products I" + using big_UNIV + by (metis admits_small_tupling has_products_iff_admits_tupling) + obtain V_of where V_of: "inj_on V_of I \ V_of ` I \ range elts" + using assms small_def by metis + have "bij_betw (inv_into I V_of) (V_of ` I) I" + using V_of bij_betw_inv_into bij_betw_imageI by metis + moreover have "small (V_of ` I)" + using assms by auto + ultimately show ?thesis + using assms 1 has_products_preserved_by_bijection by blast + qed + + theorem has_small_limits: + assumes "small (UNIV :: 'i set)" + shows "has_limits (undefined :: 'i)" + proof - + have "\I :: 'i set. I \ UNIV \ admits_tupling I" + using assms smaller_than_small subset_UNIV admits_small_tupling by auto + thus ?thesis + using assms has_limits_iff_admits_tupling by blast + qed + + end + + text\ + We now construct the desired category of small sets and functions between them, + as a full subcategory of the category of classes and functions. To show that this + subcategory is small complete, we show that the inclusion creates small products; + that is, a small product of objects corresponding to small sets itself corresponds + to a small set. + \ + + locale ZFC_set_cat + begin + + interpretation Cls: ZFC_class_cat . + + sublocale sub_set_category Cls.comp \\A. A \ Cls.Univ\ \\A. A \ Cls.Univ \ small A\ + using small_Un smaller_than_small + apply unfold_locales + apply simp_all + apply force + by auto + + abbreviation setp + where "setp A \ A \ Cls.Univ \ small A" + + lemma is_set_category: + shows "set_category comp setp" + .. + + lemma is_full_subcategory: + shows "full_subcategory Cls.comp (\a. Cls.ide a \ setp (Cls.set a))" + using full_subcategory_axioms by blast + + interpretation incl: full_inclusion_functor Cls.comp \\a. Cls.ide a \ setp (Cls.set a)\ + .. + + text\ + The following functions establish a bijection between the identities of the category + and the elements of type \V\; which in turn are in bijective correspondence with + small \V\-sets. + \ + + definition V_of_ide :: "V setcat.arr \ V" + where "V_of_ide a \ ZFC_in_HOL.set (Cls.DN ` Cls.set a)" + + definition ide_of_V :: "V \ V setcat.arr" + where "ide_of_V A \ Cls.mkIde (Cls.UP ` elts A)" + + lemma bij_betw_ide_V: + shows "V_of_ide \ Collect ide \ UNIV" + and "ide_of_V \ UNIV \ Collect ide" + and [simp]: "ide a \ ide_of_V (V_of_ide a) = a" + and [simp]: "V_of_ide (ide_of_V A) = A" + and "bij_betw V_of_ide (Collect ide) UNIV" + and "bij_betw ide_of_V UNIV (Collect ide)" + proof - + have "Univ = Cls.Univ" + using terminal_char by presburger + show 1: "V_of_ide \ Collect ide \ UNIV" + by blast + show 2: "ide_of_V \ UNIV \ Collect ide" + proof + fix A :: V + have "small (elts A)" + by blast + have "Cls.UP ` elts A \ Univ \ small (Cls.UP ` elts A)" + using Cls.UP_mapsto terminal_char by blast + hence "ide (mkIde (Cls.UP ` elts A))" + using ide_mkIde \Univ = Cls.Univ\ by auto + thus "ide_of_V A \ Collect ide" + using ide_of_V_def ide_char + by (metis (no_types, lifting) Cls.ide_mkIde Cls.set_mkIde arr_mkIde ide_char' + mem_Collect_eq) + qed + show 3: "\a. ide a \ ide_of_V (V_of_ide a) = a" + proof - + fix a + assume a: "ide a" + have "ide_of_V (V_of_ide a) = + Cls.mkIde (Cls.UP ` elts (ZFC_in_HOL.set (Cls.DN ` Cls.set a)))" + unfolding ide_of_V_def V_of_ide_def by simp + also have "... = Cls.mkIde (Cls.UP ` Cls.DN ` Cls.set a)" + using setp_set_ide a ide_char by force + also have "... = Cls.mkIde (Cls.set a)" + proof - + have "Cls.set a \ Cls.Univ" + using a ide_char by blast + hence "Cls.UP ` Cls.DN ` Cls.set a = Cls.set a" + proof - + have "\x. x \ Cls.set a \ x \ Cls.UP ` Cls.DN ` Cls.set a" + using Cls.UP_DN \Cls.set a \ Cls.Univ\ + by (metis Cls.bij_arr_of bij_betw_def image_inv_into_cancel) + moreover have "\x. x \ Cls.UP ` Cls.DN ` Cls.set a \ x \ Cls.set a" + using \Cls.set a \ Cls.Univ\ + by (metis Cls.bij_arr_of bij_betw_def image_inv_into_cancel) + ultimately show ?thesis by blast + qed + thus ?thesis by argo + qed + also have "... = a" + using a Cls.mkIde_set ide_char by blast + finally show "ide_of_V (V_of_ide a) = a" by simp + qed + show 4: "\A. V_of_ide (ide_of_V A) = A" + proof - + fix A + have "V_of_ide (ide_of_V A) = + ZFC_in_HOL.set (Cls.DN ` Cls.set (Cls.mkIde (Cls.UP ` elts A)))" + unfolding V_of_ide_def ide_of_V_def by simp + also have "... = ZFC_in_HOL.set (Cls.DN ` Cls.UP ` elts A)" + using Cls.set_mkIde [of "Cls.UP ` elts A"] Cls.UP_mapsto by fastforce + also have "... = ZFC_in_HOL.set (elts A)" + using Cls.DN_UP by force + also have "... = A" by simp + finally show "V_of_ide (ide_of_V A) = A" by blast + qed + show "bij_betw V_of_ide (Collect ide) UNIV" + using 1 2 3 4 + by (intro bij_betwI) auto + show "bij_betw ide_of_V UNIV (Collect ide)" + using 1 2 3 4 + by (intro bij_betwI) blast+ + qed + + text\ + Next, we establish bijections between the hom-sets of the category and certain + subsets of \V\ whose elements represent functions. + \ + + definition V_of_arr :: "V setcat.arr \ V" + where "V_of_arr f \ VLambda (V_of_ide (dom f)) (Cls.DN o Cls.Fun f o Cls.UP)" + + definition arr_of_V :: "V setcat.arr \ V setcat.arr \ V \ V setcat.arr" + where "arr_of_V a b F \ Cls.mkArr (Cls.set a) (Cls.set b) (Cls.UP o app F o Cls.DN)" + + definition vfun + where "vfun A B f \ f \ elts (VPow (vtimes A B)) \ elts A = Domain (pairs f) \ + single_valued (pairs f)" + + lemma small_Collect_vfun: + shows "small (Collect (vfun A B))" + unfolding vfun_def + by (metis (full_types) small_Collect small_elts) + + lemma vfunI: + assumes "f \ elts A \ elts B" + shows "vfun A B (VLambda A f)" + proof (unfold vfun_def, intro conjI) + show "VLambda A f \ elts (VPow (vtimes A B))" + using assms VLambda_def by auto + show "elts A = Domain (pairs (VLambda A f))" + using assms VLambda_def [of A f] + by (metis Domain_fst fst_pairs_VLambda) + show "single_valued (pairs (VLambda A f))" + using assms VLambda_def single_valued_def pairs_iff_elts by fastforce + qed + + lemma app_vfun_mapsto: + assumes "vfun A B F" + shows "app F \ elts A \ elts B" + proof + have "F \ elts (VPow (vtimes A B)) \ elts A = Domain (pairs F) \ single_valued (pairs F)" + using assms vfun_def by simp + hence 1: "F \ elts (VPi A (\_. B)) \ elts A = Domain (pairs F) \ single_valued (pairs F)" + unfolding VPi_def + by (metis (no_types, lifting) down elts_of_set mem_Collect_eq subsetI) + fix x + assume x: "x \ elts A" + have "x \ Domain (pairs F)" + using assms x vfun_def by blast + thus "app F x \ elts B" + using x 1 VPi_D [of F A "\_. B" x] by blast + qed + + lemma bij_betw_hom_vfun: + shows "V_of_arr \ hom a b \ Collect (vfun (V_of_ide a) (V_of_ide b))" + and "\ide a; ide b\ \ arr_of_V a b \ Collect (vfun (V_of_ide a) (V_of_ide b)) \ hom a b" + and "f \ hom a b \ arr_of_V a b (V_of_arr f) = f" + and "\ide a; ide b; F \ Collect (vfun (V_of_ide a) (V_of_ide b))\ + \ V_of_arr (arr_of_V a b F) = F" + and "\ide a; ide b\ + \ bij_betw V_of_arr (hom a b) (Collect (vfun (V_of_ide a) (V_of_ide b)))" + and "\ide a; ide b\ + \ bij_betw (arr_of_V a b) (Collect (vfun (V_of_ide a) (V_of_ide b))) (hom a b)" + proof - + show 1: "\a b. V_of_arr \ hom a b \ Collect (vfun (V_of_ide a) (V_of_ide b))" + proof + fix a b f + assume f: "f \ hom a b" + have "V_of_arr f = VLambda (V_of_ide (dom f)) (Cls.DN \ Cls.Fun f \ Cls.UP)" + unfolding V_of_arr_def by simp + moreover have "vfun (V_of_ide a) (V_of_ide b) ..." + proof - + have "Cls.DN \ Cls.Fun f \ Cls.UP \ elts (V_of_ide a) \ elts (V_of_ide b)" + proof + fix x + assume x: "x \ elts (V_of_ide a)" + have "(Cls.DN \ Cls.Fun f \ Cls.UP) x = Cls.DN (Cls.Fun f (Cls.UP x))" + by simp + moreover have "... \ elts (V_of_ide b)" + proof - + have "Cls.UP x \ Cls.Dom f" + by (metis (no_types, lifting) Cls.arr_dom_iff_arr Cls.arr_mkIde Cls.set_mkIde + bij_betw_ide_V(3) arr_char dom_simp f ide_char ide_of_V_def image_eqI + in_homE mem_Collect_eq x) + hence "Cls.DN (Cls.Fun f (Cls.UP x)) \ Cls.DN ` Cls.Cod f" + using f in_hom_char arr_char Cls.Fun_mapsto [of f] by blast + thus "Cls.DN (Cls.Fun f (Cls.UP x)) \ elts (V_of_ide b)" + unfolding V_of_ide_def + using f ide_char in_hom_char + by (metis (no_types, lifting) arr_char cod_simp elts_of_set in_homE + mem_Collect_eq replacement) + qed + ultimately show "(Cls.DN \ Cls.Fun f \ Cls.UP) x \ elts (V_of_ide b)" by argo + qed + thus ?thesis + using f vfunI by blast + qed + ultimately show "V_of_arr f \ Collect (vfun (V_of_ide a) (V_of_ide b))" by simp + qed + show 2: "\a b. \ide a; ide b\ + \ arr_of_V a b \ Collect (vfun (V_of_ide a) (V_of_ide b)) \ hom a b" + proof - + fix a b + assume a: "ide a" and b: "ide b" + show "arr_of_V a b \ Collect (vfun (V_of_ide a) (V_of_ide b)) \ hom a b" + proof + fix F + assume F: "F \ Collect (vfun (V_of_ide a) (V_of_ide b))" + have 3: "app F \ elts (V_of_ide a) \ elts (V_of_ide b)" + using F app_vfun_mapsto [of "V_of_ide a" "V_of_ide b" F] by blast + have 4: "app F \ Cls.DN ` (Cls.set a) \ Cls.DN ` (Cls.set b)" + using 3 V_of_ide_def a b ide_char by auto + have "arr_of_V a b F = Cls.mkArr (Cls.set a) (Cls.set b) (Cls.UP \ app F \ Cls.DN)" + unfolding arr_of_V_def by simp + moreover have "... \ hom a b" + proof + show "\Cls.mkArr (Cls.set a) (Cls.set b) (Cls.UP \ app F \ Cls.DN) : a \ b\" + proof + have 4: "Cls.arr (Cls.mkArr (Cls.set a) (Cls.set b) (Cls.UP \ app F \ Cls.DN))" + proof - + have "Cls.set a \ Cls.Univ \ Cls.set b \ Cls.Univ" + using a b ide_char by blast + moreover have "Cls.UP \ app F \ Cls.DN \ Cls.set a \ Cls.set b" + proof + fix x + assume x: "x \ Cls.set a" + have "(Cls.UP \ app F \ Cls.DN) x = Cls.UP (app F (Cls.DN x))" + by simp + moreover have "... \ Cls.set b" + by (metis (no_types, lifting) 4 Cls.arr_mkIde Cls.ide_char' Cls.set_mkIde + PiE V_of_ide_def bij_betw_ide_V(3) b elts_of_set ide_char + ide_of_V_def replacement rev_image_eqI x) + ultimately show "(Cls.UP \ app F \ Cls.DN) x \ Cls.set b" + by auto + qed + ultimately show ?thesis by blast + qed + show 5: "arr (Cls.mkArr (Cls.set a) (Cls.set b) (Cls.UP \ app F \ Cls.DN))" + using a b 4 arr_char ide_char by auto + show "dom (Cls.mkArr (Cls.set a) (Cls.set b) (Cls.UP \ app F \ Cls.DN)) = a" + using a 4 5 dom_char ide_char by auto + show "cod (Cls.mkArr (Cls.set a) (Cls.set b) (Cls.UP \ app F \ Cls.DN)) = b" + using b 4 5 cod_char ide_char by auto + qed + qed + ultimately show "arr_of_V a b F \ hom a b" by auto + qed + qed + show 3: "\a b f. f \ hom a b \ arr_of_V a b (V_of_arr f) = f" + proof - + fix a b f + assume f: "f \ hom a b" + have 4: "\x. x \ Cls.set a + \ (Cls.UP \ (Cls.DN \ Cls.Fun f \ Cls.UP) \ Cls.DN) x = Cls.Fun f x" + proof - + fix x + assume x: "x \ Cls.set a" + have "(Cls.UP \ (Cls.DN \ Cls.Fun f \ Cls.UP) \ Cls.DN) x = + Cls.UP (Cls.DN (Cls.Fun f (Cls.UP (Cls.DN x))))" + by simp + also have "... = Cls.UP (Cls.DN (Cls.Fun f x))" + using x Cls.UP_DN + by (metis (no_types, lifting) Cls.elem_set_implies_incl_in Cls.incl_in_def + Cls.setp_set_ide bij_betw_def replete_setcat.bij_arr_of subset_eq) + also have "... = Cls.Fun f x" + proof - + have "x \ Cls.Dom f" + using x f dom_char by fastforce + hence "Cls.Fun f x \ Cls.Cod f" + using x f Cls.Fun_mapsto in_hom_char by blast + hence "Cls.Fun f x \ Cls.Univ" + using f cod_char in_hom_char + by (metis (no_types, lifting) Cls.elem_set_implies_incl_in Cls.incl_in_def + Cls.setp_set_ide subsetD) + thus ?thesis + by (meson bij_betw_inv_into_right replete_setcat.bij_arr_of) + qed + finally show "(Cls.UP \ (Cls.DN \ Cls.Fun f \ Cls.UP) \ Cls.DN) x = Cls.Fun f x" + by blast + qed + have 5: "Cls.arr (Cls.mkArr (Cls.set a) (Cls.set b) + (Cls.UP \ (Cls.DN \ Cls.Fun f \ Cls.UP) \ Cls.DN))" + proof - + have "Cls.set a \ Cls.Univ \ Cls.set b \ Cls.Univ" + using f ide_char codomains_def domains_def in_hom_def by force + moreover have "Cls.UP \ (Cls.DN \ Cls.Fun f \ Cls.UP) \ Cls.DN + \ Cls.set a \ Cls.set b" + proof + fix x + assume x: "x \ Cls.set a" + hence 6: "x \ Cls.Dom f" + using f by (metis (no_types, lifting) dom_char in_homE mem_Collect_eq) + have "(Cls.UP \ (Cls.DN \ Cls.Fun f \ Cls.UP) \ Cls.DN) x = Cls.Fun f x" + using 4 f x by blast + moreover have "... \ Cls.Cod f" + using 4 6 f Cls.Fun_mapsto + by (metis (no_types, lifting) Cls.arr_dom_iff_arr Cls.elem_set_implies_incl_in + Cls.ideD(1) Cls.incl_in_def IntE PiE) + moreover have "... = Cls.set b" + using f by (metis (no_types, lifting) cod_char in_homE mem_Collect_eq) + ultimately show "(Cls.UP \ (Cls.DN \ Cls.Fun f \ Cls.UP) \ Cls.DN) x \ Cls.set b" + by auto + qed + ultimately show ?thesis by blast + qed + have "arr_of_V a b (V_of_arr f) = + Cls.mkArr (Cls.set a) (Cls.set b) + (Cls.UP \ app (VLambda (V_of_ide (dom f)) (Cls.DN \ Cls.Fun f \ Cls.UP)) + \ Cls.DN)" + unfolding arr_of_V_def V_of_arr_def by simp + also have "... = Cls.mkArr (Cls.set a) (Cls.set b) + (Cls.UP \ (Cls.DN \ Cls.Fun f \ Cls.UP) \ Cls.DN)" + proof (intro Cls.mkArr_eqI') + show 6: "\x. x \ Cls.set a \ + (Cls.UP \ app (VLambda (V_of_ide (dom f)) (Cls.DN \ Cls.Fun f \ Cls.UP)) + \ Cls.DN) x = + (Cls.UP \ (Cls.DN \ Cls.Fun f \ Cls.UP) \ Cls.DN) x" + proof - + fix x + assume x: "x \ Cls.set a" + have "(Cls.UP \ app (VLambda (V_of_ide (dom f)) (Cls.DN \ Cls.Fun f \ Cls.UP)) + \ Cls.DN) x = + Cls.UP (app (VLambda (V_of_ide (dom f)) (Cls.DN \ Cls.Fun f \ Cls.UP)) + (Cls.DN x))" + by simp + also have "... = (Cls.UP \ (Cls.DN \ Cls.Fun f \ Cls.UP) \ Cls.DN) x" + proof - + have "Cls.DN x \ elts (V_of_ide (dom f))" + using f + by (metis (no_types, lifting) V_of_ide_def elts_of_set ide_char ide_dom image_eqI + in_homE mem_Collect_eq replacement x) + thus ?thesis + using beta by auto + qed + ultimately show "(Cls.UP \ app (VLambda (V_of_ide (dom f)) + (Cls.DN \ Cls.Fun f \ Cls.UP)) + \ Cls.DN) x = + (Cls.UP \ (Cls.DN \ Cls.Fun f \ Cls.UP) \ Cls.DN) x" + by argo + qed + show "Cls.arr (Cls.mkArr (Cls.set a) (Cls.set b) + (Cls.UP \ app (VLambda (V_of_ide (local.dom f)) + (Cls.DN \ Cls.Fun f \ Cls.UP)) + \ Cls.DN))" + using 5 6 Cls.mkArr_eqI by auto + qed + also have "... = Cls.mkArr (Cls.set a) (Cls.set b) (Cls.Fun f)" + using 4 5 by force + also have "... = f" + using f Cls.mkArr_Fun + by (metis (no_types, lifting) arr_char cod_simp dom_char in_homE mem_Collect_eq) + finally show "arr_of_V a b (V_of_arr f) = f" by blast + qed + show 4: "\a b F. \ide a; ide b; F \ Collect (vfun (V_of_ide a) (V_of_ide b))\ + \ V_of_arr (arr_of_V a b F) = F" + proof - + fix a b F + assume a: "ide a" and b: "ide b" + assume F: "F \ Collect (vfun (V_of_ide a) (V_of_ide b))" + have "F \ elts (VPow (vtimes (V_of_ide a) (V_of_ide b))) \ + elts (V_of_ide a) = Domain (pairs F) \ single_valued (pairs F)" + using F vfun_def by simp + hence 5: "F \ elts (VPi (V_of_ide a) (\_. V_of_ide b))" + unfolding VPi_def + by (metis (no_types, lifting) down elts_of_set mem_Collect_eq subsetI) + let ?f = "Cls.mkArr (Cls.set a) (Cls.set b) (Cls.UP \ app F \ Cls.DN)" + have 6: "Cls.arr ?f" + proof - + have "Cls.set a \ Cls.Univ \ Cls.set b \ Cls.Univ" + using a b ide_char by blast + moreover have "Cls.UP \ app F \ Cls.DN \ Cls.set a \ Cls.set b" + proof + fix x + assume x: "x \ Cls.set a" + have "(Cls.UP \ app F \ Cls.DN) x = Cls.UP (app F (Cls.DN x))" + by simp + moreover have "... \ Cls.set b" + proof - + have "app F (Cls.DN x) \ Cls.DN ` Cls.set b" + using a b ide_char x F app_vfun_mapsto [of "V_of_ide a" "V_of_ide b" F] V_of_ide_def + by auto + thus ?thesis + using \Cls.set a \ Cls.Univ \ Cls.set b \ Cls.Univ\ + by (metis Cls.bij_arr_of bij_betw_def imageI image_inv_into_cancel) + qed + ultimately show "(Cls.UP \ app F \ Cls.DN) x \ Cls.set b" by auto + qed + ultimately show ?thesis by blast + qed + have "V_of_arr (arr_of_V a b F) = VLambda (V_of_ide (dom ?f)) (Cls.DN \ Cls.Fun ?f \ Cls.UP)" + unfolding V_of_arr_def arr_of_V_def by simp + also have "... = VLambda (V_of_ide a) (Cls.DN \ Cls.Fun ?f \ Cls.UP)" + unfolding V_of_ide_def + using a b 6 ide_char V_of_ide_def dom_char Cls.dom_mkArr arrI by auto + also have "... = VLambda (V_of_ide a) + (Cls.DN \ + restrict (Cls.UP \ app F \ Cls.DN) (Cls.set a) \ Cls.UP)" + using 6 Cls.Fun_mkArr by simp + also have "... = VLambda (V_of_ide a) (app F)" + proof - + have 7: "\x. x \ elts (V_of_ide a) \ + (Cls.DN \ restrict (Cls.UP \ app F \ Cls.DN) (Cls.set a) \ Cls.UP) x = + app F x" + unfolding V_of_ide_def + using a + apply simp + by (metis (no_types, lifting) Cls.bij_arr_of a bij_betw_def empty_iff ide_char + image_eqI image_inv_into_cancel) + have 8: "\x. x \ elts (V_of_ide a) \ + (Cls.DN \ restrict (Cls.UP \ app F \ Cls.DN) (Cls.set a) \ Cls.UP) x + \ elts (V_of_ide b)" + using 5 7 VPi_D by fastforce + have "VLambda (V_of_ide a) (app F) \ elts (VPi (V_of_ide a) (\_. V_of_ide b))" + using 5 VPi_I VPi_D by auto + moreover have "VLambda (V_of_ide a) + (Cls.DN \ + restrict (Cls.UP \ app F \ Cls.DN) (Cls.set a) \ Cls.UP) + \ elts (VPi (V_of_ide a) (\_. V_of_ide b))" + using 8 VPi_I by auto + moreover have "\x. x \ elts (V_of_ide a) \ + app (VLambda (V_of_ide a) + (Cls.DN \ + restrict (Cls.UP \ app F \ Cls.DN) (Cls.set a) \ + Cls.UP)) x = + app F x" + using 7 beta by auto + ultimately show ?thesis + using fun_ext by simp + qed + also have "... = F" + using 5 eta [of F "V_of_ide a" "\_. V_of_ide b"] by auto + finally show "V_of_arr (arr_of_V a b F) = F" by blast + qed + show "\ide a; ide b\ + \ bij_betw V_of_arr (hom a b) (Collect (vfun (V_of_ide a) (V_of_ide b)))" + using 1 2 3 4 + apply (intro bij_betwI) + apply blast + apply blast + by auto + show "\ide a; ide b\ + \ bij_betw (arr_of_V a b) (Collect (vfun (V_of_ide a) (V_of_ide b))) (hom a b)" + using 1 2 3 4 + apply (intro bij_betwI) + apply blast + apply blast + by auto + qed + + lemma small_hom: + shows "small (hom a b)" + proof (cases "ide a \ ide b") + assume 1: "\ (ide a \ ide b)" + have "\f. \ \f : a \ b\" + using 1 in_hom_def ide_cod ide_dom by blast + hence "hom a b = {}" + by blast + thus ?thesis by simp + next + assume 1: "ide a \ ide b" + show ?thesis + using 1 bij_betw_hom_vfun(5) small_Collect_vfun small_def + by (metis (no_types, lifting) bij_betw_def small_iff_range) + qed + + text\ + We can now show that the inclusion of the subcategory into the ambient category \Cls\ + creates small products. To do this, we consider a product in \Cls\ of objects of the + subcategory indexed by a small set \I\. Since \Cls\ is a replete set category, + by a previous result we know that the elements of a product object \p\ in \Cls\ + correspond to its points; that is, to the elements of \hom unity p\. + The elements of \hom unity p\ in turn correspond to \I\-tuples. By carrying out + the construction of the set of \I\-tuples in \V\ and exploiting the bijections between + homs of the subcatgory and \V\-sets, we can obtain an injection of \hom unity p\ + to the extension of a \V\-set, thus showing \hom unity p\ is small. + Since \hom unity p\ is small, it determines an object of the subcategory, + which must then be a product in the subcategory, in view of the fact that the + subcategory is full. + \ + + lemma has_small_V_products: + assumes "small (I :: V set)" + shows "has_products I" + proof (unfold has_products_def, intro conjI impI allI) + show "I \ UNIV" + using assms big_UNIV by blast + fix J D + assume D: "discrete_diagram J comp D \ Collect (partial_magma.arr J) = I" + interpret J: category J + using D discrete_diagram_def by blast + interpret D: discrete_diagram J comp D + using D by blast + interpret incloD: composite_functor J comp Cls.comp D map .. + interpret incloD: discrete_diagram J Cls.comp incloD.map + using D.is_discrete + by unfold_locales auto + interpret incloD: diagram_in_set_category J Cls.comp \\A. A \ Cls.Univ\ incloD.map + .. + have 1: "small (Collect J.ide)" + using assms D D.is_discrete by argo + show "\a. has_as_product J D a" + proof - + have 2: "\a. Cls.has_as_product J incloD.map a" + proof - + have "Collect J.ide \ UNIV" + using J.ide_def by blast + thus ?thesis + using 1 D.is_discrete Cls.has_small_products [of "Collect J.ide"] + Cls.has_products_def [of "Collect J.ide"] incloD.discrete_diagram_axioms + by presburger + qed + obtain \D where \D: "Cls.has_as_product J incloD.map \D" + using 2 by blast + interpret \D: constant_functor J Cls.comp \D + using \D Cls.product_is_ide + by unfold_locales auto + obtain \ where \: "product_cone J Cls.comp incloD.map \D \" + using \D Cls.has_as_product_def by blast + interpret \: product_cone J Cls.comp incloD.map \D \ + using \ by blast + have "small (Cls.hom Cls.unity \D)" + proof - + obtain \ where \: "bij_betw \ (Cls.hom Cls.unity \D) (incloD.cones Cls.unity)" + using incloD.limits_are_sets_of_cones \.limit_cone_axioms by blast + let ?J = "ZFC_in_HOL.set (Collect J.arr)" + let ?V_of_point = "\x. VLambda ?J (\j. V_of_arr (\ x j))" + let ?Tuples = "VPi ?J (\j. ZFC_in_HOL.set (V_of_arr ` hom Cls.unity (D j)))" + have V_of_point: "?V_of_point \ Cls.hom Cls.unity \D \ elts ?Tuples" + proof + fix x + assume x: "x \ Cls.hom Cls.unity \D" + have \x: "\ x \ incloD.cones Cls.unity" + using \ x + unfolding bij_betw_def by blast + interpret \x: cone J Cls.comp incloD.map Cls.unity \\ x\ + using \x by blast + have "\j. J.arr j \ V_of_arr (\ x j) + \ elts (ZFC_in_HOL.set (V_of_arr ` hom Cls.unity (D j)))" + proof - + fix j + assume j: "J.arr j" + have "\ x j \ hom Cls.unity (D j)" + by (metis (mono_tags, lifting) D.preserves_ide \x.component_in_hom cod_simp + ideD(1,3) in_hom_char incloD.is_discrete incloD.preserves_cod j map_simp + mem_Collect_eq o_apply terminal_char2 terminal_unity) + moreover have "V_of_arr ` hom Cls.unity (D j) = + elts (ZFC_in_HOL.set (V_of_arr ` hom Cls.unity (D j)))" + using small_hom replacement by simp + ultimately show "V_of_arr (\ x j) + \ elts (ZFC_in_HOL.set (V_of_arr ` hom Cls.unity (D j)))" + using j \x bij_betw_hom_vfun(1) by blast + qed + thus "?V_of_point x \ elts ?Tuples" + using VPi_I by fastforce + qed + have "?V_of_point ` Cls.hom Cls.unity \D \ range elts" + proof - + have "?V_of_point ` Cls.hom Cls.unity \D \ elts ?Tuples" + using V_of_point by blast + thus ?thesis + using smaller_than_small down_raw by auto + qed + moreover have "inj_on ?V_of_point (Cls.hom Cls.unity \D)" + proof + fix x y + assume x: "x \ Cls.hom Cls.unity \D" and y: "y \ Cls.hom Cls.unity \D" + and eq: "?V_of_point x = ?V_of_point y" + have \x: "\ x \ incloD.cones Cls.unity" + using \ x + unfolding bij_betw_def by blast + have \y: "\ y \ incloD.cones Cls.unity" + using \ y + unfolding bij_betw_def by blast + interpret \x: cone J Cls.comp incloD.map Cls.unity \\ x\ + using \x by blast + interpret \y: cone J Cls.comp incloD.map Cls.unity \\ y\ + using \y by blast + have "\ x = \ y" + proof - + have "\j. j \ elts ?J \ \ x j = \ y j" + proof - + fix j + assume j: "j \ elts ?J" + hence 3: "J.arr j" + by (simp add: 1 incloD.is_discrete) + have 4: "ide (D j)" + using 3 incloD.is_discrete D.preserves_ide by force + have 5:"ide Cls.unity" + using Cls.terminal_unity terminal_char terminal_def by auto + have \xj: "\ x j \ hom Cls.unity (D j)" + using 3 4 5 incloD.is_discrete \x.preserves_hom \x.A.map_simp in_hom_char + by (metis (no_types, lifting) J.ide_char \x.component_in_hom ideD(1) map_simp + mem_Collect_eq o_apply) + have \yj: "\ y j \ hom Cls.unity (D j)" + using 3 4 5 incloD.is_discrete \y.preserves_hom \x.A.map_simp in_hom_char + by (metis (no_types, lifting) J.ide_char \y.component_in_hom ideD(1) map_simp + mem_Collect_eq o_apply) + show "\ x j = \ y j" + proof - + have "\j. j \ elts ?J \ V_of_arr (\ x j) = V_of_arr (\ y j)" + using x eq VLambda_eq_D2 by blast + thus ?thesis + using V_of_arr_def + by (metis (mono_tags, lifting) j \xj \yj bij_betw_hom_vfun(3) mem_Collect_eq) + qed + qed + moreover have "elts ?J = Collect J.arr" + by (simp add: 1 incloD.is_discrete) + ultimately show ?thesis + using \x.is_extensional \y.is_extensional + by (metis HOL.ext mem_Collect_eq) + qed + thus "x = y" + using x y \ bij_betw_imp_inj_on inj_on_def + by (metis (no_types, lifting)) + qed + ultimately show "small (Cls.hom Cls.unity \D)" + using small_def by blast + qed + hence "small (Cls.set \D)" + by (simp add: Cls.set_def) + hence 2: "ide \D" + using ide_char Cls.setp_set_ide Cls.product_is_ide \D by blast + interpret \D': constant_functor J comp \D + using 2 by unfold_locales + interpret \': cone J comp D \D \ + proof - + have "\j. J.arr j \ \\ j : \D \ D j\" + proof + fix j + assume j: "J.arr j" + show 3: "arr (\ j)" + by (metis (mono_tags, lifting) 2 D.as_nat_trans.preserves_cod D.is_discrete + D.preserves_ide \.component_in_hom ideD(1) ideD(3) in_homE in_hom_char + j map_simp o_apply) + show "dom (\ j) = \D" + using 3 arr_char dom_char \.preserves_dom by auto + show "cod (\ j) = D j" + using 3 arr_char cod_char \.preserves_cod + by (metis (no_types, lifting) Cls.ideD(3) D.preserves_arr functor.preserves_ide + incloD.is_discrete incloD.is_functor incloD.preserves_cod j map_simp o_apply) + qed + moreover have "D.mkCone \ = \" + using \.is_extensional null_char by auto + ultimately show "cone J comp D \D \" + using 2 D.cone_mkCone [of \D \] by simp + qed + interpret \': product_cone J comp D \D \ + proof + fix a \' + assume \': "D.cone a \'" + interpret \': cone J comp D a \' + using \' by blast + have a: "Cls.ide a" + using ide_char \'.A.value_is_ide by blast + moreover have "\j. J.arr j \ Cls.in_hom (\' j) a (incloD.map j)" + proof - + fix j + assume j: "J.arr j" + have "\\' j : a \ D j\" + using j D.is_discrete \'.component_in_hom by force + thus "Cls.in_hom (\' j) a (incloD.map j)" + using a j D.is_discrete in_hom_char map_simp by auto + qed + ultimately have 3: "incloD.cone a (incloD.mkCone \')" + using incloD.cone_mkCone [of a \'] by blast + interpret \: cone J Cls.comp incloD.map a \incloD.mkCone \'\ + using 3 by blast + have univ: "\!f. Cls.in_hom f a \D \ incloD.cones_map f \ = incloD.mkCone \'" + using \.cone_axioms \.is_universal [of a "incloD.mkCone \'"] by blast + have 4: "incloD.mkCone \' = \'" + using D.as_nat_trans.preserves_reflects_arr D.preserves_arr Limit.cone_def + \' \'.is_extensional identity_functor.intro identity_functor.map_def + incloD.as_nat_trans.is_extensional o_apply + by fastforce + have 5: "D.mkCone \ = \" + using \.is_extensional null_char by auto + have 6: "\f. Cls.in_hom f a \D \ incloD.cones_map f \ = D.cones_map f \" + proof - + fix f + assume f: "Cls.in_hom f a \D" + have "incloD.cones_map f \ = (\j. if J.arr j then Cls.comp (\ j) f else Cls.null)" + using f \.cone_axioms by auto + also have "... = (\j. if J.arr j then comp (\ j) f else null)" + proof - + have "\j. J.arr j \ Cls.comp (\ j) f = comp (\ j) f" + using f 2 comp_char in_hom_char seq_char + by (metis (no_types, lifting) Cls.ext Cls.in_homE \'.ide_apex + \'.preserves_reflects_arr arr_char ide_char) + thus ?thesis + using null_char by auto + qed + also have "... = D.cones_map f \" + proof - + have "\ \ D.cones (cod f)" + proof - + have "\f : a \ \D\" + using f 2 in_hom_char \'.ide_apex ideD(1) by presburger + thus ?thesis + using f \'.cone_axioms by blast + qed + thus ?thesis + using \\ \ D.cones (cod f)\ by simp + qed + finally show "incloD.cones_map f \ = D.cones_map f \" by blast + qed + moreover have "\f. \f : a \ \D\ \ Cls.in_hom f a \D" + using in_hom_char by blast + show "\!f. \f : a \ \D\ \ D.cones_map f \ = \'" + proof - + have "\f. \f : a \ \D\ \ D.cones_map f \ = \'" + using 2 4 5 6 univ \'.ide_apex ideD(1) in_hom_char + by (metis (no_types, lifting)) + moreover have "\f g. \\f : a \ \D\ \ D.cones_map f \ = \'; + \g : a \ \D\ \ D.cones_map g \ = \'\ + \ f = g" + using 2 4 5 6 univ \'.ide_apex ideD(1) in_hom_char by auto + ultimately show ?thesis by blast + qed + qed + show ?thesis + using \'.product_cone_axioms has_as_product_def by blast + qed + qed + + corollary has_small_products: + assumes "small I" and "I \ UNIV" + shows "has_products I" + proof - + have 1: "\I :: V set. small I \ has_products I" + using has_small_V_products by blast + obtain \ where \: "inj_on \ I \ \ ` I \ range elts" + using assms small_def by metis + have "bij_betw (inv_into I \) (\ ` I) I" + using \ bij_betw_inv_into bij_betw_imageI by metis + moreover have "small (\ ` I)" + using assms by auto + ultimately show ?thesis + using assms 1 has_products_preserved_by_bijection by blast + qed + + theorem has_small_limits: + assumes "category (J :: 'j comp)" and "small (Collect (partial_magma.arr J))" + shows "has_limits_of_shape J" + proof - + interpret J: category J + using assms by blast + have "small (Collect J.ide)" + using assms smaller_than_small [of "Collect J.arr" "Collect J.ide"] by fastforce + moreover have "Collect J.ide \ UNIV" + using J.ide_def by blast + moreover have "Collect J.arr \ UNIV" + using J.not_arr_null by blast + ultimately show "has_limits_of_shape J" + using assms has_small_products has_limits_if_has_products [of J] by blast + qed + + sublocale concrete_set_category comp setp UNIV Cls.UP + proof + show "Cls.UP \ UNIV \ Univ" + using Cls.UP_mapsto terminal_char by presburger + show "inj Cls.UP" + using Cls.inj_UP by blast + qed + + lemma is_concrete_set_category: + shows "concrete_set_category comp setp UNIV Cls.UP" + .. + + end + + text\ + In pure HOL (without ZFC), we were able to show that every category \C\ has a ``hom functor'', + but there was necessarily a dependence of the target set category of the hom functor + on the arrow type of \C\. Using the construction of the present theory, we can now show + that every ``locally small'' category \C\ has a hom functor, whose target is the same set + category for all such \C\. To obtain such a hom functor requires a choice, for each hom-set + \hom a b\ of \C\, of an injection of \hom a b\ to the extension of a \V\-set. + \ + + locale locally_small_category = + category + + assumes locally_small: "\ide a; ide b\ \ small (hom b a)" + begin + + interpretation Cop: dual_category C .. + interpretation CopxC: product_category Cop.comp C .. + interpretation S: ZFC_set_cat . + + definition Hom + where "Hom \ \(b, a). S.UP o (SOME \. \ ` hom b a \ range elts \ inj_on \ (hom b a))" + + interpretation Hom: hom_functor C S.comp S.setp Hom + proof + have 1: "\a b. Hom (b, a) \ hom b a \ S.Univ \ inj_on (Hom (b, a)) (hom b a)" + proof - + fix a b + show "Hom (b, a) \ hom b a \ S.Univ \ inj_on (Hom (b, a)) (hom b a)" + proof (cases "ide a \ ide b") + show "\ (ide a \ ide b) \ ?thesis" + using inj_on_def by fastforce + assume ab: "ide a \ ide b" + show ?thesis + proof + have 1: "\\. \ ` hom b a \ range elts \ inj_on \ (hom b a)" + using ab locally_small [of a b] small_def [of "hom b a"] by blast + let ?\ = "SOME \. \ ` hom b a \ range elts \ inj_on \ (hom b a)" + have \: "?\ ` hom b a \ range elts \ inj_on ?\ (hom b a)" + using 1 someI_ex [of "\\. \ ` hom b a \ range elts \ inj_on \ (hom b a)"] + by blast + show "Hom (b, a) \ hom b a \ S.Univ" + unfolding Hom_def + using \ S.UP_mapsto by auto + show "inj_on (Hom (b, a)) (hom b a)" + unfolding Hom_def + apply simp + using ab \ S.inj_UP comp_inj_on injD inj_on_def + by (metis (no_types, lifting)) + qed + qed + qed + show "\f. arr f \ Hom (dom f, cod f) f \ S.Univ" + using 1 by blast + show "\b a. \ide b; ide a\ \ inj_on (Hom (b, a)) (hom b a)" + using 1 by blast + show "\b a. \ide b; ide a\ \ Hom (b, a) ` hom b a \ S.S.Univ \ small (Hom (b, a) ` hom b a)" + using 1 locally_small S.terminal_char by force + qed + + lemma has_ZFC_hom_functor: + shows "hom_functor C S.comp S.setp Hom" + .. + + text\ + Using this result, we can now state a more traditional version of the Yoneda Lemma + in which the target category of the Yoneda functor is the same for all locally small + categories. + \ + + interpretation Y: yoneda_functor C S.comp S.setp Hom + .. + + theorem ZFC_yoneda_lemma: + assumes "ide a" and "functor Cop.comp S.comp F" + shows "\\. bij_betw \ (S.set (F a)) {\. natural_transformation Cop.comp S.comp (Y.Y a) F \}" + proof - + interpret F: "functor" Cop.comp S.comp F + using assms(2) by blast + interpret F: set_valued_functor Cop.comp S.comp S.setp F + .. + interpret Ya: yoneda_functor_fixed_object C S.comp S.setp Hom a + using assms(1) by unfold_locales blast + interpret Ya: yoneda_lemma C S.comp S.setp Hom F a + .. + show ?thesis + using Ya.yoneda_lemma by blast + qed + + end + +end diff --git a/thys/MonoidalCategory/FreeMonoidalCategory.thy b/thys/MonoidalCategory/FreeMonoidalCategory.thy --- a/thys/MonoidalCategory/FreeMonoidalCategory.thy +++ b/thys/MonoidalCategory/FreeMonoidalCategory.thy @@ -1,3473 +1,3473 @@ (* Title: FreeMonoidalCategory Author: Eugene W. Stark , 2017 Maintainer: Eugene W. Stark *) chapter "The Free Monoidal Category" text_raw\ \label{fmc-chap} \ theory FreeMonoidalCategory imports Category3.Subcategory MonoidalFunctor begin text \ In this theory, we use the monoidal language of a category @{term C} defined in @{theory MonoidalCategory.MonoidalCategory} to give a construction of the free monoidal category \\C\ generated by @{term C}. The arrows of \\C\ are the equivalence classes of formal arrows obtained by declaring two formal arrows to be equivalent if they are parallel and have the same diagonalization. Composition, tensor, and the components of the associator and unitors are all defined in terms of the corresponding syntactic constructs. After defining \\C\ and showing that it does indeed have the structure of a monoidal category, we prove the freeness: every functor from @{term C} to a monoidal category @{term D} extends uniquely to a strict monoidal functor from \\C\ to @{term D}. We then consider the full subcategory \\\<^sub>SC\ of \\C\ whose objects are the equivalence classes of diagonal identity terms ({\em i.e.}~equivalence classes of lists of identity arrows of @{term C}), and we show that this category is monoidally equivalent to \\C\. In addition, we show that \\\<^sub>SC\ is the free strict monoidal category, as any functor from \C\ to a strict monoidal category @{term D} extends uniquely to a strict monoidal functor from \\\<^sub>SC\ to @{term D}. \ section "Syntactic Construction" locale free_monoidal_category = monoidal_language C for C :: "'c comp" begin no_notation C.in_hom ("\_ : _ \ _\") notation C.in_hom ("\_ : _ \\<^sub>C _\") text \ Two terms of the monoidal language of @{term C} are defined to be equivalent if they are parallel formal arrows with the same diagonalization. \ abbreviation equiv where "equiv t u \ Par t u \ \<^bold>\t\<^bold>\ = \<^bold>\u\<^bold>\" text \ Arrows of \\C\ will be the equivalence classes of formal arrows determined by the relation @{term equiv}. We define here the property of being an equivalence class of the relation @{term equiv}. Later we show that this property coincides with that of being an arrow of the category that we will construct. \ type_synonym 'a arr = "'a term set" definition ARR where "ARR f \ f \ {} \ (\t. t \ f \ f = Collect (equiv t))" lemma not_ARR_empty: shows "\ARR {}" using ARR_def by simp lemma ARR_eqI: assumes "ARR f" and "ARR g" and "f \ g \ {}" shows "f = g" using assms ARR_def by fastforce text \ We will need to choose a representative of each equivalence class as a normal form. The requirements we have of these representatives are: (1) the normal form of an arrow @{term t} is equivalent to @{term t}; (2) equivalent arrows have identical normal forms; (3) a normal form is a canonical term if and only if its diagonalization is an identity. It follows from these properties and coherence that a term and its normal form have the same evaluation in any monoidal category. We choose here as a normal form for an arrow @{term t} the particular term @{term "Inv (Cod t\<^bold>\) \<^bold>\ \<^bold>\t\<^bold>\ \<^bold>\ Dom t\<^bold>\"}. However, the only specific properties of this definition we actually use are the three we have just stated. \ definition norm ("\<^bold>\_\<^bold>\") where "\<^bold>\t\<^bold>\ = Inv (Cod t\<^bold>\) \<^bold>\ \<^bold>\t\<^bold>\ \<^bold>\ Dom t\<^bold>\" text \ If @{term t} is a formal arrow, then @{term t} is equivalent to its normal form. \ lemma equiv_norm_Arr: assumes "Arr t" shows "equiv \<^bold>\t\<^bold>\ t" proof - have "Par t (Inv (Cod t\<^bold>\) \<^bold>\ \<^bold>\t\<^bold>\ \<^bold>\ Dom t\<^bold>\)" using assms Diagonalize_in_Hom red_in_Hom Inv_in_Hom Arr_implies_Ide_Dom Arr_implies_Ide_Cod Ide_implies_Arr Can_red by auto moreover have "\<^bold>\(Inv (Cod t\<^bold>\) \<^bold>\ \<^bold>\t\<^bold>\ \<^bold>\ Dom t\<^bold>\)\<^bold>\ = \<^bold>\t\<^bold>\" using assms Arr_implies_Ide_Dom Arr_implies_Ide_Cod Diagonalize_preserves_Ide Diagonalize_in_Hom Diagonalize_Inv [of "Cod t\<^bold>\"] Diag_Diagonalize CompDiag_Diag_Dom [of "\<^bold>\t\<^bold>\"] CompDiag_Cod_Diag [of "\<^bold>\t\<^bold>\"] by (simp add: Diagonalize_red [of "Cod t"] Can_red(1)) ultimately show ?thesis using norm_def by simp qed text \ Equivalent arrows have identical normal forms. \ lemma norm_respects_equiv: assumes "equiv t u" shows "\<^bold>\t\<^bold>\ = \<^bold>\u\<^bold>\" using assms norm_def by simp text \ The normal form of an arrow is canonical if and only if its diagonalization is an identity term. \ lemma Can_norm_iff_Ide_Diagonalize: assumes "Arr t" shows "Can \<^bold>\t\<^bold>\ \ Ide \<^bold>\t\<^bold>\" using assms norm_def Can_implies_Arr Arr_implies_Ide_Dom Arr_implies_Ide_Cod Can_red Inv_preserves_Can Diagonalize_preserves_Can red_in_Hom Diagonalize_in_Hom Ide_Diagonalize_Can by fastforce text \ We now establish various additional properties of normal forms that are consequences of the three already proved. The definition \norm_def\ is not used subsequently. \ lemma norm_preserves_Can: assumes "Can t" shows "Can \<^bold>\t\<^bold>\" using assms Can_implies_Arr Can_norm_iff_Ide_Diagonalize Ide_Diagonalize_Can by simp lemma Par_Arr_norm: assumes "Arr t" shows "Par \<^bold>\t\<^bold>\ t" using assms equiv_norm_Arr by auto lemma Diagonalize_norm [simp]: assumes "Arr t" shows " \<^bold>\\<^bold>\t\<^bold>\\<^bold>\ = \<^bold>\t\<^bold>\" using assms equiv_norm_Arr by auto lemma unique_norm: assumes "ARR f" shows "\!t. \u. u \ f \ \<^bold>\u\<^bold>\ = t" proof have 1: "(SOME t. t \ f) \ f" using assms ARR_def someI_ex [of "\t. t \ f"] by auto show "\t. \u. u \ f \ \<^bold>\u\<^bold>\ = t \ t = \<^bold>\SOME t. t \ f\<^bold>\" using assms ARR_def 1 by auto show "\u. u \ f \ \<^bold>\u\<^bold>\ = \<^bold>\SOME t. t \ f\<^bold>\" using assms ARR_def 1 norm_respects_equiv by blast qed lemma Dom_norm: assumes "Arr t" shows "Dom \<^bold>\t\<^bold>\ = Dom t" using assms Par_Arr_norm by metis lemma Cod_norm: assumes "Arr t" shows "Cod \<^bold>\t\<^bold>\ = Cod t" using assms Par_Arr_norm by metis lemma norm_in_Hom: assumes "Arr t" shows "\<^bold>\t\<^bold>\ \ Hom (Dom t) (Cod t)" using assms Par_Arr_norm [of t] by simp text \ As all the elements of an equivalence class have the same normal form, we can use the normal form of an arbitrarily chosen element as a canonical representative. \ definition rep where "rep f \ \<^bold>\SOME t. t \ f\<^bold>\" lemma rep_in_ARR: assumes "ARR f" shows "rep f \ f" using assms ARR_def someI_ex [of "\t. t \ f"] equiv_norm_Arr rep_def ARR_def by fastforce lemma Arr_rep_ARR: assumes "ARR f" shows "Arr (rep f)" using assms ARR_def rep_in_ARR by auto text \ We next define a function \mkarr\ that maps formal arrows to their equivalence classes. For terms that are not formal arrows, the function yields the empty set. \ definition mkarr where "mkarr t = Collect (equiv t)" lemma mkarr_extensionality: assumes "\Arr t" shows "mkarr t = {}" using assms mkarr_def by simp lemma ARR_mkarr: assumes "Arr t" shows "ARR (mkarr t)" using assms ARR_def mkarr_def by auto lemma mkarr_memb_ARR: assumes "ARR f" and "t \ f" shows "mkarr t = f" using assms ARR_def mkarr_def by simp lemma mkarr_rep_ARR [simp]: assumes "ARR f" shows "mkarr (rep f) = f" using assms rep_in_ARR mkarr_memb_ARR by auto lemma Arr_in_mkarr: assumes "Arr t" shows "t \ mkarr t" using assms mkarr_def by simp text \ Two terms are related by @{term equiv} iff they are both formal arrows and have identical normal forms. \ lemma equiv_iff_eq_norm: shows "equiv t u \ Arr t \ Arr u \ \<^bold>\t\<^bold>\ = \<^bold>\u\<^bold>\" proof show "equiv t u \ Arr t \ Arr u \ \<^bold>\t\<^bold>\ = \<^bold>\u\<^bold>\" using mkarr_def Arr_in_mkarr ARR_mkarr unique_norm by blast show "Arr t \ Arr u \ \<^bold>\t\<^bold>\ = \<^bold>\u\<^bold>\ \ equiv t u" using Par_Arr_norm Diagonalize_norm by metis qed lemma norm_norm [simp]: assumes "Arr t" shows "\<^bold>\\<^bold>\t\<^bold>\\<^bold>\ = \<^bold>\t\<^bold>\" proof - have "t \ mkarr t" using assms Arr_in_mkarr by blast moreover have "\<^bold>\t\<^bold>\ \ mkarr t" using assms equiv_norm_Arr mkarr_def by simp ultimately show ?thesis using assms ARR_mkarr unique_norm by auto qed lemma norm_in_ARR: assumes "ARR f" and "t \ f" shows "\<^bold>\t\<^bold>\ \ f" using assms ARR_def equiv_iff_eq_norm norm_norm Par_Arr_norm by fastforce lemma norm_rep_ARR [simp]: assumes "ARR f" shows "\<^bold>\rep f\<^bold>\ = rep f" using assms ARR_def someI_ex [of "\t. t \ f"] rep_def norm_norm by fastforce lemma norm_memb_eq_rep_ARR: assumes "ARR f" and "t \ f" shows "norm t = rep f" using assms ARR_def someI_ex [of "\t. t \ f"] unique_norm rep_def by metis lemma rep_mkarr: assumes "Arr f" shows "rep (mkarr f) = \<^bold>\f\<^bold>\" using assms ARR_mkarr Arr_in_mkarr norm_memb_eq_rep_ARR by fastforce text \ To prove that two terms determine the same equivalence class, it suffices to show that they are parallel formal arrows with identical diagonalizations. \ lemma mkarr_eqI [intro]: assumes "Par f g" and "\<^bold>\f\<^bold>\ = \<^bold>\g\<^bold>\" shows "mkarr f = mkarr g" using assms by (metis ARR_mkarr equiv_iff_eq_norm rep_mkarr mkarr_rep_ARR) text \ We use canonical representatives to lift the formal domain and codomain functions from terms to equivalence classes. \ abbreviation DOM where "DOM f \ Dom (rep f)" abbreviation COD where "COD f \ Cod (rep f)" lemma DOM_mkarr: assumes "Arr t" shows "DOM (mkarr t) = Dom t" using assms rep_mkarr by (metis Par_Arr_norm) lemma COD_mkarr: assumes "Arr t" shows "COD (mkarr t) = Cod t" using assms rep_mkarr by (metis Par_Arr_norm) text \ A composition operation can now be defined on equivalence classes using the syntactic constructor \Comp\. \ definition comp (infixr "\" 55) where "comp f g \ (if ARR f \ ARR g \ DOM f = COD g then mkarr ((rep f) \<^bold>\ (rep g)) else {})" text \ We commence the task of showing that the composition \comp\ so defined determines a category. \ interpretation partial_magma comp apply unfold_locales using comp_def not_ARR_empty by metis notation in_hom ("\_ : _ \ _\") text \ The empty set serves as the null for the composition. \ lemma null_char: shows "null = {}" proof - let ?P = "\n. \f. f \ n = n \ n \ f = n" have "?P {}" using comp_def not_ARR_empty by simp moreover have "\!n. ?P n" using ex_un_null by metis ultimately show ?thesis using null_def theI_unique [of ?P "{}"] by (metis comp_null(2)) qed lemma ARR_comp: assumes "ARR f" and "ARR g" and "DOM f = COD g" shows "ARR (f \ g)" using assms comp_def Arr_rep_ARR ARR_mkarr(1) by simp lemma DOM_comp [simp]: assumes "ARR f" and "ARR g" and "DOM f = COD g" shows "DOM (f \ g) = DOM g" using assms comp_def ARR_comp Arr_rep_ARR DOM_mkarr by simp lemma COD_comp [simp]: assumes "ARR f" and "ARR g" and "DOM f = COD g" shows "COD (f \ g) = COD f" using assms comp_def ARR_comp Arr_rep_ARR COD_mkarr by simp lemma comp_assoc: assumes "g \ f \ null" and "h \ g \ null" shows "h \ (g \ f) = (h \ g) \ f" proof - have 1: "ARR f \ ARR g \ ARR h \ DOM h = COD g \ DOM g = COD f" using assms comp_def not_ARR_empty null_char by metis hence 2: "Arr (rep f) \ Arr (rep g) \ Arr (rep h) \ Dom (rep h) = Cod (rep g) \ Dom (rep g) = Cod (rep f)" using Arr_rep_ARR by simp have 3: "h \ g \ f = mkarr (rep h \<^bold>\ rep (mkarr (rep g \<^bold>\ rep f)))" using 1 comp_def ARR_comp COD_comp by simp also have "... = mkarr (rep h \<^bold>\ rep g \<^bold>\ rep f)" proof - have "equiv (rep h \<^bold>\ rep (mkarr (rep g \<^bold>\ rep f))) (rep h \<^bold>\ rep g \<^bold>\ rep f)" proof - have "Par (rep h \<^bold>\ rep (mkarr (rep g \<^bold>\ rep f))) (rep h \<^bold>\ rep g \<^bold>\ rep f)" using 1 2 3 DOM_mkarr ARR_comp COD_comp mkarr_extensionality not_ARR_empty by (metis Arr.simps(4) Cod.simps(4) Dom.simps(4) snd_map_prod) (* Here metis claims it is not using snd_map_prod, but removing it fails. *) moreover have "\<^bold>\rep h \<^bold>\ rep (mkarr (rep g \<^bold>\ rep f))\<^bold>\ = \<^bold>\rep h \<^bold>\ rep g \<^bold>\ rep f\<^bold>\" using 1 2 Arr_rep_ARR rep_mkarr rep_in_ARR assms(1) ARR_comp mkarr_extensionality comp_def equiv_iff_eq_norm norm_memb_eq_rep_ARR null_char by auto ultimately show ?thesis using equiv_iff_eq_norm by blast qed thus ?thesis using mkarr_def by force qed also have "... = mkarr ((rep h \<^bold>\ rep g) \<^bold>\ rep f)" proof - have "Par (rep h \<^bold>\ rep g \<^bold>\ rep f) ((rep h \<^bold>\ rep g) \<^bold>\ rep f)" using 2 by simp moreover have "\<^bold>\rep h \<^bold>\ rep g \<^bold>\ rep f\<^bold>\ = \<^bold>\(rep h \<^bold>\ rep g) \<^bold>\ rep f\<^bold>\" using 2 Diag_Diagonalize by (simp add: CompDiag_assoc) ultimately show ?thesis using equiv_iff_eq_norm by (simp add: mkarr_def) qed also have "... = mkarr (rep (mkarr (rep h \<^bold>\ rep g)) \<^bold>\ rep f)" proof - have "equiv (rep (mkarr (rep h \<^bold>\ rep g)) \<^bold>\ rep f) ((rep h \<^bold>\ rep g) \<^bold>\ rep f)" proof - have "Par (rep (mkarr (rep h \<^bold>\ rep g)) \<^bold>\ rep f) ((rep h \<^bold>\ rep g) \<^bold>\ rep f)" using 1 2 Arr_rep_ARR DOM_comp ARR_comp COD_comp comp_def by auto moreover have "\<^bold>\rep (mkarr (rep h \<^bold>\ rep g)) \<^bold>\ rep f\<^bold>\ = \<^bold>\(rep h \<^bold>\ rep g) \<^bold>\ rep f\<^bold>\" using assms(2) 1 2 ARR_comp Arr_rep_ARR mkarr_extensionality rep_mkarr rep_in_ARR equiv_iff_eq_norm norm_memb_eq_rep_ARR comp_def null_char by simp ultimately show ?thesis using equiv_iff_eq_norm by blast qed thus ?thesis using mkarr_def by auto qed also have "... = (h \ g) \ f" using 1 comp_def ARR_comp DOM_comp by simp finally show ?thesis by blast qed lemma Comp_in_comp_ARR: assumes "ARR f" and "ARR g" and "DOM f = COD g" and "t \ f" and "u \ g" shows "t \<^bold>\ u \ f \ g" proof - have "equiv (t \<^bold>\ u) (rep f \<^bold>\ rep g)" proof - have 1: "Par (t \<^bold>\ u) (rep f \<^bold>\ rep g)" using assms ARR_def Arr_rep_ARR COD_mkarr DOM_mkarr mkarr_memb_ARR mkarr_extensionality by (metis (no_types, lifting) Arr.simps(4) Cod.simps(4) Dom.simps(4) snd_map_prod) (* Here metis claims it is not using snd_map_prod, but removing it fails. *) moreover have "\<^bold>\t \<^bold>\ u\<^bold>\ = \<^bold>\rep f \<^bold>\ rep g\<^bold>\" using assms 1 rep_in_ARR equiv_iff_eq_norm norm_memb_eq_rep_ARR by (metis (no_types, lifting) Arr.simps(4) Diagonalize.simps(4)) ultimately show ?thesis by simp qed thus ?thesis using assms comp_def mkarr_def by simp qed text \ Ultimately, we will show that that the identities of the category are those equivalence classes, all of whose members diagonalize to formal identity arrows, having the further property that their canonical representative is a formal endo-arrow. \ definition IDE where "IDE f \ ARR f \ (\t. t \ f \ Ide \<^bold>\t\<^bold>\) \ DOM f = COD f" lemma IDE_implies_ARR: assumes "IDE f" shows "ARR f" using assms IDE_def ARR_def by auto lemma IDE_mkarr_Ide: assumes "Ide a" shows "IDE (mkarr a)" proof - have "DOM (mkarr a) = COD (mkarr a)" using assms mkarr_def equiv_iff_eq_norm Par_Arr_norm COD_mkarr DOM_mkarr Ide_in_Hom by (metis Ide_implies_Can Inv_Ide Ide_implies_Arr Inv_preserves_Can(2)) moreover have "ARR (mkarr a) \ (\t. t \ mkarr a \ Ide \<^bold>\t\<^bold>\)" proof - have "ARR (mkarr a)" using assms ARR_mkarr Ide_implies_Arr by simp moreover have "\t. t \ mkarr a \ Ide \<^bold>\t\<^bold>\" using assms mkarr_def Diagonalize_preserves_Ide by fastforce ultimately show ?thesis by blast qed ultimately show ?thesis using IDE_def by blast qed lemma IDE_implies_ide: assumes "IDE a" shows "ide a" proof (unfold ide_def) have "a \ a \ null" proof - have "rep a \<^bold>\ rep a \ a \ a" using assms IDE_def comp_def Arr_rep_ARR Arr_in_mkarr by simp thus ?thesis using null_char by auto qed moreover have "\f. (f \ a \ null \ f \ a = f) \ (a \ f \ null \ a \ f = f)" proof fix f :: "'c arr" show "a \ f \ null \ a \ f = f" proof assume f: "a \ f \ null" hence "ARR f" using comp_def null_char by auto have "rep a \<^bold>\ rep f \ a \ f" using assms f Comp_in_comp_ARR comp_def rep_in_ARR null_char by metis moreover have "rep a \<^bold>\ rep f \ f" proof - have "rep f \ f" using \ARR f\ rep_in_ARR by auto moreover have "equiv (rep a \<^bold>\ rep f) (rep f)" proof - have 1: "Par (rep a \<^bold>\ rep f) (rep f)" using assms f comp_def mkarr_extensionality Arr_rep_ARR IDE_def null_char by (metis Cod.simps(4) Dom.simps(4)) moreover have "\<^bold>\rep a \<^bold>\ rep f\<^bold>\ = \<^bold>\rep f\<^bold>\" using assms f 1 comp_def IDE_def CompDiag_Ide_Diag Diag_Diagonalize(1) Diag_Diagonalize(2) Diag_Diagonalize(3) rep_in_ARR by auto ultimately show ?thesis by auto qed ultimately show ?thesis using \ARR f\ ARR_def by auto qed ultimately show "a \ f = f" using mkarr_memb_ARR comp_def by auto qed show "f \ a \ null \ f \ a = f" proof assume f: "f \ a \ null" hence "ARR f" using comp_def null_char by auto have "rep f \<^bold>\ rep a \ f \ a" using assms f Comp_in_comp_ARR comp_def rep_in_ARR null_char by metis moreover have "rep f \<^bold>\ rep a \ f" proof - have "rep f \ f" using \ARR f\ rep_in_ARR by auto moreover have "equiv (rep f \<^bold>\ rep a) (rep f)" proof - have 1: "Par (rep f \<^bold>\ rep a) (rep f)" using assms f comp_def mkarr_extensionality Arr_rep_ARR IDE_def null_char by (metis Cod.simps(4) Dom.simps(4)) moreover have "\<^bold>\rep f \<^bold>\ rep a\<^bold>\ = \<^bold>\rep f\<^bold>\" using assms f 1 comp_def IDE_def CompDiag_Diag_Ide Diag_Diagonalize(1) Diag_Diagonalize(2) Diag_Diagonalize(3) rep_in_ARR by force ultimately show ?thesis by auto qed ultimately show ?thesis using \ARR f\ ARR_def by auto qed ultimately show "f \ a = f" using mkarr_memb_ARR comp_def by auto qed qed ultimately show "a \ a \ null \ (\f. (f \ a \ null \ f \ a = f) \ (a \ f \ null \ a \ f = f))" by blast qed lemma ARR_iff_has_domain: shows "ARR f \ domains f \ {}" proof assume f: "domains f \ {}" show "ARR f" using f domains_def comp_def null_char by auto next assume f: "ARR f" have "Ide (DOM f)" using f ARR_def by (simp add: Arr_implies_Ide_Dom Arr_rep_ARR) hence "IDE (mkarr (DOM f))" using IDE_mkarr_Ide by metis hence "ide (mkarr (DOM f))" using IDE_implies_ide by simp moreover have "f \ mkarr (DOM f) = f" proof - have 1: "rep f \<^bold>\ DOM f \ f \ mkarr (DOM f)" using f Comp_in_comp_ARR using IDE_implies_ARR Ide_in_Hom rep_in_ARR \IDE (mkarr (DOM f))\ \Ide (DOM f)\ Arr_in_mkarr COD_mkarr by fastforce moreover have "rep f \<^bold>\ DOM f \ f" proof - have 2: "rep f \ f" using f rep_in_ARR by simp moreover have "equiv (rep f \<^bold>\ DOM f) (rep f)" by (metis 1 Arr.simps(4) Arr_rep_ARR COD_mkarr Cod.simps(4) Diagonalize_Comp_Arr_Dom Dom.simps(4) IDE_def Ide_implies_Arr \IDE (mkarr (DOM f))\ \Ide (DOM f)\ all_not_in_conv DOM_mkarr comp_def) ultimately show ?thesis using f ARR_eqI 1 \ide (mkarr (DOM f))\ null_char ide_def by auto qed ultimately show ?thesis using f ARR_eqI \ide (mkarr (DOM f))\ null_char ide_def by auto qed ultimately show "domains f \ {}" using f domains_def not_ARR_empty null_char by auto qed lemma ARR_iff_has_codomain: shows "ARR f \ codomains f \ {}" proof assume f: "codomains f \ {}" show "ARR f" using f codomains_def comp_def null_char by auto next assume f: "ARR f" have "Ide (COD f)" using f ARR_def by (simp add: Arr_rep_ARR Arr_implies_Ide_Cod) hence "IDE (mkarr (COD f))" using IDE_mkarr_Ide by metis hence "ide (mkarr (COD f))" using IDE_implies_ide by simp moreover have "mkarr (COD f) \ f = f" proof - have 1: "COD f \<^bold>\ rep f \ mkarr (COD f) \ f" using f Comp_in_comp_ARR using IDE_implies_ARR Ide_in_Hom rep_in_ARR \IDE (mkarr (COD f))\ \Ide (COD f)\ Arr_in_mkarr DOM_mkarr by fastforce moreover have "COD f \<^bold>\ rep f \ f" using 1 null_char norm_rep_ARR norm_memb_eq_rep_ARR mkarr_memb_ARR \ide (mkarr (COD f))\ emptyE equiv_iff_eq_norm mkarr_extensionality ide_def by metis ultimately show ?thesis using f ARR_eqI \ide (mkarr (COD f))\ null_char ide_def by auto qed ultimately show "codomains f \ {}" using codomains_def f not_ARR_empty null_char by auto qed lemma arr_iff_ARR: shows "arr f \ ARR f" using arr_def ARR_iff_has_domain ARR_iff_has_codomain by simp text \ The arrows of the category are the equivalence classes of formal arrows. \ lemma arr_char: shows "arr f \ f \ {} \ (\t. t \ f \ f = mkarr t)" using arr_iff_ARR ARR_def mkarr_def by simp lemma seq_char: shows "seq g f \ g \ f \ null" proof show "g \ f \ null \ seq g f" using comp_def null_char Comp_in_comp_ARR rep_in_ARR ARR_mkarr Arr_rep_ARR arr_iff_ARR by auto show "seq g f \ g \ f \ null" by auto qed lemma seq_char': shows "seq g f \ ARR f \ ARR g \ DOM g = COD f" proof show "ARR f \ ARR g \ DOM g = COD f \ seq g f" using comp_def null_char Comp_in_comp_ARR rep_in_ARR ARR_mkarr Arr_rep_ARR arr_iff_ARR by auto have "\ (ARR f \ ARR g \ DOM g = COD f) \ g \ f = null" using comp_def null_char by auto thus "seq g f \ ARR f \ ARR g \ DOM g = COD f" using ext by fastforce qed text \ Finally, we can show that the composition \comp\ determines a category. \ interpretation category comp proof show "\f. domains f \ {} \ codomains f \ {}" using ARR_iff_has_domain ARR_iff_has_codomain by simp show 1: "\f g. g \ f \ null \ seq g f" using comp_def ARR_comp null_char arr_iff_ARR by metis fix f g h show "seq h g \ seq (h \ g) f \ seq g f" using seq_char' by auto show "seq h (g \ f) \ seq g f \ seq h g" using seq_char' by auto show "seq g f \ seq h g \ seq (h \ g) f" using seq_char' ARR_comp arr_iff_ARR by auto show "seq g f \ seq h g \ (h \ g) \ f = h \ g \ f" using seq_char comp_assoc by auto qed lemma mkarr_rep [simp]: assumes "arr f" shows "mkarr (rep f) = f" using assms arr_iff_ARR by simp lemma arr_mkarr [simp]: assumes "Arr t" shows "arr (mkarr t)" using assms by (simp add: ARR_mkarr arr_iff_ARR) lemma mkarr_memb: assumes "t \ f" and "arr f" shows "Arr t" and "mkarr t = f" using assms arr_char mkarr_extensionality by auto lemma rep_in_arr [simp]: assumes "arr f" shows "rep f \ f" using assms by (simp add: rep_in_ARR arr_iff_ARR) lemma Arr_rep [simp]: assumes "arr f" shows "Arr (rep f)" using assms mkarr_memb rep_in_arr by blast lemma rep_in_Hom: assumes "arr f" shows "rep f \ Hom (DOM f) (COD f)" using assms by simp lemma norm_memb_eq_rep: assumes "arr f" and "t \ f" shows "\<^bold>\t\<^bold>\ = rep f" using assms arr_iff_ARR norm_memb_eq_rep_ARR by auto lemma norm_rep: assumes "arr f" shows "\<^bold>\rep f\<^bold>\ = rep f" using assms norm_memb_eq_rep by simp text \ Composition, domain, and codomain on arrows reduce to the corresponding syntactic operations on their representative terms. \ lemma comp_mkarr [simp]: assumes "Arr t" and "Arr u" and "Dom t = Cod u" shows "mkarr t \ mkarr u = mkarr (t \<^bold>\ u)" using assms by (metis (no_types, lifting) ARR_mkarr ARR_comp ARR_def Arr_in_mkarr COD_mkarr Comp_in_comp_ARR DOM_mkarr mkarr_def) lemma dom_char: shows "dom f = (if arr f then mkarr (DOM f) else null)" proof - have "\arr f \ ?thesis" using dom_def by (simp add: arr_def) moreover have "arr f \ ?thesis" proof - assume f: "arr f" have "dom f = mkarr (DOM f)" proof (intro dom_eqI) have 1: "Ide (DOM f)" using f arr_char by (metis Arr_rep Arr_implies_Ide_Dom) hence 2: "IDE (mkarr (DOM f))" using IDE_mkarr_Ide by metis thus "ide (mkarr (DOM f))" using IDE_implies_ide by simp moreover show "seq f (mkarr (DOM f))" proof - have "f \ mkarr (DOM f) \ null" using f 1 2 ARR_def DOM_mkarr IDE_implies_ARR Ide_in_Hom ARR_comp IDE_def ARR_iff_has_codomain ARR_iff_has_domain null_char arr_def by (metis (mono_tags, lifting) mem_Collect_eq) thus ?thesis using seq_char by simp qed qed thus ?thesis using f by simp qed ultimately show ?thesis by blast qed lemma dom_simp: assumes "arr f" shows "dom f = mkarr (DOM f)" using assms dom_char by simp lemma cod_char: shows "cod f = (if arr f then mkarr (COD f) else null)" proof - have "\arr f \ ?thesis" using cod_def by (simp add: arr_def) moreover have "arr f \ ?thesis" proof - assume f: "arr f" have "cod f = mkarr (COD f)" proof (intro cod_eqI) have 1: "Ide (COD f)" using f arr_char by (metis Arr_rep Arr_implies_Ide_Cod) hence 2: "IDE (mkarr (COD f))" using IDE_mkarr_Ide by metis thus "ide (mkarr (COD f))" using IDE_implies_ide by simp moreover show "seq (mkarr (COD f)) f" proof - have "mkarr (COD f) \ f \ null" using f 1 2 ARR_def DOM_mkarr IDE_implies_ARR Ide_in_Hom ARR_comp IDE_def ARR_iff_has_codomain ARR_iff_has_domain null_char arr_def by (metis (mono_tags, lifting) mem_Collect_eq) thus ?thesis using seq_char by simp qed qed thus ?thesis using f by simp qed ultimately show ?thesis by blast qed lemma cod_simp: assumes "arr f" shows "cod f = mkarr (COD f)" using assms cod_char by simp lemma Dom_memb: assumes "arr f" and "t \ f" shows "Dom t = DOM f" using assms DOM_mkarr mkarr_extensionality arr_char by fastforce lemma Cod_memb: assumes "arr f" and "t \ f" shows "Cod t = COD f" using assms COD_mkarr mkarr_extensionality arr_char by fastforce lemma dom_mkarr [simp]: assumes "Arr t" shows "dom (mkarr t) = mkarr (Dom t)" using assms dom_char DOM_mkarr arr_mkarr by auto lemma cod_mkarr [simp]: assumes "Arr t" shows "cod (mkarr t) = mkarr (Cod t)" using assms cod_char COD_mkarr arr_mkarr by auto lemma mkarr_in_hom: assumes "Arr t" shows "\mkarr t : mkarr (Dom t) \ mkarr (Cod t)\" using assms arr_mkarr dom_mkarr cod_mkarr by auto lemma DOM_in_dom [intro]: assumes "arr f" shows "DOM f \ dom f" using assms dom_char by (metis Arr_in_mkarr mkarr_extensionality ideD(1) ide_dom not_arr_null null_char) lemma COD_in_cod [intro]: assumes "arr f" shows "COD f \ cod f" using assms cod_char by (metis Arr_in_mkarr mkarr_extensionality ideD(1) ide_cod not_arr_null null_char) lemma DOM_dom: assumes "arr f" shows "DOM (dom f) = DOM f" using assms Arr_rep Arr_implies_Ide_Dom Ide_implies_Arr dom_char rep_mkarr Par_Arr_norm Ide_in_Hom by simp lemma DOM_cod: assumes "arr f" shows "DOM (cod f) = COD f" using assms Arr_rep Arr_implies_Ide_Cod Ide_implies_Arr cod_char rep_mkarr Par_Arr_norm Ide_in_Hom by simp lemma memb_equiv: assumes "arr f" and "t \ f" and "u \ f" shows "Par t u" and "\<^bold>\t\<^bold>\ = \<^bold>\u\<^bold>\" proof - show "Par t u" using assms Cod_memb Dom_memb mkarr_memb(1) by metis show "\<^bold>\t\<^bold>\ = \<^bold>\u\<^bold>\" using assms arr_iff_ARR ARR_def by auto qed text \ Two arrows can be proved equal by showing that they are parallel and have representatives with identical diagonalizations. \ lemma arr_eqI: assumes "par f g" and "t \ f" and "u \ g" and "\<^bold>\t\<^bold>\ = \<^bold>\u\<^bold>\" shows "f = g" proof - have "Arr t \ Arr u" using assms mkarr_memb(1) by blast moreover have "Dom t = Dom u \ Cod t = Cod u" using assms Dom_memb Cod_memb comp_def arr_char comp_arr_dom comp_cod_arr by (metis (full_types)) ultimately have "Par t u" by simp thus ?thesis using assms arr_char by (metis rep_mkarr rep_in_arr equiv_iff_eq_norm) qed lemma comp_char: shows "f \ g = (if seq f g then mkarr (rep f \<^bold>\ rep g) else null)" using comp_def seq_char arr_char by meson text \ The mapping that takes identity terms to their equivalence classes is injective. \ lemma mkarr_inj_on_Ide: assumes "Ide t" and "Ide u" and "mkarr t = mkarr u" shows "t = u" using assms by (metis (mono_tags, lifting) COD_mkarr Ide_in_Hom mem_Collect_eq) lemma Comp_in_comp [intro]: assumes "arr f" and "g \ hom (dom g) (dom f)" and "t \ f" and "u \ g" shows "t \<^bold>\ u \ f \ g" proof - have "ARR f" using assms arr_iff_ARR by simp moreover have "ARR g" using assms arr_iff_ARR by auto moreover have "DOM f = COD g" using assms dom_char cod_char mkarr_inj_on_Ide Arr_implies_Ide_Cod Arr_implies_Ide_Dom by force ultimately show ?thesis using assms Comp_in_comp_ARR by simp qed text \ An arrow is defined to be ``canonical'' if some (equivalently, all) its representatives diagonalize to an identity term. \ definition can where "can f \ arr f \ (\t. t \ f \ Ide \<^bold>\t\<^bold>\)" lemma can_def_alt: shows "can f \ arr f \ (\t. t \ f \ Ide \<^bold>\t\<^bold>\)" proof assume "arr f \ (\t. t \ f \ Ide \<^bold>\t\<^bold>\)" thus "can f" using can_def arr_char by fastforce next assume f: "can f" show "arr f \ (\t. t \ f \ Ide \<^bold>\t\<^bold>\)" proof - obtain t where t: "t \ f \ Ide \<^bold>\t\<^bold>\" using f can_def by auto have "ARR f" using f can_def arr_char ARR_def mkarr_def by simp hence "\u. u \ f \ \<^bold>\u\<^bold>\ = \<^bold>\t\<^bold>\" using t unique_norm by auto hence "\u. u \ f \ \<^bold>\t\<^bold>\ = \<^bold>\u\<^bold>\" using t by (metis \ARR f\ equiv_iff_eq_norm arr_iff_ARR mkarr_memb(1)) hence "\u. u \ f \ Ide \<^bold>\u\<^bold>\" using t by metis thus ?thesis using f can_def by blast qed qed lemma can_implies_arr: assumes "can f" shows "arr f" using assms can_def by auto text \ The identities of the category are precisely the canonical endo-arrows. \ lemma ide_char: shows "ide f \ can f \ dom f = cod f" proof assume f: "ide f" show "can f \ dom f = cod f" using f can_def arr_char dom_char cod_char IDE_def Arr_implies_Ide_Cod can_def_alt Arr_rep IDE_mkarr_Ide by (metis ideD(1) ideD(3)) next assume f: "can f \ dom f = cod f" show "ide f" proof - have "f = dom f" proof (intro arr_eqI) show "par f (dom f)" using f can_def by simp show "rep f \ f" using f can_def by simp show "DOM f \ dom f" using f can_def by auto show "\<^bold>\rep f\<^bold>\ = \<^bold>\DOM f\<^bold>\" proof - have "\<^bold>\rep f\<^bold>\ \ Hom \<^bold>\DOM f\<^bold>\ \<^bold>\COD f\<^bold>\" using f can_def Diagonalize_in_Hom by simp moreover have "Ide \<^bold>\rep f\<^bold>\" using f can_def_alt rep_in_arr by simp ultimately show ?thesis using f can_def Ide_in_Hom by simp qed qed thus ?thesis using f can_implies_arr ide_dom [of f] by auto qed qed lemma ide_iff_IDE: shows "ide a \ IDE a" using ide_char IDE_def can_def_alt arr_iff_ARR dom_char cod_char mkarr_inj_on_Ide Arr_implies_Ide_Cod Arr_implies_Ide_Dom Arr_rep by auto lemma ide_mkarr_Ide: assumes "Ide a" shows "ide (mkarr a)" using assms IDE_mkarr_Ide ide_iff_IDE by simp lemma rep_dom: assumes "arr f" shows "rep (dom f) = \<^bold>\DOM f\<^bold>\" using assms dom_simp rep_mkarr Arr_rep Arr_implies_Ide_Dom by simp lemma rep_cod: assumes "arr f" shows "rep (cod f) = \<^bold>\COD f\<^bold>\" using assms cod_simp rep_mkarr Arr_rep Arr_implies_Ide_Cod by simp lemma rep_preserves_seq: assumes "seq g f" shows "Seq (rep g) (rep f)" using assms Arr_rep dom_char cod_char mkarr_inj_on_Ide Arr_implies_Ide_Dom Arr_implies_Ide_Cod by auto lemma rep_comp: assumes "seq g f" shows "rep (g \ f) = \<^bold>\rep g \<^bold>\ rep f\<^bold>\" proof - have "rep (g \ f) = rep (mkarr (rep g \<^bold>\ rep f))" using assms comp_char by metis also have "... = \<^bold>\rep g \<^bold>\ rep f\<^bold>\" using assms rep_preserves_seq rep_mkarr by simp finally show ?thesis by blast qed text \ The equivalence classes of canonical terms are canonical arrows. \ lemma can_mkarr_Can: assumes "Can t" shows "can (mkarr t)" using assms Arr_in_mkarr Can_implies_Arr Ide_Diagonalize_Can arr_mkarr can_def by blast lemma ide_implies_can: assumes "ide a" shows "can a" using assms ide_char by blast lemma Can_rep_can: assumes "can f" shows "Can (rep f)" proof - have "Can \<^bold>\rep f\<^bold>\" using assms can_def_alt Can_norm_iff_Ide_Diagonalize by auto moreover have "rep f = \<^bold>\rep f\<^bold>\" using assms can_implies_arr norm_rep by simp ultimately show ?thesis by simp qed text \ Parallel canonical arrows are identical. \ lemma can_coherence: assumes "par f g" and "can f" and "can g" shows "f = g" proof - have "\<^bold>\rep f\<^bold>\ = \<^bold>\rep g\<^bold>\" proof - have "\<^bold>\rep f\<^bold>\ = \<^bold>\DOM f\<^bold>\" using assms Ide_Diagonalize_Can Can_rep_can Diagonalize_in_Hom Ide_in_Hom by force also have "... = \<^bold>\DOM g\<^bold>\" using assms dom_char equiv_iff_eq_norm by (metis DOM_in_dom mkarr_memb(1) rep_mkarr arr_dom_iff_arr) also have "... = \<^bold>\rep g\<^bold>\" using assms Ide_Diagonalize_Can Can_rep_can Diagonalize_in_Hom Ide_in_Hom by force finally show ?thesis by blast qed hence "rep f = rep g" using assms rep_in_arr norm_memb_eq_rep equiv_iff_eq_norm by (metis (no_types, lifting) arr_eqI) thus ?thesis using assms arr_eqI [of f g] rep_in_arr [of f] rep_in_arr [of g] by metis qed text \ Canonical arrows are invertible, and their inverses can be obtained syntactically. \ lemma inverse_arrows_can: assumes "can f" shows "inverse_arrows f (mkarr (Inv (DOM f\<^bold>\) \<^bold>\ \<^bold>\rep f\<^bold>\ \<^bold>\ COD f\<^bold>\))" proof let ?t = "(Inv (DOM f\<^bold>\) \<^bold>\ \<^bold>\rep f\<^bold>\ \<^bold>\ COD f\<^bold>\)" have 1: "rep f \ f \ Arr (rep f) \ Can (rep f) \ Ide \<^bold>\rep f\<^bold>\" using assms can_def_alt rep_in_arr rep_in_arr(1) Can_rep_can by simp hence 2: "\<^bold>\DOM f\<^bold>\ = \<^bold>\COD f\<^bold>\" using Diagonalize_in_Hom [of "rep f"] Ide_in_Hom by auto have 3: "Can ?t" using assms 1 2 Can_red Ide_implies_Can Diagonalize_in_Hom Inv_preserves_Can Arr_implies_Ide_Cod Arr_implies_Ide_Dom Diag_Diagonalize by simp have 4: "DOM f = Cod ?t" using assms can_def Can_red by (simp add: Arr_implies_Ide_Dom Inv_preserves_Can(3)) have 5: "COD f = Dom ?t" using assms can_def Can_red Arr_rep Arr_implies_Ide_Cod by simp have 6: "antipar f (mkarr ?t)" using assms 3 4 5 dom_char cod_char can_def cod_mkarr dom_mkarr Can_implies_Arr by simp show "ide (f \ mkarr ?t)" proof - have 7: "par (f \ mkarr ?t) (dom (f \ mkarr ?t))" using assms 6 by auto moreover have "can (f \ mkarr ?t)" proof - have 8: "Comp (rep f) ?t \ (f \ mkarr ?t)" using assms 1 3 4 6 can_implies_arr Arr_in_mkarr COD_mkarr Comp_in_comp_ARR Can_implies_Arr arr_iff_ARR seq_char' by meson moreover have "Can (rep f \<^bold>\ ?t)" using 1 3 7 8 mkarr_memb(1) by (metis Arr.simps(4) Can.simps(4)) ultimately show ?thesis using can_mkarr_Can 7 mkarr_memb(2) by metis qed moreover have "can (dom (f \ mkarr ?t))" using 7 ide_implies_can by force ultimately have "f \ mkarr ?t = dom (f \ mkarr ?t)" using can_coherence by meson thus ?thesis using 7 ide_dom by metis qed show "ide (mkarr ?t \ f)" proof - have 7: "par (mkarr ?t \ f) (cod (mkarr ?t \ f))" using assms 6 by auto moreover have "can (mkarr ?t \ f)" proof - have 8: "Comp ?t (rep f) \ mkarr ?t \ f" using assms 1 3 6 7 Arr_in_mkarr Comp_in_comp_ARR Can_implies_Arr arr_char comp_def by meson moreover have "Can (?t \<^bold>\ rep f)" using 1 3 7 8 mkarr_memb(1) by (metis Arr.simps(4) Can.simps(4)) ultimately show ?thesis using can_mkarr_Can 7 mkarr_memb(2) by metis qed moreover have "can (cod (mkarr ?t \ f))" using 7 ide_implies_can by force ultimately have "mkarr ?t \ f = cod (mkarr ?t \ f)" using can_coherence by meson thus ?thesis using 7 can_implies_arr ide_cod by metis qed qed lemma inv_mkarr [simp]: assumes "Can t" shows "inv (mkarr t) = mkarr (Inv t)" proof - have t: "Can t \ Arr t \ Can (Inv t) \ Arr (Inv t) \ Ide (Dom t) \ Ide (Cod t)" using assms Can_implies_Arr Arr_implies_Ide_Dom Arr_implies_Ide_Cod Inv_preserves_Can by simp have "inverse_arrows (mkarr t) (mkarr (Inv t))" proof show "ide (mkarr t \ mkarr (Inv t))" proof - have "mkarr (Cod t) = mkarr (Comp t (Inv t))" using t Inv_in_Hom Ide_in_Hom Diagonalize_Inv Diag_Diagonalize Diagonalize_preserves_Can by (intro mkarr_eqI, auto) also have "... = mkarr t \ mkarr (Inv t)" using t comp_mkarr Inv_in_Hom by simp finally have "mkarr (Cod t) = mkarr t \ mkarr (Inv t)" by blast thus ?thesis using t ide_mkarr_Ide [of "Cod t"] by simp qed show "ide (mkarr (Inv t) \ mkarr t)" proof - have "mkarr (Dom t) = mkarr (Inv t \<^bold>\ t)" using t Inv_in_Hom Ide_in_Hom Diagonalize_Inv Diag_Diagonalize Diagonalize_preserves_Can by (intro mkarr_eqI, auto) also have "... = mkarr (Inv t) \ mkarr t" using t comp_mkarr Inv_in_Hom by simp finally have "mkarr (Dom t) = mkarr (Inv t) \ mkarr t" by blast thus ?thesis using t ide_mkarr_Ide [of "Dom t"] by simp qed qed thus ?thesis using inverse_unique by auto qed lemma iso_can: assumes "can f" shows "iso f" using assms inverse_arrows_can by auto text \ The following function produces the unique canonical arrow between two given objects, if such an arrow exists. \ definition mkcan where "mkcan a b = mkarr (Inv (COD b\<^bold>\) \<^bold>\ (DOM a\<^bold>\))" lemma can_mkcan: assumes "ide a" and "ide b" and "\<^bold>\DOM a\<^bold>\ = \<^bold>\COD b\<^bold>\" shows "can (mkcan a b)" and "\mkcan a b : a \ b\" proof - show "can (mkcan a b)" using assms mkcan_def Arr_rep Arr_implies_Ide_Dom Arr_implies_Ide_Cod Can_red Inv_preserves_Can can_mkarr_Can by simp show "\mkcan a b : a \ b\" using assms mkcan_def Arr_rep Arr_implies_Ide_Dom Arr_implies_Ide_Cod Can_red Inv_in_Hom dom_char [of a] cod_char [of b] mkarr_rep mkarr_in_hom can_implies_arr by auto qed lemma dom_mkcan: assumes "ide a" and "ide b" and "\<^bold>\DOM a\<^bold>\ = \<^bold>\COD b\<^bold>\" shows "dom (mkcan a b) = a" using assms can_mkcan by blast lemma cod_mkcan: assumes "ide a" and "ide b" and "\<^bold>\DOM a\<^bold>\ = \<^bold>\COD b\<^bold>\" shows "cod (mkcan a b) = b" using assms can_mkcan by blast lemma can_coherence': assumes "can f" shows "mkcan (dom f) (cod f) = f" proof - have "Ide \<^bold>\rep f\<^bold>\" using assms Ide_Diagonalize_Can Can_rep_can by simp hence "Dom \<^bold>\rep f\<^bold>\ = Cod \<^bold>\rep f\<^bold>\" using Ide_in_Hom by simp hence "\<^bold>\DOM f\<^bold>\ = \<^bold>\COD f\<^bold>\" using assms can_implies_arr Arr_rep Diagonalize_in_Hom by simp moreover have "DOM f = DOM (dom f)" using assms can_implies_arr dom_char rep_mkarr Arr_implies_Ide_Dom Ide_implies_Arr Par_Arr_norm [of "DOM f"] Ide_in_Hom by auto moreover have "COD f = COD (cod f)" using assms can_implies_arr cod_char rep_mkarr Arr_implies_Ide_Cod Ide_implies_Arr Par_Arr_norm [of "COD f"] Ide_in_Hom by auto ultimately have "can (mkcan (dom f) (cod f)) \ par f (mkcan (dom f) (cod f))" using assms can_implies_arr can_mkcan dom_mkcan cod_mkcan by simp thus ?thesis using assms can_coherence by blast qed lemma Ide_Diagonalize_rep_ide: assumes "ide a" shows "Ide \<^bold>\rep a\<^bold>\" using assms ide_implies_can can_def_alt rep_in_arr by simp lemma Diagonalize_DOM: assumes "arr f" shows "\<^bold>\DOM f\<^bold>\ = Dom \<^bold>\rep f\<^bold>\" using assms Diag_Diagonalize by simp lemma Diagonalize_COD: assumes "arr f" shows "\<^bold>\COD f\<^bold>\ = Cod \<^bold>\rep f\<^bold>\" using assms Diag_Diagonalize by simp lemma Diagonalize_rep_preserves_seq: assumes "seq g f" shows "Seq \<^bold>\rep g\<^bold>\ \<^bold>\rep f\<^bold>\" using assms Diagonalize_DOM Diagonalize_COD Diag_implies_Arr Diag_Diagonalize(1) rep_preserves_seq by force lemma Dom_Diagonalize_rep: assumes "arr f" shows "Dom \<^bold>\rep f\<^bold>\ = \<^bold>\rep (dom f)\<^bold>\" using assms Diagonalize_rep_preserves_seq [of f "dom f"] Ide_Diagonalize_rep_ide Ide_in_Hom by simp lemma Cod_Diagonalize_rep: assumes "arr f" shows "Cod \<^bold>\rep f\<^bold>\ = \<^bold>\rep (cod f)\<^bold>\" using assms Diagonalize_rep_preserves_seq [of "cod f" f] Ide_Diagonalize_rep_ide Ide_in_Hom by simp lemma mkarr_Diagonalize_rep: assumes "arr f" and "Diag (DOM f)" and "Diag (COD f)" shows "mkarr \<^bold>\rep f\<^bold>\ = f" proof - have "mkarr (rep f) = mkarr \<^bold>\rep f\<^bold>\" using assms rep_in_Hom Diagonalize_in_Hom Diag_Diagonalize Diagonalize_Diag by (intro mkarr_eqI, simp_all) thus ?thesis using assms mkarr_rep by auto qed text \ We define tensor product of arrows via the constructor @{term Tensor} on terms. \ definition tensor\<^sub>F\<^sub>M\<^sub>C (infixr "\" 53) where "f \ g \ (if arr f \ arr g then mkarr (rep f \<^bold>\ rep g) else null)" lemma arr_tensor [simp]: assumes "arr f" and "arr g" shows "arr (f \ g)" using assms tensor\<^sub>F\<^sub>M\<^sub>C_def arr_mkarr by simp lemma rep_tensor: assumes "arr f" and "arr g" shows "rep (f \ g) = \<^bold>\rep f \<^bold>\ rep g\<^bold>\" using assms tensor\<^sub>F\<^sub>M\<^sub>C_def rep_mkarr by simp lemma Par_memb_rep: assumes "arr f" and "t \ f" shows "Par t (rep f)" using assms mkarr_memb apply simp using rep_in_Hom Dom_memb Cod_memb by metis lemma Tensor_in_tensor [intro]: assumes "arr f" and "arr g" and "t \ f" and "u \ g" shows "t \<^bold>\ u \ f \ g" proof - have "equiv (t \<^bold>\ u) (rep f \<^bold>\ rep g)" proof - have 1: "Par (t \<^bold>\ u) (rep f \<^bold>\ rep g)" proof - have "Par t (rep f) \ Par u (rep g)" using assms Par_memb_rep by blast thus ?thesis by simp qed moreover have "\<^bold>\t \<^bold>\ u\<^bold>\ = \<^bold>\rep f \<^bold>\ rep g\<^bold>\" using assms 1 equiv_iff_eq_norm rep_mkarr norm_norm mkarr_memb(2) by (metis Arr.simps(3) Diagonalize.simps(3)) ultimately show ?thesis by simp qed thus ?thesis using assms tensor\<^sub>F\<^sub>M\<^sub>C_def mkarr_def by simp qed lemma DOM_tensor [simp]: assumes "arr f" and "arr g" shows "DOM (f \ g) = DOM f \<^bold>\ DOM g" by (metis (no_types, lifting) DOM_mkarr Dom.simps(3) mkarr_extensionality arr_char arr_tensor assms(1) assms(2) tensor\<^sub>F\<^sub>M\<^sub>C_def) lemma COD_tensor [simp]: assumes "arr f" and "arr g" shows "COD (f \ g) = COD f \<^bold>\ COD g" by (metis (no_types, lifting) COD_mkarr Cod.simps(3) mkarr_extensionality arr_char arr_tensor assms(1) assms(2) tensor\<^sub>F\<^sub>M\<^sub>C_def) lemma tensor_in_hom [simp]: assumes "\f : a \ b\" and "\g : c \ d\" shows "\f \ g : a \ c \ b \ d\" proof - have f: "arr f \ dom f = a \ cod f = b" using assms(1) by auto have g: "arr g \ dom g = c \ cod g = d" using assms(2) by auto have "dom (f \ g) = dom f \ dom g" using f g arr_tensor dom_char Tensor_in_tensor [of "dom f" "dom g" "DOM f" "DOM g"] DOM_in_dom mkarr_memb(2) DOM_tensor arr_dom_iff_arr by metis moreover have "cod (f \ g) = cod f \ cod g" using f g arr_tensor cod_char Tensor_in_tensor [of "cod f" "cod g" "COD f" "COD g"] COD_in_cod mkarr_memb(2) COD_tensor arr_cod_iff_arr by metis ultimately show ?thesis using assms arr_tensor by blast qed lemma dom_tensor [simp]: assumes "arr f" and "arr g" shows "dom (f \ g) = dom f \ dom g" using assms tensor_in_hom [of f] by blast lemma cod_tensor [simp]: assumes "arr f" and "arr g" shows "cod (f \ g) = cod f \ cod g" using assms tensor_in_hom [of f] by blast lemma tensor_mkarr [simp]: assumes "Arr t" and "Arr u" shows "mkarr t \ mkarr u = mkarr (t \<^bold>\ u)" using assms by (meson Tensor_in_tensor arr_char Arr_in_mkarr arr_mkarr arr_tensor) lemma tensor_preserves_ide: assumes "ide a" and "ide b" shows "ide (a \ b)" proof - have "can (a \ b)" using assms tensor\<^sub>F\<^sub>M\<^sub>C_def Can_rep_can ide_implies_can can_mkarr_Can by simp moreover have "dom (a \ b) = cod (a \ b)" using assms tensor_in_hom by simp ultimately show ?thesis using ide_char by metis qed lemma tensor_preserves_can: assumes "can f" and "can g" shows "can (f \ g)" using assms can_implies_arr Can_rep_can tensor\<^sub>F\<^sub>M\<^sub>C_def can_mkarr_Can by simp lemma comp_preserves_can: assumes "can f" and "can g" and "dom f = cod g" shows "can (f \ g)" proof - have 1: "ARR f \ ARR g \ DOM f = COD g" using assms can_implies_arr arr_iff_ARR Arr_implies_Ide_Cod Arr_implies_Ide_Dom mkarr_inj_on_Ide cod_char dom_char by simp hence "Can (rep f \<^bold>\ rep g)" using assms can_implies_arr Can_rep_can by force thus ?thesis using assms 1 can_implies_arr comp_char can_mkarr_Can seq_char' by simp qed text \ The remaining structure required of a monoidal category is also defined syntactically. \ definition unity\<^sub>F\<^sub>M\<^sub>C :: "'c arr" ("\") where "\ = mkarr \<^bold>\" definition lunit\<^sub>F\<^sub>M\<^sub>C :: "'c arr \ 'c arr" ("\[_]") where "\[a] = mkarr \<^bold>\\<^bold>[rep a\<^bold>]" definition runit\<^sub>F\<^sub>M\<^sub>C :: "'c arr \ 'c arr" ("\[_]") where "\[a] = mkarr \<^bold>\\<^bold>[rep a\<^bold>]" definition assoc\<^sub>F\<^sub>M\<^sub>C :: "'c arr \ 'c arr \ 'c arr \ 'c arr" ("\[_, _, _]") where "\[a, b, c] = mkarr \<^bold>\\<^bold>[rep a, rep b, rep c\<^bold>]" lemma can_lunit: assumes "ide a" shows "can \[a]" using assms lunit\<^sub>F\<^sub>M\<^sub>C_def can_mkarr_Can by (simp add: Can_rep_can ide_implies_can) lemma lunit_in_hom: assumes "ide a" shows "\\[a] : \ \ a \ a\" proof - have "dom \[a] = \ \ a" using assms lunit\<^sub>F\<^sub>M\<^sub>C_def unity\<^sub>F\<^sub>M\<^sub>C_def Ide_implies_Arr dom_mkarr dom_char tensor_mkarr Arr_rep by (metis Arr.simps(2) Arr.simps(5) Arr_implies_Ide_Dom Dom.simps(5) ideD(1) ideD(2)) moreover have "cod \[a] = a" using assms lunit\<^sub>F\<^sub>M\<^sub>C_def rep_in_arr(1) cod_mkarr cod_char ideD(3) by auto ultimately show ?thesis using assms arr_cod_iff_arr by (intro in_homI, fastforce) qed lemma arr_lunit [simp]: assumes "ide a" shows "arr \[a]" using assms can_lunit can_implies_arr by simp lemma dom_lunit [simp]: assumes "ide a" shows "dom \[a] = \ \ a" using assms lunit_in_hom by auto lemma cod_lunit [simp]: assumes "ide a" shows "cod \[a] = a" using assms lunit_in_hom by auto lemma can_runit: assumes "ide a" shows "can \[a]" using assms runit\<^sub>F\<^sub>M\<^sub>C_def can_mkarr_Can by (simp add: Can_rep_can ide_implies_can) lemma runit_in_hom [simp]: assumes "ide a" shows "\\[a] : a \ \ \ a\" proof - have "dom \[a] = a \ \" using assms Arr_rep Arr.simps(2) Arr.simps(7) Arr_implies_Ide_Dom Dom.simps(7) Ide_implies_Arr dom_mkarr dom_char ideD(1) ideD(2) runit\<^sub>F\<^sub>M\<^sub>C_def tensor_mkarr unity\<^sub>F\<^sub>M\<^sub>C_def by metis moreover have "cod \[a] = a" using assms runit\<^sub>F\<^sub>M\<^sub>C_def rep_in_arr(1) cod_mkarr cod_char ideD(3) by auto ultimately show ?thesis using assms arr_cod_iff_arr by (intro in_homI, fastforce) qed lemma arr_runit [simp]: assumes "ide a" shows "arr \[a]" using assms can_runit can_implies_arr by simp lemma dom_runit [simp]: assumes "ide a" shows "dom \[a] = a \ \" using assms runit_in_hom by blast lemma cod_runit [simp]: assumes "ide a" shows "cod \[a] = a" using assms runit_in_hom by blast lemma can_assoc: assumes "ide a" and "ide b" and "ide c" shows "can \[a, b, c]" using assms assoc\<^sub>F\<^sub>M\<^sub>C_def can_mkarr_Can by (simp add: Can_rep_can ide_implies_can) lemma assoc_in_hom: assumes "ide a" and "ide b" and "ide c" shows "\\[a, b, c] : (a \ b) \ c \ a \ b \ c\" proof - have "dom \[a, b, c] = (a \ b) \ c" proof - have "dom \[a, b, c] = mkarr (Dom \<^bold>\\<^bold>[rep a, rep b, rep c\<^bold>])" using assms assoc\<^sub>F\<^sub>M\<^sub>C_def rep_in_arr(1) by simp also have "... = mkarr ((DOM a \<^bold>\ DOM b) \<^bold>\ DOM c)" by simp also have "... = (a \ b) \ c" by (metis mkarr_extensionality arr_tensor assms dom_char ideD(1) ideD(2) not_arr_null null_char tensor_mkarr) finally show ?thesis by blast qed moreover have "cod \[a, b, c] = a \ b \ c" proof - have "cod \[a, b, c] = mkarr (Cod \<^bold>\\<^bold>[rep a, rep b, rep c\<^bold>])" using assms assoc\<^sub>F\<^sub>M\<^sub>C_def rep_in_arr(1) by simp also have "... = mkarr (COD a \<^bold>\ COD b \<^bold>\ COD c)" by simp also have "... = a \ b \ c" by (metis mkarr_extensionality arr_tensor assms(1) assms(2) assms(3) cod_char ideD(1) ideD(3) not_arr_null null_char tensor_mkarr) finally show ?thesis by blast qed moreover have "arr \[a, b, c]" using assms assoc\<^sub>F\<^sub>M\<^sub>C_def rep_in_arr(1) arr_mkarr by simp ultimately show ?thesis by auto qed lemma arr_assoc [simp]: assumes "ide a" and "ide b" and "ide c" shows "arr \[a, b, c]" using assms can_assoc can_implies_arr by simp lemma dom_assoc [simp]: assumes "ide a" and "ide b" and "ide c" shows "dom \[a, b, c] = (a \ b) \ c" using assms assoc_in_hom by blast lemma cod_assoc [simp]: assumes "ide a" and "ide b" and "ide c" shows "cod \[a, b, c] = a \ b \ c" using assms assoc_in_hom by blast lemma ide_unity [simp]: shows "ide \" using unity\<^sub>F\<^sub>M\<^sub>C_def Arr.simps(2) Dom.simps(2) arr_mkarr dom_mkarr ide_dom by metis lemma Unity_in_unity [simp]: shows "\<^bold>\ \ \" using unity\<^sub>F\<^sub>M\<^sub>C_def Arr_in_mkarr by simp lemma rep_unity [simp]: shows "rep \ = \<^bold>\\<^bold>\\<^bold>\" using unity\<^sub>F\<^sub>M\<^sub>C_def rep_mkarr by simp lemma Lunit_in_lunit [intro]: assumes "arr f" and "t \ f" shows "\<^bold>\\<^bold>[t\<^bold>] \ \[f]" proof - have "Arr t \ Arr (rep f) \ Dom t = DOM f \ Cod t = COD f \ \<^bold>\t\<^bold>\ = \<^bold>\rep f\<^bold>\" using assms by (metis mkarr_memb(1) mkarr_memb(2) rep_mkarr rep_in_arr(1) equiv_iff_eq_norm norm_rep) thus ?thesis by (simp add: mkarr_def lunit\<^sub>F\<^sub>M\<^sub>C_def) qed lemma Runit_in_runit [intro]: assumes "arr f" and "t \ f" shows "\<^bold>\\<^bold>[t\<^bold>] \ \[f]" proof - have "Arr t \ Arr (rep f) \ Dom t = DOM f \ Cod t = COD f \ \<^bold>\t\<^bold>\ = \<^bold>\rep f\<^bold>\" using assms by (metis mkarr_memb(1) mkarr_memb(2) rep_mkarr rep_in_arr(1) equiv_iff_eq_norm norm_rep) thus ?thesis by (simp add: mkarr_def runit\<^sub>F\<^sub>M\<^sub>C_def) qed lemma Assoc_in_assoc [intro]: assumes "arr f" and "arr g" and "arr h" and "t \ f" and "u \ g" and "v \ h" shows "\<^bold>\\<^bold>[t, u, v\<^bold>] \ \[f, g, h]" proof - have "Arr t \ Arr (rep f) \ Dom t = DOM f \ Cod t = COD f \ \<^bold>\t\<^bold>\ = \<^bold>\rep f\<^bold>\" using assms by (metis mkarr_memb(1) rep_mkarr rep_in_arr(1) equiv_iff_eq_norm mkarr_memb(2) norm_rep) moreover have "Arr u \ Arr (rep g) \ Dom u = DOM g \ Cod u = COD g \ \<^bold>\u\<^bold>\ = \<^bold>\rep g\<^bold>\" using assms by (metis mkarr_memb(1) rep_mkarr rep_in_arr(1) equiv_iff_eq_norm mkarr_memb(2) norm_rep) moreover have "Arr v \ Arr (rep h) \ Dom v = DOM h \ Cod v = COD h \ \<^bold>\v\<^bold>\ = \<^bold>\rep h\<^bold>\" using assms by (metis mkarr_memb(1) rep_mkarr rep_in_arr(1) equiv_iff_eq_norm mkarr_memb(2) norm_rep) ultimately show ?thesis using assoc\<^sub>F\<^sub>M\<^sub>C_def mkarr_def by simp qed text \ At last, we can show that we've constructed a monoidal category. \ interpretation EMC: elementary_monoidal_category comp tensor\<^sub>F\<^sub>M\<^sub>C unity\<^sub>F\<^sub>M\<^sub>C lunit\<^sub>F\<^sub>M\<^sub>C runit\<^sub>F\<^sub>M\<^sub>C assoc\<^sub>F\<^sub>M\<^sub>C proof show "ide \" using ide_unity by auto show "\a. ide a \ \\[a] : \ \ a \ a\" by auto show "\a. ide a \ \\[a] : a \ \ \ a\" by auto show "\a. ide a \ iso \[a]" using can_lunit iso_can by auto show "\a. ide a \ iso \[a]" using can_runit iso_can by auto show "\a b c. \ ide a; ide b; ide c \ \ \\[a, b, c] : (a \ b) \ c \ a \ b \ c\" by auto show "\a b c. \ ide a; ide b; ide c \ \ iso \[a, b, c]" using can_assoc iso_can by auto show "\a b. \ ide a; ide b \ \ ide (a \ b)" using tensor_preserves_ide by auto fix f a b g c d show "\ \f : a \ b\; \g : c \ d\ \ \ \f \ g : a \ c \ b \ d\" using tensor_in_hom by auto next text \Naturality of left unitor.\ fix f assume f: "arr f" show "\[cod f] \ (\ \ f) = f \ \[dom f]" proof (intro arr_eqI) show "par (\[cod f] \ (\ \ f)) (f \ \[dom f])" using f by simp show "\<^bold>\\<^bold>[COD f\<^bold>] \<^bold>\ (\<^bold>\ \<^bold>\ rep f) \ \[cod f] \ (\ \ f)" using f by fastforce show "rep f \<^bold>\ \<^bold>\\<^bold>[DOM f\<^bold>] \ f \ \[dom f]" using f by fastforce show "\<^bold>\\<^bold>\\<^bold>[COD f\<^bold>] \<^bold>\ (\<^bold>\ \<^bold>\ rep f)\<^bold>\ = \<^bold>\rep f \<^bold>\ \<^bold>\\<^bold>[DOM f\<^bold>]\<^bold>\" using f by (simp add: Diag_Diagonalize(1) Diagonalize_DOM Diagonalize_COD) qed text \Naturality of right unitor.\ show "\[cod f] \ (f \ \) = f \ \[dom f]" proof (intro arr_eqI) show "par (\[cod f] \ (f \ \)) (f \ \[dom f])" using f by simp show "\<^bold>\\<^bold>[COD f\<^bold>] \<^bold>\ (rep f \<^bold>\ \<^bold>\) \ \[cod f] \ (f \ \)" using f by fastforce show "rep f \<^bold>\ \<^bold>\\<^bold>[DOM f\<^bold>] \ f \ \[dom f]" using f by fastforce show "\<^bold>\\<^bold>\\<^bold>[COD f\<^bold>] \<^bold>\ (rep f \<^bold>\ \<^bold>\)\<^bold>\ = \<^bold>\rep f \<^bold>\ \<^bold>\\<^bold>[DOM f\<^bold>]\<^bold>\" using f by (simp add: Diag_Diagonalize(1) Diagonalize_DOM Diagonalize_COD) qed next text \Naturality of associator.\ fix f0 :: "'c arr" and f1 f2 assume f0: "arr f0" and f1: "arr f1" and f2: "arr f2" show "\[cod f0, cod f1, cod f2] \ ((f0 \ f1) \ f2) = (f0 \ f1 \ f2) \ \[dom f0, dom f1, dom f2]" proof (intro arr_eqI) show 1: "par (\[cod f0, cod f1, cod f2] \ ((f0 \ f1) \ f2)) ((f0 \ f1 \ f2) \ \[dom f0, dom f1, dom f2])" using f0 f1 f2 by force show "\<^bold>\\<^bold>[rep (cod f0), rep (cod f1), rep (cod f2)\<^bold>] \<^bold>\ ((rep f0 \<^bold>\ rep f1) \<^bold>\ rep f2) \ \[cod f0, cod f1, cod f2] \ ((f0 \ f1) \ f2)" using f0 f1 f2 by fastforce show "(rep f0 \<^bold>\ rep f1 \<^bold>\ rep f2) \<^bold>\ \<^bold>\\<^bold>[rep (dom f0), rep (dom f1), rep (dom f2)\<^bold>] \ (f0 \ f1 \ f2) \ \[dom f0, dom f1, dom f2]" using f0 f1 f2 by fastforce show "\<^bold>\\<^bold>\\<^bold>[rep (cod f0), rep (cod f1), rep (cod f2)\<^bold>] \<^bold>\ ((rep f0 \<^bold>\ rep f1) \<^bold>\ rep f2)\<^bold>\ = \<^bold>\(rep f0 \<^bold>\ rep f1 \<^bold>\ rep f2) \<^bold>\ \<^bold>\\<^bold>[rep (dom f0), rep (dom f1), rep (dom f2)\<^bold>]\<^bold>\" proof - have "\<^bold>\\<^bold>\\<^bold>[rep (cod f0), rep (cod f1), rep (cod f2)\<^bold>] \<^bold>\ ((rep f0 \<^bold>\ rep f1) \<^bold>\ rep f2)\<^bold>\ = \<^bold>\rep f0\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\rep f1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\rep f2\<^bold>\" proof - have b0: "\<^bold>\rep (cod f0)\<^bold>\ = Cod \<^bold>\rep f0\<^bold>\" using f0 Cod_Diagonalize_rep by simp have b1: "\<^bold>\rep (cod f1)\<^bold>\ = Cod \<^bold>\rep f1\<^bold>\" using f1 Cod_Diagonalize_rep by simp have b2: "\<^bold>\rep (cod f2)\<^bold>\ = Cod \<^bold>\rep f2\<^bold>\" using f2 Cod_Diagonalize_rep by simp have "\<^bold>\\<^bold>\\<^bold>[rep (cod f0), rep (cod f1), rep (cod f2)\<^bold>] \<^bold>\ ((rep f0 \<^bold>\ rep f1) \<^bold>\ rep f2)\<^bold>\ = (\<^bold>\rep (cod f0)\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\rep (cod f1)\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\rep (cod f2)\<^bold>\) \<^bold>\\<^bold>\\<^bold>\ (\<^bold>\rep f0\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\rep f1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\rep f2\<^bold>\)" using f0 f1 f2 using Diag_Diagonalize(1) TensorDiag_assoc by auto also have "... = \<^bold>\rep (cod f0)\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\rep f0\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\rep (cod f1)\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\rep f1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\rep (cod f2)\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\rep f2\<^bold>\" proof - have "Seq \<^bold>\rep (cod f0)\<^bold>\ \<^bold>\rep f0\<^bold>\ \ Seq \<^bold>\rep (cod f1)\<^bold>\ \<^bold>\rep f1\<^bold>\ \ Seq \<^bold>\rep (cod f2)\<^bold>\ \<^bold>\rep f2\<^bold>\" using f0 f1 f2 rep_in_Hom Diagonalize_in_Hom Dom_Diagonalize_rep Cod_Diagonalize_rep by auto thus ?thesis using f0 f1 f2 b0 b1 b2 TensorDiag_in_Hom TensorDiag_preserves_Diag Diag_Diagonalize Arr_implies_Ide_Dom Arr_implies_Ide_Cod CompDiag_TensorDiag by simp qed also have "... = \<^bold>\rep f0\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\rep f1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\rep f2\<^bold>\" proof - have "\<^bold>\rep (cod f0)\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\rep f0\<^bold>\ = \<^bold>\rep f0\<^bold>\" using f0 b0 CompDiag_Cod_Diag [of "\<^bold>\rep f0\<^bold>\"] Diag_Diagonalize by simp moreover have "\<^bold>\rep (cod f1)\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\rep f1\<^bold>\ = \<^bold>\rep f1\<^bold>\" using f1 b1 CompDiag_Cod_Diag [of "\<^bold>\rep f1\<^bold>\"] Diag_Diagonalize by simp moreover have "\<^bold>\rep (cod f2)\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\rep f2\<^bold>\ = \<^bold>\rep f2\<^bold>\" using f2 b2 CompDiag_Cod_Diag [of "\<^bold>\rep f2\<^bold>\"] Diag_Diagonalize by simp ultimately show ?thesis by argo qed finally show ?thesis by blast qed also have "... = \<^bold>\(rep f0 \<^bold>\ rep f1 \<^bold>\ rep f2) \<^bold>\ \<^bold>\\<^bold>[rep (dom f0), rep (dom f1), rep (dom f2)\<^bold>]\<^bold>\" proof - have a0: "\<^bold>\rep (dom f0)\<^bold>\ = Dom \<^bold>\rep f0\<^bold>\" using f0 Dom_Diagonalize_rep by simp have a1: "\<^bold>\rep (dom f1)\<^bold>\ = Dom \<^bold>\rep f1\<^bold>\" using f1 Dom_Diagonalize_rep by simp have a2: "\<^bold>\rep (dom f2)\<^bold>\ = Dom \<^bold>\rep f2\<^bold>\" using f2 Dom_Diagonalize_rep by simp have "\<^bold>\(rep f0 \<^bold>\ rep f1 \<^bold>\ rep f2) \<^bold>\ \<^bold>\\<^bold>[rep (dom f0), rep (dom f1), rep (dom f2)\<^bold>]\<^bold>\ = (\<^bold>\rep f0\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\rep f1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\rep f2\<^bold>\) \<^bold>\\<^bold>\\<^bold>\ (\<^bold>\rep (dom f0)\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\rep (dom f1)\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\rep (dom f2)\<^bold>\)" using f0 f1 f2 using Diag_Diagonalize(1) TensorDiag_assoc by auto also have "... = \<^bold>\rep f0\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\rep (dom f0)\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\rep f1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\rep (dom f1)\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\rep f2\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\rep (dom f2)\<^bold>\" proof - have "Seq \<^bold>\rep f0\<^bold>\ \<^bold>\rep (dom f0)\<^bold>\ \ Seq \<^bold>\rep f1\<^bold>\ \<^bold>\rep (dom f1)\<^bold>\ \ Seq \<^bold>\rep f2\<^bold>\ \<^bold>\rep (dom f2)\<^bold>\" using f0 f1 f2 rep_in_Hom Diagonalize_in_Hom Dom_Diagonalize_rep Cod_Diagonalize_rep by auto thus ?thesis using f0 f1 f2 a0 a1 a2 TensorDiag_in_Hom TensorDiag_preserves_Diag Diag_Diagonalize Arr_implies_Ide_Dom Arr_implies_Ide_Cod CompDiag_TensorDiag by force qed also have "... = \<^bold>\rep f0\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\rep f1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\rep f2\<^bold>\" proof - have "\<^bold>\rep f0\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\rep (dom f0)\<^bold>\ = \<^bold>\rep f0\<^bold>\" using f0 a0 CompDiag_Diag_Dom [of "Diagonalize (rep f0)"] Diag_Diagonalize by simp moreover have "\<^bold>\rep f1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\rep (dom f1)\<^bold>\ = \<^bold>\rep f1\<^bold>\" using f1 a1 CompDiag_Diag_Dom [of "Diagonalize (rep f1)"] Diag_Diagonalize by simp moreover have "\<^bold>\rep f2\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\rep (dom f2)\<^bold>\ = \<^bold>\rep f2\<^bold>\" using f2 a2 CompDiag_Diag_Dom [of "Diagonalize (rep f2)"] Diag_Diagonalize by simp ultimately show ?thesis by argo qed finally show ?thesis by argo qed finally show ?thesis by blast qed qed next text \Tensor preserves composition (interchange).\ fix f g f' g' show "\ seq g f; seq g' f' \ \ (g \ g') \ (f \ f') = g \ f \ g' \ f'" proof - assume gf: "seq g f" assume gf': "seq g' f'" show ?thesis proof (intro arr_eqI) show "par ((g \ g') \ (f \ f')) (g \ f \ g' \ f')" using gf gf' by fastforce show "(rep g \<^bold>\ rep g') \<^bold>\ (rep f \<^bold>\ rep f') \ (g \ g') \ (f \ f')" using gf gf' by force show "rep g \<^bold>\ rep f \<^bold>\ rep g' \<^bold>\ rep f' \ g \ f \ g' \ f'" using gf gf' by (meson Comp_in_comp_ARR Tensor_in_tensor rep_in_arr seqE seq_char') show "\<^bold>\(rep g \<^bold>\ rep g') \<^bold>\ (rep f \<^bold>\ rep f')\<^bold>\ = \<^bold>\rep g \<^bold>\ rep f \<^bold>\ rep g' \<^bold>\ rep f'\<^bold>\" proof - have "\<^bold>\(rep g \<^bold>\ rep g') \<^bold>\ (rep f \<^bold>\ rep f')\<^bold>\ = (\<^bold>\rep g\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\rep g'\<^bold>\) \<^bold>\\<^bold>\\<^bold>\ (\<^bold>\rep f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\rep f'\<^bold>\)" by auto also have "... = \<^bold>\rep g\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\rep f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\rep g'\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\rep f'\<^bold>\" using gf gf' Arr_rep Diagonalize_rep_preserves_seq CompDiag_TensorDiag [of "\<^bold>\rep g\<^bold>\" " \<^bold>\rep g'\<^bold>\" "\<^bold>\rep f\<^bold>\" "\<^bold>\rep f'\<^bold>\"] Diag_Diagonalize Diagonalize_DOM Diagonalize_COD by force also have "... = \<^bold>\rep g \<^bold>\ rep f \<^bold>\ rep g' \<^bold>\ rep f'\<^bold>\" by auto finally show ?thesis by blast qed qed qed next text \The triangle.\ fix a b assume a: "ide a" assume b: "ide b" show "(a \ \[b]) \ \[a, \, b] = \[a] \ b" proof - have "par ((a \ \[b]) \ \[a, \, b]) (\[a] \ b)" using a b by simp moreover have "can ((a \ \[b]) \ \[a, \, b])" using a b ide_implies_can comp_preserves_can tensor_preserves_can can_assoc can_lunit by simp moreover have "can (\[a] \ b)" using a b ide_implies_can can_runit tensor_preserves_can by simp ultimately show ?thesis using can_coherence by blast qed next text \The pentagon.\ fix a b c d assume a: "ide a" assume b: "ide b" assume c: "ide c" assume d: "ide d" show "(a \ \[b, c, d]) \ \[a, b \ c, d] \ (\[a, b, c] \ d) = \[a, b, c \ d] \ \[a \ b, c, d]" proof - let ?LHS = "(a \ \[b, c, d]) \ \[a, b \ c, d] \ (\[a, b, c] \ d)" let ?RHS = "\[a, b, c \ d] \ \[a \ b, c, d]" have "par ?LHS ?RHS" using a b c d can_assoc tensor_preserves_ide by auto moreover have "can ?LHS" using a b c d ide_implies_can comp_preserves_can tensor_preserves_can can_assoc tensor_preserves_ide by simp moreover have "can ?RHS" using a b c d comp_preserves_can tensor_preserves_can can_assoc tensor_in_hom tensor_preserves_ide by simp ultimately show ?thesis using can_coherence by blast qed qed lemma is_elementary_monoidal_category: shows "elementary_monoidal_category comp tensor\<^sub>F\<^sub>M\<^sub>C unity\<^sub>F\<^sub>M\<^sub>C lunit\<^sub>F\<^sub>M\<^sub>C runit\<^sub>F\<^sub>M\<^sub>C assoc\<^sub>F\<^sub>M\<^sub>C" .. abbreviation T\<^sub>F\<^sub>M\<^sub>C where "T\<^sub>F\<^sub>M\<^sub>C \ EMC.T" abbreviation \\<^sub>F\<^sub>M\<^sub>C where "\\<^sub>F\<^sub>M\<^sub>C \ EMC.\" abbreviation \\<^sub>F\<^sub>M\<^sub>C where "\\<^sub>F\<^sub>M\<^sub>C \ EMC.\" interpretation MC: monoidal_category comp T\<^sub>F\<^sub>M\<^sub>C \\<^sub>F\<^sub>M\<^sub>C \\<^sub>F\<^sub>M\<^sub>C using EMC.induces_monoidal_category by auto lemma induces_monoidal_category: shows "monoidal_category comp T\<^sub>F\<^sub>M\<^sub>C \\<^sub>F\<^sub>M\<^sub>C \\<^sub>F\<^sub>M\<^sub>C" .. end sublocale free_monoidal_category \ elementary_monoidal_category comp tensor\<^sub>F\<^sub>M\<^sub>C unity\<^sub>F\<^sub>M\<^sub>C lunit\<^sub>F\<^sub>M\<^sub>C runit\<^sub>F\<^sub>M\<^sub>C assoc\<^sub>F\<^sub>M\<^sub>C using is_elementary_monoidal_category by auto sublocale free_monoidal_category \ monoidal_category comp T\<^sub>F\<^sub>M\<^sub>C \\<^sub>F\<^sub>M\<^sub>C \\<^sub>F\<^sub>M\<^sub>C using induces_monoidal_category by auto section "Proof of Freeness" text \ Now we proceed on to establish the freeness of \\C\: each functor from @{term C} to a monoidal category @{term D} extends uniquely to a strict monoidal functor from \\C\ to D. \ context free_monoidal_category begin lemma rep_lunit: assumes "ide a" shows "rep \[a] = \<^bold>\\<^bold>\\<^bold>[rep a\<^bold>]\<^bold>\" using assms Lunit_in_lunit [of a "rep a"] rep_in_arr norm_memb_eq_rep [of "\[a]"] by simp lemma rep_runit: assumes "ide a" shows "rep \[a] = \<^bold>\\<^bold>\\<^bold>[rep a\<^bold>]\<^bold>\" using assms Runit_in_runit [of a "rep a"] rep_in_arr norm_memb_eq_rep [of "\[a]"] by simp lemma rep_assoc: assumes "ide a" and "ide b" and "ide c" shows "rep \[a, b, c] = \<^bold>\\<^bold>\\<^bold>[rep a, rep b, rep c\<^bold>]\<^bold>\" using assms Assoc_in_assoc [of a b c "rep a" "rep b" "rep c"] rep_in_arr norm_memb_eq_rep [of "\[a, b, c]"] by simp lemma mkarr_Unity: shows "mkarr \<^bold>\ = \" using unity\<^sub>F\<^sub>M\<^sub>C_def by simp text \ The unitors and associator were given syntactic definitions in terms of corresponding terms, but these were only for the special case of identity arguments (\emph{i.e.}~the components of the natural transformations). We need to show that @{term mkarr} gives the correct result for \emph{all} terms. \ lemma mkarr_Lunit: assumes "Arr t" shows "mkarr \<^bold>\\<^bold>[t\<^bold>] = \ (mkarr t)" proof - have "mkarr \<^bold>\\<^bold>[t\<^bold>] = mkarr (t \<^bold>\ \<^bold>\\<^bold>[\<^bold>\Dom t\<^bold>\\<^bold>])" using assms Arr_implies_Ide_Dom Ide_in_Hom Diagonalize_preserves_Ide Diag_Diagonalize Par_Arr_norm by (intro mkarr_eqI) simp_all also have "... = mkarr t \ mkarr \<^bold>\\<^bold>[\<^bold>\Dom t\<^bold>\\<^bold>]" using assms Arr_implies_Ide_Dom Par_Arr_norm Ide_in_Hom by simp also have "... = mkarr t \ \[dom (mkarr t)]" proof - have "arr \[mkarr (Dom t)]" using assms Arr_implies_Ide_Dom ide_mkarr_Ide by simp moreover have "\<^bold>\\<^bold>[\<^bold>\Dom t\<^bold>\\<^bold>] \ \[mkarr (Dom t)]" using assms Arr_implies_Ide_Dom Lunit_in_lunit rep_mkarr rep_in_arr [of "mkarr (Dom t)"] by simp ultimately show ?thesis using assms mkarr_memb(2) by simp qed also have "... = \ (mkarr t)" using assms Arr_implies_Ide_Dom ide_mkarr_Ide lunit_agreement by simp finally show ?thesis by blast qed lemma mkarr_Lunit': assumes "Arr t" shows "mkarr \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = \' (mkarr t)" proof - have "mkarr \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = mkarr (\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\Cod t\<^bold>\\<^bold>] \<^bold>\ t)" using assms Arr_implies_Ide_Cod Ide_in_Hom Diagonalize_preserves_Ide Diag_Diagonalize Par_Arr_norm by (intro mkarr_eqI) simp_all also have "... = mkarr \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\Cod t\<^bold>\\<^bold>] \ mkarr t" using assms Arr_implies_Ide_Cod Ide_in_Hom Par_Arr_norm by simp also have "... = mkarr (Inv \<^bold>\\<^bold>[\<^bold>\Cod t\<^bold>\\<^bold>]) \ mkarr t" proof - have "mkarr \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\Cod t\<^bold>\\<^bold>] = mkarr (Inv \<^bold>\\<^bold>[\<^bold>\Cod t\<^bold>\\<^bold>])" using assms Arr_implies_Ide_Cod Ide_in_Hom Par_Arr_norm Inv_in_Hom Ide_implies_Can norm_preserves_Can Diagonalize_Inv Diagonalize_preserves_Ide by (intro mkarr_eqI, simp_all) thus ?thesis by argo qed also have "... = \' (cod (mkarr t)) \ mkarr t" proof - have "mkarr (Inv \<^bold>\\<^bold>[\<^bold>\Cod t\<^bold>\\<^bold>]) \ mkarr t = lunit' (cod (mkarr t)) \ mkarr t" using assms Arr_implies_Ide_Cod rep_mkarr Par_Arr_norm inv_mkarr norm_preserves_Can Ide_implies_Can lunit_agreement \'_ide_simp Can_implies_Arr arr_mkarr cod_mkarr ide_cod lunit\<^sub>F\<^sub>M\<^sub>C_def by (metis (no_types, lifting) Can.simps(5)) also have "... = \' (cod (mkarr t)) \ mkarr t" using assms \'_ide_simp arr_mkarr ide_cod by presburger finally show ?thesis by blast qed also have "... = \' (mkarr t)" using assms \'.is_natural_2 [of "mkarr t"] by simp finally show ?thesis by blast qed lemma mkarr_Runit: assumes "Arr t" shows "mkarr \<^bold>\\<^bold>[t\<^bold>] = \ (mkarr t)" proof - have "mkarr \<^bold>\\<^bold>[t\<^bold>] = mkarr (t \<^bold>\ \<^bold>\\<^bold>[\<^bold>\Dom t\<^bold>\\<^bold>])" proof - have "\ Diag (Dom t \<^bold>\ \<^bold>\)" by (cases "Dom t") simp_all thus ?thesis using assms Par_Arr_norm Arr_implies_Ide_Dom Ide_in_Hom Diag_Diagonalize Diagonalize_preserves_Ide by (intro mkarr_eqI) simp_all qed also have "... = mkarr t \ mkarr \<^bold>\\<^bold>[\<^bold>\Dom t\<^bold>\\<^bold>]" using assms Arr_implies_Ide_Dom Par_Arr_norm Ide_in_Hom by simp also have "... = mkarr t \ \[dom (mkarr t)]" proof - have "arr \[mkarr (Dom t)]" using assms Arr_implies_Ide_Dom ide_mkarr_Ide by simp moreover have "\<^bold>\\<^bold>[\<^bold>\Dom t\<^bold>\\<^bold>] \ \[mkarr (Dom t)]" using assms Arr_implies_Ide_Dom Runit_in_runit rep_mkarr rep_in_arr [of "mkarr (Dom t)"] by simp moreover have "mkarr (Dom t) = mkarr \<^bold>\Dom t\<^bold>\" using assms mkarr_rep rep_mkarr arr_mkarr Ide_implies_Arr Arr_implies_Ide_Dom by metis ultimately show ?thesis using assms mkarr_memb(2) by simp qed also have "... = \ (mkarr t)" using assms Arr_implies_Ide_Dom ide_mkarr_Ide runit_agreement by simp finally show ?thesis by blast qed lemma mkarr_Runit': assumes "Arr t" shows "mkarr \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = \' (mkarr t)" proof - have "mkarr \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = mkarr (\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\Cod t\<^bold>\\<^bold>] \<^bold>\ t)" proof - have "\ Diag (Cod t \<^bold>\ \<^bold>\)" by (cases "Cod t") simp_all thus ?thesis using assms Par_Arr_norm Arr_implies_Ide_Cod Ide_in_Hom Diagonalize_preserves_Ide Diag_Diagonalize by (intro mkarr_eqI) simp_all qed also have "... = mkarr \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\Cod t\<^bold>\\<^bold>] \ mkarr t" using assms Arr_implies_Ide_Cod Ide_in_Hom Par_Arr_norm by simp also have "... = mkarr (Inv \<^bold>\\<^bold>[\<^bold>\Cod t\<^bold>\\<^bold>]) \ mkarr t" proof - have "mkarr (Runit' (norm (Cod t))) = mkarr (Inv (Runit (norm (Cod t))))" using assms Arr_implies_Ide_Cod Ide_in_Hom Par_Arr_norm Inv_in_Hom Ide_implies_Can norm_preserves_Can Diagonalize_Inv Diagonalize_preserves_Ide by (intro mkarr_eqI) simp_all thus ?thesis by argo qed also have "... = \' (cod (mkarr t)) \ mkarr t" proof - have "mkarr (Inv \<^bold>\\<^bold>[\<^bold>\Cod t\<^bold>\\<^bold>]) \ mkarr t = runit' (cod (mkarr t)) \ mkarr t" using assms Arr_implies_Ide_Cod rep_mkarr inv_mkarr norm_preserves_Can Ide_implies_Can runit_agreement Can_implies_Arr arr_mkarr cod_mkarr ide_cod runit\<^sub>F\<^sub>M\<^sub>C_def by (metis (no_types, lifting) Can.simps(7)) also have "... = \' (cod (mkarr t)) \ mkarr t" proof - have "runit' (cod (mkarr t)) = \' (cod (mkarr t))" using assms \'_ide_simp arr_mkarr ide_cod by blast thus ?thesis by argo qed finally show ?thesis by blast qed also have "... = \' (mkarr t)" using assms \'.is_natural_2 [of "mkarr t"] by simp finally show ?thesis by blast qed lemma mkarr_Assoc: assumes "Arr t" and "Arr u" and "Arr v" shows "mkarr \<^bold>\\<^bold>[t, u, v\<^bold>] = \ (mkarr t, mkarr u, mkarr v)" proof - have "mkarr \<^bold>\\<^bold>[t, u, v\<^bold>] = mkarr ((t \<^bold>\ u \<^bold>\ v) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\Dom t\<^bold>\, \<^bold>\Dom u\<^bold>\, \<^bold>\Dom v\<^bold>\\<^bold>])" using assms Arr_implies_Ide_Dom Arr_implies_Ide_Cod Ide_in_Hom Diag_Diagonalize Diagonalize_preserves_Ide TensorDiag_preserves_Ide TensorDiag_preserves_Diag TensorDiag_assoc Par_Arr_norm by (intro mkarr_eqI, simp_all) also have "... = \ (mkarr t, mkarr u, mkarr v)" using assms Arr_implies_Ide_Dom rep_mkarr Ide_in_Hom assoc\<^sub>F\<^sub>M\<^sub>C_def Par_Arr_norm [of "Dom t"] Par_Arr_norm [of "Dom u"] Par_Arr_norm [of "Dom v"] \_simp by simp finally show ?thesis by blast qed lemma mkarr_Assoc': assumes "Arr t" and "Arr u" and "Arr v" shows "mkarr \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] = \' (mkarr t, mkarr u, mkarr v)" proof - have "mkarr \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] = mkarr (\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\Cod t\<^bold>\, \<^bold>\Cod u\<^bold>\, \<^bold>\Cod v\<^bold>\\<^bold>] \<^bold>\ (t \<^bold>\ u \<^bold>\ v))" using assms Par_Arr_norm Arr_implies_Ide_Cod Ide_in_Hom Diag_Diagonalize TensorDiag_preserves_Diag CompDiag_Cod_Diag [of "\<^bold>\t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\u\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\v\<^bold>\"] by (intro mkarr_eqI, simp_all) also have "... = mkarr \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\Cod t\<^bold>\, \<^bold>\Cod u\<^bold>\, \<^bold>\Cod v\<^bold>\\<^bold>] \ mkarr (t \<^bold>\ u \<^bold>\ v)" using assms Arr_implies_Ide_Cod Ide_in_Hom Par_Arr_norm by simp also have "... = mkarr (Inv \<^bold>\\<^bold>[\<^bold>\Cod t\<^bold>\, \<^bold>\Cod u\<^bold>\, \<^bold>\Cod v\<^bold>\\<^bold>]) \ mkarr (t \<^bold>\ u \<^bold>\ v)" proof - have "mkarr \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\Cod t\<^bold>\, \<^bold>\Cod u\<^bold>\, \<^bold>\Cod v\<^bold>\\<^bold>] = mkarr (Inv \<^bold>\\<^bold>[\<^bold>\Cod t\<^bold>\, \<^bold>\Cod u\<^bold>\, \<^bold>\Cod v\<^bold>\\<^bold>])" using assms Arr_implies_Ide_Cod Ide_in_Hom Par_Arr_norm Inv_in_Hom Ide_implies_Can norm_preserves_Can Diagonalize_Inv Diagonalize_preserves_Ide by (intro mkarr_eqI, simp_all) thus ?thesis by argo qed also have "... = inv (mkarr \<^bold>\\<^bold>[\<^bold>\Cod t\<^bold>\, \<^bold>\Cod u\<^bold>\, \<^bold>\Cod v\<^bold>\\<^bold>]) \ mkarr (t \<^bold>\ u \<^bold>\ v)" using assms Arr_implies_Ide_Cod Ide_implies_Can norm_preserves_Can by simp also have "... = \' (mkarr t, mkarr u, mkarr v)" proof - have "mkarr (\<^bold>\\<^sup>-\<^sup>1\<^bold>[Inv \<^bold>\Cod t\<^bold>\, Inv \<^bold>\Cod u\<^bold>\, Inv \<^bold>\Cod v\<^bold>\\<^bold>] \<^bold>\ (Cod t \<^bold>\ Cod u \<^bold>\ Cod v)) = mkarr \<^bold>\\<^sup>-\<^sup>1\<^bold>[Inv \<^bold>\Cod t\<^bold>\, Inv \<^bold>\Cod u\<^bold>\, Inv \<^bold>\Cod v\<^bold>\\<^bold>]" using assms Arr_implies_Ide_Cod Inv_in_Hom norm_preserves_Can Diagonalize_Inv Ide_implies_Can Diag_Diagonalize Ide_in_Hom Diagonalize_preserves_Ide Par_Arr_norm TensorDiag_preserves_Diag CompDiag_Cod_Diag [of "\<^bold>\Cod t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Cod u\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Cod v\<^bold>\"] by (intro mkarr_eqI) simp_all thus ?thesis using assms Arr_implies_Ide_Cod rep_mkarr assoc\<^sub>F\<^sub>M\<^sub>C_def \'.map_simp by simp qed finally show ?thesis by blast qed text \ Next, we define the ``inclusion of generators'' functor from @{term C} to \\C\. \ definition inclusion_of_generators where "inclusion_of_generators \ \f. if C.arr f then mkarr \<^bold>\f\<^bold>\ else null" lemma inclusion_is_functor: shows "functor C comp inclusion_of_generators" unfolding inclusion_of_generators_def apply unfold_locales apply auto[4] by (elim C.seqE, simp, intro mkarr_eqI, auto) end text \ We now show that, given a functor @{term V} from @{term C} to a a monoidal category @{term D}, the evaluation map that takes formal arrows of the monoidal language of @{term C} to arrows of @{term D} induces a strict monoidal functor from \\C\ to @{term D}. \ locale evaluation_functor = C: category C + D: monoidal_category D T\<^sub>D \\<^sub>D \\<^sub>D + evaluation_map C D T\<^sub>D \\<^sub>D \\<^sub>D V + \C: free_monoidal_category C for C :: "'c comp" (infixr "\\<^sub>C" 55) and D :: "'d comp" (infixr "\\<^sub>D" 55) and T\<^sub>D :: "'d * 'd \ 'd" and \\<^sub>D :: "'d * 'd * 'd \ 'd" and \\<^sub>D :: "'d" and V :: "'c \ 'd" begin notation eval ("\_\") definition map where "map f \ if \C.arr f then \\C.rep f\ else D.null" text \ It follows from the coherence theorem that a formal arrow and its normal form always have the same evaluation. \ lemma eval_norm: assumes "Arr t" shows "\\<^bold>\t\<^bold>\\ = \t\" using assms \C.Par_Arr_norm \C.Diagonalize_norm coherence canonical_factorization by simp interpretation "functor" \C.comp D map proof fix f show "\\C.arr f \ map f = D.null" using map_def by simp assume f: "\C.arr f" show "D.arr (map f)" using f map_def \C.arr_char by simp show "D.dom (map f) = map (\C.dom f)" using f map_def eval_norm \C.rep_dom Arr_implies_Ide_Dom by auto show "D.cod (map f) = map (\C.cod f)" using f map_def eval_norm \C.rep_cod Arr_implies_Ide_Cod by auto next fix f g assume fg: "\C.seq g f" show "map (\C.comp g f) = D (map g) (map f)" using fg map_def \C.rep_comp \C.rep_preserves_seq eval_norm by auto qed lemma is_functor: shows "functor \C.comp D map" .. interpretation FF: product_functor \C.comp \C.comp D D map map .. interpretation FoT: composite_functor \C.CC.comp \C.comp D \C.T\<^sub>F\<^sub>M\<^sub>C map .. interpretation ToFF: composite_functor \C.CC.comp D.CC.comp D FF.map T\<^sub>D .. interpretation strict_monoidal_functor \C.comp \C.T\<^sub>F\<^sub>M\<^sub>C \C.\ \C.\ D T\<^sub>D \\<^sub>D \\<^sub>D map proof show "map \C.\ = \\<^sub>D" using \C.\_def \C.lunit_agreement map_def \C.rep_lunit \C.Arr_rep [of \] eval_norm \C.lunit_agreement D.unitor_coincidence D.comp_cod_arr D.\_in_hom by auto show "\f g. \ \C.arr f; \C.arr g \ \ map (\C.tensor f g) = D.tensor (map f) (map g)" using map_def \C.rep_tensor \C.Arr_rep eval_norm by simp show "\a b c. \ \C.ide a; \C.ide b; \C.ide c \ \ map (\C.assoc a b c) = D.assoc (map a) (map b) (map c)" using map_def \C.assoc\<^sub>F\<^sub>M\<^sub>C_def \C.rep_mkarr eval_norm by auto qed lemma is_strict_monoidal_functor: shows "strict_monoidal_functor \C.comp \C.T\<^sub>F\<^sub>M\<^sub>C \C.\ \C.\ D T\<^sub>D \\<^sub>D \\<^sub>D map" .. end sublocale evaluation_functor \ strict_monoidal_functor \C.comp \C.T\<^sub>F\<^sub>M\<^sub>C \C.\\<^sub>F\<^sub>M\<^sub>C \C.\\<^sub>F\<^sub>M\<^sub>C D T\<^sub>D \\<^sub>D \\<^sub>D map using is_strict_monoidal_functor by auto text \ The final step in proving freeness is to show that the evaluation functor is the \emph{unique} strict monoidal extension of the functor @{term V} to \\C\. This is done by induction, exploiting the syntactic construction of \\C\. \ text \ To ease the statement and proof of the result, we define a locale that expresses that @{term F} is a strict monoidal extension to monoidal category @{term C}, of a functor @{term "V"} from @{term "C\<^sub>0"} to a monoidal category @{term D}, along a functor @{term I} from @{term "C\<^sub>0"} to @{term C}. \ locale strict_monoidal_extension = C\<^sub>0: category C\<^sub>0 + C: monoidal_category C T\<^sub>C \\<^sub>C \\<^sub>C + D: monoidal_category D T\<^sub>D \\<^sub>D \\<^sub>D + I: "functor" C\<^sub>0 C I + V: "functor" C\<^sub>0 D V + strict_monoidal_functor C T\<^sub>C \\<^sub>C \\<^sub>C D T\<^sub>D \\<^sub>D \\<^sub>D F for C\<^sub>0 :: "'c\<^sub>0 comp" and C :: "'c comp" (infixr "\\<^sub>C" 55) and T\<^sub>C :: "'c * 'c \ 'c" and \\<^sub>C :: "'c * 'c * 'c \ 'c" and \\<^sub>C :: "'c" and D :: "'d comp" (infixr "\\<^sub>D" 55) and T\<^sub>D :: "'d * 'd \ 'd" and \\<^sub>D :: "'d * 'd * 'd \ 'd" and \\<^sub>D :: "'d" and I :: "'c\<^sub>0 \ 'c" and V :: "'c\<^sub>0 \ 'd" and F :: "'c \ 'd" + assumes is_extension: "\f. C\<^sub>0.arr f \ F (I f) = V f" sublocale evaluation_functor \ strict_monoidal_extension C \C.comp \C.T\<^sub>F\<^sub>M\<^sub>C \C.\ \C.\ D T\<^sub>D \\<^sub>D \\<^sub>D \C.inclusion_of_generators V map proof - interpret inclusion: "functor" C \C.comp \C.inclusion_of_generators using \C.inclusion_is_functor by auto show "strict_monoidal_extension C \C.comp \C.T\<^sub>F\<^sub>M\<^sub>C \C.\ \C.\ D T\<^sub>D \\<^sub>D \\<^sub>D \C.inclusion_of_generators V map" apply unfold_locales using map_def \C.rep_mkarr eval_norm \C.inclusion_of_generators_def by simp qed text \ A special case of interest is a strict monoidal extension to \\C\, of a functor @{term V} from a category @{term C} to a monoidal category @{term D}, along the inclusion of generators from @{term C} to \\C\. The evaluation functor induced by @{term V} is such an extension. \ locale strict_monoidal_extension_to_free_monoidal_category = C: category C + monoidal_language C + \C: free_monoidal_category C + strict_monoidal_extension C \C.comp \C.T\<^sub>F\<^sub>M\<^sub>C \C.\ \C.\ D T\<^sub>D \\<^sub>D \\<^sub>D \C.inclusion_of_generators V F for C :: "'c comp" (infixr "\\<^sub>C" 55) and D :: "'d comp" (infixr "\\<^sub>D" 55) and T\<^sub>D :: "'d * 'd \ 'd" and \\<^sub>D :: "'d * 'd * 'd \ 'd" and \\<^sub>D :: "'d" and V :: "'c \ 'd" and F :: "'c free_monoidal_category.arr \ 'd" begin lemma strictly_preserves_everything: shows "C.arr f \ F (\C.mkarr \<^bold>\f\<^bold>\) = V f" and "F (\C.mkarr \<^bold>\) = \\<^sub>D" and "\ Arr t; Arr u \ \ F (\C.mkarr (t \<^bold>\ u)) = F (\C.mkarr t) \\<^sub>D F (\C.mkarr u)" and "\ Arr t; Arr u; Dom t = Cod u \ \ F (\C.mkarr (t \<^bold>\ u)) = F (\C.mkarr t) \\<^sub>D F (\C.mkarr u)" and "Arr t \ F (\C.mkarr \<^bold>\\<^bold>[t\<^bold>]) = D.\ (F (\C.mkarr t))" and "Arr t \ F (\C.mkarr \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]) = D.\'.map (F (\C.mkarr t))" and "Arr t \ F (\C.mkarr \<^bold>\\<^bold>[t\<^bold>]) = D.\ (F (\C.mkarr t))" and "Arr t \ F (\C.mkarr \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]) = D.\'.map (F (\C.mkarr t))" and "\ Arr t; Arr u; Arr v \ \ F (\C.mkarr \<^bold>\\<^bold>[t, u, v\<^bold>]) = \\<^sub>D (F (\C.mkarr t), F (\C.mkarr u), F (\C.mkarr v))" and "\ Arr t; Arr u; Arr v \ \ F (\C.mkarr \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]) = D.\' (F (\C.mkarr t), F (\C.mkarr u), F (\C.mkarr v))" proof - show "C.arr f \ F (\C.mkarr \<^bold>\f\<^bold>\) = V f" using is_extension \C.inclusion_of_generators_def by simp show "F (\C.mkarr \<^bold>\) = \\<^sub>D" using \C.mkarr_Unity \C.\_def strictly_preserves_unity \C.\_agreement by auto show tensor_case: "\t u.\ Arr t; Arr u \ \ F (\C.mkarr (t \<^bold>\ u)) = F (\C.mkarr t) \\<^sub>D F (\C.mkarr u)" proof - fix t u assume t: "Arr t" and u: "Arr u" have "F (\C.mkarr (t \<^bold>\ u)) = F (\C.tensor (\C.mkarr t) (\C.mkarr u))" using t u \C.tensor_mkarr \C.arr_mkarr by simp also have "... = F (\C.mkarr t) \\<^sub>D F (\C.mkarr u)" using t u \C.arr_mkarr strictly_preserves_tensor by blast finally show "F (\C.mkarr (t \<^bold>\ u)) = F (\C.mkarr t) \\<^sub>D F (\C.mkarr u)" by fast qed show "\ Arr t; Arr u; Dom t = Cod u \ \ F (\C.mkarr (t \<^bold>\ u)) = F (\C.mkarr t) \\<^sub>D F (\C.mkarr u)" proof - fix t u assume t: "Arr t" and u: "Arr u" and tu: "Dom t = Cod u" show "F (\C.mkarr (t \<^bold>\ u)) = F (\C.mkarr t) \\<^sub>D F (\C.mkarr u)" proof - have "F (\C.mkarr (t \<^bold>\ u)) = F (\C.mkarr t \ \C.mkarr u)" using t u tu \C.comp_mkarr by simp also have "... = F (\C.mkarr t) \\<^sub>D F (\C.mkarr u)" using t u tu \C.arr_mkarr by fastforce finally show ?thesis by blast qed qed show "Arr t \ F (\C.mkarr \<^bold>\\<^bold>[t\<^bold>]) = D.\ (F (\C.mkarr t))" using \C.mkarr_Lunit Arr_implies_Ide_Dom \C.ide_mkarr_Ide strictly_preserves_lunit by simp show "Arr t \ F (\C.mkarr \<^bold>\\<^bold>[t\<^bold>]) = D.\ (F (\C.mkarr t))" using \C.mkarr_Runit Arr_implies_Ide_Dom \C.ide_mkarr_Ide strictly_preserves_runit by simp show "\ Arr t; Arr u; Arr v \ \ F (\C.mkarr \<^bold>\\<^bold>[t, u, v\<^bold>]) = \\<^sub>D (F (\C.mkarr t), F (\C.mkarr u), F (\C.mkarr v))" using \C.mkarr_Assoc strictly_preserves_assoc \C.ide_mkarr_Ide tensor_case by simp show "Arr t \ F (\C.mkarr \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]) = D.\'.map (F (\C.mkarr t))" proof - assume t: "Arr t" have "F (\C.mkarr \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]) = F (\C.lunit' (\C.mkarr (Cod t))) \\<^sub>D F (\C.mkarr t)" using t \C.mkarr_Lunit' Arr_implies_Ide_Cod \C.ide_mkarr_Ide \C.\'.map_simp \C.comp_cod_arr by simp also have "... = D.lunit' (D.cod (F (\C.mkarr t))) \\<^sub>D F (\C.mkarr t)" using t Arr_implies_Ide_Cod \C.ide_mkarr_Ide strictly_preserves_lunit preserves_inv by simp also have "... = D.\'.map (F (\C.mkarr t))" using t D.\'.map_simp D.comp_cod_arr by simp finally show ?thesis by blast qed show "Arr t \ F (\C.mkarr \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]) = D.\'.map (F (\C.mkarr t))" proof - assume t: "Arr t" have "F (\C.mkarr \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]) = F (\C.runit' (\C.mkarr (Cod t))) \\<^sub>D F (\C.mkarr t)" using t \C.mkarr_Runit' Arr_implies_Ide_Cod \C.ide_mkarr_Ide \C.\'.map_simp \C.comp_cod_arr by simp also have "... = D.runit' (D.cod (F (\C.mkarr t))) \\<^sub>D F (\C.mkarr t)" using t Arr_implies_Ide_Cod \C.ide_mkarr_Ide strictly_preserves_runit preserves_inv by simp also have "... = D.\'.map (F (\C.mkarr t))" using t D.\'.map_simp D.comp_cod_arr by simp finally show ?thesis by blast qed show "\ Arr t; Arr u; Arr v \ \ F (\C.mkarr \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]) = D.\'.map (F (\C.mkarr t), F (\C.mkarr u), F (\C.mkarr v))" proof - assume t: "Arr t" and u: "Arr u" and v: "Arr v" have "F (\C.mkarr \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]) = F (\C.assoc' (\C.mkarr (Cod t)) (\C.mkarr (Cod u)) (\C.mkarr (Cod v))) \\<^sub>D (F (\C.mkarr t) \\<^sub>D F (\C.mkarr u) \\<^sub>D F (\C.mkarr v))" using t u v \C.mkarr_Assoc' Arr_implies_Ide_Cod \C.ide_mkarr_Ide \C.\'.map_simp tensor_case \C.iso_assoc by simp also have "... = D.assoc' (D.cod (F (\C.mkarr t))) (D.cod (F (\C.mkarr u))) (D.cod (F (\C.mkarr v))) \\<^sub>D (F (\C.mkarr t) \\<^sub>D F (\C.mkarr u) \\<^sub>D F (\C.mkarr v))" using t u v \C.ide_mkarr_Ide Arr_implies_Ide_Cod preserves_inv \C.iso_assoc strictly_preserves_assoc [of "\C.mkarr (Cod t)" "\C.mkarr (Cod u)" "\C.mkarr (Cod v)"] by simp also have "... = D.\'.map (F (\C.mkarr t), F (\C.mkarr u), F (\C.mkarr v))" using t u v D.\'.map_simp by simp finally show ?thesis by blast qed qed end sublocale evaluation_functor \ strict_monoidal_extension_to_free_monoidal_category C D T\<^sub>D \\<^sub>D \\<^sub>D V map .. context free_monoidal_category begin text \ The evaluation functor induced by @{term V} is the unique strict monoidal extension of @{term V} to \\C\. \ theorem is_free: assumes "strict_monoidal_extension_to_free_monoidal_category C D T\<^sub>D \\<^sub>D \\<^sub>D V F" shows "F = evaluation_functor.map C D T\<^sub>D \\<^sub>D \\<^sub>D V" proof - interpret F: strict_monoidal_extension_to_free_monoidal_category C D T\<^sub>D \\<^sub>D \\<^sub>D V F using assms by auto interpret E: evaluation_functor C D T\<^sub>D \\<^sub>D \\<^sub>D V .. have Ide_case: "\a. Ide a \ F (mkarr a) = E.map (mkarr a)" proof - fix a show "Ide a \ F (mkarr a) = E.map (mkarr a)" using E.strictly_preserves_everything F.strictly_preserves_everything Ide_implies_Arr by (induct a) auto qed show ?thesis proof fix f have "\arr f \ F f = E.map f" using E.is_extensional F.is_extensional by simp moreover have "arr f \ F f = E.map f" proof - assume f: "arr f" have "Arr (rep f) \ f = mkarr (rep f)" using f mkarr_rep by simp moreover have "\t. Arr t \ F (mkarr t) = E.map (mkarr t)" proof - fix t show "Arr t \ F (mkarr t) = E.map (mkarr t)" using Ide_case E.strictly_preserves_everything F.strictly_preserves_everything Arr_implies_Ide_Dom Arr_implies_Ide_Cod by (induct t) auto qed ultimately show "F f = E.map f" by metis qed ultimately show "F f = E.map f" by blast qed qed end section "Strict Subcategory" context free_monoidal_category begin text \ In this section we show that \\C\ is monoidally equivalent to its full subcategory \\\<^sub>SC\ whose objects are the equivalence classes of diagonal identity terms, and that this subcategory is the free strict monoidal category generated by @{term C}. \ interpretation \\<^sub>SC: full_subcategory comp \\f. ide f \ Diag (DOM f)\ by (unfold_locales) auto text \ The mapping defined on equivalence classes by diagonalizing their representatives is a functor from the free monoidal category to the subcategory @{term "\\<^sub>SC"}. \ definition D where "D \ \f. if arr f then mkarr \<^bold>\rep f\<^bold>\ else \\<^sub>SC.null" text \ The arrows of \\\<^sub>SC\ are those equivalence classes whose canonical representative term has diagonal formal domain and codomain. \ lemma strict_arr_char: shows "\\<^sub>SC.arr f \ arr f \ Diag (DOM f) \ Diag (COD f)" proof show "arr f \ Diag (DOM f) \ Diag (COD f) \ \\<^sub>SC.arr f" using \\<^sub>SC.arr_char DOM_dom DOM_cod by simp show "\\<^sub>SC.arr f \ arr f \ Diag (DOM f) \ Diag (COD f)" using \\<^sub>SC.arr_char Arr_rep Arr_implies_Ide_Cod Ide_implies_Arr DOM_dom DOM_cod by force qed text \ Alternatively, the arrows of \\\<^sub>SC\ are those equivalence classes that are preserved by diagonalization of representatives. \ lemma strict_arr_char': shows "\\<^sub>SC.arr f \ arr f \ D f = f" proof fix f assume f: "\\<^sub>SC.arr f" show "arr f \ D f = f" proof show "arr f" using f \\<^sub>SC.arr_char by blast show "D f = f" using f strict_arr_char mkarr_Diagonalize_rep D_def by simp qed next assume f: "arr f \ D f = f" show "\\<^sub>SC.arr f" proof - have "arr f" using f by simp moreover have "Diag (DOM f)" proof - have "DOM f = DOM (mkarr \<^bold>\rep f\<^bold>\)" using f D_def by auto also have "... = Dom \<^bold>\\<^bold>\rep f\<^bold>\\<^bold>\" using f Arr_rep Diagonalize_in_Hom rep_mkarr by simp also have "... = Dom \<^bold>\rep f\<^bold>\" using f Arr_rep Diagonalize_in_Hom Par_Arr_norm [of "\<^bold>\rep f\<^bold>\"] by force finally have "DOM f = Dom \<^bold>\rep f\<^bold>\" by blast thus ?thesis using f Arr_rep Diag_Diagonalize Dom_preserves_Diag by metis qed moreover have "Diag (COD f)" proof - have "COD f = COD (mkarr \<^bold>\rep f\<^bold>\)" using f D_def by auto also have "... = Cod \<^bold>\\<^bold>\rep f\<^bold>\\<^bold>\" using f Arr_rep Diagonalize_in_Hom rep_mkarr by simp also have "... = Cod \<^bold>\rep f\<^bold>\" using f Arr_rep Diagonalize_in_Hom Par_Arr_norm [of "\<^bold>\rep f\<^bold>\"] by force finally have "COD f = Cod \<^bold>\rep f\<^bold>\" by blast thus ?thesis using f Arr_rep Diag_Diagonalize Cod_preserves_Diag by metis qed ultimately show ?thesis using strict_arr_char by auto qed qed interpretation D: "functor" comp \\<^sub>SC.comp D proof - have 1: "\f. arr f \ \\<^sub>SC.arr (D f)" unfolding strict_arr_char D_def using arr_mkarr Diagonalize_in_Hom Arr_rep rep_mkarr Par_Arr_norm Arr_implies_Ide_Dom Arr_implies_Ide_Cod Diag_Diagonalize by force show "functor comp \\<^sub>SC.comp D" proof show "\f. \ arr f \ D f = \\<^sub>SC.null" using D_def by simp show "\f. arr f \ \\<^sub>SC.arr (D f)" by fact show "\f. arr f \ \\<^sub>SC.dom (D f) = D (dom f)" using D_def Diagonalize_in_Hom \\<^sub>SC.dom_char \\<^sub>SC.arr_char rep_mkarr rep_dom Arr_implies_Ide_Dom Arr_implies_Ide_Cod Diagonalize_preserves_Ide ide_mkarr_Ide Diag_Diagonalize Dom_norm by simp show 2: "\f. arr f \ \\<^sub>SC.cod (D f) = D (cod f)" using D_def Diagonalize_in_Hom \\<^sub>SC.cod_char \\<^sub>SC.arr_char rep_mkarr rep_cod Arr_implies_Ide_Dom Arr_implies_Ide_Cod Diagonalize_preserves_Ide ide_mkarr_Ide Diag_Diagonalize Dom_norm by simp fix f g assume fg: "seq g f" hence fg': "arr f \ arr g \ dom g = cod f" by blast show "D (g \ f) = \\<^sub>SC.comp (D g) (D f)" proof - have seq: "\\<^sub>SC.seq (mkarr \<^bold>\rep g\<^bold>\) (mkarr \<^bold>\rep f\<^bold>\)" proof - have 3: "\\<^sub>SC.arr (mkarr \<^bold>\rep g\<^bold>\) \ \\<^sub>SC.arr (mkarr \<^bold>\rep f\<^bold>\)" using fg' 1 arr_char D_def by force moreover have "\\<^sub>SC.dom (mkarr \<^bold>\rep g\<^bold>\) = \\<^sub>SC.cod (mkarr \<^bold>\rep f\<^bold>\)" using fg' 2 3 \\<^sub>SC.dom_char rep_in_Hom mkarr_in_hom D_def Dom_Diagonalize_rep Diag_implies_Arr Diag_Diagonalize(1) \\<^sub>SC.arr_char by force ultimately show ?thesis using \\<^sub>SC.seqI by auto qed have "mkarr \<^bold>\rep (g \ f)\<^bold>\ = \\<^sub>SC.comp (mkarr \<^bold>\rep g\<^bold>\) (mkarr \<^bold>\rep f\<^bold>\)" proof - have Seq: "Seq \<^bold>\rep g\<^bold>\ \<^bold>\rep f\<^bold>\" using fg rep_preserves_seq Diagonalize_in_Hom by force hence 4: "\<^bold>\rep g\<^bold>\ \<^bold>\ \<^bold>\rep f\<^bold>\ \ Hom \<^bold>\DOM f\<^bold>\ \<^bold>\COD g\<^bold>\" using fg' Seq Diagonalize_in_Hom by auto have "\\<^sub>SC.comp (mkarr \<^bold>\rep g\<^bold>\) (mkarr \<^bold>\rep f\<^bold>\) = mkarr \<^bold>\rep g\<^bold>\ \ mkarr \<^bold>\rep f\<^bold>\" using seq \\<^sub>SC.comp_char \\<^sub>SC.seq_char by meson also have "... = mkarr (\<^bold>\rep g\<^bold>\ \<^bold>\ \<^bold>\rep f\<^bold>\)" using Seq comp_mkarr by fastforce also have "... = mkarr \<^bold>\rep (g \ f)\<^bold>\" proof (intro mkarr_eqI) show "Par (\<^bold>\rep g\<^bold>\ \<^bold>\ \<^bold>\rep f\<^bold>\) \<^bold>\rep (g \ f)\<^bold>\" using fg 4 rep_in_Hom rep_preserves_seq rep_in_Hom Diagonalize_in_Hom Par_Arr_norm apply (elim seqE, auto) by (simp_all add: rep_comp) show "\<^bold>\\<^bold>\rep g\<^bold>\ \<^bold>\ \<^bold>\rep f\<^bold>\\<^bold>\ = \<^bold>\\<^bold>\rep (g \ f)\<^bold>\\<^bold>\" using fg rep_preserves_seq norm_in_Hom Diag_Diagonalize Diagonalize_Diag apply auto by (simp add: rep_comp) qed finally show ?thesis by blast qed thus ?thesis using fg D_def by auto qed qed qed lemma diagonalize_is_functor: shows "functor comp \\<^sub>SC.comp D" .. lemma diagonalize_strict_arr: assumes "\\<^sub>SC.arr f" shows "D f = f" using assms arr_char D_def strict_arr_char Arr_rep Arr_implies_Ide_Dom Ide_implies_Arr mkarr_Diagonalize_rep [of f] by auto lemma diagonalize_is_idempotent: shows "D o D = D" unfolding D_def using D.is_extensional \\<^sub>SC.null_char Arr_rep Diagonalize_in_Hom mkarr_Diagonalize_rep strict_arr_char rep_mkarr by fastforce lemma diagonalize_tensor: assumes "arr f" and "arr g" shows "D (f \ g) = D (D f \ D g)" unfolding D_def using assms strict_arr_char rep_in_Hom Diagonalize_in_Hom tensor_mkarr rep_tensor Diagonalize_in_Hom rep_mkarr Diagonalize_norm Diagonalize_Tensor by force lemma ide_diagonalize_can: assumes "can f" shows "ide (D f)" using assms D_def Can_rep_can Ide_Diagonalize_Can ide_mkarr_Ide can_implies_arr by simp text \ We next show that the diagonalization functor and the inclusion of the full subcategory \\\<^sub>SC\ underlie an equivalence of categories. The arrows @{term "mkarr (DOM a\<^bold>\)"}, determined by reductions of canonical representatives, are the components of a natural isomorphism. \ interpretation S: full_inclusion_functor comp \\f. ide f \ Diag (DOM f)\ .. interpretation DoS: composite_functor \\<^sub>SC.comp comp \\<^sub>SC.comp \\<^sub>SC.map D .. interpretation SoD: composite_functor comp \\<^sub>SC.comp comp D \\<^sub>SC.map .. interpretation \: transformation_by_components comp comp map SoD.map \\a. mkarr (DOM a\<^bold>\)\ proof fix a assume a: "ide a" show "\mkarr (DOM a\<^bold>\) : map a \ SoD.map a\" proof - have "\mkarr (DOM a\<^bold>\) : a \ mkarr \<^bold>\DOM a\<^bold>\\" using a Arr_implies_Ide_Dom red_in_Hom dom_char [of a] by auto moreover have "map a = a" using a map_simp by simp moreover have "SoD.map a = mkarr \<^bold>\DOM a\<^bold>\" using a D.preserves_ide \\<^sub>SC.ideD \\<^sub>SC.map_simp D_def Ide_Diagonalize_rep_ide Ide_in_Hom Diagonalize_in_Hom by force ultimately show ?thesis by simp qed next fix f assume f: "arr f" show "mkarr (DOM (cod f)\<^bold>\) \ map f = SoD.map f \ mkarr (DOM (dom f)\<^bold>\)" proof - have "SoD.map f \ mkarr (DOM (dom f)\<^bold>\) = mkarr \<^bold>\rep f\<^bold>\ \ mkarr (DOM f\<^bold>\)" using f DOM_dom D.preserves_arr \\<^sub>SC.map_simp D_def by simp also have "... = mkarr (\<^bold>\rep f\<^bold>\ \<^bold>\ DOM f\<^bold>\)" using f Diagonalize_in_Hom red_in_Hom comp_mkarr Arr_implies_Ide_Dom by simp also have "... = mkarr (COD f\<^bold>\ \<^bold>\ rep f)" proof (intro mkarr_eqI) show "Par (\<^bold>\rep f\<^bold>\ \<^bold>\ DOM f\<^bold>\) (COD f\<^bold>\ \<^bold>\ rep f)" using f Diagonalize_in_Hom red_in_Hom Arr_implies_Ide_Dom Arr_implies_Ide_Cod by simp show "\<^bold>\\<^bold>\rep f\<^bold>\ \<^bold>\ DOM f\<^bold>\\<^bold>\ = \<^bold>\COD f\<^bold>\ \<^bold>\ rep f\<^bold>\" proof - have "\<^bold>\\<^bold>\rep f\<^bold>\ \<^bold>\ DOM f\<^bold>\\<^bold>\ = \<^bold>\rep f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\DOM f\<^bold>\\<^bold>\" using f by simp also have "... = \<^bold>\rep f\<^bold>\" using f Arr_implies_Ide_Dom Can_red Ide_Diagonalize_Can [of "DOM f\<^bold>\"] Diag_Diagonalize CompDiag_Diag_Ide by force also have "... = \<^bold>\COD f\<^bold>\\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\rep f\<^bold>\" using f Arr_implies_Ide_Cod Can_red Ide_Diagonalize_Can [of "COD f\<^bold>\"] Diag_Diagonalize CompDiag_Diag_Ide by force also have "... = \<^bold>\COD f\<^bold>\ \<^bold>\ rep f\<^bold>\" by simp finally show ?thesis by blast qed qed also have "... = mkarr (COD f\<^bold>\) \ mkarr (rep f)" using f comp_mkarr rep_in_Hom red_in_Hom Arr_implies_Ide_Cod by blast also have "... = mkarr (DOM (cod f)\<^bold>\) \ map f" using f DOM_cod by simp finally show ?thesis by blast qed qed interpretation \: natural_isomorphism comp comp map SoD.map \.map apply unfold_locales using \.map_simp_ide rep_in_Hom Arr_implies_Ide_Dom Can_red can_mkarr_Can iso_can by simp text \ The restriction of the diagonalization functor to the subcategory \\\<^sub>SC\ is the identity. \ lemma DoS_eq_\\<^sub>SC: shows "DoS.map = \\<^sub>SC.map" proof fix f have "\ \\<^sub>SC.arr f \ DoS.map f = \\<^sub>SC.map f" using DoS.is_extensional \\<^sub>SC.map_def by simp moreover have "\\<^sub>SC.arr f \ DoS.map f = \\<^sub>SC.map f" using \\<^sub>SC.map_simp strict_arr_char Diagonalize_Diag D_def mkarr_Diagonalize_rep by simp ultimately show "DoS.map f = \\<^sub>SC.map f" by blast qed interpretation \: transformation_by_components \\<^sub>SC.comp \\<^sub>SC.comp DoS.map \\<^sub>SC.map \\a. a\ using \\<^sub>SC.ideD \\<^sub>SC.map_simp DoS_eq_\\<^sub>SC \\<^sub>SC.map_simp \\<^sub>SC.comp_cod_arr \\<^sub>SC.comp_arr_dom apply unfold_locales by (intro \\<^sub>SC.in_homI) auto interpretation \: natural_isomorphism \\<^sub>SC.comp \\<^sub>SC.comp DoS.map \\<^sub>SC.map \.map apply unfold_locales using \.map_simp_ide \\<^sub>SC.ide_is_iso by simp interpretation equivalence_of_categories \\<^sub>SC.comp comp D \\<^sub>SC.map \.map \.map .. text \ We defined the natural isomorphisms @{term \} and @{term \} by giving their components (\emph{i.e.}~their values at objects). However, it is helpful in exporting these facts to have simple characterizations of their values for all arrows. \ definition \ where "\ \ \f. if \\<^sub>SC.arr f then f else \\<^sub>SC.null" definition \ where "\ \ \f. if arr f then mkarr (COD f\<^bold>\) \ f else null" lemma \_char: shows "\.map = \" proof (intro NaturalTransformation.eqI) show "natural_transformation \\<^sub>SC.comp \\<^sub>SC.comp DoS.map \\<^sub>SC.map \.map" .. have "natural_transformation \\<^sub>SC.comp \\<^sub>SC.comp \\<^sub>SC.map \\<^sub>SC.map \\<^sub>SC.map" using DoS.as_nat_trans.natural_transformation_axioms DoS_eq_\\<^sub>SC by simp moreover have "\\<^sub>SC.map = \" unfolding \_def using \\<^sub>SC.map_def by blast ultimately show "natural_transformation \\<^sub>SC.comp \\<^sub>SC.comp DoS.map \\<^sub>SC.map \" using \\<^sub>SC.as_nat_trans.natural_transformation_axioms DoS_eq_\\<^sub>SC by simp show "\a. \\<^sub>SC.ide a \ \.map a = \ a" using \.map_simp_ide \\<^sub>SC.ideD \_def by simp qed lemma \_char: shows "\.map = \" unfolding \.map_def \_def using map_simp DOM_cod by fastforce lemma is_equivalent_to_strict_subcategory: shows "equivalence_of_categories \\<^sub>SC.comp comp D \\<^sub>SC.map \ \" proof - have "equivalence_of_categories \\<^sub>SC.comp comp D \\<^sub>SC.map \.map \.map" .. thus "equivalence_of_categories \\<^sub>SC.comp comp D \\<^sub>SC.map \ \" using \_char \_char by simp qed text \ The inclusion of generators functor from @{term C} to \\C\ corestricts to a functor from @{term C} to \\\<^sub>SC\. \ interpretation I: "functor" C comp inclusion_of_generators using inclusion_is_functor by auto interpretation DoI: composite_functor C comp \\<^sub>SC.comp inclusion_of_generators D .. lemma DoI_eq_I: shows "DoI.map = inclusion_of_generators" proof fix f have "\ C.arr f \ DoI.map f = inclusion_of_generators f" using DoI.is_extensional I.is_extensional \\<^sub>SC.null_char by blast moreover have "C.arr f \ DoI.map f = inclusion_of_generators f" proof - assume f: "C.arr f" have "DoI.map f = D (inclusion_of_generators f)" using f by simp also have "... = inclusion_of_generators f" proof - have "\\<^sub>SC.arr (inclusion_of_generators f)" using f arr_mkarr rep_mkarr Par_Arr_norm [of "\<^bold>\f\<^bold>\"] strict_arr_char inclusion_of_generators_def by simp thus ?thesis using f strict_arr_char' by blast qed finally show "DoI.map f = inclusion_of_generators f" by blast qed ultimately show "DoI.map f = inclusion_of_generators f" by blast qed end text \ Next, we show that the subcategory \\\<^sub>SC\ inherits monoidal structure from the ambient category \\C\, and that this monoidal structure is strict. \ locale free_strict_monoidal_category = monoidal_language C + \C: free_monoidal_category C + full_subcategory \C.comp "\f. \C.ide f \ Diag (\C.DOM f)" for C :: "'c comp" begin interpretation D: "functor" \C.comp comp \C.D using \C.diagonalize_is_functor by auto notation comp (infixr "\\<^sub>S" 55) definition tensor\<^sub>S (infixr "\\<^sub>S" 53) where "f \\<^sub>S g \ \C.D (\C.tensor f g)" definition assoc\<^sub>S ("\\<^sub>S[_, _, _]") where "assoc\<^sub>S a b c \ a \\<^sub>S b \\<^sub>S c" lemma tensor_char: assumes "arr f" and "arr g" shows "f \\<^sub>S g = \C.mkarr (\<^bold>\\C.rep f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\\C.rep g\<^bold>\)" unfolding \C.D_def tensor\<^sub>S_def using assms arr_char \C.rep_tensor by simp lemma tensor_in_hom [simp]: assumes "\f : a \ b\" and "\g : c \ d\" shows "\f \\<^sub>S g : a \\<^sub>S c \ b \\<^sub>S d\" unfolding tensor\<^sub>S_def - using assms D.preserves_hom D.preserves_arr arr_char in_hom_char by simp + using assms D.preserves_hom arr_char in_hom_char + by (metis (no_types, lifting) \C.T_simp \C.tensor_in_hom in_homE) lemma arr_tensor [simp]: assumes "arr f" and "arr g" shows "arr (f \\<^sub>S g)" using assms arr_iff_in_hom [of f] arr_iff_in_hom [of g] tensor_in_hom by blast lemma dom_tensor [simp]: assumes "arr f" and "arr g" shows "dom (f \\<^sub>S g) = dom f \\<^sub>S dom g" using assms arr_iff_in_hom [of f] arr_iff_in_hom [of g] tensor_in_hom by blast lemma cod_tensor [simp]: assumes "arr f" and "arr g" shows "cod (f \\<^sub>S g) = cod f \\<^sub>S cod g" using assms arr_iff_in_hom [of f] arr_iff_in_hom [of g] tensor_in_hom by blast lemma tensor_preserves_ide: assumes "ide a" and "ide b" shows "ide (a \\<^sub>S b)" using assms tensor\<^sub>S_def D.preserves_ide \C.tensor_preserves_ide ide_char by fastforce lemma tensor_tensor: assumes "arr f" and "arr g" and "arr h" shows "(f \\<^sub>S g) \\<^sub>S h = \C.mkarr (\<^bold>\\C.rep f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\\C.rep g\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\\C.rep h\<^bold>\)" and "f \\<^sub>S g \\<^sub>S h = \C.mkarr (\<^bold>\\C.rep f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\\C.rep g\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\\C.rep h\<^bold>\)" proof - show "(f \\<^sub>S g) \\<^sub>S h = \C.mkarr (\<^bold>\\C.rep f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\\C.rep g\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\\C.rep h\<^bold>\)" proof - have "(f \\<^sub>S g) \\<^sub>S h = \C.mkarr (\<^bold>\\C.rep (f \\<^sub>S g)\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\\C.rep h\<^bold>\)" using assms Diag_Diagonalize TensorDiag_preserves_Diag Diag_implies_Arr \C.COD_mkarr \C.DOM_mkarr \C.strict_arr_char tensor_char by simp also have "... = \C.mkarr (\<^bold>\\C.rep (\C.mkarr (\<^bold>\\C.rep f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\\C.rep g\<^bold>\))\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\\C.rep h\<^bold>\)" using assms arr_char tensor_char by simp also have "... = \C.mkarr (\<^bold>\\<^bold>\\C.rep f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\\C.rep g\<^bold>\\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\\C.rep h\<^bold>\)" using assms \C.rep_mkarr TensorDiag_in_Hom Diag_Diagonalize TensorDiag_preserves_Diag arr_char by force also have "... = \C.mkarr (\<^bold>\\C.rep f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\\C.rep g\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\\C.rep h\<^bold>\)" using assms Diag_Diagonalize TensorDiag_preserves_Diag TensorDiag_assoc arr_char by force finally show ?thesis by blast qed show "f \\<^sub>S g \\<^sub>S h = \C.mkarr (\<^bold>\\C.rep f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\\C.rep g\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\\C.rep h\<^bold>\)" proof - have "... = \C.mkarr (\<^bold>\\C.rep f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\\<^bold>\\C.rep g\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\\C.rep h\<^bold>\\<^bold>\)" using assms Diag_Diagonalize TensorDiag_preserves_Diag arr_char by force also have "... = \C.mkarr (\<^bold>\\C.rep f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ (\<^bold>\\C.rep (\C.mkarr (\<^bold>\\C.rep g\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\\C.rep h\<^bold>\))\<^bold>\))" using assms \C.rep_mkarr TensorDiag_in_Hom Diag_Diagonalize arr_char by force also have "... = \C.mkarr (\<^bold>\\C.rep f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\\C.rep (g \\<^sub>S h)\<^bold>\)" using assms tensor_char by simp also have "... = f \\<^sub>S g \\<^sub>S h" using assms Diag_Diagonalize TensorDiag_preserves_Diag Diag_implies_Arr \C.COD_mkarr \C.DOM_mkarr \C.strict_arr_char tensor_char by simp finally show ?thesis by blast qed qed lemma tensor_assoc: assumes "arr f" and "arr g" and "arr h" shows "(f \\<^sub>S g) \\<^sub>S h = f \\<^sub>S g \\<^sub>S h" using assms tensor_tensor by presburger lemma arr_unity: shows "arr \" using \C.rep_unity \C.Par_Arr_norm \C.\_agreement \C.strict_arr_char by force lemma tensor_unity_arr: assumes "arr f" shows "\ \\<^sub>S f = f" using assms arr_unity tensor_char \C.strict_arr_char \C.mkarr_Diagonalize_rep by simp lemma tensor_arr_unity: assumes "arr f" shows "f \\<^sub>S \ = f" using assms arr_unity tensor_char \C.strict_arr_char \C.mkarr_Diagonalize_rep by simp lemma assoc_char: assumes "ide a" and "ide b" and "ide c" shows "\\<^sub>S[a, b, c] = \C.mkarr (\<^bold>\\C.rep a\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\\C.rep b\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\\C.rep c\<^bold>\)" using assms tensor_tensor(2) assoc\<^sub>S_def ideD(1) by simp lemma assoc_in_hom: assumes "ide a" and "ide b" and "ide c" shows "\\\<^sub>S[a, b, c] : (a \\<^sub>S b) \\<^sub>S c \ a \\<^sub>S b \\<^sub>S c\" using assms tensor_preserves_ide ideD tensor_assoc assoc\<^sub>S_def by (metis (no_types, lifting) ide_in_hom) text \The category \\\<^sub>SC\ is a monoidal category.\ interpretation EMC: elementary_monoidal_category comp tensor\<^sub>S \ \\a. a\ \\a. a\ assoc\<^sub>S proof show "ide \" using ide_char arr_char \C.rep_mkarr \C.Dom_norm \C.Cod_norm \C.\_agreement by auto show "\a. ide a \ iso a" using ide_char arr_char iso_char by auto show "\f a b g c d. \ in_hom a b f; in_hom c d g \ \ in_hom (a \\<^sub>S c) (b \\<^sub>S d) (f \\<^sub>S g)" using tensor_in_hom by blast show "\a b. \ ide a; ide b \ \ ide (a \\<^sub>S b)" using tensor_preserves_ide by blast show "\a b c. \ ide a; ide b; ide c\ \ iso \\<^sub>S[a, b, c]" using tensor_preserves_ide ide_is_iso assoc\<^sub>S_def by presburger show "\a b c. \ ide a; ide b; ide c\ \ \\\<^sub>S[a, b, c] : (a \\<^sub>S b) \\<^sub>S c \ a \\<^sub>S b \\<^sub>S c\" using assoc_in_hom by blast show "\a b. \ ide a; ide b \ \ (a \\<^sub>S b) \\<^sub>S \\<^sub>S[a, \, b] = a \\<^sub>S b" using ide_def tensor_unity_arr assoc\<^sub>S_def ideD(1) tensor_preserves_ide comp_ide_self by simp show "\f. arr f \ cod f \\<^sub>S (\ \\<^sub>S f) = f \\<^sub>S dom f" using tensor_unity_arr comp_arr_dom comp_cod_arr by presburger show "\f. arr f \ cod f \\<^sub>S (f \\<^sub>S \) = f \\<^sub>S dom f" using tensor_arr_unity comp_arr_dom comp_cod_arr by presburger next fix a assume a: "ide a" show "\a : \ \\<^sub>S a \ a\" using a tensor_unity_arr ide_in_hom [of a] by fast show "\a : a \\<^sub>S \ \ a\" using a tensor_arr_unity ide_in_hom [of a] by fast next fix f g f' g' assume fg: "seq g f" assume fg': "seq g' f'" show "(g \\<^sub>S g') \\<^sub>S (f \\<^sub>S f') = g \\<^sub>S f \\<^sub>S g' \\<^sub>S f'" proof - have A: "\C.seq g f" and B: "\C.seq g' f'" using fg fg' seq_char by auto have "(g \\<^sub>S g') \\<^sub>S (f \\<^sub>S f') = \C.D ((g \ g') \ (f \ f'))" using A B tensor\<^sub>S_def by fastforce also have "... = \C.D (g \ f \ g' \ f')" using A B \C.interchange \C.T_simp \C.seqE by metis also have "... = \C.D (g \ f) \\<^sub>S \C.D (g' \ f')" using A B tensor\<^sub>S_def \C.T_simp \C.seqE \C.diagonalize_tensor arr_char by (metis (no_types, lifting) D.preserves_reflects_arr) also have "... = \C.D g \\<^sub>S \C.D f \\<^sub>S \C.D g' \\<^sub>S \C.D f'" using A B by simp also have "... = g \\<^sub>S f \\<^sub>S g' \\<^sub>S f'" using fg fg' \C.diagonalize_strict_arr by (elim seqE, simp) finally show ?thesis by blast qed next fix f0 f1 f2 assume f0: "arr f0" and f1: "arr f1" and f2: "arr f2" show "\\<^sub>S[cod f0, cod f1, cod f2] \\<^sub>S ((f0 \\<^sub>S f1) \\<^sub>S f2) = (f0 \\<^sub>S f1 \\<^sub>S f2) \\<^sub>S \\<^sub>S[dom f0, dom f1, dom f2]" using f0 f1 f2 assoc\<^sub>S_def tensor_assoc dom_tensor cod_tensor arr_tensor comp_cod_arr [of "f0 \\<^sub>S f1 \\<^sub>S f2" "cod f0 \\<^sub>S cod f1 \\<^sub>S cod f2"] comp_arr_dom [of "f0 \\<^sub>S f1 \\<^sub>S f2" "dom f0 \\<^sub>S dom f1 \\<^sub>S dom f2"] by presburger next fix a b c d assume a: "ide a" and b: "ide b" and c: "ide c" and d: "ide d" show "(a \\<^sub>S \\<^sub>S[b, c, d]) \\<^sub>S \\<^sub>S[a, b \\<^sub>S c, d] \\<^sub>S (\\<^sub>S[a, b, c] \\<^sub>S d) = \\<^sub>S[a, b, c \\<^sub>S d] \\<^sub>S \\<^sub>S[a \\<^sub>S b, c, d]" unfolding assoc\<^sub>S_def using a b c d tensor_assoc tensor_preserves_ide ideD tensor_in_hom comp_arr_dom [of "a \\<^sub>S b \\<^sub>S c \\<^sub>S d"] by simp qed lemma is_elementary_monoidal_category: shows "elementary_monoidal_category comp tensor\<^sub>S \ (\a. a) (\a. a) assoc\<^sub>S" .. abbreviation T\<^sub>F\<^sub>S\<^sub>M\<^sub>C where "T\<^sub>F\<^sub>S\<^sub>M\<^sub>C \ EMC.T" abbreviation \\<^sub>F\<^sub>S\<^sub>M\<^sub>C where "\\<^sub>F\<^sub>S\<^sub>M\<^sub>C \ EMC.\" abbreviation \\<^sub>F\<^sub>S\<^sub>M\<^sub>C where "\\<^sub>F\<^sub>S\<^sub>M\<^sub>C \ EMC.\" lemma is_monoidal_category: shows "monoidal_category comp T\<^sub>F\<^sub>S\<^sub>M\<^sub>C \\<^sub>F\<^sub>S\<^sub>M\<^sub>C \\<^sub>F\<^sub>S\<^sub>M\<^sub>C" using EMC.induces_monoidal_category by auto end sublocale free_strict_monoidal_category \ elementary_monoidal_category comp tensor\<^sub>S \ "\a. a" "\a. a" assoc\<^sub>S using is_elementary_monoidal_category by auto sublocale free_strict_monoidal_category \ monoidal_category comp T\<^sub>F\<^sub>S\<^sub>M\<^sub>C \\<^sub>F\<^sub>S\<^sub>M\<^sub>C \\<^sub>F\<^sub>S\<^sub>M\<^sub>C using is_monoidal_category by auto sublocale free_strict_monoidal_category \ strict_monoidal_category comp T\<^sub>F\<^sub>S\<^sub>M\<^sub>C \\<^sub>F\<^sub>S\<^sub>M\<^sub>C \\<^sub>F\<^sub>S\<^sub>M\<^sub>C using tensor_preserves_ide lunit_agreement runit_agreement \_ide_simp assoc\<^sub>S_def by unfold_locales auto context free_strict_monoidal_category begin text \ The inclusion of generators functor from @{term C} to \\\<^sub>SC\ is the composition of the inclusion of generators from @{term C} to \\C\ and the diagonalization functor, which projects \\C\ to \\\<^sub>SC\. As the diagonalization functor is the identity map on the image of @{term C}, the composite functor amounts to the corestriction to \\\<^sub>SC\ of the inclusion of generators of \\C\. \ interpretation D: "functor" \C.comp comp \C.D using \C.diagonalize_is_functor by auto interpretation I: composite_functor C \C.comp comp \C.inclusion_of_generators \C.D proof - interpret "functor" C \C.comp \C.inclusion_of_generators using \C.inclusion_is_functor by blast show "composite_functor C \C.comp comp \C.inclusion_of_generators \C.D" .. qed definition inclusion_of_generators where "inclusion_of_generators \ \C.inclusion_of_generators" lemma inclusion_is_functor: shows "functor C comp inclusion_of_generators" using \C.DoI_eq_I I.functor_axioms inclusion_of_generators_def by auto text \ The diagonalization functor is strict monoidal. \ interpretation D: strict_monoidal_functor \C.comp \C.T\<^sub>F\<^sub>M\<^sub>C \C.\\<^sub>F\<^sub>M\<^sub>C \C.\\<^sub>F\<^sub>M\<^sub>C comp T\<^sub>F\<^sub>S\<^sub>M\<^sub>C \\<^sub>F\<^sub>S\<^sub>M\<^sub>C \\<^sub>F\<^sub>S\<^sub>M\<^sub>C \C.D proof show "\C.D \C.\ = \" proof - have "\C.D \C.\ = \C.mkarr \<^bold>\\C.rep \C.\\<^bold>\" unfolding \C.D_def using \C.\_in_hom by auto also have "... = \C.mkarr \<^bold>\\<^bold>\\<^bold>[\<^bold>\\<^bold>\\<^bold>\\<^bold>]\<^bold>\" using \C.\_def \C.rep_unity \C.rep_lunit \C.Par_Arr_norm \C.Diagonalize_norm by auto also have "... = \" using \C.unity\<^sub>F\<^sub>M\<^sub>C_def \C.\_agreement \_def by simp finally show ?thesis by blast qed show "\f g. \ \C.arr f; \C.arr g \ \ \C.D (\C.tensor f g) = tensor (\C.D f) (\C.D g)" proof - fix f g assume f: "\C.arr f" and g: "\C.arr g" have fg: "arr (\C.D f) \ arr (\C.D g)" using f g D.preserves_arr by blast have "\C.D (\C.tensor f g) = f \\<^sub>S g" using tensor\<^sub>S_def by simp also have "f \\<^sub>S g = \C.D (f \ g)" using f g tensor\<^sub>S_def by simp also have "... = \C.D f \\<^sub>S \C.D g" using f g fg tensor\<^sub>S_def \C.T_simp \C.diagonalize_tensor arr_char by (metis (no_types, lifting)) also have "... = tensor (\C.D f) (\C.D g)" using fg T_simp by simp finally show "\C.D (\C.tensor f g) = tensor (\C.D f) (\C.D g)" by blast qed show "\a b c. \ \C.ide a; \C.ide b; \C.ide c \ \ \C.D (\C.assoc a b c) = assoc (\C.D a) (\C.D b) (\C.D c)" proof - fix a b c assume a: "\C.ide a" and b: "\C.ide b" and c: "\C.ide c" have abc: "ide (\C.D a) \ ide (\C.D b) \ ide (\C.D c)" using a b c D.preserves_ide by blast have abc': "\C.ide (\C.D a) \ \C.ide (\C.D b) \ \C.ide (\C.D c)" using a b c D.preserves_ide ide_char by simp have 1: "\f g. \C.arr f \ \C.arr g \ f \\<^sub>S g = \C.D (f \ g)" using tensor\<^sub>S_def by simp have 2: "\f. ide f \ \C.ide f" using ide_char by blast have "assoc (\C.D a) (\C.D b) (\C.D c) = \C.D a \\<^sub>S \C.D b \\<^sub>S \C.D c" using abc \_ide_simp assoc\<^sub>S_def by simp also have "... = \C.D a \\<^sub>S \C.D (\C.D b \ \C.D c)" using abc' 1 by auto also have "... = \C.D a \\<^sub>S \C.D (b \ c)" using b c \C.diagonalize_tensor by force also have "... = \C.D (\C.D a \ \C.D (b \ c))" using 1 b c abc D.preserves_ide \C.tensor_preserves_ide ide_char by simp also have "... = \C.D (a \ b \ c)" using a b c \C.diagonalize_tensor by force also have "... = \C.D \[a, b, c]" proof - have "\C.can \[a, b, c]" using a b c \C.can_assoc by simp hence "\C.ide (\C.D \[a, b, c])" using a b c \C.ide_diagonalize_can by simp moreover have "\C.cod (\C.D \[a, b, c]) = \C.D (a \ b \ c)" using a b c \C.assoc_in_hom D.preserves_hom by (metis (no_types, lifting) cod_char in_homE) ultimately show ?thesis by simp qed also have "... = \C.D (\C.assoc a b c)" using a b c by simp finally show "\C.D (\C.assoc a b c) = assoc (\C.D a) (\C.D b) (\C.D c)" by blast qed qed lemma diagonalize_is_strict_monoidal_functor: shows "strict_monoidal_functor \C.comp \C.T\<^sub>F\<^sub>M\<^sub>C \C.\\<^sub>F\<^sub>M\<^sub>C \C.\\<^sub>F\<^sub>M\<^sub>C comp T\<^sub>F\<^sub>S\<^sub>M\<^sub>C \\<^sub>F\<^sub>S\<^sub>M\<^sub>C \\<^sub>F\<^sub>S\<^sub>M\<^sub>C \C.D" .. interpretation \: natural_isomorphism \C.CC.comp comp D.T\<^sub>DoFF.map D.FoT\<^sub>C.map D.\ using D.structure_is_natural_isomorphism by simp text \ The diagonalization functor is part of a monoidal equivalence between the free monoidal category and the subcategory @{term "\\<^sub>SC"}. \ interpretation E: equivalence_of_categories comp \C.comp \C.D map \C.\ \C.\ using \C.is_equivalent_to_strict_subcategory by auto interpretation D: monoidal_functor \C.comp \C.T\<^sub>F\<^sub>M\<^sub>C \C.\\<^sub>F\<^sub>M\<^sub>C \C.\\<^sub>F\<^sub>M\<^sub>C comp T\<^sub>F\<^sub>S\<^sub>M\<^sub>C \\<^sub>F\<^sub>S\<^sub>M\<^sub>C \\<^sub>F\<^sub>S\<^sub>M\<^sub>C \C.D D.\ using D.monoidal_functor_axioms by metis interpretation equivalence_of_monoidal_categories comp T\<^sub>F\<^sub>S\<^sub>M\<^sub>C \\<^sub>F\<^sub>S\<^sub>M\<^sub>C \\<^sub>F\<^sub>S\<^sub>M\<^sub>C \C.comp \C.T\<^sub>F\<^sub>M\<^sub>C \C.\\<^sub>F\<^sub>M\<^sub>C \C.\\<^sub>F\<^sub>M\<^sub>C \C.D D.\ \ map \C.\ \C.\ .. text \ The category @{term "\C"} is monoidally equivalent to its subcategory @{term "\\<^sub>SC"}. \ theorem monoidally_equivalent_to_free_monoidal_category: shows "equivalence_of_monoidal_categories comp T\<^sub>F\<^sub>S\<^sub>M\<^sub>C \\<^sub>F\<^sub>S\<^sub>M\<^sub>C \\<^sub>F\<^sub>S\<^sub>M\<^sub>C \C.comp \C.T\<^sub>F\<^sub>M\<^sub>C \C.\\<^sub>F\<^sub>M\<^sub>C \C.\\<^sub>F\<^sub>M\<^sub>C \C.D D.\ map \C.\ \C.\" .. end text \ We next show that the evaluation functor induced on the free monoidal category generated by @{term C} by a functor @{term V} from @{term C} to a strict monoidal category @{term D} restricts to a strict monoidal functor on the subcategory @{term "\\<^sub>SC"}. \ locale strict_evaluation_functor = D: strict_monoidal_category D T\<^sub>D \\<^sub>D \\<^sub>D + evaluation_map C D T\<^sub>D \\<^sub>D \\<^sub>D V + \C: free_monoidal_category C + E: evaluation_functor C D T\<^sub>D \\<^sub>D \\<^sub>D V + \\<^sub>SC: free_strict_monoidal_category C for C :: "'c comp" (infixr "\\<^sub>C" 55) and D :: "'d comp" (infixr "\\<^sub>D" 55) and T\<^sub>D :: "'d * 'd \ 'd" and \\<^sub>D :: "'d * 'd * 'd \ 'd" and \\<^sub>D :: "'d" and V :: "'c \ 'd" begin notation \C.in_hom ("\_ : _ \ _\") notation \\<^sub>SC.in_hom ("\_ : _ \\<^sub>S _\") (* TODO: This is just the restriction of the evaluation functor to a subcategory. It would be useful to define a restriction_of_functor locale that does this in general and gives the lemma that it yields a functor. *) definition map where "map \ \f. if \\<^sub>SC.arr f then E.map f else D.null" interpretation "functor" \\<^sub>SC.comp D map unfolding map_def apply unfold_locales apply simp using \\<^sub>SC.arr_char E.preserves_arr apply simp using \\<^sub>SC.arr_char \\<^sub>SC.dom_char E.preserves_dom apply simp using \\<^sub>SC.arr_char \\<^sub>SC.cod_char E.preserves_cod apply simp using \\<^sub>SC.arr_char \\<^sub>SC.dom_char \\<^sub>SC.cod_char \\<^sub>SC.comp_char E.preserves_comp by (elim \\<^sub>SC.seqE, auto) lemma is_functor: shows "functor \\<^sub>SC.comp D map" .. text \ Every canonical arrow is an equivalence class of canonical terms. The evaluations in \D\ of all such terms are identities, due to the strictness of \D\. \ lemma ide_eval_Can: shows "Can t \ D.ide \t\" proof (induct t) show "\x. Can \<^bold>\x\<^bold>\ \ D.ide \\<^bold>\x\<^bold>\\" by simp show "Can \<^bold>\ \ D.ide \\<^bold>\\" by simp show "\t1 t2. \ Can t1 \ D.ide \t1\; Can t2 \ D.ide \t2\; Can (t1 \<^bold>\ t2) \ \ D.ide \t1 \<^bold>\ t2\" by simp show "\t1 t2. \ Can t1 \ D.ide \t1\; Can t2 \ D.ide \t2\; Can (t1 \<^bold>\ t2) \ \ D.ide \t1 \<^bold>\ t2\" proof - fix t1 t2 assume t1: "Can t1 \ D.ide \t1\" and t2: "Can t2 \ D.ide \t2\" and t12: "Can (t1 \<^bold>\ t2)" show "D.ide \t1 \<^bold>\ t2\" using t1 t2 t12 Can_implies_Arr eval_in_hom [of t1] eval_in_hom [of t2] D.comp_ide_arr by fastforce qed show "\t. (Can t \ D.ide \t\) \ Can \<^bold>\\<^bold>[t\<^bold>] \ D.ide \\<^bold>\\<^bold>[t\<^bold>]\" using D.strict_lunit by simp show "\t. (Can t \ D.ide \t\) \ Can \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] \ D.ide \\<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\" using D.strict_lunit by simp show "\t. (Can t \ D.ide \t\) \ Can \<^bold>\\<^bold>[t\<^bold>] \ D.ide \\<^bold>\\<^bold>[t\<^bold>]\" using D.strict_runit by simp show "\t. (Can t \ D.ide \t\) \ Can \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] \ D.ide \\<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\" using D.strict_runit by simp fix t1 t2 t3 assume t1: "Can t1 \ D.ide \t1\" and t2: "Can t2 \ D.ide \t2\" and t3: "Can t3 \ D.ide \t3\" show "Can \<^bold>\\<^bold>[t1, t2, t3\<^bold>] \ D.ide \\<^bold>\\<^bold>[t1, t2, t3\<^bold>]\" proof - assume "Can \<^bold>\\<^bold>[t1, t2, t3\<^bold>]" hence t123: "D.ide \t1\ \ D.ide \t2\ \ D.ide \t3\" using t1 t2 t3 by simp have "\\<^bold>\\<^bold>[t1, t2, t3\<^bold>]\ = \t1\ \\<^sub>D \t2\ \\<^sub>D \t3\" using t123 D.strict_assoc D.assoc_in_hom [of "\t1\" "\t2\" "\t3\"] apply simp by (elim D.in_homE, simp) thus ?thesis using t123 by simp qed show "Can \<^bold>\\<^sup>-\<^sup>1\<^bold>[t1, t2, t3\<^bold>] \ D.ide \\<^bold>\\<^sup>-\<^sup>1\<^bold>[t1, t2, t3\<^bold>]\" proof - assume "Can \<^bold>\\<^sup>-\<^sup>1\<^bold>[t1, t2, t3\<^bold>]" hence t123: "Can t1 \ Can t2 \ Can t3 \ D.ide \t1\ \ D.ide \t2\ \ D.ide \t3\" using t1 t2 t3 by simp have "\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t1, t2, t3\<^bold>]\ = D.inv \\<^sub>D[D.cod \t1\, D.cod \t2\, D.cod \t3\] \\<^sub>D (\t1\ \\<^sub>D \t2\ \\<^sub>D \t3\)" using t123 eval_Assoc' [of t1 t2 t3] Can_implies_Arr by simp also have "... = \t1\ \\<^sub>D \t2\ \\<^sub>D \t3\" proof - have "D.dom \\<^sub>D[\t1\, \t2\, \t3\] = \t1\ \\<^sub>D \t2\ \\<^sub>D \t3\" proof - have "D.dom \\<^sub>D[\t1\, \t2\, \t3\] = D.cod \\<^sub>D[\t1\, \t2\, \t3\]" using t123 D.strict_assoc by simp also have "... = \t1\ \\<^sub>D \t2\ \\<^sub>D \t3\" using t123 by simp finally show ?thesis by blast qed thus ?thesis using t123 D.strict_assoc D.comp_arr_dom by auto qed finally have "\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t1, t2, t3\<^bold>]\ = \t1\ \\<^sub>D \t2\ \\<^sub>D \t3\" by blast thus ?thesis using t123 by auto qed qed lemma ide_eval_can: assumes "\C.can f" shows "D.ide (E.map f)" proof - have "f = \C.mkarr (\C.rep f)" using assms \C.can_implies_arr \C.mkarr_rep by blast moreover have 1: "Can (\C.rep f)" using assms \C.Can_rep_can by simp moreover have "D.ide \\C.rep f\" using assms ide_eval_Can by (simp add: 1) ultimately show ?thesis using assms \C.can_implies_arr E.map_def by force qed text \ Diagonalization transports formal arrows naturally along reductions, which are canonical terms and therefore evaluate to identities of \D\. It follows that the evaluation in \D\ of a formal arrow is equal to the evaluation of its diagonalization. \ lemma map_diagonalize: assumes f: "\C.arr f" shows "E.map (\C.D f) = E.map f" proof - interpret EQ: equivalence_of_categories \\<^sub>SC.comp \C.comp \C.D \\<^sub>SC.map \C.\ \C.\ using \C.is_equivalent_to_strict_subcategory by auto have 1: "\C.seq (\\<^sub>SC.map (\C.D f)) (\C.\ (\C.dom f))" proof show "\\C.\ (\C.dom f) : \C.dom f \ \C.D (\C.dom f)\" using f \\<^sub>SC.map_simp EQ.F.preserves_arr by (intro \C.in_homI, simp_all) show "\\\<^sub>SC.map (\C.D f) : \C.D (\C.dom f) \ \C.cod (\C.D f)\" - using f \\<^sub>SC.map_simp \C.arr_iff_in_hom EQ.F.preserves_hom \\<^sub>SC.arr_char - \\<^sub>SC.in_hom_char [of "\C.D f" "\C.D (\C.dom f)" "\C.D (\C.cod f)"] - by (intro \C.in_homI, auto) + by (metis (no_types, lifting) EQ.F.preserves_dom EQ.F.preserves_reflects_arr + \\<^sub>SC.arr_iff_in_hom \\<^sub>SC.cod_simp \\<^sub>SC.in_hom_char \\<^sub>SC.map_simp f) qed have "E.map (\C.\ (\C.cod f)) \\<^sub>D E.map f = E.map (\C.D f) \\<^sub>D E.map (\C.\ (\C.dom f))" proof - have "E.map (\C.\ (\C.cod f)) \\<^sub>D E.map f = E.map (\C.\ (\C.cod f) \ f)" using f by simp also have "... = E.map (\C.D f \ \C.\ (\C.dom f))" using f EQ.\.naturality \\<^sub>SC.map_simp EQ.F.preserves_arr by simp also have "... = E.map (\\<^sub>SC.map (\C.D f)) \\<^sub>D E.map (\C.\ (\C.dom f))" using f 1 E.as_nat_trans.preserves_comp_2 EQ.F.preserves_arr \\<^sub>SC.map_simp by (metis (no_types, lifting)) also have "... = E.map (\C.D f) \\<^sub>D E.map (\C.\ (\C.dom f))" using f EQ.F.preserves_arr \\<^sub>SC.map_simp by simp finally show ?thesis by blast qed moreover have "\a. \C.ide a \ D.ide (E.map (\C.\ a))" using \C.\_def \C.Arr_rep Arr_implies_Ide_Cod Can_red \C.can_mkarr_Can ide_eval_can by (metis (no_types, lifting) EQ.\.preserves_reflects_arr \C.seqE \C.comp_preserves_can \C.ideD(1) \C.ide_implies_can) moreover have "D.cod (E.map f) = D.dom (E.map (\C.\ (\C.cod f)))" using f E.preserves_hom EQ.\.preserves_hom by simp moreover have "D.dom (E.map (\C.D f)) = D.cod (E.map (\C.\ (\C.dom f)))" using f 1 E.preserves_seq EQ.F.preserves_arr \\<^sub>SC.map_simp by auto ultimately show ?thesis using f D.comp_arr_dom D.ideD D.arr_dom_iff_arr E.as_nat_trans.is_natural_2 by (metis E.preserves_cod \C.ide_cod \C.ide_dom) qed lemma strictly_preserves_tensor: assumes "\\<^sub>SC.arr f" and "\\<^sub>SC.arr g" shows "map (\\<^sub>SC.tensor f g) = map f \\<^sub>D map g" proof - have 1: "\C.arr (f \ g)" using assms \\<^sub>SC.arr_char \C.tensor_in_hom by auto have 2: "\\<^sub>SC.arr (\\<^sub>SC.tensor f g)" using assms \\<^sub>SC.tensor_in_hom [of f g] \\<^sub>SC.T_simp by fastforce have "map (\\<^sub>SC.tensor f g) = E.map (f \ g)" proof - have "map (\\<^sub>SC.tensor f g) = map (f \\<^sub>S g)" using assms \\<^sub>SC.T_simp by simp also have "... = map (\C.D (f \ g))" using assms \C.tensor\<^sub>F\<^sub>M\<^sub>C_def \\<^sub>SC.tensor\<^sub>S_def \\<^sub>SC.arr_char by force also have "... = E.map (f \ g)" proof - interpret Diag: "functor" \C.comp \\<^sub>SC.comp \C.D using \C.diagonalize_is_functor by auto show ?thesis using assms 1 map_diagonalize [of "f \ g"] Diag.preserves_arr map_def by simp qed finally show ?thesis by blast qed thus ?thesis using assms \\<^sub>SC.arr_char E.strictly_preserves_tensor map_def by simp qed lemma is_strict_monoidal_functor: shows "strict_monoidal_functor \\<^sub>SC.comp \\<^sub>SC.T\<^sub>F\<^sub>S\<^sub>M\<^sub>C \\<^sub>SC.\ \\<^sub>SC.\ D T\<^sub>D \\<^sub>D \\<^sub>D map" proof show "\f g. \\<^sub>SC.arr f \ \\<^sub>SC.arr g \ map (\\<^sub>SC.tensor f g) = map f \\<^sub>D map g" using strictly_preserves_tensor by fast show "map \\<^sub>SC.\ = \\<^sub>D" using \\<^sub>SC.arr_unity \\<^sub>SC.\_def map_def E.map_def \C.rep_mkarr E.eval_norm D.strict_unit by auto fix a b c assume a: "\\<^sub>SC.ide a" and b: "\\<^sub>SC.ide b" and c: "\\<^sub>SC.ide c" show "map (\\<^sub>SC.assoc a b c) = \\<^sub>D[map a, map b, map c]" proof - have "map (\\<^sub>SC.assoc a b c) = map a \\<^sub>D map b \\<^sub>D map c" using a b c \\<^sub>SC.\_def \\<^sub>SC.assoc\<^sub>S_def \\<^sub>SC.arr_tensor \\<^sub>SC.T_simp \\<^sub>SC.ideD(1) strictly_preserves_tensor \\<^sub>SC.\_ide_simp by presburger also have "... = \\<^sub>D[map a, map b, map c]" using a b c D.strict_assoc D.assoc_in_hom [of "map a" "map b" "map c"] by auto finally show ?thesis by blast qed qed end sublocale strict_evaluation_functor \ strict_monoidal_functor \\<^sub>SC.comp \\<^sub>SC.T\<^sub>F\<^sub>S\<^sub>M\<^sub>C \\<^sub>SC.\ \\<^sub>SC.\ D T\<^sub>D \\<^sub>D \\<^sub>D map using is_strict_monoidal_functor by auto locale strict_monoidal_extension_to_free_strict_monoidal_category = C: category C + monoidal_language C + \\<^sub>SC: free_strict_monoidal_category C + strict_monoidal_extension C \\<^sub>SC.comp \\<^sub>SC.T\<^sub>F\<^sub>S\<^sub>M\<^sub>C \\<^sub>SC.\ \\<^sub>SC.\ D T\<^sub>D \\<^sub>D \\<^sub>D \\<^sub>SC.inclusion_of_generators V F for C :: "'c comp" (infixr "\\<^sub>C" 55) and D :: "'d comp" (infixr "\\<^sub>D" 55) and T\<^sub>D :: "'d * 'd \ 'd" and \\<^sub>D :: "'d * 'd * 'd \ 'd" and \\<^sub>D :: "'d" and V :: "'c \ 'd" and F :: "'c free_monoidal_category.arr \ 'd" sublocale strict_evaluation_functor \ strict_monoidal_extension C \\<^sub>SC.comp \\<^sub>SC.T\<^sub>F\<^sub>S\<^sub>M\<^sub>C \\<^sub>SC.\ \\<^sub>SC.\ D T\<^sub>D \\<^sub>D \\<^sub>D \\<^sub>SC.inclusion_of_generators V map proof - interpret V: "functor" C \\<^sub>SC.comp \\<^sub>SC.inclusion_of_generators using \\<^sub>SC.inclusion_is_functor by auto show "strict_monoidal_extension C \\<^sub>SC.comp \\<^sub>SC.T\<^sub>F\<^sub>S\<^sub>M\<^sub>C \\<^sub>SC.\ \\<^sub>SC.\ D T\<^sub>D \\<^sub>D \\<^sub>D \\<^sub>SC.inclusion_of_generators V map" proof show "\f. C.arr f \ map (\\<^sub>SC.inclusion_of_generators f) = V f" using V.preserves_arr E.is_extension map_def \\<^sub>SC.inclusion_of_generators_def by simp qed qed context free_strict_monoidal_category begin text \ We now have the main result of this section: the evaluation functor on \\\<^sub>SC\ induced by a functor @{term V} from @{term C} to a strict monoidal category @{term D} is the unique strict monoidal extension of @{term V} to \\\<^sub>SC\. \ theorem is_free: assumes "strict_monoidal_category D T\<^sub>D \\<^sub>D \\<^sub>D" and "strict_monoidal_extension_to_free_strict_monoidal_category C D T\<^sub>D \\<^sub>D \\<^sub>D V F" shows "F = strict_evaluation_functor.map C D T\<^sub>D \\<^sub>D \\<^sub>D V" proof - interpret D: strict_monoidal_category D T\<^sub>D \\<^sub>D \\<^sub>D using assms(1) by auto text \ Let @{term F} be a given extension of V to a strict monoidal functor defined on \\\<^sub>SC\. \ interpret F: strict_monoidal_extension_to_free_strict_monoidal_category C D T\<^sub>D \\<^sub>D \\<^sub>D V F using assms(2) by auto text \ Let @{term E\<^sub>S} be the evaluation functor from \\\<^sub>SC\ to @{term D} induced by @{term V}. Then @{term E\<^sub>S} is also a strict monoidal extension of @{term V}. \ interpret E\<^sub>S: strict_evaluation_functor C D T\<^sub>D \\<^sub>D \\<^sub>D V .. text \ Let @{term D} be the strict monoidal functor @{term "\C.D"} that projects \\C\ to the subcategory \\\<^sub>SC\. \ interpret D: "functor" \C.comp comp \C.D using \C.diagonalize_is_functor by auto interpret D: strict_monoidal_functor \C.comp \C.T\<^sub>F\<^sub>M\<^sub>C \C.\ \C.\ comp T\<^sub>F\<^sub>S\<^sub>M\<^sub>C \ \ \C.D using diagonalize_is_strict_monoidal_functor by blast text \ The composite functor \F o D\ is also an extension of @{term V} to a strict monoidal functor on \\C\. \ interpret FoD: composite_functor \C.comp comp D \C.D F .. interpret FoD: strict_monoidal_functor \C.comp \C.T\<^sub>F\<^sub>M\<^sub>C \C.\ \C.\ D T\<^sub>D \\<^sub>D \\<^sub>D \F o \C.D\ using D.strict_monoidal_functor_axioms F.strict_monoidal_functor_axioms strict_monoidal_functors_compose by fast interpret FoD: strict_monoidal_extension_to_free_monoidal_category C D T\<^sub>D \\<^sub>D \\<^sub>D V FoD.map proof show "\f. C.arr f \ FoD.map (\C.inclusion_of_generators f) = V f" proof - have "\f. C.arr f \ FoD.map (\C.inclusion_of_generators f) = V f" proof - fix f assume f: "C.arr f" have "FoD.map (\C.inclusion_of_generators f) = F (\C.D (\C.inclusion_of_generators f))" using f by simp also have "... = F (inclusion_of_generators f)" using f \C.strict_arr_char' F.I.preserves_arr inclusion_of_generators_def by simp also have "... = V f" using f F.is_extension by simp finally show "FoD.map (\C.inclusion_of_generators f) = V f" by blast qed thus ?thesis by blast qed qed text \ By the freeness of \\C\, we have that \F o D\ is equal to the evaluation functor @{term "E\<^sub>S.E.map"} induced by @{term V} on \\C\. Moreover, @{term E\<^sub>S.map} coincides with @{term "E\<^sub>S.E.map"} on \\\<^sub>SC\ and \F o D\ coincides with @{term F} on \\\<^sub>SC\. Therefore, @{term F} coincides with @{term E} on their common domain \\\<^sub>SC\, showing @{term "F = E\<^sub>S.map"}. \ have "\f. arr f \ F f = E\<^sub>S.map f" using \C.strict_arr_char' \C.is_free [of D] E\<^sub>S.E.evaluation_functor_axioms FoD.strict_monoidal_extension_to_free_monoidal_category_axioms E\<^sub>S.map_def by simp moreover have "\f. \arr f \ F f = E\<^sub>S.map f" using F.is_extensional E\<^sub>S.is_extensional arr_char by auto ultimately show "F = E\<^sub>S.map" by blast qed end end