begin
:: deftheorem defines Funcs ENS_1:def 1 :
for V being non empty set holds Funcs V = union { (Funcs (A,B)) where A, B is Element of V : verum } ;
theorem Th1:
theorem
theorem Th3:
definition
let V be non
empty set ;
func Maps V -> set equals
{ [[A,B],f] where A, B is Element of V, f is Element of Funcs V : ( ( B = {} implies A = {} ) & f is Function of A,B ) } ;
coherence
{ [[A,B],f] where A, B is Element of V, f is Element of Funcs V : ( ( B = {} implies A = {} ) & f is Function of A,B ) } is set
;
end;
:: deftheorem defines Maps ENS_1:def 2 :
for V being non empty set holds Maps V = { [[A,B],f] where A, B is Element of V, f is Element of Funcs V : ( ( B = {} implies A = {} ) & f is Function of A,B ) } ;
theorem Th4:
theorem Th5:
theorem
theorem Th7:
Lm1:
for x1, y1, x2, y2, x3, y3 being set st [[x1,x2],x3] = [[y1,y2],y3] holds
( x1 = y1 & x2 = y2 )
:: deftheorem ENS_1:def 3 :
canceled;
:: deftheorem defines dom ENS_1:def 4 :
for V being non empty set
for m being Element of Maps V holds dom m = (m `1) `1 ;
:: deftheorem defines cod ENS_1:def 5 :
for V being non empty set
for m being Element of Maps V holds cod m = (m `1) `2 ;
theorem Th8:
Lm2:
for V being non empty set
for m, m9 being Element of Maps V st m `2 = m9 `2 & dom m = dom m9 & cod m = cod m9 holds
m = m9
theorem Th9:
Lm3:
for V being non empty set
for m being Element of Maps V holds
( dom (m `2) = dom m & rng (m `2) c= cod m )
theorem Th10:
:: deftheorem defines id$ ENS_1:def 6 :
for V being non empty set
for A being Element of V holds id$ A = [[A,A],(id A)];
theorem Th11:
:: deftheorem Def7 defines * ENS_1:def 7 :
for V being non empty set
for m1, m2 being Element of Maps V st cod m1 = dom m2 holds
m2 * m1 = [[(dom m1),(cod m2)],((m2 `2) * (m1 `2))];
theorem Th12:
theorem Th13:
theorem Th14:
definition
let V be non
empty set ;
let A,
B be
Element of
V;
func Maps (
A,
B)
-> set equals
{ [[A,B],f] where f is Element of Funcs V : [[A,B],f] in Maps V } ;
correctness
coherence
{ [[A,B],f] where f is Element of Funcs V : [[A,B],f] in Maps V } is set ;
;
end;
:: deftheorem defines Maps ENS_1:def 8 :
for V being non empty set
for A, B being Element of V holds Maps (A,B) = { [[A,B],f] where f is Element of Funcs V : [[A,B],f] in Maps V } ;
theorem Th15:
theorem Th16:
theorem Th17:
Lm4:
for V being non empty set
for A, B being Element of V
for f being Function st [[A,B],f] in Maps (A,B) holds
( ( B = {} implies A = {} ) & f is Function of A,B )
theorem
theorem Th19:
theorem Th20:
Lm5:
for V being non empty set
for W being non empty Subset of V
for A, B being Element of W
for A9, B9 being Element of V st A = A9 & B = B9 holds
Maps (A,B) = Maps (A9,B9)
:: deftheorem defines surjective ENS_1:def 9 :
for V being non empty set
for m being Element of Maps V holds
( m is surjective iff rng (m `2) = cod m );
begin
definition
let V be non
empty set ;
func fDom V -> Function of
(Maps V),
V means :
Def10:
for
m being
Element of
Maps V holds
it . m = dom m;
existence
ex b1 being Function of (Maps V),V st
for m being Element of Maps V holds b1 . m = dom m
uniqueness
for b1, b2 being Function of (Maps V),V st ( for m being Element of Maps V holds b1 . m = dom m ) & ( for m being Element of Maps V holds b2 . m = dom m ) holds
b1 = b2
func fCod V -> Function of
(Maps V),
V means :
Def11:
for
m being
Element of
Maps V holds
it . m = cod m;
existence
ex b1 being Function of (Maps V),V st
for m being Element of Maps V holds b1 . m = cod m
uniqueness
for b1, b2 being Function of (Maps V),V st ( for m being Element of Maps V holds b1 . m = cod m ) & ( for m being Element of Maps V holds b2 . m = cod m ) holds
b1 = b2
func fComp V -> PartFunc of
[:(Maps V),(Maps V):],
(Maps V) means :
Def12:
( ( for
m2,
m1 being
Element of
Maps V holds
(
[m2,m1] in dom it iff
dom m2 = cod m1 ) ) & ( for
m2,
m1 being
Element of
Maps V st
dom m2 = cod m1 holds
it . [m2,m1] = m2 * m1 ) );
existence
ex b1 being PartFunc of [:(Maps V),(Maps V):],(Maps V) st
( ( for m2, m1 being Element of Maps V holds
( [m2,m1] in dom b1 iff dom m2 = cod m1 ) ) & ( for m2, m1 being Element of Maps V st dom m2 = cod m1 holds
b1 . [m2,m1] = m2 * m1 ) )
uniqueness
for b1, b2 being PartFunc of [:(Maps V),(Maps V):],(Maps V) st ( for m2, m1 being Element of Maps V holds
( [m2,m1] in dom b1 iff dom m2 = cod m1 ) ) & ( for m2, m1 being Element of Maps V st dom m2 = cod m1 holds
b1 . [m2,m1] = m2 * m1 ) & ( for m2, m1 being Element of Maps V holds
( [m2,m1] in dom b2 iff dom m2 = cod m1 ) ) & ( for m2, m1 being Element of Maps V st dom m2 = cod m1 holds
b2 . [m2,m1] = m2 * m1 ) holds
b1 = b2
func fId V -> Function of
V,
(Maps V) means :
Def13:
for
A being
Element of
V holds
it . A = id$ A;
existence
ex b1 being Function of V,(Maps V) st
for A being Element of V holds b1 . A = id$ A
uniqueness
for b1, b2 being Function of V,(Maps V) st ( for A being Element of V holds b1 . A = id$ A ) & ( for A being Element of V holds b2 . A = id$ A ) holds
b1 = b2
end;
:: deftheorem Def10 defines fDom ENS_1:def 10 :
for V being non empty set
for b2 being Function of (Maps V),V holds
( b2 = fDom V iff for m being Element of Maps V holds b2 . m = dom m );
:: deftheorem Def11 defines fCod ENS_1:def 11 :
for V being non empty set
for b2 being Function of (Maps V),V holds
( b2 = fCod V iff for m being Element of Maps V holds b2 . m = cod m );
:: deftheorem Def12 defines fComp ENS_1:def 12 :
for V being non empty set
for b2 being PartFunc of [:(Maps V),(Maps V):],(Maps V) holds
( b2 = fComp V iff ( ( for m2, m1 being Element of Maps V holds
( [m2,m1] in dom b2 iff dom m2 = cod m1 ) ) & ( for m2, m1 being Element of Maps V st dom m2 = cod m1 holds
b2 . [m2,m1] = m2 * m1 ) ) );
:: deftheorem Def13 defines fId ENS_1:def 13 :
for V being non empty set
for b2 being Function of V,(Maps V) holds
( b2 = fId V iff for A being Element of V holds b2 . A = id$ A );
definition
let V be non
empty set ;
func Ens V -> CatStr equals
CatStr(#
V,
(Maps V),
(fDom V),
(fCod V),
(fComp V),
(fId V) #);
coherence
CatStr(# V,(Maps V),(fDom V),(fCod V),(fComp V),(fId V) #) is CatStr
;
end;
:: deftheorem defines Ens ENS_1:def 14 :
for V being non empty set holds Ens V = CatStr(# V,(Maps V),(fDom V),(fCod V),(fComp V),(fId V) #);
theorem Th21:
theorem
:: deftheorem defines @ ENS_1:def 15 :
for V being non empty set
for A being Element of V holds @ A = A;
theorem
:: deftheorem defines @ ENS_1:def 16 :
for V being non empty set
for a being Object of (Ens V) holds @ a = a;
theorem
:: deftheorem defines @ ENS_1:def 17 :
for V being non empty set
for m being Element of Maps V holds @ m = m;
theorem
:: deftheorem defines @ ENS_1:def 18 :
for V being non empty set
for f being Morphism of (Ens V) holds @ f = f;
theorem
theorem Th27:
Lm6:
for V being non empty set
for a, b being Object of (Ens V) st Hom (a,b) <> {} holds
Funcs ((@ a),(@ b)) <> {}
theorem Th28:
theorem
theorem
theorem Th31:
theorem
theorem
theorem Th34:
theorem
theorem
theorem Th37:
theorem
theorem
theorem
begin
:: deftheorem defines Hom ENS_1:def 19 :
for C being Category holds Hom C = { (Hom (a,b)) where a, b is Object of C : verum } ;
theorem
theorem
definition
let C be
Category;
let a be
Object of
C;
let f be
Morphism of
C;
func hom (
a,
f)
-> Function of
(Hom (a,(dom f))),
(Hom (a,(cod f))) means :
Def20:
for
g being
Morphism of
C st
g in Hom (
a,
(dom f)) holds
it . g = f * g;
existence
ex b1 being Function of (Hom (a,(dom f))),(Hom (a,(cod f))) st
for g being Morphism of C st g in Hom (a,(dom f)) holds
b1 . g = f * g
uniqueness
for b1, b2 being Function of (Hom (a,(dom f))),(Hom (a,(cod f))) st ( for g being Morphism of C st g in Hom (a,(dom f)) holds
b1 . g = f * g ) & ( for g being Morphism of C st g in Hom (a,(dom f)) holds
b2 . g = f * g ) holds
b1 = b2
func hom (
f,
a)
-> Function of
(Hom ((cod f),a)),
(Hom ((dom f),a)) means :
Def21:
for
g being
Morphism of
C st
g in Hom (
(cod f),
a) holds
it . g = g * f;
existence
ex b1 being Function of (Hom ((cod f),a)),(Hom ((dom f),a)) st
for g being Morphism of C st g in Hom ((cod f),a) holds
b1 . g = g * f
uniqueness
for b1, b2 being Function of (Hom ((cod f),a)),(Hom ((dom f),a)) st ( for g being Morphism of C st g in Hom ((cod f),a) holds
b1 . g = g * f ) & ( for g being Morphism of C st g in Hom ((cod f),a) holds
b2 . g = g * f ) holds
b1 = b2
end;
:: deftheorem Def20 defines hom ENS_1:def 20 :
for C being Category
for a being Object of C
for f being Morphism of C
for b4 being Function of (Hom (a,(dom f))),(Hom (a,(cod f))) holds
( b4 = hom (a,f) iff for g being Morphism of C st g in Hom (a,(dom f)) holds
b4 . g = f * g );
:: deftheorem Def21 defines hom ENS_1:def 21 :
for C being Category
for a being Object of C
for f being Morphism of C
for b4 being Function of (Hom ((cod f),a)),(Hom ((dom f),a)) holds
( b4 = hom (f,a) iff for g being Morphism of C st g in Hom ((cod f),a) holds
b4 . g = g * f );
theorem Th43:
theorem Th44:
theorem Th45:
theorem Th46:
theorem Th47:
theorem Th48:
definition
let C be
Category;
let a be
Object of
C;
func hom?- a -> Function of the
carrier' of
C,
(Maps (Hom C)) means :
Def22:
for
f being
Morphism of
C holds
it . f = [[(Hom (a,(dom f))),(Hom (a,(cod f)))],(hom (a,f))];
existence
ex b1 being Function of the carrier' of C,(Maps (Hom C)) st
for f being Morphism of C holds b1 . f = [[(Hom (a,(dom f))),(Hom (a,(cod f)))],(hom (a,f))]
uniqueness
for b1, b2 being Function of the carrier' of C,(Maps (Hom C)) st ( for f being Morphism of C holds b1 . f = [[(Hom (a,(dom f))),(Hom (a,(cod f)))],(hom (a,f))] ) & ( for f being Morphism of C holds b2 . f = [[(Hom (a,(dom f))),(Hom (a,(cod f)))],(hom (a,f))] ) holds
b1 = b2
func hom-? a -> Function of the
carrier' of
C,
(Maps (Hom C)) means :
Def23:
for
f being
Morphism of
C holds
it . f = [[(Hom ((cod f),a)),(Hom ((dom f),a))],(hom (f,a))];
existence
ex b1 being Function of the carrier' of C,(Maps (Hom C)) st
for f being Morphism of C holds b1 . f = [[(Hom ((cod f),a)),(Hom ((dom f),a))],(hom (f,a))]
uniqueness
for b1, b2 being Function of the carrier' of C,(Maps (Hom C)) st ( for f being Morphism of C holds b1 . f = [[(Hom ((cod f),a)),(Hom ((dom f),a))],(hom (f,a))] ) & ( for f being Morphism of C holds b2 . f = [[(Hom ((cod f),a)),(Hom ((dom f),a))],(hom (f,a))] ) holds
b1 = b2
end;
:: deftheorem Def22 defines hom?- ENS_1:def 22 :
for C being Category
for a being Object of C
for b3 being Function of the carrier' of C,(Maps (Hom C)) holds
( b3 = hom?- a iff for f being Morphism of C holds b3 . f = [[(Hom (a,(dom f))),(Hom (a,(cod f)))],(hom (a,f))] );
:: deftheorem Def23 defines hom-? ENS_1:def 23 :
for C being Category
for a being Object of C
for b3 being Function of the carrier' of C,(Maps (Hom C)) holds
( b3 = hom-? a iff for f being Morphism of C holds b3 . f = [[(Hom ((cod f),a)),(Hom ((dom f),a))],(hom (f,a))] );
Lm7:
for V being non empty set
for C being Category
for T being Function of the carrier' of C,(Maps (Hom C)) st Hom C c= V holds
T is Function of the carrier' of C, the carrier' of (Ens V)
Lm8:
for V being non empty set
for C being Category
for a, c being Object of C st Hom C c= V holds
for d being Object of (Ens V) st d = Hom (a,c) holds
(hom?- a) . (id c) = id d
Lm9:
for V being non empty set
for C being Category
for c, a being Object of C st Hom C c= V holds
for d being Object of (Ens V) st d = Hom (c,a) holds
(hom-? a) . (id c) = id d
theorem Th49:
theorem Th50:
theorem Th51:
definition
let C be
Category;
let f,
g be
Morphism of
C;
func hom (
f,
g)
-> Function of
(Hom ((cod f),(dom g))),
(Hom ((dom f),(cod g))) means :
Def24:
for
h being
Morphism of
C st
h in Hom (
(cod f),
(dom g)) holds
it . h = (g * h) * f;
existence
ex b1 being Function of (Hom ((cod f),(dom g))),(Hom ((dom f),(cod g))) st
for h being Morphism of C st h in Hom ((cod f),(dom g)) holds
b1 . h = (g * h) * f
uniqueness
for b1, b2 being Function of (Hom ((cod f),(dom g))),(Hom ((dom f),(cod g))) st ( for h being Morphism of C st h in Hom ((cod f),(dom g)) holds
b1 . h = (g * h) * f ) & ( for h being Morphism of C st h in Hom ((cod f),(dom g)) holds
b2 . h = (g * h) * f ) holds
b1 = b2
end;
:: deftheorem Def24 defines hom ENS_1:def 24 :
for C being Category
for f, g being Morphism of C
for b4 being Function of (Hom ((cod f),(dom g))),(Hom ((dom f),(cod g))) holds
( b4 = hom (f,g) iff for h being Morphism of C st h in Hom ((cod f),(dom g)) holds
b4 . h = (g * h) * f );
theorem Th52:
theorem Th53:
theorem Th54:
theorem
theorem Th56:
definition
let C be
Category;
func hom?? C -> Function of the
carrier' of
[:C,C:],
(Maps (Hom C)) means :
Def25:
for
f,
g being
Morphism of
C holds
it . [f,g] = [[(Hom ((cod f),(dom g))),(Hom ((dom f),(cod g)))],(hom (f,g))];
existence
ex b1 being Function of the carrier' of [:C,C:],(Maps (Hom C)) st
for f, g being Morphism of C holds b1 . [f,g] = [[(Hom ((cod f),(dom g))),(Hom ((dom f),(cod g)))],(hom (f,g))]
uniqueness
for b1, b2 being Function of the carrier' of [:C,C:],(Maps (Hom C)) st ( for f, g being Morphism of C holds b1 . [f,g] = [[(Hom ((cod f),(dom g))),(Hom ((dom f),(cod g)))],(hom (f,g))] ) & ( for f, g being Morphism of C holds b2 . [f,g] = [[(Hom ((cod f),(dom g))),(Hom ((dom f),(cod g)))],(hom (f,g))] ) holds
b1 = b2
end;
:: deftheorem Def25 defines hom?? ENS_1:def 25 :
for C being Category
for b2 being Function of the carrier' of [:C,C:],(Maps (Hom C)) holds
( b2 = hom?? C iff for f, g being Morphism of C holds b2 . [f,g] = [[(Hom ((cod f),(dom g))),(Hom ((dom f),(cod g)))],(hom (f,g))] );
theorem Th57:
Lm10:
for V being non empty set
for C being Category
for a, b being Object of C st Hom C c= V holds
for d being Object of (Ens V) st d = Hom (a,b) holds
(hom?? C) . [(id a),(id b)] = id d
theorem Th58:
:: deftheorem Def26 defines hom?- ENS_1:def 26 :
for V being non empty set
for C being Category
for a being Object of C st Hom C c= V holds
hom?- (V,a) = hom?- a;
:: deftheorem Def27 defines hom-? ENS_1:def 27 :
for V being non empty set
for C being Category
for a being Object of C st Hom C c= V holds
hom-? (V,a) = hom-? a;
:: deftheorem Def28 defines hom?? ENS_1:def 28 :
for V being non empty set
for C being Category st Hom C c= V holds
hom?? (V,C) = hom?? C;
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem