:: ANALMETR semantic presentation
Lemma37:
for V being RealLinearSpace
for v1, w, y, v2 being VECTOR of V
for b1, b2, c1, c2 being Real st v1 = (b1 * w) + (b2 * y) & v2 = (c1 * w) + (c2 * y) holds
( v1 + v2 = ((b1 + c1) * w) + ((b2 + c2) * y) & v1 - v2 = ((b1 - c1) * w) + ((b2 - c2) * y) )
Lemma40:
for V being RealLinearSpace
for w, y being VECTOR of V holds (0 * w) + (0 * y) = 0. V
Lemma41:
for V being RealLinearSpace
for v, w, y being VECTOR of V
for b1, b2, a being Real st v = (b1 * w) + (b2 * y) holds
a * v = ((a * b1) * w) + ((a * b2) * y)
:: deftheorem Def1 defines Gen ANALMETR:def 1 :
:: deftheorem Def2 defines are_Ort_wrt ANALMETR:def 2 :
Lemma44:
for V being RealLinearSpace
for w, y being VECTOR of V
for a1, a2, b1, b2 being Real st Gen w,y & (a1 * w) + (a2 * y) = (b1 * w) + (b2 * y) holds
( a1 = b1 & a2 = b2 )
theorem Th1: :: ANALMETR:1
canceled;
theorem Th2: :: ANALMETR:2
canceled;
theorem Th3: :: ANALMETR:3
canceled;
theorem Th4: :: ANALMETR:4
canceled;
theorem Th5: :: ANALMETR:5
Lemma54:
for V being RealLinearSpace
for w, y being VECTOR of V st Gen w,y holds
( w <> 0. V & y <> 0. V )
theorem Th6: :: ANALMETR:6
theorem Th7: :: ANALMETR:7
theorem Th8: :: ANALMETR:8
theorem Th9: :: ANALMETR:9
theorem Th10: :: ANALMETR:10
theorem Th11: :: ANALMETR:11
for
V being
RealLinearSpace for
u,
v,
w,
y being
VECTOR of
V for
a,
b being
Real st
u,
v are_Ort_wrt w,
y holds
(
a * u,
v are_Ort_wrt w,
y &
u,
b * v are_Ort_wrt w,
y )
theorem Th12: :: ANALMETR:12
theorem Th13: :: ANALMETR:13
for
V being
RealLinearSpace for
w,
y,
v,
u1,
u2 being
VECTOR of
V st
Gen w,
y &
v,
u1 are_Ort_wrt w,
y &
v,
u2 are_Ort_wrt w,
y &
v <> 0. V holds
ex
a,
b being
Real st
(
a * u1 = b * u2 & (
a <> 0 or
b <> 0 ) )
theorem Th14: :: ANALMETR:14
for
V being
RealLinearSpace for
w,
y,
u,
v1,
v2 being
VECTOR of
V st
Gen w,
y &
u,
v1 are_Ort_wrt w,
y &
u,
v2 are_Ort_wrt w,
y holds
(
u,
v1 + v2 are_Ort_wrt w,
y &
u,
v1 - v2 are_Ort_wrt w,
y )
theorem Th15: :: ANALMETR:15
theorem Th16: :: ANALMETR:16
for
V being
RealLinearSpace for
w,
y,
u,
u1,
u2 being
VECTOR of
V st
Gen w,
y &
u,
u1 - u2 are_Ort_wrt w,
y &
u1,
u2 - u are_Ort_wrt w,
y holds
u2,
u - u1 are_Ort_wrt w,
y
theorem Th17: :: ANALMETR:17
theorem Th18: :: ANALMETR:18
theorem Th19: :: ANALMETR:19
:: deftheorem Def3 defines are_Ort_wrt ANALMETR:def 3 :
for
V being
RealLinearSpace for
u,
u1,
v,
v1,
w,
y being
VECTOR of
V holds
(
u,
u1,
v,
v1 are_Ort_wrt w,
y iff
u1 - u,
v1 - v are_Ort_wrt w,
y );
definition
let V be
RealLinearSpace;
let w be
VECTOR of
V;
let y be
VECTOR of
V;
func Orthogonality c1,
c2,
c3 -> Relation of
[:the carrier of a1,the carrier of a1:] means :
Def4:
:: ANALMETR:def 4
for
x,
z being
set holds
(
[x,z] in it iff ex
u,
u1,
v,
v1 being
VECTOR of
V st
(
x = [u,u1] &
z = [v,v1] &
u,
u1,
v,
v1 are_Ort_wrt w,
y ) );
existence
ex b1 being Relation of [:the carrier of V,the carrier of V:] st
for x, z being set holds
( [x,z] in b1 iff ex u, u1, v, v1 being VECTOR of V st
( x = [u,u1] & z = [v,v1] & u,u1,v,v1 are_Ort_wrt w,y ) )
uniqueness
for b1, b2 being Relation of [:the carrier of V,the carrier of V:] st ( for x, z being set holds
( [x,z] in b1 iff ex u, u1, v, v1 being VECTOR of V st
( x = [u,u1] & z = [v,v1] & u,u1,v,v1 are_Ort_wrt w,y ) ) ) & ( for x, z being set holds
( [x,z] in b2 iff ex u, u1, v, v1 being VECTOR of V st
( x = [u,u1] & z = [v,v1] & u,u1,v,v1 are_Ort_wrt w,y ) ) ) holds
b1 = b2
end;
:: deftheorem Def4 defines Orthogonality ANALMETR:def 4 :
for
V being
RealLinearSpace for
w,
y being
VECTOR of
V for
b4 being
Relation of
[:the carrier of V,the carrier of V:] holds
(
b4 = Orthogonality V,
w,
y iff for
x,
z being
set holds
(
[x,z] in b4 iff ex
u,
u1,
v,
v1 being
VECTOR of
V st
(
x = [u,u1] &
z = [v,v1] &
u,
u1,
v,
v1 are_Ort_wrt w,
y ) ) );
theorem Th20: :: ANALMETR:20
canceled;
theorem Th21: :: ANALMETR:21
canceled;
theorem Th22: :: ANALMETR:22
theorem Th23: :: ANALMETR:23
theorem Th24: :: ANALMETR:24
:: deftheorem Def5 ANALMETR:def 5 :
canceled;
:: deftheorem Def6 defines _|_ ANALMETR:def 6 :
definition
let V be
RealLinearSpace;
let w be
VECTOR of
V;
let y be
VECTOR of
V;
func AMSpace c1,
c2,
c3 -> strict ParOrtStr equals :: ANALMETR:def 7
ParOrtStr(# the
carrier of
V,
(lambda (DirPar V)),
(Orthogonality V,w,y) #);
correctness
coherence
ParOrtStr(# the carrier of V,(lambda (DirPar V)),(Orthogonality V,w,y) #) is strict ParOrtStr ;
;
end;
:: deftheorem Def7 defines AMSpace ANALMETR:def 7 :
theorem Th25: :: ANALMETR:25
canceled;
theorem Th26: :: ANALMETR:26
canceled;
theorem Th27: :: ANALMETR:27
canceled;
theorem Th28: :: ANALMETR:28
:: deftheorem Def8 defines Af ANALMETR:def 8 :
theorem Th29: :: ANALMETR:29
canceled;
theorem Th30: :: ANALMETR:30
theorem Th31: :: ANALMETR:31
for
V being
RealLinearSpace for
u,
u1,
v,
v1,
w,
y being
VECTOR of
V for
p,
p1,
q,
q1 being
Element of
(AMSpace V,w,y) st
p = u &
p1 = u1 &
q = v &
q1 = v1 holds
(
p,
q _|_ p1,
q1 iff
u,
v,
u1,
v1 are_Ort_wrt w,
y )
theorem Th32: :: ANALMETR:32
for
V being
RealLinearSpace for
w,
y,
u,
v,
u1,
v1 being
VECTOR of
V for
p,
q,
p1,
q1 being
Element of
(AMSpace V,w,y) st
p = u &
q = v &
p1 = u1 &
q1 = v1 holds
(
p,
q // p1,
q1 iff ex
a,
b being
Real st
(
a * (v - u) = b * (v1 - u1) & (
a <> 0 or
b <> 0 ) ) )
theorem Th33: :: ANALMETR:33
for
V being
RealLinearSpace for
w,
y being
VECTOR of
V for
p,
q,
p1,
q1 being
Element of
(AMSpace V,w,y) st
p,
q _|_ p1,
q1 holds
p1,
q1 _|_ p,
q
theorem Th34: :: ANALMETR:34
for
V being
RealLinearSpace for
w,
y being
VECTOR of
V for
p,
q,
p1,
q1 being
Element of
(AMSpace V,w,y) st
p,
q _|_ p1,
q1 holds
p,
q _|_ q1,
p1
theorem Th35: :: ANALMETR:35
theorem Th36: :: ANALMETR:36
for
V being
RealLinearSpace for
w,
y being
VECTOR of
V for
p,
p1,
q,
q1,
r,
r1 being
Element of
(AMSpace V,w,y) st
p,
p1 _|_ q,
q1 &
p,
p1 // r,
r1 & not
p = p1 holds
q,
q1 _|_ r,
r1
theorem Th37: :: ANALMETR:37
for
V being
RealLinearSpace for
w,
y being
VECTOR of
V st
Gen w,
y holds
for
p,
q,
r being
Element of
(AMSpace V,w,y) ex
r1 being
Element of
(AMSpace V,w,y) st
(
p,
q _|_ r,
r1 &
r <> r1 )
theorem Th38: :: ANALMETR:38
for
V being
RealLinearSpace for
w,
y being
VECTOR of
V for
p,
p1,
q,
q1,
r,
r1 being
Element of
(AMSpace V,w,y) st
Gen w,
y &
p,
p1 _|_ q,
q1 &
p,
p1 _|_ r,
r1 & not
p = p1 holds
q,
q1 // r,
r1
theorem Th39: :: ANALMETR:39
for
V being
RealLinearSpace for
w,
y being
VECTOR of
V for
p,
q,
r,
r1,
r2 being
Element of
(AMSpace V,w,y) st
Gen w,
y &
p,
q _|_ r,
r1 &
p,
q _|_ r,
r2 holds
p,
q _|_ r1,
r2
theorem Th40: :: ANALMETR:40
theorem Th41: :: ANALMETR:41
for
V being
RealLinearSpace for
w,
y being
VECTOR of
V for
p,
q,
p1,
p2 being
Element of
(AMSpace V,w,y) st
Gen w,
y &
p,
q _|_ p1,
p2 &
p1,
q _|_ p2,
p holds
p2,
q _|_ p,
p1
theorem Th42: :: ANALMETR:42
for
V being
RealLinearSpace for
w,
y being
VECTOR of
V for
p,
p1 being
Element of
(AMSpace V,w,y) st
Gen w,
y &
p <> p1 holds
for
q being
Element of
(AMSpace V,w,y) ex
q1 being
Element of
(AMSpace V,w,y) st
(
p,
p1 // p,
q1 &
p,
p1 _|_ q1,
q )
consider V0 being RealLinearSpace such that
Lemma130:
ex w, y being VECTOR of V0 st Gen w,y
by ;
consider w0 being VECTOR of V0, y0 being VECTOR of V0 such that
Lemma133:
Gen w0,y0
by ;
E134:
now
set X =
AffinStruct(# the
carrier of
(AMSpace V0,w0,y0),the
CONGR of
(AMSpace V0,w0,y0) #);
AffinStruct(# the
carrier of
(AMSpace V0,w0,y0),the
CONGR of
(AMSpace V0,w0,y0) #)
= Af (AMSpace V0,w0,y0)
;
then E38:
AffinStruct(# the
carrier of
(AMSpace V0,w0,y0),the
CONGR of
(AMSpace V0,w0,y0) #)
= Lambda (OASpace V0)
by ;
for
a,
b being
Real st
(a * w0) + (b * y0) = 0. V0 holds
(
a = 0 &
b = 0 )
by , Lemma40;
then
OASpace V0 is
OAffinSpace
by ANALOAF:38;
hence
(
AffinStruct(# the
carrier of
(AMSpace V0,w0,y0),the
CONGR of
(AMSpace V0,w0,y0) #) is
AffinSpace & ( for
a,
b,
c,
d,
p,
q,
r,
s being
Element of
(AMSpace V0,w0,y0) holds
( (
a,
b _|_ a,
b implies
a = b ) &
a,
b _|_ c,
c & (
a,
b _|_ c,
d implies (
a,
b _|_ d,
c &
c,
d _|_ a,
b ) ) & (
a,
b _|_ p,
q &
a,
b // r,
s & not
p,
q _|_ r,
s implies
a = b ) & (
a,
b _|_ p,
q &
a,
b _|_ p,
s implies
a,
b _|_ q,
s ) ) ) & ( for
a,
b,
c being
Element of
(AMSpace V0,w0,y0) st
a <> b holds
ex
x being
Element of
(AMSpace V0,w0,y0) st
(
a,
b // a,
x &
a,
b _|_ x,
c ) ) & ( for
a,
b,
c being
Element of
(AMSpace V0,w0,y0) ex
x being
Element of
(AMSpace V0,w0,y0) st
(
a,
b _|_ c,
x &
c <> x ) ) )
by Def1, Lemma40, , , , , , , Lemma37, E38, DIRAF:48;
end;
definition
let IT be non
empty ParOrtStr ;
attr a1 is
OrtAfSp-like means :
Def9:
:: ANALMETR:def 9
(
AffinStruct(# the
carrier of
V0,the
CONGR of
V0 #) is
AffinSpace & ( for
a,
b,
c,
d,
p,
q,
r,
s being
Element of
V0 holds
( (
a,
b _|_ a,
b implies
a = b ) &
a,
b _|_ c,
c & (
a,
b _|_ c,
d implies (
a,
b _|_ d,
c &
c,
d _|_ a,
b ) ) & (
a,
b _|_ p,
q &
a,
b // r,
s & not
p,
q _|_ r,
s implies
a = b ) & (
a,
b _|_ p,
q &
a,
b _|_ p,
s implies
a,
b _|_ q,
s ) ) ) & ( for
a,
b,
c being
Element of
V0 st
a <> b holds
ex
x being
Element of
V0 st
(
a,
b // a,
x &
a,
b _|_ x,
c ) ) & ( for
a,
b,
c being
Element of
V0 ex
x being
Element of
V0 st
(
a,
b _|_ c,
x &
c <> x ) ) );
end;
:: deftheorem Def9 defines OrtAfSp-like ANALMETR:def 9 :
for
IT being non
empty ParOrtStr holds
(
IT is
OrtAfSp-like iff (
AffinStruct(# the
carrier of
IT,the
CONGR of
IT #) is
AffinSpace & ( for
a,
b,
c,
d,
p,
q,
r,
s being
Element of
IT holds
( (
a,
b _|_ a,
b implies
a = b ) &
a,
b _|_ c,
c & (
a,
b _|_ c,
d implies (
a,
b _|_ d,
c &
c,
d _|_ a,
b ) ) & (
a,
b _|_ p,
q &
a,
b // r,
s & not
p,
q _|_ r,
s implies
a = b ) & (
a,
b _|_ p,
q &
a,
b _|_ p,
s implies
a,
b _|_ q,
s ) ) ) & ( for
a,
b,
c being
Element of
IT st
a <> b holds
ex
x being
Element of
IT st
(
a,
b // a,
x &
a,
b _|_ x,
c ) ) & ( for
a,
b,
c being
Element of
IT ex
x being
Element of
IT st
(
a,
b _|_ c,
x &
c <> x ) ) ) );
theorem Th43: :: ANALMETR:43
canceled;
theorem Th44: :: ANALMETR:44
consider V0 being RealLinearSpace such that
Lemma139:
ex w, y being VECTOR of V0 st Gen w,y
by ;
consider w0 being VECTOR of V0, y0 being VECTOR of V0 such that
Lemma140:
Gen w0,y0
by Def2;
E141:
now
set X =
AffinStruct(# the
carrier of
(AMSpace V0,w0,y0),the
CONGR of
(AMSpace V0,w0,y0) #);
AffinStruct(# the
carrier of
(AMSpace V0,w0,y0),the
CONGR of
(AMSpace V0,w0,y0) #)
= Af (AMSpace V0,w0,y0)
;
then E38:
AffinStruct(# the
carrier of
(AMSpace V0,w0,y0),the
CONGR of
(AMSpace V0,w0,y0) #)
= Lambda (OASpace V0)
by ;
( ( for
a,
b being
Real st
(a * w0) + (b * y0) = 0. V0 holds
(
a = 0 &
b = 0 ) ) & ( for
w1 being
VECTOR of
V0 ex
a,
b being
Real st
w1 = (a * w0) + (b * y0) ) )
by , Lemma44;
then
OASpace V0 is
OAffinPlane
by ANALOAF:51;
hence
(
AffinStruct(# the
carrier of
(AMSpace V0,w0,y0),the
CONGR of
(AMSpace V0,w0,y0) #) is
AffinPlane & ( for
a,
b,
c,
d,
p,
q,
r,
s being
Element of
(AMSpace V0,w0,y0) holds
( (
a,
b _|_ a,
b implies
a = b ) &
a,
b _|_ c,
c & (
a,
b _|_ c,
d implies (
a,
b _|_ d,
c &
c,
d _|_ a,
b ) ) & (
a,
b _|_ p,
q &
a,
b // r,
s & not
p,
q _|_ r,
s implies
a = b ) & (
a,
b _|_ p,
q &
a,
b _|_ r,
s & not
p,
q // r,
s implies
a = b ) ) ) & ( for
a,
b,
c being
Element of
(AMSpace V0,w0,y0) ex
x being
Element of
(AMSpace V0,w0,y0) st
(
a,
b _|_ c,
x &
c <> x ) ) )
by , Lemma44, , , , , , , Lemma37, DIRAF:53;
end;
definition
let IT be non
empty ParOrtStr ;
attr a1 is
OrtAfPl-like means :
Def10:
:: ANALMETR:def 10
(
AffinStruct(# the
carrier of
V0,the
CONGR of
V0 #) is
AffinPlane & ( for
a,
b,
c,
d,
p,
q,
r,
s being
Element of
V0 holds
( (
a,
b _|_ a,
b implies
a = b ) &
a,
b _|_ c,
c & (
a,
b _|_ c,
d implies (
a,
b _|_ d,
c &
c,
d _|_ a,
b ) ) & (
a,
b _|_ p,
q &
a,
b // r,
s & not
p,
q _|_ r,
s implies
a = b ) & (
a,
b _|_ p,
q &
a,
b _|_ r,
s & not
p,
q // r,
s implies
a = b ) ) ) & ( for
a,
b,
c being
Element of
V0 ex
x being
Element of
V0 st
(
a,
b _|_ c,
x &
c <> x ) ) );
end;
:: deftheorem Def10 defines OrtAfPl-like ANALMETR:def 10 :
for
IT being non
empty ParOrtStr holds
(
IT is
OrtAfPl-like iff (
AffinStruct(# the
carrier of
IT,the
CONGR of
IT #) is
AffinPlane & ( for
a,
b,
c,
d,
p,
q,
r,
s being
Element of
IT holds
( (
a,
b _|_ a,
b implies
a = b ) &
a,
b _|_ c,
c & (
a,
b _|_ c,
d implies (
a,
b _|_ d,
c &
c,
d _|_ a,
b ) ) & (
a,
b _|_ p,
q &
a,
b // r,
s & not
p,
q _|_ r,
s implies
a = b ) & (
a,
b _|_ p,
q &
a,
b _|_ r,
s & not
p,
q // r,
s implies
a = b ) ) ) & ( for
a,
b,
c being
Element of
IT ex
x being
Element of
IT st
(
a,
b _|_ c,
x &
c <> x ) ) ) );
theorem Th45: :: ANALMETR:45
canceled;
theorem Th46: :: ANALMETR:46
theorem Th47: :: ANALMETR:47
theorem Th48: :: ANALMETR:48
for
POS being non
empty ParOrtStr for
a,
b,
c,
d being
Element of
POS for
a',
b',
c',
d' being
Element of
(Af POS) st
a = a' &
b = b' &
c = c' &
d = d' holds
(
a,
b // c,
d iff
a',
b' // c',
d' )
Lemma149:
for POS being OrtAfSp holds Af POS is AffinSpace
by Def1;
Lemma150:
for POS being OrtAfPl holds Af POS is AffinPlane
by ;
theorem Th49: :: ANALMETR:49
theorem Th50: :: ANALMETR:50
theorem Th51: :: ANALMETR:51
for
POS being non
empty ParOrtStr holds
(
POS is
OrtAfPl-like iff ( ex
a,
b being
Element of
POS st
a <> b & ( for
a,
b,
c,
d,
p,
q,
r,
s being
Element of
POS holds
(
a,
b // b,
a &
a,
b // c,
c & (
a,
b // p,
q &
a,
b // r,
s & not
p,
q // r,
s implies
a = b ) & (
a,
b // a,
c implies
b,
a // b,
c ) & ex
x being
Element of
POS st
(
a,
b // c,
x &
a,
c // b,
x ) & not for
x,
y,
z being
Element of
POS holds
x,
y // x,
z & ex
x being
Element of
POS st
(
a,
b // c,
x &
c <> x ) & (
a,
b // b,
d &
b <> a implies ex
x being
Element of
POS st
(
c,
b // b,
x &
c,
a // d,
x ) ) & (
a,
b _|_ a,
b implies
a = b ) &
a,
b _|_ c,
c & (
a,
b _|_ c,
d implies (
a,
b _|_ d,
c &
c,
d _|_ a,
b ) ) & (
a,
b _|_ p,
q &
a,
b // r,
s & not
p,
q _|_ r,
s implies
a = b ) & (
a,
b _|_ p,
q &
a,
b _|_ r,
s & not
p,
q // r,
s implies
a = b ) & ex
x being
Element of
POS st
(
a,
b _|_ c,
x &
c <> x ) & ( not
a,
b // c,
d implies ex
x being
Element of
POS st
(
a,
b // a,
x &
c,
d // c,
x ) ) ) ) ) )
:: deftheorem Def11 defines LIN ANALMETR:def 11 :
definition
let POS be non
empty ParOrtStr ;
let a be
Element of
POS;
let b be
Element of
POS;
func Line c2,
c3 -> Subset of
a1 means :
Def12:
:: ANALMETR:def 12
for
x being
Element of
V0 holds
(
x in it iff
LIN w0,
y0,
x );
existence
ex b1 being Subset of POS st
for x being Element of POS holds
( x in b1 iff LIN a,b,x )
uniqueness
for b1, b2 being Subset of POS st ( for x being Element of POS holds
( x in b1 iff LIN a,b,x ) ) & ( for x being Element of POS holds
( x in b2 iff LIN a,b,x ) ) holds
b1 = b2
end;
:: deftheorem Def12 defines Line ANALMETR:def 12 :
:: deftheorem Def13 defines being_line ANALMETR:def 13 :
theorem Th52: :: ANALMETR:52
canceled;
theorem Th53: :: ANALMETR:53
canceled;
theorem Th54: :: ANALMETR:54
canceled;
theorem Th55: :: ANALMETR:55
for
POS being
OrtAfSp for
a,
b,
c being
Element of
POS for
a',
b',
c' being
Element of
(Af POS) st
a = a' &
b = b' &
c = c' holds
(
LIN a,
b,
c iff
LIN a',
b',
c' )
theorem Th56: :: ANALMETR:56
theorem Th57: :: ANALMETR:57
theorem Th58: :: ANALMETR:58
:: deftheorem Def14 defines _|_ ANALMETR:def 14 :
:: deftheorem Def15 defines _|_ ANALMETR:def 15 :
:: deftheorem Def16 defines // ANALMETR:def 16 :
theorem Th59: :: ANALMETR:59
canceled;
theorem Th60: :: ANALMETR:60
canceled;
theorem Th61: :: ANALMETR:61
canceled;
theorem Th62: :: ANALMETR:62
theorem Th63: :: ANALMETR:63
theorem Th64: :: ANALMETR:64
for
POS being
OrtAfSp for
M,
N being
Subset of
POS for
M',
N' being
Subset of
(Af POS) st
M = M' &
N = N' holds
(
M // N iff
M' // N' )
theorem Th65: :: ANALMETR:65
theorem Th66: :: ANALMETR:66
theorem Th67: :: ANALMETR:67
theorem Th68: :: ANALMETR:68
canceled;
theorem Th69: :: ANALMETR:69
theorem Th70: :: ANALMETR:70
canceled;
theorem Th71: :: ANALMETR:71
theorem Th72: :: ANALMETR:72
canceled;
theorem Th73: :: ANALMETR:73
theorem Th74: :: ANALMETR:74
canceled;
theorem Th75: :: ANALMETR:75
theorem Th76: :: ANALMETR:76
theorem Th77: :: ANALMETR:77
theorem Th78: :: ANALMETR:78
theorem Th79: :: ANALMETR:79
theorem Th80: :: ANALMETR:80
for
POS being
OrtAfSp for
b,
c,
a being
Element of
POS holds
(
b,
c _|_ a,
a &
a,
a _|_ b,
c &
b,
c // a,
a &
a,
a // b,
c )
theorem Th81: :: ANALMETR:81
for
POS being
OrtAfSp for
a,
b,
c,
d being
Element of
POS st
a,
b // c,
d holds
(
a,
b // d,
c &
b,
a // c,
d &
b,
a // d,
c &
c,
d // a,
b &
c,
d // b,
a &
d,
c // a,
b &
d,
c // b,
a )
theorem Th82: :: ANALMETR:82
for
POS being
OrtAfSp for
p,
q,
a,
b,
c,
d being
Element of
POS st
p <> q & ( (
p,
q // a,
b &
p,
q // c,
d ) or (
p,
q // a,
b &
c,
d // p,
q ) or (
a,
b // p,
q &
c,
d // p,
q ) or (
a,
b // p,
q &
p,
q // c,
d ) ) holds
a,
b // c,
d
theorem Th83: :: ANALMETR:83
for
POS being
OrtAfSp for
a,
b,
c,
d being
Element of
POS st
a,
b _|_ c,
d holds
(
a,
b _|_ d,
c &
b,
a _|_ c,
d &
b,
a _|_ d,
c &
c,
d _|_ a,
b &
c,
d _|_ b,
a &
d,
c _|_ a,
b &
d,
c _|_ b,
a )
theorem Th84: :: ANALMETR:84
for
POS being
OrtAfSp for
p,
q,
a,
b,
c,
d being
Element of
POS st
p <> q & ( (
p,
q // a,
b &
p,
q _|_ c,
d ) or (
p,
q // c,
d &
p,
q _|_ a,
b ) or (
p,
q // a,
b &
c,
d _|_ p,
q ) or (
p,
q // c,
d &
a,
b _|_ p,
q ) or (
a,
b // p,
q &
c,
d _|_ p,
q ) or (
c,
d // p,
q &
a,
b _|_ p,
q ) or (
a,
b // p,
q &
p,
q _|_ c,
d ) or (
c,
d // p,
q &
p,
q _|_ a,
b ) ) holds
a,
b _|_ c,
d
theorem Th85: :: ANALMETR:85
for
POS being
OrtAfPl for
p,
q,
a,
b,
c,
d being
Element of
POS st
p <> q & ( (
p,
q _|_ a,
b &
p,
q _|_ c,
d ) or (
p,
q _|_ a,
b &
c,
d _|_ p,
q ) or (
a,
b _|_ p,
q &
c,
d _|_ p,
q ) or (
a,
b _|_ p,
q &
p,
q _|_ c,
d ) ) holds
a,
b // c,
d
theorem Th86: :: ANALMETR:86
theorem Th87: :: ANALMETR:87
theorem Th88: :: ANALMETR:88
theorem Th89: :: ANALMETR:89
theorem Th90: :: ANALMETR:90
theorem Th91: :: ANALMETR:91
theorem Th92: :: ANALMETR:92