:: AFVECT01 semantic presentation

registration
let A be non empty set ;
let C be Relation of [:A,A:];
cluster AffinStruct(# a1,a2 #) -> non empty ;
coherence
not AffinStruct(# A,C #) is empty
by STRUCT_0:def 1;
end;

Lemma25: for AFV being WeakAffVect
for a, b, c being Element of AFV st a,b '||' b,c & a <> c holds
a,b // b,c
proof end;

Lemma27: for AFV being WeakAffVect
for a, b being Element of AFV holds
( a,b // b,a iff ex p, q being Element of AFV st
( a,b '||' p,q & a,p '||' p,b & a,q '||' q,b ) )
proof end;

Lemma42: for AFV being WeakAffVect
for a, b, c, d being Element of AFV st a,b '||' c,d holds
b,a '||' c,d
proof end;

Lemma43: for AFV being WeakAffVect
for a, b being Element of AFV holds a,b '||' b,a
proof end;

Lemma44: for AFV being WeakAffVect
for a, b, p being Element of AFV st a,b '||' p,p holds
a = b
proof end;

Lemma45: for AFV being WeakAffVect
for a, b, c, d, p, q being Element of AFV st a,b '||' p,q & c,d '||' p,q holds
a,b '||' c,d
proof end;

Lemma46: for AFV being WeakAffVect
for a, c being Element of AFV ex b being Element of AFV st a,b '||' b,c
proof end;

Lemma47: for AFV being WeakAffVect
for a, a', b, b', p being Element of AFV st a <> a' & b <> b' & p,a '||' p,a' & p,b '||' p,b' holds
a,b '||' a',b'
proof end;

Lemma48: for AFV being WeakAffVect
for a, b being Element of AFV holds
( a = b or ex c being Element of AFV st
( ( a <> c & a,b '||' b,c ) or ex p, p' being Element of AFV st
( p <> p' & a,b '||' p,p' & a,p '||' p,b & a,p' '||' p',b ) ) )
proof end;

Lemma49: for AFV being WeakAffVect
for a, b, b', p, p', c being Element of AFV st a,b '||' b,c & b,b' '||' p,p' & b,p '||' p,b' & b,p' '||' p',b' holds
a,b' '||' b',c
proof end;

Lemma50: for AFV being WeakAffVect
for a, b, b', c being Element of AFV st a <> c & b <> b' & a,b '||' b,c & a,b' '||' b',c holds
ex p, p' being Element of AFV st
( p <> p' & b,b' '||' p,p' & b,p '||' p,b' & b,p' '||' p',b' )
proof end;

Lemma51: for AFV being WeakAffVect
for a, b, c, p, p', q, q' being Element of AFV st a,b '||' p,p' & a,c '||' q,q' & a,p '||' p,b & a,q '||' q,c & a,p' '||' p',b & a,q' '||' q',c holds
ex r, r' being Element of AFV st
( b,c '||' r,r' & b,r '||' r,c & b,r' '||' r',c )
proof end;

consider AFV0 being WeakAffVect;

set X = the carrier of AFV0;

set XX = [:the carrier of AFV0,the carrier of AFV0:];

defpred S1[ set , set ] means ex a, b, c, d being Element of the carrier of AFV0 st
( a1 = [a,b] & a2 = [c,d] & a,b '||' c,d );

consider P being Relation of [:the carrier of AFV0,the carrier of AFV0:],[:the carrier of AFV0,the carrier of AFV0:] such that
Lemma56: for x, y being set holds
( [x,y] in P iff ( x in [:the carrier of AFV0,the carrier of AFV0:] & y in [:the carrier of AFV0,the carrier of AFV0:] & S1[x,y] ) ) from RELSET_1:sch 1();

Lemma59: for a, b, c, d being Element of the carrier of AFV0 holds
( [[a,b],[c,d]] in P iff a,b '||' c,d )
proof end;

set WAS = AffinStruct(# the carrier of AFV0,P #);

Lemma63: for a, b, c, d being Element of AFV0
for a', b', c', d' being Element of the carrier of AffinStruct(# the carrier of AFV0,P #) st a = a' & b = b' & c = c' & d = d' holds
( a,b '||' c,d iff a',b' // c',d' )
proof end;

E64: now
thus ex a', b' being Element of AffinStruct(# the carrier of AFV0,P #) st a' <> b' by REALSET2:def 7;
thus for a', b' being Element of AffinStruct(# the carrier of AFV0,P #) holds a',b' // b',a'
proof
let a' be Element of AffinStruct(# the carrier of AFV0,P #), b' be Element of AffinStruct(# the carrier of AFV0,P #);
reconsider a = a', b = b' as Element of the carrier of AFV0 ;
a,b '||' b,a by ;
hence a',b' // b',a' by ;
end;
thus for a', b' being Element of AffinStruct(# the carrier of AFV0,P #) st a',b' // a',a' holds
a' = b'
proof
let a' be Element of AffinStruct(# the carrier of AFV0,P #), b' be Element of AffinStruct(# the carrier of AFV0,P #);
assume E26: a',b' // a',a' ;
reconsider a = a', b = b' as Element of AFV0 ;
a,b '||' a,a by , ;
hence a' = b' by ;
end;
thus for a, b, c, d, p, q being Element of AffinStruct(# the carrier of AFV0,P #) st a,b // p,q & c,d // p,q holds
a,b // c,d
proof
let a be Element of AffinStruct(# the carrier of AFV0,P #), b be Element of AffinStruct(# the carrier of AFV0,P #), c be Element of AffinStruct(# the carrier of AFV0,P #), d be Element of AffinStruct(# the carrier of AFV0,P #), p be Element of AffinStruct(# the carrier of AFV0,P #), q be Element of AffinStruct(# the carrier of AFV0,P #);
assume E28: ( a,b // p,q & c,d // p,q ) ;
reconsider a1 = a, b1 = b, c1 = c, d1 = d, p1 = p, q1 = q as Element of AFV0 ;
( a1,b1 '||' p1,q1 & c1,d1 '||' p1,q1 ) by , ;
then a1,b1 '||' c1,d1 by ;
hence a,b // c,d by ;
end;
thus for a, c being Element of AffinStruct(# the carrier of AFV0,P #) ex b being Element of the carrier of AffinStruct(# the carrier of AFV0,P #) st a,b // b,c
proof
let a be Element of AffinStruct(# the carrier of AFV0,P #), c be Element of AffinStruct(# the carrier of AFV0,P #);
reconsider a1 = a, c1 = c as Element of AFV0 ;
consider b1 being Element of AFV0 such that
E29: a1,b1 '||' b1,c1 by ;
reconsider b = b1 as Element of AffinStruct(# the carrier of AFV0,P #) ;
a,b // b,c by , ;
hence ex b being Element of the carrier of AffinStruct(# the carrier of AFV0,P #) st a,b // b,c ;
end;
thus for a, a', b, b', p being Element of AffinStruct(# the carrier of AFV0,P #) st a <> a' & b <> b' & p,a // p,a' & p,b // p,b' holds
a,b // a',b'
proof
let a be Element of AffinStruct(# the carrier of AFV0,P #), a' be Element of AffinStruct(# the carrier of AFV0,P #), b be Element of AffinStruct(# the carrier of AFV0,P #), b' be Element of AffinStruct(# the carrier of AFV0,P #), p be Element of AffinStruct(# the carrier of AFV0,P #);
assume that
E30: ( a <> a' & b <> b' ) and
E31: ( p,a // p,a' & p,b // p,b' ) ;
reconsider a1 = a, a1' = a', b1 = b, b1' = b', p1 = p as Element of AFV0 ;
( p1,a1 '||' p1,a1' & p1,b1 '||' p1,b1' ) by , ;
then a1,b1 '||' a1',b1' by , ;
hence a,b // a',b' by ;
end;
thus for a, b being Element of AffinStruct(# the carrier of AFV0,P #) holds
( a = b or ex c being Element of AffinStruct(# the carrier of AFV0,P #) st
( ( a <> c & a,b // b,c ) or ex p, p' being Element of AffinStruct(# the carrier of AFV0,P #) st
( p <> p' & a,b // p,p' & a,p // p,b & a,p' // p',b ) ) )
proof
let a be Element of AffinStruct(# the carrier of AFV0,P #), b be Element of AffinStruct(# the carrier of AFV0,P #);
assume E32: not a = b ;
reconsider a1 = a, b1 = b as Element of AFV0 ;
E33: now
given c1 being Element of AFV0 such that E34: ( a1 <> c1 & a1,b1 '||' b1,c1 ) ;
reconsider c = c1 as Element of AffinStruct(# the carrier of AFV0,P #) ;
( a <> c & a,b // b,c ) by , ;
hence ex c being Element of AffinStruct(# the carrier of AFV0,P #) st
( a <> c & a,b // b,c ) ;
end;
now
given p1 being Element of AFV0, p1' being Element of AFV0 such that E35: ( p1 <> p1' & a1,b1 '||' p1,p1' & a1,p1 '||' p1,b1 & a1,p1' '||' p1',b1 ) ;
reconsider p = p1, p' = p1' as Element of AffinStruct(# the carrier of AFV0,P #) ;
( p <> p' & a,b // p,p' & a,p // p,b & a,p' // p',b ) by , ;
hence ex p, p' being Element of AffinStruct(# the carrier of AFV0,P #) st
( p <> p' & a,b // p,p' & a,p // p,b & a,p' // p',b ) ;
end;
hence ex c being Element of AffinStruct(# the carrier of AFV0,P #) st
( ( a <> c & a,b // b,c ) or ex p, p' being Element of AffinStruct(# the carrier of AFV0,P #) st
( p <> p' & a,b // p,p' & a,p // p,b & a,p' // p',b ) ) by , , ;
end;
thus for a, b, b', p, p', c being Element of AffinStruct(# the carrier of AFV0,P #) st a,b // b,c & b,b' // p,p' & b,p // p,b' & b,p' // p',b' holds
a,b' // b',c
proof
let a be Element of AffinStruct(# the carrier of AFV0,P #), b be Element of AffinStruct(# the carrier of AFV0,P #), b' be Element of AffinStruct(# the carrier of AFV0,P #), p be Element of AffinStruct(# the carrier of AFV0,P #), p' be Element of AffinStruct(# the carrier of AFV0,P #), c be Element of AffinStruct(# the carrier of AFV0,P #);
assume E36: ( a,b // b,c & b,b' // p,p' & b,p // p,b' & b,p' // p',b' ) ;
reconsider a1 = a, b1 = b, b1' = b', p1 = p, p1' = p', c1 = c as Element of AFV0 ;
( a1,b1 '||' b1,c1 & b1,b1' '||' p1,p1' & b1,p1 '||' p1,b1' & b1,p1' '||' p1',b1' ) by , ;
then a1,b1' '||' b1',c1 by ;
hence a,b' // b',c by ;
end;
thus for a, b, b', c being Element of AffinStruct(# the carrier of AFV0,P #) st a <> c & b <> b' & a,b // b,c & a,b' // b',c holds
ex p, p' being Element of AffinStruct(# the carrier of AFV0,P #) st
( p <> p' & b,b' // p,p' & b,p // p,b' & b,p' // p',b' )
proof
let a be Element of AffinStruct(# the carrier of AFV0,P #), b be Element of AffinStruct(# the carrier of AFV0,P #), b' be Element of AffinStruct(# the carrier of AFV0,P #), c be Element of AffinStruct(# the carrier of AFV0,P #);
assume that
E37: ( a <> c & b <> b' ) and
E38: ( a,b // b,c & a,b' // b',c ) ;
reconsider a1 = a, b1 = b, b1' = b', c1 = c as Element of the carrier of AFV0 ;
( a1,b1 '||' b1,c1 & a1,b1' '||' b1',c1 ) by , ;
then consider p1 being Element of AFV0, p1' being Element of AFV0 such that
E39: ( p1 <> p1' & b1,b1' '||' p1,p1' & b1,p1 '||' p1,b1' & b1,p1' '||' p1',b1' ) by , ;
reconsider p = p1, p' = p1' as Element of AffinStruct(# the carrier of AFV0,P #) ;
( p <> p' & b,b' // p,p' & b,p // p,b' & b,p' // p',b' ) by , ;
hence ex p, p' being Element of AffinStruct(# the carrier of AFV0,P #) st
( p <> p' & b,b' // p,p' & b,p // p,b' & b,p' // p',b' ) ;
end;
thus for a, b, c, p, p', q, q' being Element of AffinStruct(# the carrier of AFV0,P #) st a,b // p,p' & a,c // q,q' & a,p // p,b & a,q // q,c & a,p' // p',b & a,q' // q',c holds
ex r, r' being Element of AffinStruct(# the carrier of AFV0,P #) st
( b,c // r,r' & b,r // r,c & b,r' // r',c )
proof
let a be Element of AffinStruct(# the carrier of AFV0,P #), b be Element of AffinStruct(# the carrier of AFV0,P #), c be Element of AffinStruct(# the carrier of AFV0,P #), p be Element of AffinStruct(# the carrier of AFV0,P #), p' be Element of AffinStruct(# the carrier of AFV0,P #), q be Element of AffinStruct(# the carrier of AFV0,P #), q' be Element of AffinStruct(# the carrier of AFV0,P #);
assume E40: ( a,b // p,p' & a,c // q,q' & a,p // p,b & a,q // q,c & a,p' // p',b & a,q' // q',c ) ;
reconsider a1 = a, b1 = b, c1 = c, p1 = p, p1' = p', q1 = q, q1' = q' as Element of the carrier of AFV0 ;
( a1,b1 '||' p1,p1' & a1,c1 '||' q1,q1' & a1,p1 '||' p1,b1 & a1,q1 '||' q1,c1 & a1,p1' '||' p1',b1 & a1,q1' '||' q1',c1 ) by , ;
then consider r1 being Element of AFV0, r1' being Element of AFV0 such that
E41: ( b1,c1 '||' r1,r1' & b1,r1 '||' r1,c1 & b1,r1' '||' r1',c1 ) by ;
reconsider r = r1, r' = r1' as Element of AffinStruct(# the carrier of AFV0,P #) ;
( b,c // r,r' & b,r // r,c & b,r' // r',c ) by , ;
hence ex r, r' being Element of AffinStruct(# the carrier of AFV0,P #) st
( b,c // r,r' & b,r // r,c & b,r' // r',c ) ;
end;
end;

definition
let IT be non empty AffinStruct ;
canceled;
attr a1 is WeakAffSegm-like means :Def2: :: AFVECT01:def 2
( ( for a, b being Element of AFV0 holds a,b // b,a ) & ( for a, b being Element of AFV0 st a,b // a,a holds
a = b ) & ( for a, b, c, d, p, q being Element of AFV0 st a,b // p,q & c,d // p,q holds
a,b // c,d ) & ( for a, c being Element of AFV0 ex b being Element of AFV0 st a,b // b,c ) & ( for a, a', b, b', p being Element of AFV0 st a <> a' & b <> b' & p,a // p,a' & p,b // p,b' holds
a,b // a',b' ) & ( for a, b being Element of AFV0 holds
( a = b or ex c being Element of AFV0 st
( ( a <> c & a,b // b,c ) or ex p, p' being Element of AFV0 st
( p <> p' & a,b // p,p' & a,p // p,b & a,p' // p',b ) ) ) ) & ( for a, b, b', p, p', c being Element of AFV0 st a,b // b,c & b,b' // p,p' & b,p // p,b' & b,p' // p',b' holds
a,b' // b',c ) & ( for a, b, b', c being Element of AFV0 st a <> c & b <> b' & a,b // b,c & a,b' // b',c holds
ex p, p' being Element of AFV0 st
( p <> p' & b,b' // p,p' & b,p // p,b' & b,p' // p',b' ) ) & ( for a, b, c, p, p', q, q' being Element of AFV0 st a,b // p,p' & a,c // q,q' & a,p // p,b & a,q // q,c & a,p' // p',b & a,q' // q',c holds
ex r, r' being Element of AFV0 st
( b,c // r,r' & b,r // r,c & b,r' // r',c ) ) );
end;

:: deftheorem Def1 AFVECT01:def 1 :
canceled;

:: deftheorem Def2 defines WeakAffSegm-like AFVECT01:def 2 :
for IT being non empty AffinStruct holds
( IT is WeakAffSegm-like iff ( ( for a, b being Element of IT holds a,b // b,a ) & ( for a, b being Element of IT st a,b // a,a holds
a = b ) & ( for a, b, c, d, p, q being Element of IT st a,b // p,q & c,d // p,q holds
a,b // c,d ) & ( for a, c being Element of IT ex b being Element of IT st a,b // b,c ) & ( for a, a', b, b', p being Element of IT st a <> a' & b <> b' & p,a // p,a' & p,b // p,b' holds
a,b // a',b' ) & ( for a, b being Element of IT holds
( a = b or ex c being Element of IT st
( ( a <> c & a,b // b,c ) or ex p, p' being Element of IT st
( p <> p' & a,b // p,p' & a,p // p,b & a,p' // p',b ) ) ) ) & ( for a, b, b', p, p', c being Element of IT st a,b // b,c & b,b' // p,p' & b,p // p,b' & b,p' // p',b' holds
a,b' // b',c ) & ( for a, b, b', c being Element of IT st a <> c & b <> b' & a,b // b,c & a,b' // b',c holds
ex p, p' being Element of IT st
( p <> p' & b,b' // p,p' & b,p // p,b' & b,p' // p',b' ) ) & ( for a, b, c, p, p', q, q' being Element of IT st a,b // p,p' & a,c // q,q' & a,p // p,b & a,q // q,c & a,p' // p',b & a,q' // q',c holds
ex r, r' being Element of IT st
( b,c // r,r' & b,r // r,c & b,r' // r',c ) ) ) );

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

definition
mode WeakAffSegm is non empty non trivial WeakAffSegm-like AffinStruct ;
end;

theorem Th1: :: AFVECT01:1
for AFV being WeakAffSegm
for a, b being Element of AFV holds a,b // a,b
proof end;

theorem Th2: :: AFVECT01:2
for AFV being WeakAffSegm
for a, b, c, d being Element of AFV st a,b // c,d holds
c,d // a,b
proof end;

theorem Th3: :: AFVECT01:3
for AFV being WeakAffSegm
for a, b, c, d being Element of AFV st a,b // c,d holds
a,b // d,c
proof end;

theorem Th4: :: AFVECT01:4
for AFV being WeakAffSegm
for a, b, c, d being Element of AFV st a,b // c,d holds
b,a // c,d
proof end;

theorem Th5: :: AFVECT01:5
for AFV being WeakAffSegm
for a, b being Element of AFV holds a,a // b,b
proof end;

theorem Th6: :: AFVECT01:6
for AFV being WeakAffSegm
for a, b, c being Element of AFV st a,b // c,c holds
a = b
proof end;

theorem Th7: :: AFVECT01:7
for AFV being WeakAffSegm
for a, b, p, p', c being Element of AFV st a,b // p,p' & a,b // b,c & a,p // p,b & a,p' // p',b holds
a = c
proof end;

theorem Th8: :: AFVECT01:8
for AFV being WeakAffSegm
for a, b', b'', b being Element of AFV st a,b' // a,b'' & a,b // a,b'' & not b = b' & not b = b'' holds
b' = b''
proof end;

definition
let AFV be WeakAffSegm;
let a be Element of AFV;
let b be Element of AFV;
canceled;
pred MDist c2,c3 means :Def4: :: AFVECT01:def 4
ex p, p' being Element of AFV0 st
( p <> p' & X,XX // p,p' & X,p // p,XX & X,p' // p',XX );
end;

:: deftheorem Def3 AFVECT01:def 3 :
canceled;

:: deftheorem Def4 defines MDist AFVECT01:def 4 :
for AFV being WeakAffSegm
for a, b being Element of AFV holds
( MDist a,b iff ex p, p' being Element of AFV st
( p <> p' & a,b // p,p' & a,p // p,b & a,p' // p',b ) );

definition
let AFV be WeakAffSegm;
let a be Element of AFV;
let b be Element of AFV;
let c be Element of AFV;
pred Mid c2,c3,c4 means :: AFVECT01:def 5
( ( X = XX & XX = P & X = P ) or ( X = P & MDist X,XX ) or ( X <> P & X,XX // XX,P ) );
end;

:: deftheorem Def5 defines Mid AFVECT01:def 5 :
for AFV being WeakAffSegm
for a, b, c being Element of AFV holds
( Mid a,b,c iff ( ( a = b & b = c & a = c ) or ( a = c & MDist a,b ) or ( a <> c & a,b // b,c ) ) );

theorem Th9: :: AFVECT01:9
canceled;

theorem Th10: :: AFVECT01:10
canceled;

theorem Th11: :: AFVECT01:11
for AFV being WeakAffSegm
for a, b being Element of AFV st a <> b & not MDist a,b holds
ex c being Element of AFV st
( a <> c & a,b // b,c )
proof end;

theorem Th12: :: AFVECT01:12
for AFV being WeakAffSegm
for a, b, c being Element of AFV st MDist a,b & a,b // b,c holds
a = c
proof end;

theorem Th13: :: AFVECT01:13
for AFV being WeakAffSegm
for a, b being Element of AFV st MDist a,b holds
a <> b
proof end;