begin
:: deftheorem defines cpx2euc EUCLID_3:def 1 :
for z being complex number holds cpx2euc z = |[(Re z),(Im z)]|;
:: deftheorem defines euc2cpx EUCLID_3:def 2 :
for p being Point of (TOP-REAL 2) holds euc2cpx p = (p `1) + ((p `2) * <i>);
theorem Th1:
theorem Th2:
theorem
canceled;
theorem
canceled;
theorem
theorem Th6:
theorem
canceled;
theorem
canceled;
theorem Th9:
theorem Th10:
theorem Th11:
theorem Th12:
theorem Th13:
theorem Th14:
theorem Th15:
theorem Th16:
theorem Th17:
theorem
theorem Th19:
theorem Th20:
theorem Th21:
theorem
theorem
theorem
canceled;
theorem
theorem Th26:
theorem
theorem
theorem
theorem
canceled;
theorem Th31:
:: deftheorem defines Arg EUCLID_3:def 3 :
for p being Point of (TOP-REAL 2) holds Arg p = Arg (euc2cpx p);
theorem
theorem
canceled;
theorem
theorem
theorem
theorem
theorem Th38:
theorem
theorem
theorem
canceled;
theorem
:: deftheorem defines angle EUCLID_3:def 4 :
for p1, p2, p3 being Point of (TOP-REAL 2) holds angle (p1,p2,p3) = angle ((euc2cpx p1),(euc2cpx p2),(euc2cpx p3));
theorem
canceled;
theorem
theorem
theorem
theorem
theorem Th48:
theorem Th49:
theorem Th50:
theorem Th51:
theorem
theorem
theorem
theorem
theorem
for
p1,
p2,
p3 being
Point of
(TOP-REAL 2) st
p2 <> p1 &
p1 <> p3 &
p3 <> p2 &
angle (
p2,
p1,
p3)
< PI holds
((angle (p2,p1,p3)) + (angle (p1,p3,p2))) + (angle (p3,p2,p1)) = PI
definition
let n be
Element of
NAT ;
let p1,
p2,
p3 be
Point of
(TOP-REAL n);
func Triangle (
p1,
p2,
p3)
-> Subset of
(TOP-REAL n) equals
((LSeg (p1,p2)) \/ (LSeg (p2,p3))) \/ (LSeg (p3,p1));
correctness
coherence
((LSeg (p1,p2)) \/ (LSeg (p2,p3))) \/ (LSeg (p3,p1)) is Subset of (TOP-REAL n);
;
end;
:: deftheorem defines Triangle EUCLID_3:def 5 :
for n being Element of NAT
for p1, p2, p3 being Point of (TOP-REAL n) holds Triangle (p1,p2,p3) = ((LSeg (p1,p2)) \/ (LSeg (p2,p3))) \/ (LSeg (p3,p1));
:: deftheorem defines closed_inside_of_triangle EUCLID_3:def 6 :
for n being Element of NAT
for p1, p2, p3 being Point of (TOP-REAL n) holds closed_inside_of_triangle (p1,p2,p3) = { p where p is Point of (TOP-REAL n) : ex a1, a2, a3 being Real st
( 0 <= a1 & 0 <= a2 & 0 <= a3 & (a1 + a2) + a3 = 1 & p = ((a1 * p1) + (a2 * p2)) + (a3 * p3) ) } ;
definition
let n be
Element of
NAT ;
let p1,
p2,
p3 be
Point of
(TOP-REAL n);
func inside_of_triangle (
p1,
p2,
p3)
-> Subset of
(TOP-REAL n) equals
(closed_inside_of_triangle (p1,p2,p3)) \ (Triangle (p1,p2,p3));
correctness
coherence
(closed_inside_of_triangle (p1,p2,p3)) \ (Triangle (p1,p2,p3)) is Subset of (TOP-REAL n);
;
end;
:: deftheorem defines inside_of_triangle EUCLID_3:def 7 :
for n being Element of NAT
for p1, p2, p3 being Point of (TOP-REAL n) holds inside_of_triangle (p1,p2,p3) = (closed_inside_of_triangle (p1,p2,p3)) \ (Triangle (p1,p2,p3));
:: deftheorem defines outside_of_triangle EUCLID_3:def 8 :
for n being Element of NAT
for p1, p2, p3 being Point of (TOP-REAL n) holds outside_of_triangle (p1,p2,p3) = { p where p is Point of (TOP-REAL n) : ex a1, a2, a3 being Real st
( ( 0 > a1 or 0 > a2 or 0 > a3 ) & (a1 + a2) + a3 = 1 & p = ((a1 * p1) + (a2 * p2)) + (a3 * p3) ) } ;
definition
let n be
Element of
NAT ;
let p1,
p2,
p3 be
Point of
(TOP-REAL n);
func plane (
p1,
p2,
p3)
-> Subset of
(TOP-REAL n) equals
(outside_of_triangle (p1,p2,p3)) \/ (closed_inside_of_triangle (p1,p2,p3));
correctness
coherence
(outside_of_triangle (p1,p2,p3)) \/ (closed_inside_of_triangle (p1,p2,p3)) is Subset of (TOP-REAL n);
;
end;
:: deftheorem defines plane EUCLID_3:def 9 :
for n being Element of NAT
for p1, p2, p3 being Point of (TOP-REAL n) holds plane (p1,p2,p3) = (outside_of_triangle (p1,p2,p3)) \/ (closed_inside_of_triangle (p1,p2,p3));
theorem Th57:
theorem
:: deftheorem Def10 defines are_lindependent2 EUCLID_3:def 10 :
for n being Element of NAT
for q1, q2 being Point of (TOP-REAL n) holds
( q1,q2 are_lindependent2 iff for a1, a2 being Real st (a1 * q1) + (a2 * q2) = 0. (TOP-REAL n) holds
( a1 = 0 & a2 = 0 ) );
theorem Th59:
theorem Th60:
for
n being
Element of
NAT for
p1,
p2,
p3,
p0 being
Point of
(TOP-REAL n) st
p2 - p1,
p3 - p1 are_lindependent2 &
p0 in plane (
p1,
p2,
p3) holds
ex
a1,
a2,
a3 being
Real st
(
p0 = ((a1 * p1) + (a2 * p2)) + (a3 * p3) &
(a1 + a2) + a3 = 1 & ( for
b1,
b2,
b3 being
Real st
p0 = ((b1 * p1) + (b2 * p2)) + (b3 * p3) &
(b1 + b2) + b3 = 1 holds
(
b1 = a1 &
b2 = a2 &
b3 = a3 ) ) )
theorem Th61:
theorem
theorem Th63:
definition
let n be
Element of
NAT ;
let p1,
p2,
p3,
p be
Point of
(TOP-REAL n);
assume A1:
(
p2 - p1,
p3 - p1 are_lindependent2 &
p in plane (
p1,
p2,
p3) )
;
func tricord1 (
p1,
p2,
p3,
p)
-> Real means :
Def11:
ex
a2,
a3 being
Real st
(
(it + a2) + a3 = 1 &
p = ((it * p1) + (a2 * p2)) + (a3 * p3) );
existence
ex b1, a2, a3 being Real st
( (b1 + a2) + a3 = 1 & p = ((b1 * p1) + (a2 * p2)) + (a3 * p3) )
uniqueness
for b1, b2 being Real st ex a2, a3 being Real st
( (b1 + a2) + a3 = 1 & p = ((b1 * p1) + (a2 * p2)) + (a3 * p3) ) & ex a2, a3 being Real st
( (b2 + a2) + a3 = 1 & p = ((b2 * p1) + (a2 * p2)) + (a3 * p3) ) holds
b1 = b2
end;
:: deftheorem Def11 defines tricord1 EUCLID_3:def 11 :
for n being Element of NAT
for p1, p2, p3, p being Point of (TOP-REAL n) st p2 - p1,p3 - p1 are_lindependent2 & p in plane (p1,p2,p3) holds
for b6 being Real holds
( b6 = tricord1 (p1,p2,p3,p) iff ex a2, a3 being Real st
( (b6 + a2) + a3 = 1 & p = ((b6 * p1) + (a2 * p2)) + (a3 * p3) ) );
definition
let n be
Element of
NAT ;
let p1,
p2,
p3,
p be
Point of
(TOP-REAL n);
assume A1:
(
p2 - p1,
p3 - p1 are_lindependent2 &
p in plane (
p1,
p2,
p3) )
;
func tricord2 (
p1,
p2,
p3,
p)
-> Real means :
Def12:
ex
a1,
a3 being
Real st
(
(a1 + it) + a3 = 1 &
p = ((a1 * p1) + (it * p2)) + (a3 * p3) );
existence
ex b1, a1, a3 being Real st
( (a1 + b1) + a3 = 1 & p = ((a1 * p1) + (b1 * p2)) + (a3 * p3) )
uniqueness
for b1, b2 being Real st ex a1, a3 being Real st
( (a1 + b1) + a3 = 1 & p = ((a1 * p1) + (b1 * p2)) + (a3 * p3) ) & ex a1, a3 being Real st
( (a1 + b2) + a3 = 1 & p = ((a1 * p1) + (b2 * p2)) + (a3 * p3) ) holds
b1 = b2
end;
:: deftheorem Def12 defines tricord2 EUCLID_3:def 12 :
for n being Element of NAT
for p1, p2, p3, p being Point of (TOP-REAL n) st p2 - p1,p3 - p1 are_lindependent2 & p in plane (p1,p2,p3) holds
for b6 being Real holds
( b6 = tricord2 (p1,p2,p3,p) iff ex a1, a3 being Real st
( (a1 + b6) + a3 = 1 & p = ((a1 * p1) + (b6 * p2)) + (a3 * p3) ) );
definition
let n be
Element of
NAT ;
let p1,
p2,
p3,
p be
Point of
(TOP-REAL n);
assume A1:
(
p2 - p1,
p3 - p1 are_lindependent2 &
p in plane (
p1,
p2,
p3) )
;
func tricord3 (
p1,
p2,
p3,
p)
-> Real means :
Def13:
ex
a1,
a2 being
Real st
(
(a1 + a2) + it = 1 &
p = ((a1 * p1) + (a2 * p2)) + (it * p3) );
existence
ex b1, a1, a2 being Real st
( (a1 + a2) + b1 = 1 & p = ((a1 * p1) + (a2 * p2)) + (b1 * p3) )
uniqueness
for b1, b2 being Real st ex a1, a2 being Real st
( (a1 + a2) + b1 = 1 & p = ((a1 * p1) + (a2 * p2)) + (b1 * p3) ) & ex a1, a2 being Real st
( (a1 + a2) + b2 = 1 & p = ((a1 * p1) + (a2 * p2)) + (b2 * p3) ) holds
b1 = b2
end;
:: deftheorem Def13 defines tricord3 EUCLID_3:def 13 :
for n being Element of NAT
for p1, p2, p3, p being Point of (TOP-REAL n) st p2 - p1,p3 - p1 are_lindependent2 & p in plane (p1,p2,p3) holds
for b6 being Real holds
( b6 = tricord3 (p1,p2,p3,p) iff ex a1, a2 being Real st
( (a1 + a2) + b6 = 1 & p = ((a1 * p1) + (a2 * p2)) + (b6 * p3) ) );
definition
let p1,
p2,
p3 be
Point of
(TOP-REAL 2);
func trcmap1 (
p1,
p2,
p3)
-> Function of
(TOP-REAL 2),
R^1 means
for
p being
Point of
(TOP-REAL 2) holds
it . p = tricord1 (
p1,
p2,
p3,
p);
existence
ex b1 being Function of (TOP-REAL 2),R^1 st
for p being Point of (TOP-REAL 2) holds b1 . p = tricord1 (p1,p2,p3,p)
uniqueness
for b1, b2 being Function of (TOP-REAL 2),R^1 st ( for p being Point of (TOP-REAL 2) holds b1 . p = tricord1 (p1,p2,p3,p) ) & ( for p being Point of (TOP-REAL 2) holds b2 . p = tricord1 (p1,p2,p3,p) ) holds
b1 = b2
end;
:: deftheorem defines trcmap1 EUCLID_3:def 14 :
for p1, p2, p3 being Point of (TOP-REAL 2)
for b4 being Function of (TOP-REAL 2),R^1 holds
( b4 = trcmap1 (p1,p2,p3) iff for p being Point of (TOP-REAL 2) holds b4 . p = tricord1 (p1,p2,p3,p) );
definition
let p1,
p2,
p3 be
Point of
(TOP-REAL 2);
func trcmap2 (
p1,
p2,
p3)
-> Function of
(TOP-REAL 2),
R^1 means
for
p being
Point of
(TOP-REAL 2) holds
it . p = tricord2 (
p1,
p2,
p3,
p);
existence
ex b1 being Function of (TOP-REAL 2),R^1 st
for p being Point of (TOP-REAL 2) holds b1 . p = tricord2 (p1,p2,p3,p)
uniqueness
for b1, b2 being Function of (TOP-REAL 2),R^1 st ( for p being Point of (TOP-REAL 2) holds b1 . p = tricord2 (p1,p2,p3,p) ) & ( for p being Point of (TOP-REAL 2) holds b2 . p = tricord2 (p1,p2,p3,p) ) holds
b1 = b2
end;
:: deftheorem defines trcmap2 EUCLID_3:def 15 :
for p1, p2, p3 being Point of (TOP-REAL 2)
for b4 being Function of (TOP-REAL 2),R^1 holds
( b4 = trcmap2 (p1,p2,p3) iff for p being Point of (TOP-REAL 2) holds b4 . p = tricord2 (p1,p2,p3,p) );
definition
let p1,
p2,
p3 be
Point of
(TOP-REAL 2);
func trcmap3 (
p1,
p2,
p3)
-> Function of
(TOP-REAL 2),
R^1 means
for
p being
Point of
(TOP-REAL 2) holds
it . p = tricord3 (
p1,
p2,
p3,
p);
existence
ex b1 being Function of (TOP-REAL 2),R^1 st
for p being Point of (TOP-REAL 2) holds b1 . p = tricord3 (p1,p2,p3,p)
uniqueness
for b1, b2 being Function of (TOP-REAL 2),R^1 st ( for p being Point of (TOP-REAL 2) holds b1 . p = tricord3 (p1,p2,p3,p) ) & ( for p being Point of (TOP-REAL 2) holds b2 . p = tricord3 (p1,p2,p3,p) ) holds
b1 = b2
end;
:: deftheorem defines trcmap3 EUCLID_3:def 16 :
for p1, p2, p3 being Point of (TOP-REAL 2)
for b4 being Function of (TOP-REAL 2),R^1 holds
( b4 = trcmap3 (p1,p2,p3) iff for p being Point of (TOP-REAL 2) holds b4 . p = tricord3 (p1,p2,p3,p) );
theorem
for
p1,
p2,
p3,
p being
Point of
(TOP-REAL 2) st
p2 - p1,
p3 - p1 are_lindependent2 holds
(
p in outside_of_triangle (
p1,
p2,
p3) iff (
tricord1 (
p1,
p2,
p3,
p)
< 0 or
tricord2 (
p1,
p2,
p3,
p)
< 0 or
tricord3 (
p1,
p2,
p3,
p)
< 0 ) )
theorem Th65:
for
p1,
p2,
p3,
p being
Point of
(TOP-REAL 2) st
p2 - p1,
p3 - p1 are_lindependent2 holds
(
p in Triangle (
p1,
p2,
p3) iff (
tricord1 (
p1,
p2,
p3,
p)
>= 0 &
tricord2 (
p1,
p2,
p3,
p)
>= 0 &
tricord3 (
p1,
p2,
p3,
p)
>= 0 & (
tricord1 (
p1,
p2,
p3,
p)
= 0 or
tricord2 (
p1,
p2,
p3,
p)
= 0 or
tricord3 (
p1,
p2,
p3,
p)
= 0 ) ) )
theorem
for
p1,
p2,
p3,
p being
Point of
(TOP-REAL 2) st
p2 - p1,
p3 - p1 are_lindependent2 holds
(
p in Triangle (
p1,
p2,
p3) iff ( (
tricord1 (
p1,
p2,
p3,
p)
= 0 &
tricord2 (
p1,
p2,
p3,
p)
>= 0 &
tricord3 (
p1,
p2,
p3,
p)
>= 0 ) or (
tricord1 (
p1,
p2,
p3,
p)
>= 0 &
tricord2 (
p1,
p2,
p3,
p)
= 0 &
tricord3 (
p1,
p2,
p3,
p)
>= 0 ) or (
tricord1 (
p1,
p2,
p3,
p)
>= 0 &
tricord2 (
p1,
p2,
p3,
p)
>= 0 &
tricord3 (
p1,
p2,
p3,
p)
= 0 ) ) )
by Th65;
theorem Th67:
for
p1,
p2,
p3,
p being
Point of
(TOP-REAL 2) st
p2 - p1,
p3 - p1 are_lindependent2 holds
(
p in inside_of_triangle (
p1,
p2,
p3) iff (
tricord1 (
p1,
p2,
p3,
p)
> 0 &
tricord2 (
p1,
p2,
p3,
p)
> 0 &
tricord3 (
p1,
p2,
p3,
p)
> 0 ) )
theorem