begin
Lm1:
for g being quaternion number ex r, s, t, u being Element of REAL st g = [*r,s,t,u*]
by QUATERNI:7;
Lm2:
for a, b, c, d being real number holds (((a ^2) + (b ^2)) + (c ^2)) + (d ^2) >= 0
by QUATERNI:74;
Lm3:
for a, b, c, d being real number st (((a ^2) + (b ^2)) + (c ^2)) + (d ^2) = 0 holds
( a = 0 & b = 0 & c = 0 & d = 0 )
by QUATERNI:96;
theorem Th1:
theorem Th2:
theorem Th3:
theorem Th4:
theorem
for
x1,
x2,
x3,
x4,
y1,
y2,
y3,
y4 being
Element of
REAL holds
[*x1,x2,x3,x4*] - [*y1,y2,y3,y4*] = [*(x1 - y1),(x2 - y2),(x3 - y3),(x4 - y4)*]
theorem
theorem Th7:
theorem
theorem Th9:
Lm4:
for r being quaternion number st |.r.| = 0 holds
r = 0
by QUATERNI:66;
theorem Th10:
theorem Th11:
theorem
theorem
theorem Th14:
theorem Th15:
theorem Th16:
theorem Th17:
theorem Th18:
theorem Th19:
theorem Th20:
theorem Th21:
theorem Th22:
theorem Th23:
theorem Th24:
theorem Th25:
theorem
definition
let q,
r be
quaternion number ;
consider q0,
q1,
q2,
q3 being
Element of
REAL such that A1:
q = [*q0,q1,q2,q3*]
by Lm1;
consider r0,
r1,
r2,
r3 being
Element of
REAL such that A2:
r = [*r0,r1,r2,r3*]
by Lm1;
func q / r -> set means :
Def1:
ex
q0,
q1,
q2,
q3,
r0,
r1,
r2,
r3 being
Element of
REAL st
(
q = [*q0,q1,q2,q3*] &
r = [*r0,r1,r2,r3*] &
it = [*(((((r0 * q0) + (r1 * q1)) + (r2 * q2)) + (r3 * q3)) / (|.r.| ^2)),(((((r0 * q1) - (r1 * q0)) - (r2 * q3)) + (r3 * q2)) / (|.r.| ^2)),(((((r0 * q2) + (r1 * q3)) - (r2 * q0)) - (r3 * q1)) / (|.r.| ^2)),(((((r0 * q3) - (r1 * q2)) + (r2 * q1)) - (r3 * q0)) / (|.r.| ^2))*] );
existence
ex b1 being set ex q0, q1, q2, q3, r0, r1, r2, r3 being Element of REAL st
( q = [*q0,q1,q2,q3*] & r = [*r0,r1,r2,r3*] & b1 = [*(((((r0 * q0) + (r1 * q1)) + (r2 * q2)) + (r3 * q3)) / (|.r.| ^2)),(((((r0 * q1) - (r1 * q0)) - (r2 * q3)) + (r3 * q2)) / (|.r.| ^2)),(((((r0 * q2) + (r1 * q3)) - (r2 * q0)) - (r3 * q1)) / (|.r.| ^2)),(((((r0 * q3) - (r1 * q2)) + (r2 * q1)) - (r3 * q0)) / (|.r.| ^2))*] )
uniqueness
for b1, b2 being set st ex q0, q1, q2, q3, r0, r1, r2, r3 being Element of REAL st
( q = [*q0,q1,q2,q3*] & r = [*r0,r1,r2,r3*] & b1 = [*(((((r0 * q0) + (r1 * q1)) + (r2 * q2)) + (r3 * q3)) / (|.r.| ^2)),(((((r0 * q1) - (r1 * q0)) - (r2 * q3)) + (r3 * q2)) / (|.r.| ^2)),(((((r0 * q2) + (r1 * q3)) - (r2 * q0)) - (r3 * q1)) / (|.r.| ^2)),(((((r0 * q3) - (r1 * q2)) + (r2 * q1)) - (r3 * q0)) / (|.r.| ^2))*] ) & ex q0, q1, q2, q3, r0, r1, r2, r3 being Element of REAL st
( q = [*q0,q1,q2,q3*] & r = [*r0,r1,r2,r3*] & b2 = [*(((((r0 * q0) + (r1 * q1)) + (r2 * q2)) + (r3 * q3)) / (|.r.| ^2)),(((((r0 * q1) - (r1 * q0)) - (r2 * q3)) + (r3 * q2)) / (|.r.| ^2)),(((((r0 * q2) + (r1 * q3)) - (r2 * q0)) - (r3 * q1)) / (|.r.| ^2)),(((((r0 * q3) - (r1 * q2)) + (r2 * q1)) - (r3 * q0)) / (|.r.| ^2))*] ) holds
b1 = b2
end;
:: deftheorem Def1 defines / QUATERN2:def 1 :
for q, r being quaternion number
for b3 being set holds
( b3 = q / r iff ex q0, q1, q2, q3, r0, r1, r2, r3 being Element of REAL st
( q = [*q0,q1,q2,q3*] & r = [*r0,r1,r2,r3*] & b3 = [*(((((r0 * q0) + (r1 * q1)) + (r2 * q2)) + (r3 * q3)) / (|.r.| ^2)),(((((r0 * q1) - (r1 * q0)) - (r2 * q3)) + (r3 * q2)) / (|.r.| ^2)),(((((r0 * q2) + (r1 * q3)) - (r2 * q0)) - (r3 * q1)) / (|.r.| ^2)),(((((r0 * q3) - (r1 * q2)) + (r2 * q1)) - (r3 * q0)) / (|.r.| ^2))*] ) );
theorem
:: deftheorem defines " QUATERN2:def 2 :
for c being quaternion number holds c " = 1q / c;
theorem
theorem
theorem
theorem Th31:
theorem
theorem Th33:
theorem
definition
func compquaternion -> UnOp of
QUATERNION means
for
c being
Element of
QUATERNION holds
it . c = - c;
existence
ex b1 being UnOp of QUATERNION st
for c being Element of QUATERNION holds b1 . c = - c
from FUNCT_2:sch 4();
uniqueness
for b1, b2 being UnOp of QUATERNION st ( for c being Element of QUATERNION holds b1 . c = - c ) & ( for c being Element of QUATERNION holds b2 . c = - c ) holds
b1 = b2
from BINOP_2:sch 1();
func addquaternion -> BinOp of
QUATERNION means :
Def4:
for
c1,
c2 being
Element of
QUATERNION holds
it . (
c1,
c2)
= c1 + c2;
existence
ex b1 being BinOp of QUATERNION st
for c1, c2 being Element of QUATERNION holds b1 . (c1,c2) = c1 + c2
from BINOP_1:sch 4();
uniqueness
for b1, b2 being BinOp of QUATERNION st ( for c1, c2 being Element of QUATERNION holds b1 . (c1,c2) = c1 + c2 ) & ( for c1, c2 being Element of QUATERNION holds b2 . (c1,c2) = c1 + c2 ) holds
b1 = b2
from BINOP_2:sch 2();
func diffquaternion -> BinOp of
QUATERNION means
for
c1,
c2 being
Element of
QUATERNION holds
it . (
c1,
c2)
= c1 - c2;
existence
ex b1 being BinOp of QUATERNION st
for c1, c2 being Element of QUATERNION holds b1 . (c1,c2) = c1 - c2
from BINOP_1:sch 4();
uniqueness
for b1, b2 being BinOp of QUATERNION st ( for c1, c2 being Element of QUATERNION holds b1 . (c1,c2) = c1 - c2 ) & ( for c1, c2 being Element of QUATERNION holds b2 . (c1,c2) = c1 - c2 ) holds
b1 = b2
from BINOP_2:sch 2();
func multquaternion -> BinOp of
QUATERNION means :
Def6:
for
c1,
c2 being
Element of
QUATERNION holds
it . (
c1,
c2)
= c1 * c2;
existence
ex b1 being BinOp of QUATERNION st
for c1, c2 being Element of QUATERNION holds b1 . (c1,c2) = c1 * c2
from BINOP_1:sch 4();
uniqueness
for b1, b2 being BinOp of QUATERNION st ( for c1, c2 being Element of QUATERNION holds b1 . (c1,c2) = c1 * c2 ) & ( for c1, c2 being Element of QUATERNION holds b2 . (c1,c2) = c1 * c2 ) holds
b1 = b2
from BINOP_2:sch 2();
func divquaternion -> BinOp of
QUATERNION means
for
c1,
c2 being
Element of
QUATERNION holds
it . (
c1,
c2)
= c1 / c2;
existence
ex b1 being BinOp of QUATERNION st
for c1, c2 being Element of QUATERNION holds b1 . (c1,c2) = c1 / c2
from BINOP_1:sch 4();
uniqueness
for b1, b2 being BinOp of QUATERNION st ( for c1, c2 being Element of QUATERNION holds b1 . (c1,c2) = c1 / c2 ) & ( for c1, c2 being Element of QUATERNION holds b2 . (c1,c2) = c1 / c2 ) holds
b1 = b2
from BINOP_2:sch 2();
func invquaternion -> UnOp of
QUATERNION means
for
c being
Element of
QUATERNION holds
it . c = c " ;
existence
ex b1 being UnOp of QUATERNION st
for c being Element of QUATERNION holds b1 . c = c "
from FUNCT_2:sch 4();
uniqueness
for b1, b2 being UnOp of QUATERNION st ( for c being Element of QUATERNION holds b1 . c = c " ) & ( for c being Element of QUATERNION holds b2 . c = c " ) holds
b1 = b2
from BINOP_2:sch 1();
end;
:: deftheorem defines compquaternion QUATERN2:def 3 :
for b1 being UnOp of QUATERNION holds
( b1 = compquaternion iff for c being Element of QUATERNION holds b1 . c = - c );
:: deftheorem Def4 defines addquaternion QUATERN2:def 4 :
for b1 being BinOp of QUATERNION holds
( b1 = addquaternion iff for c1, c2 being Element of QUATERNION holds b1 . (c1,c2) = c1 + c2 );
:: deftheorem defines diffquaternion QUATERN2:def 5 :
for b1 being BinOp of QUATERNION holds
( b1 = diffquaternion iff for c1, c2 being Element of QUATERNION holds b1 . (c1,c2) = c1 - c2 );
:: deftheorem Def6 defines multquaternion QUATERN2:def 6 :
for b1 being BinOp of QUATERNION holds
( b1 = multquaternion iff for c1, c2 being Element of QUATERNION holds b1 . (c1,c2) = c1 * c2 );
:: deftheorem defines divquaternion QUATERN2:def 7 :
for b1 being BinOp of QUATERNION holds
( b1 = divquaternion iff for c1, c2 being Element of QUATERNION holds b1 . (c1,c2) = c1 / c2 );
:: deftheorem defines invquaternion QUATERN2:def 8 :
for b1 being UnOp of QUATERNION holds
( b1 = invquaternion iff for c being Element of QUATERNION holds b1 . c = c " );
:: deftheorem Def9 defines G_Quaternion QUATERN2:def 9 :
for b1 being strict addLoopStr holds
( b1 = G_Quaternion iff ( the carrier of b1 = QUATERNION & the addF of b1 = addquaternion & 0. b1 = 0q ) );
theorem Th35:
theorem
:: deftheorem Def10 defines R_Quaternion QUATERN2:def 10 :
for b1 being strict doubleLoopStr holds
( b1 = R_Quaternion iff ( the carrier of b1 = QUATERNION & the addF of b1 = addquaternion & the multF of b1 = multquaternion & 1. b1 = 1q & 0. b1 = 0q ) );
theorem
theorem
theorem Th39:
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
:: deftheorem defines .|. QUATERN2:def 11 :
for x, y being quaternion number holds x .|. y = x * (y *');
theorem Th48:
theorem Th49:
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem