:: ANALMETR semantic presentation

Lemma1: for b1 being RealLinearSpace
for b2, b3, b4, b5 being VECTOR of b1
for b6, b7, b8, b9 being Real holds
( b2 = (b6 * b3) + (b7 * b4) & b5 = (b8 * b3) + (b9 * b4) implies ( b2 + b5 = ((b6 + b8) * b3) + ((b7 + b9) * b4) & b2 - b5 = ((b6 - b8) * b3) + ((b7 - b9) * b4) ) )
proof end;

Lemma2: for b1 being RealLinearSpace
for b2, b3 being VECTOR of b1 holds (0 * b2) + (0 * b3) = 0. b1
proof end;

Lemma3: for b1 being RealLinearSpace
for b2, b3, b4 being VECTOR of b1
for b5, b6, b7 being Real holds
( b2 = (b5 * b3) + (b6 * b4) implies b7 * b2 = ((b7 * b5) * b3) + ((b7 * b6) * b4) )
proof end;

definition
let c1 be RealLinearSpace;
let c2, c3 be VECTOR of c1;
pred Gen c2,c3 means :Def1: :: ANALMETR:def 1
( ( for b1 being VECTOR of a1 holds
ex b2, b3 being Real st b1 = (b2 * a2) + (b3 * a3) ) & ( for b1, b2 being Real holds
( (b1 * a2) + (b2 * a3) = 0. a1 implies ( b1 = 0 & b2 = 0 ) ) ) );
end;

:: deftheorem Def1 defines Gen ANALMETR:def 1 :
for b1 being RealLinearSpace
for b2, b3 being VECTOR of b1 holds
( Gen b2,b3 iff ( ( for b4 being VECTOR of b1 holds
ex b5, b6 being Real st b4 = (b5 * b2) + (b6 * b3) ) & ( for b4, b5 being Real holds
( (b4 * b2) + (b5 * b3) = 0. b1 implies ( b4 = 0 & b5 = 0 ) ) ) ) );

definition
let c1 be RealLinearSpace;
let c2, c3, c4, c5 be VECTOR of c1;
pred c2,c3 are_Ort_wrt c4,c5 means :Def2: :: ANALMETR:def 2
ex b1, b2, b3, b4 being Real st
( a2 = (b1 * a4) + (b2 * a5) & a3 = (b3 * a4) + (b4 * a5) & (b1 * b3) + (b2 * b4) = 0 );
end;

:: deftheorem Def2 defines are_Ort_wrt ANALMETR:def 2 :
for b1 being RealLinearSpace
for b2, b3, b4, b5 being VECTOR of b1 holds
( b2,b3 are_Ort_wrt b4,b5 iff ex b6, b7, b8, b9 being Real st
( b2 = (b6 * b4) + (b7 * b5) & b3 = (b8 * b4) + (b9 * b5) & (b6 * b8) + (b7 * b9) = 0 ) );

Lemma6: for b1 being RealLinearSpace
for b2, b3 being VECTOR of b1
for b4, b5, b6, b7 being Real holds
( Gen b2,b3 & (b4 * b2) + (b5 * b3) = (b6 * b2) + (b7 * b3) implies ( b4 = b6 & b5 = b7 ) )
proof end;

theorem Th1: :: ANALMETR:1
canceled;

theorem Th2: :: ANALMETR:2
canceled;

theorem Th3: :: ANALMETR:3
canceled;

theorem Th4: :: ANALMETR:4
canceled;

theorem Th5: :: ANALMETR:5
for b1 being RealLinearSpace
for b2, b3, b4, b5 being VECTOR of b1 holds
( Gen b4,b5 implies ( b2,b3 are_Ort_wrt b4,b5 iff for b6, b7, b8, b9 being Real holds
( b2 = (b6 * b4) + (b7 * b5) & b3 = (b8 * b4) + (b9 * b5) implies (b6 * b8) + (b7 * b9) = 0 ) ) )
proof end;

Lemma8: for b1 being RealLinearSpace
for b2, b3 being VECTOR of b1 holds
( Gen b2,b3 implies ( b2 <> 0. b1 & b3 <> 0. b1 ) )
proof end;

theorem Th6: :: ANALMETR:6
for b1 being RealLinearSpace
for b2, b3 being VECTOR of b1 holds b2,b3 are_Ort_wrt b2,b3
proof end;

theorem Th7: :: ANALMETR:7
ex b1 being RealLinearSpaceex b2, b3 being VECTOR of b1 st Gen b2,b3
proof end;

theorem Th8: :: ANALMETR:8
for b1 being RealLinearSpace
for b2, b3, b4, b5 being VECTOR of b1 holds
( b2,b3 are_Ort_wrt b4,b5 implies b3,b2 are_Ort_wrt b4,b5 )
proof end;

theorem Th9: :: ANALMETR:9
for b1 being RealLinearSpace
for b2, b3 being VECTOR of b1 holds
( Gen b2,b3 implies for b4, b5 being VECTOR of b1 holds
( b4, 0. b1 are_Ort_wrt b2,b3 & 0. b1,b5 are_Ort_wrt b2,b3 ) )
proof end;

theorem Th10: :: ANALMETR:10
for b1 being RealLinearSpace
for b2, b3, b4, b5 being VECTOR of b1
for b6, b7 being Real holds
( b2,b3 are_Ort_wrt b4,b5 implies b6 * b2,b7 * b3 are_Ort_wrt b4,b5 )
proof end;

theorem Th11: :: ANALMETR:11
for b1 being RealLinearSpace
for b2, b3, b4, b5 being VECTOR of b1
for b6, b7 being Real holds
( b2,b3 are_Ort_wrt b4,b5 implies ( b6 * b2,b3 are_Ort_wrt b4,b5 & b2,b7 * b3 are_Ort_wrt b4,b5 ) )
proof end;

theorem Th12: :: ANALMETR:12
for b1 being RealLinearSpace
for b2, b3 being VECTOR of b1 holds
( Gen b2,b3 implies for b4 being VECTOR of b1 holds
ex b5 being VECTOR of b1 st
( b4,b5 are_Ort_wrt b2,b3 & b5 <> 0. b1 ) )
proof end;

theorem Th13: :: ANALMETR:13
for b1 being RealLinearSpace
for b2, b3, b4, b5, b6 being VECTOR of b1 holds
not ( Gen b2,b3 & b4,b5 are_Ort_wrt b2,b3 & b4,b6 are_Ort_wrt b2,b3 & b4 <> 0. b1 & ( for b7, b8 being Real holds
not ( b7 * b5 = b8 * b6 & not ( not b7 <> 0 & not b8 <> 0 ) ) ) )
proof end;

theorem Th14: :: ANALMETR:14
for b1 being RealLinearSpace
for b2, b3, b4, b5, b6 being VECTOR of b1 holds
( Gen b2,b3 & b4,b5 are_Ort_wrt b2,b3 & b4,b6 are_Ort_wrt b2,b3 implies ( b4,b5 + b6 are_Ort_wrt b2,b3 & b4,b5 - b6 are_Ort_wrt b2,b3 ) )
proof end;

theorem Th15: :: ANALMETR:15
for b1 being RealLinearSpace
for b2, b3, b4 being VECTOR of b1 holds
( Gen b2,b3 & b4,b4 are_Ort_wrt b2,b3 implies b4 = 0. b1 )
proof end;

theorem Th16: :: ANALMETR:16
for b1 being RealLinearSpace
for b2, b3, b4, b5, b6 being VECTOR of b1 holds
( Gen b2,b3 & b4,b5 - b6 are_Ort_wrt b2,b3 & b5,b6 - b4 are_Ort_wrt b2,b3 implies b6,b4 - b5 are_Ort_wrt b2,b3 )
proof end;

theorem Th17: :: ANALMETR:17
for b1 being RealLinearSpace
for b2, b3, b4, b5 being VECTOR of b1 holds
not ( Gen b2,b3 & b4 <> 0. b1 & ( for b6 being Real holds
not b5 - (b6 * b4),b4 are_Ort_wrt b2,b3 ) )
proof end;

