begin
:: deftheorem RINGCAT1:def 1 :
canceled;
:: deftheorem Def2 defines linear RINGCAT1:def 2 :
for G, H being non empty doubleLoopStr
for IT being Function of G,H holds
( IT is linear iff ( IT is additive & IT is multiplicative & IT is unity-preserving ) );
theorem
canceled;
theorem
canceled;
theorem Th3:
:: deftheorem defines dom RINGCAT1:def 3 :
for f being RingMorphismStr holds dom f = the Dom of f;
:: deftheorem defines cod RINGCAT1:def 4 :
for f being RingMorphismStr holds cod f = the Cod of f;
:: deftheorem defines fun RINGCAT1:def 5 :
for f being RingMorphismStr holds fun f = the Fun of f;
Lm1:
for G being non empty doubleLoopStr holds id G is linear
:: deftheorem Def6 defines RingMorphism-like RINGCAT1:def 6 :
for IT being RingMorphismStr holds
( IT is RingMorphism-like iff fun IT is linear );
:: deftheorem defines ID RINGCAT1:def 7 :
for G being Ring holds ID G = RingMorphismStr(# G,G,(id G) #);
:: deftheorem Def8 defines <= RINGCAT1:def 8 :
for G, H being Ring holds
( G <= H iff ex F being RingMorphism st
( dom F = G & cod F = H ) );
Lm2:
for F being RingMorphism holds RingMorphismStr(# the Dom of F, the Cod of F, the Fun of F #) is RingMorphism-like
:: deftheorem Def9 defines Morphism RINGCAT1:def 9 :
for G, H being Ring st G <= H holds
for b3 being strict RingMorphism holds
( b3 is Morphism of G,H iff ( dom b3 = G & cod b3 = H ) );
Lm3:
for F being RingMorphism holds the Fun of F is linear
Lm4:
for G, H being Ring
for f being strict RingMorphismStr st dom f = G & cod f = H & fun f is linear holds
f is Morphism of G,H
Lm5:
for G, H being Ring
for f being Function of G,H st f is linear holds
RingMorphismStr(# G,H,f #) is Morphism of G,H
Lm6:
for F being RingMorphism ex G, H being Ring st
( G <= H & dom F = G & cod F = H & RingMorphismStr(# the Dom of F, the Cod of F, the Fun of F #) is Morphism of G,H )
theorem
canceled;
theorem Th5:
theorem Th6:
Lm7:
for G, H being Ring
for F being Morphism of G,H st G <= H holds
ex f being Function of G,H st
( F = RingMorphismStr(# G,H,f #) & f is linear )
Lm8:
for G, H being Ring
for F being Morphism of G,H st G <= H holds
ex f being Function of G,H st F = RingMorphismStr(# G,H,f #)
theorem
definition
let G,
F be
RingMorphism;
assume A1:
dom G = cod F
;
func G * F -> strict RingMorphism means :
Def10:
for
G1,
G2,
G3 being
Ring for
g being
Function of
G2,
G3 for
f being
Function of
G1,
G2 st
RingMorphismStr(# the
Dom of
G, the
Cod of
G, the
Fun of
G #)
= RingMorphismStr(#
G2,
G3,
g #) &
RingMorphismStr(# the
Dom of
F, the
Cod of
F, the
Fun of
F #)
= RingMorphismStr(#
G1,
G2,
f #) holds
it = RingMorphismStr(#
G1,
G3,
(g * f) #);
existence
ex b1 being strict RingMorphism st
for G1, G2, G3 being Ring
for g being Function of G2,G3
for f being Function of G1,G2 st RingMorphismStr(# the Dom of G, the Cod of G, the Fun of G #) = RingMorphismStr(# G2,G3,g #) & RingMorphismStr(# the Dom of F, the Cod of F, the Fun of F #) = RingMorphismStr(# G1,G2,f #) holds
b1 = RingMorphismStr(# G1,G3,(g * f) #)
uniqueness
for b1, b2 being strict RingMorphism st ( for G1, G2, G3 being Ring
for g being Function of G2,G3
for f being Function of G1,G2 st RingMorphismStr(# the Dom of G, the Cod of G, the Fun of G #) = RingMorphismStr(# G2,G3,g #) & RingMorphismStr(# the Dom of F, the Cod of F, the Fun of F #) = RingMorphismStr(# G1,G2,f #) holds
b1 = RingMorphismStr(# G1,G3,(g * f) #) ) & ( for G1, G2, G3 being Ring
for g being Function of G2,G3
for f being Function of G1,G2 st RingMorphismStr(# the Dom of G, the Cod of G, the Fun of G #) = RingMorphismStr(# G2,G3,g #) & RingMorphismStr(# the Dom of F, the Cod of F, the Fun of F #) = RingMorphismStr(# G1,G2,f #) holds
b2 = RingMorphismStr(# G1,G3,(g * f) #) ) holds
b1 = b2
end;
:: deftheorem Def10 defines * RINGCAT1:def 10 :
for G, F being RingMorphism st dom G = cod F holds
for b3 being strict RingMorphism holds
( b3 = G * F iff for G1, G2, G3 being Ring
for g being Function of G2,G3
for f being Function of G1,G2 st RingMorphismStr(# the Dom of G, the Cod of G, the Fun of G #) = RingMorphismStr(# G2,G3,g #) & RingMorphismStr(# the Dom of F, the Cod of F, the Fun of F #) = RingMorphismStr(# G1,G2,f #) holds
b3 = RingMorphismStr(# G1,G3,(g * f) #) );
theorem Th8:
for
G1,
G2,
G3 being
Ring st
G1 <= G2 &
G2 <= G3 holds
G1 <= G3
theorem Th9:
:: deftheorem Def11 defines *' RINGCAT1:def 11 :
for G1, G2, G3 being Ring
for G being Morphism of G2,G3
for F being Morphism of G1,G2 st G1 <= G2 & G2 <= G3 holds
G *' F = G * F;
theorem Th10:
for
f,
g being
strict RingMorphism st
dom g = cod f holds
ex
G1,
G2,
G3 being
Ring ex
f0 being
Function of
G1,
G2 ex
g0 being
Function of
G2,
G3 st
(
f = RingMorphismStr(#
G1,
G2,
f0 #) &
g = RingMorphismStr(#
G2,
G3,
g0 #) &
g * f = RingMorphismStr(#
G1,
G3,
(g0 * f0) #) )
theorem Th11:
theorem Th12:
theorem Th13:
theorem Th14:
:: deftheorem Def12 defines Ring_DOMAIN-like RINGCAT1:def 12 :
for IT being set holds
( IT is Ring_DOMAIN-like iff for x being Element of IT holds x is strict Ring );
:: deftheorem Def13 defines RingMorphism_DOMAIN-like RINGCAT1:def 13 :
for IT being set holds
( IT is RingMorphism_DOMAIN-like iff for x being set st x in IT holds
x is strict RingMorphism );
theorem
canceled;
theorem
canceled;
theorem Th17:
:: deftheorem Def14 defines RingMorphism_DOMAIN RINGCAT1:def 14 :
for G, H being Ring
for b3 being RingMorphism_DOMAIN holds
( b3 is RingMorphism_DOMAIN of G,H iff for x being Element of b3 holds x is Morphism of G,H );
theorem Th18:
theorem
definition
let G,
H be
Ring;
assume A1:
G <= H
;
func Morphs (
G,
H)
-> RingMorphism_DOMAIN of
G,
H means :
Def15:
for
x being
set holds
(
x in it iff
x is
Morphism of
G,
H );
existence
ex b1 being RingMorphism_DOMAIN of G,H st
for x being set holds
( x in b1 iff x is Morphism of G,H )
uniqueness
for b1, b2 being RingMorphism_DOMAIN of G,H st ( for x being set holds
( x in b1 iff x is Morphism of G,H ) ) & ( for x being set holds
( x in b2 iff x is Morphism of G,H ) ) holds
b1 = b2
end;
:: deftheorem Def15 defines Morphs RINGCAT1:def 15 :
for G, H being Ring st G <= H holds
for b3 being RingMorphism_DOMAIN of G,H holds
( b3 = Morphs (G,H) iff for x being set holds
( x in b3 iff x is Morphism of G,H ) );
definition
let x,
y be
set ;
pred GO x,
y means :
Def16:
ex
x1,
x2,
x3,
x4,
x5,
x6 being
set st
(
x = [[x1,x2,x3,x4],x5,x6] & ex
G being
strict Ring st
(
y = G &
x1 = the
carrier of
G &
x2 = the
addF of
G &
x3 = comp G &
x4 = 0. G &
x5 = the
multF of
G &
x6 = 1. G ) );
end;
:: deftheorem Def16 defines GO RINGCAT1:def 16 :
for x, y being set holds
( GO x,y iff ex x1, x2, x3, x4, x5, x6 being set st
( x = [[x1,x2,x3,x4],x5,x6] & ex G being strict Ring st
( y = G & x1 = the carrier of G & x2 = the addF of G & x3 = comp G & x4 = 0. G & x5 = the multF of G & x6 = 1. G ) ) );
theorem Th20:
for
x,
y1,
y2 being
set st
GO x,
y1 &
GO x,
y2 holds
y1 = y2
theorem Th21:
:: deftheorem Def17 defines RingObjects RINGCAT1:def 17 :
for UN being Universe
for b2 being set holds
( b2 = RingObjects UN iff for y being set holds
( y in b2 iff ex x being set st
( x in UN & GO x,y ) ) );
theorem Th22:
theorem Th23:
:: deftheorem Def18 defines Morphs RINGCAT1:def 18 :
for V being Ring_DOMAIN
for b2 being RingMorphism_DOMAIN holds
( b2 = Morphs V iff for x being set holds
( x in b2 iff ex G, H being Element of V st
( G <= H & x is Morphism of G,H ) ) );
:: deftheorem defines ID RINGCAT1:def 19 :
for V being Ring_DOMAIN
for G being Element of V holds ID G = ID G;
definition
let V be
Ring_DOMAIN;
func dom V -> Function of
(Morphs V),
V means :
Def20:
for
f being
Element of
Morphs V holds
it . f = dom f;
existence
ex b1 being Function of (Morphs V),V st
for f being Element of Morphs V holds b1 . f = dom f
uniqueness
for b1, b2 being Function of (Morphs V),V st ( for f being Element of Morphs V holds b1 . f = dom f ) & ( for f being Element of Morphs V holds b2 . f = dom f ) holds
b1 = b2
func cod V -> Function of
(Morphs V),
V means :
Def21:
for
f being
Element of
Morphs V holds
it . f = cod f;
existence
ex b1 being Function of (Morphs V),V st
for f being Element of Morphs V holds b1 . f = cod f
uniqueness
for b1, b2 being Function of (Morphs V),V st ( for f being Element of Morphs V holds b1 . f = cod f ) & ( for f being Element of Morphs V holds b2 . f = cod f ) holds
b1 = b2
func ID V -> Function of
V,
(Morphs V) means :
Def22:
for
G being
Element of
V holds
it . G = ID G;
existence
ex b1 being Function of V,(Morphs V) st
for G being Element of V holds b1 . G = ID G
uniqueness
for b1, b2 being Function of V,(Morphs V) st ( for G being Element of V holds b1 . G = ID G ) & ( for G being Element of V holds b2 . G = ID G ) holds
b1 = b2
end;
:: deftheorem Def20 defines dom RINGCAT1:def 20 :
for V being Ring_DOMAIN
for b2 being Function of (Morphs V),V holds
( b2 = dom V iff for f being Element of Morphs V holds b2 . f = dom f );
:: deftheorem Def21 defines cod RINGCAT1:def 21 :
for V being Ring_DOMAIN
for b2 being Function of (Morphs V),V holds
( b2 = cod V iff for f being Element of Morphs V holds b2 . f = cod f );
:: deftheorem Def22 defines ID RINGCAT1:def 22 :
for V being Ring_DOMAIN
for b2 being Function of V,(Morphs V) holds
( b2 = ID V iff for G being Element of V holds b2 . G = ID G );
theorem Th24:
theorem Th25:
definition
let V be
Ring_DOMAIN;
func comp V -> PartFunc of
[:(Morphs V),(Morphs V):],
(Morphs V) means :
Def23:
( ( for
g,
f being
Element of
Morphs V holds
(
[g,f] in dom it iff
dom g = cod f ) ) & ( for
g,
f being
Element of
Morphs V st
[g,f] in dom it holds
it . (
g,
f)
= g * f ) );
existence
ex b1 being PartFunc of [:(Morphs V),(Morphs V):],(Morphs V) st
( ( for g, f being Element of Morphs V holds
( [g,f] in dom b1 iff dom g = cod f ) ) & ( for g, f being Element of Morphs V st [g,f] in dom b1 holds
b1 . (g,f) = g * f ) )
uniqueness
for b1, b2 being PartFunc of [:(Morphs V),(Morphs V):],(Morphs V) st ( for g, f being Element of Morphs V holds
( [g,f] in dom b1 iff dom g = cod f ) ) & ( for g, f being Element of Morphs V st [g,f] in dom b1 holds
b1 . (g,f) = g * f ) & ( for g, f being Element of Morphs V holds
( [g,f] in dom b2 iff dom g = cod f ) ) & ( for g, f being Element of Morphs V st [g,f] in dom b2 holds
b2 . (g,f) = g * f ) holds
b1 = b2
end;
:: deftheorem Def23 defines comp RINGCAT1:def 23 :
for V being Ring_DOMAIN
for b2 being PartFunc of [:(Morphs V),(Morphs V):],(Morphs V) holds
( b2 = comp V iff ( ( for g, f being Element of Morphs V holds
( [g,f] in dom b2 iff dom g = cod f ) ) & ( for g, f being Element of Morphs V st [g,f] in dom b2 holds
b2 . (g,f) = g * f ) ) );
definition
let UN be
Universe;
func RingCat UN -> CatStr equals
CatStr(#
(RingObjects UN),
(Morphs (RingObjects UN)),
(dom (RingObjects UN)),
(cod (RingObjects UN)),
(comp (RingObjects UN)),
(ID (RingObjects UN)) #);
coherence
CatStr(# (RingObjects UN),(Morphs (RingObjects UN)),(dom (RingObjects UN)),(cod (RingObjects UN)),(comp (RingObjects UN)),(ID (RingObjects UN)) #) is CatStr
;
end;
:: deftheorem defines RingCat RINGCAT1:def 24 :
for UN being Universe holds RingCat UN = CatStr(# (RingObjects UN),(Morphs (RingObjects UN)),(dom (RingObjects UN)),(cod (RingObjects UN)),(comp (RingObjects UN)),(ID (RingObjects UN)) #);
theorem Th26:
theorem Th27:
theorem Th28:
theorem
canceled;
theorem Th30:
Lm9:
for UN being Universe
for f, g being Morphism of (RingCat UN) st dom g = cod f holds
( dom (g * f) = dom f & cod (g * f) = cod g )
Lm10:
for UN being Universe
for f, g, h being Morphism of (RingCat UN) st dom h = cod g & dom g = cod f holds
h * (g * f) = (h * g) * f
Lm11:
for UN being Universe
for b being Object of (RingCat UN) holds
( dom (id b) = b & cod (id b) = b & ( for f being Morphism of (RingCat UN) st cod f = b holds
(id b) * f = f ) & ( for g being Morphism of (RingCat UN) st dom g = b holds
g * (id b) = g ) )