begin
theorem
theorem
theorem
theorem
theorem
theorem
begin
:: deftheorem Def1 defines cc= ALTCAT_2:def 1 :
for f, g being Function holds
( f cc= g iff ( dom f c= dom g & ( for i being set st i in dom f holds
f . i c= g . i ) ) );
:: deftheorem Def2 defines cc= ALTCAT_2:def 2 :
for I, J being set
for A being ManySortedSet of I
for B being ManySortedSet of J holds
( A cc= B iff ( I c= J & ( for i being set st i in I holds
A . i c= B . i ) ) );
theorem
canceled;
theorem Th8:
theorem Th9:
theorem
begin
begin
theorem Th11:
theorem Th12:
definition
let C be non
empty non
void CatStr ;
func the_hom_sets_of C -> ManySortedSet of
[: the carrier of C, the carrier of C:] means :
Def3:
for
i,
j being
Object of
C holds
it . (
i,
j)
= Hom (
i,
j);
existence
ex b1 being ManySortedSet of [: the carrier of C, the carrier of C:] st
for i, j being Object of C holds b1 . (i,j) = Hom (i,j)
uniqueness
for b1, b2 being ManySortedSet of [: the carrier of C, the carrier of C:] st ( for i, j being Object of C holds b1 . (i,j) = Hom (i,j) ) & ( for i, j being Object of C holds b2 . (i,j) = Hom (i,j) ) holds
b1 = b2
end;
:: deftheorem Def3 defines the_hom_sets_of ALTCAT_2:def 3 :
for C being non empty non void CatStr
for b2 being ManySortedSet of [: the carrier of C, the carrier of C:] holds
( b2 = the_hom_sets_of C iff for i, j being Object of C holds b2 . (i,j) = Hom (i,j) );
theorem Th13:
definition
let C be
Category;
func the_comps_of C -> BinComp of
(the_hom_sets_of C) means :
Def4:
for
i,
j,
k being
Object of
C holds
it . (
i,
j,
k)
= the
Comp of
C | [:((the_hom_sets_of C) . (j,k)),((the_hom_sets_of C) . (i,j)):];
existence
ex b1 being BinComp of (the_hom_sets_of C) st
for i, j, k being Object of C holds b1 . (i,j,k) = the Comp of C | [:((the_hom_sets_of C) . (j,k)),((the_hom_sets_of C) . (i,j)):]
uniqueness
for b1, b2 being BinComp of (the_hom_sets_of C) st ( for i, j, k being Object of C holds b1 . (i,j,k) = the Comp of C | [:((the_hom_sets_of C) . (j,k)),((the_hom_sets_of C) . (i,j)):] ) & ( for i, j, k being Object of C holds b2 . (i,j,k) = the Comp of C | [:((the_hom_sets_of C) . (j,k)),((the_hom_sets_of C) . (i,j)):] ) holds
b1 = b2
end;
:: deftheorem Def4 defines the_comps_of ALTCAT_2:def 4 :
for C being Category
for b2 being BinComp of (the_hom_sets_of C) holds
( b2 = the_comps_of C iff for i, j, k being Object of C holds b2 . (i,j,k) = the Comp of C | [:((the_hom_sets_of C) . (j,k)),((the_hom_sets_of C) . (i,j)):] );
theorem Th14:
theorem Th15:
theorem Th16:
begin
:: deftheorem defines Alter ALTCAT_2:def 5 :
for C being Category holds Alter C = AltCatStr(# the carrier of C,(the_hom_sets_of C),(the_comps_of C) #);
theorem Th17:
theorem Th18:
theorem Th19:
begin
:: deftheorem Def6 defines reflexive ALTCAT_2:def 6 :
for C being AltGraph holds
( C is reflexive iff for x being set st x in the carrier of C holds
the Arrows of C . (x,x) <> {} );
:: deftheorem defines reflexive ALTCAT_2:def 7 :
for C being non empty AltGraph holds
( C is reflexive iff for o being object of C holds <^o,o^> <> {} );
definition
let C be non
empty transitive AltCatStr ;
redefine attr C is
associative means :
Def8:
for
o1,
o2,
o3,
o4 being
object of
C for
f being
Morphism of
o1,
o2 for
g being
Morphism of
o2,
o3 for
h being
Morphism of
o3,
o4 st
<^o1,o2^> <> {} &
<^o2,o3^> <> {} &
<^o3,o4^> <> {} holds
(h * g) * f = h * (g * f);
compatibility
( C is associative iff for o1, o2, o3, o4 being object of C
for f being Morphism of o1,o2
for g being Morphism of o2,o3
for h being Morphism of o3,o4 st <^o1,o2^> <> {} & <^o2,o3^> <> {} & <^o3,o4^> <> {} holds
(h * g) * f = h * (g * f) )
end;
:: deftheorem Def8 defines associative ALTCAT_2:def 8 :
for C being non empty transitive AltCatStr holds
( C is associative iff for o1, o2, o3, o4 being object of C
for f being Morphism of o1,o2
for g being Morphism of o2,o3
for h being Morphism of o3,o4 st <^o1,o2^> <> {} & <^o2,o3^> <> {} & <^o3,o4^> <> {} holds
(h * g) * f = h * (g * f) );
definition
let C be non
empty AltCatStr ;
redefine attr C is
with_units means
for
o being
object of
C holds
(
<^o,o^> <> {} & ex
i being
Morphism of
o,
o st
for
o9 being
object of
C for
m9 being
Morphism of
o9,
o for
m99 being
Morphism of
o,
o9 holds
( (
<^o9,o^> <> {} implies
i * m9 = m9 ) & (
<^o,o9^> <> {} implies
m99 * i = m99 ) ) );
compatibility
( C is with_units iff for o being object of C holds
( <^o,o^> <> {} & ex i being Morphism of o,o st
for o9 being object of C
for m9 being Morphism of o9,o
for m99 being Morphism of o,o9 holds
( ( <^o9,o^> <> {} implies i * m9 = m9 ) & ( <^o,o9^> <> {} implies m99 * i = m99 ) ) ) )
end;
:: deftheorem defines with_units ALTCAT_2:def 9 :
for C being non empty AltCatStr holds
( C is with_units iff for o being object of C holds
( <^o,o^> <> {} & ex i being Morphism of o,o st
for o9 being object of C
for m9 being Morphism of o9,o
for m99 being Morphism of o,o9 holds
( ( <^o9,o^> <> {} implies i * m9 = m9 ) & ( <^o,o9^> <> {} implies m99 * i = m99 ) ) ) );
begin
Lm1:
for E1, E2 being strict AltCatStr st the carrier of E1 is empty & the carrier of E2 is empty holds
E1 = E2
:: deftheorem Def10 defines the_empty_category ALTCAT_2:def 10 :
for b1 being strict AltCatStr holds
( b1 = the_empty_category iff the carrier of b1 is empty );
theorem
begin
:: deftheorem Def11 defines SubCatStr ALTCAT_2:def 11 :
for C, b2 being AltCatStr holds
( b2 is SubCatStr of C iff ( the carrier of b2 c= the carrier of C & the Arrows of b2 cc= the Arrows of C & the Comp of b2 cc= the Comp of C ) );
theorem Th21:
theorem
theorem Th23:
definition
let C be non
empty AltCatStr ;
let o be
object of
C;
func ObCat o -> strict SubCatStr of
C means :
Def12:
( the
carrier of
it = {o} & the
Arrows of
it = (
o,
o)
:-> <^o,o^> & the
Comp of
it = [o,o,o] .--> ( the Comp of C . (o,o,o)) );
existence
ex b1 being strict SubCatStr of C st
( the carrier of b1 = {o} & the Arrows of b1 = (o,o) :-> <^o,o^> & the Comp of b1 = [o,o,o] .--> ( the Comp of C . (o,o,o)) )
uniqueness
for b1, b2 being strict SubCatStr of C st the carrier of b1 = {o} & the Arrows of b1 = (o,o) :-> <^o,o^> & the Comp of b1 = [o,o,o] .--> ( the Comp of C . (o,o,o)) & the carrier of b2 = {o} & the Arrows of b2 = (o,o) :-> <^o,o^> & the Comp of b2 = [o,o,o] .--> ( the Comp of C . (o,o,o)) holds
b1 = b2
;
end;
:: deftheorem Def12 defines ObCat ALTCAT_2:def 12 :
for C being non empty AltCatStr
for o being object of C
for b3 being strict SubCatStr of C holds
( b3 = ObCat o iff ( the carrier of b3 = {o} & the Arrows of b3 = (o,o) :-> <^o,o^> & the Comp of b3 = [o,o,o] .--> ( the Comp of C . (o,o,o)) ) );
theorem Th24:
theorem Th25:
:: deftheorem Def13 defines full ALTCAT_2:def 13 :
for C being AltCatStr
for D being SubCatStr of C holds
( D is full iff the Arrows of D = the Arrows of C || the carrier of D );
:: deftheorem Def14 defines id-inheriting ALTCAT_2:def 14 :
for C being non empty with_units AltCatStr
for D being SubCatStr of C holds
( ( not D is empty implies ( D is id-inheriting iff for o being object of D
for o9 being object of C st o = o9 holds
idm o9 in <^o,o^> ) ) & ( D is empty implies ( D is id-inheriting iff verum ) ) );
theorem Th26:
theorem Th27:
theorem
theorem Th29:
theorem Th30:
theorem
theorem Th32:
theorem Th33:
for
C being non
empty transitive AltCatStr for
D being non
empty transitive SubCatStr of
C for
p1,
p2,
p3 being
object of
D st
<^p1,p2^> <> {} &
<^p2,p3^> <> {} holds
for
o1,
o2,
o3 being
object of
C st
o1 = p1 &
o2 = p2 &
o3 = p3 holds
for
f being
Morphism of
o1,
o2 for
g being
Morphism of
o2,
o3 for
ff being
Morphism of
p1,
p2 for
gg being
Morphism of
p2,
p3 st
f = ff &
g = gg holds
g * f = gg * ff
theorem Th34:
theorem