begin
theorem Th1:
for
i,
j,
k being
Element of
NAT for
f being
FinSequence of the
carrier of
(TOP-REAL 2) for
G being
Go-board st
f is_sequence_on G &
LSeg (
(G * (i,j)),
(G * (i,k)))
meets L~ f &
[i,j] in Indices G &
[i,k] in Indices G &
j <= k holds
ex
n being
Element of
NAT st
(
j <= n &
n <= k &
(G * (i,n)) `2 = lower_bound (proj2 .: ((LSeg ((G * (i,j)),(G * (i,k)))) /\ (L~ f))) )
theorem
for
i,
j,
k being
Element of
NAT for
f being
FinSequence of the
carrier of
(TOP-REAL 2) for
G being
Go-board st
f is_sequence_on G &
LSeg (
(G * (i,j)),
(G * (i,k)))
meets L~ f &
[i,j] in Indices G &
[i,k] in Indices G &
j <= k holds
ex
n being
Element of
NAT st
(
j <= n &
n <= k &
(G * (i,n)) `2 = upper_bound (proj2 .: ((LSeg ((G * (i,j)),(G * (i,k)))) /\ (L~ f))) )
theorem
for
j,
i,
k being
Element of
NAT for
f being
FinSequence of the
carrier of
(TOP-REAL 2) for
G being
Go-board st
f is_sequence_on G &
LSeg (
(G * (j,i)),
(G * (k,i)))
meets L~ f &
[j,i] in Indices G &
[k,i] in Indices G &
j <= k holds
ex
n being
Element of
NAT st
(
j <= n &
n <= k &
(G * (n,i)) `1 = lower_bound (proj1 .: ((LSeg ((G * (j,i)),(G * (k,i)))) /\ (L~ f))) )
theorem
for
j,
i,
k being
Element of
NAT for
f being
FinSequence of the
carrier of
(TOP-REAL 2) for
G being
Go-board st
f is_sequence_on G &
LSeg (
(G * (j,i)),
(G * (k,i)))
meets L~ f &
[j,i] in Indices G &
[k,i] in Indices G &
j <= k holds
ex
n being
Element of
NAT st
(
j <= n &
n <= k &
(G * (n,i)) `1 = upper_bound (proj1 .: ((LSeg ((G * (j,i)),(G * (k,i)))) /\ (L~ f))) )
theorem Th5:
theorem Th6:
theorem Th7:
theorem Th8:
theorem Th9:
theorem Th10:
theorem Th11:
for
G being
Go-board for
p being
Point of
(TOP-REAL 2) for
f being
FinSequence of
(TOP-REAL 2) st
f is_sequence_on G & ex
i,
j being
Element of
NAT st
(
[i,j] in Indices G &
p = G * (
i,
j) ) & ( for
i1,
j1,
i2,
j2 being
Element of
NAT st
[i1,j1] in Indices G &
[i2,j2] in Indices G &
p = G * (
i1,
j1) &
f /. 1
= G * (
i2,
j2) holds
(abs (i2 - i1)) + (abs (j2 - j1)) = 1 ) holds
<*p*> ^ f is_sequence_on G
theorem Th12:
theorem
for
i being
Element of
NAT for
C being non
empty being_simple_closed_curve compact non
horizontal non
vertical Subset of
(TOP-REAL 2) for
p being
Point of
(TOP-REAL 2) st
p `1 = ((W-bound C) + (E-bound C)) / 2 &
p `2 = lower_bound (proj2 .: ((LSeg (((Gauge (C,1)) * ((Center (Gauge (C,1))),1)),((Gauge (C,1)) * ((Center (Gauge (C,1))),(width (Gauge (C,1))))))) /\ (Upper_Arc (L~ (Cage (C,(i + 1))))))) holds
ex
j being
Element of
NAT st
( 1
<= j &
j <= width (Gauge (C,(i + 1))) &
p = (Gauge (C,(i + 1))) * (
(Center (Gauge (C,(i + 1)))),
j) )