begin
theorem Th1:
theorem
theorem Th3:
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem Th7:
theorem Th8:
theorem Th9:
:: deftheorem defines Functor ISOCAT_1:def 1 :
for C, D being Category
for b3 being M2( bool [: the carrier' of b1, the carrier' of b2:]) holds
( b3 is Functor of C,D iff ( ( for c being Object of C ex d being Object of D st b3 . (id c) = id d ) & ( for f being Morphism of C holds
( b3 . (id (dom f)) = id (dom (b3 . f)) & b3 . (id (cod f)) = id (cod (b3 . f)) ) ) & ( for f, g being Morphism of C st dom g = cod f holds
b3 . (g * f) = (b3 . g) * (b3 . f) ) ) );
theorem Th10:
theorem Th11:
theorem Th12:
:: deftheorem Def2 defines " ISOCAT_1:def 2 :
for A, B being Category
for F being Functor of A,B st F is isomorphic holds
F " = F " ;
:: deftheorem Def3 defines isomorphic ISOCAT_1:def 3 :
for A, B being Category
for F being Functor of A,B holds
( F is isomorphic iff ( F is one-to-one & rng F = the carrier' of B ) );
theorem Th13:
theorem
theorem
theorem Th16:
theorem Th17:
:: deftheorem defines are_isomorphic ISOCAT_1:def 4 :
for A, B being Category holds
( A,B are_isomorphic iff ex F being Functor of A,B st F is isomorphic );
theorem
canceled;
theorem
canceled;
theorem
theorem
theorem
theorem
theorem
:: deftheorem Def5 defines * ISOCAT_1:def 5 :
for A, B, C being Category
for F1, F2 being Functor of A,B st F1 is_transformable_to F2 holds
for t being transformation of F1,F2
for G being Functor of B,C holds G * t = G * t;
:: deftheorem Def6 defines * ISOCAT_1:def 6 :
for A, B, C being Category
for G1, G2 being Functor of B,C st G1 is_transformable_to G2 holds
for F being Functor of A,B
for t being transformation of G1,G2 holds t * F = t * (Obj F);
theorem Th25:
theorem Th26:
theorem Th27:
:: deftheorem Def7 defines * ISOCAT_1:def 7 :
for A, B, C being Category
for F1, F2 being Functor of A,B st F1 is_naturally_transformable_to F2 holds
for t being natural_transformation of F1,F2
for G being Functor of B,C holds G * t = G * t;
theorem Th28:
:: deftheorem Def8 defines * ISOCAT_1:def 8 :
for A, B, C being Category
for G1, G2 being Functor of B,C st G1 is_naturally_transformable_to G2 holds
for F being Functor of A,B
for t being natural_transformation of G1,G2 holds t * F = t * F;
theorem Th29:
theorem Th30:
theorem Th31:
theorem Th32:
theorem Th33:
theorem Th34:
theorem Th35:
theorem Th36:
theorem Th37:
theorem Th38:
theorem Th39:
theorem Th40:
definition
let A,
B,
C be
Category;
let F1,
F2 be
Functor of
A,
B;
let G1,
G2 be
Functor of
B,
C;
let s be
natural_transformation of
F1,
F2;
let t be
natural_transformation of
G1,
G2;
func t (#) s -> natural_transformation of
G1 * F1,
G2 * F2 equals
(t * F2) `*` (G1 * s);
correctness
coherence
(t * F2) `*` (G1 * s) is natural_transformation of G1 * F1,G2 * F2;
;
end;
:: deftheorem defines (#) ISOCAT_1:def 9 :
for A, B, C being Category
for F1, F2 being Functor of A,B
for G1, G2 being Functor of B,C
for s being natural_transformation of F1,F2
for t being natural_transformation of G1,G2 holds t (#) s = (t * F2) `*` (G1 * s);
theorem Th41:
theorem
theorem
theorem
for
A,
B,
C,
D being
Category for
F1,
F2 being
Functor of
A,
B for
G1,
G2 being
Functor of
B,
C for
H1,
H2 being
Functor of
C,
D for
s being
natural_transformation of
F1,
F2 for
t being
natural_transformation of
G1,
G2 for
u being
natural_transformation of
H1,
H2 st
F1 is_naturally_transformable_to F2 &
G1 is_naturally_transformable_to G2 &
H1 is_naturally_transformable_to H2 holds
u (#) (t (#) s) = (u (#) t) (#) s
theorem
theorem
theorem
for
A,
B,
C being
Category for
F1,
F2,
F3 being
Functor of
A,
B for
G1,
G2,
G3 being
Functor of
B,
C for
s being
natural_transformation of
F1,
F2 for
s9 being
natural_transformation of
F2,
F3 for
t being
natural_transformation of
G1,
G2 for
t9 being
natural_transformation of
G2,
G3 st
F1 is_naturally_transformable_to F2 &
F2 is_naturally_transformable_to F3 &
G1 is_naturally_transformable_to G2 &
G2 is_naturally_transformable_to G3 holds
(t9 `*` t) (#) (s9 `*` s) = (t9 (#) s9) `*` (t (#) s)
theorem Th48:
theorem Th49:
definition
let A,
B be
Category;
pred A is_equivalent_with B means :
Def10:
ex
F being
Functor of
A,
B ex
G being
Functor of
B,
A st
(
G * F ~= id A &
F * G ~= id B );
reflexivity
for A being Category ex F, G being Functor of A,A st
( G * F ~= id A & F * G ~= id A )
symmetry
for A, B being Category st ex F being Functor of A,B ex G being Functor of B,A st
( G * F ~= id A & F * G ~= id B ) holds
ex F being Functor of B,A ex G being Functor of A,B st
( G * F ~= id B & F * G ~= id A )
;
end;
:: deftheorem Def10 defines is_equivalent_with ISOCAT_1:def 10 :
for A, B being Category holds
( A is_equivalent_with B iff ex F being Functor of A,B ex G being Functor of B,A st
( G * F ~= id A & F * G ~= id B ) );
theorem
theorem
canceled;
theorem
canceled;
theorem Th53:
:: deftheorem Def11 defines Equivalence ISOCAT_1:def 11 :
for A, B being Category st A,B are_equivalent holds
for b3 being Functor of A,B holds
( b3 is Equivalence of A,B iff ex G being Functor of B,A st
( G * b3 ~= id A & b3 * G ~= id B ) );
theorem
theorem
theorem Th56:
theorem Th57:
theorem