theorem Th18: :: ANALMETR:18
for b1 being RealLinearSpace
for b2, b3, b4, b5 being VECTOR of b1 holds
( not ( ( b2,b3 // b4,b5 or b2,b3 // b5,b4 ) & ( for b6, b7 being Real holds
not ( b6 * (b3 - b2) = b7 * (b5 - b4) & not ( not b6 <> 0 & not b7 <> 0 ) ) ) ) & not ( ex b6, b7 being Real st
( b6 * (b3 - b2) = b7 * (b5 - b4) & not ( not b6 <> 0 & not b7 <> 0 ) ) & not b2,b3 // b4,b5 & not b2,b3 // b5,b4 ) )
proof end;

theorem Th19: :: ANALMETR:19
for b1 being RealLinearSpace
for b2, b3, b4, b5 being VECTOR of b1 holds
( [[b2,b3],[b4,b5]] in lambda (DirPar b1) iff ex b6, b7 being Real st
( b6 * (b3 - b2) = b7 * (b5 - b4) & not ( not b6 <> 0 & not b7 <> 0 ) ) )
proof end;

definition
let c1 be RealLinearSpace;
let c2, c3, c4, c5, c6, c7 be VECTOR of c1;
pred c2,c3,c4,c5 are_Ort_wrt c6,c7 means :Def3: :: ANALMETR:def 3
a3 - a2,a5 - a4 are_Ort_wrt a6,a7;
end;

:: deftheorem Def3 defines are_Ort_wrt ANALMETR:def 3 :
for b1 being RealLinearSpace
for b2, b3, b4, b5, b6, b7 being VECTOR of b1 holds
( b2,b3,b4,b5 are_Ort_wrt b6,b7 iff b3 - b2,b5 - b4 are_Ort_wrt b6,b7 );

definition
let c1 be RealLinearSpace;
let c2, c3 be VECTOR of c1;
func Orthogonality c1,c2,c3 -> Relation of [:the carrier of a1,the carrier of a1:] means :Def4: :: ANALMETR:def 4
for b1, b2 being set holds
( [b1,b2] in a4 iff ex b3, b4, b5, b6 being VECTOR of a1 st
( b1 = [b3,b4] & b2 = [b5,b6] & b3,b4,b5,b6 are_Ort_wrt a2,a3 ) );
existence
ex b1 being Relation of [:the carrier of c1,the carrier of c1:] st
for b2, b3 being set holds
( [b2,b3] in b1 iff ex b4, b5, b6, b7 being VECTOR of c1 st
( b2 = [b4,b5] & b3 = [b6,b7] & b4,b5,b6,b7 are_Ort_wrt c2,c3 ) )
proof end;
uniqueness
for b1, b2 being Relation of [:the carrier of c1,the carrier of c1:] holds
( ( for b3, b4 being set holds
( [b3,b4] in b1 iff ex b5, b6, b7, b8 being VECTOR of c1 st
( b3 = [b5,b6] & b4 = [b7,b8] & b5,b6,b7,b8 are_Ort_wrt c2,c3 ) ) ) & ( for b3, b4 being set holds
( [b3,b4] in b2 iff ex b5, b6, b7, b8 being VECTOR of c1 st
( b3 = [b5,b6] & b4 = [b7,b8] & b5,b6,b7,b8 are_Ort_wrt c2,c3 ) ) ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def4 defines Orthogonality ANALMETR:def 4 :
for b1 being RealLinearSpace
for b2, b3 being VECTOR of b1
for b4 being Relation of [:the carrier of b1,the carrier of b1:] holds
( b4 = Orthogonality b1,b2,b3 iff for b5, b6 being set holds
( [b5,b6] in b4 iff ex b7, b8, b9, b10 being VECTOR of b1 st
( b5 = [b7,b8] & b6 = [b9,b10] & b7,b8,b9,b10 are_Ort_wrt b2,b3 ) ) );

theorem Th20: :: ANALMETR:20
canceled;

theorem Th21: :: ANALMETR:21
canceled;

theorem Th22: :: ANALMETR:22
for b1 being RealLinearSpace holds the carrier of (Lambda (OASpace b1)) = the carrier of b1
proof end;

theorem Th23: :: ANALMETR:23
for b1 being RealLinearSpace holds the CONGR of (Lambda (OASpace b1)) = lambda (DirPar b1)
proof end;

theorem Th24: :: ANALMETR:24
for b1 being RealLinearSpace
for b2, b3, b4, b5 being VECTOR of b1
for b6, b7, b8, b9 being Element of (Lambda (OASpace b1)) holds
( b6 = b2 & b7 = b3 & b8 = b4 & b9 = b5 implies ( b6,b7 // b8,b9 iff ex b10, b11 being Real st
( b10 * (b3 - b2) = b11 * (b5 - b4) & not ( not b10 <> 0 & not b11 <> 0 ) ) ) )
proof end;

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

registration
cluster non empty ParOrtStr ;
existence
not for b1 being ParOrtStr holds b1 is empty
proof end;
end;

definition
let c1 be non empty ParOrtStr ;
let c2, c3, c4, c5 be Element of c1;
canceled;
pred c2,c3 _|_ c4,c5 means :Def6: :: ANALMETR:def 6
[[a2,a3],[a4,a5]] in the orthogonality of a1;
end;

:: deftheorem Def5 ANALMETR:def 5 :
canceled;

:: deftheorem Def6 defines _|_ ANALMETR:def 6 :
for b1 being non empty ParOrtStr
for b2, b3, b4, b5 being Element of b1 holds
( b2,b3 _|_ b4,b5 iff [[b2,b3],[b4,b5]] in the orthogonality of b1 );

definition
let c1 be RealLinearSpace;
let c2, c3 be VECTOR of c1;
func AMSpace c1,c2,c3 -> strict ParOrtStr equals :: ANALMETR:def 7
ParOrtStr(# the carrier of a1,(lambda (DirPar a1)),(Orthogonality a1,a2,a3) #);
correctness
coherence
ParOrtStr(# the carrier of c1,(lambda (DirPar c1)),(Orthogonality c1,c2,c3) #) is strict ParOrtStr
;
;
end;

:: deftheorem Def7 defines AMSpace ANALMETR:def 7 :
for b1 being RealLinearSpace
for b2, b3 being VECTOR of b1 holds AMSpace b1,b2,b3 = ParOrtStr(# the carrier of b1,(lambda (DirPar b1)),(Orthogonality b1,b2,b3) #);

registration
let c1 be RealLinearSpace;
let c2, c3 be VECTOR of c1;
cluster AMSpace a1,a2,a3 -> non empty strict ;
coherence
not AMSpace c1,c2,c3 is empty
proof end;
end;

theorem Th25: :: ANALMETR:25
canceled;

theorem Th26: :: ANALMETR:26
canceled;

theorem Th27: :: ANALMETR:27
canceled;

theorem Th28: :: ANALMETR:28
for b1 being RealLinearSpace
for b2, b3 being VECTOR of b1 holds
( the carrier of (AMSpace b1,b2,b3) = the carrier of b1 & the CONGR of (AMSpace b1,b2,b3) = lambda (DirPar b1) & the orthogonality of (AMSpace b1,b2,b3) = Orthogonality b1,b2,b3 ) ;

definition
let c1 be non empty ParOrtStr ;
func Af c1 -> strict AffinStruct equals :: ANALMETR:def 8
AffinStruct(# the carrier of a1,the CONGR of a1 #);
correctness
coherence
AffinStruct(# the carrier of c1,the CONGR of c1 #) is strict AffinStruct
;
;
end;

:: deftheorem Def8 defines Af ANALMETR:def 8 :
for b1 being non empty ParOrtStr holds Af b1 = AffinStruct(# the carrier of b1,the CONGR of b1 #);

registration
let c1 be non empty ParOrtStr ;
cluster Af a1 -> non empty strict ;
coherence
not Af c1 is empty
proof end;
end;

theorem Th29: :: ANALMETR:29
canceled;

theorem Th30: :: ANALMETR:30
for b1 being RealLinearSpace
for b2, b3 being VECTOR of b1 holds Af (AMSpace b1,b2,b3) = Lambda (OASpace b1)
proof end;

theorem Th31: :: ANALMETR:31
for b1 being RealLinearSpace
for b2, b3, b4, b5, b6, b7 being VECTOR of b1
for b8, b9, b10, b11 being Element of (AMSpace b1,b6,b7) holds
( b8 = b2 & b9 = b3 & b10 = b4 & b11 = b5 implies ( b8,b10 _|_ b9,b11 iff b2,b4,b3,b5 are_Ort_wrt b6,b7 ) )
proof end;

theorem Th32: :: ANALMETR:32
for b1 being RealLinearSpace
for b2, b3, b4, b5, b6, b7 being VECTOR of b1
for b8, b9, b10, b11 being Element of (AMSpace b1,b2,b3) holds
( b8 = b4 & b9 = b5 & b10 = b6 & b11 = b7 implies ( b8,b9 // b10,b11 iff ex b12, b13 being Real st
( b12 * (b5 - b4) = b13 * (b7 - b6) & not ( not b12 <> 0 & not b13 <> 0 ) ) ) )
proof end;

theorem Th33: :: ANALMETR:33
for b1 being RealLinearSpace
for b2, b3 being VECTOR of b1
for b4, b5, b6, b7 being Element of (AMSpace b1,b2,b3) holds
( b4,b5 _|_ b6,b7 implies b6,b7 _|_ b4,b5 )
proof end;

theorem Th34: :: ANALMETR:34
for b1 being RealLinearSpace
for b2, b3 being VECTOR of b1
for b4, b5, b6, b7 being Element of (AMSpace b1,b2,b3) holds
( b4,b5 _|_ b6,b7 implies b4,b5 _|_ b7,b6 )
proof end;

theorem Th35: :: ANALMETR:35
for b1 being RealLinearSpace
for b2, b3 being VECTOR of b1 holds
( Gen b2,b3 implies for b4, b5, b6 being Element of (AMSpace b1,b2,b3) holds b4,b5 _|_ b6,b6 )
proof end;

theorem Th36: :: ANALMETR:36
for b1 being RealLinearSpace
for b2, b3 being VECTOR of b1
for b4, b5, b6, b7, b8, b9 being Element of (AMSpace b1,b2,b3) holds
not ( b4,b5 _|_ b6,b7 & b4,b5 // b8,b9 & not b4 = b5 & not b6,b7 _|_ b8,b9 )
proof end;

theorem Th37: :: ANALMETR:37
for b1 being RealLinearSpace
for b2, b3 being VECTOR of b1 holds
( Gen b2,b3 implies for b4, b5, b6 being Element of (AMSpace b1,b2,b3) holds
ex b7 being Element of (AMSpace b1,b2,b3) st
( b4,b5 _|_ b6,b7 & b6 <> b7 ) )
proof end;

theorem Th38: :: ANALMETR:38
for b1 being RealLinearSpace
for b2, b3 being VECTOR of b1
for b4, b5, b6, b7, b8, b9 being Element of (AMSpace b1,b2,b3) holds
not ( Gen b2,b3 & b4,b5 _|_ b6,b7 & b4,b5 _|_ b8,b9 & not b4 = b5 & not b6,b7 // b8,b9 )
proof end;

theorem Th39: :: ANALMETR:39
for b1 being RealLinearSpace
for b2, b3 being VECTOR of b1
for b4, b5, b6, b7, b8 being Element of (AMSpace b1,b2,b3) holds
( Gen b2,b3 & b4,b5 _|_ b6,b7 & b4,b5 _|_ b6,b8 implies b4,b5 _|_ b7,b8 )
proof end;

theorem Th40: :: ANALMETR:40
for b1 being RealLinearSpace
for b2, b3 being VECTOR of b1
for b4, b5 being Element of (AMSpace b1,b2,b3) holds
( Gen b2,b3 & b4,b5 _|_ b4,b5 implies b4 = b5 )
proof end;

theorem Th41: :: ANALMETR:41
for b1 being RealLinearSpace
for b2, b3 being VECTOR of b1
for b4, b5, b6, b7 being Element of (AMSpace b1,b2,b3) holds
( Gen b2,b3 & b4,b5 _|_ b6,b7 & b6,b5 _|_ b7,b4 implies b7,b5 _|_ b4,b6 )
proof end;

theorem Th42: :: ANALMETR:42
for b1 being RealLinearSpace
for b2, b3 being VECTOR of b1
for b4, b5 being Element of (AMSpace b1,b2,b3) holds
( Gen b2,b3 & b4 <> b5 implies for b6 being Element of (AMSpace b1,b2,b3) holds
ex b7 being Element of (AMSpace b1,b2,b3) st
( b4,b5 // b4,b7 & b4,b5 _|_ b7,b6 ) )
proof end;

consider c1 being RealLinearSpace such that
Lemma39: ex b1, b2 being VECTOR of c1 st Gen b1,b2 by Th7;

consider c2, c3 being VECTOR of c1 such that
Lemma40: Gen c2,c3 by Lemma39;

E41: now
set c4 = AffinStruct(# the carrier of (AMSpace c1,c2,c3),the CONGR of (AMSpace c1,c2,c3) #);
AffinStruct(# the carrier of (AMSpace c1,c2,c3),the CONGR of (AMSpace c1,c2,c3) #) = Af (AMSpace c1,c2,c3) ;
then E42: AffinStruct(# the carrier of (AMSpace c1,c2,c3),the CONGR of (AMSpace c1,c2,c3) #) = Lambda (OASpace c1) by Th30;
for b1, b2 being Real holds
( (b1 * c2) + (b2 * c3) = 0. c1 implies ( b1 = 0 & b2 = 0 ) ) by Def1, Lemma40;
then OASpace c1 is OAffinSpace by ANALOAF:38;
hence ( AffinStruct(# the carrier of (AMSpace c1,c2,c3),the CONGR of (AMSpace c1,c2,c3) #) is AffinSpace & ( for b1, b2, b3, b4, b5, b6, b7, b8 being Element of (AMSpace c1,c2,c3) holds
( ( b1,b2 _|_ b1,b2 implies b1 = b2 ) & b1,b2 _|_ b3,b3 & ( b1,b2 _|_ b3,b4 implies ( b1,b2 _|_ b4,b3 & b3,b4 _|_ b1,b2 ) ) & not ( b1,b2 _|_ b5,b6 & b1,b2 // b7,b8 & not b5,b6 _|_ b7,b8 & not b1 = b2 ) & ( b1,b2 _|_ b5,b6 & b1,b2 _|_ b5,b8 implies b1,b2 _|_ b6,b8 ) ) ) & ( for b1, b2, b3 being Element of (AMSpace c1,c2,c3) holds
not ( b1 <> b2 & ( for b4 being Element of (AMSpace c1,c2,c3) holds
not ( b1,b2 // b1,b4 & b1,b2 _|_ b4,b3 ) ) ) ) & ( for b1, b2, b3 being Element of (AMSpace c1,c2,c3) holds
ex b4 being Element of (AMSpace c1,c2,c3) st
( b1,b2 _|_ b3,b4 & b3 <> b4 ) ) ) by E42, Lemma40, Th33, Th34, Th35, Th36, Th37, Th39, Th40, Th42, DIRAF:48;
end;

definition
let c4 be non empty ParOrtStr ;
attr a1 is OrtAfSp-like means :Def9: :: ANALMETR:def 9
( AffinStruct(# the carrier of a1,the CONGR of a1 #) is AffinSpace & ( for b1, b2, b3, b4, b5, b6, b7, b8 being Element of a1 holds
( ( b1,b2 _|_ b1,b2 implies b1 = b2 ) & b1,b2 _|_ b3,b3 & ( b1,b2 _|_ b3,b4 implies ( b1,b2 _|_ b4,b3 & b3,b4 _|_ b1,b2 ) ) & not ( b1,b2 _|_ b5,b6 & b1,b2 // b7,b8 & not b5,b6 _|_ b7,b8 & not b1 = b2 ) & ( b1,b2 _|_ b5,b6 & b1,b2 _|_ b5,b8 implies b1,b2 _|_ b6,b8 ) ) ) & ( for b1, b2, b3 being Element of a1 holds
not ( b1 <> b2 & ( for b4 being Element of a1 holds
not ( b1,b2 // b1,b4 & b1,b2 _|_ b4,b3 ) ) ) ) & ( for b1, b2, b3 being Element of a1 holds
ex b4 being Element of a1 st
( b1,b2 _|_ b3,b4 & b3 <> b4 ) ) );
end;

:: deftheorem Def9 defines OrtAfSp-like ANALMETR:def 9 :
for b1 being non empty ParOrtStr holds
( b1 is OrtAfSp-like iff ( AffinStruct(# the carrier of b1,the CONGR of b1 #) is AffinSpace & ( for b2, b3, b4, b5, b6, b7, b8, b9 being Element of b1 holds
( ( b2,b3 _|_ b2,b3 implies b2 = b3 ) & b2,b3 _|_ b4,b4 & ( b2,b3 _|_ b4,b5 implies ( b2,b3 _|_ b5,b4 & b4,b5 _|_ b2,b3 ) ) & not ( b2,b3 _|_ b6,b7 & b2,b3 // b8,b9 & not b6,b7 _|_ b8,b9 & not b2 = b3 ) & ( b2,b3 _|_ b6,b7 & b2,b3 _|_ b6,b9 implies b2,b3 _|_ b7,b9 ) ) ) & ( for b2, b3, b4 being Element of b1 holds
not ( b2 <> b3 & ( for b5 being Element of b1 holds
not ( b2,b3 // b2,b5 & b2,b3 _|_ b5,b4 ) ) ) ) & ( for b2, b3, b4 being Element of b1 holds
ex b5 being Element of b1 st
( b2,b3 _|_ b4,b5 & b4 <> b5 ) ) ) );

registration
cluster non empty strict OrtAfSp-like ParOrtStr ;
existence
ex b1 being non empty ParOrtStr st
( b1 is strict & b1 is OrtAfSp-like )
proof end;
end;

definition
mode OrtAfSp is non empty OrtAfSp-like ParOrtStr ;
end;

theorem Th43: :: ANALMETR:43
canceled;

theorem Th44: :: ANALMETR:44
for b1 being RealLinearSpace
for b2, b3 being VECTOR of b1 holds
( Gen b2,b3 implies AMSpace b1,b2,b3 is OrtAfSp )
proof end;

consider c4 being RealLinearSpace such that
Lemma43: ex b1, b2 being VECTOR of c4 st Gen b1,b2 by Th7;

consider c5, c6 being VECTOR of c4 such that
Lemma44: Gen c5,c6 by Lemma43;

E45: now
set c7 = AffinStruct(# the carrier of (AMSpace c4,c5,c6),the CONGR of (AMSpace c4,c5,c6) #);
AffinStruct(# the carrier of (AMSpace c4,c5,c6),the CONGR of (AMSpace c4,c5,c6) #) = Af (AMSpace c4,c5,c6) ;
then E46: AffinStruct(# the carrier of (AMSpace c4,c5,c6),the CONGR of (AMSpace c4,c5,c6) #) = Lambda (OASpace c4) by Th30;
( ( for b1, b2 being Real holds
( (b1 * c5) + (b2 * c6) = 0. c4 implies ( b1 = 0 & b2 = 0 ) ) ) & ( for b1 being VECTOR of c4 holds
ex b2, b3 being Real st b1 = (b2 * c5) + (b3 * c6) ) ) by Def1, Lemma44;
then OASpace c4 is OAffinPlane by ANALOAF:51;
hence ( AffinStruct(# the carrier of (AMSpace c4,c5,c6),the CONGR of (AMSpace c4,c5,c6) #) is AffinPlane & ( for b1, b2, b3, b4, b5, b6, b7, b8 being Element of (AMSpace c4,c5,c6) holds
( ( b1,b2 _|_ b1,b2 implies b1 = b2 ) & b1,b2 _|_ b3,b3 & ( b1,b2 _|_ b3,b4 implies ( b1,b2 _|_ b4,b3 & b3,b4 _|_ b1,b2 ) ) & not ( b1,b2 _|_ b5,b6 & b1,b2 // b7,b8 & not b5,b6 _|_ b7,b8 & not b1 = b2 ) & not ( b1,b2 _|_ b5,b6 & b1,b2 _|_ b7,b8 & not b5,b6 // b7,b8 & not b1 = b2 ) ) ) & ( for b1, b2, b3 being Element of (AMSpace c4,c5,c6) holds
ex b4 being Element of (AMSpace c4,c5,c6) st
( b1,b2 _|_ b3,b4 & b3 <> b4 ) ) ) by E46, Lemma44, Th33, Th34, Th35, Th36, Th37, Th38, Th40, DIRAF:53;
end;

definition
let c7 be non empty ParOrtStr ;
attr a1 is OrtAfPl-like means :Def10: :: ANALMETR:def 10
( AffinStruct(# the carrier of a1,the CONGR of a1 #) is AffinPlane & ( for b1, b2, b3, b4, b5, b6, b7, b8 being Element of a1 holds
( ( b1,b2 _|_ b1,b2 implies b1 = b2 ) & b1,b2 _|_ b3,b3 & ( b1,b2 _|_ b3,b4 implies ( b1,b2 _|_ b4,b3 & b3,b4 _|_ b1,b2 ) ) & not ( b1,b2 _|_ b5,b6 & b1,b2 // b7,b8 & not b5,b6 _|_ b7,b8 & not b1 = b2 ) & not ( b1,b2 _|_ b5,b6 & b1,b2 _|_ b7,b8 & not b5,b6 // b7,b8 & not b1 = b2 ) ) ) & ( for b1, b2, b3 being Element of a1 holds
ex b4 being Element of a1 st
( b1,b2 _|_ b3,b4 & b3 <> b4 ) ) );
end;

:: deftheorem Def10 defines OrtAfPl-like ANALMETR:def 10 :
for b1 being non empty ParOrtStr holds
( b1 is OrtAfPl-like iff ( AffinStruct(# the carrier of b1,the CONGR of b1 #) is AffinPlane & ( for b2, b3, b4, b5, b6, b7, b8, b9 being Element of b1 holds
( ( b2,b3 _|_ b2,b3 implies b2 = b3 ) & b2,b3 _|_ b4,b4 & ( b2,b3 _|_ b4,b5 implies ( b2,b3 _|_ b5,b4 & b4,b5 _|_ b2,b3 ) ) & not ( b2,b3 _|_ b6,b7 & b2,b3 // b8,b9 & not b6,b7 _|_ b8,b9 & not b2 = b3 ) & not ( b2,b3 _|_ b6,b7 & b2,b3 _|_ b8,b9 & not b6,b7 // b8,b9 & not b2 = b3 ) ) ) & ( for b2, b3, b4 being Element of b1 holds
ex b5 being Element of b1 st
( b2,b3 _|_ b4,b5 & b4 <> b5 ) ) ) );

registration
cluster non empty strict OrtAfPl-like ParOrtStr ;
existence
ex b1 being non empty ParOrtStr st
( b1 is strict & b1 is OrtAfPl-like )
proof end;
end;

definition
mode OrtAfPl is non empty OrtAfPl-like ParOrtStr ;
end;

theorem Th45: :: ANALMETR:45
canceled;

theorem Th46: :: ANALMETR:46
for b1 being RealLinearSpace
for b2, b3 being VECTOR of b1 holds
( Gen b2,b3 implies AMSpace b1,b2,b3 is OrtAfPl )
proof end;

theorem Th47: :: ANALMETR:47
for b1 being non empty ParOrtStr
for b2 being set holds
( b2 is Element of b1 iff b2 is Element of (Af b1) ) ;

theorem Th48: :: ANALMETR:48
for b1 being non empty ParOrtStr
for b2, b3, b4, b5 being Element of b1
for b6, b7, b8, b9 being Element of (Af b1) holds
( b2 = b6 & b3 = b7 & b4 = b8 & b5 = b9 implies ( b2,b3 // b4,b5 iff b6,b7 // b8,b9 ) )
proof end;

Lemma48: for b1 being OrtAfSp holds
Af b1 is AffinSpace
by Def9;

registration
let c7 be OrtAfSp;
cluster Af a1 -> non empty non trivial strict AffinSpace-like ;
correctness
coherence
( Af c7 is AffinSpace-like & not Af c7 is trivial )
;
by Lemma48;
end;

Lemma49: for b1 being OrtAfPl holds
Af b1 is AffinPlane
by Def10;

registration
let c7 be OrtAfPl;
cluster Af a1 -> non empty non trivial strict AffinSpace-like 2-dimensional ;
correctness
coherence
( Af c7 is 2-dimensional & Af c7 is AffinSpace-like & not Af c7 is trivial )
;
by Lemma49;
end;

theorem Th49: :: ANALMETR:49
for b1 being OrtAfPl holds
b1 is OrtAfSp
proof end;

registration
cluster non empty OrtAfPl-like -> non empty OrtAfSp-like ParOrtStr ;
coherence
for b1 being non empty ParOrtStr holds
( b1 is OrtAfPl-like implies b1 is OrtAfSp-like )
by Th49;
end;

theorem Th50: :: ANALMETR:50
for b1 being OrtAfSp holds
( Af b1 is AffinPlane implies b1 is OrtAfPl )
proof end;

theorem Th51: :: ANALMETR:51
for b1 being non empty ParOrtStr holds
( b1 is OrtAfPl-like iff ( not for b2, b3 being Element of b1 holds not b2 <> b3 & ( for b2, b3, b4, b5, b6, b7, b8, b9 being Element of b1 holds
( b2,b3 // b3,b2 & b2,b3 // b4,b4 & not ( b2,b3 // b6,b7 & b2,b3 // b8,b9 & not b6,b7 // b8,b9 & not b2 = b3 ) & ( b2,b3 // b2,b4 implies b3,b2 // b3,b4 ) & ex b10 being Element of b1 st
( b2,b3 // b4,b10 & b2,b4 // b3,b10 ) & not for b10, b11, b12 being Element of b1 holds b10,b11 // b10,b12 & ex b10 being Element of b1 st
( b2,b3 // b4,b10 & b4 <> b10 ) & not ( b2,b3 // b3,b5 & b3 <> b2 & ( for b10 being Element of b1 holds
not ( b4,b3 // b3,b10 & b4,b2 // b5,b10 ) ) ) & ( b2,b3 _|_ b2,b3 implies b2 = b3 ) & b2,b3 _|_ b4,b4 & ( b2,b3 _|_ b4,b5 implies ( b2,b3 _|_ b5,b4 & b4,b5 _|_ b2,b3 ) ) & not ( b2,b3 _|_ b6,b7 & b2,b3 // b8,b9 & not b6,b7 _|_ b8,b9 & not b2 = b3 ) & not ( b2,b3 _|_ b6,b7 & b2,b3 _|_ b8,b9 & not b6,b7 // b8,b9 & not b2 = b3 ) & ex b10 being Element of b1 st
( b2,b3 _|_ b4,b10 & b4 <> b10 ) & not ( not b2,b3 // b4,b5 & ( for b10 being Element of b1 holds
not ( b2,b3 // b2,b10 & b4,b5 // b4,b10 ) ) ) ) ) ) )
proof end;

definition
let c7 be non empty ParOrtStr ;
let c8, c9, c10 be Element of c7;
pred LIN c2,c3,c4 means :Def11: :: ANALMETR:def 11
a2,a3 // a2,a4;
end;

:: deftheorem Def11 defines LIN ANALMETR:def 11 :
for b1 being non empty ParOrtStr
for b2, b3, b4 being Element of b1 holds
( LIN b2,b3,b4 iff b2,b3 // b2,b4 );

definition
let c7 be non empty ParOrtStr ;
let c8, c9 be Element of c7;
func Line c2,c3 -> Subset of a1 means :Def12: :: ANALMETR:def 12
for b1 being Element of a1 holds
( b1 in a4 iff LIN a2,a3,b1 );
existence
ex b1 being Subset of c7 st
for b2 being Element of c7 holds
( b2 in b1 iff LIN c8,c9,b2 )
proof end;
uniqueness
for b1, b2 being Subset of c7 holds
( ( for b3 being Element of c7 holds
( b3 in b1 iff LIN c8,c9,b3 ) ) & ( for b3 being Element of c7 holds
( b3 in b2 iff LIN c8,c9,b3 ) ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def12 defines Line ANALMETR:def 12 :
for b1 being non empty ParOrtStr
for b2, b3 being Element of b1
for b4 being Subset of b1 holds
( b4 = Line b2,b3 iff for b5 being Element of b1 holds
( b5 in b4 iff LIN b2,b3,b5 ) );

definition
let c7 be non empty ParOrtStr ;
let c8 be Subset of c7;
attr a2 is being_line means :Def13: :: ANALMETR:def 13
ex b1, b2 being Element of a1 st
( b1 <> b2 & a2 = Line b1,b2 );
end;

:: deftheorem Def13 defines being_line ANALMETR:def 13 :
for b1 being non empty ParOrtStr
for b2 being Subset of b1 holds
( b2 is being_line iff ex b3, b4 being Element of b1 st
( b3 <> b4 & b2 = Line b3,b4 ) );

notation
let c7 be non empty ParOrtStr ;
let c8 be Subset of c7;
synonym c2 is_line for being_line c2c2 is_lineis_line ;
end;

theorem Th52: :: ANALMETR:52
canceled;

theorem Th53: :: ANALMETR:53
canceled;

theorem Th54: :: ANALMETR:54
canceled;

theorem Th55: :: ANALMETR:55
for b1 being OrtAfSp
for b2, b3, b4 being Element of b1
for b5, b6, b7 being Element of (Af b1) holds
( b2 = b5 & b3 = b6 & b4 = b7 implies ( LIN b2,b3,b4 iff LIN b5,b6,b7 ) )
proof end;

theorem Th56: :: ANALMETR:56
for b1 being OrtAfSp
for b2, b3 being Element of b1
for b4, b5 being Element of (Af b1) holds
( b2 = b4 & b3 = b5 implies Line b2,b3 = Line b4,b5 )
proof end;

theorem Th57: :: ANALMETR:57
for b1 being non empty ParOrtStr
for b2 being set holds
( b2 is Subset of b1 iff b2 is Subset of (Af b1) ) ;

theorem Th58: :: ANALMETR:58
for b1 being OrtAfSp
for b2 being Subset of b1
for b3 being Subset of (Af b1) holds
( b2 = b3 implies ( b2 is being_line iff b3 is_line ) )
proof end;

definition
let c7 be non empty ParOrtStr ;
let c8, c9 be Element of c7;
let c10 be Subset of c7;
pred c2,c3 _|_ c4 means :Def14: :: ANALMETR:def 14
ex b1, b2 being Element of a1 st
( b1 <> b2 & a4 = Line b1,b2 & a2,a3 _|_ b1,b2 );
end;

:: deftheorem Def14 defines _|_ ANALMETR:def 14 :
for b1 being non empty ParOrtStr
for b2, b3 being Element of b1
for b4 being Subset of b1 holds
( b2,b3 _|_ b4 iff ex b5, b6 being Element of b1 st
( b5 <> b6 & b4 = Line b5,b6 & b2,b3 _|_ b5,b6 ) );

definition
let c7 be non empty ParOrtStr ;
let c8, c9 be Subset of c7;
pred c2 _|_ c3 means :Def15: :: ANALMETR:def 15
ex b1, b2 being Element of a1 st
( b1 <> b2 & a2 = Line b1,b2 & b1,b2 _|_ a3 );
end;

:: deftheorem Def15 defines _|_ ANALMETR:def 15 :
for b1 being non empty ParOrtStr
for b2, b3 being Subset of b1 holds
( b2 _|_ b3 iff ex b4, b5 being Element of b1 st
( b4 <> b5 & b2 = Line b4,b5 & b4,b5 _|_ b3 ) );

definition
let c7 be non empty ParOrtStr ;
let c8, c9 be Subset of c7;
pred c2 // c3 means :Def16: :: ANALMETR:def 16
ex b1, b2, b3, b4 being Element of a1 st
( b1 <> b2 & b3 <> b4 & a2 = Line b1,b2 & a3 = Line b3,b4 & b1,b2 // b3,b4 );
end;

:: deftheorem Def16 defines // ANALMETR:def 16 :
for b1 being non empty ParOrtStr
for b2, b3 being Subset of b1 holds
( b2 // b3 iff ex b4, b5, b6, b7 being Element of b1 st
( b4 <> b5 & b6 <> b7 & b2 = Line b4,b5 & b3 = Line b6,b7 & b4,b5 // b6,b7 ) );

theorem Th59: :: ANALMETR:59
canceled;

theorem Th60: :: ANALMETR:60
canceled;

theorem Th61: :: ANALMETR:61
canceled;

theorem Th62: :: ANALMETR:62
for b1 being non empty ParOrtStr
for b2, b3 being Element of b1
for b4, b5 being Subset of b1 holds
( ( b2,b3 _|_ b4 implies b4 is being_line ) & ( b4 _|_ b5 implies ( b4 is being_line & b5 is being_line ) ) )
proof end;

theorem Th63: :: ANALMETR:63
for b1 being non empty ParOrtStr
for b2, b3 being Subset of b1 holds
( b2 _|_ b3 iff ex b4, b5, b6, b7 being Element of b1 st
( b4 <> b5 & b6 <> b7 & b2 = Line b4,b5 & b3 = Line b6,b7 & b4,b5 _|_ b6,b7 ) )
proof end;

theorem Th64: :: ANALMETR:64
for b1 being OrtAfSp
for b2, b3 being Subset of b1
for b4, b5 being Subset of (Af b1) holds
( b2 = b4 & b3 = b5 implies ( b2 // b3 iff b4 // b5 ) )
proof end;

theorem Th65: :: ANALMETR:65
for b1 being OrtAfSp
for b2 being Subset of b1
for b3 being Element of b1 holds
( b2 is being_line implies b3,b3 _|_ b2 )
proof end;

theorem Th66: :: ANALMETR:66
for b1 being OrtAfSp
for b2 being Subset of b1
for b3, b4, b5, b6 being Element of b1 holds
( b3,b4 _|_ b2 & ( b3,b4 // b5,b6 or b5,b6 // b3,b4 ) & b3 <> b4 implies b5,b6 _|_ b2 )
proof end;

theorem Th67: :: ANALMETR:67
for b1 being OrtAfSp
for b2 being Subset of b1
for b3, b4 being Element of b1 holds
( b3,b4 _|_ b2 implies b4,b3 _|_ b2 )
proof end;

definition
let c7 be OrtAfSp;
let c8, c9 be Subset of c7;
redefine pred // as c2 // c3;
symmetry
for b1, b2 being Subset of c7 holds
( b1 // b2 implies b2 // b1 )
proof end;
end;

theorem Th68: :: ANALMETR:68
canceled;

theorem Th69: :: ANALMETR:69
for b1 being OrtAfSp
for b2, b3 being Subset of b1
for b4, b5 being Element of b1 holds
( b4,b5 _|_ b2 & b2 // b3 implies b4,b5 _|_ b3 )
proof end;

theorem Th70: :: ANALMETR:70
canceled;

theorem Th71: :: ANALMETR:71
for b1 being OrtAfSp
for b2 being Subset of b1
for b3, b4 being Element of b1 holds
( b3 in b2 & b4 in b2 & b3,b4 _|_ b2 implies b3 = b4 )
proof end;

definition
let c7 be OrtAfSp;
let c8, c9 be Subset of c7;
redefine pred _|_ as c2 _|_ c3;
irreflexivity
for b1 being Subset of c7 holds
not b1 _|_ b1
proof end;
symmetry
for b1, b2 being Subset of c7 holds
( b1 _|_ b2 implies b2 _|_ b1 )
proof end;
end;

theorem Th72: :: ANALMETR:72
canceled;

theorem Th73: :: ANALMETR:73
for b1 being OrtAfSp
for b2, b3, b4 being Subset of b1 holds
( b2 _|_ b3 & b2 // b4 implies b4 _|_ b3 )
proof end;

theorem Th74: :: ANALMETR:74
canceled;

theorem Th75: :: ANALMETR:75
for b1 being OrtAfSp
for b2 being Subset of b1
for b3, b4, b5, b6 being Element of b1 holds
( b3 in b2 & b4 in b2 & b5,b6 _|_ b2 implies ( b5,b6 _|_ b3,b4 & b3,b4 _|_ b5,b6 ) )
proof end;

theorem Th76: :: ANALMETR:76
for b1 being OrtAfSp
for b2 being Subset of b1
for b3, b4 being Element of b1 holds
( b3 in b2 & b4 in b2 & b3 <> b4 & b2 is being_line implies b2 = Line b3,b4 )
proof end;

theorem Th77: :: ANALMETR:77
for b1 being OrtAfSp
for b2 being Subset of b1
for b3, b4, b5, b6 being Element of b1 holds
( b3 in b2 & b4 in b2 & b3 <> b4 & b2 is being_line & ( b3,b4 _|_ b5,b6 or b5,b6 _|_ b3,b4 ) implies b5,b6 _|_ b2 )
proof end;

theorem Th78: :: ANALMETR:78
for b1 being OrtAfSp
for b2, b3 being Subset of b1
for b4, b5, b6, b7 being Element of b1 holds
( b4 in b2 & b5 in b2 & b6 in b3 & b7 in b3 & b2 _|_ b3 implies b4,b5 _|_ b6,b7 )
proof end;

theorem Th79: :: ANALMETR:79
for b1 being OrtAfSp
for b2, b3, b4, b5 being Subset of b1
for b6, b7, b8 being Element of b1 holds
( b6 in b2 & b6 in b3 & b7 in b2 & b8 in b3 & b7 <> b8 & b7 in b4 & b8 in b4 & b5 _|_ b2 & b5 _|_ b3 & b4 is being_line implies b5 _|_ b4 )
proof end;

theorem Th80: :: ANALMETR:80
for b1 being OrtAfSp
for b2, b3, b4 being Element of b1 holds
( b2,b3 _|_ b4,b4 & b4,b4 _|_ b2,b3 & b2,b3 // b4,b4 & b4,b4 // b2,b3 )
proof end;

theorem Th81: :: ANALMETR:81
for b1 being OrtAfSp
for b2, b3, b4, b5 being Element of b1 holds
( b2,b3 // b4,b5 implies ( b2,b3 // b5,b4 & b3,b2 // b4,b5 & b3,b2 // b5,b4 & b4,b5 // b2,b3 & b4,b5 // b3,b2 & b5,b4 // b2,b3 & b5,b4 // b3,b2 ) )
proof end;

theorem Th82: :: ANALMETR:82
for b1 being OrtAfSp
for b2, b3, b4, b5, b6, b7 being Element of b1 holds
( b2 <> b3 & not ( not ( b2,b3 // b4,b5 & b2,b3 // b6,b7 ) & not ( b2,b3 // b4,b5 & b6,b7 // b2,b3 ) & not ( b4,b5 // b2,b3 & b6,b7 // b2,b3 ) & not ( b4,b5 // b2,b3 & b2,b3 // b6,b7 ) ) implies b4,b5 // b6,b7 )
proof end;

theorem Th83: :: ANALMETR:83
for b1 being OrtAfSp
for b2, b3, b4, b5 being Element of b1 holds
( b2,b3 _|_ b4,b5 implies ( b2,b3 _|_ b5,b4 & b3,b2 _|_ b4,b5 & b3,b2 _|_ b5,b4 & b4,b5 _|_ b2,b3 & b4,b5 _|_ b3,b2 & b5,b4 _|_ b2,b3 & b5,b4 _|_ b3,b2 ) )
proof end;

theorem Th84: :: ANALMETR:84
for b1 being OrtAfSp
for b2, b3, b4, b5, b6, b7 being Element of b1 holds
( b2 <> b3 & not ( not ( b2,b3 // b4,b5 & b2,b3 _|_ b6,b7 ) & not ( b2,b3 // b6,b7 & b2,b3 _|_ b4,b5 ) & not ( b2,b3 // b4,b5 & b6,b7 _|_ b2,b3 ) & not ( b2,b3 // b6,b7 & b4,b5 _|_ b2,b3 ) & not ( b4,b5 // b2,b3 & b6,b7 _|_ b2,b3 ) & not ( b6,b7 // b2,b3 & b4,b5 _|_ b2,b3 ) & not ( b4,b5 // b2,b3 & b2,b3 _|_ b6,b7 ) & not ( b6,b7 // b2,b3 & b2,b3 _|_ b4,b5 ) ) implies b4,b5 _|_ b6,b7 )
proof end;

theorem Th85: :: ANALMETR:85
for b1 being OrtAfPl
for b2, b3, b4, b5, b6, b7 being Element of b1 holds
( b2 <> b3 & not ( not ( b2,b3 _|_ b4,b5 & b2,b3 _|_ b6,b7 ) & not ( b2,b3 _|_ b4,b5 & b6,b7 _|_ b2,b3 ) & not ( b4,b5 _|_ b2,b3 & b6,b7 _|_ b2,b3 ) & not ( b4,b5 _|_ b2,b3 & b2,b3 _|_ b6,b7 ) ) implies b4,b5 // b6,b7 )
proof end;

theorem Th86: :: ANALMETR:86
for b1 being OrtAfPl
for b2, b3 being Subset of b1
for b4, b5, b6, b7 being Element of b1 holds
( b4 in b2 & b5 in b2 & b4 <> b5 & b2 is being_line & b6 in b3 & b7 in b3 & b6 <> b7 & b3 is being_line & b4,b5 // b6,b7 implies b2 // b3 )
proof end;

theorem Th87: :: ANALMETR:87
for b1 being OrtAfPl
for b2, b3, b4 being Subset of b1 holds
( b2 _|_ b3 & b4 _|_ b3 implies b2 // b4 )
proof end;

theorem Th88: :: ANALMETR:88
for b1 being OrtAfPl
for b2, b3 being Subset of b1 holds
not ( b2 _|_ b3 & ( for b4 being Element of b1 holds
not ( b4 in b2 & b4 in b3 ) ) )
proof end;

theorem Th89: :: ANALMETR:89
for b1 being OrtAfPl
for b2, b3, b4, b5 being Element of b1 holds
not ( b2,b3 _|_ b4,b5 & ( for b6 being Element of b1 holds
not ( LIN b2,b3,b6 & LIN b4,b5,b6 ) ) )
proof end;

theorem Th90: :: ANALMETR:90
for b1 being OrtAfPl
for b2 being Subset of b1
for b3, b4 being Element of b1 holds
not ( b3,b4 _|_ b2 & ( for b5 being Element of b1 holds
not ( LIN b3,b4,b5 & b5 in b2 ) ) )
proof end;

theorem Th91: :: ANALMETR:91
for b1 being OrtAfPl
for b2, b3, b4 being Element of b1 holds
ex b5 being Element of b1 st
( b2,b5 _|_ b3,b4 & LIN b3,b4,b5 )
proof end;

theorem Th92: :: ANALMETR:92
for b1 being OrtAfPl
for b2 being Subset of b1
for b3 being Element of b1 holds
not ( b2 is being_line & ( for b4 being Element of b1 holds
not ( b3,b4 _|_ b2 & b4 in b2 ) ) )
proof end;