deffunc H1( CatStr ) -> set = the carrier of $1;
deffunc H2( CatStr ) -> set = the carrier' of $1;
definition
let C,
D,
E be
Category;
let F be
Functor of
C,
E;
let G be
Functor of
D,
E;
given c1 being
Object of
C,
d1 being
Object of
D,
f1 being
Morphism of
E such that A1:
f1 in Hom (
(F . c1),
(G . d1))
;
coherence
{ [[c,d],f] where c is Object of C, d is Object of D, f is Morphism of E : f in Hom ((F . c),(G . d)) } is non empty Subset of [:[: the carrier of C, the carrier of D:], the carrier' of E:]
end;
::
deftheorem Def1 defines
commaObjs COMMACAT:def 1 :
for C, D, E being Category
for F being Functor of C,E
for G being Functor of D,E st ex c1 being Object of C ex d1 being Object of D ex f1 being Morphism of E st f1 in Hom ((F . c1),(G . d1)) holds
commaObjs (F,G) = { [[c,d],f] where c is Object of C, d is Object of D, f is Morphism of E : f in Hom ((F . c),(G . d)) } ;
definition
let C,
D,
E be
Category;
let F be
Functor of
C,
E;
let G be
Functor of
D,
E;
given c1 being
Object of
C,
d1 being
Object of
D,
f1 being
Morphism of
E such that A1:
f1 in Hom (
(F . c1),
(G . d1))
;
func commaMorphs (
F,
G)
-> non
empty Subset of
[:[:(commaObjs (F,G)),(commaObjs (F,G)):],[: the carrier' of C, the carrier' of D:]:] equals :
Def2:
{ [[o1,o2],[g,h]] where g is Morphism of C, h is Morphism of D, o1, o2 is Element of commaObjs (F,G) : ( dom g = o1 `11 & cod g = o2 `11 & dom h = o1 `12 & cod h = o2 `12 & (o2 `2) (*) (F . g) = (G . h) (*) (o1 `2) ) } ;
coherence
{ [[o1,o2],[g,h]] where g is Morphism of C, h is Morphism of D, o1, o2 is Element of commaObjs (F,G) : ( dom g = o1 `11 & cod g = o2 `11 & dom h = o1 `12 & cod h = o2 `12 & (o2 `2) (*) (F . g) = (G . h) (*) (o1 `2) ) } is non empty Subset of [:[:(commaObjs (F,G)),(commaObjs (F,G)):],[: the carrier' of C, the carrier' of D:]:]
end;
::
deftheorem Def2 defines
commaMorphs COMMACAT:def 2 :
for C, D, E being Category
for F being Functor of C,E
for G being Functor of D,E st ex c1 being Object of C ex d1 being Object of D ex f1 being Morphism of E st f1 in Hom ((F . c1),(G . d1)) holds
commaMorphs (F,G) = { [[o1,o2],[g,h]] where g is Morphism of C, h is Morphism of D, o1, o2 is Element of commaObjs (F,G) : ( dom g = o1 `11 & cod g = o2 `11 & dom h = o1 `12 & cod h = o2 `12 & (o2 `2) (*) (F . g) = (G . h) (*) (o1 `2) ) } ;
definition
let C,
D,
E be
Category;
let F be
Functor of
C,
E;
let G be
Functor of
D,
E;
let k1,
k2 be
Element of
commaMorphs (
F,
G);
given c1 being
Object of
C,
d1 being
Object of
D,
f1 being
Morphism of
E such that A1:
f1 in Hom (
(F . c1),
(G . d1))
;
assume A2:
k1 `12 = k2 `11
;
coherence
[[(k1 `11),(k2 `12)],[((k2 `21) (*) (k1 `21)),((k2 `22) (*) (k1 `22))]] is Element of commaMorphs (F,G)
end;
::
deftheorem Def3 defines
* COMMACAT:def 3 :
for C, D, E being Category
for F being Functor of C,E
for G being Functor of D,E
for k1, k2 being Element of commaMorphs (F,G) st ex c1 being Object of C ex d1 being Object of D ex f1 being Morphism of E st f1 in Hom ((F . c1),(G . d1)) & k1 `12 = k2 `11 holds
k2 * k1 = [[(k1 `11),(k2 `12)],[((k2 `21) (*) (k1 `21)),((k2 `22) (*) (k1 `22))]];
definition
let C,
D,
E be
Category;
let F be
Functor of
C,
E;
let G be
Functor of
D,
E;
func commaComp (
F,
G)
-> PartFunc of
[:(commaMorphs (F,G)),(commaMorphs (F,G)):],
(commaMorphs (F,G)) means :
Def4:
(
dom it = { [k1,k2] where k1, k2 is Element of commaMorphs (F,G) : k1 `11 = k2 `12 } & ( for
k,
k9 being
Element of
commaMorphs (
F,
G) st
[k,k9] in dom it holds
it . [k,k9] = k * k9 ) );
existence
ex b1 being PartFunc of [:(commaMorphs (F,G)),(commaMorphs (F,G)):],(commaMorphs (F,G)) st
( dom b1 = { [k1,k2] where k1, k2 is Element of commaMorphs (F,G) : k1 `11 = k2 `12 } & ( for k, k9 being Element of commaMorphs (F,G) st [k,k9] in dom b1 holds
b1 . [k,k9] = k * k9 ) )
uniqueness
for b1, b2 being PartFunc of [:(commaMorphs (F,G)),(commaMorphs (F,G)):],(commaMorphs (F,G)) st dom b1 = { [k1,k2] where k1, k2 is Element of commaMorphs (F,G) : k1 `11 = k2 `12 } & ( for k, k9 being Element of commaMorphs (F,G) st [k,k9] in dom b1 holds
b1 . [k,k9] = k * k9 ) & dom b2 = { [k1,k2] where k1, k2 is Element of commaMorphs (F,G) : k1 `11 = k2 `12 } & ( for k, k9 being Element of commaMorphs (F,G) st [k,k9] in dom b2 holds
b2 . [k,k9] = k * k9 ) holds
b1 = b2
end;
::
deftheorem Def4 defines
commaComp COMMACAT:def 4 :
for C, D, E being Category
for F being Functor of C,E
for G being Functor of D,E
for b6 being PartFunc of [:(commaMorphs (F,G)),(commaMorphs (F,G)):],(commaMorphs (F,G)) holds
( b6 = commaComp (F,G) iff ( dom b6 = { [k1,k2] where k1, k2 is Element of commaMorphs (F,G) : k1 `11 = k2 `12 } & ( for k, k9 being Element of commaMorphs (F,G) st [k,k9] in dom b6 holds
b6 . [k,k9] = k * k9 ) ) );
definition
let C,
D,
E be
Category;
let F be
Functor of
C,
E;
let G be
Functor of
D,
E;
given c1 being
Object of
C,
d1 being
Object of
D,
f1 being
Morphism of
E such that A1:
f1 in Hom (
(F . c1),
(G . d1))
;
existence
ex b1 being strict Category st
( the carrier of b1 = commaObjs (F,G) & the carrier' of b1 = commaMorphs (F,G) & ( for k being Element of commaMorphs (F,G) holds the Source of b1 . k = k `11 ) & ( for k being Element of commaMorphs (F,G) holds the Target of b1 . k = k `12 ) & the Comp of b1 = commaComp (F,G) )
uniqueness
for b1, b2 being strict Category st the carrier of b1 = commaObjs (F,G) & the carrier' of b1 = commaMorphs (F,G) & ( for k being Element of commaMorphs (F,G) holds the Source of b1 . k = k `11 ) & ( for k being Element of commaMorphs (F,G) holds the Target of b1 . k = k `12 ) & the Comp of b1 = commaComp (F,G) & the carrier of b2 = commaObjs (F,G) & the carrier' of b2 = commaMorphs (F,G) & ( for k being Element of commaMorphs (F,G) holds the Source of b2 . k = k `11 ) & ( for k being Element of commaMorphs (F,G) holds the Target of b2 . k = k `12 ) & the Comp of b2 = commaComp (F,G) holds
b1 = b2
end;
::
deftheorem defines
comma COMMACAT:def 5 :
for C, D, E being Category
for F being Functor of C,E
for G being Functor of D,E st ex c1 being Object of C ex d1 being Object of D ex f1 being Morphism of E st f1 in Hom ((F . c1),(G . d1)) holds
for b6 being strict Category holds
( b6 = F comma G iff ( the carrier of b6 = commaObjs (F,G) & the carrier' of b6 = commaMorphs (F,G) & ( for k being Element of commaMorphs (F,G) holds the Source of b6 . k = k `11 ) & ( for k being Element of commaMorphs (F,G) holds the Target of b6 . k = k `12 ) & the Comp of b6 = commaComp (F,G) ) );