:: JORDAN20 semantic presentation
theorem Th1: :: JORDAN20:1
theorem Th2: :: JORDAN20:2
theorem Th3: :: JORDAN20:3
theorem Th4: :: JORDAN20:4
theorem Th5: :: JORDAN20:5
theorem Th6: :: JORDAN20:6
theorem Th7: :: JORDAN20:7
definition
let c
1 be
Subset of
(TOP-REAL 2);
let c
2, c
3, c
4 be
Point of
(TOP-REAL 2);
let c
5 be
Real;
pred c
4 is_Lin c
1,c
2,c
3,c
5 means :
Def1:
:: JORDAN20:def 1
( a
1 is_an_arc_of a
2,a
3 & a
4 in a
1 & a
4 `1 = a
5 & ex b
1 being
Point of
(TOP-REAL 2) st
( b
1 `1 < a
5 &
LE b
1,a
4,a
1,a
2,a
3 & ( for b
2 being
Point of
(TOP-REAL 2) holds
(
LE b
1,b
2,a
1,a
2,a
3 &
LE b
2,a
4,a
1,a
2,a
3 implies b
2 `1 <= a
5 ) ) ) );
pred c
4 is_Rin c
1,c
2,c
3,c
5 means :
Def2:
:: JORDAN20:def 2
( a
1 is_an_arc_of a
2,a
3 & a
4 in a
1 & a
4 `1 = a
5 & ex b
1 being
Point of
(TOP-REAL 2) st
( b
1 `1 > a
5 &
LE b
1,a
4,a
1,a
2,a
3 & ( for b
2 being
Point of
(TOP-REAL 2) holds
(
LE b
1,b
2,a
1,a
2,a
3 &
LE b
2,a
4,a
1,a
2,a
3 implies b
2 `1 >= a
5 ) ) ) );
pred c
4 is_Lout c
1,c
2,c
3,c
5 means :
Def3:
:: JORDAN20:def 3
( a
1 is_an_arc_of a
2,a
3 & a
4 in a
1 & a
4 `1 = a
5 & ex b
1 being
Point of
(TOP-REAL 2) st
( b
1 `1 < a
5 &
LE a
4,b
1,a
1,a
2,a
3 & ( for b
2 being
Point of
(TOP-REAL 2) holds
(
LE b
2,b
1,a
1,a
2,a
3 &
LE a
4,b
2,a
1,a
2,a
3 implies b
2 `1 <= a
5 ) ) ) );
pred c
4 is_Rout c
1,c
2,c
3,c
5 means :
Def4:
:: JORDAN20:def 4
( a
1 is_an_arc_of a
2,a
3 & a
4 in a
1 & a
4 `1 = a
5 & ex b
1 being
Point of
(TOP-REAL 2) st
( b
1 `1 > a
5 &
LE a
4,b
1,a
1,a
2,a
3 & ( for b
2 being
Point of
(TOP-REAL 2) holds
(
LE b
2,b
1,a
1,a
2,a
3 &
LE a
4,b
2,a
1,a
2,a
3 implies b
2 `1 >= a
5 ) ) ) );
pred c
4 is_OSin c
1,c
2,c
3,c
5 means :
Def5:
:: JORDAN20:def 5
( a
1 is_an_arc_of a
2,a
3 & a
4 in a
1 & a
4 `1 = a
5 & ex b
1 being
Point of
(TOP-REAL 2) st
(
LE b
1,a
4,a
1,a
2,a
3 & ( for b
2 being
Point of
(TOP-REAL 2) holds
(
LE b
1,b
2,a
1,a
2,a
3 &
LE b
2,a
4,a
1,a
2,a
3 implies b
2 `1 = a
5 ) ) & ( for b
2 being
Point of
(TOP-REAL 2) holds
(
LE b
2,b
1,a
1,a
2,a
3 & b
2 <> b
1 implies ( ex b
3 being
Point of
(TOP-REAL 2) st
(
LE b
2,b
3,a
1,a
2,a
3 &
LE b
3,b
1,a
1,a
2,a
3 & b
3 `1 > a
5 ) & ex b
3 being
Point of
(TOP-REAL 2) st
(
LE b
2,b
3,a
1,a
2,a
3 &
LE b
3,b
1,a
1,a
2,a
3 & b
3 `1 < a
5 ) ) ) ) ) );
pred c
4 is_OSout c
1,c
2,c
3,c
5 means :
Def6:
:: JORDAN20:def 6
( a
1 is_an_arc_of a
2,a
3 & a
4 in a
1 & a
4 `1 = a
5 & ex b
1 being
Point of
(TOP-REAL 2) st
(
LE a
4,b
1,a
1,a
2,a
3 & ( for b
2 being
Point of
(TOP-REAL 2) holds
(
LE b
2,b
1,a
1,a
2,a
3 &
LE a
4,b
2,a
1,a
2,a
3 implies b
2 `1 = a
5 ) ) & ( for b
2 being
Point of
(TOP-REAL 2) holds
(
LE b
1,b
2,a
1,a
2,a
3 & b
2 <> b
1 implies ( ex b
3 being
Point of
(TOP-REAL 2) st
(
LE b
3,b
2,a
1,a
2,a
3 &
LE b
1,b
3,a
1,a
2,a
3 & b
3 `1 > a
5 ) & ex b
3 being
Point of
(TOP-REAL 2) st
(
LE b
3,b
2,a
1,a
2,a
3 &
LE b
1,b
3,a
1,a
2,a
3 & b
3 `1 < a
5 ) ) ) ) ) );
correctness
;
end;
:: deftheorem Def1 defines is_Lin JORDAN20:def 1 :
for b
1 being
Subset of
(TOP-REAL 2)for b
2, b
3, b
4 being
Point of
(TOP-REAL 2)for b
5 being
Real holds
( b
4 is_Lin b
1,b
2,b
3,b
5 iff ( b
1 is_an_arc_of b
2,b
3 & b
4 in b
1 & b
4 `1 = b
5 & ex b
6 being
Point of
(TOP-REAL 2) st
( b
6 `1 < b
5 &
LE b
6,b
4,b
1,b
2,b
3 & ( for b
7 being
Point of
(TOP-REAL 2) holds
(
LE b
6,b
7,b
1,b
2,b
3 &
LE b
7,b
4,b
1,b
2,b
3 implies b
7 `1 <= b
5 ) ) ) ) );
:: deftheorem Def2 defines is_Rin JORDAN20:def 2 :
for b
1 being
Subset of
(TOP-REAL 2)for b
2, b
3, b
4 being
Point of
(TOP-REAL 2)for b
5 being
Real holds
( b
4 is_Rin b
1,b
2,b
3,b
5 iff ( b
1 is_an_arc_of b
2,b
3 & b
4 in b
1 & b
4 `1 = b
5 & ex b
6 being
Point of
(TOP-REAL 2) st
( b
6 `1 > b
5 &
LE b
6,b
4,b
1,b
2,b
3 & ( for b
7 being
Point of
(TOP-REAL 2) holds
(
LE b
6,b
7,b
1,b
2,b
3 &
LE b
7,b
4,b
1,b
2,b
3 implies b
7 `1 >= b
5 ) ) ) ) );
:: deftheorem Def3 defines is_Lout JORDAN20:def 3 :
for b
1 being
Subset of
(TOP-REAL 2)for b
2, b
3, b
4 being
Point of
(TOP-REAL 2)for b
5 being
Real holds
( b
4 is_Lout b
1,b
2,b
3,b
5 iff ( b
1 is_an_arc_of b
2,b
3 & b
4 in b
1 & b
4 `1 = b
5 & ex b
6 being
Point of
(TOP-REAL 2) st
( b
6 `1 < b
5 &
LE b
4,b
6,b
1,b
2,b
3 & ( for b
7 being
Point of
(TOP-REAL 2) holds
(
LE b
7,b
6,b
1,b
2,b
3 &
LE b
4,b
7,b
1,b
2,b
3 implies b
7 `1 <= b
5 ) ) ) ) );
:: deftheorem Def4 defines is_Rout JORDAN20:def 4 :
for b
1 being
Subset of
(TOP-REAL 2)for b
2, b
3, b
4 being
Point of
(TOP-REAL 2)for b
5 being
Real holds
( b
4 is_Rout b
1,b
2,b
3,b
5 iff ( b
1 is_an_arc_of b
2,b
3 & b
4 in b
1 & b
4 `1 = b
5 & ex b
6 being
Point of
(TOP-REAL 2) st
( b
6 `1 > b
5 &
LE b
4,b
6,b
1,b
2,b
3 & ( for b
7 being
Point of
(TOP-REAL 2) holds
(
LE b
7,b
6,b
1,b
2,b
3 &
LE b
4,b
7,b
1,b
2,b
3 implies b
7 `1 >= b
5 ) ) ) ) );
:: deftheorem Def5 defines is_OSin JORDAN20:def 5 :
for b
1 being
Subset of
(TOP-REAL 2)for b
2, b
3, b
4 being
Point of
(TOP-REAL 2)for b
5 being
Real holds
( b
4 is_OSin b
1,b
2,b
3,b
5 iff ( b
1 is_an_arc_of b
2,b
3 & b
4 in b
1 & b
4 `1 = b
5 & ex b
6 being
Point of
(TOP-REAL 2) st
(
LE b
6,b
4,b
1,b
2,b
3 & ( for b
7 being
Point of
(TOP-REAL 2) holds
(
LE b
6,b
7,b
1,b
2,b
3 &
LE b
7,b
4,b
1,b
2,b
3 implies b
7 `1 = b
5 ) ) & ( for b
7 being
Point of
(TOP-REAL 2) holds
(
LE b
7,b
6,b
1,b
2,b
3 & b
7 <> b
6 implies ( ex b
8 being
Point of
(TOP-REAL 2) st
(
LE b
7,b
8,b
1,b
2,b
3 &
LE b
8,b
6,b
1,b
2,b
3 & b
8 `1 > b
5 ) & ex b
8 being
Point of
(TOP-REAL 2) st
(
LE b
7,b
8,b
1,b
2,b
3 &
LE b
8,b
6,b
1,b
2,b
3 & b
8 `1 < b
5 ) ) ) ) ) ) );
:: deftheorem Def6 defines is_OSout JORDAN20:def 6 :
for b
1 being
Subset of
(TOP-REAL 2)for b
2, b
3, b
4 being
Point of
(TOP-REAL 2)for b
5 being
Real holds
( b
4 is_OSout b
1,b
2,b
3,b
5 iff ( b
1 is_an_arc_of b
2,b
3 & b
4 in b
1 & b
4 `1 = b
5 & ex b
6 being
Point of
(TOP-REAL 2) st
(
LE b
4,b
6,b
1,b
2,b
3 & ( for b
7 being
Point of
(TOP-REAL 2) holds
(
LE b
7,b
6,b
1,b
2,b
3 &
LE b
4,b
7,b
1,b
2,b
3 implies b
7 `1 = b
5 ) ) & ( for b
7 being
Point of
(TOP-REAL 2) holds
(
LE b
6,b
7,b
1,b
2,b
3 & b
7 <> b
6 implies ( ex b
8 being
Point of
(TOP-REAL 2) st
(
LE b
8,b
7,b
1,b
2,b
3 &
LE b
6,b
8,b
1,b
2,b
3 & b
8 `1 > b
5 ) & ex b
8 being
Point of
(TOP-REAL 2) st
(
LE b
8,b
7,b
1,b
2,b
3 &
LE b
6,b
8,b
1,b
2,b
3 & b
8 `1 < b
5 ) ) ) ) ) ) );
theorem Th8: :: JORDAN20:8
theorem Th9: :: JORDAN20:9
for b
1 being non
empty Subset of
(TOP-REAL 2)for b
2, b
3, b
4 being
Point of
(TOP-REAL 2)for b
5 being
Real holds
not ( b
1 is_an_arc_of b
2,b
3 & b
2 `1 < b
5 & b
3 `1 > b
5 & b
4 in b
1 & b
4 `1 = b
5 & not b
4 is_Lin b
1,b
2,b
3,b
5 & not b
4 is_Rin b
1,b
2,b
3,b
5 & not b
4 is_OSin b
1,b
2,b
3,b
5 )
theorem Th10: :: JORDAN20:10
for b
1 being non
empty Subset of
(TOP-REAL 2)for b
2, b
3, b
4 being
Point of
(TOP-REAL 2)for b
5 being
Real holds
not ( b
1 is_an_arc_of b
2,b
3 & b
2 `1 < b
5 & b
3 `1 > b
5 & b
4 in b
1 & b
4 `1 = b
5 & not b
4 is_Lout b
1,b
2,b
3,b
5 & not b
4 is_Rout b
1,b
2,b
3,b
5 & not b
4 is_OSout b
1,b
2,b
3,b
5 )
theorem Th11: :: JORDAN20:11
theorem Th12: :: JORDAN20:12
theorem Th13: :: JORDAN20:13
theorem Th14: :: JORDAN20:14
theorem Th15: :: JORDAN20:15
theorem Th16: :: JORDAN20:16
theorem Th17: :: JORDAN20:17
theorem Th18: :: JORDAN20:18
theorem Th19: :: JORDAN20:19
for b
1 being non
empty Subset of
(TOP-REAL 2)for b
2, b
3, b
4, b
5 being
Point of
(TOP-REAL 2) holds
not ( b
1 is_an_arc_of b
2,b
3 & b
4 in b
1 & b
5 in b
1 & not
LE b
4,b
5,b
1,b
2,b
3 & not
LE b
5,b
4,b
1,b
2,b
3 )
theorem Th20: :: JORDAN20:20
theorem Th21: :: JORDAN20:21
theorem Th22: :: JORDAN20:22
theorem Th23: :: JORDAN20:23
for b
1 being non
empty Subset of
(TOP-REAL 2)for b
2, b
3, b
4, b
5, b
6 being
Point of
(TOP-REAL 2) holds
( b
1 is_an_arc_of b
2,b
3 &
LE b
4,b
5,b
1,b
2,b
3 &
LE b
5,b
6,b
1,b
2,b
3 implies
(Segment b1,b2,b3,b4,b5) \/ (Segment b1,b2,b3,b5,b6) = Segment b
1,b
2,b
3,b
4,b
6 )
theorem Th24: :: JORDAN20:24
for b
1 being non
empty Subset of
(TOP-REAL 2)for b
2, b
3, b
4, b
5, b
6 being
Point of
(TOP-REAL 2) holds
( b
1 is_an_arc_of b
2,b
3 &
LE b
4,b
5,b
1,b
2,b
3 &
LE b
5,b
6,b
1,b
2,b
3 implies
(Segment b1,b2,b3,b4,b5) /\ (Segment b1,b2,b3,b5,b6) = {b5} )
theorem Th25: :: JORDAN20:25
theorem Th26: :: JORDAN20:26
canceled;
theorem Th27: :: JORDAN20:27
canceled;
Lemma26:
for b1 being non empty TopSpace
for b2, b3, b4 being Point of b1
for b5 being Path of b2,b3
for b6 being Path of b3,b4 holds
( b5 is continuous & b6 is continuous & b5 . 0 = b2 & b5 . 1 = b3 & b6 . 0 = b3 & b6 . 1 = b4 implies ( b5 + b6 is continuous & (b5 + b6) . 0 = b2 & (b5 + b6) . 1 = b4 ) )
by BORSUK_2:17;
theorem Th28: :: JORDAN20:28
for b
1, b
2 being non
empty Subset of
(TOP-REAL 2)for b
3, b
4, b
5, b
6 being
Point of
(TOP-REAL 2) holds
( b
1 is_an_arc_of b
3,b
4 & b
2 is_an_arc_of b
5,b
6 &
LE b
5,b
6,b
1,b
3,b
4 & b
2 c= b
1 implies b
2 = Segment b
1,b
3,b
4,b
5,b
6 )
theorem Th29: :: JORDAN20:29
for b
1 being non
empty Subset of
(TOP-REAL 2)for b
2, b
3, b
4, b
5, b
6 being
Point of
(TOP-REAL 2)for b
7 being
Real holds
( b
2 `1 < b
7 & b
3 `1 > b
7 & b
4 is_Lin b
1,b
2,b
3,b
7 & b
5 `1 = b
7 &
LSeg b
4,b
5 c= b
1 & b
6 in LSeg b
4,b
5 implies b
6 is_Lin b
1,b
2,b
3,b
7 )
theorem Th30: :: JORDAN20:30
for b
1 being non
empty Subset of
(TOP-REAL 2)for b
2, b
3, b
4, b
5, b
6 being
Point of
(TOP-REAL 2)for b
7 being
Real holds
( b
2 `1 < b
7 & b
3 `1 > b
7 & b
4 is_Rin b
1,b
2,b
3,b
7 & b
5 `1 = b
7 &
LSeg b
4,b
5 c= b
1 & b
6 in LSeg b
4,b
5 implies b
6 is_Rin b
1,b
2,b
3,b
7 )
theorem Th31: :: JORDAN20:31
for b
1 being non
empty Subset of
(TOP-REAL 2)for b
2, b
3, b
4, b
5, b
6 being
Point of
(TOP-REAL 2)for b
7 being
Real holds
( b
2 `1 < b
7 & b
3 `1 > b
7 & b
4 is_Lout b
1,b
2,b
3,b
7 & b
5 `1 = b
7 &
LSeg b
4,b
5 c= b
1 & b
6 in LSeg b
4,b
5 implies b
6 is_Lout b
1,b
2,b
3,b
7 )
theorem Th32: :: JORDAN20:32
for b
1 being non
empty Subset of
(TOP-REAL 2)for b
2, b
3, b
4, b
5, b
6 being
Point of
(TOP-REAL 2)for b
7 being
Real holds
( b
2 `1 < b
7 & b
3 `1 > b
7 & b
4 is_Rout b
1,b
2,b
3,b
7 & b
5 `1 = b
7 &
LSeg b
4,b
5 c= b
1 & b
6 in LSeg b
4,b
5 implies b
6 is_Rout b
1,b
2,b
3,b
7 )
theorem Th33: :: JORDAN20:33
theorem Th34: :: JORDAN20:34