:: Cages, external approximation of Jordan's curve
:: by Czes{\l}aw Byli\'nski and Mariusz \.Zynel
::
:: Received June 22, 1999
:: Copyright (c) 1999 Association of Mizar Users


begin

Lm1: for n being Element of NAT st 1 <= n holds
(n -' 1) + 2 = n + 1
proof end;

theorem :: JORDAN9:1
canceled;

theorem :: JORDAN9:2
canceled;

theorem Th3: :: JORDAN9:3
for T being non empty TopSpace
for B, C1, C2, D being Subset of T st B is connected & C1 is_a_component_of D & C2 is_a_component_of D & B meets C1 & B meets C2 & B c= D holds
C1 = C2
proof end;

theorem Th4: :: JORDAN9:4
for D being non empty set
for f, g being FinSequence of D st ( for n being Element of NAT holds f | n = g | n ) holds
f = g
proof end;

theorem Th5: :: JORDAN9:5
for n being Element of NAT
for D being non empty set
for f being FinSequence of D st n in dom f holds
ex k being Element of NAT st
( k in dom (Rev f) & n + k = (len f) + 1 & f /. n = (Rev f) /. k )
proof end;

theorem Th6: :: JORDAN9:6
for n being Element of NAT
for D being non empty set
for f being FinSequence of D st n in dom (Rev f) holds
ex k being Element of NAT st
( k in dom f & n + k = (len f) + 1 & (Rev f) /. n = f /. k )
proof end;

begin

theorem Th7: :: JORDAN9:7
for D being non empty set
for G being Matrix of D
for f being FinSequence of D holds
( f is_sequence_on G iff Rev f is_sequence_on G )
proof end;

theorem Th8: :: JORDAN9:8
for k being Element of NAT
for G being Matrix of (TOP-REAL 2)
for f being FinSequence of (TOP-REAL 2) st f is_sequence_on G & 1 <= k & k <= len f holds
f /. k in Values G
proof end;

Lm2: for k being Element of NAT
for G being Go-board
for f being FinSequence of (TOP-REAL 2) st f is_sequence_on G & 1 <= k & k <= len f holds
ex i, j being Element of NAT st
( [i,j] in Indices G & f /. k = G * i,j )
proof end;

theorem Th9: :: JORDAN9:9
for n being Element of NAT
for f being FinSequence of (TOP-REAL 2)
for x being set st n <= len f & x in L~ (f /^ n) holds
ex i being Element of NAT st
( n + 1 <= i & i + 1 <= len f & x in LSeg f,i )
proof end;

theorem Th10: :: JORDAN9:10
for k being Element of NAT
for G being Go-board
for f being FinSequence of (TOP-REAL 2) st f is_sequence_on G & 1 <= k & k + 1 <= len f holds
( f /. k in left_cell f,k,G & f /. k in right_cell f,k,G )
proof end;

theorem Th11: :: JORDAN9:11
for k being Element of NAT
for G being Go-board
for f being FinSequence of (TOP-REAL 2) st f is_sequence_on G & 1 <= k & k + 1 <= len f holds
( Int (left_cell f,k,G) <> {} & Int (right_cell f,k,G) <> {} )
proof end;

theorem Th12: :: JORDAN9:12
for k being Element of NAT
for G being Go-board
for f being FinSequence of (TOP-REAL 2) st f is_sequence_on G & 1 <= k & k + 1 <= len f holds
( Int (left_cell f,k,G) is connected & Int (right_cell f,k,G) is connected )
proof end;

theorem Th13: :: JORDAN9:13
for k being Element of NAT
for G being Go-board
for f being FinSequence of (TOP-REAL 2) st f is_sequence_on G & 1 <= k & k + 1 <= len f holds
( Cl (Int (left_cell f,k,G)) = left_cell f,k,G & Cl (Int (right_cell f,k,G)) = right_cell f,k,G )
proof end;

theorem Th14: :: JORDAN9:14
for k being Element of NAT
for G being Go-board
for f being FinSequence of (TOP-REAL 2) st f is_sequence_on G & LSeg f,k is horizontal holds
ex j being Element of NAT st
( 1 <= j & j <= width G & ( for p being Point of (TOP-REAL 2) st p in LSeg f,k holds
p `2 = (G * 1,j) `2 ) )
proof end;

theorem Th15: :: JORDAN9:15
for k being Element of NAT
for G being Go-board
for f being FinSequence of (TOP-REAL 2) st f is_sequence_on G & LSeg f,k is vertical holds
ex i being Element of NAT st
( 1 <= i & i <= len G & ( for p being Point of (TOP-REAL 2) st p in LSeg f,k holds
p `1 = (G * i,1) `1 ) )
proof end;

theorem Th16: :: JORDAN9:16
for i, j being Element of NAT
for G being Go-board
for f being FinSequence of (TOP-REAL 2) st f is_sequence_on G & f is special & i <= len G & j <= width G holds
Int (cell G,i,j) misses L~ f
proof end;

theorem Th17: :: JORDAN9:17
for k being Element of NAT
for G being Go-board
for f being FinSequence of (TOP-REAL 2) st f is_sequence_on G & f is special & 1 <= k & k + 1 <= len f holds
( Int (left_cell f,k,G) misses L~ f & Int (right_cell f,k,G) misses L~ f )
proof end;

theorem Th18: :: JORDAN9:18
for i, j being Element of NAT
for G being Go-board st 1 <= i & i + 1 <= len G & 1 <= j & j + 1 <= width G holds
( (G * i,j) `1 = (G * i,(j + 1)) `1 & (G * i,j) `2 = (G * (i + 1),j) `2 & (G * (i + 1),(j + 1)) `1 = (G * (i + 1),j) `1 & (G * (i + 1),(j + 1)) `2 = (G * i,(j + 1)) `2 )
proof end;

theorem Th19: :: JORDAN9:19
for G being Go-board
for p being Point of (TOP-REAL 2)
for i, j being Element of NAT st 1 <= i & i + 1 <= len G & 1 <= j & j + 1 <= width G holds
( p in cell G,i,j iff ( (G * i,j) `1 <= p `1 & p `1 <= (G * (i + 1),j) `1 & (G * i,j) `2 <= p `2 & p `2 <= (G * i,(j + 1)) `2 ) )
proof end;

theorem :: JORDAN9:20
for i, j being Element of NAT
for G being Go-board st 1 <= i & i + 1 <= len G & 1 <= j & j + 1 <= width G holds
cell G,i,j = { |[r,s]| where r, s is Real : ( (G * i,j) `1 <= r & r <= (G * (i + 1),j) `1 & (G * i,j) `2 <= s & s <= (G * i,(j + 1)) `2 ) }
proof end;

theorem Th21: :: JORDAN9:21
for i, j being Element of NAT
for G being Go-board
for p being Point of (TOP-REAL 2) st 1 <= i & i + 1 <= len G & 1 <= j & j + 1 <= width G & p in Values G & p in cell G,i,j & not p = G * i,j & not p = G * i,(j + 1) & not p = G * (i + 1),(j + 1) holds
p = G * (i + 1),j
proof end;

theorem Th22: :: JORDAN9:22
for i, j being Element of NAT
for G being Go-board st 1 <= i & i + 1 <= len G & 1 <= j & j + 1 <= width G holds
( G * i,j in cell G,i,j & G * i,(j + 1) in cell G,i,j & G * (i + 1),(j + 1) in cell G,i,j & G * (i + 1),j in cell G,i,j )
proof end;

theorem Th23: :: JORDAN9:23
for i, j being Element of NAT
for G being Go-board
for p being Point of (TOP-REAL 2) st 1 <= i & i + 1 <= len G & 1 <= j & j + 1 <= width G & p in Values G & p in cell G,i,j holds
p is_extremal_in cell G,i,j
proof end;

theorem Th24: :: JORDAN9:24
for k being Element of NAT
for G being Go-board
for f being FinSequence of (TOP-REAL 2) st 2 <= len G & 2 <= width G & f is_sequence_on G & 1 <= k & k + 1 <= len f holds
ex i, j being Element of NAT st
( 1 <= i & i + 1 <= len G & 1 <= j & j + 1 <= width G & LSeg f,k c= cell G,i,j )
proof end;

theorem Th25: :: JORDAN9:25
for k being Element of NAT
for G being Go-board
for f being FinSequence of (TOP-REAL 2)
for p being Point of (TOP-REAL 2) st 2 <= len G & 2 <= width G & f is_sequence_on G & 1 <= k & k + 1 <= len f & p in Values G & p in LSeg f,k & not p = f /. k holds
p = f /. (k + 1)
proof end;

theorem :: JORDAN9:26
for i, j, k being Element of NAT
for G being Go-board st [i,j] in Indices G & 1 <= k & k <= width G holds
(G * i,j) `1 <= (G * (len G),k) `1
proof end;

theorem :: JORDAN9:27
for i, j, k being Element of NAT
for G being Go-board st [i,j] in Indices G & 1 <= k & k <= len G holds
(G * i,j) `2 <= (G * k,(width G)) `2
proof end;

theorem Th28: :: JORDAN9:28
for k being Element of NAT
for G being Go-board
for f, g being FinSequence of (TOP-REAL 2) st f is_sequence_on G & f is special & L~ g c= L~ f & 1 <= k & k + 1 <= len f holds
for A being Subset of (TOP-REAL 2) st ( A = (right_cell f,k,G) \ (L~ g) or A = (left_cell f,k,G) \ (L~ g) ) holds
A is connected
proof end;

theorem Th29: :: JORDAN9:29
for G being Go-board
for f being non constant standard special_circular_sequence st f is_sequence_on G holds
for k being Element of NAT st 1 <= k & k + 1 <= len f holds
( (right_cell f,k,G) \ (L~ f) c= RightComp f & (left_cell f,k,G) \ (L~ f) c= LeftComp f )
proof end;

begin

theorem Th30: :: JORDAN9:30
for C being non empty compact non horizontal non vertical Subset of (TOP-REAL 2)
for n being Nat ex i being Element of NAT st
( 1 <= i & i + 1 <= len (Gauge C,n) & N-min C in cell (Gauge C,n),i,((width (Gauge C,n)) -' 1) & N-min C <> (Gauge C,n) * i,((width (Gauge C,n)) -' 1) )
proof end;

theorem Th31: :: JORDAN9:31
for C being non empty compact non horizontal non vertical Subset of (TOP-REAL 2)
for n, i1, i2 being Nat st 1 <= i1 & i1 + 1 <= len (Gauge C,n) & N-min C in cell (Gauge C,n),i1,((width (Gauge C,n)) -' 1) & N-min C <> (Gauge C,n) * i1,((width (Gauge C,n)) -' 1) & 1 <= i2 & i2 + 1 <= len (Gauge C,n) & N-min C in cell (Gauge C,n),i2,((width (Gauge C,n)) -' 1) & N-min C <> (Gauge C,n) * i2,((width (Gauge C,n)) -' 1) holds
i1 = i2
proof end;

theorem Th32: :: JORDAN9:32
for C being non empty compact non horizontal non vertical Subset of (TOP-REAL 2)
for n being Nat
for f being non constant standard special_circular_sequence st f is_sequence_on Gauge C,n & ( for k being Element of NAT st 1 <= k & k + 1 <= len f holds
( left_cell f,k,(Gauge C,n) misses C & right_cell f,k,(Gauge C,n) meets C ) ) & ex i being Element of NAT st
( 1 <= i & i + 1 <= len (Gauge C,n) & f /. 1 = (Gauge C,n) * i,(width (Gauge C,n)) & f /. 2 = (Gauge C,n) * (i + 1),(width (Gauge C,n)) & N-min C in cell (Gauge C,n),i,((width (Gauge C,n)) -' 1) & N-min C <> (Gauge C,n) * i,((width (Gauge C,n)) -' 1) ) holds
N-min (L~ f) = f /. 1
proof end;

definition
let C be non empty compact non horizontal non vertical Subset of (TOP-REAL 2);
let n be Nat;
assume A1: C is connected ;
func Cage C,n -> non constant standard clockwise_oriented special_circular_sequence means :Def1: :: JORDAN9:def 1
( it is_sequence_on Gauge C,n & ex i being Element of NAT st
( 1 <= i & i + 1 <= len (Gauge C,n) & it /. 1 = (Gauge C,n) * i,(width (Gauge C,n)) & it /. 2 = (Gauge C,n) * (i + 1),(width (Gauge C,n)) & N-min C in cell (Gauge C,n),i,((width (Gauge C,n)) -' 1) & N-min C <> (Gauge C,n) * i,((width (Gauge C,n)) -' 1) ) & ( for k being Element of NAT st 1 <= k & k + 2 <= len it holds
( ( front_left_cell it,k,(Gauge C,n) misses C & front_right_cell it,k,(Gauge C,n) misses C implies it turns_right k, Gauge C,n ) & ( front_left_cell it,k,(Gauge C,n) misses C & front_right_cell it,k,(Gauge C,n) meets C implies it goes_straight k, Gauge C,n ) & ( front_left_cell it,k,(Gauge C,n) meets C implies it turns_left k, Gauge C,n ) ) ) );
existence
ex b1 being non constant standard clockwise_oriented special_circular_sequence st
( b1 is_sequence_on Gauge C,n & ex i being Element of NAT st
( 1 <= i & i + 1 <= len (Gauge C,n) & b1 /. 1 = (Gauge C,n) * i,(width (Gauge C,n)) & b1 /. 2 = (Gauge C,n) * (i + 1),(width (Gauge C,n)) & N-min C in cell (Gauge C,n),i,((width (Gauge C,n)) -' 1) & N-min C <> (Gauge C,n) * i,((width (Gauge C,n)) -' 1) ) & ( for k being Element of NAT st 1 <= k & k + 2 <= len b1 holds
( ( front_left_cell b1,k,(Gauge C,n) misses C & front_right_cell b1,k,(Gauge C,n) misses C implies b1 turns_right k, Gauge C,n ) & ( front_left_cell b1,k,(Gauge C,n) misses C & front_right_cell b1,k,(Gauge C,n) meets C implies b1 goes_straight k, Gauge C,n ) & ( front_left_cell b1,k,(Gauge C,n) meets C implies b1 turns_left k, Gauge C,n ) ) ) )
proof end;
uniqueness
for b1, b2 being non constant standard clockwise_oriented special_circular_sequence st b1 is_sequence_on Gauge C,n & ex i being Element of NAT st
( 1 <= i & i + 1 <= len (Gauge C,n) & b1 /. 1 = (Gauge C,n) * i,(width (Gauge C,n)) & b1 /. 2 = (Gauge C,n) * (i + 1),(width (Gauge C,n)) & N-min C in cell (Gauge C,n),i,((width (Gauge C,n)) -' 1) & N-min C <> (Gauge C,n) * i,((width (Gauge C,n)) -' 1) ) & ( for k being Element of NAT st 1 <= k & k + 2 <= len b1 holds
( ( front_left_cell b1,k,(Gauge C,n) misses C & front_right_cell b1,k,(Gauge C,n) misses C implies b1 turns_right k, Gauge C,n ) & ( front_left_cell b1,k,(Gauge C,n) misses C & front_right_cell b1,k,(Gauge C,n) meets C implies b1 goes_straight k, Gauge C,n ) & ( front_left_cell b1,k,(Gauge C,n) meets C implies b1 turns_left k, Gauge C,n ) ) ) & b2 is_sequence_on Gauge C,n & ex i being Element of NAT st
( 1 <= i & i + 1 <= len (Gauge C,n) & b2 /. 1 = (Gauge C,n) * i,(width (Gauge C,n)) & b2 /. 2 = (Gauge C,n) * (i + 1),(width (Gauge C,n)) & N-min C in cell (Gauge C,n),i,((width (Gauge C,n)) -' 1) & N-min C <> (Gauge C,n) * i,((width (Gauge C,n)) -' 1) ) & ( for k being Element of NAT st 1 <= k & k + 2 <= len b2 holds
( ( front_left_cell b2,k,(Gauge C,n) misses C & front_right_cell b2,k,(Gauge C,n) misses C implies b2 turns_right k, Gauge C,n ) & ( front_left_cell b2,k,(Gauge C,n) misses C & front_right_cell b2,k,(Gauge C,n) meets C implies b2 goes_straight k, Gauge C,n ) & ( front_left_cell b2,k,(Gauge C,n) meets C implies b2 turns_left k, Gauge C,n ) ) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def1 defines Cage JORDAN9:def 1 :
for C being non empty compact non horizontal non vertical Subset of (TOP-REAL 2)
for n being Nat st C is connected holds
for b3 being non constant standard clockwise_oriented special_circular_sequence holds
( b3 = Cage C,n iff ( b3 is_sequence_on Gauge C,n & ex i being Element of NAT st
( 1 <= i & i + 1 <= len (Gauge C,n) & b3 /. 1 = (Gauge C,n) * i,(width (Gauge C,n)) & b3 /. 2 = (Gauge C,n) * (i + 1),(width (Gauge C,n)) & N-min C in cell (Gauge C,n),i,((width (Gauge C,n)) -' 1) & N-min C <> (Gauge C,n) * i,((width (Gauge C,n)) -' 1) ) & ( for k being Element of NAT st 1 <= k & k + 2 <= len b3 holds
( ( front_left_cell b3,k,(Gauge C,n) misses C & front_right_cell b3,k,(Gauge C,n) misses C implies b3 turns_right k, Gauge C,n ) & ( front_left_cell b3,k,(Gauge C,n) misses C & front_right_cell b3,k,(Gauge C,n) meets C implies b3 goes_straight k, Gauge C,n ) & ( front_left_cell b3,k,(Gauge C,n) meets C implies b3 turns_left k, Gauge C,n ) ) ) ) );

theorem Th33: :: JORDAN9:33
for k, n being Element of NAT
for C being non empty compact non horizontal non vertical Subset of (TOP-REAL 2) st C is connected & 1 <= k & k + 1 <= len (Cage C,n) holds
( left_cell (Cage C,n),k,(Gauge C,n) misses C & right_cell (Cage C,n),k,(Gauge C,n) meets C )
proof end;

theorem :: JORDAN9:34
for n being Element of NAT
for C being non empty compact non horizontal non vertical Subset of (TOP-REAL 2) st C is connected holds
N-min (L~ (Cage C,n)) = (Cage C,n) /. 1
proof end;