:: GOBOARD1 semantic presentation  Show TPTP formulae Show IDV graph for whole article:: Showing IDV graph ... (Click the Palm Trees again to close it)

theorem Th1: :: GOBOARD1:1  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for r, s being Real holds
( abs (r - s) = 1 iff ( ( r > s & r = s + 1 ) or ( r < s & s = r + 1 ) ) )
proof end;

theorem Th2: :: GOBOARD1:2  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for i, j, n, m being Nat holds
( (abs (i - j)) + (abs (n - m)) = 1 iff ( ( abs (i - j) = 1 & n = m ) or ( abs (n - m) = 1 & i = j ) ) )
proof end;

theorem Th3: :: GOBOARD1:3  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for n being Nat holds
( n > 1 iff ex m being Nat st
( n = m + 1 & m > 0 ) )
proof end;

scheme :: GOBOARD1:sch 1
FinSeqDChoice{ F1() -> non empty set , F2() -> Nat, P1[ set , set ] } :
ex f being FinSequence of F1() st
( len f = F2() & ( for n being Nat st n in Seg F2() holds
P1[n,f /. n] ) )
provided
A1: for n being Nat st n in Seg F2() holds
ex d being Element of F1() st P1[n,d]
proof end;

theorem :: GOBOARD1:4  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
canceled;

theorem :: GOBOARD1:5  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
canceled;

theorem :: GOBOARD1:6  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
canceled;

theorem Th7: :: GOBOARD1:7  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for f being FinSequence
for n, m, k being Nat st len f = m + 1 & n in dom f & k in Seg m & not (Del f,n) . k = f . k holds
(Del f,n) . k = f . (k + 1)
proof end;

theorem :: GOBOARD1:8  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
canceled;

theorem :: GOBOARD1:9  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
canceled;

definition
let f be FinSequence of REAL ;
let k be Nat;
:: original: .
redefine func f . k -> Real;
coherence
f . k is Real
proof end;
end;

definition
let IT be FinSequence of REAL ;
attr IT is increasing means :Def1: :: GOBOARD1:def 1
for n, m being Nat st n in dom IT & m in dom IT & n < m holds
IT . n < IT . m;
end;

:: deftheorem Def1 defines increasing GOBOARD1:def 1 :
for IT being FinSequence of REAL holds
( IT is increasing iff for n, m being Nat st n in dom IT & m in dom IT & n < m holds
IT . n < IT . m );

definition
let f be FinSequence;
redefine attr f is constant means :Def2: :: GOBOARD1:def 2
for n, m being Nat st n in dom f & m in dom f holds
f . n = f . m;
compatibility
( f is constant iff for n, m being Nat st n in dom f & m in dom f holds
f . n = f . m )
proof end;
end;

:: deftheorem Def2 defines constant GOBOARD1:def 2 :
for f being FinSequence holds
( f is constant iff for n, m being Nat st n in dom f & m in dom f holds
f . n = f . m );

registration
cluster non empty increasing FinSequence of REAL ;
existence
ex b1 being FinSequence of REAL st
( not b1 is empty & b1 is increasing )
proof end;
end;

registration
let D be non empty set ;
cluster non empty FinSequence of D;
existence
not for b1 being FinSequence of D holds b1 is empty
proof end;
end;

registration
cluster constant FinSequence of REAL ;
existence
ex b1 being FinSequence of REAL st b1 is constant
proof end;
end;

definition
let f be FinSequence of (TOP-REAL 2);
func X_axis f -> FinSequence of REAL means :Def3: :: GOBOARD1:def 3
( len it = len f & ( for n being Nat st n in dom it holds
it . n = (f /. n) `1 ) );
existence
ex b1 being FinSequence of REAL st
( len b1 = len f & ( for n being Nat st n in dom b1 holds
b1 . n = (f /. n) `1 ) )
proof end;
uniqueness
for b1, b2 being FinSequence of REAL st len b1 = len f & ( for n being Nat st n in dom b1 holds
b1 . n = (f /. n) `1 ) & len b2 = len f & ( for n being Nat st n in dom b2 holds
b2 . n = (f /. n) `1 ) holds
b1 = b2
proof end;
func Y_axis f -> FinSequence of REAL means :Def4: :: GOBOARD1:def 4
( len it = len f & ( for n being Nat st n in dom it holds
it . n = (f /. n) `2 ) );
existence
ex b1 being FinSequence of REAL st
( len b1 = len f & ( for n being Nat st n in dom b1 holds
b1 . n = (f /. n) `2 ) )
proof end;
uniqueness
for b1, b2 being FinSequence of REAL st len b1 = len f & ( for n being Nat st n in dom b1 holds
b1 . n = (f /. n) `2 ) & len b2 = len f & ( for n being Nat st n in dom b2 holds
b2 . n = (f /. n) `2 ) holds
b1 = b2
proof end;
end;

:: deftheorem Def3 defines X_axis GOBOARD1:def 3 :
for f being FinSequence of (TOP-REAL 2)
for b2 being FinSequence of REAL holds
( b2 = X_axis f iff ( len b2 = len f & ( for n being Nat st n in dom b2 holds
b2 . n = (f /. n) `1 ) ) );

:: deftheorem Def4 defines Y_axis GOBOARD1:def 4 :
for f being FinSequence of (TOP-REAL 2)
for b2 being FinSequence of REAL holds
( b2 = Y_axis f iff ( len b2 = len f & ( for n being Nat st n in dom b2 holds
b2 . n = (f /. n) `2 ) ) );

theorem :: GOBOARD1:10  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
canceled;

theorem :: GOBOARD1:11  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
canceled;

theorem :: GOBOARD1:12  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
canceled;

theorem :: GOBOARD1:13  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
canceled;

theorem Th14: :: GOBOARD1:14  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for v being FinSequence of REAL
for n, i, m being Nat st v <> {} & rng v c= Seg n & v . (len v) = n & ( for k being Nat st 1 <= k & k <= (len v) - 1 holds
for r, s being Real st r = v . k & s = v . (k + 1) & not abs (r - s) = 1 holds
r = s ) & i in Seg n & i + 1 in Seg n & m in dom v & v . m = i & ( for k being Nat st k in dom v & v . k = i holds
k <= m ) holds
( m + 1 in dom v & v . (m + 1) = i + 1 )
proof end;

theorem :: GOBOARD1:15  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for v being FinSequence of REAL
for n being Nat st v <> {} & rng v c= Seg n & v . 1 = 1 & v . (len v) = n & ( for k being Nat st 1 <= k & k <= (len v) - 1 holds
for r, s being Real st r = v . k & s = v . (k + 1) & not abs (r - s) = 1 holds
r = s ) holds
( ( for i being Nat st i in Seg n holds
ex k being Nat st
( k in dom v & v . k = i ) ) & ( for m, k, i being Nat
for r being Real st m in dom v & v . m = i & ( for j being Nat st j in dom v & v . j = i holds
j <= m ) & m < k & k in dom v & r = v . k holds
i < r ) )
proof end;

theorem Th16: :: GOBOARD1:16  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for f being FinSequence of (TOP-REAL 2)
for i being Nat st i in dom f & 2 <= len f holds
f /. i in L~ f
proof end;

theorem Th17: :: GOBOARD1:17  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for D being non empty set
for M being Matrix of D
for i, j being Nat st j in dom M & i in Seg (width M) holds
(Col M,i) . j = (Line M,j) . i
proof end;

definition
let D be non empty set ;
let M be Matrix of D;
:: original: empty-yielding
redefine attr M is empty-yielding means :Def5: :: GOBOARD1:def 5
( 0 = len M or 0 = width M );
compatibility
( M is empty-yielding iff ( 0 = len M or 0 = width M ) )
proof end;
end;

:: deftheorem Def5 defines empty-yielding GOBOARD1:def 5 :
for D being non empty set
for M being Matrix of D holds
( M is empty-yielding iff ( 0 = len M or 0 = width M ) );

definition
let M be Matrix of (TOP-REAL 2);
attr M is X_equal-in-line means :Def6: :: GOBOARD1:def 6
for n being Nat st n in dom M holds
X_axis (Line M,n) is constant;
attr M is Y_equal-in-column means :Def7: :: GOBOARD1:def 7
for n being Nat st n in Seg (width M) holds
Y_axis (Col M,n) is constant;
attr M is Y_increasing-in-line means :Def8: :: GOBOARD1:def 8
for n being Nat st n in dom M holds
Y_axis (Line M,n) is increasing;
attr M is X_increasing-in-column means :Def9: :: GOBOARD1:def 9
for n being Nat st n in Seg (width M) holds
X_axis (Col M,n) is increasing;
end;

:: deftheorem Def6 defines X_equal-in-line GOBOARD1:def 6 :
for M being Matrix of (TOP-REAL 2) holds
( M is X_equal-in-line iff for n being Nat st n in dom M holds
X_axis (Line M,n) is constant );

:: deftheorem Def7 defines Y_equal-in-column GOBOARD1:def 7 :
for M being Matrix of (TOP-REAL 2) holds
( M is Y_equal-in-column iff for n being Nat st n in Seg (width M) holds
Y_axis (Col M,n) is constant );

:: deftheorem Def8 defines Y_increasing-in-line GOBOARD1:def 8 :
for M being Matrix of (TOP-REAL 2) holds
( M is Y_increasing-in-line iff for n being Nat st n in dom M holds
Y_axis (Line M,n) is increasing );

:: deftheorem Def9 defines X_increasing-in-column GOBOARD1:def 9 :
for M being Matrix of (TOP-REAL 2) holds
( M is X_increasing-in-column iff for n being Nat st n in Seg (width M) holds
X_axis (Col M,n) is increasing );

Lm1: for D being non empty set
for M being Matrix of D holds
( not M is empty-yielding iff ( 0 < len M & 0 < width M ) )
proof end;

registration
cluster V4 X_equal-in-line Y_equal-in-column Y_increasing-in-line X_increasing-in-column FinSequence of the carrier of (TOP-REAL 2) * ;
existence
ex b1 being Matrix of (TOP-REAL 2) st
( not b1 is empty-yielding & b1 is X_equal-in-line & b1 is Y_equal-in-column & b1 is Y_increasing-in-line & b1 is X_increasing-in-column )
proof end;
end;

theorem :: GOBOARD1:18  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
canceled;

theorem Th19: :: GOBOARD1:19  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for M being X_equal-in-line X_increasing-in-column Matrix of (TOP-REAL 2)
for x being set
for n, m being Nat st x in rng (Line M,n) & x in rng (Line M,m) & n in dom M & m in dom M holds
n = m
proof end;

theorem Th20: :: GOBOARD1:20  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for M being Y_equal-in-column Y_increasing-in-line Matrix of (TOP-REAL 2)
for x being set
for n, m being Nat st x in rng (Col M,n) & x in rng (Col M,m) & n in Seg (width M) & m in Seg (width M) holds
n = m
proof end;

definition
mode Go-board is V4 X_equal-in-line Y_equal-in-column Y_increasing-in-line X_increasing-in-column Matrix of (TOP-REAL 2);
end;

theorem :: GOBOARD1:21  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for m, k, i, j being Nat
for x being set
for G being Go-board st x = G * m,k & x = G * i,j & [m,k] in Indices G & [i,j] in Indices G holds
( m = i & k = j )
proof end;

theorem :: GOBOARD1:22  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for f being FinSequence of (TOP-REAL 2)
for m being Nat
for G being Go-board st m in dom f & f /. 1 in rng (Col G,1) holds
(f | m) /. 1 in rng (Col G,1)
proof end;

theorem :: GOBOARD1:23  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for f being FinSequence of (TOP-REAL 2)
for m being Nat
for G being Go-board st m in dom f & f /. m in rng (Col G,(width G)) holds
(f | m) /. (len (f | m)) in rng (Col G,(width G))
proof end;

theorem Th24: :: GOBOARD1:24  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for f being FinSequence of (TOP-REAL 2)
for i, n, m, k being Nat
for G being Go-board st rng f misses rng (Col G,i) & f /. n = G * m,k & n in dom f & m in dom G holds
i <> k
proof end;

definition
let G be Go-board;
let i be Nat;
assume A1: ( i in Seg (width G) & width G > 1 ) ;
func DelCol G,i -> Go-board means :Def10: :: GOBOARD1:def 10
( len it = len G & ( for k being Nat st k in dom G holds
it . k = Del (Line G,k),i ) );
existence
ex b1 being Go-board st
( len b1 = len G & ( for k being Nat st k in dom G holds
b1 . k = Del (Line G,k),i ) )
proof end;
uniqueness
for b1, b2 being Go-board st len b1 = len G & ( for k being Nat st k in dom G holds
b1 . k = Del (Line G,k),i ) & len b2 = len G & ( for k being Nat st k in dom G holds
b2 . k = Del (Line G,k),i ) holds
b1 = b2
proof end;
end;

:: deftheorem Def10 defines DelCol GOBOARD1:def 10 :
for G being Go-board
for i being Nat st i in Seg (width G) & width G > 1 holds
for b3 being Go-board holds
( b3 = DelCol G,i iff ( len b3 = len G & ( for k being Nat st k in dom G holds
b3 . k = Del (Line G,k),i ) ) );

theorem Th25: :: GOBOARD1:25  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for i, k being Nat
for G being Go-board st i in Seg (width G) & width G > 1 & k in dom G holds
Line (DelCol G,i),k = Del (Line G,k),i
proof end;

theorem Th26: :: GOBOARD1:26  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for i, m being Nat
for G being Go-board st i in Seg (width G) & width G = m + 1 & m > 0 holds
width (DelCol G,i) = m
proof end;

theorem :: GOBOARD1:27  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for i being Nat
for G being Go-board st i in Seg (width G) & width G > 1 holds
width G = (width (DelCol G,i)) + 1
proof end;

theorem Th28: :: GOBOARD1:28  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for i, n, m being Nat
for G being Go-board st i in Seg (width G) & width G > 1 & n in dom G & m in Seg (width (DelCol G,i)) holds
(DelCol G,i) * n,m = (Del (Line G,n),i) . m
proof end;

theorem Th29: :: GOBOARD1:29  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for i, m, k being Nat
for G being Go-board st i in Seg (width G) & width G = m + 1 & m > 0 & 1 <= k & k < i holds
( Col (DelCol G,i),k = Col G,k & k in Seg (width (DelCol G,i)) & k in Seg (width G) )
proof end;

theorem Th30: :: GOBOARD1:30  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for i, m, k being Nat
for G being Go-board st i in Seg (width G) & width G = m + 1 & m > 0 & i <= k & k <= m holds
( Col (DelCol G,i),k = Col G,(k + 1) & k in Seg (width (DelCol G,i)) & k + 1 in Seg (width G) )
proof end;

theorem Th31: :: GOBOARD1:31  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for i, m, n, k being Nat
for G being Go-board st i in Seg (width G) & width G = m + 1 & m > 0 & n in dom G & 1 <= k & k < i holds
( (DelCol G,i) * n,k = G * n,k & k in Seg (width G) )
proof end;

theorem Th32: :: GOBOARD1:32  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for i, m, n, k being Nat
for G being Go-board st i in Seg (width G) & width G = m + 1 & m > 0 & n in dom G & i <= k & k <= m holds
( (DelCol G,i) * n,k = G * n,(k + 1) & k + 1 in Seg (width G) )
proof end;

theorem :: GOBOARD1:33  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for m, k being Nat
for G being Go-board st width G = m + 1 & m > 0 & k in Seg m holds
( Col (DelCol G,1),k = Col G,(k + 1) & k in Seg (width (DelCol G,1)) & k + 1 in Seg (width G) )
proof end;

theorem :: GOBOARD1:34  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for m, k, n being Nat
for G being Go-board st width G = m + 1 & m > 0 & k in Seg m & n in dom G holds
( (DelCol G,1) * n,k = G * n,(k + 1) & 1 in Seg (width G) )
proof end;

theorem :: GOBOARD1:35  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for m, k being Nat
for G being Go-board st width G = m + 1 & m > 0 & k in Seg m holds
( Col (DelCol G,(width G)),k = Col G,k & k in Seg (width (DelCol G,(width G))) )
proof end;

theorem Th36: :: GOBOARD1:36  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for m, k, n being Nat
for G being Go-board st width G = m + 1 & m > 0 & k in Seg m & n in dom G holds
( k in Seg (width G) & (DelCol G,(width G)) * n,k = G * n,k & width G in Seg (width G) )
proof end;

theorem :: GOBOARD1:37  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for f being FinSequence of (TOP-REAL 2)
for i, n, m being Nat
for G being Go-board st rng f misses rng (Col G,i) & f /. n in rng (Line G,m) & n in dom f & i in Seg (width G) & m in dom G & width G > 1 holds
f /. n in rng (Line (DelCol G,i),m)
proof end;

definition
let D be set ;
let f be FinSequence of D;
let M be Matrix of D;
pred f is_sequence_on M means :Def11: :: GOBOARD1:def 11
( ( for n being Nat st n in dom f holds
ex i, j being Nat st
( [i,j] in Indices M & f /. n = M * i,j ) ) & ( for n being Nat st n in dom f & n + 1 in dom f holds
for m, k, i, j being Nat st [m,k] in Indices M & [i,j] in Indices M & f /. n = M * m,k & f /. (n + 1) = M * i,j holds
(abs (m - i)) + (abs (k - j)) = 1 ) );
end;

:: deftheorem Def11 defines is_sequence_on GOBOARD1:def 11 :
for D being set
for f being FinSequence of D
for M being Matrix of D holds
( f is_sequence_on M iff ( ( for n being Nat st n in dom f holds
ex i, j being Nat st
( [i,j] in Indices M & f /. n = M * i,j ) ) & ( for n being Nat st n in dom f & n + 1 in dom f holds
for m, k, i, j being Nat st [m,k] in Indices M & [i,j] in Indices M & f /. n = M * m,k & f /. (n + 1) = M * i,j holds
(abs (m - i)) + (abs (k - j)) = 1 ) ) );

Lm2: for D being set
for M being Matrix of D holds <*> D is_sequence_on M
proof end;

theorem :: GOBOARD1:38  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for m being Nat
for D being set
for f being FinSequence of D
for M being Matrix of D holds
( ( m in dom f implies 1 <= len (f | m) ) & ( f is_sequence_on M implies f | m is_sequence_on M ) )
proof end;

theorem :: GOBOARD1:39  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for f1, f2 being FinSequence of (TOP-REAL 2)
for D being set
for M being Matrix of D st ( for n being Nat st n in dom f1 holds
ex i, j being Nat st
( [i,j] in Indices M & f1 /. n = M * i,j ) ) & ( for n being Nat st n in dom f2 holds
ex i, j being Nat st
( [i,j] in Indices M & f2 /. n = M * i,j ) ) holds
for n being Nat st n in dom (f1 ^ f2) holds
ex i, j being Nat st
( [i,j] in Indices M & (f1 ^ f2) /. n = M * i,j )
proof end;

