Theory Limit
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 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: set_category S setp +
F: "functor" Cop.comp S F +
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 ‹⋅⇩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: 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 ‹⋅⇩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 ‹⋅⇩J› 55)
and C :: "'c comp" (infixr ‹⋅› 55)
and D :: "'j ⇒ 'c"
begin
notation J.in_hom (‹«_ : _ →⇩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 ‹⋅⇩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
lemma cod_determines_component:
assumes "J.arr j"
shows "χ j = χ (J.cod j)"
using assms is_natural_2 A.map_simp C.comp_arr_ide ide_apex preserves_reflects_arr
by metis
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_composition.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 ‹⋅⇩J› 55)
and C :: "'c comp" (infixr ‹⋅› 55)
and D :: "'j ⇒ 'c"
and S :: "'s comp" (infixr ‹⋅⇩S› 55)
and ι :: "('j ⇒ 'c) ⇒ 's"
begin
notation S.in_hom (‹«_ : _ →⇩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.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.UP_mapsto D.cones_map_ide by force
also have "... = S.mkIde (ι ` D.cones a)"
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 ⋅⇧o⇧p f) = map g ⋅⇩S map f"
proof -
have "map (g ⋅⇧o⇧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 ⋅⇧o⇧p f))"
using assms arr_map [of "C f g"] by simp
have "map (g ⋅⇧o⇧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 ⋅⇩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 ‹λ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 ‹⋅⇩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 ‹λA. A ⊆ S.Univ› φ +
representation_of_functor C S S.setp φ Cones.map a Φ
for J :: "'j comp" (infixr ‹⋅⇩J› 55)
and C :: "'c comp" (infixr ‹⋅› 55)
and D :: "'j ⇒ 'c"
and S :: "'s comp" (infixr ‹⋅⇩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 S.setp φ +
Φ: representation_of_functor C S S.setp φ Cones.map a Φ +
χ: limit_cone J C D a χ
for J :: "'j comp" (infixr ‹⋅⇩J› 55)
and C :: "'c comp" (infixr ‹⋅› 55)
and D :: "'j ⇒ 'c"
and S :: "'s comp" (infixr ‹⋅⇩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.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 ‹TYPE('c + ('j ⇒ 'c))› .
notation S.comp (infixr ‹⋅⇩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.setp ‹λ_. S.UP o Inl›
apply (unfold_locales)
using S.UP_mapsto
apply auto[1]
using S.inj_UP injD inj_onI inj_Inl inj_compose
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 is_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)) →⇩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)) →⇩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 is_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' →⇩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) ⋅⇩S Y a g = Cones.map g ⋅⇩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) ⋅⇩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 ⋅⇩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) ⋅⇩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 -
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 is_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 is_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 ⋅⇩S Φo (Cop.dom g)"
using 2 by auto
finally show ?thesis by auto
qed
qed
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' is_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_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' is_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.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.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.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 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] is_cone by simp
finally have "Φ.FUN a (φ (a, a) a) = ι χ" by blast
hence "𝗈 (Φ.FUN a (φ (a, a) a)) = 𝗈 (ι χ)" by simp
thus ?thesis using is_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 ‹⋅⇩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 S.setp ‹Y a› Cones.map Φ ..
interpretation Ψ: inverse_transformation Cop.comp S ‹Y a› Cones.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.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.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 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' ⋅⇩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)) →⇩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 χ χ.is_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 χ.is_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' χ.is_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 ‹⋅⇩J› 55)
and C :: "'c comp" (infixr ‹⋅› 55)
begin
notation J.in_hom (‹«_ : _ →⇩J _»›)
notation J_C.comp (infixr ‹⋅⇩[⇩J⇩,⇩C⇩]› 55)
notation J_C.in_hom (‹«_ : _ →⇩[⇩J⇩,⇩C⇩] _»›)
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 ⋅⇩[⇩J⇩,⇩C⇩] 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 ⋅⇩[⇩J⇩,⇩C⇩]
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 ⋅⇩[⇩J⇩,⇩C⇩] 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 →⇩[⇩J⇩,⇩C⇩] 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 →⇩[⇩J⇩,⇩C⇩] 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 χ ⋅⇩[⇩J⇩,⇩C⇩] 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 →⇩[⇩J⇩,⇩C⇩] 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' →⇩[⇩J⇩,⇩C⇩] 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 ⋅⇩[⇩J⇩,⇩C⇩] 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 ⋅⇩[⇩J⇩,⇩C⇩] map g))"
using assms(1) 2 by auto
also have "... = J_C.Dom (?x ⋅⇩[⇩J⇩,⇩C⇩] 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 ⋅⇩[⇩J⇩,⇩C⇩] 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 ⋅⇩[⇩J⇩,⇩C⇩] map g))"
using 2 3 J_C.cod_comp J_C.in_homE by metis
also have "... = J_C.Cod (?x ⋅⇩[⇩J⇩,⇩C⇩] 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 ⋅⇩[⇩J⇩,⇩C⇩] 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 ⋅⇩[⇩J⇩,⇩C⇩] 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' →⇩[⇩J⇩,⇩C⇩] 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 ⋅⇩[⇩J⇩,⇩C⇩] 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 ⋅⇩[⇩J⇩,⇩C⇩] 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 ⋅⇩[⇩J⇩,⇩C⇩] 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 ‹⋅⇩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 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' →⇩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' →⇩C ?a»"
using x.is_coext_def by simp
show "∃!g. «g : a' →⇩C ?a» ∧ D.cones_map g ?χ = χ'"
proof -
have "∃g. «g : a' →⇩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) limits_are_isomorphic:
fixes J :: "'j comp"
assumes "limit_cone J C D a χ" and "limit_cone J C D a' χ'"
shows "isomorphic a a'" and "iso (limit_cone.induced_arrow J C D a χ a' χ')"
proof -
interpret J: category J
using assms(1) limit_cone.axioms(2) by metis
interpret C: category C
using assms(1) limit_cone.axioms(1) by metis
interpret D: diagram J C D
using assms(1) limit_cone.axioms(3) by metis
interpret χ: limit_cone J C D a χ
using assms(1) by blast
interpret χ': limit_cone J C D a' χ'
using assms(2) by blast
have 1: "∃!f. «f : a → a'» ∧ D.cones_map f χ' = χ"
using χ'.is_universal [of a χ] χ.cone_axioms by simp
have 2: "∃!g. «g : a' → a» ∧ D.cones_map g χ = χ'"
using χ.is_universal [of a' χ'] χ'.cone_axioms by simp
define f where "f = χ'.induced_arrow a χ"
define g where "g = χ.induced_arrow a' χ'"
have f: "«f : a → a'» ∧ D.cones_map f χ' = χ"
using f_def χ'.induced_arrowI(1-2) χ.is_cone by blast
have g: "«g : a' → a» ∧ D.cones_map g χ = χ'"
using g_def χ.induced_arrowI(1-2) χ'.is_cone by blast
have *: "inverse_arrows f g"
proof
show "ide (g ⋅ f)"
proof -
have "g ⋅ f = a"
proof -
have "∃!h. «h : a → a» ∧ D.cones_map h χ = χ"
using χ.is_universal [of a χ] χ.cone_axioms by blast
moreover have "«g ⋅ f : a → a»"
using f g by blast
moreover have "D.cones_map (g ⋅ f) χ = χ"
proof
fix j :: 'j
show "D.cones_map (g ⋅ f) χ j = χ j"
proof (cases "J.arr j")
show "¬ J.arr j ⟹ ?thesis"
using f g χ.cone_axioms χ.is_extensional by fastforce
assume j: "J.arr j"
have "D.cone (dom g) (D.cones_map g χ)"
using g D.cones_map_mapsto χ.cone_axioms by blast
thus ?thesis
using f g χ.cone_axioms D.cones_map_comp [of g f] by fastforce
qed
qed
moreover have "«a : a → a»"
using χ.ide_apex by auto
moreover have "D.cones_map a χ = χ"
using f χ.cone_axioms D.cones_map_ide by blast
ultimately show ?thesis by blast
qed
thus ?thesis
using χ.ide_apex by blast
qed
show "ide (f ⋅ g)"
proof -
have "f ⋅ g = a'"
proof -
have "∃!h. «h : a' → a'» ∧ D.cones_map h χ' = χ'"
using χ'.is_universal [of a' χ'] χ'.cone_axioms by blast
moreover have "«f ⋅ g : a' → a'»"
using f g by blast
moreover have "D.cones_map (f ⋅ g) χ' = χ'"
proof
fix j :: 'j
show "D.cones_map (f ⋅ g) χ' j = χ' j"
proof (cases "J.arr j")
show "¬ J.arr j ⟹ ?thesis"
using f g χ'.cone_axioms χ'.is_extensional by fastforce
assume j: "J.arr j"
have "D.cone (dom f) (D.cones_map f χ')"
using f D.cones_map_mapsto χ'.cone_axioms by blast
thus ?thesis
using f g χ'.cone_axioms D.cones_map_comp [of f g] by fastforce
qed
qed
moreover have "«a' : a' → a'»"
using χ'.ide_apex by auto
moreover have "D.cones_map a' χ' = χ'"
using g χ'.cone_axioms D.cones_map_ide by blast
ultimately show ?thesis by blast
qed
thus ?thesis
using χ'.ide_apex by blast
qed
qed
show "isomorphic a a'"
using * f g by blast
show "iso (χ.induced_arrow a' χ')"
using * g_def by blast
qed
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 ⋅⇩[⇩J⇩,⇩C⇩] Δ.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 ⋅⇩[⇩J⇩,⇩C⇩] Δ.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 ⋅⇩[⇩J⇩,⇩C⇩] Δ.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 ⋅⇩[⇩J⇩,⇩C⇩] Δ.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 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)) ⋅⇩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 ⋅⇩C ?χ' (J.dom j) = ?χ' j"
proof -
have "E j ⋅⇩C ?χ' (J.dom j) = (E j ⋅⇩C Adj.ε (E (J.dom j))) ⋅⇩C F (κ (J.dom j))"
using j C.comp_assoc by simp
also have "... = Adj.ε (E (J.cod j)) ⋅⇩C F (κ j)"
proof -
have "(E j ⋅⇩C Adj.ε (E (J.dom j))) ⋅⇩C F (κ (J.dom j))
= (Adj.ε (C.cod (E j)) ⋅⇩C Adj.FG.map (E j)) ⋅⇩C F (κ (J.dom j))"
using j Adj.ε.naturality [of "E j"] by fastforce
also have "... = Adj.ε (C.cod (E j)) ⋅⇩C Adj.FG.map (E j) ⋅⇩C F (κ (J.dom j))"
using C.comp_assoc by simp
also have "... = Adj.ε (E (J.cod j)) ⋅⇩C F (κ j)"
proof -
have "Adj.FG.map (E j) ⋅⇩C F (κ (J.dom j)) = F (GE.map j ⋅⇩D κ (J.dom j))"
using j by simp
hence "Adj.FG.map (E j) ⋅⇩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) ⋅⇩C Fb.map j = ?χ' j"
proof -
have "?χ' (J.cod j) ⋅⇩C Fb.map j = Adj.ε (E (J.cod j)) ⋅⇩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)) ⋅⇩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 →⇩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 ⋅⇩D Adj.η b"
have 1: "«?g : b →⇩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) ⋅⇩D ?g"
using j 1 Gχ.cone_axioms mem_Collect_eq restrict_apply by auto
also have "... = G (χ j ⋅⇩C f) ⋅⇩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) ⋅⇩D Adj.η b"
proof -
have "χ j ⋅⇩C f = Adj.ε (C.cod (E j)) ⋅⇩C F (κ j)"
proof -
have "χ j ⋅⇩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)) ⋅⇩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))) ⋅⇩D Adj.η (D.cod (GE.map j))) ⋅⇩D κ j"
using j f Adj.η.naturality [of "κ j"] D.comp_assoc by auto
also have "... = D.cod (κ j) ⋅⇩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 →⇩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 →⇩D G a» ∧ GE.cones_map g' ?Gχ = κ ⟹ g' = ?g"
proof -
fix g'
assume g': "«g' : b →⇩D G a» ∧ GE.cones_map g' ?Gχ = κ"
have 1: "«ψ a g' : F b →⇩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 ⋅⇩C ψ a g'"
using 1 χ.cone_axioms χ.is_extensional by auto
also have "... = (χ j ⋅⇩C Adj.ε a) ⋅⇩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)) ⋅⇩C F (G (χ j))) ⋅⇩C F g'"
using j a g' Adj.ε.naturality [of "χ j"] by simp
also have "... = Adj.ε (C.cod (E j)) ⋅⇩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 →⇩C a» ∧ E.cones_map f χ = ?χ'"
using cone_χ' χ.is_universal by simp
moreover have "«ψ a g' : F b →⇩C a» ∧ E.cones_map (ψ a g') χ = ?χ'"
using 1 2 by simp
ultimately show ?thesis
using ex1E [of "λf. «f : F b →⇩C a» ∧ E.cones_map f χ = ?χ'" "ψ a g' = f"]
using 1 2 Adj.ε.is_extensional C.null_is_zero(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 →⇩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 ‹⋅⇩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 ‹⋅⇩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 ‹⋅⇩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_composition.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_composition.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_composition.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_composition.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_composition.arr J') = I'
⟹ ∃a. has_as_product J' D' a"
proof -
fix J' D'
assume 1: "discrete_diagram J' C D' ∧ Collect (partial_composition.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 π.is_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_composition.arr J) = {1}
⟶ (∃a. has_as_product J D a)"
proof
fix J :: "nat comp"
show "∀D. discrete_diagram J (⋅) D ∧ Collect (partial_composition.arr J) = {1}
⟶ (∃a. has_as_product J D a)"
proof (intro allI impI)
fix D
assume D: "discrete_diagram J (⋅) D ∧ Collect (partial_composition.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 ‹⋅⇩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' ⋅⇩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' is_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 is_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 is_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))"
lemma has_as_equalizerI [intro]:
assumes "par f g" and "seq f e" and "f ⋅ e = g ⋅ e"
and "⋀e'. ⟦seq f e'; f ⋅ e' = g ⋅ e'⟧ ⟹ ∃!h. e ⋅ h = e'"
shows "has_as_equalizer f g e"
proof (unfold has_as_equalizer_def, intro conjI)
show "arr f" and "arr g" and "dom f = dom g" and "cod f = cod g"
using assms(1) by auto
interpret J: parallel_pair .
interpret D: parallel_pair_diagram C f g
using assms(1) by unfold_locales
show "D.has_as_equalizer e"
proof -
let ?χ = "D.mkCone e"
let ?a = "dom e"
interpret χ: cone J.comp C D.map ?a ?χ
using assms(2-3) D.cone_mkCone [of e] by simp
interpret χ: limit_cone J.comp C D.map ?a ?χ
proof
fix a' χ'
assume χ': "D.cone a' χ'"
interpret χ': cone J.comp C D.map a' χ'
using χ' by blast
have "seq f (χ' J.Zero)"
using J.ide_char J.arr_char χ'.preserves_hom
by (metis (no_types, lifting) D.map_simp(3) χ'.is_natural_1
χ'.natural_transformation_axioms natural_transformation.preserves_reflects_arr
parallel_pair.dom_simp(3))
moreover have "f ⋅ (χ' J.Zero) = g ⋅ (χ' J.Zero)"
using χ' D.is_equalized_by_cone by blast
ultimately have 1: "∃!h. e ⋅ h = χ' J.Zero"
using assms by blast
obtain h where h: "e ⋅ h = χ' J.Zero"
using 1 by blast
have 2: "D.is_equalized_by e"
using assms(2-3) by blast
have "«h : a' → dom e» ∧ D.cones_map h (D.mkCone e) = χ'"
proof
show 3: "«h : a' → dom e»"
using h χ'.preserves_dom
by (metis (no_types, lifting) χ'.component_in_hom ‹seq f (χ' J.Zero)›
category.has_codomain_iff_arr category.seqE category_axioms cod_in_codomains
domains_comp in_hom_def parallel_pair.arr_char)
show "D.cones_map h (D.mkCone e) = χ'"
proof
fix j
have "D.cone (cod h) (D.mkCone e)"
using 2 3 D.cone_mkCone by auto
thus "D.cones_map h (D.mkCone e) j = χ' j"
using h 2 3 D.cone_mkCone [of e] J.arr_char D.mkCone_def comp_assoc
apply (cases "J.arr j")
apply simp_all
apply (metis (no_types, lifting) D.mkCone_cone χ')
using χ'.is_extensional
by presburger
qed
qed
hence "∃h. «h : a' → dom e» ∧ D.cones_map h (D.mkCone e) = χ'"
by blast
moreover have "⋀h'. «h' : a' → dom e» ∧ D.cones_map h' (D.mkCone e) = χ' ⟹ h' = h"
proof (elim conjE)
fix h'
assume h': "«h' : a' → dom e»"
assume eq: "D.cones_map h' (D.mkCone e) = χ'"
have "e ⋅ h' = χ' J.Zero"
using eq D.cone_mkCone D.mkCone_def χ'.preserves_reflects_arr χ.cone_axioms
‹seq f (χ' J.Zero)› eq h' in_homE mem_Collect_eq restrict_apply seqE
apply simp
by fastforce
moreover have "∃!h. e ⋅ h = χ' J.Zero"
using assms(2,4) 1 category.seqI by blast
ultimately show "h' = h"
using h by auto
qed
ultimately show "∃!h. «h : a' → dom e» ∧ D.cones_map h (D.mkCone e) = χ'"
by blast
qed
show "D.has_as_equalizer e"
using assms χ.limit_cone_axioms by blast
qed
qed
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: Π⇩i F⇩i → Π⇩u F⇩c⇩o⇩d ⇩u (u ∈ arr J, i ∈ J)›
where ‹p⇩u f = p⇩c⇩o⇩d ⇩u›, ‹p⇩u g = F⇩u o p⇩d⇩o⇩m ⇩u›;
the limiting cone ‹μ› is ‹μ⇩j = p⇩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 ‹⋅⇩J› 55)
assumes "category J" and "has_products (Collect (partial_composition.ide J))"
and "has_products (Collect (partial_composition.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.is_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.is_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 is_set +
diagram J S D
for J :: "'j comp" (infixr ‹⋅⇩J› 55)
and S :: "'s comp" (infixr ‹⋅› 55)
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'⇩S⇩C [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'⇩S⇩C [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'⇩S⇩C [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'⇩S⇩C [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 ‹⋅⇩J› 55)
and S :: "'s comp" (infixr ‹⋅› 55)
and D :: "'j ⇒ 's"
begin
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⇩S⇩C:
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 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: "setp (img ` {e. e ∈ hom unity (dom f0) ∧ f0 ⋅ e = f1 ⋅ e})"
proof -
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⇩S⇩C 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_set_category =
S: set_category S 𝔖 +
discrete_diagram J S D +
diagram_in_set_category J S 𝔖 D
for J :: "'j comp" (infixr ‹⋅⇩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 "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 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⇩S⇩C 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⇩S⇩C 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_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 metis
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 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 →⇩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_composition.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_composition.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_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 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)"
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 (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_composition.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_composition.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 ‹⋅⇩J› 55)
and A :: "'a comp" (infixr ‹⋅⇩A› 55)
and B :: "'b comp" (infixr ‹⋅⇩B› 55)
and D :: "'j * 'a ⇒ 'b"
begin
notation J.in_hom (‹«_ : _ →⇩J _»›)
notation JxA.comp (infixr ‹⋅⇩J⇩x⇩A› 55)
notation JxA.in_hom (‹«_ : _ →⇩J⇩x⇩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) →⇩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) →⇩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 →⇩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' ⋅⇩A a))› ‹l (A.dom (a' ⋅⇩A a))›
‹χ (A.dom (a' ⋅⇩A a))›
using a'a assms by auto
interpret χ_cod_a'a: limit_cone J B ‹λj. D (j, A.cod (a' ⋅⇩A a))› ‹l (A.cod (a' ⋅⇩A a))›
‹χ (A.cod (a' ⋅⇩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' ⋅⇩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' ⋅⇩A a))›
‹χ (A.cod a)› ‹λj. D (j, a')›..
interpret Da'oχ_cod_a: cone J B ‹λj. D (j, A.cod (a' ⋅⇩A a))› ‹l (A.cod a)› Da'oχ_cod_a.map
..
interpret Da'a: natural_transformation J B
‹λj. D (j, A.dom (a' ⋅⇩A a))› ‹λj. D (j, A.cod (a' ⋅⇩A a))›
‹λj. D (j, a' ⋅⇩A a)›
using a'a fixing_arr_gives_natural_transformation_2 [of "a' ⋅⇩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' ⋅⇩A a))›
‹λj. D (j, A.cod (a' ⋅⇩A a))› ‹χ (A.dom (a' ⋅⇩A a))›
‹λj. D (j, a' ⋅⇩A a)› ..
interpret Da'aoχ_dom_a'a: cone J B ‹λj. D (j, A.cod (a' ⋅⇩A a))›
‹l (A.dom (a' ⋅⇩A a))› Da'aoχ_dom_a'a.map ..
show "?L (a' ⋅⇩A a) = ?L a' ⋅⇩B ?L a"
proof -
have "?P (a' ⋅⇩A a) (?L (a' ⋅⇩A a))" using assms a'a L_arr [of l χ "a' ⋅⇩A a"] by fastforce
moreover have "?P (a' ⋅⇩A a) (?L a' ⋅⇩B ?L a)"
proof
have La: "«?L a : l (A.dom a) →⇩B l (A.cod a)»"
using assms a L_arr by fast
moreover have La': "«?L a' : l (A.cod a) →⇩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' ⋅⇩B ?L a : l (A.dom (a' ⋅⇩A a)) →⇩B l (A.cod (a' ⋅⇩A a))»"
using a a' 1 La La' by (intro B.comp_in_homI, auto)
show "χ_cod_a'a.D.cones_map (?L a' ⋅⇩B ?L a) (χ (A.cod (a' ⋅⇩A a)))
= Da'aoχ_dom_a'a.map"
proof -
have "χ_cod_a'a.D.cones_map (?L a' ⋅⇩B ?L a) (χ (A.cod (a' ⋅⇩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' ⋅⇩B ?L a) (χ (A.cod (a' ⋅⇩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' ⋅⇩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' ⋅⇩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' ⋅⇩A a))) (λj. D (j, a' ⋅⇩A a))"
using a'a preserves_comp_2 by simp
finally show ?thesis by auto
qed
qed
moreover have "∃!f. ?P (a' ⋅⇩A a) f"
using χ_cod_a'a.is_universal
[of "l (A.dom (a' ⋅⇩A a))"
"vertical_composite.map J B (χ (A.dom (a' ⋅⇩A a))) (λj. D (j, a' ⋅⇩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 ‹⋅⇩A› 55)
and B :: "'b comp" (infixr ‹⋅⇩B› 55)
and J :: "'j comp" (infixr ‹⋅⇩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 ‹⋅⇩J⇩x⇩A› 55)
notation JxA.in_hom (‹«_ : _ →⇩J⇩x⇩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' ⋅⇩A a) D›
using assms functor_at_arr_is_transformation [of D "a' ⋅⇩A a"] functor_axioms by simp
show "at (a' ⋅⇩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' ⋅⇩A a) D)" ..
show "⋀j. J.ide j ⟹ at (a' ⋅⇩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' ⋅⇩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 →⇩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' →⇩[⇩A⇩,⇩B⇩] x» ∧ cones_map f χ = χ'"
proof
let ?P = "λa fa. «fa : A_B.Map x' a →⇩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 →⇩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 →⇩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 →⇩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 →⇩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 →⇩B ?x_dom_a»" using a AaPa by simp
have f_cod_a: "«?f_cod_a : ?x'_cod_a →⇩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 ⋅⇩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 ⋅⇩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) ⋅⇩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) ⋅⇩B A_B.Map x a"
using a j at_simp by simp
also have "... = at (A.cod a) χ j ⋅⇩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 ⋅⇩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 ⋅⇩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) ⋅⇩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) ⋅⇩B A_B.Map x' a"
using a j at_simp by simp
also have "... = at (A.cod a) χ' j ⋅⇩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 ⋅⇩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 ⋅⇩B at (A.dom a) χ j) ⋅⇩B ?f_dom_a"
using j Daoχ_dom_a.map_simp_ide by simp
also have "... = at a D j ⋅⇩B at (A.dom a) χ j ⋅⇩B ?f_dom_a"
using B.comp_assoc by simp
also have "... = at a D j ⋅⇩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 ⋅⇩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 →⇩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 ⋅⇩B ?x'_a ∈ B.hom ?x'_dom_a ?x_cod_a ∧
D_cod_a.cones_map (?f_cod_a ⋅⇩B ?x'_a) (at (A.cod a) χ) = Daoχ'_dom_a.map"
proof
show "?f_cod_a ⋅⇩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 ⋅⇩B ?x'_a) (at (A.cod a) χ) = Daoχ'_dom_a.map"
proof -
have "D_cod_a.cones_map (?f_cod_a ⋅⇩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 ⋅⇩B ?f_dom_a ∈ B.hom ?x'_dom_a ?x_cod_a ∧
D_cod_a.cones_map (?x_a ⋅⇩B ?f_dom_a) (at (A.cod a) χ) = Daoχ'_dom_a.map"
proof
show "?x_a ⋅⇩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 ⋅⇩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 ⋅⇩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 ⋅⇩B ?x'_a = ?x_a ⋅⇩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' →⇩[⇩A⇩,⇩B⇩] x» ∧ cones_map ?f χ = χ'" using f_in_hom cones_map_f by auto
show "⋀f'. «f' : x' →⇩[⇩A⇩,⇩B⇩] x» ∧ cones_map f' χ = χ' ⟹ f' = ?f"
proof -
fix f'
assume f': "«f' : x' →⇩[⇩A⇩,⇩B⇩] 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 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 →⇩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) ⋅⇩B ?L' ja = B (Curry.uncurry D ja) (?χ' (JxA.dom ja))"
proof -
have "?χ' (JxA.cod ja) ⋅⇩B ?L' ja = ?χ (A.cod ?a) (J.cod ?j) ⋅⇩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) ⋅⇩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) ⋅⇩B D.at (A.dom ?a) D ?j ⋅⇩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) ⋅⇩B D.at (A.dom ?a) D ?j) ⋅⇩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 ⋅⇩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
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 →⇩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' ⋅⇩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 →⇩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
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 ⋅⇩S ?f"
using f j Fun_map_a_a' Yoχ_a'.cone_axioms by fastforce
also have "... = σ j"
proof (intro S.arr_eqI⇩S⇩C)
show "S.par (YoD.at a' (map o χ) j ⋅⇩S ?f) (σ j)"
using 1 f j x YoD_a'.preserves_hom by fastforce
show "S.Fun (YoD.at a' (map o χ) j ⋅⇩S ?f) = S.Fun (σ j)"
proof
fix e
have "e ∉ S.set x ⟹ S.Fun (YoD.at a' (map o χ) j ⋅⇩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 ⋅⇩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 ⋅⇩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
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 ⋅⇩S ?f) e = S.Fun (σ j) e"
by auto
qed
ultimately show "S.Fun (YoD.at a' (map o χ) j ⋅⇩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 →⇩S Hom.map (a', a)» ∧
YoD_a'.cones_map f' (YoD.at a' (map o χ)) = σ
⟹ f' = ?f"
proof -
fix f'
assume f': "«f' : x →⇩S Hom.map (a', a)» ∧
YoD_a'.cones_map f' (YoD.at a' (map o χ)) = σ"
show "f' = ?f"
proof (intro S.arr_eqI⇩S⇩C)
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 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 ⋅⇩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 ⋅⇩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 →⇩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 →⇩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