begin
theorem Th1:
theorem Th2:
theorem Th3:
theorem
canceled;
theorem Th5:
theorem Th6:
theorem Th7:
theorem Th8:
theorem Th9:
begin
definition
let A,
B be
Category;
let a be
Object of
A;
func a |-> B -> Functor of
Functors (
A,
B),
B means :
Def1:
for
F1,
F2 being
Functor of
A,
B for
t being
natural_transformation of
F1,
F2 st
F1 is_naturally_transformable_to F2 holds
it . [[F1,F2],t] = t . a;
existence
ex b1 being Functor of Functors (A,B),B st
for F1, F2 being Functor of A,B
for t being natural_transformation of F1,F2 st F1 is_naturally_transformable_to F2 holds
b1 . [[F1,F2],t] = t . a
uniqueness
for b1, b2 being Functor of Functors (A,B),B st ( for F1, F2 being Functor of A,B
for t being natural_transformation of F1,F2 st F1 is_naturally_transformable_to F2 holds
b1 . [[F1,F2],t] = t . a ) & ( for F1, F2 being Functor of A,B
for t being natural_transformation of F1,F2 st F1 is_naturally_transformable_to F2 holds
b2 . [[F1,F2],t] = t . a ) holds
b1 = b2
end;
:: deftheorem Def1 defines |-> ISOCAT_2:def 1 :
for A, B being Category
for a being Object of A
for b4 being Functor of Functors (A,B),B holds
( b4 = a |-> B iff for F1, F2 being Functor of A,B
for t being natural_transformation of F1,F2 st F1 is_naturally_transformable_to F2 holds
b4 . [[F1,F2],t] = t . a );
theorem
canceled;
theorem
begin
theorem Th12:
theorem Th13:
theorem Th14:
for
A,
B being
Category for
a1,
a2 being
Object of
A for
b1,
b2 being
Object of
B st
Hom (
[a1,b1],
[a2,b2])
<> {} holds
for
f being
Morphism of
A for
g being
Morphism of
B holds
(
[f,g] is
Morphism of
[a1,b1],
[a2,b2] iff (
f is
Morphism of
a1,
a2 &
g is
Morphism of
b1,
b2 ) )
theorem Th15:
:: deftheorem defines curry ISOCAT_2:def 2 :
for A, B, C being Category
for F being Functor of [:A,B:],C
for f being Morphism of A holds curry (F,f) = (curry F) . f;
theorem Th16:
theorem Th17:
:: deftheorem defines ?- ISOCAT_2:def 3 :
for A, B, C being Category
for F being Functor of [:A,B:],C
for f being Morphism of A holds F ?- f = (curry (F,f)) * the Id of B;
theorem Th18:
theorem Th19:
theorem Th20:
theorem Th21:
definition
let A,
B,
C be
Category;
let F be
Functor of
[:A,B:],
C;
func export F -> Functor of
A,
Functors (
B,
C)
means :
Def4:
for
f being
Morphism of
A holds
it . f = [[(F ?- (dom f)),(F ?- (cod f))],(F ?- f)];
existence
ex b1 being Functor of A, Functors (B,C) st
for f being Morphism of A holds b1 . f = [[(F ?- (dom f)),(F ?- (cod f))],(F ?- f)]
uniqueness
for b1, b2 being Functor of A, Functors (B,C) st ( for f being Morphism of A holds b1 . f = [[(F ?- (dom f)),(F ?- (cod f))],(F ?- f)] ) & ( for f being Morphism of A holds b2 . f = [[(F ?- (dom f)),(F ?- (cod f))],(F ?- f)] ) holds
b1 = b2
end;
:: deftheorem Def4 defines export ISOCAT_2:def 4 :
for A, B, C being Category
for F being Functor of [:A,B:],C
for b5 being Functor of A, Functors (B,C) holds
( b5 = export F iff for f being Morphism of A holds b5 . f = [[(F ?- (dom f)),(F ?- (cod f))],(F ?- f)] );
Lm1:
for A, B 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 a being Object of A holds t . a in Hom ((F1 . a),(F2 . a))
theorem
canceled;
theorem
canceled;
theorem Th24:
theorem Th25:
theorem Th26:
theorem Th27:
for
A,
B,
C being
Category for
F1,
F2 being
Functor of
[:A,B:],
C st
F1 is_naturally_transformable_to F2 holds
for
t being
natural_transformation of
F1,
F2 holds
(
export F1 is_naturally_transformable_to export F2 & ex
G being
natural_transformation of
export F1,
export F2 st
for
s being
Function of
[: the carrier of A, the carrier of B:], the
carrier' of
C st
t = s holds
for
a being
Object of
A holds
G . a = [[((export F1) . a),((export F2) . a)],((curry s) . a)] )
definition
let A,
B,
C be
Category;
let F1,
F2 be
Functor of
[:A,B:],
C;
assume A1:
F1 is_naturally_transformable_to F2
;
let t be
natural_transformation of
F1,
F2;
func export t -> natural_transformation of
export F1,
export F2 means :
Def5:
for
s being
Function of
[: the carrier of A, the carrier of B:], the
carrier' of
C st
t = s holds
for
a being
Object of
A holds
it . a = [[((export F1) . a),((export F2) . a)],((curry s) . a)];
existence
ex b1 being natural_transformation of export F1, export F2 st
for s being Function of [: the carrier of A, the carrier of B:], the carrier' of C st t = s holds
for a being Object of A holds b1 . a = [[((export F1) . a),((export F2) . a)],((curry s) . a)]
by A1, Th27;
uniqueness
for b1, b2 being natural_transformation of export F1, export F2 st ( for s being Function of [: the carrier of A, the carrier of B:], the carrier' of C st t = s holds
for a being Object of A holds b1 . a = [[((export F1) . a),((export F2) . a)],((curry s) . a)] ) & ( for s being Function of [: the carrier of A, the carrier of B:], the carrier' of C st t = s holds
for a being Object of A holds b2 . a = [[((export F1) . a),((export F2) . a)],((curry s) . a)] ) holds
b1 = b2
end;
:: deftheorem Def5 defines export ISOCAT_2:def 5 :
for A, B, C being Category
for F1, F2 being Functor of [:A,B:],C st F1 is_naturally_transformable_to F2 holds
for t being natural_transformation of F1,F2
for b7 being natural_transformation of export F1, export F2 holds
( b7 = export t iff for s being Function of [: the carrier of A, the carrier of B:], the carrier' of C st t = s holds
for a being Object of A holds b7 . a = [[((export F1) . a),((export F2) . a)],((curry s) . a)] );
theorem Th28:
theorem Th29:
theorem Th30:
theorem Th31:
theorem Th32:
definition
let A,
B,
C be
Category;
func export (
A,
B,
C)
-> Functor of
Functors (
[:A,B:],
C),
Functors (
A,
(Functors (B,C)))
means :
Def6:
for
F1,
F2 being
Functor of
[:A,B:],
C st
F1 is_naturally_transformable_to F2 holds
for
t being
natural_transformation of
F1,
F2 holds
it . [[F1,F2],t] = [[(export F1),(export F2)],(export t)];
existence
ex b1 being Functor of Functors ([:A,B:],C), Functors (A,(Functors (B,C))) st
for F1, F2 being Functor of [:A,B:],C st F1 is_naturally_transformable_to F2 holds
for t being natural_transformation of F1,F2 holds b1 . [[F1,F2],t] = [[(export F1),(export F2)],(export t)]
uniqueness
for b1, b2 being Functor of Functors ([:A,B:],C), Functors (A,(Functors (B,C))) st ( for F1, F2 being Functor of [:A,B:],C st F1 is_naturally_transformable_to F2 holds
for t being natural_transformation of F1,F2 holds b1 . [[F1,F2],t] = [[(export F1),(export F2)],(export t)] ) & ( for F1, F2 being Functor of [:A,B:],C st F1 is_naturally_transformable_to F2 holds
for t being natural_transformation of F1,F2 holds b2 . [[F1,F2],t] = [[(export F1),(export F2)],(export t)] ) holds
b1 = b2
end;
:: deftheorem Def6 defines export ISOCAT_2:def 6 :
for A, B, C being Category
for b4 being Functor of Functors ([:A,B:],C), Functors (A,(Functors (B,C))) holds
( b4 = export (A,B,C) iff for F1, F2 being Functor of [:A,B:],C st F1 is_naturally_transformable_to F2 holds
for t being natural_transformation of F1,F2 holds b4 . [[F1,F2],t] = [[(export F1),(export F2)],(export t)] );
theorem
canceled;
theorem
begin
theorem Th35:
definition
let A,
B,
C be
Category;
let F be
Functor of
A,
B;
let G be
Functor of
A,
C;
<:redefine func <:F,G:> -> Functor of
A,
[:B,C:];
coherence
<:F,G:> is Functor of A,[:B,C:]
end;
definition
let A,
B,
C be
Category;
let F be
Functor of
A,
[:B,C:];
func Pr1 F -> Functor of
A,
B equals
(pr1 (B,C)) * F;
correctness
coherence
(pr1 (B,C)) * F is Functor of A,B;
;
func Pr2 F -> Functor of
A,
C equals
(pr2 (B,C)) * F;
correctness
coherence
(pr2 (B,C)) * F is Functor of A,C;
;
end;
:: deftheorem defines Pr1 ISOCAT_2:def 7 :
for A, B, C being Category
for F being Functor of A,[:B,C:] holds Pr1 F = (pr1 (B,C)) * F;
:: deftheorem defines Pr2 ISOCAT_2:def 8 :
for A, B, C being Category
for F being Functor of A,[:B,C:] holds Pr2 F = (pr2 (B,C)) * F;
theorem Th36:
theorem Th37:
definition
let A,
B,
C be
Category;
let F1,
F2 be
Functor of
A,
[:B,C:];
let t be
natural_transformation of
F1,
F2;
func Pr1 t -> natural_transformation of
Pr1 F1,
Pr1 F2 equals
(pr1 (B,C)) * t;
coherence
(pr1 (B,C)) * t is natural_transformation of Pr1 F1, Pr1 F2
;
func Pr2 t -> natural_transformation of
Pr2 F1,
Pr2 F2 equals
(pr2 (B,C)) * t;
coherence
(pr2 (B,C)) * t is natural_transformation of Pr2 F1, Pr2 F2
;
end;
:: deftheorem defines Pr1 ISOCAT_2:def 9 :
for A, B, C being Category
for F1, F2 being Functor of A,[:B,C:]
for t being natural_transformation of F1,F2 holds Pr1 t = (pr1 (B,C)) * t;
:: deftheorem defines Pr2 ISOCAT_2:def 10 :
for A, B, C being Category
for F1, F2 being Functor of A,[:B,C:]
for t being natural_transformation of F1,F2 holds Pr2 t = (pr2 (B,C)) * t;
theorem
canceled;
theorem Th39:
Lm2:
for A, B, C being Category
for F1, G1 being Functor of A,B
for F2, G2 being Functor of A,C st F1 is_transformable_to G1 & F2 is_transformable_to G2 holds
for t1 being transformation of F1,G1
for t2 being transformation of F2,G2
for a being Object of A holds <:t1,t2:> . a = [(t1 . a),(t2 . a)]
theorem
canceled;
theorem
canceled;
theorem Th42:
theorem Th43:
Lm3:
for A, B, C being Category
for F1, G1 being Functor of A,B
for F2, G2 being Functor of A,C st F1 is_transformable_to G1 & F2 is_transformable_to G2 holds
for t1 being transformation of F1,G1
for t2 being transformation of F2,G2
for a being Object of A holds <:t1,t2:> . a in Hom ((<:F1,F2:> . a),(<:G1,G2:> . a))
theorem Th44:
definition
let A,
B,
C be
Category;
let F1,
G1 be
Functor of
A,
B;
let F2,
G2 be
Functor of
A,
C;
assume A1:
(
F1 is_transformable_to G1 &
F2 is_transformable_to G2 )
;
let t1 be
transformation of
F1,
G1;
let t2 be
transformation of
F2,
G2;
func <:t1,t2:> -> transformation of
<:F1,F2:>,
<:G1,G2:> equals :
Def11:
<:t1,t2:>;
coherence
<:t1,t2:> is transformation of <:F1,F2:>,<:G1,G2:>
end;
:: deftheorem Def11 defines <: ISOCAT_2:def 11 :
for A, B, C being Category
for F1, G1 being Functor of A,B
for F2, G2 being Functor of A,C st F1 is_transformable_to G1 & F2 is_transformable_to G2 holds
for t1 being transformation of F1,G1
for t2 being transformation of F2,G2 holds <:t1,t2:> = <:t1,t2:>;
theorem Th45:
for
A,
B,
C being
Category for
F1,
G1 being
Functor of
A,
B for
F2,
G2 being
Functor of
A,
C st
F1 is_transformable_to G1 &
F2 is_transformable_to G2 holds
for
t1 being
transformation of
F1,
G1 for
t2 being
transformation of
F2,
G2 for
a being
Object of
A holds
<:t1,t2:> . a = [(t1 . a),(t2 . a)]
Lm4:
for A, B, C being Category
for F1, G1 being Functor of A,B
for F2, G2 being Functor of A,C st F1 is_naturally_transformable_to G1 & F2 is_naturally_transformable_to G2 holds
for t1 being natural_transformation of F1,G1
for t2 being natural_transformation of F2,G2
for a, b being Object of A st Hom (a,b) <> {} holds
for f being Morphism of a,b holds (<:t1,t2:> . b) * (<:F1,F2:> . f) = (<:G1,G2:> . f) * (<:t1,t2:> . a)
theorem Th46:
definition
let A,
B,
C be
Category;
let F1,
G1 be
Functor of
A,
B;
let F2,
G2 be
Functor of
A,
C;
assume A1:
(
F1 is_naturally_transformable_to G1 &
F2 is_naturally_transformable_to G2 )
;
let t1 be
natural_transformation of
F1,
G1;
let t2 be
natural_transformation of
F2,
G2;
func <:t1,t2:> -> natural_transformation of
<:F1,F2:>,
<:G1,G2:> equals :
Def12:
<:t1,t2:>;
coherence
<:t1,t2:> is natural_transformation of <:F1,F2:>,<:G1,G2:>
end;
:: deftheorem Def12 defines <: ISOCAT_2:def 12 :
for A, B, C being Category
for F1, G1 being Functor of A,B
for F2, G2 being Functor of A,C st F1 is_naturally_transformable_to G1 & F2 is_naturally_transformable_to G2 holds
for t1 being natural_transformation of F1,G1
for t2 being natural_transformation of F2,G2 holds <:t1,t2:> = <:t1,t2:>;
theorem Th47:
for
A,
B,
C being
Category for
F1,
G1 being
Functor of
A,
B for
F2,
G2 being
Functor of
A,
C st
F1 is_naturally_transformable_to G1 &
F2 is_naturally_transformable_to G2 holds
for
t1 being
natural_transformation of
F1,
G1 for
t2 being
natural_transformation of
F2,
G2 holds
(
Pr1 <:t1,t2:> = t1 &
Pr2 <:t1,t2:> = t2 )
definition
let A,
B,
C be
Category;
func distribute (
A,
B,
C)
-> Functor of
Functors (
A,
[:B,C:]),
[:(Functors (A,B)),(Functors (A,C)):] means :
Def13:
for
F1,
F2 being
Functor of
A,
[:B,C:] st
F1 is_naturally_transformable_to F2 holds
for
t being
natural_transformation of
F1,
F2 holds
it . [[F1,F2],t] = [[[(Pr1 F1),(Pr1 F2)],(Pr1 t)],[[(Pr2 F1),(Pr2 F2)],(Pr2 t)]];
existence
ex b1 being Functor of Functors (A,[:B,C:]),[:(Functors (A,B)),(Functors (A,C)):] st
for F1, F2 being Functor of A,[:B,C:] st F1 is_naturally_transformable_to F2 holds
for t being natural_transformation of F1,F2 holds b1 . [[F1,F2],t] = [[[(Pr1 F1),(Pr1 F2)],(Pr1 t)],[[(Pr2 F1),(Pr2 F2)],(Pr2 t)]]
uniqueness
for b1, b2 being Functor of Functors (A,[:B,C:]),[:(Functors (A,B)),(Functors (A,C)):] st ( for F1, F2 being Functor of A,[:B,C:] st F1 is_naturally_transformable_to F2 holds
for t being natural_transformation of F1,F2 holds b1 . [[F1,F2],t] = [[[(Pr1 F1),(Pr1 F2)],(Pr1 t)],[[(Pr2 F1),(Pr2 F2)],(Pr2 t)]] ) & ( for F1, F2 being Functor of A,[:B,C:] st F1 is_naturally_transformable_to F2 holds
for t being natural_transformation of F1,F2 holds b2 . [[F1,F2],t] = [[[(Pr1 F1),(Pr1 F2)],(Pr1 t)],[[(Pr2 F1),(Pr2 F2)],(Pr2 t)]] ) holds
b1 = b2
end;
:: deftheorem Def13 defines distribute ISOCAT_2:def 13 :
for A, B, C being Category
for b4 being Functor of Functors (A,[:B,C:]),[:(Functors (A,B)),(Functors (A,C)):] holds
( b4 = distribute (A,B,C) iff for F1, F2 being Functor of A,[:B,C:] st F1 is_naturally_transformable_to F2 holds
for t being natural_transformation of F1,F2 holds b4 . [[F1,F2],t] = [[[(Pr1 F1),(Pr1 F2)],(Pr1 t)],[[(Pr2 F1),(Pr2 F2)],(Pr2 t)]] );
theorem
canceled;
theorem