theorem :: GOBOARD1:40  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for f1, f2 being FinSequence of (TOP-REAL 2)
for D being set
for M being Matrix of D st ( for n being Nat st n in dom f1 & n + 1 in dom f1 holds
for m, k, i, j being Nat st [m,k] in Indices M & [i,j] in Indices M & f1 /. n = M * m,k & f1 /. (n + 1) = M * i,j holds
(abs (m - i)) + (abs (k - j)) = 1 ) & ( for n being Nat st n in dom f2 & n + 1 in dom f2 holds
for m, k, i, j being Nat st [m,k] in Indices M & [i,j] in Indices M & f2 /. n = M * m,k & f2 /. (n + 1) = M * i,j holds
(abs (m - i)) + (abs (k - j)) = 1 ) & ( for m, k, i, j being Nat st [m,k] in Indices M & [i,j] in Indices M & f1 /. (len f1) = M * m,k & f2 /. 1 = M * i,j & len f1 in dom f1 & 1 in dom f2 holds
(abs (m - i)) + (abs (k - j)) = 1 ) holds
for n being Nat st n in dom (f1 ^ f2) & n + 1 in dom (f1 ^ f2) holds
for m, k, i, j being Nat st [m,k] in Indices M & [i,j] in Indices M & (f1 ^ f2) /. n = M * m,k & (f1 ^ f2) /. (n + 1) = M * i,j holds
(abs (m - i)) + (abs (k - j)) = 1
proof end;

theorem :: GOBOARD1:41  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for i being Nat
for G being Go-board
for f being FinSequence of (TOP-REAL 2) st f is_sequence_on G & i in Seg (width G) & rng f misses rng (Col G,i) & width G > 1 holds
f is_sequence_on DelCol G,i
proof end;

theorem Th42: :: GOBOARD1:42  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for i being Nat
for G being Go-board
for f being FinSequence of (TOP-REAL 2) st f is_sequence_on G & i in dom f holds
ex n being Nat st
( n in dom G & f /. i in rng (Line G,n) )
proof end;

theorem Th43: :: GOBOARD1:43  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for i, n being Nat
for G being Go-board
for f being FinSequence of (TOP-REAL 2) st f is_sequence_on G & i in dom f & i + 1 in dom f & n in dom G & f /. i in rng (Line G,n) & not f /. (i + 1) in rng (Line G,n) holds
for k being Nat st f /. (i + 1) in rng (Line G,k) & k in dom G holds
abs (n - k) = 1
proof end;

theorem Th44: :: GOBOARD1:44  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for i, m being Nat
for G being Go-board
for f being FinSequence of (TOP-REAL 2) st 1 <= len f & f /. (len f) in rng (Line G,(len G)) & f is_sequence_on G & i in dom G & i + 1 in dom G & m in dom f & f /. m in rng (Line G,i) & ( for k being Nat st k in dom f & f /. k in rng (Line G,i) holds
k <= m ) holds
( m + 1 in dom f & f /. (m + 1) in rng (Line G,(i + 1)) )
proof end;

theorem :: GOBOARD1:45  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for G being Go-board
for f being FinSequence of (TOP-REAL 2) st 1 <= len f & f /. 1 in rng (Line G,1) & f /. (len f) in rng (Line G,(len G)) & f is_sequence_on G holds
( ( for i being Nat st 1 <= i & i <= len G holds
ex k being Nat st
( k in dom f & f /. k in rng (Line G,i) ) ) & ( for i being Nat st 1 <= i & i <= len G & 2 <= len f holds
L~ f meets rng (Line G,i) ) & ( for i, j, k, m being Nat st 1 <= i & i <= len G & 1 <= j & j <= len G & k in dom f & m in dom f & f /. k in rng (Line G,i) & ( for n being Nat st n in dom f & f /. n in rng (Line G,i) holds
n <= k ) & k < m & f /. m in rng (Line G,j) holds
i < j ) )
proof end;

theorem Th46: :: GOBOARD1:46  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for i being Nat
for G being Go-board
for f being FinSequence of (TOP-REAL 2) st f is_sequence_on G & i in dom f holds
ex n being Nat st
( n in Seg (width G) & f /. i in rng (Col G,n) )
proof end;

theorem Th47: :: GOBOARD1:47  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for i, n being Nat
for G being Go-board
for f being FinSequence of (TOP-REAL 2) st f is_sequence_on G & i in dom f & i + 1 in dom f & n in Seg (width G) & f /. i in rng (Col G,n) & not f /. (i + 1) in rng (Col G,n) holds
for k being Nat st f /. (i + 1) in rng (Col G,k) & k in Seg (width G) holds
abs (n - k) = 1
proof end;

