begin
theorem Th1:
theorem Th2:
theorem Th3:
theorem
theorem Th5:
theorem
:: deftheorem defines AfLines AFPROJ:def 1 :
for AS being AffinSpace holds AfLines AS = { A where A is Subset of AS : A is being_line } ;
:: deftheorem defines AfPlanes AFPROJ:def 2 :
for AS being AffinSpace holds AfPlanes AS = { A where A is Subset of AS : A is being_plane } ;
theorem
theorem
:: deftheorem defines LinesParallelity AFPROJ:def 3 :
for AS being AffinSpace holds LinesParallelity AS = { [K,M] where K, M is Subset of AS : ( K is being_line & M is being_line & K '||' M ) } ;
:: deftheorem defines PlanesParallelity AFPROJ:def 4 :
for AS being AffinSpace holds PlanesParallelity AS = { [X,Y] where X, Y is Subset of AS : ( X is being_plane & Y is being_plane & X '||' Y ) } ;
:: deftheorem defines LDir AFPROJ:def 5 :
for AS being AffinSpace
for X being Subset of AS holds LDir X = Class ((LinesParallelity AS),X);
:: deftheorem defines PDir AFPROJ:def 6 :
for AS being AffinSpace
for X being Subset of AS holds PDir X = Class ((PlanesParallelity AS),X);
theorem Th9:
theorem Th10:
theorem Th11:
theorem Th12:
theorem Th13:
:: deftheorem defines Dir_of_Lines AFPROJ:def 7 :
for AS being AffinSpace holds Dir_of_Lines AS = Class (LinesParallelity AS);
:: deftheorem defines Dir_of_Planes AFPROJ:def 8 :
for AS being AffinSpace holds Dir_of_Planes AS = Class (PlanesParallelity AS);
theorem Th14:
theorem Th15:
theorem Th16:
theorem
theorem Th18:
theorem Th19:
:: deftheorem defines ProjectivePoints AFPROJ:def 9 :
for AS being AffinSpace holds ProjectivePoints AS = the carrier of AS \/ (Dir_of_Lines AS);
:: deftheorem defines ProjectiveLines AFPROJ:def 10 :
for AS being AffinSpace holds ProjectiveLines AS = [:(AfLines AS),{1}:] \/ [:(Dir_of_Planes AS),{2}:];
definition
let AS be
AffinSpace;
func Proj_Inc AS -> Relation of
(ProjectivePoints AS),
(ProjectiveLines AS) means :
Def11:
for
x,
y being
set holds
(
[x,y] in it iff ( ex
K being
Subset of
AS st
(
K is
being_line &
y = [K,1] & ( (
x in the
carrier of
AS &
x in K ) or
x = LDir K ) ) or ex
K,
X being
Subset of
AS st
(
K is
being_line &
X is
being_plane &
x = LDir K &
y = [(PDir X),2] &
K '||' X ) ) );
existence
ex b1 being Relation of (ProjectivePoints AS),(ProjectiveLines AS) st
for x, y being set holds
( [x,y] in b1 iff ( ex K being Subset of AS st
( K is being_line & y = [K,1] & ( ( x in the carrier of AS & x in K ) or x = LDir K ) ) or ex K, X being Subset of AS st
( K is being_line & X is being_plane & x = LDir K & y = [(PDir X),2] & K '||' X ) ) )
uniqueness
for b1, b2 being Relation of (ProjectivePoints AS),(ProjectiveLines AS) st ( for x, y being set holds
( [x,y] in b1 iff ( ex K being Subset of AS st
( K is being_line & y = [K,1] & ( ( x in the carrier of AS & x in K ) or x = LDir K ) ) or ex K, X being Subset of AS st
( K is being_line & X is being_plane & x = LDir K & y = [(PDir X),2] & K '||' X ) ) ) ) & ( for x, y being set holds
( [x,y] in b2 iff ( ex K being Subset of AS st
( K is being_line & y = [K,1] & ( ( x in the carrier of AS & x in K ) or x = LDir K ) ) or ex K, X being Subset of AS st
( K is being_line & X is being_plane & x = LDir K & y = [(PDir X),2] & K '||' X ) ) ) ) holds
b1 = b2
end;
:: deftheorem Def11 defines Proj_Inc AFPROJ:def 11 :
for AS being AffinSpace
for b2 being Relation of (ProjectivePoints AS),(ProjectiveLines AS) holds
( b2 = Proj_Inc AS iff for x, y being set holds
( [x,y] in b2 iff ( ex K being Subset of AS st
( K is being_line & y = [K,1] & ( ( x in the carrier of AS & x in K ) or x = LDir K ) ) or ex K, X being Subset of AS st
( K is being_line & X is being_plane & x = LDir K & y = [(PDir X),2] & K '||' X ) ) ) );
definition
let AS be
AffinSpace;
func Inc_of_Dir AS -> Relation of
(Dir_of_Lines AS),
(Dir_of_Planes AS) means :
Def12:
for
x,
y being
set holds
(
[x,y] in it iff ex
A,
X being
Subset of
AS st
(
x = LDir A &
y = PDir X &
A is
being_line &
X is
being_plane &
A '||' X ) );
existence
ex b1 being Relation of (Dir_of_Lines AS),(Dir_of_Planes AS) st
for x, y being set holds
( [x,y] in b1 iff ex A, X being Subset of AS st
( x = LDir A & y = PDir X & A is being_line & X is being_plane & A '||' X ) )
uniqueness
for b1, b2 being Relation of (Dir_of_Lines AS),(Dir_of_Planes AS) st ( for x, y being set holds
( [x,y] in b1 iff ex A, X being Subset of AS st
( x = LDir A & y = PDir X & A is being_line & X is being_plane & A '||' X ) ) ) & ( for x, y being set holds
( [x,y] in b2 iff ex A, X being Subset of AS st
( x = LDir A & y = PDir X & A is being_line & X is being_plane & A '||' X ) ) ) holds
b1 = b2
end;
:: deftheorem Def12 defines Inc_of_Dir AFPROJ:def 12 :
for AS being AffinSpace
for b2 being Relation of (Dir_of_Lines AS),(Dir_of_Planes AS) holds
( b2 = Inc_of_Dir AS iff for x, y being set holds
( [x,y] in b2 iff ex A, X being Subset of AS st
( x = LDir A & y = PDir X & A is being_line & X is being_plane & A '||' X ) ) );
:: deftheorem defines IncProjSp_of AFPROJ:def 13 :
for AS being AffinSpace holds IncProjSp_of AS = IncProjStr(# (ProjectivePoints AS),(ProjectiveLines AS),(Proj_Inc AS) #);
:: deftheorem defines ProjHorizon AFPROJ:def 14 :
for AS being AffinSpace holds ProjHorizon AS = IncProjStr(# (Dir_of_Lines AS),(Dir_of_Planes AS),(Inc_of_Dir AS) #);
theorem Th20:
theorem
theorem Th22:
theorem Th23:
theorem
theorem Th25:
theorem Th26:
theorem Th27:
theorem Th28:
theorem Th29:
theorem Th30:
theorem Th31:
theorem Th32:
theorem
theorem Th34:
theorem Th35:
theorem Th36:
theorem Th37:
theorem Th38:
theorem Th39:
theorem Th40:
Lm1:
for AS being AffinSpace st AS is not AffinPlane holds
ex a being Element of the Points of (ProjHorizon AS) ex A being Element of the Lines of (ProjHorizon AS) st not a on A
Lm2:
for AS being AffinSpace
for a, b being POINT of (IncProjSp_of AS)
for A, K being LINE of (IncProjSp_of AS) st a on A & a on K & b on A & b on K & not a = b holds
A = K
Lm3:
for AS being AffinSpace
for A being LINE of (IncProjSp_of AS) ex a, b, c being POINT of (IncProjSp_of AS) st
( a on A & b on A & c on A & a <> b & b <> c & c <> a )
Lm4:
for AS being AffinSpace
for a, b being POINT of (IncProjSp_of AS) ex A being LINE of (IncProjSp_of AS) st
( a on A & b on A )
Lm5:
for AS being AffinSpace
for A, K being LINE of (IncProjSp_of AS) st AS is AffinPlane holds
ex a being POINT of (IncProjSp_of AS) st
( a on A & a on K )
Lm6:
for AS being AffinSpace holds
not for a being POINT of (IncProjSp_of AS)
for A being LINE of (IncProjSp_of AS) holds a on A
theorem Th41:
theorem Th42:
Lm7:
for AS being AffinSpace
for X, X9, Y, Y9 being Subset of AS
for b, c being POINT of (IncProjSp_of AS)
for A, K, M being LINE of (IncProjSp_of AS) st X is being_line & X9 is being_line & Y is being_plane & X c= Y & X9 c= Y & A = [X,1] & K = [X9,1] & b on A & c on K & b on M & c on M & b <> c & M = [Y9,1] & Y9 is being_line holds
Y9 c= Y
Lm8:
for AS being AffinSpace
for X, X9, Y, Y9 being Subset of AS
for b, c being POINT of (IncProjSp_of AS)
for A, K, M being LINE of (IncProjSp_of AS) st X is being_line & X9 is being_line & Y is being_plane & X c= Y & X9 c= Y & A = [X,1] & K = [X9,1] & b on A & c on K & b on M & c on M & b <> c & M = [(PDir Y9),2] & Y9 is being_plane holds
( Y9 '||' Y & Y '||' Y9 )
theorem Th43:
Lm9:
for AS being AffinSpace
for a, b, c, d, p being POINT of (IncProjSp_of AS)
for M, N, P, Q being LINE of (IncProjSp_of AS) st a on M & b on M & c on N & d on N & p on M & p on N & a on P & c on P & b on Q & d on Q & not p on P & not p on Q & M <> N & p is Element of AS holds
ex q being POINT of (IncProjSp_of AS) st
( q on P & q on Q )
Lm10:
for AS being AffinSpace
for a, b, c, d, p being POINT of (IncProjSp_of AS)
for M, N, P, Q being LINE of (IncProjSp_of AS) st a on M & b on M & c on N & d on N & p on M & p on N & a on P & c on P & b on Q & d on Q & not p on P & not p on Q & M <> N & p is not Element of AS & a is Element of AS holds
ex q being POINT of (IncProjSp_of AS) st
( q on P & q on Q )
Lm11:
for AS being AffinSpace
for a, b, c, d, p being POINT of (IncProjSp_of AS)
for M, N, P, Q being LINE of (IncProjSp_of AS) st a on M & b on M & c on N & d on N & p on M & p on N & a on P & c on P & b on Q & d on Q & not p on P & not p on Q & M <> N & p is not Element of AS & ( a is Element of AS or b is Element of AS or c is Element of AS or d is Element of AS ) holds
ex q being POINT of (IncProjSp_of AS) st
( q on P & q on Q )
Lm12:
for AS being AffinSpace
for a, b, c, d, p being POINT of (IncProjSp_of AS)
for M, N, P, Q being LINE of (IncProjSp_of AS) st a on M & b on M & c on N & d on N & p on M & p on N & a on P & c on P & b on Q & d on Q & not p on P & not p on Q & M <> N holds
ex q being POINT of (IncProjSp_of AS) st
( q on P & q on Q )
theorem Th44:
theorem
theorem
theorem
theorem Th48:
for
AS being
AffinSpace for
M,
N being
Subset of
AS for
o,
a,
b,
c,
a9,
b9,
c9 being
Element of
AS st
M is
being_line &
N is
being_line &
M <> N &
o in M &
o in N &
o <> b &
o <> b9 &
o <> c9 &
b in M &
c in M &
a9 in N &
b9 in N &
c9 in N &
a,
b9 // b,
a9 &
b,
c9 // c,
b9 & (
a = b or
b = c or
a = c ) holds
a,
c9 // c,
a9
theorem
theorem Th50:
for
AS being
AffinSpace for
A,
P,
C being
Subset of
AS for
o,
a,
b,
c,
a9,
b9,
c9 being
Element of
AS st
o in A &
o in P &
o in C &
o <> a &
o <> b &
o <> c &
a in A &
b in P &
b9 in P &
c in C &
c9 in C &
A is
being_line &
P is
being_line &
C is
being_line &
A <> P &
A <> C &
a,
b // a9,
b9 &
a,
c // a9,
c9 & (
o = a9 or
a = a9 ) holds
b,
c // b9,
c9
theorem
theorem