Lm1:
for L being non empty multLoopStr st L is well-unital holds
1. L = 1_ L
by GROUP_1:def 4;
Lm2:
for R being non empty right_complementable Abelian add-associative right_zeroed addLoopStr
for x, y, z being Scalar of R st x + y = z holds
x = z - y
Lm3:
for R being non empty right_complementable Abelian add-associative right_zeroed addLoopStr
for x, y, z being Scalar of R st x = z - y holds
x + y = z
canceled;
deffunc H1( Ring) -> ModuleStr over $1 = ModuleStr(# the carrier of $1, the addF of $1,(0. $1), the multF of $1 #);
Lm4:
for R being Ring holds
( H1(R) is Abelian & H1(R) is add-associative & H1(R) is right_zeroed & H1(R) is right_complementable )
deffunc H2( Ring) -> RightModStr over $1 = RightModStr(# the carrier of $1, the addF of $1,(0. $1), the multF of $1 #);
Lm5:
for R being Ring holds
( H2(R) is Abelian & H2(R) is add-associative & H2(R) is right_zeroed & H2(R) is right_complementable )
deffunc H3( Ring, Ring) -> BiModStr over $1,$2 = BiModStr(# {0},op2,op0,(pr2 ( the carrier of $1,{0})),(pr1 ({0}, the carrier of $2)) #);
Lm6:
for R1, R2 being Ring holds
( H3(R1,R2) is Abelian & H3(R1,R2) is add-associative & H3(R1,R2) is right_zeroed & H3(R1,R2) is right_complementable )
Lm7:
for R being Ring holds
( LeftModule R is vector-distributive & LeftModule R is scalar-distributive & LeftModule R is scalar-associative & LeftModule R is scalar-unital )
by Th23;
Lm8:
for R being Ring holds RightModule R is RightMod-like
by Th24;
Lm9:
for R1, R2 being Ring
for x, y being Scalar of R1
for p, q being Scalar of R2
for v, w being Vector of (BiModule (R1,R2)) holds
( x * (v + w) = (x * v) + (x * w) & (x + y) * v = (x * v) + (y * v) & (x * y) * v = x * (y * v) & (1_ R1) * v = v & (v + w) * p = (v * p) + (w * p) & v * (p + q) = (v * p) + (v * q) & v * (q * p) = (v * q) * p & v * (1_ R2) = v & x * (v * p) = (x * v) * p )
registration
let R1,
R2 be
Ring;
coherence
( BiModule (R1,R2) is Abelian & BiModule (R1,R2) is add-associative & BiModule (R1,R2) is right_zeroed & BiModule (R1,R2) is right_complementable & BiModule (R1,R2) is RightMod-like & BiModule (R1,R2) is vector-distributive & BiModule (R1,R2) is scalar-distributive & BiModule (R1,R2) is scalar-associative & BiModule (R1,R2) is scalar-unital & BiModule (R1,R2) is BiMod-like )
by Th26;
end;