:: Analytical Ordered Affine Spaces
:: by Henryk Oryszczyszyn and Krzysztof Pra\.zmowski
::
:: Received April 11, 1990
:: Copyright (c) 1990-2021 Association of Mizar Users


definition
let V be RealLinearSpace;
let u, v, w, y be VECTOR of V;
pred u,v // w,y means :: ANALOAF:def 1
( u = v or w = y or ex a, b being Real st
( 0 < a & 0 < b & a * (v - u) = b * (y - w) ) );
end;

:: deftheorem defines // ANALOAF:def 1 :
for V being RealLinearSpace
for u, v, w, y being VECTOR of V holds
( u,v // w,y iff ( u = v or w = y or ex a, b being Real st
( 0 < a & 0 < b & a * (v - u) = b * (y - w) ) ) );

theorem Th1: :: ANALOAF:1
for V being RealLinearSpace
for u, v, w being VECTOR of V holds (w - v) + (v - u) = w - u
proof end;

theorem Th2: :: ANALOAF:2
for V being RealLinearSpace
for u, v, w, y being VECTOR of V st y + u = v + w holds
y - w = v - u
proof end;

theorem Th3: :: ANALOAF:3
for V being RealLinearSpace
for u, v being VECTOR of V
for a being Real holds a * (u - v) = - (a * (v - u))
proof end;

theorem Th4: :: ANALOAF:4
for V being RealLinearSpace
for u, v being VECTOR of V
for a, b being Real holds (a - b) * (u - v) = (b - a) * (v - u)
proof end;

