:: MOD_4 semantic presentation

definition
let A be non empty set ;
let B be non empty set ;
let C be non empty set ;
let f be Function of [:A,B:],C;
redefine func ~ as ~ c4 -> Function of [:a2,a1:],a3;
coherence
~ f is Function of [:B,A:],C
by FUNCT_4:51;
end;

theorem Th1: :: MOD_4:1
for C, A, B being non empty set
for f being Function of [:A,B:],C
for x being Element of A
for y being Element of B holds f . x,y = (~ f) . y,x
proof end;

definition
let K be non empty doubleLoopStr ;
func opp c1 -> strict doubleLoopStr equals :: MOD_4:def 1
doubleLoopStr(# the carrier of K,the add of K,(~ the mult of K),the unity of K,the Zero of K #);
correctness
coherence
doubleLoopStr(# the carrier of K,the add of K,(~ the mult of K),the unity of K,the Zero of K #) is strict doubleLoopStr
;
;
end;

:: deftheorem Def1 defines opp MOD_4:def 1 :
for K being non empty doubleLoopStr holds opp K = doubleLoopStr(# the carrier of K,the add of K,(~ the mult of K),the unity of K,the Zero of K #);

registration
let K be non empty doubleLoopStr ;
cluster opp a1 -> non empty strict ;
coherence
not opp K is empty
proof end;
end;

Lemma28: for K being non empty doubleLoopStr
for x, y being Scalar of K
for a, b being Scalar of (opp K) st x = a & y = b holds
( x + y = a + b & x * y = b * a )
by ;

Lemma31: for K being non empty unital doubleLoopStr
for h, e being Element of (opp K) st e = 1. K holds
( h * e = h & e * h = h )
proof end;

registration
let K be non empty unital doubleLoopStr ;
cluster opp a1 -> non empty unital strict ;
coherence
opp K is unital
proof end;
end;

Lemma34: for K being non empty unital doubleLoopStr holds 1. K = 1. (opp K)
proof end;

E36: now
let K be non empty add-associative right_zeroed right_complementable doubleLoopStr ;
set L = opp K;
thus for x, y, z being Scalar of (opp K) holds
( (x + y) + z = x + (y + z) & x + (0. (opp K)) = x )
proof
let x be Scalar of (opp K), y be Scalar of (opp K), z be Scalar of (opp K);
reconsider a = x, b = y, c = z as Scalar of K ;
thus (x + y) + z = (a + b) + c
.= a + (b + c) by RLVECT_1:def 6
.= x + (y + z) ;
thus x + (0. (opp K)) = a + (0. K)
.= x by RLVECT_1:def 7 ;
end;
end;

registration
let K be non empty add-associative right_zeroed right_complementable doubleLoopStr ;
cluster opp a1 -> non empty add-associative right_zeroed right_complementable strict ;
coherence
( opp K is add-associative & opp K is right_zeroed & opp K is right_complementable )
proof end;
end;

Lemma40: for K being non empty add-associative right_zeroed right_complementable doubleLoopStr
for x, y being Scalar of K
for a, b being Scalar of (opp K) st x = a & y = b holds
( x + y = a + b & x * y = b * a & - x = - a )
proof end;

theorem Th2: :: MOD_4:2
for K being non empty doubleLoopStr holds
( LoopStr(# the carrier of (opp K),the add of (opp K),the Zero of (opp K) #) = LoopStr(# the carrier of K,the add of K,the Zero of K #) & ( K is add-associative & K is right_zeroed & K is right_complementable implies comp (opp K) = comp K ) & ( for x being set holds
( x is Scalar of (opp K) iff x is Scalar of K ) ) )
proof end;

theorem Th3: :: MOD_4:3
for K being non empty doubleLoopStr holds opp (opp K) = doubleLoopStr(# the carrier of K,the add of K,the mult of K,the unity of K,the Zero of K #) by FUNCT_4:54;

Lemma45: for K being non empty add-associative right_zeroed right_complementable doubleLoopStr
for x, y, z, u being Scalar of K
for a, b, c, d being Scalar of (opp K) st x = a & y = b & z = c & u = d holds
( (x + y) + z = (a + b) + c & x + (y + z) = a + (b + c) & (x * y) * z = c * (b * a) & x * (y * z) = (c * b) * a & x * (y + z) = (b + c) * a & (y + z) * x = a * (b + c) & (x * y) + (z * u) = (b * a) + (d * c) )
proof end;

theorem Th4: :: MOD_4:4
( ( for K being non empty unital doubleLoopStr holds 1. K = 1. (opp K) ) & ( for K being non empty add-associative right_zeroed right_complementable doubleLoopStr holds
( 0. K = 0. (opp K) & ( for x, y, z, u being Scalar of K
for a, b, c, d being Scalar of (opp K) st x = a & y = b & z = c & u = d holds
( x + y = a + b & x * y = b * a & - x = - a & (x + y) + z = (a + b) + c & x + (y + z) = a + (b + c) & (x * y) * z = c * (b * a) & x * (y * z) = (c * b) * a & x * (y + z) = (b + c) * a & (y + z) * x = a * (b + c) & (x * y) + (z * u) = (b * a) + (d * c) ) ) ) ) ) by , , ;

theorem Th5: :: MOD_4:5
for K being Ring holds opp K is strict Ring
proof end;

registration
let K be Ring;
cluster opp a1 -> non empty Abelian add-associative right_zeroed right_complementable unital strict distributive ;
coherence
( opp K is Abelian & opp K is add-associative & opp K is right_zeroed & opp K is right_complementable & opp K is unital & opp K is distributive )
by ;
end;

theorem Th6: :: MOD_4:6
for K being Ring holds opp K is Ring
proof end;

registration
let K be Ring;
cluster opp a1 -> non empty Abelian add-associative right_zeroed right_complementable unital associative strict distributive ;
coherence
opp K is associative
by ;
end;

theorem Th7: :: MOD_4:7
for K being Skew-Field holds opp K is Skew-Field
proof end;

registration
let K be Skew-Field;
cluster opp a1 -> non empty Abelian add-associative right_zeroed right_complementable unital associative strict distributive Field-like non degenerated ;
coherence
( not opp K is degenerated & opp K is Field-like & opp K is associative & opp K is Abelian & opp K is add-associative & opp K is right_zeroed & opp K is right_complementable & opp K is unital & opp K is distributive )
by ;
end;

Lemma51: for F being non empty commutative doubleLoopStr
for x, y being Element of F holds x * y = y * x
;

theorem Th8: :: MOD_4:8
for K being Field holds opp K is strict Field
proof end;

registration
let K be Field;
cluster opp a1 -> non empty Abelian add-associative right_zeroed right_complementable unital associative strict distributive Field-like non degenerated ;
coherence
( opp K is strict & opp K is Field-like )
;
end;

definition
let K be non empty doubleLoopStr ;
let V be non empty VectSpStr of K;
func opp c2 -> strict RightModStr of opp a1 means :Def2: :: MOD_4:def 2
for o being Function of [:the carrier of V,the carrier of (opp K):],the carrier of V st o = ~ the lmult of V holds
it = RightModStr(# the carrier of V,the add of V,the Zero of V,o #);
existence
ex b1 being strict RightModStr of opp K st
for o being Function of [:the carrier of V,the carrier of (opp K):],the carrier of V st o = ~ the lmult of V holds
b1 = RightModStr(# the carrier of V,the add of V,the Zero of V,o #)
proof end;
uniqueness
for b1, b2 being strict RightModStr of opp K st ( for o being Function of [:the carrier of V,the carrier of (opp K):],the carrier of V st o = ~ the lmult of V holds
b1 = RightModStr(# the carrier of V,the add of V,the Zero of V,o #) ) & ( for o being Function of [:the carrier of V,the carrier of (opp K):],the carrier of V st o = ~ the lmult of V holds
b2 = RightModStr(# the carrier of V,the add of V,the Zero of V,o #) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def2 defines opp MOD_4:def 2 :
for K being non empty doubleLoopStr
for V being non empty VectSpStr of K
for b3 being strict RightModStr of opp K holds
( b3 = opp V iff for o being Function of [:the carrier of V,the carrier of (opp K):],the carrier of V st o = ~ the lmult of V holds
b3 = RightModStr(# the carrier of V,the add of V,the Zero of V,o #) );

registration
let K be non empty doubleLoopStr ;
let V be non empty VectSpStr of K;
cluster opp a2 -> non empty strict ;
coherence
not opp V is empty
proof end;
end;

theorem Th9: :: MOD_4:9
for K being non empty doubleLoopStr
for V being non empty VectSpStr of K holds
( LoopStr(# the carrier of (opp V),the add of (opp V),the Zero of (opp V) #) = LoopStr(# the carrier of V,the add of V,the Zero of V #) & ( for x being set holds
( x is Vector of V iff x is Vector of (opp V) ) ) )
proof end;

definition
let K be non empty doubleLoopStr ;
let V be non empty VectSpStr of K;
let o be Function of [:the carrier of K,the carrier of V:],the carrier of V;
func opp c3 -> Function of [:the carrier of (opp a2),the carrier of (opp a1):],the carrier of (opp a2) equals :: MOD_4:def 3
~ o;
coherence
~ o is Function of [:the carrier of (opp V),the carrier of (opp K):],the carrier of (opp V)
proof end;
end;

:: deftheorem Def3 defines opp MOD_4:def 3 :
for K being non empty doubleLoopStr
for V being non empty VectSpStr of K
for o being Function of [:the carrier of K,the carrier of V:],the carrier of V holds opp o = ~ o;

theorem Th10: :: MOD_4:10
for K being non empty doubleLoopStr
for V being non empty VectSpStr of K holds the rmult of (opp V) = opp the lmult of V
proof end;

definition
let K be non empty doubleLoopStr ;
let W be non empty RightModStr of K;
func opp c2 -> strict VectSpStr of opp a1 means :Def4: :: MOD_4:def 4
for o being Function of [:the carrier of (opp K),the carrier of W:],the carrier of W st o = ~ the rmult of W holds
it = VectSpStr(# the carrier of W,the add of W,the Zero of W,o #);
existence
ex b1 being strict VectSpStr of opp K st
for o being Function of [:the carrier of (opp K),the carrier of W:],the carrier of W st o = ~ the rmult of W holds
b1 = VectSpStr(# the carrier of W,the add of W,the Zero of W,o #)
proof end;
uniqueness
for b1, b2 being strict VectSpStr of opp K st ( for o being Function of [:the carrier of (opp K),the carrier of W:],the carrier of W st o = ~ the rmult of W holds
b1 = VectSpStr(# the carrier of W,the add of W,the Zero of W,o #) ) & ( for o being Function of [:the carrier of (opp K),the carrier of W:],the carrier of W st o = ~ the rmult of W holds
b2 = VectSpStr(# the carrier of W,the add of W,the Zero of W,o #) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def4 defines opp MOD_4:def 4 :
for K being non empty doubleLoopStr
for W being non empty RightModStr of K
for b3 being strict VectSpStr of opp K holds
( b3 = opp W iff for o being Function of [:the carrier of (opp K),the carrier of W:],the carrier of W st o = ~ the rmult of W holds
b3 = VectSpStr(# the carrier of W,the add of W,the Zero of W,o #) );

registration
let K be non empty doubleLoopStr ;
let W be non empty RightModStr of K;
cluster opp a2 -> non empty strict ;
coherence
not opp W is empty
proof end;
end;

theorem Th11: :: MOD_4:11
canceled;

theorem Th12: :: MOD_4:12
for K being non empty doubleLoopStr
for W being non empty RightModStr of K holds
( LoopStr(# the carrier of (opp W),the add of (opp W),the Zero of (opp W) #) = LoopStr(# the carrier of W,the add of W,the Zero of W #) & ( for x being set holds
( x is Vector of W iff x is Vector of (opp W) ) ) )
proof end;

definition
let K be non empty doubleLoopStr ;
let W be non empty RightModStr of K;
let o be Function of [:the carrier of W,the carrier of K:],the carrier of W;
func opp c3 -> Function of [:the carrier of (opp a1),the carrier of (opp a2):],the carrier of (opp a2) equals :: MOD_4:def 5
~ o;
coherence
~ o is Function of [:the carrier of (opp K),the carrier of (opp W):],the carrier of (opp W)
proof end;
end;

:: deftheorem Def5 defines opp MOD_4:def 5 :
for K being non empty doubleLoopStr
for W being non empty RightModStr of K
for o being Function of [:the carrier of W,the carrier of K:],the carrier of W holds opp o = ~ o;

theorem Th13: :: MOD_4:13
for K being non empty doubleLoopStr
for W being non empty RightModStr of K holds the lmult of (opp W) = opp the rmult of W
proof end;

theorem Th14: :: MOD_4:14
canceled;

theorem Th15: :: MOD_4:15
for K being non empty doubleLoopStr
for V being non empty VectSpStr of K
for o being Function of [:the carrier of K,the carrier of V:],the carrier of V holds opp (opp o) = o by FUNCT_4:54;

theorem Th16: :: MOD_4:16
for K being non empty doubleLoopStr
for V being non empty VectSpStr of K
for o being Function of [:the carrier of K,the carrier of V:],the carrier of V
for x being Scalar of K
for y being Scalar of (opp K)
for v being Vector of V
for w being Vector of (opp V) st x = y & v = w holds
(opp o) . w,y = o . x,v by ;

theorem Th17: :: MOD_4:17
for K, L being Ring
for V being non empty VectSpStr of K
for W being non empty RightModStr of L
for x being Scalar of K
for y being Scalar of L
for v being Vector of V
for w being Vector of W st L = opp K & W = opp V & x = y & v = w holds
w * y = x * v
proof end;

theorem Th18: :: MOD_4:18
for K, L being Ring
for V being non empty VectSpStr of K
for W being non empty RightModStr of L
for v1, v2 being Vector of V
for w1, w2 being Vector of W st L = opp K & W = opp V & v1 = w1 & v2 = w2 holds
w1 + w2 = v1 + v2
proof end;

theorem Th19: :: MOD_4:19
for K being non empty doubleLoopStr
for W being non empty RightModStr of K
for o being Function of [:the carrier of W,the carrier of K:],the carrier of W holds opp (opp o) = o by FUNCT_4:54;

theorem Th20: :: MOD_4:20
for K being non empty doubleLoopStr
for W being non empty RightModStr of K
for o being Function of [:the carrier of W,the carrier of K:],the carrier of W
for x being Scalar of K
for y being Scalar of (opp K)
for v being Vector of W
for w being Vector of (opp W) st x = y & v = w holds
(opp o) . y,w = o . v,x by ;

theorem Th21: :: MOD_4:21
for K, L being Ring
for V being non empty VectSpStr of K
for W being non empty RightModStr of L
for x being Scalar of K
for y being Scalar of L
for v being Vector of V
for w being Vector of W st K = opp L & V = opp W & x = y & v = w holds
w * y = x * v
proof end;

theorem Th22: :: MOD_4:22
for K, L being Ring
for V being non empty VectSpStr of K
for W being non empty RightModStr of L
for v1, v2 being Vector of V
for w1, w2 being Vector of W st K = opp L & V = opp W & v1 = w1 & v2 = w2 holds
w1 + w2 = v1 + v2
proof end;

theorem Th23: :: MOD_4:23
for K being non empty strict doubleLoopStr
for V being non empty VectSpStr of K holds opp (opp V) = VectSpStr(# the carrier of V,the add of V,the Zero of V,the lmult of V #)
proof end;

theorem Th24: :: MOD_4:24
for K being non empty strict doubleLoopStr
for W being non empty RightModStr of K holds opp (opp W) = RightModStr(# the carrier of W,the add of W,the Zero of W,the rmult of W #)
proof end;

theorem Th25: :: MOD_4:25
for K being Ring
for V being LeftMod of K holds opp V is strict RightMod of opp K
proof end;

registration
let K be Ring;
let V be LeftMod of K;
cluster opp a2 -> non empty Abelian add-associative right_zeroed right_complementable strict RightMod-like ;
coherence
( opp V is Abelian & opp V is add-associative & opp V is right_zeroed & opp V is right_complementable & opp V is RightMod-like )
by ;
end;

theorem Th26: :: MOD_4:26
for K being Ring
for W being RightMod of K holds opp W is strict LeftMod of opp K
proof end;

registration
let K be Ring;
let W be RightMod of K;
cluster opp a2 -> non empty Abelian add-associative right_zeroed right_complementable strict VectSp-like ;
coherence
( opp W is Abelian & opp W is add-associative & opp W is right_zeroed & opp W is right_complementable & opp W is VectSp-like )
by Lemma28;
end;

definition
let K be non empty doubleLoopStr , L be non empty doubleLoopStr ;
let IT be Function of K,L;
attr a3 is antilinear means :Def6: :: MOD_4:def 6
( ( for x, y being Scalar of K holds IT . (x + y) = (IT . x) + (IT . y) ) & ( for x, y being Scalar of K holds IT . (x * y) = (IT . y) * (IT . x) ) & IT . (1. K) = 1. L );
end;

:: deftheorem Def6 defines antilinear MOD_4:def 6 :
for K, L being non empty doubleLoopStr
for IT being Function of K,L holds
( IT is antilinear iff ( ( for x, y being Scalar of K holds IT . (x + y) = (IT . x) + (IT . y) ) & ( for x, y being Scalar of K holds IT . (x * y) = (IT . y) * (IT . x) ) & IT . (1. K) = 1. L ) );

definition
let K be non empty doubleLoopStr , L be non empty doubleLoopStr ;
let IT be Function of K,L;
attr a3 is monomorphism means :Def7: :: MOD_4:def 7
( IT is linear & IT is one-to-one );
attr a3 is antimonomorphism means :Def8: :: MOD_4:def 8
( IT is antilinear & IT is one-to-one );
end;

:: deftheorem Def7 defines monomorphism MOD_4:def 7 :
for K, L being non empty doubleLoopStr
for IT being Function of K,L holds
( IT is monomorphism iff ( IT is linear & IT is one-to-one ) );

:: deftheorem Def8 defines antimonomorphism MOD_4:def 8 :
for K, L being non empty doubleLoopStr
for IT being Function of K,L holds
( IT is antimonomorphism iff ( IT is antilinear & IT is one-to-one ) );

definition
let K be non empty doubleLoopStr , L be non empty doubleLoopStr ;
let IT be Function of K,L;
attr a3 is epimorphism means :Def9: :: MOD_4:def 9
( IT is linear & rng IT = the carrier of L );
attr a3 is antiepimorphism means :Def10: :: MOD_4:def 10
( IT is antilinear & rng IT = the carrier of L );
end;

:: deftheorem Def9 defines epimorphism MOD_4:def 9 :
for K, L being non empty doubleLoopStr
for IT being Function of K,L holds
( IT is epimorphism iff ( IT is linear & rng IT = the carrier of L ) );

:: deftheorem Def10 defines antiepimorphism MOD_4:def 10 :
for K, L being non empty doubleLoopStr
for IT being Function of K,L holds
( IT is antiepimorphism iff ( IT is antilinear & rng IT = the carrier of L ) );

definition
let K be non empty doubleLoopStr , L be non empty doubleLoopStr ;
let IT be Function of K,L;
attr a3 is isomorphism means :Def11: :: MOD_4:def 11
( IT is monomorphism & rng IT = the carrier of L );
attr a3 is antiisomorphism means :Def12: :: MOD_4:def 12
( IT is antimonomorphism & rng IT = the carrier of L );
end;

:: deftheorem Def11 defines isomorphism MOD_4:def 11 :
for K, L being non empty doubleLoopStr
for IT being Function of K,L holds
( IT is isomorphism iff ( IT is monomorphism & rng IT = the carrier of L ) );

:: deftheorem Def12 defines antiisomorphism MOD_4:def 12 :
for K, L being non empty doubleLoopStr
for IT being Function of K,L holds
( IT is antiisomorphism iff ( IT is antimonomorphism & rng IT = the carrier of L ) );

definition
let K be non empty doubleLoopStr ;
let IT be Function of K,K;
attr a2 is endomorphism means :Def13: :: MOD_4:def 13
IT is linear;
attr a2 is antiendomorphism means :Def14: :: MOD_4:def 14
IT is antilinear;
attr a2 is automorphism means :Def15: :: MOD_4:def 15
IT is isomorphism;
attr a2 is antiautomorphism means :Def16: :: MOD_4:def 16
IT is antiisomorphism;
end;

:: deftheorem Def13 defines endomorphism MOD_4:def 13 :
for K being non empty doubleLoopStr
for IT being Function of K,K holds
( IT is endomorphism iff IT is linear );

:: deftheorem Def14 defines antiendomorphism MOD_4:def 14 :
for K being non empty doubleLoopStr
for IT being Function of K,K holds
( IT is antiendomorphism iff IT is antilinear );

:: deftheorem Def15 defines automorphism MOD_4:def 15 :
for K being non empty doubleLoopStr
for IT being Function of K,K holds
( IT is automorphism iff IT is isomorphism );

:: deftheorem Def16 defines antiautomorphism MOD_4:def 16 :
for K being non empty doubleLoopStr
for IT being Function of K,K holds
( IT is antiautomorphism iff IT is antiisomorphism );

theorem Th27: :: MOD_4:27
for K being non empty doubleLoopStr
for J being Function of K,K holds
( J is automorphism iff ( ( for x, y being Scalar of K holds J . (x + y) = (J . x) + (J . y) ) & ( for x, y being Scalar of K holds J . (x * y) = (J . x) * (J . y) ) & J . (1. K) = 1. K & J is one-to-one & rng J = the carrier of K ) )
proof end;

theorem Th28: :: MOD_4:28
for K being non empty doubleLoopStr
for J being Function of K,K holds
( J is antiautomorphism iff ( ( for x, y being Scalar of K holds J . (x + y) = (J . x) + (J . y) ) & ( for x, y being Scalar of K holds J . (x * y) = (J . y) * (J . x) ) & J . (1. K) = 1. K & J is one-to-one & rng J = the carrier of K ) )
proof end;

Lemma104: for K being non empty doubleLoopStr holds
( ( for x, y being Scalar of K holds (id K) . (x + y) = ((id K) . x) + ((id K) . y) ) & ( for x, y being Scalar of K holds (id K) . (x * y) = ((id K) . x) * ((id K) . y) ) & (id K) . (1. K) = 1. K & id K is one-to-one & rng (id K) = the carrier of K )
proof end;

theorem Th29: :: MOD_4:29
for K being non empty doubleLoopStr holds id K is automorphism
proof end;

Lemma105: for K, L being Ring
for J being Function of K,L st J is linear holds
J . (0. K) = 0. L
proof end;

Lemma106: for L, K being Ring
for J being Function of K,L
for x being Scalar of K st J is linear holds
J . (- x) = - (J . x)
proof end;

Lemma107: for L, K being Ring
for J being Function of K,L
for x, y being Scalar of K st J is linear holds
J . (x - y) = (J . x) - (J . y)
proof end;

theorem Th30: :: MOD_4:30
for K, L being Ring
for J being Function of K,L
for x, y being Scalar of K st J is linear holds
( J . (0. K) = 0. L & J . (- x) = - (J . x) & J . (x - y) = (J . x) - (J . y) ) by , , Th3;

Lemma108: for K, L being Ring
for J being Function of K,L st J is antilinear holds
J . (0. K) = 0. L
proof end;

Lemma109: for L, K being Ring
for J being Function of K,L
for x being Scalar of K st J is antilinear holds
J . (- x) = - (J . x)
proof end;

Lemma110: for L, K being Ring
for J being Function of K,L
for x, y being Scalar of K st J is antilinear holds
J . (x - y) = (J . x) - (J . y)
proof end;

theorem Th31: :: MOD_4:31
for K, L being Ring
for J being Function of K,L
for x, y being Scalar of K st J is antilinear holds
( J . (0. K) = 0. L & J . (- x) = - (J . x) & J . (x - y) = (J . x) - (J . y) ) by Lemma45, , ;

theorem Th32: :: MOD_4:32
for K being Ring holds
( id K is antiautomorphism iff K is comRing )
proof end;

theorem Th33: :: MOD_4:33
for K being Skew-Field holds
( id K is antiautomorphism iff K is Field )
proof end;

definition
let K be non empty doubleLoopStr , L be non empty doubleLoopStr ;
let J be Function of K,L;
func opp c3 -> Function of a1,(opp a2) equals :: MOD_4:def 17
J;
coherence
J is Function of K,(opp L)
;
end;

:: deftheorem Def17 defines opp MOD_4:def 17 :
for K, L being non empty doubleLoopStr
for J being Function of K,L holds opp J = J;

theorem Th34: :: MOD_4:34
for K, L being non empty add-associative right_zeroed right_complementable doubleLoopStr
for J being Function of K,L holds opp (opp J) = J ;

theorem Th35: :: MOD_4:35
for K being non empty add-associative right_zeroed right_complementable doubleLoopStr
for L being non empty add-associative right_zeroed right_complementable unital doubleLoopStr
for J being Function of K,L holds
( J is linear iff opp J is antilinear )
proof end;

theorem Th36: :: MOD_4:36
for K being non empty add-associative right_zeroed right_complementable doubleLoopStr
for L being non empty add-associative right_zeroed right_complementable unital doubleLoopStr
for J being Function of K,L holds
( J is antilinear iff opp J is linear )
proof end;

theorem Th37: :: MOD_4:37
for K being non empty add-associative right_zeroed right_complementable doubleLoopStr
for L being non empty add-associative right_zeroed right_complementable unital doubleLoopStr
for J being Function of K,L holds
( J is monomorphism iff opp J is antimonomorphism )
proof end;

theorem Th38: :: MOD_4:38
for K being non empty add-associative right_zeroed right_complementable doubleLoopStr
for L being non empty add-associative right_zeroed right_complementable unital doubleLoopStr
for J being Function of K,L holds
( J is antimonomorphism iff opp J is monomorphism )
proof end;

theorem Th39: :: MOD_4:39
for K being non empty add-associative right_zeroed right_complementable doubleLoopStr
for L being non empty add-associative right_zeroed right_complementable unital doubleLoopStr
for J being Function of K,L holds
( J is epimorphism iff opp J is antiepimorphism )
proof end;

theorem Th40: :: MOD_4:40
for K being non empty add-associative right_zeroed right_complementable doubleLoopStr
for L being non empty add-associative right_zeroed right_complementable unital doubleLoopStr
for J being Function of K,L holds
( J is antiepimorphism iff opp J is epimorphism )
proof end;

theorem Th41: :: MOD_4:41
for K being non empty add-associative right_zeroed right_complementable doubleLoopStr
for L being non empty add-associative right_zeroed right_complementable unital doubleLoopStr
for J being Function of K,L holds
( J is isomorphism iff opp J is antiisomorphism )
proof end;

theorem Th42: :: MOD_4:42
for K being non empty add-associative right_zeroed right_complementable doubleLoopStr
for L being non empty add-associative right_zeroed right_complementable unital doubleLoopStr
for J being Function of K,L holds
( J is antiisomorphism iff opp J is isomorphism )
proof end;

theorem Th43: :: MOD_4:43
for K being non empty add-associative right_zeroed right_complementable unital doubleLoopStr
for J being Function of K,K holds
( J is endomorphism iff opp J is antilinear )
proof end;

theorem Th44: :: MOD_4:44
for K being non empty add-associative right_zeroed right_complementable unital doubleLoopStr
for J being Function of K,K holds
( J is antiendomorphism iff opp J is linear )
proof end;

theorem Th45: :: MOD_4:45
for K being non empty add-associative right_zeroed right_complementable unital doubleLoopStr
for J being Function of K,K holds
( J is automorphism iff opp J is antiisomorphism )
proof end;

theorem Th46: :: MOD_4:46
for K being non empty add-associative right_zeroed right_complementable unital doubleLoopStr
for J being Function of K,K holds
( J is antiautomorphism iff opp J is isomorphism )
proof end;

Lemma120: for G, H being AddGroup
for x, y being Element of G holds (ZeroMap G,H) . (x + y) = ((ZeroMap G,H) . x) + ((ZeroMap G,H) . y)
proof end;

definition
let G be AddGroup;
let H be AddGroup;
mode Homomorphism of c1,c2 -> Function of a1,a2 means :Def18: :: MOD_4:def 18
for x, y being Element of G holds it . (x + y) = (it . x) + (it . y);
existence
ex b1 being Function of G,H st
for x, y being Element of G holds b1 . (x + y) = (b1 . x) + (b1 . y)
proof end;
end;

:: deftheorem Def18 defines Homomorphism MOD_4:def 18 :
for G, H being AddGroup
for b3 being Function of G,H holds
( b3 is Homomorphism of G,H iff for x, y being Element of G holds b3 . (x + y) = (b3 . x) + (b3 . y) );

definition
let G be AddGroup;
let H be AddGroup;
redefine func ZeroMap as ZeroMap c1,c2 -> Homomorphism of a1,a2;
coherence
ZeroMap G,H is Homomorphism of G,H
proof end;
end;

definition
let G be AddGroup;
let H be AddGroup;
let IT be Homomorphism of G,H;
attr a3 is monomorphism means :: MOD_4:def 19
IT is one-to-one;
end;

:: deftheorem Def19 defines monomorphism MOD_4:def 19 :
for G, H being AddGroup
for IT being Homomorphism of G,H holds
( IT is monomorphism iff IT is one-to-one );

definition
let G be AddGroup;
let H be AddGroup;
let IT be Homomorphism of G,H;
attr a3 is epimorphism means :: MOD_4:def 20
rng IT = the carrier of H;
end;

:: deftheorem Def20 defines epimorphism MOD_4:def 20 :
for G, H being AddGroup
for IT being Homomorphism of G,H holds
( IT is epimorphism iff rng IT = the carrier of H );

definition
let G be AddGroup;
let H be AddGroup;
let IT be Homomorphism of G,H;
attr a3 is isomorphism means :Def21: :: MOD_4:def 21
( IT is one-to-one & rng IT = the carrier of H );
end;

:: deftheorem Def21 defines isomorphism MOD_4:def 21 :
for G, H being AddGroup
for IT being Homomorphism of G,H holds
( IT is isomorphism iff ( IT is one-to-one & rng IT = the carrier of H ) );

definition
let G be AddGroup;
mode Endomorphism of a1 is Homomorphism of a1,a1;
end;

Lemma123: for G being AddGroup
for x being Element of G holds (id G) . x = x
by FUNCT_1:35;

Lemma124: for G being AddGroup holds
( ( for x, y being Element of G holds (id G) . (x + y) = ((id G) . x) + ((id G) . y) ) & id G is one-to-one & rng (id G) = the carrier of G )
proof end;

registration
let G be AddGroup;
cluster isomorphism Homomorphism of a1,a1;
existence
ex b1 being Endomorphism of G st b1 is isomorphism
proof end;
end;

definition
let G be AddGroup;
mode Automorphism of a1 is isomorphism Endomorphism of a1;
end;

definition
let G be AddGroup;
redefine func id as id c1 -> Automorphism of a1;
coherence
id G is Automorphism of G
proof end;
end;

Lemma125: for G, H being AddGroup
for f being Homomorphism of G,H holds f . (0. G) = 0. H
proof end;

Lemma126: for H, G being AddGroup
for f being Homomorphism of G,H
for x being Element of G holds f . (- x) = - (f . x)
proof end;

Lemma127: for H, G being AddGroup
for f being Homomorphism of G,H
for x, y being Element of G holds f . (x - y) = (f . x) - (f . y)
proof end;

theorem Th47: :: MOD_4:47
canceled;

theorem Th48: :: MOD_4:48
for G, H being AddGroup
for f being Homomorphism of G,H
for x, y being Element of G holds
( f . (0. G) = 0. H & f . (- x) = - (f . x) & f . (x - y) = (f . x) - (f . y) ) by , Th10, ;

theorem Th49: :: MOD_4:49
for H, G being AbGroup
for f being Homomorphism of G,H
for x, y being Element of G holds f . (x - y) = (f . x) - (f . y)
proof end;

Lemma128: for K, L being Ring
for V being LeftMod of K
for W being LeftMod of L
for x, y being Vector of V holds (ZeroMap V,W) . (x + y) = ((ZeroMap V,W) . x) + ((ZeroMap V,W) . y)
proof end;

Lemma129: for L, K being Ring
for J being Function of K,L
for V being LeftMod of K
for W being LeftMod of L
for a being Scalar of K
for x being Vector of V holds (ZeroMap V,W) . (a * x) = (J . a) * ((ZeroMap V,W) . x)
proof end;

definition
let K be Ring;
let L be Ring;
let J be Function of K,L;
let V be LeftMod of K;
let W be LeftMod of L;
canceled;
mode Homomorphism of c3,c4,c5 -> Function of a4,a5 means :Def23: :: MOD_4:def 23
( ( for x, y being Vector of V holds it . (x + y) = (it . x) + (it . y) ) & ( for a being Scalar of K
for x being Vector of V holds it . (a * x) = (J . a) * (it . x) ) );
existence
ex b1 being Function of V,W st
( ( for x, y being Vector of V holds b1 . (x + y) = (b1 . x) + (b1 . y) ) & ( for a being Scalar of K
for x being Vector of V holds b1 . (a * x) = (J . a) * (b1 . x) ) )
proof end;
end;

:: deftheorem Def22 MOD_4:def 22 :
canceled;

:: deftheorem Def23 defines Homomorphism MOD_4:def 23 :
for K, L being Ring
for J being Function of K,L
for V being LeftMod of K
for W being LeftMod of L
for b6 being Function of V,W holds
( b6 is Homomorphism of J,V,W iff ( ( for x, y being Vector of V holds b6 . (x + y) = (b6 . x) + (b6 . y) ) & ( for a being Scalar of K
for x being Vector of V holds b6 . (a * x) = (J . a) * (b6 . x) ) ) );

theorem Th50: :: MOD_4:50
for K, L being Ring
for J being Function of K,L
for V being LeftMod of K
for W being LeftMod of L holds ZeroMap V,W is Homomorphism of J,V,W
proof end;

definition
let K be Ring;
let L be Ring;
let J be Function of K,L;
let V be LeftMod of K;
let W be LeftMod of L;
let f be Homomorphism of J,V,W;
pred c6 is_monomorphism_wrp c3 means :: MOD_4:def 24
f is one-to-one;
pred c6 is_epimorphism_wrp c3 means :: MOD_4:def 25
rng f = the carrier of W;
pred c6 is_isomorphism_wrp c3 means :: MOD_4:def 26
( f is one-to-one & rng f = the carrier of W );
end;

:: deftheorem Def24 defines is_monomorphism_wrp MOD_4:def 24 :
for K, L being Ring
for J being Function of K,L
for V being LeftMod of K
for W being LeftMod of L
for f being Homomorphism of J,V,W holds
( f is_monomorphism_wrp J iff f is one-to-one );

:: deftheorem Def25 defines is_epimorphism_wrp MOD_4:def 25 :
for K, L being Ring
for J being Function of K,L
for V being LeftMod of K
for W being LeftMod of L
for f being Homomorphism of J,V,W holds
( f is_epimorphism_wrp J iff rng f = the carrier of W );

:: deftheorem Def26 defines is_isomorphism_wrp MOD_4:def 26 :
for K, L being Ring
for J being Function of K,L
for V being LeftMod of K
for W being LeftMod of L
for f being Homomorphism of J,V,W holds
( f is_isomorphism_wrp J iff ( f is one-to-one & rng f = the carrier of W ) );

definition
let K be Ring;
let J be Function of K,K;
let V be LeftMod of K;
mode Endomorphism of a2,a3 is Homomorphism of a2,a3,a3;
end;

definition
let K be Ring;
let J be Function of K,K;
let V be LeftMod of K;
let f be Homomorphism of J,V,V;
pred c4 is_automorphism_wrp c2 means :: MOD_4:def 27
( f is one-to-one & rng f = the carrier of V );
end;

:: deftheorem Def27 defines is_automorphism_wrp MOD_4:def 27 :
for K being Ring
for J being Function of K,K
for V being LeftMod of K
for f being Homomorphism of J,V,V holds
( f is_automorphism_wrp J iff ( f is one-to-one & rng f = the carrier of V ) );

definition
let K be Ring;
let V be LeftMod of K;
let W be LeftMod of K;
mode Homomorphism of a2,a3 is Homomorphism of id a1,a2,a3;
end;

theorem Th51: :: MOD_4:51
for K being Ring
for V, W being LeftMod of K
for f being Function of V,W holds
( f is Homomorphism of V,W iff ( ( for x, y being Vector of V holds f . (x + y) = (f . x) + (f . y) ) & ( for a being Scalar of K
for x being Vector of V holds f . (a * x) = a * (f . x) ) ) )
proof end;

definition
let K be Ring;
let V be LeftMod of K;
let W be LeftMod of K;
let IT be Homomorphism of V,W;
attr a4 is monomorphism means :: MOD_4:def 28
IT is one-to-one;
attr a4 is epimorphism means :: MOD_4:def 29
rng IT = the carrier of W;
attr a4 is isomorphism means :: MOD_4:def 30
( IT is one-to-one & rng IT = the carrier of W );
end;

:: deftheorem Def28 defines monomorphism MOD_4:def 28 :
for K being Ring
for V, W being LeftMod of K
for IT being Homomorphism of V,W holds
( IT is monomorphism iff IT is one-to-one );

:: deftheorem Def29 defines epimorphism MOD_4:def 29 :
for K being Ring
for V, W being LeftMod of K
for IT being Homomorphism of V,W holds
( IT is epimorphism iff rng IT = the carrier of W );

:: deftheorem Def30 defines isomorphism MOD_4:def 30 :
for K being Ring
for V, W being LeftMod of K
for IT being Homomorphism of V,W holds
( IT is isomorphism iff ( IT is one-to-one & rng IT = the carrier of W ) );

definition
let K be Ring;
let V be LeftMod of K;
mode Endomorphism of a2 is Homomorphism of a2,a2;
end;

definition
let K be Ring;
let V be LeftMod of K;
let IT be Endomorphism of V;
attr a3 is automorphism means :: MOD_4:def 31
( IT is one-to-one & rng IT = the carrier of V );
end;

:: deftheorem Def31 defines automorphism MOD_4:def 31 :
for K being Ring
for V being LeftMod of K
for IT being Endomorphism of V holds
( IT is automorphism iff ( IT is one-to-one & rng IT = the carrier of V ) );