begin
:: deftheorem defines @ MIDSP_1:def 1 :
for MS being non empty MidStr
for a, b being Element of MS holds a @ b = the MIDPOINT of MS . (a,b);
:: deftheorem MIDSP_1:def 2 :
canceled;
:: deftheorem defines Example MIDSP_1:def 3 :
Example = MidStr(# 1,op2 #);
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
theorem
theorem
canceled;
theorem
theorem
canceled;
theorem Th10:
:: deftheorem Def4 defines MidSp-like MIDSP_1:def 4 :
for IT being non empty MidStr holds
( IT is MidSp-like iff for a, b, c, d being Element of IT holds
( a @ a = a & a @ b = b @ a & (a @ b) @ (c @ d) = (a @ c) @ (b @ d) & ex x being Element of IT st x @ a = b ) );
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem Th15:
theorem Th16:
theorem Th17:
theorem Th18:
for
M being
MidSp for
x,
a,
x9 being
Element of
M st
x @ a = x9 @ a holds
x = x9
theorem
:: deftheorem Def5 defines @@ MIDSP_1:def 5 :
for M being MidSp
for a, b, c, d being Element of M holds
( a,b @@ c,d iff a @ d = b @ c );
theorem
canceled;
theorem Th21:
theorem Th22:
for
M being
MidSp for
a,
b,
c,
d being
Element of
M st
a,
b @@ c,
d holds
c,
d @@ a,
b
theorem Th23:
theorem Th24:
theorem Th25:
theorem Th26:
theorem Th27:
for
M being
MidSp for
a,
b,
c,
d,
d9 being
Element of
M st
a,
b @@ c,
d &
a,
b @@ c,
d9 holds
d = d9
theorem Th28:
for
M being
MidSp for
x,
y,
a,
b,
c,
d being
Element of
M st
x,
y @@ a,
b &
x,
y @@ c,
d holds
a,
b @@ c,
d
theorem Th29:
for
M being
MidSp for
a,
b,
a9,
b9,
c,
c9 being
Element of
M st
a,
b @@ a9,
b9 &
b,
c @@ b9,
c9 holds
a,
c @@ a9,
c9
definition
let M be
MidSp;
let p,
q be
Element of
[: the carrier of M, the carrier of M:];
pred p ## q means :
Def6:
p `1 ,
p `2 @@ q `1 ,
q `2 ;
reflexivity
for p being Element of [: the carrier of M, the carrier of M:] holds p `1 ,p `2 @@ p `1 ,p `2
by Th25;
symmetry
for p, q being Element of [: the carrier of M, the carrier of M:] st p `1 ,p `2 @@ q `1 ,q `2 holds
q `1 ,q `2 @@ p `1 ,p `2
by Th22;
end;
:: deftheorem Def6 defines ## MIDSP_1:def 6 :
for M being MidSp
for p, q being Element of [: the carrier of M, the carrier of M:] holds
( p ## q iff p `1 ,p `2 @@ q `1 ,q `2 );
theorem
canceled;
theorem Th31:
theorem Th32:
theorem
canceled;
theorem
canceled;
theorem Th35:
theorem
theorem
theorem
theorem Th39:
:: deftheorem defines ~ MIDSP_1:def 7 :
for M being MidSp
for p being Element of [: the carrier of M, the carrier of M:] holds p ~ = { q where q is Element of [: the carrier of M, the carrier of M:] : q ## p } ;
theorem
canceled;
theorem Th41:
theorem Th42:
theorem Th43:
theorem Th44:
theorem
:: deftheorem Def8 defines Vector MIDSP_1:def 8 :
for M being MidSp
for b2 being non empty Subset of [: the carrier of M, the carrier of M:] holds
( b2 is Vector of M iff ex p being Element of [: the carrier of M, the carrier of M:] st b2 = p ~ );
theorem
canceled;
theorem
canceled;
theorem Th48:
:: deftheorem defines ID MIDSP_1:def 9 :
for M being MidSp holds ID M = { p where p is Element of [: the carrier of M, the carrier of M:] : p `1 = p `2 } ;
theorem
canceled;
theorem Th50:
theorem Th51:
theorem Th52:
definition
let M be
MidSp;
let u,
v be
Vector of
M;
func u + v -> Vector of
M means :
Def10:
ex
p,
q being
Element of
[: the carrier of M, the carrier of M:] st
(
u = p ~ &
v = q ~ &
p `2 = q `1 &
it = [(p `1),(q `2)] ~ );
existence
ex b1 being Vector of M ex p, q being Element of [: the carrier of M, the carrier of M:] st
( u = p ~ & v = q ~ & p `2 = q `1 & b1 = [(p `1),(q `2)] ~ )
by Th51;
uniqueness
for b1, b2 being Vector of M st ex p, q being Element of [: the carrier of M, the carrier of M:] st
( u = p ~ & v = q ~ & p `2 = q `1 & b1 = [(p `1),(q `2)] ~ ) & ex p, q being Element of [: the carrier of M, the carrier of M:] st
( u = p ~ & v = q ~ & p `2 = q `1 & b2 = [(p `1),(q `2)] ~ ) holds
b1 = b2
by Th52;
end;
:: deftheorem Def10 defines + MIDSP_1:def 10 :
for M being MidSp
for u, v, b4 being Vector of M holds
( b4 = u + v iff ex p, q being Element of [: the carrier of M, the carrier of M:] st
( u = p ~ & v = q ~ & p `2 = q `1 & b4 = [(p `1),(q `2)] ~ ) );
theorem Th53:
:: deftheorem defines vect MIDSP_1:def 11 :
for M being MidSp
for a, b being Element of M holds vect (a,b) = [a,b] ~ ;
theorem
canceled;
theorem Th55:
theorem
theorem
theorem
theorem
theorem Th60:
theorem Th61:
theorem
theorem Th63:
theorem Th64:
theorem Th65:
theorem Th66:
theorem Th67:
for
M being
MidSp for
u,
v,
w being
Vector of
M st
u + v = u + w holds
v = w
:: deftheorem defines - MIDSP_1:def 12 :
for M being MidSp
for u, b3 being Vector of M holds
( b3 = - u iff u + b3 = ID M );
:: deftheorem defines setvect MIDSP_1:def 13 :
for M being MidSp holds setvect M = { X where X is Subset of [: the carrier of M, the carrier of M:] : X is Vector of M } ;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem Th71:
:: deftheorem Def14 defines + MIDSP_1:def 14 :
for M being MidSp
for u1, v1, b4 being Element of setvect M holds
( b4 = u1 + v1 iff for u, v being Vector of M st u1 = u & v1 = v holds
b4 = u + v );
theorem
canceled;
theorem
canceled;
theorem Th74:
theorem Th75:
:: deftheorem Def15 defines addvect MIDSP_1:def 15 :
for M being MidSp
for b2 being BinOp of (setvect M) holds
( b2 = addvect M iff for u1, v1 being Element of setvect M holds b2 . (u1,v1) = u1 + v1 );
theorem
canceled;
theorem Th77:
theorem Th78:
:: deftheorem Def16 defines complvect MIDSP_1:def 16 :
for M being MidSp
for b2 being UnOp of (setvect M) holds
( b2 = complvect M iff for W being Element of setvect M holds W + (b2 . W) = ID M );
:: deftheorem defines zerovect MIDSP_1:def 17 :
for M being MidSp holds zerovect M = ID M;
:: deftheorem defines vectgroup MIDSP_1:def 18 :
for M being MidSp holds vectgroup M = addLoopStr(# (setvect M),(addvect M),(zerovect M) #);
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
theorem
theorem
canceled;
theorem
theorem