theorem Th5: :: ANALOAF:5
for V being RealLinearSpace
for u, v being VECTOR of V
for a being Real st a <> 0 & a * u = v holds
u = (a ") * v
proof end;

theorem Th6: :: ANALOAF:6
for V being RealLinearSpace
for u, v being VECTOR of V
for a being Real holds
( ( a <> 0 & a * u = v implies u = (a ") * v ) & ( a <> 0 & u = (a ") * v implies a * u = v ) )
proof end;

theorem :: ANALOAF:7
for V being RealLinearSpace
for u, v, w, y being VECTOR of V st u,v // w,y & u <> v & w <> y holds
ex a, b being Real st
( a * (v - u) = b * (y - w) & 0 < a & 0 < b ) ;

reconsider jj = 1 as Real ;

theorem Th8: :: ANALOAF:8
for V being RealLinearSpace
for u, v being VECTOR of V holds u,v // u,v
proof end;

theorem :: ANALOAF:9
for V being RealLinearSpace
for u, v, w being VECTOR of V holds
( u,v // w,w & u,u // v,w ) ;

theorem Th10: :: ANALOAF:10
for V being RealLinearSpace
for u, v being VECTOR of V st u,v // v,u holds
u = v
proof end;

theorem Th11: :: ANALOAF:11
for V being RealLinearSpace
for p, q, u, v, w, y being VECTOR of V st p <> q & p,q // u,v & p,q // w,y holds
u,v // w,y
proof end;

theorem Th12: :: ANALOAF:12
for V being RealLinearSpace
for u, v, w, y being VECTOR of V st u,v // w,y holds
( v,u // y,w & w,y // u,v )
proof end;

theorem Th13: :: ANALOAF:13
for V being RealLinearSpace
for u, v, w being VECTOR of V st u,v // v,w holds
u,v // u,w
proof end;

theorem Th14: :: ANALOAF:14
for V being RealLinearSpace
for u, v, w being VECTOR of V holds
( not u,v // u,w or u,v // v,w or u,w // w,v )
proof end;

theorem Th15: :: ANALOAF:15
for V being RealLinearSpace
for u, v, w, y being VECTOR of V st v - u = y - w holds
u,v // w,y
proof end;

theorem Th16: :: ANALOAF:16
for V being RealLinearSpace
for u, v, w, y being VECTOR of V st y = (v + w) - u holds
( u,v // w,y & u,w // v,y )
proof end;

theorem Th17: :: ANALOAF:17
for V being RealLinearSpace st ex p, q being VECTOR of V st p <> q holds
for u, v, w being VECTOR of V ex y being VECTOR of V st
( u,v // w,y & u,w // v,y & v <> y )
proof end;

theorem Th18: :: ANALOAF:18
for V being RealLinearSpace
for p, u, v, w being VECTOR of V st p <> v & v,p // p,w holds
ex y being VECTOR of V st
( u,p // p,y & u,v // w,y )
proof end;

theorem Th19: :: ANALOAF:19
for V being RealLinearSpace
for u, v being VECTOR of V st ( for a, b being Real st (a * u) + (b * v) = 0. V holds
( a = 0 & b = 0 ) ) holds
( u <> v & u <> 0. V & v <> 0. V )
proof end;

theorem Th20: :: ANALOAF:20
for V being RealLinearSpace st ex u, v being VECTOR of V st
for a, b being Real st (a * u) + (b * v) = 0. V holds
( a = 0 & b = 0 ) holds
ex u, v, w, y being VECTOR of V st
( not u,v // w,y & not u,v // y,w )
proof end;

Lm1: for V being RealLinearSpace
for u, v, w, y being VECTOR of V
for a, b being Real st a * (v - u) = b * (w - y) & ( a <> 0 or b <> 0 ) & not u,v // w,y holds
u,v // y,w

proof end;

theorem Th21: :: ANALOAF:21
for V being RealLinearSpace st ex p, q being VECTOR of V st
for w being VECTOR of V ex a, b being Real st (a * p) + (b * q) = w holds
for u, v, w, y being VECTOR of V st not u,v // w,y & not u,v // y,w holds
ex z being VECTOR of V st
( ( u,v // u,z or u,v // z,u ) & ( w,y // w,z or w,y // z,w ) )
proof end;

definition
attr c1 is strict ;
struct AffinStruct -> 1-sorted ;
aggr AffinStruct(# carrier, CONGR #) -> AffinStruct ;
sel CONGR c1 -> Relation of [: the carrier of c1, the carrier of c1:];
end;

registration
cluster non trivial strict for AffinStruct ;
existence
ex b1 being AffinStruct st
( not b1 is trivial & b1 is strict )
proof end;
end;

definition
let AS be non empty AffinStruct ;
let a, b, c, d be Element of AS;
pred a,b // c,d means :: ANALOAF:def 2
[[a,b],[c,d]] in the CONGR of AS;
end;

:: deftheorem defines // ANALOAF:def 2 :
for AS being non empty AffinStruct
for a, b, c, d being Element of AS holds
( a,b // c,d iff [[a,b],[c,d]] in the CONGR of AS );

definition
let V be RealLinearSpace;
func DirPar V -> Relation of [: the carrier of V, the carrier of V:] means :Def3: :: ANALOAF:def 3
for x, z being object holds
( [x,z] in it iff ex u, v, w, y being VECTOR of V st
( x = [u,v] & z = [w,y] & u,v // w,y ) );
existence
ex b1 being Relation of [: the carrier of V, the carrier of V:] st
for x, z being object holds
( [x,z] in b1 iff ex u, v, w, y being VECTOR of V st
( x = [u,v] & z = [w,y] & u,v // w,y ) )
proof end;
uniqueness
for b1, b2 being Relation of [: the carrier of V, the carrier of V:] st ( for x, z being object holds
( [x,z] in b1 iff ex u, v, w, y being VECTOR of V st
( x = [u,v] & z = [w,y] & u,v // w,y ) ) ) & ( for x, z being object holds
( [x,z] in b2 iff ex u, v, w, y being VECTOR of V st
( x = [u,v] & z = [w,y] & u,v // w,y ) ) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def3 defines DirPar ANALOAF:def 3 :
for V being RealLinearSpace
for b2 being Relation of [: the carrier of V, the carrier of V:] holds
( b2 = DirPar V iff for x, z being object holds
( [x,z] in b2 iff ex u, v, w, y being VECTOR of V st
( x = [u,v] & z = [w,y] & u,v // w,y ) ) );

theorem Th22: :: ANALOAF:22
for V being RealLinearSpace
for u, v, w, y being VECTOR of V holds
( [[u,v],[w,y]] in DirPar V iff u,v // w,y )
proof end;

definition
let V be RealLinearSpace;
func OASpace V -> strict AffinStruct equals :: ANALOAF:def 4
AffinStruct(# the carrier of V,(DirPar V) #);
correctness
coherence
AffinStruct(# the carrier of V,(DirPar V) #) is strict AffinStruct
;
;
end;

:: deftheorem defines OASpace ANALOAF:def 4 :
for V being RealLinearSpace holds OASpace V = AffinStruct(# the carrier of V,(DirPar V) #);

registration
let V be RealLinearSpace;
cluster OASpace V -> non empty strict ;
coherence
not OASpace V is empty
;
end;

theorem Th23: :: ANALOAF:23
for V being RealLinearSpace st ex u, v being VECTOR of V st
for a, b being Real st (a * u) + (b * v) = 0. V holds
( a = 0 & b = 0 ) holds
( ex a, b being Element of (OASpace V) st a <> b & ( for a, b, c, d, p, q, r, s being Element of (OASpace V) holds
( a,b // c,c & ( a,b // b,a implies a = b ) & ( a <> b & a,b // p,q & a,b // r,s implies p,q // r,s ) & ( a,b // c,d implies b,a // d,c ) & ( a,b // b,c implies a,b // a,c ) & ( not a,b // a,c or a,b // b,c or a,c // c,b ) ) ) & ex a, b, c, d being Element of (OASpace V) st
( not a,b // c,d & not a,b // d,c ) & ( for a, b, c being Element of (OASpace V) ex d being Element of (OASpace V) st
( a,b // c,d & a,c // b,d & b <> d ) ) & ( for p, a, b, c being Element of (OASpace V) st p <> b & b,p // p,c holds
ex d being Element of (OASpace V) st
( a,p // p,d & a,b // c,d ) ) )
proof end;

theorem Th24: :: ANALOAF:24
for V being RealLinearSpace st ex p, q being VECTOR of V st
for w being VECTOR of V ex a, b being Real st (a * p) + (b * q) = w holds
for a, b, c, d being Element of (OASpace V) st not a,b // c,d & not a,b // d,c holds
ex t being Element of (OASpace V) st
( ( a,b // a,t or a,b // t,a ) & ( c,d // c,t or c,d // t,c ) )
proof end;

definition
let IT be non empty AffinStruct ;
attr IT is OAffinSpace-like means :Def5: :: ANALOAF:def 5
( ( for a, b, c, d, p, q, r, s being Element of IT holds
( a,b // c,c & ( a,b // b,a implies a = b ) & ( a <> b & a,b // p,q & a,b // r,s implies p,q // r,s ) & ( a,b // c,d implies b,a // d,c ) & ( a,b // b,c implies a,b // a,c ) & ( not a,b // a,c or a,b // b,c or a,c // c,b ) ) ) & ex a, b, c, d being Element of IT st
( not a,b // c,d & not a,b // d,c ) & ( for a, b, c being Element of IT ex d being Element of IT st
( a,b // c,d & a,c // b,d & b <> d ) ) & ( for p, a, b, c being Element of IT st p <> b & b,p // p,c holds
ex d being Element of IT st
( a,p // p,d & a,b // c,d ) ) );
end;

:: deftheorem Def5 defines OAffinSpace-like ANALOAF:def 5 :
for IT being non empty AffinStruct holds
( IT is OAffinSpace-like iff ( ( for a, b, c, d, p, q, r, s being Element of IT holds
( a,b // c,c & ( a,b // b,a implies a = b ) & ( a <> b & a,b // p,q & a,b // r,s implies p,q // r,s ) & ( a,b // c,d implies b,a // d,c ) & ( a,b // b,c implies a,b // a,c ) & ( not a,b // a,c or a,b // b,c or a,c // c,b ) ) ) & ex a, b, c, d being Element of IT st
( not a,b // c,d & not a,b // d,c ) & ( for a, b, c being Element of IT ex d being Element of IT st
( a,b // c,d & a,c // b,d & b <> d ) ) & ( for p, a, b, c being Element of IT st p <> b & b,p // p,c holds
ex d being Element of IT st
( a,p // p,d & a,b // c,d ) ) ) );

registration
cluster non empty non trivial strict OAffinSpace-like for AffinStruct ;
existence
ex b1 being non trivial AffinStruct st
( b1 is strict & b1 is OAffinSpace-like )
proof end;
end;

definition
mode OAffinSpace is non trivial OAffinSpace-like AffinStruct ;
end;

theorem :: ANALOAF:25
for AS being non empty AffinStruct holds
( ( ex a, b being Element of AS st a <> b & ( for a, b, c, d, p, q, r, s being Element of AS holds
( a,b // c,c & ( a,b // b,a implies a = b ) & ( a <> b & a,b // p,q & a,b // r,s implies p,q // r,s ) & ( a,b // c,d implies b,a // d,c ) & ( a,b // b,c implies a,b // a,c ) & ( not a,b // a,c or a,b // b,c or a,c // c,b ) ) ) & ex a, b, c, d being Element of AS st
( not a,b // c,d & not a,b // d,c ) & ( for a, b, c being Element of AS ex d being Element of AS st
( a,b // c,d & a,c // b,d & b <> d ) ) & ( for p, a, b, c being Element of AS st p <> b & b,p // p,c holds
ex d being Element of AS st
( a,p // p,d & a,b // c,d ) ) ) iff AS is OAffinSpace ) by Def5, STRUCT_0:def 10;

theorem Th26: :: ANALOAF:26
for V being RealLinearSpace st ex u, v being VECTOR of V st
for a, b being Real st (a * u) + (b * v) = 0. V holds
( a = 0 & b = 0 ) holds
OASpace V is OAffinSpace
proof end;

definition
let IT be OAffinSpace;
attr IT is 2-dimensional means :Def6: :: ANALOAF:def 6
for a, b, c, d being Element of IT st not a,b // c,d & not a,b // d,c holds
ex p being Element of IT st
( ( a,b // a,p or a,b // p,a ) & ( c,d // c,p or c,d // p,c ) );
end;

:: deftheorem Def6 defines 2-dimensional ANALOAF:def 6 :
for IT being OAffinSpace holds
( IT is 2-dimensional iff for a, b, c, d being Element of IT st not a,b // c,d & not a,b // d,c holds
ex p being Element of IT st
( ( a,b // a,p or a,b // p,a ) & ( c,d // c,p or c,d // p,c ) ) );

registration
cluster non empty non trivial strict OAffinSpace-like 2-dimensional for AffinStruct ;
existence
ex b1 being OAffinSpace st
( b1 is strict & b1 is 2-dimensional )
proof end;
end;

definition
mode OAffinPlane is 2-dimensional OAffinSpace;
end;

theorem :: ANALOAF:27
for AS being non empty AffinStruct holds
( ( ex a, b being Element of AS st a <> b & ( for a, b, c, d, p, q, r, s being Element of AS holds
( a,b // c,c & ( a,b // b,a implies a = b ) & ( a <> b & a,b // p,q & a,b // r,s implies p,q // r,s ) & ( a,b // c,d implies b,a // d,c ) & ( a,b // b,c implies a,b // a,c ) & ( not a,b // a,c or a,b // b,c or a,c // c,b ) ) ) & ex a, b, c, d being Element of AS st
( not a,b // c,d & not a,b // d,c ) & ( for a, b, c being Element of AS ex d being Element of AS st
( a,b // c,d & a,c // b,d & b <> d ) ) & ( for p, a, b, c being Element of AS st p <> b & b,p // p,c holds
ex d being Element of AS st
( a,p // p,d & a,b // c,d ) ) & ( for a, b, c, d being Element of AS st not a,b // c,d & not a,b // d,c holds
ex p being Element of AS st
( ( a,b // a,p or a,b // p,a ) & ( c,d // c,p or c,d // p,c ) ) ) ) iff AS is OAffinPlane ) by Def5, Def6, STRUCT_0:def 10;

theorem :: ANALOAF:28
for V being RealLinearSpace st ex u, v being VECTOR of V st
( ( for a, b being Real st (a * u) + (b * v) = 0. V holds
( a = 0 & b = 0 ) ) & ( for w being VECTOR of V ex a, b being Real st w = (a * u) + (b * v) ) ) holds
OASpace V is OAffinPlane
proof end;