theorem Th48: :: GOBOARD1:48  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for i, m being Nat
for G being Go-board
for f being FinSequence of (TOP-REAL 2) st 1 <= len f & f /. (len f) in rng (Col G,(width G)) & f is_sequence_on G & i in Seg (width G) & i + 1 in Seg (width G) & m in dom f & f /. m in rng (Col G,i) & ( for k being Nat st k in dom f & f /. k in rng (Col G,i) holds
k <= m ) holds
( m + 1 in dom f & f /. (m + 1) in rng (Col G,(i + 1)) )
proof end;

theorem Th49: :: GOBOARD1:49  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for G being Go-board
for f being FinSequence of (TOP-REAL 2) st 1 <= len f & f /. 1 in rng (Col G,1) & f /. (len f) in rng (Col G,(width G)) & f is_sequence_on G holds
( ( for i being Nat st 1 <= i & i <= width G holds
ex k being Nat st
( k in dom f & f /. k in rng (Col G,i) ) ) & ( for i being Nat st 1 <= i & i <= width G & 2 <= len f holds
L~ f meets rng (Col G,i) ) & ( for i, j, k, m being Nat st 1 <= i & i <= width G & 1 <= j & j <= width G & k in dom f & m in dom f & f /. k in rng (Col G,i) & ( for n being Nat st n in dom f & f /. n in rng (Col G,i) holds
n <= k ) & k < m & f /. m in rng (Col G,j) holds
i < j ) )
proof end;

theorem Th50: :: GOBOARD1:50  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for n, k being Nat
for G being Go-board
for f being FinSequence of (TOP-REAL 2) st n in dom f & f /. n in rng (Col G,k) & k in Seg (width G) & f /. 1 in rng (Col G,1) & f is_sequence_on G & ( for i being Nat st i in dom f & f /. i in rng (Col G,k) holds
n <= i ) holds
for i being Nat st i in dom f & i <= n holds
for m being Nat st m in Seg (width G) & f /. i in rng (Col G,m) holds
m <= k
proof end;

theorem :: GOBOARD1:51  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for G being Go-board
for f being FinSequence of (TOP-REAL 2) st f is_sequence_on G & f /. 1 in rng (Col G,1) & f /. (len f) in rng (Col G,(width G)) & width G > 1 & 1 <= len f holds
ex g being FinSequence of (TOP-REAL 2) st
( g /. 1 in rng (Col (DelCol G,(width G)),1) & g /. (len g) in rng (Col (DelCol G,(width G)),(width (DelCol G,(width G)))) & 1 <= len g & g is_sequence_on DelCol G,(width G) & rng g c= rng f )
proof end;

theorem :: GOBOARD1:52  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for G being Go-board
for f being FinSequence of (TOP-REAL 2) st f is_sequence_on G & (rng f) /\ (rng (Col G,1)) <> {} & (rng f) /\ (rng (Col G,(width G))) <> {} holds
ex g being FinSequence of (TOP-REAL 2) st
( rng g c= rng f & g /. 1 in rng (Col G,1) & g /. (len g) in rng (Col G,(width G)) & 1 <= len g & g is_sequence_on G )
proof end;

theorem :: GOBOARD1:53  Show TPTP formulae Show IDV graph:: Showing IDV graph ... (Click the Palm Tree again to close it) Show TPTP problem
for k, n being Nat
for G being Go-board
for f being FinSequence of (TOP-REAL 2) st k in dom G & f is_sequence_on G & f /. (len f) in rng (Line G,(len G)) & n in dom f & f /. n in rng (Line G,k) holds
( ( for i being Nat st k <= i & i <= len G holds
ex j being Nat st
( j in dom f & n <= j & f /. j in rng (Line G,i) ) ) & ( for i being Nat st k < i & i <= len G holds
ex j being Nat st
( j in dom f & n < j & f /. j in rng (Line G,i) ) ) )
proof end;