:: EUCLID_3 semantic presentation
:: deftheorem Def1 defines cpx2euc EUCLID_3:def 1 :
:: deftheorem Def2 defines euc2cpx EUCLID_3:def 2 :
theorem Th1: :: EUCLID_3:1
theorem Th2: :: EUCLID_3:2
theorem Th3: :: EUCLID_3:3
theorem Th4: :: EUCLID_3:4
theorem Th5: :: EUCLID_3:5
theorem Th6: :: EUCLID_3:6
theorem Th7: :: EUCLID_3:7
theorem Th8: :: EUCLID_3:8
theorem Th9: :: EUCLID_3:9
theorem Th10: :: EUCLID_3:10
theorem Th11: :: EUCLID_3:11
theorem Th12: :: EUCLID_3:12
theorem Th13: :: EUCLID_3:13
theorem Th14: :: EUCLID_3:14
theorem Th15: :: EUCLID_3:15
theorem Th16: :: EUCLID_3:16
theorem Th17: :: EUCLID_3:17
theorem Th18: :: EUCLID_3:18
theorem Th19: :: EUCLID_3:19
theorem Th20: :: EUCLID_3:20
theorem Th21: :: EUCLID_3:21
theorem Th22: :: EUCLID_3:22
theorem Th23: :: EUCLID_3:23
theorem Th24: :: EUCLID_3:24
theorem Th25: :: EUCLID_3:25
theorem Th26: :: EUCLID_3:26
theorem Th27: :: EUCLID_3:27
theorem Th28: :: EUCLID_3:28
theorem Th29: :: EUCLID_3:29
theorem Th30: :: EUCLID_3:30
theorem Th31: :: EUCLID_3:31
:: deftheorem Def3 defines Arg EUCLID_3:def 3 :
theorem Th32: :: EUCLID_3:32
theorem Th33: :: EUCLID_3:33
theorem Th34: :: EUCLID_3:34
theorem Th35: :: EUCLID_3:35
theorem Th36: :: EUCLID_3:36
theorem Th37: :: EUCLID_3:37
theorem Th38: :: EUCLID_3:38
theorem Th39: :: EUCLID_3:39
theorem Th40: :: EUCLID_3:40
theorem Th41: :: EUCLID_3:41
theorem Th42: :: EUCLID_3:42
:: deftheorem Def4 defines angle EUCLID_3:def 4 :
theorem Th43: :: EUCLID_3:43
theorem Th44: :: EUCLID_3:44
theorem Th45: :: EUCLID_3:45
theorem Th46: :: EUCLID_3:46
theorem Th47: :: EUCLID_3:47
theorem Th48: :: EUCLID_3:48
theorem Th49: :: EUCLID_3:49
theorem Th50: :: EUCLID_3:50
theorem Th51: :: EUCLID_3:51
theorem Th52: :: EUCLID_3:52
theorem Th53: :: EUCLID_3:53
theorem Th54: :: EUCLID_3:54
theorem Th55: :: EUCLID_3:55
theorem Th56: :: EUCLID_3:56
for b
1, b
2, b
3 being
Point of
(TOP-REAL 2) holds
( b
2 <> b
1 & b
1 <> b
3 & b
3 <> b
2 &
angle b
2,b
1,b
3 < PI &
angle b
1,b
3,b
2 < PI &
angle b
3,b
2,b
1 < PI implies
((angle b2,b1,b3) + (angle b1,b3,b2)) + (angle b3,b2,b1) = PI )
definition
let c
1 be
Nat;
let c
2, c
3, c
4 be
Point of
(TOP-REAL c1);
func Triangle c
2,c
3,c
4 -> Subset of
(TOP-REAL a1) equals :: EUCLID_3:def 5
((LSeg a2,a3) \/ (LSeg a3,a4)) \/ (LSeg a4,a2);
correctness
coherence
((LSeg c2,c3) \/ (LSeg c3,c4)) \/ (LSeg c4,c2) is Subset of (TOP-REAL c1);
;
end;
:: deftheorem Def5 defines Triangle EUCLID_3:def 5 :
:: deftheorem Def6 defines closed_inside_of_triangle EUCLID_3:def 6 :
definition
let c
1 be
Nat;
let c
2, c
3, c
4 be
Point of
(TOP-REAL c1);
func inside_of_triangle c
2,c
3,c
4 -> Subset of
(TOP-REAL a1) equals :: EUCLID_3:def 7
(closed_inside_of_triangle a2,a3,a4) \ (Triangle a2,a3,a4);
correctness
coherence
(closed_inside_of_triangle c2,c3,c4) \ (Triangle c2,c3,c4) is Subset of (TOP-REAL c1);
;
end;
:: deftheorem Def7 defines inside_of_triangle EUCLID_3:def 7 :
:: deftheorem Def8 defines outside_of_triangle EUCLID_3:def 8 :
definition
let c
1 be
Nat;
let c
2, c
3, c
4 be
Point of
(TOP-REAL c1);
func plane c
2,c
3,c
4 -> Subset of
(TOP-REAL a1) equals :: EUCLID_3:def 9
(outside_of_triangle a2,a3,a4) \/ (closed_inside_of_triangle a2,a3,a4);
correctness
coherence
(outside_of_triangle c2,c3,c4) \/ (closed_inside_of_triangle c2,c3,c4) is Subset of (TOP-REAL c1);
;
end;
:: deftheorem Def9 defines plane EUCLID_3:def 9 :
theorem Th57: :: EUCLID_3:57
for b
1 being
Natfor b
2, b
3, b
4, b
5 being
Point of
(TOP-REAL b1) holds
not ( b
5 in plane b
2,b
3,b
4 & ( for b
6, b
7, b
8 being
Real holds
not (
(b6 + b7) + b
8 = 1 & b
5 = ((b6 * b2) + (b7 * b3)) + (b8 * b4) ) ) )
theorem Th58: :: EUCLID_3:58
:: deftheorem Def10 defines are_lindependent2 EUCLID_3:def 10 :
theorem Th59: :: EUCLID_3:59
theorem Th60: :: EUCLID_3:60
for b
1 being
Natfor b
2, b
3, b
4, b
5 being
Point of
(TOP-REAL b1) holds
not ( b
3 - b
2,b
4 - b
2 are_lindependent2 & b
5 in plane b
2,b
3,b
4 & ( for b
6, b
7, b
8 being
Real holds
not ( b
5 = ((b6 * b2) + (b7 * b3)) + (b8 * b4) &
(b6 + b7) + b
8 = 1 & ( for b
9, b
10, b
11 being
Real holds
( b
5 = ((b9 * b2) + (b10 * b3)) + (b11 * b4) &
(b9 + b10) + b
11 = 1 implies ( b
9 = b
6 & b
10 = b
7 & b
11 = b
8 ) ) ) ) ) )
theorem Th61: :: EUCLID_3:61
for b
1 being
Natfor b
2, b
3, b
4, b
5 being
Point of
(TOP-REAL b1) holds
( ex b
6, b
7, b
8 being
Real st
( b
5 = ((b6 * b2) + (b7 * b3)) + (b8 * b4) &
(b6 + b7) + b
8 = 1 ) implies b
5 in plane b
2,b
3,b
4 )
theorem Th62: :: EUCLID_3:62
theorem Th63: :: EUCLID_3:63
definition
let c
1 be
Nat;
let c
2, c
3, c
4, c
5 be
Point of
(TOP-REAL c1);
assume E33:
( c
3 - c
2,c
4 - c
2 are_lindependent2 & c
5 in plane c
2,c
3,c
4 )
;
func tricord1 c
2,c
3,c
4,c
5 -> Real means :
Def11:
:: EUCLID_3:def 11
ex b
1, b
2 being
Real st
(
(a6 + b1) + b
2 = 1 & a
5 = ((a6 * a2) + (b1 * a3)) + (b2 * a4) );
existence
ex b1, b2, b3 being Real st
( (b1 + b2) + b3 = 1 & c5 = ((b1 * c2) + (b2 * c3)) + (b3 * c4) )
uniqueness
for b1, b2 being Real holds
( ex b3, b4 being Real st
( (b1 + b3) + b4 = 1 & c5 = ((b1 * c2) + (b3 * c3)) + (b4 * c4) ) & ex b3, b4 being Real st
( (b2 + b3) + b4 = 1 & c5 = ((b2 * c2) + (b3 * c3)) + (b4 * c4) ) implies b1 = b2 )
end;
:: deftheorem Def11 defines tricord1 EUCLID_3:def 11 :
for b
1 being
Natfor b
2, b
3, b
4, b
5 being
Point of
(TOP-REAL b1) holds
( b
3 - b
2,b
4 - b
2 are_lindependent2 & b
5 in plane b
2,b
3,b
4 implies for b
6 being
Real holds
( b
6 = tricord1 b
2,b
3,b
4,b
5 iff ex b
7, b
8 being
Real st
(
(b6 + b7) + b
8 = 1 & b
5 = ((b6 * b2) + (b7 * b3)) + (b8 * b4) ) ) );
definition
let c
1 be
Nat;
let c
2, c
3, c
4, c
5 be
Point of
(TOP-REAL c1);
assume E34:
( c
3 - c
2,c
4 - c
2 are_lindependent2 & c
5 in plane c
2,c
3,c
4 )
;
func tricord2 c
2,c
3,c
4,c
5 -> Real means :
Def12:
:: EUCLID_3:def 12
ex b
1, b
2 being
Real st
(
(b1 + a6) + b
2 = 1 & a
5 = ((b1 * a2) + (a6 * a3)) + (b2 * a4) );
existence
ex b1, b2, b3 being Real st
( (b2 + b1) + b3 = 1 & c5 = ((b2 * c2) + (b1 * c3)) + (b3 * c4) )
uniqueness
for b1, b2 being Real holds
( ex b3, b4 being Real st
( (b3 + b1) + b4 = 1 & c5 = ((b3 * c2) + (b1 * c3)) + (b4 * c4) ) & ex b3, b4 being Real st
( (b3 + b2) + b4 = 1 & c5 = ((b3 * c2) + (b2 * c3)) + (b4 * c4) ) implies b1 = b2 )
end;
:: deftheorem Def12 defines tricord2 EUCLID_3:def 12 :
for b
1 being
Natfor b
2, b
3, b
4, b
5 being
Point of
(TOP-REAL b1) holds
( b
3 - b
2,b
4 - b
2 are_lindependent2 & b
5 in plane b
2,b
3,b
4 implies for b
6 being
Real holds
( b
6 = tricord2 b
2,b
3,b
4,b
5 iff ex b
7, b
8 being
Real st
(
(b7 + b6) + b
8 = 1 & b
5 = ((b7 * b2) + (b6 * b3)) + (b8 * b4) ) ) );
definition
let c
1 be
Nat;
let c
2, c
3, c
4, c
5 be
Point of
(TOP-REAL c1);
assume E35:
( c
3 - c
2,c
4 - c
2 are_lindependent2 & c
5 in plane c
2,c
3,c
4 )
;
func tricord3 c
2,c
3,c
4,c
5 -> Real means :
Def13:
:: EUCLID_3:def 13
ex b
1, b
2 being
Real st
(
(b1 + b2) + a
6 = 1 & a
5 = ((b1 * a2) + (b2 * a3)) + (a6 * a4) );
existence
ex b1, b2, b3 being Real st
( (b2 + b3) + b1 = 1 & c5 = ((b2 * c2) + (b3 * c3)) + (b1 * c4) )
uniqueness
for b1, b2 being Real holds
( ex b3, b4 being Real st
( (b3 + b4) + b1 = 1 & c5 = ((b3 * c2) + (b4 * c3)) + (b1 * c4) ) & ex b3, b4 being Real st
( (b3 + b4) + b2 = 1 & c5 = ((b3 * c2) + (b4 * c3)) + (b2 * c4) ) implies b1 = b2 )
end;
:: deftheorem Def13 defines tricord3 EUCLID_3:def 13 :
for b
1 being
Natfor b
2, b
3, b
4, b
5 being
Point of
(TOP-REAL b1) holds
( b
3 - b
2,b
4 - b
2 are_lindependent2 & b
5 in plane b
2,b
3,b
4 implies for b
6 being
Real holds
( b
6 = tricord3 b
2,b
3,b
4,b
5 iff ex b
7, b
8 being
Real st
(
(b7 + b8) + b
6 = 1 & b
5 = ((b7 * b2) + (b8 * b3)) + (b6 * b4) ) ) );
definition
let c
1, c
2, c
3 be
Point of
(TOP-REAL 2);
func trcmap1 c
1,c
2,c
3 -> Function of
(TOP-REAL 2),
R^1 means :: EUCLID_3:def 14
for b
1 being
Point of
(TOP-REAL 2) holds a
4 . b
1 = tricord1 a
1,a
2,a
3,b
1;
existence
ex b1 being Function of (TOP-REAL 2),R^1 st
for b2 being Point of (TOP-REAL 2) holds b1 . b2 = tricord1 c1,c2,c3,b2
uniqueness
for b1, b2 being Function of (TOP-REAL 2),R^1 holds
( ( for b3 being Point of (TOP-REAL 2) holds b1 . b3 = tricord1 c1,c2,c3,b3 ) & ( for b3 being Point of (TOP-REAL 2) holds b2 . b3 = tricord1 c1,c2,c3,b3 ) implies b1 = b2 )
end;
:: deftheorem Def14 defines trcmap1 EUCLID_3:def 14 :
definition
let c
1, c
2, c
3 be
Point of
(TOP-REAL 2);
func trcmap2 c
1,c
2,c
3 -> Function of
(TOP-REAL 2),
R^1 means :: EUCLID_3:def 15
for b
1 being
Point of
(TOP-REAL 2) holds a
4 . b
1 = tricord2 a
1,a
2,a
3,b
1;
existence
ex b1 being Function of (TOP-REAL 2),R^1 st
for b2 being Point of (TOP-REAL 2) holds b1 . b2 = tricord2 c1,c2,c3,b2
uniqueness
for b1, b2 being Function of (TOP-REAL 2),R^1 holds
( ( for b3 being Point of (TOP-REAL 2) holds b1 . b3 = tricord2 c1,c2,c3,b3 ) & ( for b3 being Point of (TOP-REAL 2) holds b2 . b3 = tricord2 c1,c2,c3,b3 ) implies b1 = b2 )
end;
:: deftheorem Def15 defines trcmap2 EUCLID_3:def 15 :
definition
let c
1, c
2, c
3 be
Point of
(TOP-REAL 2);
func trcmap3 c
1,c
2,c
3 -> Function of
(TOP-REAL 2),
R^1 means :: EUCLID_3:def 16
for b
1 being
Point of
(TOP-REAL 2) holds a
4 . b
1 = tricord3 a
1,a
2,a
3,b
1;
existence
ex b1 being Function of (TOP-REAL 2),R^1 st
for b2 being Point of (TOP-REAL 2) holds b1 . b2 = tricord3 c1,c2,c3,b2
uniqueness
for b1, b2 being Function of (TOP-REAL 2),R^1 holds
( ( for b3 being Point of (TOP-REAL 2) holds b1 . b3 = tricord3 c1,c2,c3,b3 ) & ( for b3 being Point of (TOP-REAL 2) holds b2 . b3 = tricord3 c1,c2,c3,b3 ) implies b1 = b2 )
end;
:: deftheorem Def16 defines trcmap3 EUCLID_3:def 16 :
theorem Th64: :: EUCLID_3:64
for b
1, b
2, b
3, b
4 being
Point of
(TOP-REAL 2) holds
( b
2 - b
1,b
3 - b
1 are_lindependent2 implies ( b
4 in outside_of_triangle b
1,b
2,b
3 iff not ( not
tricord1 b
1,b
2,b
3,b
4 < 0 & not
tricord2 b
1,b
2,b
3,b
4 < 0 & not
tricord3 b
1,b
2,b
3,b
4 < 0 ) ) )
theorem Th65: :: EUCLID_3:65
for b
1, b
2, b
3, b
4 being
Point of
(TOP-REAL 2) holds
( b
2 - b
1,b
3 - b
1 are_lindependent2 implies ( b
4 in Triangle b
1,b
2,b
3 iff (
tricord1 b
1,b
2,b
3,b
4 >= 0 &
tricord2 b
1,b
2,b
3,b
4 >= 0 &
tricord3 b
1,b
2,b
3,b
4 >= 0 & not ( not
tricord1 b
1,b
2,b
3,b
4 = 0 & not
tricord2 b
1,b
2,b
3,b
4 = 0 & not
tricord3 b
1,b
2,b
3,b
4 = 0 ) ) ) )
theorem Th66: :: EUCLID_3:66
for b
1, b
2, b
3, b
4 being
Point of
(TOP-REAL 2) holds
( b
2 - b
1,b
3 - b
1 are_lindependent2 implies ( b
4 in Triangle b
1,b
2,b
3 iff not ( not (
tricord1 b
1,b
2,b
3,b
4 = 0 &
tricord2 b
1,b
2,b
3,b
4 >= 0 &
tricord3 b
1,b
2,b
3,b
4 >= 0 ) & not (
tricord1 b
1,b
2,b
3,b
4 >= 0 &
tricord2 b
1,b
2,b
3,b
4 = 0 &
tricord3 b
1,b
2,b
3,b
4 >= 0 ) & not (
tricord1 b
1,b
2,b
3,b
4 >= 0 &
tricord2 b
1,b
2,b
3,b
4 >= 0 &
tricord3 b
1,b
2,b
3,b
4 = 0 ) ) ) )
by Th65;
theorem Th67: :: EUCLID_3:67
for b
1, b
2, b
3, b
4 being
Point of
(TOP-REAL 2) holds
( b
2 - b
1,b
3 - b
1 are_lindependent2 implies ( b
4 in inside_of_triangle b
1,b
2,b
3 iff (
tricord1 b
1,b
2,b
3,b
4 > 0 &
tricord2 b
1,b
2,b
3,b
4 > 0 &
tricord3 b
1,b
2,b
3,b
4 > 0 ) ) )
theorem Th68: :: EUCLID_3:68