begin
:: deftheorem Def1 defines Relation3 COLLSP:def 1 :
for X, b2 being set holds
( b2 is Relation3 of X iff b2 c= [:X,X,X:] );
theorem
canceled;
theorem Th2:
for
X being
set holds
(
X = {} or ex
a being
set st
(
{a} = X or ex
a,
b being
set st
(
a <> b &
a in X &
b in X ) ) )
:: deftheorem Def2 defines is_collinear COLLSP:def 2 :
for CS being non empty CollStr
for a, b, c being Point of CS holds
( a,b,c is_collinear iff [a,b,c] in the Collinearity of CS );
set Z = {1};
Lm1:
1 in {1}
by TARSKI:def 1;
Lm2:
{[1,1,1]} c= [:{1},{1},{1}:]
reconsider Z = {1} as non empty set ;
reconsider RR = {[1,1,1]} as Relation3 of Z by Def1, Lm2;
reconsider CLS = CollStr(# Z,RR #) as non empty CollStr ;
Lm3:
now
let a,
b,
c,
p,
q,
r be
Point of
CLS;
( ( ( a = b or a = c or b = c ) implies [a,b,c] in the Collinearity of CLS ) & ( a <> b & [a,b,p] in the Collinearity of CLS & [a,b,q] in the Collinearity of CLS & [a,b,r] in the Collinearity of CLS implies [p,q,r] in the Collinearity of CLS ) )A1:
for
z1,
z2,
z3 being
Point of
CLS holds
[z1,z2,z3] in the
Collinearity of
CLS
hence
( (
a = b or
a = c or
b = c ) implies
[a,b,c] in the
Collinearity of
CLS )
;
( a <> b & [a,b,p] in the Collinearity of CLS & [a,b,q] in the Collinearity of CLS & [a,b,r] in the Collinearity of CLS implies [p,q,r] in the Collinearity of CLS )thus
(
a <> b &
[a,b,p] in the
Collinearity of
CLS &
[a,b,q] in the
Collinearity of
CLS &
[a,b,r] in the
Collinearity of
CLS implies
[p,q,r] in the
Collinearity of
CLS )
by A1;
verum
end;
:: deftheorem Def3 defines reflexive COLLSP:def 3 :
for IT being non empty CollStr holds
( IT is reflexive iff for a, b, c being Point of IT st ( a = b or a = c or b = c ) holds
[a,b,c] in the Collinearity of IT );
definition
let IT be non
empty CollStr ;
attr IT is
transitive means :
Def4:
for
a,
b,
p,
q,
r being
Point of
IT st
a <> b &
[a,b,p] in the
Collinearity of
IT &
[a,b,q] in the
Collinearity of
IT &
[a,b,r] in the
Collinearity of
IT holds
[p,q,r] in the
Collinearity of
IT;
end;
:: deftheorem Def4 defines transitive COLLSP:def 4 :
for IT being non empty CollStr holds
( IT is transitive iff for a, b, p, q, r being Point of IT st a <> b & [a,b,p] in the Collinearity of IT & [a,b,q] in the Collinearity of IT & [a,b,r] in the Collinearity of IT holds
[p,q,r] in the Collinearity of IT );
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem Th7:
theorem Th8:
for
CLSP being
CollSp for
a,
b,
p,
q,
r being
Point of
CLSP st
a <> b &
a,
b,
p is_collinear &
a,
b,
q is_collinear &
a,
b,
r is_collinear holds
p,
q,
r is_collinear
theorem Th9:
theorem
theorem Th11:
theorem
theorem Th13:
theorem Th14:
for
CLSP being
CollSp for
p,
q,
a,
b,
r being
Point of
CLSP st
p <> q &
a,
b,
p is_collinear &
a,
b,
q is_collinear &
p,
q,
r is_collinear holds
a,
b,
r is_collinear
:: deftheorem defines Line COLLSP:def 5 :
for CLSP being CollSp
for a, b being Point of CLSP holds Line (a,b) = { p where p is Point of CLSP : a,b,p is_collinear } ;
theorem
canceled;
theorem Th16:
theorem Th17:
set Z = {1,2,3};
set RR = { [i,j,k] where i, j, k is Element of NAT : ( ( i = j or j = k or k = i ) & i in {1,2,3} & j in {1,2,3} & k in {1,2,3} ) } ;
Lm4:
{ [i,j,k] where i, j, k is Element of NAT : ( ( i = j or j = k or k = i ) & i in {1,2,3} & j in {1,2,3} & k in {1,2,3} ) } c= [:{1,2,3},{1,2,3},{1,2,3}:]
reconsider Z = {1,2,3} as non empty set by ENUMSET1:def 1;
reconsider RR = { [i,j,k] where i, j, k is Element of NAT : ( ( i = j or j = k or k = i ) & i in {1,2,3} & j in {1,2,3} & k in {1,2,3} ) } as Relation3 of Z by Def1, Lm4;
reconsider CLS = CollStr(# Z,RR #) as non empty CollStr ;
Lm5:
for a, b, c being Point of CLS holds
( [a,b,c] in RR iff ( ( a = b or b = c or c = a ) & a in Z & b in Z & c in Z ) )
Lm6:
for a, b, c, p, q, r being Point of CLS st a <> b & [a,b,p] in the Collinearity of CLS & [a,b,q] in the Collinearity of CLS & [a,b,r] in the Collinearity of CLS holds
[p,q,r] in the Collinearity of CLS
Lm7:
not for a, b, c being Point of CLS holds a,b,c is_collinear
Lm8:
CLS is CollSp
:: deftheorem Def6 defines proper COLLSP:def 6 :
for IT being non empty CollStr holds
( IT is proper iff not for a, b, c being Point of IT holds a,b,c is_collinear );
theorem
canceled;
theorem Th19:
:: deftheorem Def7 defines LINE COLLSP:def 7 :
for CLSP being proper CollSp
for b2 being set holds
( b2 is LINE of CLSP iff ex a, b being Point of CLSP st
( a <> b & b2 = Line (a,b) ) );
theorem
canceled;
theorem
canceled;
theorem
theorem
theorem
theorem Th25:
Lm9:
for CLSP being proper CollSp
for P being LINE of CLSP
for x being set st x in P holds
x is Point of CLSP
theorem Th26:
theorem Th27:
theorem Th28:
theorem Th29:
theorem
theorem