The Mizar article:

Cages - the External Approximation of Jordan's Curve

by
Czeslaw Bylinski, and
Mariusz Zynel

Received June 22, 1999

Copyright (c) 1999 Association of Mizar Users

MML identifier: JORDAN9
[ MML identifier index ]


environ

 vocabulary FINSEQ_1, BOOLE, PRE_TOPC, RELAT_2, CONNSP_1, RELAT_1, FINSEQ_5,
      ARYTM_1, GOBOARD1, EUCLID, MATRIX_1, ABSVALUE, GOBRD13, FUNCT_1,
      TOPREAL1, RFINSEQ, GOBOARD5, TOPS_1, TREES_1, SPPOL_1, MCART_1, TARSKI,
      SUBSET_1, SEQM_3, GOBOARD9, COMPTS_1, JORDAN8, PSCOMP_1, GROUP_1,
      ARYTM_3, SPRECT_2, CARD_1, PCOMPS_1, METRIC_1, JORDAN9, FINSEQ_4, ARYTM;
 notation TARSKI, XBOOLE_0, SUBSET_1, GOBOARD5, STRUCT_0, ORDINAL1, NUMBERS,
      XREAL_0, REAL_1, NAT_1, BINARITH, ABSVALUE, RELAT_1, FUNCT_1, FUNCT_2,
      CARD_1, CARD_4, FINSEQ_1, FINSEQ_2, FINSEQ_4, FINSEQ_5, RFINSEQ,
      MATRIX_1, METRIC_1, PRE_TOPC, TOPS_1, COMPTS_1, CONNSP_1, PCOMPS_1,
      EUCLID, TOPREAL1, GOBOARD1, SPPOL_1, PSCOMP_1, SPRECT_2, GOBOARD9,
      JORDAN8, GOBRD13;
 constructors REAL_1, FINSEQ_4, CARD_4, RFINSEQ, BINARITH, TOPS_1, CONNSP_1,
      COMPTS_1, SPPOL_1, PSCOMP_1, REALSET1, SPRECT_2, GOBOARD9, JORDAN8,
      GOBRD13, GROUP_1, MEMBERED;
 clusters PSCOMP_1, RELSET_1, FINSEQ_1, SPPOL_2, REVROT_1, SPRECT_1, SPRECT_2,
      XREAL_0, GOBOARD9, JORDAN8, GOBRD13, EUCLID, TOPREAL1, NAT_1, MEMBERED,
      ZFMISC_1, NUMBERS, ORDINAL2;
 requirements REAL, NUMERALS, SUBSET, BOOLE, ARITHM;
 definitions TARSKI, GOBOARD1, GOBOARD5, GOBRD13, XBOOLE_0;
 theorems NAT_1, FINSEQ_1, GOBOARD1, FINSEQ_4, EUCLID, FINSEQ_3, AXIOMS,
      REAL_1, REAL_2, HEINE, TOPREAL4, SPPOL_2, TARSKI, JORDAN3, SQUARE_1,
      PSCOMP_1, FINSEQ_5, FINSEQ_6, CQC_THE1, GOBOARD7, TOPREAL1, BINARITH,
      AMI_5, JORDAN5B, RLVECT_1, GOBOARD5, SPRECT_2, SPPOL_1, ABSVALUE,
      FUNCT_1, FUNCT_2, GROUP_5, GOBOARD9, RELAT_1, FINSEQ_2, UNIFORM1,
      SUBSET_1, GOBRD11, JORDAN4, GOBOARD2, SPRECT_3, CARD_1, RFINSEQ,
      GOBOARD6, TOPREAL3, TOPMETR, TOPS_1, JORDAN8, GOBRD13, SPRECT_4,
      CONNSP_1, PARTFUN2, RELSET_1, SCMFSA_7, SPRECT_1, XBOOLE_0, XBOOLE_1,
      XREAL_0, XCMPLX_0, XCMPLX_1;
 schemes NAT_1, LATTICE5;

begin :: Generalities

reserve i,j,k,n for Nat,
  D for non empty set,
  f, g for FinSequence of D;

Lm1: for n st 1 <= n holds n-'1+2 = n+1
proof
  let n;
  assume 1 <= n;
  hence n-'1+2 = n+2-'1 by JORDAN4:3
            .= n+(1+1) - 1 by JORDAN4:2
            .= n+1+1 - 1 by XCMPLX_1:1
            .= n+1 by XCMPLX_1:26;
end;

canceled 2;

theorem Th3:
  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
  let T be non empty TopSpace;
  let B,C1,C2,D be Subset of T;
  assume that
A1:  B is connected and
A2:  C1 is_a_component_of D and
A3:  C2 is_a_component_of D and
A4:  B meets C1 and
A5:  B meets C2 and
A6:  B c= D;
A7:  B c= C1 by A1,A2,A4,A6,GOBOARD9:6;
A8:  B c= C2 by A1,A3,A5,A6,GOBOARD9:6;
   B <> {} by A4,XBOOLE_1:65;
 then C1 meets C2 by A7,A8,XBOOLE_1:68;
 hence C1 = C2 by A2,A3,GOBOARD9:3;
end;

theorem Th4:
  (for n holds f|n = g|n) implies f = g
proof
  assume
A1:  for n holds f|n = g|n;
A2: now assume
A3:  len f <> len g;
   per cases by A3,REAL_1:def 5;
    suppose
A4:    len f < len g;
      then f|len g = f & g|len g = g by TOPREAL1:2;
     hence contradiction by A1,A4;
    suppose
A5:    len g < len f;
      then f|len f = f & g|len f = g by TOPREAL1:2;
     hence contradiction by A1,A5;
  end;
   f|len f = f & g|len g = g by TOPREAL1:2;
 hence thesis by A1,A2;
end;

theorem Th5:
  n in dom f implies
   ex k st k in dom Rev f & n+k = len f+1 & f/.n = (Rev f)/.k
proof
  assume
A1: n in dom f;
  then A2: 1 <= n & n <= len f by FINSEQ_3:27;
    take k = len f+1-'n;
      n <= len f+1 by A2,SPPOL_1:5;
  then A3: k+n = len f+1 by AMI_5:4;
  then A4: k+n-'1 = len f by BINARITH:39;
      n+1 <= len f+1 by A2,AXIOMS:24;
  then A5: 1 <= k by SPRECT_3:8;
      k+1 <= k+n by A2,AXIOMS:24;
    then k+1-'1 <= k+n -'1 by JORDAN3:5;
    then k <= len f by A4,BINARITH:39;
    then k in dom f by A5,FINSEQ_3:27;
    hence thesis by A1,A3,FINSEQ_5:60,69;
end;

theorem Th6:
  n in dom Rev f implies
   ex k st k in dom f & n+k = len f+1 & (Rev f)/.n = f/.k
proof
  assume
    n in dom Rev f;
  then n in dom f by FINSEQ_5:60;
  then consider k such that
A1: k in dom Rev f and
A2: n+k = len f+1 and
      f/.n = (Rev f)/.k by Th5;
A3: len f = len Rev f & dom f = dom Rev f by FINSEQ_5:60,def 3;
  then (Rev f)/.n = f/.k by A1,A2,FINSEQ_5:69;
  hence thesis by A1,A2,A3;
end;

begin :: Go-board preliminaries

reserve G for Go-board,
  f, g for FinSequence of TOP-REAL 2,
  p for Point of TOP-REAL 2,
  r, s for Real,
  x for set;

theorem Th7:
  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
  let D be non empty set;
  let G be Matrix of D;
  let f be FinSequence of D;
  hereby
    assume
  A1: f is_sequence_on G;
  A2:for n st n in dom Rev f ex i,j st [i,j] in
 Indices G & (Rev f)/.n = G*(i,j)
  proof
    let n;
    assume n in dom Rev f;
    then consider k such that
  A3: k in dom f and
        n+k = len f+1 and
  A4: (Rev f)/.n = f/.k by Th6;
    consider i,j such that
  A5: [i,j] in Indices G and
  A6: f/.k = G*(i,j) by A1,A3,GOBOARD1:def 11;
    take i,j;
    thus thesis by A4,A5,A6;
  end;
    for n st n in dom Rev f & n+1 in dom Rev f holds
    for m,k,i,j being Nat st [m,k] in Indices G & [i,j] in Indices G &
                     (Rev f)/.n = G*(m,k) & (Rev f)/.(n+1) = G*(i,j) holds
      abs(m-i)+abs(k-j) = 1
  proof
    let n such that
  A7: n in dom Rev f and
  A8: n+1 in dom Rev f;
    let m,k,i,j be Nat such that
  A9: [m,k] in Indices G and
  A10: [i,j] in Indices G and
  A11: (Rev f)/.n = G*(m,k) and
  A12: (Rev f)/.(n+1) = G*(i,j);
    consider l being Nat such that
  A13: l in dom f and
  A14: n+l = len f+1 and
  A15: (Rev f)/.n = f/.l by A7,Th6;
    consider l' being Nat such that
  A16: l' in dom f and
  A17: n+1+l' = len f+1 and
  A18: (Rev f)/.(n+1) = f/.l' by A8,Th6;
      n+(1+l') = n+l by A14,A17,XCMPLX_1:1;
then A19: l'+1 = l by XCMPLX_1:2;
      abs(i-m) = abs(m-i) & abs(j-k) = abs(k-j) by UNIFORM1:13;
    hence abs(m-i)+abs(k-j) = 1
            by A1,A9,A10,A11,A12,A13,A15,A16,A18,A19,GOBOARD1:def 11;
  end;
    hence Rev f is_sequence_on G by A2,GOBOARD1:def 11;
  end;
  assume
A20: Rev f is_sequence_on G;
  A21:for n st n in dom f ex i,j st [i,j] in Indices G & f/.n = G*(i,j)
  proof
    let n;
    assume n in dom f;
    then consider k such that
  A22: k in dom Rev f and
        n+k = len f+1 and
  A23: f/.n = (Rev f)/.k by Th5;
    consider i,j such that
  A24: [i,j] in Indices G and
  A25: (Rev f)/.k = G*(i,j) by A20,A22,GOBOARD1:def 11;
    take i,j;
    thus thesis by A23,A24,A25;
  end;

    for n st n in dom f & n+1 in dom f holds
    for m,k,i,j being Nat st [m,k] in Indices G & [i,j] in Indices G &
                     f/.n = G*(m,k) & f/.(n+1) = G*(i,j) holds
      abs(m-i)+abs(k-j) = 1
  proof
    let n such that
  A26: n in dom f and
  A27: n+1 in dom f;
    let m,k,i,j be Nat such that
  A28: [m,k] in Indices G and
  A29: [i,j] in Indices G and
  A30: f/.n = G*(m,k) and
  A31: f/.(n+1) = G*(i,j);
    consider l being Nat such that
  A32: l in dom Rev f and
  A33: n+l = len f+1 and
  A34: f/.n = (Rev f)/.l by A26,Th5;

    consider l' being Nat such that
  A35: l' in dom Rev f and
  A36: n+1+l' = len f+1 and
  A37: f/.(n+1) = (Rev f)/.l' by A27,Th5;
      n+(1+l') = n+l by A33,A36,XCMPLX_1:1;
then A38: l'+1 = l by XCMPLX_1:2;
      abs(i-m) = abs(m-i) & abs(j-k) = abs(k-j) by UNIFORM1:13;
    hence abs(m-i)+abs(k-j) = 1
        by A20,A28,A29,A30,A31,A32,A34,A35,A37,A38,GOBOARD1:def 11;
  end;
  hence f is_sequence_on G by A21,GOBOARD1:def 11;
end;

theorem Th8:
  for G being Matrix of TOP-REAL 2
  for f being FinSequence of TOP-REAL 2 holds
  f is_sequence_on G & 1 <= k & k <= len f implies f/.k in Values G
proof
  let G be Matrix of TOP-REAL 2;
  let f be FinSequence of TOP-REAL 2;
  assume that
A1: f is_sequence_on G and
A2: 1 <= k and
A3: k <= len f;
A4: k in dom f by A2,A3,FINSEQ_3:27;
  then f/.k = f.k by FINSEQ_4:def 4;
then A5: f/.k in rng f by A4,FUNCT_1:def 5;
    rng f c= Values G by A1,GOBRD13:9;
  hence f/.k in Values G by A5;
end;

Lm2:
  f is_sequence_on G & 1 <= k & k <= len f implies
   ex i,j being Nat st [i,j] in Indices G & f/.k = G*(i,j)
proof
  assume that
A1: f is_sequence_on G and
A2: 1 <= k & k <= len f;
    k in dom f by A2,FINSEQ_3:27;
  then consider i,j such that
A3: [i,j] in Indices G & f/.k = G*(i,j) by A1,GOBOARD1:def 11;
   take i,j;
   thus thesis by A3;
end;

theorem Th9:
  n <= len f & x in L~(f/^n)
    implies
   ex i being Nat st n+1 <= i & i+1 <= len f & x in LSeg(f,i)
proof
  assume that
A1: n <= len f and
A2: x in L~(f/^n);
  consider j such that
A3: 1 <= j and
A4: j+1 <= len(f/^n) and
A5: x in LSeg(f/^n,j) by A2,SPPOL_2:13;
A6:  j+1 <= len f - n by A1,A4,RFINSEQ:def 2;
  take n+j;
    j+1 <= len f - n by A1,A4,RFINSEQ:def 2;
  then n+(j+1) <= len f by REAL_1:84;
  hence thesis by A3,A5,A6,AXIOMS:24,SPPOL_2:5,XCMPLX_1:1;
end;

theorem Th10:
  f is_sequence_on G & 1 <= k & k+1 <= len f
    implies
   f/.k in left_cell(f,k,G) & f/.k in right_cell(f,k,G)
proof
  assume that
A1: f is_sequence_on G and
A2: 1 <= k & k+1 <= len f;
  set p = f/.k;
    LSeg(f,k) = LSeg(f/.k, f/.(k+1)) by A2,TOPREAL1:def 5;
  then p in LSeg(f,k) by TOPREAL1:6;
  then p in left_cell(f,k,G) /\ right_cell(f,k,G) by A1,A2,GOBRD13:30;
  hence thesis by XBOOLE_0:def 3;
end;

theorem Th11:
  f is_sequence_on G & 1 <= k & k+1 <= len f
    implies
   Int left_cell(f,k,G) <> {} & Int right_cell(f,k,G) <> {}
proof
  assume that
A1: f is_sequence_on G and
A2: 1 <= k and
A3: k+1 <= len f;

  consider i1,j1,i2,j2 being Nat such that
A4: [i1,j1] in Indices G and
A5: f/.k = G*(i1,j1) and
A6: [i2,j2] in Indices G and
A7: f/.(k+1) = G*(i2,j2) and
A8: i1 = i2 & j1+1 = j2 or i1+1 = i2 & j1 = j2 or
      i1 = i2+1 & j1 = j2 or i1 = i2 & j1 = j2+1 by A1,A2,A3,JORDAN8:6;

A9: i1+1 > i1 & j1+1 > j1 & i2+1 > i2 & j2+1 > j2 by NAT_1:38;
A10: i1 <= len G & j1 <= width G by A4,GOBOARD5:1;
A11: i2 <= len G & j2 <= width G by A6,GOBOARD5:1;
A12: i1-'1 <= len G & j1-'1 <= width G by A10,JORDAN3:7;
A13: i2-'1 <= len G & j2-'1 <= width G by A11,JORDAN3:7;

  per cases by A8;
  suppose
  A14: i1 = i2 & j1+1 = j2;
  then A15: right_cell(f,k,G) = cell(G,i1,j1)
       by A1,A2,A3,A4,A5,A6,A7,A9,GOBRD13:def 2;
      left_cell(f,k,G) = cell(G,i1-'1,j1)
       by A1,A2,A3,A4,A5,A6,A7,A9,A14,GOBRD13:def 3;
    hence thesis by A10,A12,A15,GOBOARD9:17;

  suppose
  A16: i1+1 = i2 & j1 = j2;
  then A17: right_cell(f,k,G) = cell(G,i1,j1-'1)
        by A1,A2,A3,A4,A5,A6,A7,A9,GOBRD13:def 2;
      left_cell(f,k,G) = cell(G,i1,j1)
        by A1,A2,A3,A4,A5,A6,A7,A9,A16,GOBRD13:def 3;
    hence thesis by A10,A12,A17,GOBOARD9:17;

  suppose
  A18: i1 = i2+1 & j1 = j2;
  then A19: right_cell(f,k,G) = cell(G,i2,j2)
        by A1,A2,A3,A4,A5,A6,A7,A9,GOBRD13:def 2;
      left_cell(f,k,G) = cell(G,i2,j2-'1)
        by A1,A2,A3,A4,A5,A6,A7,A9,A18,GOBRD13:def 3;
    hence thesis by A11,A13,A19,GOBOARD9:17;

  suppose
  A20: i1 = i2 & j1 = j2+1;
  then A21: right_cell(f,k,G) = cell(G,i1-'1,j2)
        by A1,A2,A3,A4,A5,A6,A7,A9,GOBRD13:def 2;
      left_cell(f,k,G) = cell(G,i1,j2)
        by A1,A2,A3,A4,A5,A6,A7,A9,A20,GOBRD13:def 3;
    hence thesis by A10,A11,A12,A21,GOBOARD9:17;
end;

theorem Th12:
  f is_sequence_on G & 1 <= k & k+1 <= len f
    implies
   Int left_cell(f,k,G) is connected & Int right_cell(f,k,G) is connected
proof
  assume that
A1: f is_sequence_on G and
A2: 1 <= k and
A3: k+1 <= len f;

  consider i1,j1,i2,j2 being Nat such that
A4: [i1,j1] in Indices G and
A5: f/.k = G*(i1,j1) and
A6: [i2,j2] in Indices G and
A7: f/.(k+1) = G*(i2,j2) and
A8: i1 = i2 & j1+1 = j2 or i1+1 = i2 & j1 = j2 or
      i1 = i2+1 & j1 = j2 or i1 = i2 & j1 = j2+1 by A1,A2,A3,JORDAN8:6;

A9: i1+1 > i1 & j1+1 > j1 & i2+1 > i2 & j2+1 > j2 by NAT_1:38;
A10: i1 <= len G & j1 <= width G by A4,GOBOARD5:1;
A11: i2 <= len G & j2 <= width G by A6,GOBOARD5:1;
A12: i1-'1 <= len G & j1-'1 <= width G by A10,JORDAN3:7;
A13: i2-'1 <= len G & j2-'1 <= width G by A11,JORDAN3:7;

  per cases by A8;
  suppose
  A14: i1 = i2 & j1+1 = j2;
  then A15: right_cell(f,k,G) = cell(G,i1,j1)
       by A1,A2,A3,A4,A5,A6,A7,A9,GOBRD13:def 2;
      left_cell(f,k,G) = cell(G,i1-'1,j1)
       by A1,A2,A3,A4,A5,A6,A7,A9,A14,GOBRD13:def 3;
    hence thesis by A10,A12,A15,GOBOARD9:21;

  suppose
  A16: i1+1 = i2 & j1 = j2;
  then A17: right_cell(f,k,G) = cell(G,i1,j1-'1)
        by A1,A2,A3,A4,A5,A6,A7,A9,GOBRD13:def 2;
      left_cell(f,k,G) = cell(G,i1,j1)
        by A1,A2,A3,A4,A5,A6,A7,A9,A16,GOBRD13:def 3;
    hence thesis by A10,A12,A17,GOBOARD9:21;

  suppose
  A18: i1 = i2+1 & j1 = j2;
  then A19: right_cell(f,k,G) = cell(G,i2,j2)
        by A1,A2,A3,A4,A5,A6,A7,A9,GOBRD13:def 2;
      left_cell(f,k,G) = cell(G,i2,j2-'1)
        by A1,A2,A3,A4,A5,A6,A7,A9,A18,GOBRD13:def 3;
    hence thesis by A11,A13,A19,GOBOARD9:21;

  suppose
  A20: i1 = i2 & j1 = j2+1;
  then A21: right_cell(f,k,G) = cell(G,i1-'1,j2)
        by A1,A2,A3,A4,A5,A6,A7,A9,GOBRD13:def 2;
      left_cell(f,k,G) = cell(G,i1,j2)
        by A1,A2,A3,A4,A5,A6,A7,A9,A20,GOBRD13:def 3;
    hence thesis by A10,A11,A12,A21,GOBOARD9:21;
end;

theorem Th13:
  f is_sequence_on G & 1 <= k & k+1 <= len f
    implies
   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
  assume that
A1: f is_sequence_on G and
A2: 1 <= k and
A3: k+1 <= len f;

  consider i1,j1,i2,j2 being Nat such that
A4: [i1,j1] in Indices G and
A5: f/.k = G*(i1,j1) and
A6: [i2,j2] in Indices G and
A7: f/.(k+1) = G*(i2,j2) and
A8: i1 = i2 & j1+1 = j2 or i1+1 = i2 & j1 = j2 or
      i1 = i2+1 & j1 = j2 or i1 = i2 & j1 = j2+1 by A1,A2,A3,JORDAN8:6;

A9: i1+1 > i1 & j1+1 > j1 & i2+1 > i2 & j2+1 > j2 by NAT_1:38;
A10: i1 <= len G & j1 <= width G by A4,GOBOARD5:1;
A11: i2 <= len G & j2 <= width G by A6,GOBOARD5:1;
A12: i1-'1 <= len G & j1-'1 <= width G by A10,JORDAN3:7;
A13: i2-'1 <= len G & j2-'1 <= width G by A11,JORDAN3:7;

  per cases by A8;
  suppose
  A14: i1 = i2 & j1+1 = j2;
  then A15: right_cell(f,k,G) = cell(G,i1,j1)
       by A1,A2,A3,A4,A5,A6,A7,A9,GOBRD13:def 2;
      left_cell(f,k,G) = cell(G,i1-'1,j1)
       by A1,A2,A3,A4,A5,A6,A7,A9,A14,GOBRD13:def 3;
    hence thesis by A10,A12,A15,GOBRD11:35;

  suppose
  A16: i1+1 = i2 & j1 = j2;
  then A17: right_cell(f,k,G) = cell(G,i1,j1-'1)
        by A1,A2,A3,A4,A5,A6,A7,A9,GOBRD13:def 2;
      left_cell(f,k,G) = cell(G,i1,j1)
        by A1,A2,A3,A4,A5,A6,A7,A9,A16,GOBRD13:def 3;
    hence thesis by A10,A12,A17,GOBRD11:35;

  suppose
  A18: i1 = i2+1 & j1 = j2;
  then A19: right_cell(f,k,G) = cell(G,i2,j2)
        by A1,A2,A3,A4,A5,A6,A7,A9,GOBRD13:def 2;
      left_cell(f,k,G) = cell(G,i2,j2-'1)
        by A1,A2,A3,A4,A5,A6,A7,A9,A18,GOBRD13:def 3;
    hence thesis by A11,A13,A19,GOBRD11:35;

  suppose
  A20: i1 = i2 & j1 = j2+1;
  then A21: right_cell(f,k,G) = cell(G,i1-'1,j2)
        by A1,A2,A3,A4,A5,A6,A7,A9,GOBRD13:def 2;
      left_cell(f,k,G) = cell(G,i1,j2)
        by A1,A2,A3,A4,A5,A6,A7,A9,A20,GOBRD13:def 3;
    hence thesis by A10,A11,A12,A21,GOBRD11:35;
end;

theorem Th14:
  f is_sequence_on G & LSeg(f,k) is horizontal
    implies
   ex j st 1 <= j & j <= width G & for p st p in
 LSeg(f,k) holds p`2 = G*(1,j)`2
proof
  assume that
A1: f is_sequence_on G and
A2: LSeg(f,k) is horizontal;
 per cases;
 suppose
A3: 1 <= k & k+1 <= len f;
    k <= k+1 by NAT_1:29;
  then k <= len f by A3,AXIOMS:22;
  then consider i,j such that
A4: [i,j] in Indices G and
A5: f/.k = G*(i,j) by A1,A3,Lm2;
 take j;
 thus
A6: 1 <= j & j <= width G by A4,GOBOARD5:1;
 let p;
A7: 1 <= i & i <= len G by A4,GOBOARD5:1;
A8: f/.k in LSeg(f,k) by A3,TOPREAL1:27;
 assume p in LSeg(f,k);
 hence p`2 = (f/.k)`2 by A2,A8,SPPOL_1:def 2
          .= G*(1,j)`2 by A5,A6,A7,GOBOARD5:2;
 suppose
A9: not(1 <= k & k+1 <= len f);
 take 1;
    width G <> 0 by GOBOARD1:def 5;
 hence 1 <= 1 & 1 <= width G by RLVECT_1:99;
 thus thesis by A9,TOPREAL1:def 5;
end;

theorem Th15:
  f is_sequence_on G & LSeg(f,k) is vertical
    implies
   ex i st 1 <= i & i <= len G & for p st p in LSeg(f,k) holds p`1 = G*(i,1)`1
proof
  assume that
A1: f is_sequence_on G and
A2: LSeg(f,k) is vertical;
 per cases;
 suppose
A3: 1 <= k & k+1 <= len f;
    k <= k+1 by NAT_1:29;
  then k <= len f by A3,AXIOMS:22;
  then consider i,j such that
A4: [i,j] in Indices G and
A5: f/.k = G*(i,j) by A1,A3,Lm2;
 take i;
 thus
A6: 1 <= i & i <= len G by A4,GOBOARD5:1;
 let p;
A7: 1 <= j & j <= width G by A4,GOBOARD5:1;
A8: f/.k in LSeg(f,k) by A3,TOPREAL1:27;
 assume p in LSeg(f,k);
 hence p`1 = (f/.k)`1 by A2,A8,SPPOL_1:def 3
          .= G*(i,1)`1 by A5,A6,A7,GOBOARD5:3;
 suppose
A9: not(1 <= k & k+1 <= len f);
 take 1;
    0 <> len G by GOBOARD1:def 5;
 hence 1 <= 1 & 1 <= len G by RLVECT_1:99;
 thus thesis by A9,TOPREAL1:def 5;
end;

theorem Th16:
  f is_sequence_on G & f is special & i <= len G & j <= width G
    implies
   Int cell(G,i,j) misses L~f
proof
  assume that
A1: f is_sequence_on G and
A2: f is special and
A3: i <= len G and
A4: j <= width G;
 assume Int cell(G,i,j) meets L~f;
 then consider x being set such that
A5: x in Int cell(G,i,j) and
A6: x in L~f by XBOOLE_0:3;
    L~f = union { LSeg(f,k) : 1 <= k & k+1 <= len f } by TOPREAL1:def 6;
  then consider X being set such that
A7: x in X and
A8: X in { LSeg(f,k) : 1 <= k & k+1 <= len f } by A6,TARSKI:def 4;
  consider k such that
A9: X = LSeg(f,k) and
      1 <= k & k+1 <= len f by A8;
  reconsider p = x as Point of TOP-REAL 2 by A7,A9;
A10: Int cell(G,i,j)
           = Int(v_strip(G,i) /\ h_strip(G,j)) by GOBOARD5:def 3
          .= Int v_strip(G,i) /\ Int h_strip(G,j) by TOPS_1:46;
 per cases by A2,SPPOL_1:41;
 suppose LSeg(f,k) is horizontal;
  then consider j0 being Nat such that
A11: 1 <= j0 & j0 <= width G and
A12: for p being Point of TOP-REAL 2
       st p in LSeg(f,k) holds p`2 = G*(1,j0)`2 by A1,Th14;
     now assume
A13: p in Int h_strip(G,j);
A14: j0 > j implies j0 >= j+1 by NAT_1:38;
    per cases by A14,REAL_1:def 5;
    suppose
A15:    j0 < j;
      then j >= 1 by A11,AXIOMS:22;
      then A16:    p`2 > G*(1,j)`2 by A4,A13,GOBOARD6:30;
        0 <> len G by GOBOARD1:def 5;
      then 1 <= len G by RLVECT_1:99;
      then G*(1,j)`2 > G*(1,j0)`2 by A4,A11,A15,GOBOARD5:5;
     hence contradiction by A7,A9,A12,A16;
    suppose j0 = j;
      then p`2 > G*(1,j0)`2 by A11,A13,GOBOARD6:30;
     hence contradiction by A7,A9,A12;
    suppose
A17:    j0 > j+1;
      then j+1 <= width G by A11,AXIOMS:22;
      then j < width G by NAT_1:38;
      then A18:    p`2 < G*(1,j+1)`2 by A13,GOBOARD6:31;
        0 <> len G by GOBOARD1:def 5;
      then 1 <= len G & j+1 >= 1 by RLVECT_1:99;
      then G*(1,j+1)`2 < G*(1,j0)`2 by A11,A17,GOBOARD5:5;
     hence contradiction by A7,A9,A12,A18;
    suppose
A19:   j0 = j+1;
      then j < width G by A11,NAT_1:38;
      then p`2 < G*(1,j0)`2 by A13,A19,GOBOARD6:31;
     hence contradiction by A7,A9,A12;
   end;
 hence contradiction by A5,A10,XBOOLE_0:def 3;
 suppose LSeg(f,k) is vertical;
  then consider i0 being Nat such that
A20: 1 <= i0 & i0 <= len G and
A21: for p being Point of TOP-REAL 2
  st p in LSeg(f,k) holds p`1 = G*(i0,1)`1 by A1,Th15;
     now assume
A22: p in Int v_strip(G,i);
A23: i0 > i implies i0 >= i+1 by NAT_1:38;
    per cases by A23,REAL_1:def 5;
    suppose
A24:    i0 < i;
      then i >= 1 by A20,AXIOMS:22;
      then A25:    p`1 > G*(i,1)`1 by A3,A22,GOBOARD6:32;
        0 <> width G by GOBOARD1:def 5;
      then 1 <= width G by RLVECT_1:99;
      then G*(i,1)`1 > G*(i0,1)`1 by A3,A20,A24,GOBOARD5:4;
     hence contradiction by A7,A9,A21,A25;
    suppose i0 = i;
      then p`1 > G*(i0,1)`1 by A20,A22,GOBOARD6:32;
     hence contradiction by A7,A9,A21;
    suppose
A26:    i0 > i+1;
      then i+1 <= len G by A20,AXIOMS:22;
      then i < len G by NAT_1:38;
      then A27:    p`1 < G*(i+1,1)`1 by A22,GOBOARD6:33;
        0 <> width G by GOBOARD1:def 5;
      then 1 <= width G & i+1 >= 1 by RLVECT_1:99;
      then G*(i+1,1)`1 < G*(i0,1)`1 by A20,A26,GOBOARD5:4;
     hence contradiction by A7,A9,A21,A27;
    suppose
A28:   i0 = i+1;
      then i < len G by A20,NAT_1:38;
      then p`1 < G*(i0,1)`1 by A22,A28,GOBOARD6:33;
     hence contradiction by A7,A9,A21;
   end;
 hence contradiction by A5,A10,XBOOLE_0:def 3;
end;

theorem Th17:
  f is_sequence_on G & f is special & 1 <= k & k+1 <= len f
    implies
   Int left_cell(f,k,G) misses L~f & Int right_cell(f,k,G) misses L~f
proof
  assume that
A1: f is_sequence_on G & f is special and
A2: 1 <= k and
A3: k+1 <= len f;

  consider i1,j1,i2,j2 being Nat such that
A4: [i1,j1] in Indices G and
A5: f/.k = G*(i1,j1) and
A6: [i2,j2] in Indices G and
A7: f/.(k+1) = G*(i2,j2) and
A8: i1 = i2 & j1+1 = j2 or i1+1 = i2 & j1 = j2 or
      i1 = i2+1 & j1 = j2 or i1 = i2 & j1 = j2+1 by A1,A2,A3,JORDAN8:6;

A9: i1+1 > i1 & j1+1 > j1 & i2+1 > i2 & j2+1 > j2 by NAT_1:38;
A10: i1 <= len G & j1 <= width G by A4,GOBOARD5:1;
A11: i2 <= len G & j2 <= width G by A6,GOBOARD5:1;
A12: i1-'1 <= len G & j1-'1 <= width G by A10,JORDAN3:7;
A13: i2-'1 <= len G & j2-'1 <= width G by A11,JORDAN3:7;

  per cases by A8;
  suppose
  A14: i1 = i2 & j1+1 = j2;
  then A15: right_cell(f,k,G) = cell(G,i1,j1)
       by A1,A2,A3,A4,A5,A6,A7,A9,GOBRD13:def 2;
      left_cell(f,k,G) = cell(G,i1-'1,j1)
       by A1,A2,A3,A4,A5,A6,A7,A9,A14,GOBRD13:def 3;
    hence thesis by A1,A10,A12,A15,Th16;

  suppose
  A16: i1+1 = i2 & j1 = j2;
  then A17: right_cell(f,k,G) = cell(G,i1,j1-'1)
        by A1,A2,A3,A4,A5,A6,A7,A9,GOBRD13:def 2;
      left_cell(f,k,G) = cell(G,i1,j1)
        by A1,A2,A3,A4,A5,A6,A7,A9,A16,GOBRD13:def 3;
    hence thesis by A1,A10,A12,A17,Th16;

  suppose
  A18: i1 = i2+1 & j1 = j2;
  then A19: right_cell(f,k,G) = cell(G,i2,j2)
        by A1,A2,A3,A4,A5,A6,A7,A9,GOBRD13:def 2;
      left_cell(f,k,G) = cell(G,i2,j2-'1)
        by A1,A2,A3,A4,A5,A6,A7,A9,A18,GOBRD13:def 3;
    hence thesis by A1,A11,A13,A19,Th16;

  suppose
  A20: i1 = i2 & j1 = j2+1;
  then A21: right_cell(f,k,G) = cell(G,i1-'1,j2)
        by A1,A2,A3,A4,A5,A6,A7,A9,GOBRD13:def 2;
      left_cell(f,k,G) = cell(G,i1,j2)
        by A1,A2,A3,A4,A5,A6,A7,A9,A20,GOBRD13:def 3;
    hence thesis by A1,A10,A11,A12,A21,Th16;
end;

theorem Th18:
  1 <= i & i+1 <= len G & 1 <= j & j+1 <= width G
    implies
   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
  assume that
A1: 1 <= i and
A2: i+1 <= len G and
A3: 1 <= j and
A4: j+1 <= width G;

A5: i < len G by A2,NAT_1:38;
A6: j < width G by A4,NAT_1:38;
A7: 1 <= i+1 by NAT_1:29;
A8: 1 <= j+1 by NAT_1:29;
  thus G*(i,j)`1 = G*(i,1)`1 by A1,A3,A5,A6,GOBOARD5:3
                .= G*(i,j+1)`1 by A1,A4,A5,A8,GOBOARD5:3;
  thus G*(i,j)`2 = G*(1,j)`2 by A1,A3,A5,A6,GOBOARD5:2
                .= G*(i+1,j)`2 by A2,A3,A6,A7,GOBOARD5:2;
  thus G*(i+1,j+1)`1 = G*(i+1,1)`1 by A2,A4,A7,A8,GOBOARD5:3
                    .= G*(i+1,j)`1 by A2,A3,A6,A7,GOBOARD5:3;
  thus G*(i+1,j+1)`2 = G*(1,j+1)`2 by A2,A4,A7,A8,GOBOARD5:2
                    .= G*(i,j+1)`2 by A1,A4,A5,A8,GOBOARD5:2;
end;

theorem Th19:
  for i,j being 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
  let i,j be Nat such that
A1: 1 <= i & i+1 <= len G and
A2: 1 <= j & j+1 <= width G;

A3: i < len G by A1,NAT_1:38;
A4: j < width G by A2,NAT_1:38;
then A5: v_strip(G,i) = { |[r,s]| : G*(i,j)`1 <= r & r <= G*(i+1,j)`1 } by A1,
A2,A3,GOBOARD5:9;
A6: h_strip(G,j) = { |[r,s]| : G*(i,j)`2 <= s & s <= G*(i,j+1)`2 } by A1,A2,A3,
A4,GOBOARD5:6;
  hereby
    assume p in cell(G,i,j);
    then p in v_strip(G,i) /\ h_strip(G,j) by GOBOARD5:def 3;
  then A7: p in v_strip(G,i) & p in h_strip(G,j) by XBOOLE_0:def 3;
     then ex r,s st |[r,s]| = p & G*(i,j)`1 <= r & r <= G*(i+1,j)`1 by A5;
    hence G*(i,j)`1 <= p`1 & p`1 <= G*(i+1,j)`1 by EUCLID:56;
       ex r,s st |[r,s]| = p & G*(i,j)`2 <= s & s <= G*(i,j+1)`2 by A6,A7;
    hence G*(i,j)`2 <= p`2 & p`2 <= G*(i,j+1)`2 by EUCLID:56;
  end;
  assume that
A8: G*(i,j)`1 <= p`1 & p`1 <= G*(i+1,j)`1 and
A9: G*(i,j)`2 <= p`2 & p`2 <= G*(i,j+1)`2;

A10: p = |[p`1,p`2]| by EUCLID:57;
then A11: p in v_strip(G,i) by A5,A8;
    p in h_strip(G,j) by A6,A9,A10;
  then p in v_strip(G,i) /\ h_strip(G,j) by A11,XBOOLE_0:def 3;
  hence thesis by GOBOARD5:def 3;
end;

theorem
    1 <= i & i+1 <= len G & 1 <= j & j+1 <= width G
    implies
   cell(G,i,j) = { |[r,s]| : G*(i,j)`1 <= r & r <= G*(i+1,j)`1 &
     G*(i,j)`2 <= s & s <= G*(i,j+1)`2 }
proof
  assume that
A1: 1 <= i & i+1 <= len G and
A2: 1 <= j & j+1 <= width G;

  set A = { |[r,s]| : G*(i,j)`1 <= r & r <= G*(i+1,j)`1 &
    G*(i,j)`2 <= s & s <= G*(i,j+1)`2 };

    now
    let p be set;
    assume
  A3: p in cell(G,i,j);
    then reconsider q=p as Point of TOP-REAL 2;
  A4: p = |[q`1,q`2]| by EUCLID:57;
      G*(i,j)`1 <= q`1 & q`1 <= G*(i+1,j)`1 & G*(i,j)`2 <= q`2 & q`2 <=
 G*(i,j+1)`2
      by A1,A2,A3,Th19;
    hence p in A by A4;
  end;
  hence cell(G,i,j) c= A by TARSKI:def 3;

    now
    let p be set;
    assume p in A;
    then consider r,s such that
  A5: |[r,s]| = p and
  A6: G*(i,j)`1 <= r & r <= G*(i+1,j)`1 and
  A7: G*(i,j)`2 <= s & s <= G*(i,j+1)`2;
    reconsider q=p as Point of TOP-REAL 2 by A5;
      r = q`1 & s = q`2 by A5,EUCLID:56;
    hence p in cell(G,i,j) by A1,A2,A6,A7,Th19;
  end;
  hence A c= cell(G,i,j) by TARSKI:def 3;
end;

theorem Th21:
  1 <= i & i+1 <= len G & 1 <= j & j+1 <= width G &
    p in Values G & p in cell(G,i,j)
      implies
   p = G*(i,j) or p = G*(i,j+1) or p = G*(i+1,j+1) or p = G*(i+1,j)
proof
  assume that
A1: 1 <= i and
A2: i+1 <= len G and
A3: 1 <= j and
A4: j+1 <= width G and
A5: p in Values G and
A6: p in cell(G,i,j);

A7: Values G = { G*(k,l) where k,l is Nat: [k,l] in Indices G } by GOBRD13:7;
A8: i < len G by A2,NAT_1:38;
A9: j < width G by A4,NAT_1:38;
A10: 1 <= i+1 by NAT_1:29;
A11: 1 <= j+1 by NAT_1:29;

  consider k,l being Nat such that
A12: p = G*(k,l) and
A13: [k,l] in Indices G by A5,A7;
A14: 1 <= k & k <= len G & 1 <= l & l <= width G by A13,GOBOARD5:1;
A15: now
    assume
A16:  k <> i & k <> i+1;
    per cases by A16,NAT_1:27;
    suppose
        k < i;
      then G*(k,l)`1 < G*(i,l)`1 by A8,A14,GOBOARD5:4;
      then G*(k,l)`1 < G*(i,1)`1 by A1,A8,A14,GOBOARD5:3;
      then G*(k,l)`1 < G*(i,j)`1 by A1,A3,A8,A9,GOBOARD5:3;
      hence contradiction by A1,A2,A3,A4,A6,A12,Th19;
    suppose
        i+1 < k;
      then G*(i+1,l)`1 < G*(k,l)`1 by A10,A14,GOBOARD5:4;
      then G*(i+1,1)`1 < G*(k,l)`1 by A2,A10,A14,GOBOARD5:3;
      then G*(i+1,j)`1 < G*(k,l)`1 by A2,A3,A9,A10,GOBOARD5:3;
      hence contradiction by A1,A2,A3,A4,A6,A12,Th19;
  end;
    now
    assume
A17: l <> j & l <> j+1;
    per cases by A17,NAT_1:27;
    suppose
        l < j;
      then G*(k,l)`2 < G*(k,j)`2 by A9,A14,GOBOARD5:5;
      then G*(k,l)`2 < G*(1,j)`2 by A3,A9,A14,GOBOARD5:2;
      then G*(k,l)`2 < G*(i,j)`2 by A1,A3,A8,A9,GOBOARD5:2;
      hence contradiction by A1,A2,A3,A4,A6,A12,Th19;
    suppose
        j+1 < l;
      then G*(k,j+1)`2 < G*(k,l)`2 by A11,A14,GOBOARD5:5;
      then G*(1,j+1)`2 < G*(k,l)`2 by A4,A11,A14,GOBOARD5:2;
      then G*(i,j+1)`2 < G*(k,l)`2 by A1,A4,A8,A11,GOBOARD5:2;
      hence contradiction by A1,A2,A3,A4,A6,A12,Th19;
  end;
  hence thesis by A12,A15;
end;

theorem Th22:
  1 <= i & i+1 <= len G & 1 <= j & j+1 <= width G
    implies
   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
  assume that
A1: 1 <= i and
A2: i+1 <= len G and
A3: 1 <= j and
A4: j+1 <= width G;
A5: i < i+1 & j < j+1 by NAT_1:38;
A6: i < len G by A2,NAT_1:38;
A7: j < width G by A4,NAT_1:38;

then A8: G*(i,j)`1 <= G*(i+1,j)`1 by A1,A2,A3,A5,GOBOARD5:4;
    G*(i,j)`2 <= G*(i,j+1)`2 by A1,A3,A4,A5,A6,GOBOARD5:5;
  hence G*(i,j) in cell(G,i,j) by A1,A2,A3,A4,A8,Th19;

A9: G*(i,j)`1 = G*(i,j+1)`1 by A1,A2,A3,A4,Th18;
then A10: G*(i,j+1)`1 <= G*(i+1,j)`1 by A1,A2,A3,A5,A7,GOBOARD5:4;
    G*(i,j)`2 <= G*(i,j+1)`2 by A1,A3,A4,A5,A6,GOBOARD5:5;
  hence G*(i,j+1) in cell(G,i,j) by A1,A2,A3,A4,A9,A10,Th19;

A11: G*(i+1,j+1)`1 = G*(i+1,j)`1 by A1,A2,A3,A4,Th18;
then A12: G*(i,j)`1 <= G*(i+1,j+1)`1 by A1,A2,A3,A5,A7,GOBOARD5:4;
A13: G*(i+1,j+1)`2 = G*(i,j+1)`2 by A1,A2,A3,A4,Th18;
  then G*(i,j)`2 <= G*(i+1,j+1)`2 by A1,A3,A4,A5,A6,GOBOARD5:5;
  hence G*(i+1,j+1) in cell(G,i,j) by A1,A2,A3,A4,A11,A12,A13,Th19;

A14: G*(i,j)`1 <= G*(i+1,j)`1 by A1,A2,A3,A5,A7,GOBOARD5:4;
A15: G*(i,j)`2 = G*(i+1,j)`2 by A1,A2,A3,A4,Th18;
  then G*(i+1,j)`2 <= G*(i,j+1)`2 by A1,A3,A4,A5,A6,GOBOARD5:5;
  hence G*(i+1,j) in cell(G,i,j) by A1,A2,A3,A4,A14,A15,Th19;
end;

theorem Th23:
  1 <= i & i+1 <= len G & 1 <= j & j+1 <= width G &
    p in Values G & p in cell(G,i,j)
      implies
   p is_extremal_in cell(G,i,j)
proof
  assume that
A1: 1 <= i and
A2: i+1 <= len G and
A3: 1 <= j and
A4: j+1 <= width G and
A5: p in Values G and
A6: p in cell(G,i,j);
    for a,b being Point of TOP-REAL 2 st p in
 LSeg(a,b) & LSeg(a,b) c= cell(G,i,j)
    holds p = a or p = b
  proof
    let a,b be Point of TOP-REAL 2 such that
  A7: p in LSeg(a,b) and
  A8: LSeg(a,b) c= cell(G,i,j);
    assume
  A9: a <> p & b <> p;
  A10: a in LSeg(a,b) & b in LSeg(a,b) by TOPREAL1:6;
    per cases by A1,A2,A3,A4,A5,A6,Th21;
    suppose
    A11: p = G*(i,j);
    then A12: p`1 <= a`1 & p`2 <= a`2 by A1,A2,A3,A4,A8,A10,Th19;
    A13: p`1 <= b`1 & p`2 <= b`2 by A1,A2,A3,A4,A8,A10,A11,Th19;
        now
        per cases;
        suppose
            a`1 <= b`1 & a`2 <= b`2;
          then a`1 <= p`1 & a`2 <= p`2 by A7,TOPREAL1:9,10;
          then a`1 = p`1 & a`2 = p`2 by A12,REAL_1:def 5;
          hence contradiction by A9,TOPREAL3:11;
        suppose
            a`1 <= b`1 & b`2 < a`2;
          then a`1 <= p`1 & b`2 <= p`2 by A7,TOPREAL1:9,10;
        then A14: a`1 = p`1 & b`2 = p`2 by A12,A13,REAL_1:def 5;
          then a`2 <> p`2 by A9,TOPREAL3:11;
          then LSeg(a,b) is vertical by A7,A10,A14,SPPOL_1:39;
          then a`1 = b`1 by SPPOL_1:37;
          hence contradiction by A9,A14,TOPREAL3:11;
        suppose
            b`1 < a`1 & a`2 <= b`2;
          then b`1 <= p`1 & a`2 <= p`2 by A7,TOPREAL1:9,10;
        then A15: b`1 = p`1 & a`2 = p`2 by A12,A13,REAL_1:def 5;
          then b`2 <> p`2 by A9,TOPREAL3:11;
          then LSeg(a,b) is vertical by A7,A10,A15,SPPOL_1:39;
          then a`1 = b`1 by SPPOL_1:37;
          hence contradiction by A9,A15,TOPREAL3:11;
        suppose
            b`1 < a`1 & b`2 < a`2;
          then b`1 <= p`1 & b`2 <= p`2 by A7,TOPREAL1:9,10;
          then b`1 = p`1 & b`2 = p`2 by A13,REAL_1:def 5;
          hence contradiction by A9,TOPREAL3:11;
      end;
      hence contradiction;

    suppose
    A16: p = G*(i,j+1);
    then A17: p`1 = G*(i,j)`1 by A1,A2,A3,A4,Th18;
    then A18: p`1 <= a`1 & a`2 <= p`2 by A1,A2,A3,A4,A8,A10,A16,Th19;
    A19: p`1 <= b`1 & b`2 <= p`2 by A1,A2,A3,A4,A8,A10,A16,A17,Th19;
        now
        per cases;
        suppose
            a`1 <= b`1 & a`2 <= b`2;
          then a`1 <= p`1 & p`2 <= b`2 by A7,TOPREAL1:9,10;
        then A20: a`1 = p`1 & b`2 = p`2 by A18,A19,REAL_1:def 5;
          then a`2 <> p`2 by A9,TOPREAL3:11;
          then LSeg(a,b) is vertical by A7,A10,A20,SPPOL_1:39;
          then a`1 = b`1 by SPPOL_1:37;
          hence contradiction by A9,A20,TOPREAL3:11;
        suppose
            a`1 <= b`1 & b`2 < a`2;
          then a`1 <= p`1 & p`2 <= a`2 by A7,TOPREAL1:9,10;
          then a`1 = p`1 & a`2 = p`2 by A18,REAL_1:def 5;
          hence contradiction by A9,TOPREAL3:11;
        suppose
            b`1 < a`1 & a`2 <= b`2;
          then b`1 <= p`1 & p`2 <= b`2 by A7,TOPREAL1:9,10;
          then b`1 = p`1 & b`2 = p`2 by A19,REAL_1:def 5;
          hence contradiction by A9,TOPREAL3:11;
        suppose
            b`1 < a`1 & b`2 < a`2;
          then b`1 <= p`1 & p`2 <= a`2 by A7,TOPREAL1:9,10;
        then A21: b`1 = p`1 & a`2 = p`2 by A18,A19,REAL_1:def 5;
          then b`2 <> p`2 by A9,TOPREAL3:11;
          then LSeg(a,b) is vertical by A7,A10,A21,SPPOL_1:39;
          then a`1 = b`1 by SPPOL_1:37;
          hence contradiction by A9,A21,TOPREAL3:11;
      end;
      hence contradiction;

    suppose
        p = G*(i+1,j+1);
    then A22: p`1 = G*(i+1,j)`1 & p`2 = G*(i,j+1)`2 by A1,A2,A3,A4,Th18;
    then A23: a`1 <= p`1 & a`2 <= p`2 by A1,A2,A3,A4,A8,A10,Th19;
    A24: b`1 <= p`1 & b`2 <= p`2 by A1,A2,A3,A4,A8,A10,A22,Th19;
        now
        per cases;
        suppose
            a`1 <= b`1 & a`2 <= b`2;
          then p`1 <= b`1 & p`2 <= b`2 by A7,TOPREAL1:9,10;
          then b`1 = p`1 & b`2 = p`2 by A24,REAL_1:def 5;
          hence contradiction by A9,TOPREAL3:11;
        suppose
            a`1 <= b`1 & b`2 < a`2;
          then p`1 <= b`1 & p`2 <= a`2 by A7,TOPREAL1:9,10;
        then A25: b`1 = p`1 & a`2 = p`2 by A23,A24,REAL_1:def 5;
          then b`2 <> p`2 by A9,TOPREAL3:11;
          then LSeg(a,b) is vertical by A7,A10,A25,SPPOL_1:39;
          then a`1 = b`1 by SPPOL_1:37;
          hence contradiction by A9,A25,TOPREAL3:11;
        suppose
            b`1 < a`1 & a`2 <= b`2;
          then p`1 <= a`1 & p`2 <= b`2 by A7,TOPREAL1:9,10;
        then A26: a`1 = p`1 & b`2 = p`2 by A23,A24,REAL_1:def 5;
          then a`2 <> p`2 by A9,TOPREAL3:11;
          then LSeg(a,b) is vertical by A7,A10,A26,SPPOL_1:39;
          then a`1 = b`1 by SPPOL_1:37;
          hence contradiction by A9,A26,TOPREAL3:11;
        suppose
            b`1 < a`1 & b`2 < a`2;
          then p`1 <= a`1 & p`2 <= a`2 by A7,TOPREAL1:9,10;
          then a`1 = p`1 & a`2 = p`2 by A23,REAL_1:def 5;
          hence contradiction by A9,TOPREAL3:11;
      end;
      hence contradiction;
    suppose
    A27: p = G*(i+1,j);
    then A28: p`2 = G*(i,j)`2 by A1,A2,A3,A4,Th18;
    then A29: a`1 <= p`1 & p`2 <= a`2 by A1,A2,A3,A4,A8,A10,A27,Th19;
    A30: b`1 <= p`1 & p`2 <= b`2 by A1,A2,A3,A4,A8,A10,A27,A28,Th19;
        now
        per cases;
        suppose
            a`1 <= b`1 & a`2 <= b`2;
          then p`1 <= b`1 & a`2 <= p`2 by A7,TOPREAL1:9,10;
        then A31: b`1 = p`1 & a`2 = p`2 by A29,A30,REAL_1:def 5;
          then b`2 <> p`2 by A9,TOPREAL3:11;
          then LSeg(a,b) is vertical by A7,A10,A31,SPPOL_1:39;
          then a`1 = b`1 by SPPOL_1:37;
          hence contradiction by A9,A31,TOPREAL3:11;
        suppose
            a`1 <= b`1 & b`2 < a`2;
          then p`1 <= b`1 & b`2 <= p`2 by A7,TOPREAL1:9,10;
          then b`1 = p`1 & b`2 = p`2 by A30,REAL_1:def 5;
          hence contradiction by A9,TOPREAL3:11;
        suppose
            b`1 < a`1 & a`2 <= b`2;
          then p`1 <= a`1 & a`2 <= p`2 by A7,TOPREAL1:9,10;
          then a`1 = p`1 & a`2 = p`2 by A29,REAL_1:def 5;
          hence contradiction by A9,TOPREAL3:11;
        suppose
            b`1 < a`1 & b`2 < a`2;
          then p`1 <= a`1 & b`2 <= p`2 by A7,TOPREAL1:9,10;
        then A32: a`1 = p`1 & b`2 = p`2 by A29,A30,REAL_1:def 5;
          then a`2 <> p`2 by A9,TOPREAL3:11;
          then LSeg(a,b) is vertical by A7,A10,A32,SPPOL_1:39;
          then a`1 = b`1 by SPPOL_1:37;
          hence contradiction by A9,A32,TOPREAL3:11;
      end;
      hence contradiction;
  end;
  hence thesis by A6,SPPOL_1:def 1;
end;

theorem Th24:
  2 <= len G & 2 <= width G & f is_sequence_on G & 1 <= k & k+1 <= len f
    implies
   ex i,j st 1 <= i & i+1 <= len G & 1 <= j & j+1 <= width G &
     LSeg(f,k) c= cell(G,i,j)
proof
  assume that
A1: 2 <= len G and
A2: 2 <= width G and
A3: f is_sequence_on G and
A4: 1 <= k and
A5: k+1 <= len f;

A6: LSeg(f,k) = LSeg(f/.k, f/.(k+1)) by A4,A5,TOPREAL1:def 5;

  consider i1,j1,i2,j2 being Nat such that
A7: [i1,j1] in Indices G and
A8: f/.k = G*(i1,j1) and
A9: [i2,j2] in Indices G and
A10: f/.(k+1) = G*(i2,j2) and
A11: i1 = i2 & j1+1 = j2 or i1+1 = i2 & j1 = j2 or i1 = i2+1 & j1 = j2 or
      i1 = i2 & j1 = j2+1 by A3,A4,A5,JORDAN8:6;
A12: 1 <= i1 & i1 <= len G & 1 <= j1 & j1 <= width G by A7,GOBOARD5:1;
A13: 1 <= i2 & i2 <= len G & 1 <= j2 & j2 <= width G by A9,GOBOARD5:1;

  per cases by A11;
  suppose
  A14: i1 = i2 & j1+1 = j2;
  then A15: j1 < width G by A13,SPPOL_1:5;
      now
      per cases by A12,REAL_1:def 5;
      suppose
          i1 < len G;
      then A16: i1+1 <= len G by NAT_1:38;
        take i1,j1;
          LSeg(f,k) c= cell(G,i1,j1) by A6,A8,A10,A12,A14,A15,GOBOARD5:20;
        hence thesis by A12,A13,A14,A16;
      suppose
      A17: i1 = len G;
        take i1'=i1-'1,j1;
          2-1 <= 2-'1 & 2-'1 <= i1' by A1,A17,JORDAN3:3,5;
      then A18: 1 <= i1' by AXIOMS:22;
      A19: i1'+1 = i1 by A12,AMI_5:4;
           then i1' < len G by A12,NAT_1:38;
        then LSeg(f,k) c= cell(G,i1',j1) by A6,A8,A10,A12,A14,A15,A19,GOBOARD5:
19;
        hence thesis by A12,A13,A14,A18,A19;
    end;
    hence thesis;
  suppose
  A20: i1+1 = i2 & j1 = j2;
  then A21: i1 < len G by A13,SPPOL_1:5;
      now
      per cases by A12,REAL_1:def 5;
      suppose
          j1 < width G;
      then A22: j1+1 <= width G by NAT_1:38;
        take i1,j1;
          LSeg(f,k) c= cell(G,i1,j1) by A6,A8,A10,A12,A20,A21,GOBOARD5:23;
        hence thesis by A12,A13,A20,A22;
      suppose
      A23: j1 = width G;
        take i1,j1'=j1-'1;
          2-1 <= 2-'1 & 2-'1 <= j1' by A2,A23,JORDAN3:3,5;
      then A24: 1 <= j1' by AXIOMS:22;
      A25: j1'+1=j1 by A12,AMI_5:4;
      then j1' < width G by A23,NAT_1:38;
        then LSeg(f,k) c= cell(G,i1,j1') by A6,A8,A10,A12,A20,A21,A25,GOBOARD5:
22;
        hence thesis by A12,A13,A20,A24,A25;
    end;
    hence thesis;
  suppose
  A26: i1 = i2+1 & j1 = j2;
  then A27: i2 < len G by A12,SPPOL_1:5;
      now
      per cases by A12,REAL_1:def 5;
      suppose
          j1 < width G;
      then A28: j1+1 <= width G by NAT_1:38;
        take i2,j1;
          LSeg(f,k) c= cell(G,i2,j1) by A6,A8,A10,A13,A26,A27,GOBOARD5:23;
        hence thesis by A12,A13,A26,A28;
      suppose
      A29: j1 = width G;
        take i2,j1'=j1-'1;
          2-1 <= 2-'1 & 2-'1 <= j1' by A2,A29,JORDAN3:3,5;
      then A30: 1 <= j1' by AXIOMS:22;
      A31: j1'+1=j1 by A12,AMI_5:4;
      then j1' < width G by A29,NAT_1:38;
        then LSeg(f,k) c= cell(G,i2,j1') by A6,A8,A10,A13,A26,A27,A31,GOBOARD5:
22;
        hence thesis by A12,A13,A26,A30,A31;
    end;
    hence thesis;
  suppose
  A32: i1 = i2 & j1 = j2+1;
  then A33: j2 < width G by A12,SPPOL_1:5;
      now
      per cases by A12,REAL_1:def 5;
      suppose
          i1 < len G;
      then A34: i1+1 <= len G by NAT_1:38;
        take i1,j2;
          LSeg(f,k) c= cell(G,i1,j2) by A6,A8,A10,A13,A32,A33,GOBOARD5:20;
        hence thesis by A12,A13,A32,A34;
      suppose
      A35: i1 = len G;
        take i1'=i1-'1,j2;
          2-1 <= 2-'1 & 2-'1 <= i1' by A1,A35,JORDAN3:3,5;
      then A36: 1 <= i1' by AXIOMS:22;
      A37: i1'+1 = i1 by A12,AMI_5:4;
           then i1' < len G by A12,NAT_1:38;
        then LSeg(f,k) c= cell(G,i1',j2) by A6,A8,A10,A13,A32,A33,A37,GOBOARD5:
19;
        hence thesis by A12,A13,A32,A36,A37;
    end;
    hence thesis;
end;

theorem Th25:
  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)
      implies
   p = f/.k or p = f/.(k+1)
proof
  assume that
A1: 2 <= len G and
A2: 2 <= width G and
A3: f is_sequence_on G and
A4: 1 <= k and
A5: k+1 <= len f and
A6: p in Values G and
A7: p in LSeg(f,k);

  consider i,j such that
A8: 1 <= i & i+1 <= len G and
A9: 1 <= j & j+1 <= width G and
A10: LSeg(f,k) c= cell(G,i,j) by A1,A2,A3,A4,A5,Th24;
A11: LSeg(f,k) = LSeg(f/.k, f/.(k+1)) by A4,A5,TOPREAL1:def 5;
    p is_extremal_in cell(G,i,j) by A6,A7,A8,A9,A10,Th23;
  hence thesis by A7,A10,A11,SPPOL_1:def 1;
end;

theorem
    [i,j] in Indices G & 1 <= k & k <= width G implies G*(i,j)`1 <= G*
(len G,k)`1
proof
  assume that
A1: [i,j] in Indices G and
A2: 1 <= k & k <= width G;
A3: 1 <= i & i <= len G & 1 <= j & j <= width G by A1,GOBOARD5:1;
then A4: G*(i,j)`1 = G*(i,1)`1 by GOBOARD5:3
             .= G*(i,k)`1 by A2,A3,GOBOARD5:3;
    i < len G or i = len G by A3,REAL_1:def 5;
 hence thesis by A2,A3,A4,GOBOARD5:4;
end;

theorem
    [i,j] in Indices G & 1 <= k & k <= len G implies G*(i,j)`2 <= G*
(k,width G)`2
proof
  assume that
A1: [i,j] in Indices G and
A2: 1 <= k & k <= len G;
A3: 1 <= i & i <= len G & 1 <= j & j <= width G by A1,GOBOARD5:1;
then A4: G*(i,j)`2 = G*(1,j)`2 by GOBOARD5:2
             .= G*(k,j)`2 by A2,A3,GOBOARD5:2;
    j < width G or j = width G by A3,REAL_1:def 5;
 hence thesis by A2,A3,A4,GOBOARD5:5;
end;

theorem Th28:
  f is_sequence_on G & f is special & L~g c= L~f & 1 <= k & k+1 <= len f
    implies
   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
  assume that
A1: f is_sequence_on G & f is special & L~g c= L~f and
A2: 1 <= k and
A3: k+1 <= len f;
  let A be Subset of TOP-REAL 2 such that
A4: A = right_cell(f,k,G)\L~g or A = left_cell(f,k,G)\L~g;

  per cases by A4;
  suppose
  A5: A = right_cell(f,k,G)\L~g;

then A6: A = right_cell(f,k,G) /\ (L~g)` by SUBSET_1:32;
A7: A c= right_cell(f,k,G) by A5,XBOOLE_1:36;

  A8: Int right_cell(f,k,G) is connected by A1,A2,A3,Th12;
      Int right_cell(f,k,G) misses L~f by A1,A2,A3,Th17;
    then Int right_cell(f,k,G) misses L~g by A1,XBOOLE_1:63;
  then A9: Int right_cell(f,k,G) c= (L~g)` by SUBSET_1:43;
      Int right_cell(f,k,G) c= right_cell(f,k,G) by TOPS_1:44;
  then A10: Int right_cell(f,k,G) c= A by A6,A9,XBOOLE_1:19;
      A c= Cl Int right_cell(f,k,G) by A1,A2,A3,A7,Th13;
  hence A is connected by A8,A10,CONNSP_1:19;

  suppose
  A11: A = left_cell(f,k,G)\L~g;

then A12: A = left_cell(f,k,G) /\ (L~g)` by SUBSET_1:32;
A13: A c= left_cell(f,k,G) by A11,XBOOLE_1:36;

  A14: Int left_cell(f,k,G) is connected by A1,A2,A3,Th12;
      Int left_cell(f,k,G) misses L~f by A1,A2,A3,Th17;
    then Int left_cell(f,k,G) misses L~g by A1,XBOOLE_1:63;
  then A15: Int left_cell(f,k,G) c= (L~g)` by SUBSET_1:43;
      Int left_cell(f,k,G) c= left_cell(f,k,G) by TOPS_1:44;
  then A16: Int left_cell(f,k,G) c= A by A12,A15,XBOOLE_1:19;
      A c= Cl Int left_cell(f,k,G) by A1,A2,A3,A13,Th13;
  hence A is connected by A14,A16,CONNSP_1:19;
end;

theorem Th29:
  for f being non constant standard special_circular_sequence st
    f is_sequence_on G
  for k 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
  let f be non constant standard special_circular_sequence such that
A1: f is_sequence_on G;
  let k such that
A2: 1 <= k & k+1 <= len f;

  set rc = right_cell(f,k,G)\L~f;
A3: rc = right_cell(f,k,G) /\ (L~f)` by SUBSET_1:32;
A4: rc c= right_cell(f,k,G) by XBOOLE_1:36;
    right_cell(f,k,G) c= right_cell(f,k) by A1,A2,GOBRD13:34;
  then rc c= right_cell(f,k) by A4,XBOOLE_1:1;
then A5: Int rc c= Int right_cell(f,k) by TOPS_1:48;
    Int right_cell(f,k) c= RightComp f by A2,GOBOARD9:28;
then A6: Int rc c= RightComp f by A5,XBOOLE_1:1;
A7: Int right_cell(f,k,G) <> {} by A1,A2,Th11;
A8:  Int right_cell(f,k,G) misses L~f by A1,A2,Th17;
A9: Int right_cell(f,k,G) c= right_cell(f,k,G) by TOPS_1:44;
    rc \/ L~f = right_cell(f,k,G) \/ L~f by XBOOLE_1:39;
  then right_cell(f,k,G) c= rc \/ L~f by XBOOLE_1:7;
then Int right_cell(f,k,G) c= rc \/ L~f by A9,XBOOLE_1:1;
then A10: Int right_cell(f,k,G) c= rc by A8,XBOOLE_1:73;
  then Int Int right_cell(f,k,G) c= Int rc by TOPS_1:48;
  then Int right_cell(f,k,G) c= Int rc by TOPS_1:45;
then A11: rc meets Int rc by A7,A10,XBOOLE_1:68;
  A12: rc is connected by A1,A2,Th28;
  A13: RightComp f is_a_component_of (L~f)` by GOBOARD9:def 2;
    rc c= (L~f)` by A3,XBOOLE_1:17;
  hence right_cell(f,k,G)\L~f c= RightComp f by A6,A11,A12,A13,GOBOARD9:6;
  set lc = left_cell(f,k,G)\L~f;
A14: lc = left_cell(f,k,G) /\ (L~f)` by SUBSET_1:32;
A15: lc c= left_cell(f,k,G) by XBOOLE_1:36;
    left_cell(f,k,G) c= left_cell(f,k) by A1,A2,GOBRD13:34;
  then lc c= left_cell(f,k) by A15,XBOOLE_1:1;
then A16: Int lc c= Int left_cell(f,k) by TOPS_1:48;
    Int left_cell(f,k) c= LeftComp f by A2,GOBOARD9:24;
then A17: Int lc c= LeftComp f by A16,XBOOLE_1:1;
A18: Int left_cell(f,k,G) <> {} by A1,A2,Th11;
A19:  Int left_cell(f,k,G) misses L~f by A1,A2,Th17;
A20: Int left_cell(f,k,G) c= left_cell(f,k,G) by TOPS_1:44;
    lc \/ L~f = left_cell(f,k,G) \/ L~f by XBOOLE_1:39;
  then left_cell(f,k,G) c= lc \/ L~f by XBOOLE_1:7;
then Int left_cell(f,k,G) c= lc \/ L~f by A20,XBOOLE_1:1;
then A21: Int left_cell(f,k,G) c= lc by A19,XBOOLE_1:73;
  then Int Int left_cell(f,k,G) c= Int lc by TOPS_1:48;
  then Int left_cell(f,k,G) c= Int lc by TOPS_1:45;
then A22: lc meets Int lc by A18,A21,XBOOLE_1:68;
  A23: lc is connected by A1,A2,Th28;
  A24: LeftComp f is_a_component_of (L~f)` by GOBOARD9:def 1;
    lc c= (L~f)` by A14,XBOOLE_1:17;
  hence left_cell(f,k,G)\L~f c= LeftComp f by A17,A22,A23,A24,GOBOARD9:6;
end;

begin :: Cages

reserve
  C for compact non vertical non horizontal non empty Subset of TOP-REAL 2,
  l, m, i1, i2, j1, j2 for Nat;

theorem Th30:
  ex i 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
 set G = Gauge(C,n);
A1: len G = width G by JORDAN8:def 1;
 defpred P[Nat]
   means 1 <= $1 & $1 < len G & G*($1,(width G)-'1)`1 < (N-min C)`1;
A2:  for k st P[k] holds k <= len G;
A3:  len G = 2|^n+3 & len G = width G by JORDAN8:def 1;
A4:  len G >= 4 by JORDAN8:13;
then A5:  1 < len G by AXIOMS:22;
then A6:  1 <= (len G)-'1 by JORDAN3:12;
A7:  (len G)-'1 <= len G by GOBOARD9:2;
    (NW-corner C)`1 <= (N-min C)`1 by PSCOMP_1:97;
then A8:  W-bound C <= (N-min C)`1 by PSCOMP_1:74;
A9:  2 <= len G by A4,AXIOMS:22;
    G*(2,(width G)-'1)`1 = W-bound C by A1,A6,A7,JORDAN8:14;
  then G*(1,(width G)-'1)`1 < W-bound C by A3,A6,A7,A9,GOBOARD5:4;
  then G*(1,(width G)-'1)`1 < (N-min C)`1 by A8,AXIOMS:22;
then A10:  ex k st P[k] by A5;
  ex i st P[i] & for n st P[n] holds n <= i from Max(A2,A10);
  then consider i such that
A11:  1 <= i & i < len G and
A12:  G*(i,(width G)-'1)`1 < (N-min C)`1 and
A13:  for n st P[n] holds n <= i;
 take i;
 thus 1 <= i & i+1 <= len G by A11,NAT_1:38;
A14:  LSeg(G*(i,(width G)-'1),G*(i+1,(width G)-'1)) c= cell(G,i,(width G)-'1)
      by A3,A6,A7,A11,GOBOARD5:23;
A15:  i+1 <= len G by A11,NAT_1:38;
A16:  1 <= i+1 by NAT_1:37;
   (N-min C)`2 = N-bound C by PSCOMP_1:94;
then A17:  G*(i,(width G)-'1)`2 = (N-min C)`2 & (N-min C)`2 = G*
(i+1,(width G)-'1)`2
      by A1,A11,A15,A16,JORDAN8:17;
    now assume i+1 = len G;
   then len G-'1 = i by BINARITH:39;
then A18:   G*(i,(width G)-'1)`1 = E-bound C by A1,A6,A7,JORDAN8:15;
     (NE-corner C)`1 >= (N-min C)`1 by PSCOMP_1:97;
   hence contradiction by A12,A18,PSCOMP_1:76;
  end;
 then i+1 < len G & i < i+1 by A15,NAT_1:38,REAL_1:def 5;
 then (N-min C)`1 <= G*(i+1,(width G)-'1)`1 by A13,A16;
 then N-min C in LSeg(G*(i,(width G)-'1),G*(i+1,(width G)-'1))
   by A12,A17,GOBOARD7:9;
 hence N-min C in cell(G,i,(width G)-'1) by A14;
 thus N-min C <> G*(i,(width G)-'1) by A12;
end;

theorem Th31:
  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)
    implies
   i1 = i2
proof
  set G = Gauge(C,n), j = width G-'1;
  assume that
A1:  1 <= i1 & i1+1 <= len G and
A2:  N-min C in cell(G,i1,j) and
A3:  N-min C <> G*(i1,j) and
A4:  1 <= i2 & i2+1 <= len G and
A5:  N-min C in cell(G,i2,j) and
A6:  N-min C <> G*(i2,j) and
A7:  i1 <> i2;
A8: i1 < len G & i2 < len G & len G = width G by A1,A4,JORDAN8:def 1,NAT_1:38;
A9:  cell(G,i1,j) meets cell(G,i2,j) by A2,A5,XBOOLE_0:3;
A10:  len G = 2|^n+3 & len G = width G by JORDAN8:def 1;
A11: n+3 > 0 by NAT_1:19;
A12:  2|^n >= n+1 by HEINE:7;
   then len G >= n+1+3 by A10,AXIOMS:24;
   then len G >= 1+ (n+3) & 1+ (n+3) > 1+0 by A11,REAL_1:53,XCMPLX_1:1;
then A13:  len G > 1 by AXIOMS:22;
   then len G >= 1+1 by NAT_1:38;
then A14:  1 <= j by A8,JORDAN5B:2;
A15:  j+1 = len G by A8,A13,AMI_5:4;
then A16:  j < len G by NAT_1:38;
A17:  (N-min C)`2 = N-bound C by PSCOMP_1:94;
 per cases by A7,REAL_1:def 5;
  suppose
A18:   i1 < i2;
then A19:   i2-'i1+i1 = i2 by AMI_5:4;
     then i2-'i1 <= 1 by A8,A9,A14,A16,JORDAN8:10;
     then i2-'i1 < 1 or i2-'i1 = 1 by REAL_1:def 5;
     then i2-'i1 = 0 or i2-'i1 = 1 by RLVECT_1:98;
    then cell(G,i1,j) /\ cell(G,i2,j) = LSeg(G*(i2,j),G*(i2,j+1))
        by A8,A14,A16,A18,A19,GOBOARD5:26;
    then A20:  N-min C in LSeg(G*(i2,j),G*(i2,j+1)) by A2,A5,XBOOLE_0:def 3;
A21:  [i2,j] in Indices G by A4,A8,A14,A16,GOBOARD7:10;
      1 <= j+1 by NAT_1:37;
then A22:  [i2,j+1] in Indices G by A4,A8,A15,GOBOARD7:10;
     set x = (W-bound C)+(((E-bound C)-(W-bound C))/(2|^n))*(i2-2);
     set y1 = (S-bound C)+(((N-bound C)-(S-bound C))/(2|^n))*(j-2);
     set y2 = (S-bound C)+(((N-bound C)-(S-bound C))/(2|^n))*(j-1);
A23:   j+1-(1+1) = j+1-1-1 by XCMPLX_1:36
               .= j-1 by XCMPLX_1:26;
A24:   G*(i2,j) = |[x,y1]| by A21,JORDAN8:def 1;
    G*(i2,j+1) = |[x,y2]| by A22,A23,JORDAN8:def 1;
then A25:  G*(i2,j)`1 = x & G*(i2,j+1)`1 = x by A24,EUCLID:56;
     then LSeg(G*(i2,j),G*(i2,j+1)) is vertical by SPPOL_1:37;
then A26:  (N-min C)`1 = G*(i2,j)`1 by A20,SPRECT_3:20;
A27:   2|^n > 0 by A12,NAT_1:37;
       j = (2|^n+(2+1))-'1 by A10
      .= (2|^n+2+1)-'1 by XCMPLX_1:1
      .= (2|^n+2) by BINARITH:39;
     then j-2 = 2|^n by XCMPLX_1:26;
     then (((N-bound C)-(S-bound C))/(2|^n))*(j-2) = (N-bound C)-(S-bound C)
       by A27,XCMPLX_1:88;
     then y1 = N-bound C by XCMPLX_1:27;
   hence contradiction by A6,A17,A24,A25,A26,EUCLID:57;
  suppose
A28:   i2 < i1;
then A29:   i1-'i2+i2 = i1 by AMI_5:4;
     then i1-'i2 <= 1 by A8,A9,A14,A16,JORDAN8:10;
     then i1-'i2 < 1 or i1-'i2 = 1 by REAL_1:def 5;
     then i1-'i2 = 0 or i1-'i2 = 1 by RLVECT_1:98;
    then cell(G,i2,j) /\ cell(G,i1,j) = LSeg(G*(i1,j),G*(i1,j+1))
        by A8,A14,A16,A28,A29,GOBOARD5:26;
    then A30:  N-min C in LSeg(G*(i1,j),G*(i1,j+1)) by A2,A5,XBOOLE_0:def 3;
A31:  [i1,j] in Indices G by A1,A8,A14,A16,GOBOARD7:10;
      1 <= j+1 by NAT_1:37;
then A32:  [i1,j+1] in Indices G by A1,A8,A15,GOBOARD7:10;
     set x = (W-bound C)+(((E-bound C)-(W-bound C))/(2|^n))*(i1-2);
     set y1 = (S-bound C)+(((N-bound C)-(S-bound C))/(2|^n))*(j-2);
     set y2 = (S-bound C)+(((N-bound C)-(S-bound C))/(2|^n))*(j-1);
A33:   j+1-(1+1) = j+1-1-1 by XCMPLX_1:36
               .= j-1 by XCMPLX_1:26;
A34:   G*(i1,j) = |[x,y1]| by A31,JORDAN8:def 1;
    G*(i1,j+1) = |[x,y2]| by A32,A33,JORDAN8:def 1;
then A35:  G*(i1,j)`1 = x & G*(i1,j+1)`1 = x by A34,EUCLID:56;
    then LSeg(G*(i1,j),G*(i1,j+1)) is vertical by SPPOL_1:37;
then A36:  (N-min C)`1 = G*(i1,j)`1 by A30,SPRECT_3:20;
A37:   2|^n > 0 by A12,NAT_1:37;
       j = (2|^n+(2+1))-'1 by A10
      .= (2|^n+2+1)-'1 by XCMPLX_1:1
      .= (2|^n+2) by BINARITH:39;
     then j-2 = 2|^n by XCMPLX_1:26;
     then (((N-bound C)-(S-bound C))/(2|^n))*(j-2) = (N-bound C)-(S-bound C)
       by A37,XCMPLX_1:88;
     then y1 = N-bound C by XCMPLX_1:27;
   hence contradiction by A3,A17,A34,A35,A36,EUCLID:57;
end;

theorem Th32:
  for f being standard non constant special_circular_sequence
    st f is_sequence_on Gauge(C,n) &
      (for k 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 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
  set G = Gauge(C,n);
  let f be standard non constant special_circular_sequence such that
A1: f is_sequence_on G and
A2: for k st 1 <= k & k+1 <= len f
       holds left_cell(f,k,G) misses C & right_cell(f,k,G) meets C;
  given i' being Nat such that
A3:    1 <= i' & i'+1 <= len G and
A4:    f/.1 = G*(i',width G) and
A5:    f/.2 = G*(i'+1,width G) and
A6:    N-min C in cell(G,i',width G-'1) and
A7:    N-min C <> G*(i',width G-'1);
A8: i' < len G & len G = width G by A3,JORDAN8:def 1,NAT_1:38;
   len f > 4 by GOBOARD7:36;
then A9: len f >= 2 by AXIOMS:22;
A10:  len G = 2|^n+3 & len G = width G by JORDAN8:def 1;
    then len G >= 3 by NAT_1:37;
then A11:  1 < len G by AXIOMS:22;
     set W = W-bound C, S = S-bound C, E = E-bound C, N = N-bound C;
       N-min L~f in rng f by SPRECT_2:43;
     then consider m such that
A12:    m in dom f and
A13:    f.m = N-min L~f by FINSEQ_2:11;
A14:    f/.m = f.m by A12,FINSEQ_4:def 4;
A15:    (N-min L~f)`2 = N-bound L~f by PSCOMP_1:94;
A16:    2|^n > 0 by HEINE:5;
       N > S by JORDAN8:12;
     then N-S > 0 by SQUARE_1:11;
then A17:    (N-S)/(2|^n) > 0 by A16,REAL_2:127;
A18:    1 <= m & m <= len f by A12,FINSEQ_3:27;
     consider i,j such that
A19:     [i,j] in Indices G and
A20:     f/.m = G*(i,j) by A1,A12,GOBOARD1:def 11;
A21:     1 <= i & i <= len G & 1 <= j & j <= width G by A19,GOBOARD5:1;
       G*(i,j) = |[W+((E-W)/(2|^n))*(i-2), S+((N-S)/(2|^n))*(j-2)]|
       by A19,JORDAN8:def 1;
then A22:     S+((N-S)/(2|^n))*(j-2) = N-bound L~f by A13,A14,A15,A20,EUCLID:56
;
       1 in dom f by FINSEQ_5:6;
     then A23:    f/.1 in L~f by A9,GOBOARD1:16;
then A24:    N-bound L~f >= (f/.1)`2 by PSCOMP_1:71;
       G*(i',j)`2 = G*(1,j)`2 & G*(i,j)`2 = G*
(1,j)`2 by A3,A8,A21,GOBOARD5:2;
     then G*(i,j)`2 <= G*(i',len G)`2 by A3,A8,A21,SPRECT_3:24;
then A25:    N-bound L~f = (f/.1)`2 by A4,A8,A13,A14,A15,A20,A24,AXIOMS:21;
       [i',len G] in Indices G by A3,A8,A11,GOBOARD7:10;
     then G*(i',len G)=|[W+((E-W)/(2|^n))*(i'-2),S+((N-S)/(2|^n))*(len G-2)]|
      by JORDAN8:def 1;
     then S+((N-S)/(2|^n))*(len G-2) = N-bound L~f by A4,A8,A25,EUCLID:56;
     then ((N-S)/(2|^n))*(j-2) = ((N-S)/(2|^n))*(len G-2) by A22,XCMPLX_1:2;
     then len G-2 = j-2 by A17,XCMPLX_1:5;
then A26:    len G = j by XCMPLX_1:19;
A27:    (NW-corner C)`1 = W & (NE-corner C)`1 = E &
         (NW-corner C)`2 = N & (NE-corner C)`2 = N
           by PSCOMP_1:74,75,76,77;
A28:    G*(i',len G)`1 = G*(i',1)`1 by A3,A8,A21,A26,GOBOARD5:3;
A29:    G*(i,len G)`1 = G*(i,1)`1 by A21,A26,GOBOARD5:3;
A30:    (NW-corner L~f)`1 = W-bound L~f & (NE-corner L~f)`1 = E-bound L~f &
         (NW-corner L~f)`2 = N-bound L~f & (NE-corner L~f)`2 = N-bound L~f
          by PSCOMP_1:74,75,76,77;
       W-bound L~f <= (f/.1)`1 & (f/.1)`1 <= E-bound L~f by A23,PSCOMP_1:71;
     then f/.1 in LSeg(NW-corner L~f, NE-corner L~f) by A25,A30,GOBOARD7:9;
     then f/.1 in LSeg(NW-corner L~f, NE-corner L~f) /\ L~f
       by A23,XBOOLE_0:def 3;
     then f/.1 in N-most L~f by PSCOMP_1:def 39;
then A31:    (N-min L~f)`1 <= (f/.1)`1 by PSCOMP_1:98;
A32:    1 <= len G-'1 by A11,JORDAN3:12;
then A33:    len G-'1 < len G by JORDAN3:14;
A34:    len G-'1+1 = len G by A11,AMI_5:4;
     then N-min C in { |[r',s']| where r',s' is Real:
               G*(i',1)`1 <= r' & r' <= G*(i'+1,1)`1 &
                   G*(1,len G-'1)`2 <= s' & s' <= G*(1,len G)`2 }
         by A3,A6,A8,A32,A33,GOBRD11:32;
     then consider r',s' being Real such that
A35:    N-min C = |[r',s']| and
A36:    G*(i',1)`1 <= r' & r' <= G*(i'+1,1)`1 and
          G*(1,len G-'1)`2 <= s' & s' <= G*(1,len G)`2;
A37:    (f/.1)`1 <= (N-min C)`1 by A4,A8,A28,A35,A36,EUCLID:56;
then A38:    (N-min L~f)`1 <= (N-min C)`1 by A31,AXIOMS:22;
A39:    G*(i',len G-'1)`1 = G*(i',1)`1 by A3,A8,A32,A33,GOBOARD5:3;
A40:    N-min C = |[(N-min C)`1,(N-min C)`2]| by EUCLID:57;
A41:    G*(i',len G-'1) = |[G*(i',len G-'1)`1,G*(i',len G-'1)`2]|
          by EUCLID:57;
A42:    (N-min C)`2 = N by PSCOMP_1:94;
     G*(i',len G-'1)`2 = N by A3,A8,JORDAN8:17;
then A43:    G*(i',len G-'1)`1 < (N-min C)`1
         by A4,A7,A8,A28,A37,A39,A40,A41,A42,REAL_1:def 5;
A44: i <= i' by A3,A4,A8,A13,A14,A20,A21,A26,A31,GOBOARD5:4;
     then A45:    i < len G by A8,AXIOMS:22;
then A46:   i+1 <= len G by NAT_1:38;
     per cases by A18,REAL_1:def 5;
     suppose m = len f;
       hence N-min L~f = f/.1 by A13,A14,FINSEQ_6:def 1;
     suppose m < len f;
then A47:    m+1 <= len f by NAT_1:38;
then A48:    right_cell(f,m,G) meets C by A2,A18;
      then consider p being set such that
A49:     p in right_cell(f,m,G) & p in C by XBOOLE_0:3;
      reconsider p as Point of TOP-REAL 2 by A49;
      consider i1,j1,i2,j2 being Nat such that
A50:     [i1,j1] in Indices G & f/.m = G*(i1,j1) and
A51:     [i2,j2] in Indices G & f/.(m+1) = G*(i2,j2) and
A52:     i1 = i2 & j1+1 = j2 or i1+1 = i2 & j1 = j2 or
        i1 = i2+1 & j1 = j2 or i1 = i2 & j1 = j2+1 by A1,A18,A47,JORDAN8:6;
A53:     1 <= i2 & i2 <= len G & 1 <= j2 & j2 <= len G by A10,A51,GOBOARD5:1;
A54:     (N-min C)`2 = N by PSCOMP_1:94;
then A55:     p`2 <= (N-min C)`2 by A49,PSCOMP_1:71;
A56:     G*(1,len G-'1)`2 = N by A11,JORDAN8:17;
A57:     G*(1,len G-'1)`2 < G*(1,len G)`2 by A10,A11,A32,A33,GOBOARD5:5;
A58:     W <= p`1 & p`1 <= E by A49,PSCOMP_1:71;
       now per cases by A19,A20,A26,A50,A52,GOBOARD1:21;
      suppose i = i2 & len G+1 = j2;
       hence N-min L~f = f/.1 by A53,NAT_1:38;
      suppose i+1 = i2 & len G = j2;
then A59:     right_cell(f,m,G) = cell(G,i,len G-'1)
           by A1,A18,A19,A20,A26,A47,A51,GOBRD13:25;
A60:     cell(G,i,len G-'1) = {|[r,s]| : G*(i,1)`1 <= r & r <= G*(i+1,1)`1 &
           G*(1,len G-'1)`2 <= s & s <= G*(1,len G-'1+1)`2 }
          by A10,A21,A32,A33,A45,GOBRD11:32;
       then consider r,s such that
A61:     p = |[r,s]| and
A62:     G*(i,1)`1 <= r & r <= G*(i+1,1)`1 and
A63:     G*(1,len G-'1)`2 <= s & s <= G*(1,len G-'1+1)`2 by A49,A59;
A64:     p`1 = r & p`2 = s by A61,EUCLID:56;
       then p`2 = N by A54,A55,A56,A63,AXIOMS:21;
       then p in LSeg(NW-corner C, NE-corner C) by A27,A58,GOBOARD7:9;
       then p in LSeg(NW-corner C, NE-corner C) /\ C
         by A49,XBOOLE_0:def 3;
       then p in N-most C by PSCOMP_1:def 39;
       then (N-min C)`1 <= p`1 by PSCOMP_1:98;
         then (N-min C)`1 <= G*(i+1,1)`1 by A62,A64,AXIOMS:22;
       then A65:     N-min C in cell(G,i,width G-'1)
               by A8,A13,A14,A20,A26,A29,A34,A38,A40,A54,A56,A57,A60;
         N-min C <> G*(i,len G-'1) by A8,A21,A32,A33,A43,A44,SPRECT_3:25;
       hence N-min L~f = f/.1
          by A3,A4,A6,A7,A8,A13,A14,A20,A21,A26,A46,A65,Th31;
      suppose
A66:     i = i2+1 & len G = j2;
then A67:     right_cell(f,m,G) = cell(G,i2,len G)
           by A1,A18,A19,A20,A26,A47,A51,GOBRD13:27;
         i2 < len G by A21,A66,NAT_1:38;
       hence N-min L~f = f/.1 by A48,A67,JORDAN8:18;
      suppose
A68:     i = i2 & len G = j2+1;
then A69:     j2 = len G-'1 by BINARITH:39;
then A70:     right_cell(f,m,G) = cell(G,i-'1,len G-'1)
          by A1,A18,A19,A20,A26,A47,A51,A68,GOBRD13:29;
         now assume
A71:      m = 1;
         1 <= i'+1 & i'+1<= len G by A3,NAT_1:38;
       then G*(i',len G)`2 = G*(1,len G)`2 & G*(i'+1,len G)`2 = G*
(1,len G)`2 by A3,A8,A11,GOBOARD5:2;
        hence contradiction
         by A4,A5,A8,A20,A21,A26,A32,A33,A51,A68,A69,A71,GOBOARD5:5;
       end;
       then m > 1 by A18,REAL_1:def 5;
then A72:   m-'1 >= 1 by JORDAN3:12;
A73:   m-'1+1 = m by A18,AMI_5:4;
         m-'1 <= m by GOBOARD9:2;
then A74:   m-'1 <= len f by A18,AXIOMS:22;
       consider i1',j1',i2',j2' being Nat such that
A75:     [i1',j1'] in Indices G & f/.(m-'1) = G*(i1',j1') and
A76:     [i2',j2'] in Indices G & f/.m = G*(i2',j2') and
A77:     i1' = i2' & j1'+1 = j2' or i1'+1 = i2' & j1' = j2' or
          i1' = i2'+1 & j1' = j2' or i1' = i2' & j1' = j2'+1
           by A1,A18,A72,A73,JORDAN8:6;
A78:     1 <= i1' & i1' <= len G & 1 <= j1' &
           j1' <= len G by A10,A75,GOBOARD5:1;
         now per cases by A19,A20,A26,A76,A77,GOBOARD1:21;
        suppose
A79:        i1' = i & j1'+1 = len G;
         then j1' = len G-'1 by BINARITH:39;
         then left_cell(f,m-'1,G) = cell(G,i-'1,len G-'1)
           by A1,A18,A19,A20,A26,A72,A73,A75,A79,GOBRD13:22;
         hence contradiction by A2,A18,A48,A70,A72,A73;
        suppose
A80:     i1'+1 = i & j1' = len G;
         then i1' < i by NAT_1:38;
then A81:     (f/.(m-'1))`1 < (f/.m)`1 by A20,A21,A26,A75,A78,A80,GOBOARD5:4;
A82:     G*(i1',j)`2 = G*(1,j)`2 & G*(i,j)`2 = G*(1,j)`2
           by A21,A78,GOBOARD5:2;
           m-'1 in dom f by A72,A74,FINSEQ_3:27;
then A83:     f/.(m-'1) in L~f by A9,GOBOARD1:16;
         then W-bound L~f <= (f/.(m-'1))`1 &
           (f/.(m-'1))`1 <= E-bound L~f by PSCOMP_1:71;
         then f/.(m-'1) in LSeg(NW-corner L~f, NE-corner L~f)
          by A13,A14,A15,A20,A26,A30,A75,A80,A82,GOBOARD7:9;
         then f/.(m-'1) in LSeg(NW-corner L~f, NE-corner L~f) /\ L~f
          by A83,XBOOLE_0:def 3;
         then f/.(m-'1) in N-most L~f by PSCOMP_1:def 39;
         hence contradiction by A13,A14,A81,PSCOMP_1:98;
        suppose i1' = i+1 & j1' = len G;
         then right_cell(f,m-'1,G) = cell(G,i,len G)
           by A1,A18,A19,A20,A26,A72,A73,A75,GOBRD13:27;
         then cell(G,i,len G) meets C by A2,A18,A72,A73;
         hence contradiction by A21,JORDAN8:18;
        suppose i1' = i & j1' = len G+1;
         then len G+1 <= len G+0 by A10,A75,GOBOARD5:1;
         hence contradiction by REAL_1:53;
       end;
       hence N-min L~f = f/.1;
     end;
     hence N-min L~f = f/.1;
end;

definition
  let C be compact non vertical non horizontal non empty Subset of TOP-REAL 2;
  let n be Nat;
  assume
A1:    C is connected;
  func Cage(C,n) -> clockwise_oriented
    (standard non constant special_circular_sequence) means
:Def1:
  it is_sequence_on Gauge(C,n) &
    (ex i 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 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
proof
   set G = Gauge(C,n);
   set W = W-bound C, E = E-bound C, S = S-bound C, N = N-bound C;
A2:   len G = 2|^n+3 & len G = width G by JORDAN8:def 1;
   defpred P[Nat,set,set] means
      ($1 = 0 implies
        ex i st 1 <= i & i+1 <= len G &
             N-min C in cell(G,i,width G-'1) & N-min C <> G*(i,width G-'1) &
            $3 = <*G*(i,width G)*>) &
      ($1 = 1 implies
         ex i st 1 <= i & i+1 <= len G &
             N-min C in cell(G,i,width G-'1) & N-min C <> G*(i,width G-'1) &
            $3 = <*G*(i,width G),G*(i+1,width G)*>) &
      ($1 > 1 & $2 is FinSequence of TOP-REAL 2 implies
        for f being FinSequence of TOP-REAL 2 st $2 = f holds
         (len f = $1 implies
          (f is_sequence_on G & right_cell(f,len f-'1,G) meets C implies
           (front_left_cell(f,(len f)-'1,G) misses C &
              front_right_cell(f,(len f)-'1,G) misses C
             implies ex i,j st
                f^<*G*(i,j)*> turns_right (len f)-'1,G &
              $3 = f^<*G*(i,j)*>) &
           (front_left_cell(f,(len f)-'1,G) misses C &
              front_right_cell(f,(len f)-'1,G) meets C
             implies ex i,j st
               f^<*G*(i,j)*> goes_straight (len f)-'1,G &
              $3 = f^<*G*(i,j)*>) &
           (front_left_cell(f,(len f)-'1,G) meets C
             implies ex i,j st
               f^<*G*(i,j)*> turns_left (len f)-'1,G &
              $3 = f^<*G*(i,j)*>)) &
          (not f is_sequence_on G or right_cell(f,len f-'1,G) misses C
           implies $3 = f^<*G*(1,1)*>)) &
         (len f <> $1 implies $3 = {})) &
       ($1 > 1 & $2 is not FinSequence of TOP-REAL 2 implies $3 = {});
A3:  for k being Nat, x being set ex y being set st P[k,x,y]
   proof
    let k be Nat, x be set;
    consider m being Nat such that
A4:   1 <= m & m+1 <= len G and
A5:   N-min C in cell(G,m,width G-'1) &
        N-min C <> G*(m,width G-'1) by Th30;
    per cases by CQC_THE1:2;
    suppose
A6:    k=0;
     take <*G*(m,width G)*>;
     thus thesis by A4,A5,A6;
    suppose
A7:    k = 1;
     take <*G*(m,width G),G*(m+1,width G)*>;
     thus thesis by A4,A5,A7;
    suppose that
A8:    k > 1 and
A9:    x is FinSequence of TOP-REAL 2;
     reconsider f = x as FinSequence of TOP-REAL 2 by A9;
     thus thesis
     proof per cases;
      suppose
A10:      len f = k;
       thus thesis
        proof per cases;
         suppose
A11:         f is_sequence_on G & right_cell(f,len f-'1,G) meets C;
A12:         1 <= (len f)-'1 by A8,A10,JORDAN3:12;
A13:         (len f) -'1 +1 = len f by A8,A10,AMI_5:4;
         then consider i1,j1,i2,j2 being Nat such that
A14:         [i1,j1] in Indices G & f/.((len f) -'1) = G*(i1,j1) and
A15:         [i2,j2] in Indices G & f/.((len f) -'1+1) = G*(i2,j2) and
A16:         i1 = i2 & j1+1 = j2 or i1+1 = i2 & j1 = j2 or
             i1 = i2+1 & j1 = j2 or i1 = i2 & j1 = j2+1 by A11,A12,JORDAN8:6;
A17:        1 <= i1 & i1 <= len G & 1 <= j1 & j1 <= width G by A14,GOBOARD5:1;
A18:        1 <= i2 & i2 <= len G & 1 <= j2 & j2 <= width G by A15,GOBOARD5:1;
A19:       i1-'1 <= len G & j1-'1 <= width G by A17,JORDAN3:7;
A20:        1 <= i2+1 & 1 <= j2+1 by NAT_1:37;
A21:       i2-'1 <= len G & j2-'1 <= width G by A18,JORDAN3:7;
             (len f)-'1 <= len f by GOBOARD9:2;
then A22:        (len f)-'1 in dom f & (len f)-'1+1 in dom f
              by A8,A10,A12,A13,FINSEQ_3:27;
A23:        (len f)-'1+(1+1) = (len f)+1 by A13,XCMPLX_1:1;
          thus thesis
          proof per cases;
           suppose
A24:          front_left_cell(f,(len f)-'1,G) misses C &
              front_right_cell(f,(len f)-'1,G) misses C;
            thus thesis
             proof per cases by A16;
              suppose
A25:              i1 = i2 & j1+1 = j2;
               take f1 = f^<*G*(i2+1,j2)*>;
                 now
                take i=i2+1 ,j=j2;
                thus f1 turns_right (len f)-'1,G
                 proof let i1',j1',i2',j2' be Nat;
                  assume that
A26:                 [i1',j1'] in Indices G & [i2',j2'] in Indices G and
A27:                 f1/.((len f)-'1) = G*(i1',j1') &
                    f1/.((len f)-'1+1) = G*(i2',j2');
                    f/.((len f)-'1) = G*(i1',j1') &
               f/.((len f)-'1+1) = G*(i2',j2') by A22,A27,GROUP_5:95;
then A28:                 i1 = i1' & j1 = j1' & i2 = i2' & j2 = j2'
                     by A14,A15,A26,GOBOARD1:21;
                  per cases by A16,A28;
                  case i1' = i2' & j1'+1 = j2';
                    now assume i2+1 > len G;
                   then i2+1 <= (len G)+1 & (len G)+1 <= i2+1
                    by A18,AXIOMS:24,NAT_1:38;
                   then i2+1 = (len G)+1 by AXIOMS:21;
                   then i2 = len G by XCMPLX_1:2;
                   then cell(G,len G,j1) meets C
                    by A11,A12,A13,A14,A15,A25,GOBRD13:23;
                   hence contradiction by A2,A17,JORDAN8:19;
                  end;
                  hence [i2'+1,j2'] in Indices G by A18,A20,A28,GOBOARD7:10;
                  thus thesis by A23,A28,TOPREAL4:1;
                 case i1'+1 = i2' & j1' = j2';
                  hence thesis by A25,A28,NAT_1:38;
                 case i1' = i2'+1 & j1' = j2';
                  hence thesis by A25,A28,NAT_1:38;
                 case
A29:               i1' = i2' & j1' = j2'+1;
                     j1+1+1 = j1+(1+1) by XCMPLX_1:1;
                  hence thesis by A25,A28,A29,REAL_1:69;
                 end;
               end;
               hence thesis by A8,A10,A11,A24;
              suppose
A30:             i1+1 = i2 & j1 = j2;
               take f1 = f^<*G*(i2,j2-'1)*>;
                 now
                take i=i2 ,j=j2-'1;
                thus f1 turns_right (len f)-'1,G
                 proof let i1',j1',i2',j2' be Nat;
                  assume that
A31:                 [i1',j1'] in Indices G & [i2',j2'] in Indices G and
A32:                 f1/.((len f)-'1) = G*(i1',j1') &
                    f1/.((len f)-'1+1) = G*(i2',j2');
                    f/.((len f)-'1) = G*(i1',j1') &
            f/.((len f)-'1+1) = G*(i2',j2') by A22,A32,GROUP_5:95;
then A33:                 i1 = i1' & j1 = j1' & i2 = i2' & j2 = j2'
                     by A14,A15,A31,GOBOARD1:21;
                  per cases by A16,A33;
                  case i1' = i2' & j1'+1 = j2';
                   hence thesis by A30,A33,NAT_1:38;
                  case i1'+1 = i2' & j1' = j2';
                    now assume j2-'1 < 1;
                   then j2-'1 = 0 by RLVECT_1:98;
                   then j2 <= 1 by JORDAN4:1;
                   then j2 = 1 by A18,AXIOMS:21;
                   then cell(G,i1,1-'1) meets C
                    by A11,A12,A13,A14,A15,A30,GOBRD13:25;
                   then cell(G,i1,0) meets C by GOBOARD9:1;
                   hence contradiction by A17,JORDAN8:20;
                  end;
                  hence [i2',j2'-'1] in Indices G by A18,A21,A33,GOBOARD7:10;
                  thus f1/.((len f)-'1+2) = G*(i2',j2'-'1)
                         by A23,A33,TOPREAL4:1;
                  case
A34:                i1' = i2'+1 & j1' = j2';
                     i1+1+1 = i1+(1+1) by XCMPLX_1:1;
                  hence thesis by A30,A33,A34,REAL_1:69;
                  case i1' = i2' & j1' = j2'+1;
                   hence thesis by A30,A33,NAT_1:38;
                 end;
               end;
               hence thesis by A8,A10,A11,A24;
              suppose
A35:             i1 = i2+1 & j1 = j2;
               take f1 = f^<*G*(i2,j2+1)*>;
                 now
                take i=i2 ,j=j2+1;
                thus f1 turns_right (len f)-'1,G
                 proof let i1',j1',i2',j2' be Nat;
                  assume that
A36:                 [i1',j1'] in Indices G & [i2',j2'] in Indices G and
A37:                 f1/.((len f)-'1) = G*(i1',j1') &
                    f1/.((len f)-'1+1) = G*(i2',j2');
                    f/.((len f)-'1) = G*(i1',j1') &
               f/.((len f)-'1+1) = G*(i2',j2') by A22,A37,GROUP_5:95;
then A38:                 i1 = i1' & j1 = j1' & i2 = i2' & j2 = j2'
                     by A14,A15,A36,GOBOARD1:21;
                  per cases by A16,A38;
                  case i1' = i2' & j1'+1 = j2';
                   hence thesis by A35,A38,NAT_1:38;
                  case
A39:                i1'+1 = i2' & j1' = j2';
                     i1+1+1 = i1+(1+1) by XCMPLX_1:1;
                  hence thesis by A35,A38,A39,REAL_1:69;
                  case i1' = i2'+1 & j1' = j2';
                    now assume j2+1 > len G;
                   then j2+1 <= (len G)+1 & (len G)+1 <= j2+1
                     by A2,A18,AXIOMS:24,NAT_1:38;
                   then j2+1 = (len G)+1 by AXIOMS:21;
                   then j2 = len G by XCMPLX_1:2;
                   then cell(G,i2,len G) meets C
                    by A11,A12,A13,A14,A15,A35,GOBRD13:27;
                   hence contradiction by A18,JORDAN8:18;
                  end;
                  hence [i2',j2'+1] in Indices G by A2,A18,A20,A38,GOBOARD7:10
;
                  thus f1/.((len f)-'1+2) = G*(i2',j2'+1) by A23,A38,TOPREAL4:1
;
                  case i1' = i2' & j1' = j2'+1;
                   hence thesis by A35,A38,NAT_1:38;
                 end;
               end;
               hence thesis by A8,A10,A11,A24;
              suppose
A40:             i1 = i2 & j1 = j2+1;
               take f1 = f^<*G*(i2-'1,j2)*>;
                 now
                take i=i2-'1 ,j=j2;
                thus f1 turns_right (len f)-'1,G
                 proof let i1',j1',i2',j2' be Nat;
                  assume that
A41:                 [i1',j1'] in Indices G & [i2',j2'] in Indices G and
A42:                 f1/.((len f)-'1) = G*(i1',j1') &
                    f1/.((len f)-'1+1) = G*(i2',j2');
                    f/.((len f)-'1) = G*(i1',j1') &
               f/.((len f)-'1+1) = G*(i2',j2') by A22,A42,GROUP_5:95;
then A43:                 i1 = i1' & j1 = j1' & i2 = i2' & j2 = j2'
                     by A14,A15,A41,GOBOARD1:21;
                  per cases by A16,A43;
                  case
A44:                i1' = i2' & j1'+1 = j2';
                     j1+1+1 = j1+(1+1) by XCMPLX_1:1;
                  hence thesis by A40,A43,A44,REAL_1:69;
                  case i1'+1 = i2' & j1' = j2';
                   hence thesis by A40,A43,NAT_1:38;
                  case i1' = i2'+1 & j1' = j2';
                   hence thesis by A40,A43,NAT_1:38;
                  case i1' = i2' & j1' = j2'+1;
                    now assume i2-'1 < 1;
                   then i2-'1 = 0 by RLVECT_1:98;
                   then i2 <= 1 by JORDAN4:1;
                   then i2 = 1 by A18,AXIOMS:21;
                   then cell(G,1-'1,j2) meets C
                    by A11,A12,A13,A14,A15,A40,GOBRD13:29;
                   then cell(G,0,j2) meets C by GOBOARD9:1;
                   hence contradiction by A2,A18,JORDAN8:21;
                  end;
                  hence [i2'-'1,j2'] in Indices G by A18,A21,A43,GOBOARD7:10;
                  thus f1/.((len f)-'1+2) = G*(i2'-'1,j2')
                              by A23,A43,TOPREAL4:1;
                 end;
               end;
               hence thesis by A8,A10,A11,A24;
             end;
           suppose
A45:          front_left_cell(f,(len f)-'1,G) misses C &
               front_right_cell(f,(len f)-'1,G) meets C;
            thus thesis
             proof per cases by A16;
              suppose
A46:              i1 = i2 & j1+1 = j2;
               take f1 = f^<*G*(i2,j2+1)*>;
                 now
                take i=i2 ,j=j2+1;
                thus f1 goes_straight (len f)-'1,G
                 proof let i1',j1',i2',j2' be Nat;
                  assume that
A47:                 [i1',j1'] in Indices G & [i2',j2'] in Indices G and
A48:                 f1/.((len f)-'1) = G*(i1',j1') &
                    f1/.((len f)-'1+1) = G*(i2',j2');
                    f/.((len f)-'1) = G*(i1',j1') &
             f/.((len f)-'1+1) = G*(i2',j2') by A22,A48,GROUP_5:95;
then A49:                 i1 = i1' & j1 = j1' & i2 = i2' & j2 = j2'
                     by A14,A15,A47,GOBOARD1:21;
                  per cases by A16,A49;
                  case i1' = i2' & j1'+1 = j2';
                    now assume j2+1 > len G;
                   then j2+1 <= (len G)+1 & (len G)+1 <= j2+1
                     by A2,A18,AXIOMS:24,NAT_1:38;
                   then j2+1 = (len G)+1 by AXIOMS:21;
                   then j2 = len G by XCMPLX_1:2;
                   then cell(G,i1,len G) meets C
                    by A11,A12,A13,A14,A15,A45,A46,GOBRD13:36;
                   hence contradiction by A17,JORDAN8:18;
                  end;
                  hence [i2',j2'+1] in Indices G by A2,A18,A20,A49,GOBOARD7:10
;
                  thus f1/.((len f)-'1+2) = G*(i2',j2'+1)
                            by A23,A49,TOPREAL4:1;
                  case i1'+1 = i2' & j1' = j2';
                   hence thesis by A46,A49,NAT_1:38;
                  case i1' = i2'+1 & j1' = j2';
                   hence thesis by A46,A49,NAT_1:38;
                  case
A50:                i1' = i2' & j1' = j2'+1;
                     j1+1+1 = j1+(1+1) by XCMPLX_1:1;
                  hence thesis by A46,A49,A50,REAL_1:69;
                 end;
               end;
               hence thesis by A8,A10,A11,A45;
              suppose
A51:             i1+1 = i2 & j1 = j2;
               take f1 = f^<*G*(i2+1,j2)*>;
                 now
                take i=i2+1 ,j=j2;
                thus f1 goes_straight (len f)-'1,G
                 proof let i1',j1',i2',j2' be Nat;
                  assume that
A52:                 [i1',j1'] in Indices G & [i2',j2'] in Indices G and
A53:                 f1/.((len f)-'1) = G*(i1',j1') &
                    f1/.((len f)-'1+1) = G*(i2',j2');
                    f/.((len f)-'1) = G*(i1',j1') &
                 f/.((len f)-'1+1) = G*(i2',j2') by A22,A53,GROUP_5:95;
then A54:                 i1 = i1' & j1 = j1' & i2 = i2' & j2 = j2'
                     by A14,A15,A52,GOBOARD1:21;
                  per cases by A16,A54;
                  case i1' = i2' & j1'+1 = j2';
                   hence thesis by A51,A54,NAT_1:38;
                  case i1'+1 = i2' & j1' = j2';
                    now assume i2+1 > len G;
                   then i2+1 <= (len G)+1 & (len G)+1 <= i2+1
                     by A18,AXIOMS:24,NAT_1:38;
                   then i2+1 = (len G)+1 by AXIOMS:21;
                   then i2 = len G by XCMPLX_1:2;
                   then cell(G,len G,j1-'1) meets C
                     by A11,A12,A13,A14,A15,A45,A51,GOBRD13:38;
                   hence contradiction by A2,A19,JORDAN8:19;
                  end;
                  hence [i2'+1,j2'] in Indices G by A18,A20,A54,GOBOARD7:10;
                  thus f1/.((len f)-'1+2) = G*(i2'+1,j2')
                           by A23,A54,TOPREAL4:1;
                  case
A55:                i1' = i2'+1 & j1' = j2';
                     i1+1+1 = i1+(1+1) by XCMPLX_1:1;
                  hence thesis by A51,A54,A55,REAL_1:69;
                  case i1' = i2' & j1' = j2'+1;
                   hence thesis by A51,A54,NAT_1:38;
                 end;
               end;
               hence thesis by A8,A10,A11,A45;
              suppose
A56:             i1 = i2+1 & j1 = j2;
               take f1 = f^<*G*(i2-'1,j2)*>;
                 now
                take i=i2-'1 ,j=j2;
                thus f1 goes_straight (len f)-'1,G
                 proof let i1',j1',i2',j2' be Nat;
                  assume that
A57:                 [i1',j1'] in Indices G & [i2',j2'] in Indices G and
A58:                 f1/.((len f)-'1) = G*(i1',j1') &
                    f1/.((len f)-'1+1) = G*(i2',j2');
                    f/.((len f)-'1) = G*(i1',j1') &
                f/.((len f)-'1+1) = G*(i2',j2') by A22,A58,GROUP_5:95;
then A59:                 i1 = i1' & j1 = j1' & i2 = i2' & j2 = j2'
                     by A14,A15,A57,GOBOARD1:21;
                  per cases by A16,A59;
                  case i1' = i2' & j1'+1 = j2';
                   hence thesis by A56,A59,NAT_1:38;
                  case
A60:                i1'+1 = i2' & j1' = j2';
                     i1+1+1 = i1+(1+1) by XCMPLX_1:1;
                  hence thesis by A56,A59,A60,REAL_1:69;
                  case i1' = i2'+1 & j1' = j2';
                    now assume i2-'1 < 1;
                   then i2-'1 = 0 by RLVECT_1:98;
                   then i2 <= 1 by JORDAN4:1;
                   then i2 = 1 by A18,AXIOMS:21;
                   then cell(G,1-'1,j1) meets C
                    by A11,A12,A13,A14,A15,A45,A56,GOBRD13:40;
                   then cell(G,0,j1) meets C by GOBOARD9:1;
                   hence contradiction by A2,A17,JORDAN8:21;
                  end;
                  hence [i2'-'1,j2'] in Indices G by A18,A21,A59,GOBOARD7:10;
                  thus f1/.((len f)-'1+2) = G*(i2'-'1,j2')
                            by A23,A59,TOPREAL4:1;
                  case i1' = i2' & j1' = j2'+1;
                   hence thesis by A56,A59,NAT_1:38;
                 end;
               end;
               hence thesis by A8,A10,A11,A45;
              suppose
A61:             i1 = i2 & j1 = j2+1;
               take f1 = f^<*G*(i2,j2-'1)*>;
                 now
                take i=i2 ,j=j2-'1;
                thus f1 goes_straight (len f)-'1,G
                 proof let i1',j1',i2',j2' be Nat;
                  assume that
A62:                 [i1',j1'] in Indices G & [i2',j2'] in Indices G and
A63:                 f1/.((len f)-'1) = G*(i1',j1') &
                    f1/.((len f)-'1+1) = G*(i2',j2');
                    f/.((len f)-'1) = G*(i1',j1') &
                f/.((len f)-'1+1) = G*(i2',j2') by A22,A63,GROUP_5:95;
then A64:                 i1 = i1' & j1 = j1' & i2 = i2' & j2 = j2'
                     by A14,A15,A62,GOBOARD1:21;
                  per cases by A16,A64;
                  case
A65:                i1' = i2' & j1'+1 = j2';
                     j1+1+1 = j1+(1+1) by XCMPLX_1:1;
                  hence thesis by A61,A64,A65,REAL_1:69;
                  case i1'+1 = i2' & j1' = j2';
                   hence thesis by A61,A64,NAT_1:38;
                  case i1' = i2'+1 & j1' = j2';
                   hence thesis by A61,A64,NAT_1:38;
                  case i1' = i2' & j1' = j2'+1;
                    now assume j2-'1 < 1;
                   then j2-'1 = 0 by RLVECT_1:98;
                   then j2 <= 1 by JORDAN4:1;
                   then j2 = 1 by A18,AXIOMS:21;
                   then cell(G,i1-'1,1-'1) meets C
                     by A11,A12,A13,A14,A15,A45,A61,GOBRD13:42;
                   then cell(G,i1-'1,0) meets C by GOBOARD9:1;
                   hence contradiction by A19,JORDAN8:20;
                  end;
                  hence [i2',j2'-'1] in Indices G by A18,A21,A64,GOBOARD7:10;
                  thus f1/.((len f)-'1+2) = G*(i2',j2'-'1)
                            by A23,A64,TOPREAL4:1;
                 end;
               end;
               hence thesis by A8,A10,A11,A45;
             end;
           suppose
A66:          front_left_cell(f,(len f)-'1,G) meets C;
            thus thesis
             proof per cases by A16;
              suppose
A67:              i1 = i2 & j1+1 = j2;
               take f1 = f^<*G*(i2-'1,j2)*>;
                 now
                take i=i2-'1 ,j=j2;
                thus f1 turns_left (len f)-'1,G
                 proof let i1',j1',i2',j2' be Nat;
                  assume that
A68:                 [i1',j1'] in Indices G & [i2',j2'] in Indices G and
A69:                 f1/.((len f)-'1) = G*(i1',j1') &
                    f1/.((len f)-'1+1) = G*(i2',j2');
                    f/.((len f)-'1) = G*(i1',j1') &
               f/.((len f)-'1+1) = G*(i2',j2') by A22,A69,GROUP_5:95;
then A70:                 i1 = i1' & j1 = j1' & i2 = i2' & j2 = j2'
                     by A14,A15,A68,GOBOARD1:21;
                  per cases by A16,A70;
                  case i1' = i2' & j1'+1 = j2';
                    now assume i2-'1 < 1;
                   then i2-'1 = 0 by RLVECT_1:98;
                   then i2 <= 1 by JORDAN4:1;
                   then i2 = 1 by A18,AXIOMS:21;
                   then cell(G,1-'1,j2) meets C
                    by A11,A12,A13,A14,A15,A66,A67,GOBRD13:35;
                   then cell(G,0,j2) meets C by GOBOARD9:1;
                   hence contradiction by A2,A18,JORDAN8:21;
                  end;
                  hence [i2'-'1,j2'] in Indices G by A18,A21,A70,GOBOARD7:10;
                  thus f1/.((len f)-'1+2) = G*(i2'-'1,j2')
                            by A23,A70,TOPREAL4:1;
                  case i1'+1 = i2' & j1' = j2';
                   hence thesis by A67,A70,NAT_1:38;
                  case i1' = i2'+1 & j1' = j2';
                   hence thesis by A67,A70,NAT_1:38;
                  case
A71:                i1' = i2' & j1' = j2'+1;
                     j1+1+1 = j1+(1+1) by XCMPLX_1:1;
                  hence thesis by A67,A70,A71,REAL_1:69;
                 end;
               end;
               hence thesis by A8,A10,A11,A66;
              suppose
A72:             i1+1 = i2 & j1 = j2;
               take f1 = f^<*G*(i2,j2+1)*>;
                 now
                take i=i2 ,j=j2+1;
                thus f1 turns_left (len f)-'1,G
                 proof let i1',j1',i2',j2' be Nat;
                  assume that
A73:                 [i1',j1'] in Indices G & [i2',j2'] in Indices G and
A74:                 f1/.((len f)-'1) = G*(i1',j1') &
                    f1/.((len f)-'1+1) = G*(i2',j2');
                    f/.((len f)-'1) = G*(i1',j1') &
                f/.((len f)-'1+1) = G*(i2',j2') by A22,A74,GROUP_5:95;
then A75:                 i1 = i1' & j1 = j1' & i2 = i2' & j2 = j2'
                     by A14,A15,A73,GOBOARD1:21;
                  per cases by A16,A75;
                  case i1' = i2' & j1'+1 = j2';
                   hence thesis by A72,A75,NAT_1:38;
                  case i1'+1 = i2' & j1' = j2';
                    now assume j2+1 > len G;
                   then j2+1 <= (len G)+1 & (len G)+1 <= j2+1
                      by A2,A18,AXIOMS:24,NAT_1:38;
                   then j2+1 = (len G)+1 by AXIOMS:21;
                   then j2 = len G by XCMPLX_1:2;
                   then cell(G,i2,len G) meets C
                     by A11,A12,A13,A14,A15,A66,A72,GOBRD13:37;
                   hence contradiction by A18,JORDAN8:18;
                  end;
                  hence [i2',j2'+1] in Indices G by A2,A18,A20,A75,GOBOARD7:10
;
                  thus f1/.((len f)-'1+2) = G*(i2',j2'+1)
                            by A23,A75,TOPREAL4:1;
                  case
A76:                i1' = i2'+1 & j1' = j2';
                     i1+1+1 = i1+(1+1) by XCMPLX_1:1;
                  hence thesis by A72,A75,A76,REAL_1:69;
                  case i1' = i2' & j1' = j2'+1;
                   hence thesis by A72,A75,NAT_1:38;
                 end;
               end;
               hence thesis by A8,A10,A11,A66;
              suppose
A77:             i1 = i2+1 & j1 = j2;
               take f1 = f^<*G*(i2,j2-'1)*>;
                 now
                take i=i2 ,j=j2-'1;
                thus f1 turns_left (len f)-'1,G
                 proof let i1',j1',i2',j2' be Nat;
                  assume that
A78:                 [i1',j1'] in Indices G & [i2',j2'] in Indices G and
A79:                 f1/.((len f)-'1) = G*(i1',j1') &
                    f1/.((len f)-'1+1) = G*(i2',j2');
                    f/.((len f)-'1) = G*(i1',j1') &
               f/.((len f)-'1+1) = G*(i2',j2') by A22,A79,GROUP_5:95;
then A80:                 i1 = i1' & j1 = j1' & i2 = i2' & j2 = j2'
                     by A14,A15,A78,GOBOARD1:21;
                  per cases by A16,A80;
                  case i1' = i2' & j1'+1 = j2';
                   hence thesis by A77,A80,NAT_1:38;
                  case
A81:                i1'+1 = i2' & j1' = j2';
                     i1+1+1 = i1+(1+1) by XCMPLX_1:1;
                  hence thesis by A77,A80,A81,REAL_1:69;
                  case i1' = i2'+1 & j1' = j2';
                    now assume j2-'1 < 1;
                   then j2-'1 = 0 by RLVECT_1:98;
                   then j2 <= 1 by JORDAN4:1;
                   then j2 = 1 by A18,AXIOMS:21;
                   then cell(G,i2-'1,1-'1) meets C
                     by A11,A12,A13,A14,A15,A66,A77,GOBRD13:39;
                   then cell(G,i2-'1,0) meets C by GOBOARD9:1;
                   hence contradiction by A21,JORDAN8:20;
                  end;
                  hence [i2',j2'-'1] in Indices G by A18,A21,A80,GOBOARD7:10;
                  thus f1/.((len f)-'1+2) = G*(i2',j2'-'1)
                            by A23,A80,TOPREAL4:1;
                  case i1' = i2' & j1' = j2'+1;
                   hence thesis by A77,A80,NAT_1:38;
                 end;
               end;
               hence thesis by A8,A10,A11,A66;
              suppose
A82:             i1 = i2 & j1 = j2+1;
               take f1 = f^<*G*(i2+1,j2)*>;
                 now
                take i=i2+1 ,j=j2;
                thus f1 turns_left (len f)-'1,G
                 proof let i1',j1',i2',j2' be Nat;
                  assume that
A83:                 [i1',j1'] in Indices G & [i2',j2'] in Indices G and
A84:                 f1/.((len f)-'1) = G*(i1',j1') &
                    f1/.((len f)-'1+1) = G*(i2',j2');
                    f/.((len f)-'1) = G*(i1',j1') &
                f/.((len f)-'1+1) = G*(i2',j2') by A22,A84,GROUP_5:95;
then A85:                 i1 = i1' & j1 = j1' & i2 = i2' & j2 = j2'
                     by A14,A15,A83,GOBOARD1:21;
                  per cases by A16,A85;
                  case
A86:                  i1' = i2' & j1'+1 = j2';
                     j1+1+1 = j1+(1+1) by XCMPLX_1:1;
                  hence thesis by A82,A85,A86,REAL_1:69;
                  case i1'+1 = i2' & j1' = j2';
                   hence thesis by A82,A85,NAT_1:38;
                  case i1' = i2'+1 & j1' = j2';
                   hence thesis by A82,A85,NAT_1:38;
                  case i1' = i2' & j1' = j2'+1;
                    now assume i2+1 > len G;
                   then i2+1 <= (len G)+1 & (len G)+1 <= i2+1
                    by A18,AXIOMS:24,NAT_1:38;
                   then i2+1 = (len G)+1 by AXIOMS:21;
                   then i2 = len G by XCMPLX_1:2;
                   then cell(G,len G,j2-'1) meets C
                     by A11,A12,A13,A14,A15,A66,A82,GOBRD13:41;
                   hence contradiction by A2,A21,JORDAN8:19;
                  end;
                  hence [i2'+1,j2'] in Indices G by A18,A20,A85,GOBOARD7:10;
                  thus f1/.((len f)-'1+2) = G*(i2'+1,j2')
                            by A23,A85,TOPREAL4:1;
                 end;
               end;
               hence thesis by A8,A10,A11,A66;
             end;
          end;
         suppose
A87:         not f is_sequence_on G or right_cell(f,len f-'1,G) misses C;
          take f^<*G*(1,1)*>;
          thus thesis by A8,A10,A87;
       end;
      suppose
A88:      len f <> k;
       take {};
       thus thesis by A8,A88;
     end;
    suppose
A89:    k > 1 & x is not FinSequence of TOP-REAL 2;
     take {};
     thus thesis by A89;
   end;
  consider F being Function such that
A90:   dom F = NAT and
A91:   F.0 = {} and
A92:   for k being Element of NAT holds P[k,F.k,F.(k+1)] from RecChoice(A3);
defpred P[Nat] means F.$1 is FinSequence of TOP-REAL 2;
A93:   {} = <*>(the carrier of TOP-REAL 2);
then A94:  P[0] by A91;
A95:  for k st P[k] holds P[k+1]
   proof let k such that
A96:   F.k is FinSequence of TOP-REAL 2;
A97:   P[k,F.k,F.(k+1)] by A92;
    per cases by CQC_THE1:2;
     suppose k = 0;
      hence F.(k+1) is FinSequence of TOP-REAL 2 by A97;
     suppose k = 1;
      hence F.(k+1) is FinSequence of TOP-REAL 2 by A97;
     suppose
A98:     k > 1;
      thus thesis
      proof
        reconsider f = F.k as FinSequence of TOP-REAL 2 by A96;
       per cases;
        suppose
A99:        len f = k;
          thus thesis
          proof per cases;
           suppose f is_sequence_on G & right_cell(f,len f-'1,G) meets C
;
            then (front_left_cell(f,(len f)-'1,G) misses C &
              front_right_cell(f,(len f)-'1,G) misses C
             implies ex i,j st
                 f^<*G*(i,j)*> turns_right (len f)-'1,G &
               F.(k+1) = f^<*G*(i,j)*>) &
            (front_left_cell(f,(len f)-'1,G) misses C &
              front_right_cell(f,(len f)-'1,G) meets C
             implies ex i,j st
                 f^<*G*(i,j)*> goes_straight (len f)-'1,G &
               F.(k+1) = f^<*G*(i,j)*>) &
            (front_left_cell(f,(len f)-'1,G) meets C
             implies ex i,j st
                 f^<*G*(i,j)*> turns_left (len f)-'1,G &
               F.(k+1) = f^<*G*(i,j)*>) by A92,A98,A99;
            hence F.(k+1) is FinSequence of TOP-REAL 2;
           suppose
A100:           not f is_sequence_on G or right_cell(f,len f-'1,G) misses C;
               f^<*G*(1,1)*> is FinSequence of TOP-REAL 2;
            hence F.(k+1) is FinSequence of TOP-REAL 2
             by A92,A98,A99,A100;
          end;
        suppose len f <> k;
         hence F.(k+1) is FinSequence of TOP-REAL 2 by A92,A93,A98;
      end;
   end;

A101:  for k holds P[k] from Ind(A94,A95);
     rng F c= (the carrier of TOP-REAL 2)*
    proof let y be set;
     assume y in rng F;
     then ex x being set st x in dom F & F.x = y by FUNCT_1:def 5;
     then y is FinSequence of TOP-REAL 2 by A90,A101;
     hence thesis by FINSEQ_1:def 11;
    end;
  then reconsider F as Function of NAT,(the carrier of TOP-REAL 2)*
   by A90,FUNCT_2:def 1,RELSET_1:11;
   defpred P[Nat] means len(F.$1) = $1;
A102:  P[0] by A91,FINSEQ_1:25;
A103:  for k st P[k] holds P[k+1]
    proof let k such that
A104:   len(F.k) = k;
A105:   P[k,F.k,F.(k+1)] by A92;
     per cases by CQC_THE1:2;
     suppose k = 0;
      hence thesis by A105,FINSEQ_1:56;
     suppose k = 1;
      hence thesis by A105,FINSEQ_1:61;
     suppose
A106:     k > 1;
      thus thesis
      proof per cases;
       suppose
       F.k is_sequence_on G & right_cell(F.k,len(F.k)-'1,G) meets C;
            then (front_left_cell(F.k,(len(F.k))-'1,G) misses C &
              front_right_cell(F.k,(len(F.k))-'1,G) misses C
             implies ex i,j st
                 (F.k)^<*G*(i,j)*> turns_right (len(F.k))-'1,G &
               F.(k+1) = (F.k)^<*G*(i,j)*>) &
            (front_left_cell(F.k,(len(F.k))-'1,G) misses C &
              front_right_cell(F.k,(len(F.k))-'1,G) meets C
             implies ex i,j st
                 (F.k)^<*G*(i,j)*> goes_straight (len(F.k))-'1,G &
               F.(k+1) = (F.k)^<*G*(i,j)*>) &
            (front_left_cell(F.k,(len(F.k))-'1,G) meets C
             implies ex i,j st
                 (F.k)^<*G*(i,j)*> turns_left (len(F.k))-'1,G &
               F.(k+1) = (F.k)^<*G*(i,j)*>) by A92,A104,A106;
            hence thesis by A104,FINSEQ_2:19;
       suppose
          not F.k is_sequence_on G or right_cell(F.k,len(F.k)-'1,G) misses C;
        then F.(k+1) = (F.k)^<*G*(1,1)*> by A92,A104,A106;
        hence thesis by A104,FINSEQ_2:19;
      end;
    end;

A107:   for k holds P[k] from Ind(A102,A103);
defpred P[Nat] means
F.$1 is_sequence_on G &
      for m st 1 <= m & m+1 <= len(F.$1)
       holds left_cell(F.$1,m,G) misses C & right_cell(F.$1,m,G) meets C;
A108:
  P[0] proof
      (for n st n in dom(F.0) ex i,j st [i,j] in Indices G &
      (F.0)/.n = G*(i,j)) &
    (for n st n in dom(F.0) & n+1 in dom(F.0) holds
      for m,k,i,j st [m,k] in Indices G & [i,j] in Indices G &
                    (F.0)/.n = G*(m,k) & (F.0)/.(n+1) = G*(i,j) holds
      abs(m-i)+abs(k-j) = 1) by A91,RELAT_1:60;
   hence F.0 is_sequence_on G by GOBOARD1:def 11;
   let m; assume
A109:  1 <= m & m+1 <= len(F.0);
     1 <= m+1 by NAT_1:37;
   then 1 <= len(F.0) by A109,AXIOMS:22;
   then 1 <= 0 by A91,FINSEQ_1:25;
   hence left_cell(F.0,m,G) misses C &
           right_cell(F.0,m,G) meets C;
  end;
A110: now let k such that
A111:   F.k is_sequence_on G and
A112:   for m st 1 <= m & m+1 <= len(F.k)
        holds left_cell(F.k,m,G) misses C &
            right_cell(F.k,m,G) meets C and
A113:   k > 1;
     let i1,j1,i2,j2 be Nat such that
A114:   [i1,j1] in Indices G & (F.k)/.((len(F.k)) -'1) = G*(i1,j1) and
A115:   [i2,j2] in Indices G & (F.k)/.len(F.k) = G*(i2,j2);
A116:   len(F.k) = k by A107;
then A117:   1 <= (len(F.k))-'1 by A113,JORDAN3:12;
A118:   (len(F.k)) -'1 +1 = len(F.k) by A113,A116,AMI_5:4;
then A119:  left_cell(F.k,(len(F.k))-'1,G) misses C &
       right_cell(F.k,(len(F.k))-'1,G) meets C by A112,A117;
A120:   1 <= i1 & i1 <= len G & 1 <= j1 & j1 <= width G by A114,GOBOARD5:1;
A121:   1 <= i2 & i2 <= len G & 1 <= j2 & j2 <= width G by A115,GOBOARD5:1;
A122:  i1-'1 <= len G & j1-'1 <= width G by A120,JORDAN3:7;
A123:   1 <= i2+1 & 1 <= j2+1 by NAT_1:37;
A124:  i2-'1 <= len G & j2-'1 <= width G by A121,JORDAN3:7;
     hereby assume
A125:     i1 = i2 & j1+1 = j2;
         now assume i2+1 > len G;
        then i2+1 <= (len G)+1 & (len G)+1 <= i2+1 by A121,AXIOMS:24,NAT_1:38;
        then i2+1 = (len G)+1 by AXIOMS:21;
        then i2 = len G by XCMPLX_1:2;
        then cell(G,len G,j1) meets C
          by A111,A114,A115,A117,A118,A119,A125,GOBRD13:23;
        hence contradiction by A2,A120,JORDAN8:19;
       end;
      hence [i2+1,j2] in Indices G by A121,A123,GOBOARD7:10;
     end;
     hereby assume
A126:     i1+1 = i2 & j1 = j2;
         now assume j2-'1 < 1;
        then j2-'1 = 0 by RLVECT_1:98;
        then j2 <= 1 by JORDAN4:1;
        then j2 = 1 by A121,AXIOMS:21;
        then cell(G,i1,1-'1) meets C by A111,A114,A115,A117,A118,A119,A126,
GOBRD13:25;
        then cell(G,i1,0) meets C by GOBOARD9:1;
        hence contradiction by A120,JORDAN8:20;
       end;
      hence [i2,j2-'1] in Indices G by A121,A124,GOBOARD7:10;
     end;
     hereby assume
A127:     i1 = i2+1 & j1 = j2;
         now assume j2+1 > len G;
        then j2+1 <= (len G)+1 & (len G)+1 <=
 j2+1 by A2,A121,AXIOMS:24,NAT_1:38;
        then j2+1 = (len G)+1 by AXIOMS:21;
        then j2 = len G by XCMPLX_1:2;
        then cell(G,i2,len G) meets C by A111,A114,A115,A117,A118,A119,A127,
GOBRD13:27;
        hence contradiction by A121,JORDAN8:18;
       end;
      hence [i2,j2+1] in Indices G by A2,A121,A123,GOBOARD7:10;
     end;
     hereby assume
A128:     i1 = i2 & j1 = j2+1;
         now assume i2-'1 < 1;
        then i2-'1 = 0 by RLVECT_1:98;
        then i2 <= 1 by JORDAN4:1;
        then i2 = 1 by A121,AXIOMS:21;
        then cell(G,1-'1,j2) meets C by A111,A114,A115,A117,A118,A119,A128,
GOBRD13:29;
        then cell(G,0,j2) meets C by GOBOARD9:1;
        hence contradiction by A2,A121,JORDAN8:21;
       end;
      hence [i2-'1,j2] in Indices G by A121,A124,GOBOARD7:10;
     end;
     hereby assume that
A129:     front_right_cell(F.k,(len(F.k))-'1,G) meets C and
A130:     i1 = i2 & j1+1 = j2;
         now assume j2+1 > len G;
        then j2+1 <= (len G)+1 & (len G)+1 <=
 j2+1 by A2,A121,AXIOMS:24,NAT_1:38;
        then j2+1 = (len G)+1 by AXIOMS:21;
        then j2 = len G by XCMPLX_1:2;
        then cell(G,i1,len G) meets C by A111,A114,A115,A117,A118,A129,A130,
GOBRD13:36;
        hence contradiction by A120,JORDAN8:18;
       end;
      hence [i2,j2+1] in Indices G by A2,A121,A123,GOBOARD7:10;
     end;
     hereby assume that
A131:     front_right_cell(F.k,(len(F.k))-'1,G) meets C and
A132:     i1+1 = i2 & j1 = j2;
         now assume i2+1 > len G;
        then i2+1 <= (len G)+1 & (len G)+1 <= i2+1 by A121,AXIOMS:24,NAT_1:38;
        then i2+1 = (len G)+1 by AXIOMS:21;
        then i2 = len G by XCMPLX_1:2;
        then cell(G,len G,j1-'1) meets C
         by A111,A114,A115,A117,A118,A131,A132,GOBRD13:38;
        hence contradiction by A2,A122,JORDAN8:19;
       end;
      hence [i2+1,j2] in Indices G by A121,A123,GOBOARD7:10;
     end;
     hereby assume that
A133:     front_right_cell(F.k,(len(F.k))-'1,G) meets C and
A134:     i1 = i2+1 & j1 = j2;
         now assume i2-'1 < 1;
        then i2-'1 = 0 by RLVECT_1:98;
        then i2 <= 1 by JORDAN4:1;
        then i2 = 1 by A121,AXIOMS:21;
        then cell(G,1-'1,j1) meets C by A111,A114,A115,A117,A118,A133,A134,
GOBRD13:40;
        then cell(G,0,j1) meets C by GOBOARD9:1;
        hence contradiction by A2,A120,JORDAN8:21;
       end;
      hence [i2-'1,j2] in Indices G by A121,A124,GOBOARD7:10;
     end;
     hereby assume that
A135:     front_right_cell(F.k,(len(F.k))-'1,G) meets C and
A136:     i1 = i2 & j1 = j2+1;
         now assume j2-'1 < 1;
        then j2-'1 = 0 by RLVECT_1:98;
        then j2 <= 1 by JORDAN4:1;
        then j2 = 1 by A121,AXIOMS:21;
        then cell(G,i1-'1,1-'1) meets C
           by A111,A114,A115,A117,A118,A135,A136,GOBRD13:42;
        then cell(G,i1-'1,0) meets C by GOBOARD9:1;
        hence contradiction by A122,JORDAN8:20;
       end;
      hence [i2,j2-'1] in Indices G by A121,A124,GOBOARD7:10;
     end;
     hereby assume that
A137:     front_left_cell(F.k,(len(F.k))-'1,G) meets C and
A138:     i1 = i2 & j1+1 = j2;
         now assume i2-'1 < 1;
        then i2-'1 = 0 by RLVECT_1:98;
        then i2 <= 1 by JORDAN4:1;
        then i2 = 1 by A121,AXIOMS:21;
        then cell(G,1-'1,j2) meets C by A111,A114,A115,A117,A118,A137,A138,
GOBRD13:35;
        then cell(G,0,j2) meets C by GOBOARD9:1;
        hence contradiction by A2,A121,JORDAN8:21;
       end;
      hence [i2-'1,j2] in Indices G by A121,A124,GOBOARD7:10;
     end;
     hereby assume that
A139:     front_left_cell(F.k,(len(F.k))-'1,G) meets C and
A140:     i1+1 = i2 & j1 = j2;
         now assume j2+1 > len G;
        then j2+1 <= (len G)+1 & (len G)+1 <=
 j2+1 by A2,A121,AXIOMS:24,NAT_1:38;
        then j2+1 = (len G)+1 by AXIOMS:21;
        then j2 = len G by XCMPLX_1:2;
        then cell(G,i2,len G) meets C by A111,A114,A115,A117,A118,A139,A140,
GOBRD13:37;
        hence contradiction by A121,JORDAN8:18;
       end;
      hence [i2,j2+1] in Indices G by A2,A121,A123,GOBOARD7:10;
     end;
     hereby assume that
A141:     front_left_cell(F.k,(len(F.k))-'1,G) meets C and
A142:     i1 = i2+1 & j1 = j2;
         now assume j2-'1 < 1;
        then j2-'1 = 0 by RLVECT_1:98;
        then j2 <= 1 by JORDAN4:1;
        then j2 = 1 by A121,AXIOMS:21;
        then cell(G,i2-'1,1-'1) meets C
           by A111,A114,A115,A117,A118,A141,A142,GOBRD13:39;
        then cell(G,i2-'1,0) meets C by GOBOARD9:1;
        hence contradiction by A124,JORDAN8:20;
       end;
      hence [i2,j2-'1] in Indices G by A121,A124,GOBOARD7:10;
     end;
     hereby assume that
A143:     front_left_cell(F.k,(len(F.k))-'1,G) meets C and
A144:     i1 = i2 & j1 = j2+1;
         now assume i2+1 > len G;
        then i2+1 <= (len G)+1 & (len G)+1 <= i2+1 by A121,AXIOMS:24,NAT_1:38;
        then i2+1 = (len G)+1 by AXIOMS:21;
        then i2 = len G by XCMPLX_1:2;
        then cell(G,len G,j2-'1) meets C
         by A111,A114,A115,A117,A118,A143,A144,GOBRD13:41;
        hence contradiction by A2,A124,JORDAN8:19;
       end;
      hence [i2+1,j2] in Indices G by A121,A123,GOBOARD7:10;
     end;
    end;
A145: for k st P[k] holds P[k+1]
   proof let k such that
A146:   F.k is_sequence_on G and
A147:   for m st 1 <= m & m+1 <= len(F.k)
        holds left_cell(F.k,m,G) misses C & right_cell(F.k,m,G) meets C;
A148:   len(F.k) = k by A107;
A149:   len(F.(k+1)) = k+1 by A107;
A150:   1 <= len G by A2,NAT_1:37;
A151:   2|^n > 0 by HEINE:5;
    per cases by CQC_THE1:2;
    suppose
A152:   k = 0; then consider i such that
A153:   1 <= i & i+1 <= len G and
          N-min C in cell(G,i,width G-'1) & N-min C <> G*(i,width G-'1) and
A154:   F.(k+1) = <*G*(i,width G)*> by A92;
        i < len G by A153,NAT_1:38;
then A155:   [i,len G] in Indices G by A2,A150,A153,GOBOARD7:10;
A156:   now let l;
       assume l in dom(F.(k+1));
       then 1 <= l & l <= 1 by A149,A152,FINSEQ_3:27;
       then l = 1 by AXIOMS:21;
       then (F.(k+1))/.l = G*(i,width G) by A154,FINSEQ_4:25;
       hence ex i,j st [i,j] in Indices G & (F.(k+1))/.l = G*
(i,j) by A2,A155;
      end;
       now let l;
      assume l in dom(F.(k+1)) & l+1 in dom(F.(k+1));
      then 1 <= l & l <= 1 & 1 <= l+1 & l+1 <= 1 by A149,A152,FINSEQ_3:27;
      then l = 1 & 1 = l+1 by AXIOMS:21;
      hence for i1,j1,i2,j2 st [i1,j1] in Indices G & [i2,j2] in Indices G &
                    (F.(k+1))/.l = G*(i1,j1) & (F.(k+1))/.(l+1) = G*(i2,j2)
       holds abs(i1-i2)+abs(j1-j2) = 1;
     end;
     hence F.(k+1) is_sequence_on G by A156,GOBOARD1:def 11;
     let m;
     assume
A157:   1 <= m & m+1 <= len(F.(k+1));
       1 <= m+1 by NAT_1:37;
     then m+1 = 0+1 by A149,A152,A157,AXIOMS:21;
     then m = 0 by XCMPLX_1:2;
     hence left_cell(F.(k+1),m,G) misses C &
              right_cell(F.(k+1),m,G) meets C by A157;
    suppose
A158:   k = 1; then consider i such that
A159:   1 <= i & i+1 <= len G and
A160:   N-min C in cell(G,i,width G-'1) & N-min C <> G*(i,width G-'1) and
A161:   F.(k+1) = <*G*(i,width G),G*(i+1,width G)*> by A92;
A162:   (F.(k+1))/.1 = G*(i,width G) & (F.(k+1))/.2 = G*(i+1,width G)
        by A161,FINSEQ_4:26;
A163:  i < len G & 1 <= i+1 & i+1 <= len G by A159,NAT_1:38;
then A164:   [i,len G] in Indices G & [i+1,len G] in Indices G
        by A2,A150,A159,GOBOARD7:10;
A165:   now let l;
       assume A166: l in dom(F.(k+1));
       then 1 <= l & l <= 2 by A149,A158,FINSEQ_3:27;
       then l = 0 or l = 1 or l = 2 by CQC_THE1:3;
       hence ex i,j st [i,j] in Indices G & (F.(k+1))/.l = G*(i,j)
         by A2,A162,A164,A166,FINSEQ_3:27;
      end;
       now let l;
      assume A167: l in dom(F.(k+1)) & l+1 in dom(F.(k+1));
      then 1 <= l & l <= 2 & 1 <= l+1 & l+1 <= 2 by A149,A158,FINSEQ_3:27;
      then A168: l = 0 or l = 1 or l = 2 by CQC_THE1:3;
      let i1,j1,i2,j2 such that
A169:      [i1,j1] in Indices G & [i2,j2] in Indices G &
        (F.(k+1))/.l = G*(i1,j1) & (F.(k+1))/.(l+1) = G*(i2,j2);
         i1 = i & j1 = len G & i2 = i+1 & j2 = len G
         by A2,A149,A158,A162,A164,A167,A168,A169,FINSEQ_3:27,GOBOARD1:21;
      then i2-i1 = 1 & j1-j2 = 0 by XCMPLX_1:14,26;
      then abs(i2-i1) = 1 & abs(j1-j2) = 0 by ABSVALUE:def 1;
      hence abs(i1-i2)+abs(j1-j2) = 1 by UNIFORM1:13;
     end;
     hence
A170:    F.(k+1) is_sequence_on G by A165,GOBOARD1:def 11;
     let m;
     assume
A171:   1 <= m & m+1 <= len(F.(k+1));
     then 1+1 <= m+1 by AXIOMS:24;
     then m+1 = 1+1 by A149,A158,A171,AXIOMS:21;
then A172:  m = 1 by XCMPLX_1:2;
A173:  i < i+1 & i+1 < (i+1)+1 by NAT_1:38;
then A174:   left_cell(F.(k+1),m,G) = cell(G,i,len G)
        by A2,A162,A164,A170,A171,A172,GOBRD13:def 3;
       now assume left_cell(F.(k+1),m,G) meets C;
      then consider p being set such that
A175:     p in cell(G,i,len G) & p in C by A174,XBOOLE_0:3;
      reconsider p as Point of TOP-REAL 2 by A175;
       reconsider p as Element of TOP-REAL 2;
        cell(G,i,len G) = { |[r,s]|: G*(i,1)`1 <= r & r <= G*(i+1,1)`1 &
                                  G*(1,len G)`2 <= s }
       by A2,A159,A163,GOBRD11:31;
      then consider r,s such that
A176:     p = |[r,s]| and
          G*(i,1)`1 <= r & r <= G*(i+1,1)`1 and
A177:     G*(1,len G)`2 <= s by A175;
        [1,len G] in Indices G by A2,A150,GOBOARD7:10;
then A178:     G*
(1,len G) = |[W+((E-W)/(2|^n))*(1-2),S+((N-S)/(2|^n))*((len G)-2)]|
         by JORDAN8:def 1;
         2|^n+(2+1) = 2|^n+1+2 by XCMPLX_1:1;
       then (len G) - 2 = 2|^n+1 by A2,XCMPLX_1:26;
       then ((N-S)/(2|^n))*((len G)-2)
           = ((N-S)/(2|^n))*(2|^n)+((N-S)/(2|^n))*1 by XCMPLX_1:8
          .= (N-S)+(N-S)/(2|^n) by A151,XCMPLX_1:88;
then A179:     S+((N-S)/(2|^n))*((len G) -2)
          = S+(N-S)+(N-S)/(2|^n) by XCMPLX_1:1
         .= N+(N-S)/(2|^n) by XCMPLX_1:27;
A180:     G*(1,len G)`2 = S+((N-S)/(2|^n))*((len G)-2) by A178,EUCLID:56;
         N > S by JORDAN8:12;
       then N-S > S-S by REAL_1:54;
       then N-S > 0 by XCMPLX_1:14;
       then (N-S)/(2|^n) > 0 by A151,REAL_2:127;
       then N+0 < N+(N-S)/(2|^n) by REAL_1:53;
then A181:     N < s by A177,A179,A180,AXIOMS:22;
         p`2 <= N by A175,PSCOMP_1:71;
      hence contradiction by A176,A181,EUCLID:56;
     end;
     hence left_cell(F.(k+1),m,G) misses C;
A182:   N-min C in C by SPRECT_1:13;
       N-min C in right_cell(F.(k+1),m,G)
      by A2,A160,A162,A164,A170,A171,A172,A173,GOBRD13:def 2;
     hence right_cell(F.(k+1),m,G) meets C by A182,XBOOLE_0:3;
    suppose
A183:   k > 1;
    then k > 0 by AXIOMS:22;
then A184:   F.k is non empty by A148,FINSEQ_1:25;
A185:   1 <= (len(F.k))-'1 by A148,A183,JORDAN3:12;
A186:   (len(F.k)) -'1 +1 = len(F.k) by A148,A183,AMI_5:4;
     then consider i1,j1,i2,j2 being Nat such that
A187:   [i1,j1] in Indices G & (F.k)/.((len(F.k)) -'1) = G*(i1,j1) and
A188:   [i2,j2] in Indices G & (F.k)/.len(F.k) = G*(i2,j2) and
        i1 = i2 & j1+1 = j2 or i1+1 = i2 & j1 = j2 or
        i1 = i2+1 & j1 = j2 or i1 = i2 & j1 = j2+1 by A146,A185,JORDAN8:6;
A189:   1 <= i2 & i2 <= len G & 1 <= j2 & j2 <= width G by A188,GOBOARD5:1;
        (len(F.k))-'1 <= len(F.k) by GOBOARD9:2;
then A190:  (len(F.k))-'1 in dom(F.k) & len(F.k) in dom(F.k)
        by A148,A183,A185,FINSEQ_3:27;
A191:  (len(F.k))-'1+(1+1) = (len(F.k))+1 by A186,XCMPLX_1:1;
A192:  left_cell(F.k,(len(F.k))-'1,G) misses C &
       right_cell(F.k,(len(F.k))-'1,G) meets C by A147,A185,A186;
A193: i1 = i2 & j1+1 = j2 implies [i2+1,j2] in Indices G by A110,A146,A147,A183
,A187,A188;
A194: i1+1 = i2 & j1 = j2 implies [i2,j2-'1] in Indices G by A110,A146,A147,
A183,A187,A188;
A195: i1 = i2+1 & j1 = j2 implies [i2,j2+1] in Indices G by A110,A146,A147,A183
,A187,A188;
A196: i1 = i2 & j1 = j2+1 implies [i2-'1,j2] in Indices G by A110,A146,A147,
A183,A187,A188;
A197: front_right_cell(F.k,(len(F.k))-'1,G) meets C &
     i1 = i2 & j1+1 = j2 implies [i2,j2+1] in Indices G by A110,A146,A147,A183,
A187,A188;
A198: front_right_cell(F.k,(len(F.k))-'1,G) meets C &
     i1+1 = i2 & j1 = j2 implies [i2+1,j2] in Indices G by A110,A146,A147,A183,
A187,A188;
A199: front_right_cell(F.k,(len(F.k))-'1,G) meets C &
     i1 = i2+1 & j1 = j2 implies [i2-'1,j2] in Indices G by A110,A146,A147,A183
,A187,A188;
A200: front_right_cell(F.k,(len(F.k))-'1,G) meets C &
     i1 = i2 & j1 = j2+1 implies [i2,j2-'1] in Indices G by A110,A146,A147,A183
,A187,A188;
A201: front_left_cell(F.k,(len(F.k))-'1,G) meets C &
     i1 = i2 & j1+1 = j2 implies [i2-'1,j2] in Indices G by A110,A146,A147,A183
,A187,A188;
A202: front_left_cell(F.k,(len(F.k))-'1,G) meets C &
     i1+1 = i2 & j1 = j2 implies [i2,j2+1] in Indices G by A110,A146,A147,A183,
A187,A188;
A203: front_left_cell(F.k,(len(F.k))-'1,G) meets C &
     i1 = i2+1 & j1 = j2 implies [i2,j2-'1] in Indices G by A110,A146,A147,A183
,A187,A188;
A204: front_left_cell(F.k,(len(F.k))-'1,G) meets C &
     i1 = i2 & j1 = j2+1 implies [i2+1,j2] in Indices G by A110,A146,A147,A183,
A187,A188;
   thus
A205: F.(k+1) is_sequence_on G
      proof
       per cases;
       suppose front_left_cell(F.k,(len(F.k))-'1,G) misses C &
            front_right_cell(F.k,(len(F.k))-'1,G) misses C;
        then consider i,j such that
A206:       (F.k)^<*G*(i,j)*> turns_right (len(F.k))-'1,G and
A207:       F.(k+1) = (F.k)^<*G*(i,j)*> by A92,A146,A148,A183,A192;
        thus thesis
         proof set f = (F.k)^<*G*(i,j)*>;
A208:       f/.(len(F.k)-'1) = G*(i1,j1) & f/.len(F.k) = G*(i2,j2)
           by A187,A188,A190,GROUP_5:95;
A209:        f/.(len(F.k)+1) = G*(i,j) by TOPREAL4:1;
          per cases by A186,A187,A188,A191,A206,A208,GOBRD13:def 6;
          suppose that
A210:         i1 = i2 & j1+1 = j2 and
A211:         f/.(len(F.k)+1) = G*(i2+1,j2);
             now let i1',j1',i2',j2' be Nat;
            assume [i1',j1'] in Indices G & [i2',j2'] in Indices G &
                (F.k)/.len(F.k) = G*(i1',j1') & G*(i2+1,j2) = G*(i2',j2');
             then i2 = i1' & j2 = j1' & i2+1 = i2' & j2 = j2'
                by A188,A193,A210,GOBOARD1:21;
             then i2'-i1' = 1 & j2'-j1' = 0 by XCMPLX_1:14,26;
             then abs(i2'-i1') = 1 & abs(j2'-j1') = 0 by ABSVALUE:def 1;
            hence abs(i2'-i1')+abs(j2'-j1') = 1;
           end;
           hence F.(k+1) is_sequence_on G by A146,A184,A193,A207,A209,A210,A211
,JORDAN8:9;
          suppose that
A212:          i1+1 = i2 & j1 = j2 and
A213:          f/.(len(F.k)+1) = G*(i2,j2-'1);
             now let i1',j1',i2',j2' be Nat;
            assume [i1',j1'] in Indices G & [i2',j2'] in Indices G &
              (F.k)/.len(F.k) = G*(i1',j1') & G*(i2,j2-'1) = G*
(i2',j2');
then A214:           i2 = i1' & j2 = j1' & i2 = i2' & j2-'1 = j2'
                by A188,A194,A212,GOBOARD1:21;
             then j1'-j2' = j2-(j2-1) by A189,SCMFSA_7:3;
             then i2'-i1' = 0 & j1'-j2' = 1 by A214,XCMPLX_1:14,18;
             then abs(i2'-i1') = 0 & abs(j1'-j2') = 1 by ABSVALUE:def 1;
            hence abs(i2'-i1')+abs(j2'-j1') = 1 by UNIFORM1:13;
           end;
           hence F.(k+1) is_sequence_on G by A146,A184,A194,A207,A209,A212,A213
,JORDAN8:9;
          suppose that
A215:          i1 = i2+1 & j1 = j2 and
A216:          f/.(len(F.k)+1) = G*(i2,j2+1);
             now let i1',j1',i2',j2' be Nat;
            assume [i1',j1'] in Indices G & [i2',j2'] in Indices G &
                (F.k)/.len(F.k) = G*(i1',j1') & G*(i2,j2+1) = G*(i2',j2');
             then i2 = i1' & j2 = j1' & i2 = i2' & j2+1 = j2'
                by A188,A195,A215,GOBOARD1:21;
             then i2'-i1' = 0 & j2'-j1' = 1 by XCMPLX_1:14,26;
             then abs(i2'-i1') = 0 & abs(j2'-j1') = 1 by ABSVALUE:def 1;
            hence abs(i2'-i1')+abs(j2'-j1') = 1;
           end;
           hence F.(k+1) is_sequence_on G by A146,A184,A195,A207,A209,A215,A216
,JORDAN8:9;
          suppose that
A217:          i1 = i2 & j1 = j2+1 and
A218:          f/.(len(F.k)+1) = G*(i2-'1,j2);
             now let i1',j1',i2',j2' be Nat;
            assume [i1',j1'] in Indices G & [i2',j2'] in Indices G &
              (F.k)/.len(F.k) = G*(i1',j1') & G*(i2-'1,j2) = G*
(i2',j2');
then A219:           i2 = i1' & j2 = j1' & i2-'1 = i2' & j2 = j2'
                by A188,A196,A217,GOBOARD1:21;
             then i1'-i2' = i2-(i2-1) by A189,SCMFSA_7:3;
             then j2'-j1' = 0 & i1'-i2' = 1 by A219,XCMPLX_1:14,18;
             then abs(j2'-j1') = 0 & abs(i1'-i2') = 1 by ABSVALUE:def 1;
            hence abs(i2'-i1')+abs(j2'-j1') = 1 by UNIFORM1:13;
           end;
           hence F.(k+1) is_sequence_on G by A146,A184,A196,A207,A209,A217,A218
,JORDAN8:9;
         end;
       suppose
A220:       front_left_cell(F.k,(len(F.k))-'1,G) misses C &
             front_right_cell(F.k,(len(F.k))-'1,G) meets C;
        then consider i,j such that
A221:       (F.k)^<*G*(i,j)*> goes_straight (len(F.k))-'1,G and
A222:       F.(k+1) = (F.k)^<*G*(i,j)*> by A92,A146,A148,A183,A192;
        thus thesis
         proof set f = (F.k)^<*G*(i,j)*>;
A223:       f/.(len(F.k)-'1) = G*(i1,j1) & f/.len(F.k) = G*(i2,j2)
           by A187,A188,A190,GROUP_5:95;
A224:        f/.(len(F.k)+1) = G*(i,j) by TOPREAL4:1;
          per cases by A186,A187,A188,A191,A221,A223,GOBRD13:def 8;
          suppose that
A225:         i1 = i2 & j1+1 = j2 and
A226:         f/.(len(F.k)+1) = G*(i2,j2+1);
             now let i1',j1',i2',j2' be Nat;
            assume [i1',j1'] in Indices G & [i2',j2'] in Indices G &
                (F.k)/.len(F.k) = G*(i1',j1') & G*(i2,j2+1) = G*(i2',j2');
             then i2 = i1' & j2 = j1' & i2 = i2' & j2+1 = j2'
                by A188,A197,A220,A225,GOBOARD1:21;
             then i2'-i1' = 0 & j2'-j1' = 1 by XCMPLX_1:14,26;
             then abs(i2'-i1') = 0 & abs(j2'-j1') = 1 by ABSVALUE:def 1;
            hence abs(i2'-i1')+abs(j2'-j1') = 1;
           end;
          hence F.(k+1) is_sequence_on G by A146,A184,A197,A220,A222,A224,A225,
A226,JORDAN8:9;
          suppose that
A227:          i1+1 = i2 & j1 = j2 and
A228:          f/.(len(F.k)+1) = G*(i2+1,j2);
             now let i1',j1',i2',j2' be Nat;
            assume [i1',j1'] in Indices G & [i2',j2'] in Indices G &
              (F.k)/.len(F.k) = G*(i1',j1') & G*(i2+1,j2) = G*
(i2',j2');
              then i2 = i1' & j2 = j1' & i2+1 = i2' & j2 = j2'
                by A188,A198,A220,A227,GOBOARD1:21;
             then i2'-i1' = 1 & j2'-j1' = 0 by XCMPLX_1:14,26;
             then abs(i2'-i1') = 1 & abs(j2'-j1') = 0 by ABSVALUE:def 1;
            hence abs(i2'-i1')+abs(j2'-j1') = 1;
           end;
          hence F.(k+1) is_sequence_on G by A146,A184,A198,A220,A222,A224,A227,
A228,JORDAN8:9;
          suppose that
A229:          i1 = i2+1 & j1 = j2 and
A230:          f/.(len(F.k)+1) = G*(i2-'1,j2);
             now let i1',j1',i2',j2' be Nat;
            assume [i1',j1'] in Indices G & [i2',j2'] in Indices G &
                (F.k)/.len(F.k) = G*(i1',j1') & G*(i2-'1,j2) = G*(i2',j2');
             then A231:           i2 = i1' & j2 = j1' & i2-'1 = i2' & j2 = j2'
                by A188,A199,A220,A229,GOBOARD1:21;
             then i1'-i2' = i2-(i2-1) by A189,SCMFSA_7:3;
             then j2'-j1' = 0 & i1'-i2' = 1 by A231,XCMPLX_1:14,18;
             then abs(j2'-j1') = 0 & abs(i1'-i2') = 1 by ABSVALUE:def 1;
            hence abs(i2'-i1')+abs(j2'-j1') = 1 by UNIFORM1:13;
           end;
          hence F.(k+1) is_sequence_on G by A146,A184,A199,A220,A222,A224,A229,
A230,JORDAN8:9;
          suppose that
A232:          i1 = i2 & j1 = j2+1 and
A233:          f/.(len(F.k)+1) = G*(i2,j2-'1);
             now let i1',j1',i2',j2' be Nat;
            assume [i1',j1'] in Indices G & [i2',j2'] in Indices G &
              (F.k)/.len(F.k) = G*(i1',j1') & G*(i2,j2-'1) = G*
(i2',j2');
then A234:           i2 = i1' & j2 = j1' & i2 = i2' & j2-'1 = j2'
                by A188,A200,A220,A232,GOBOARD1:21;
             then j1'-j2' = j2-(j2-1) by A189,SCMFSA_7:3;
             then i2'-i1' = 0 & j1'-j2' = 1 by A234,XCMPLX_1:14,18;
             then abs(i2'-i1') = 0 & abs(j1'-j2') = 1 by ABSVALUE:def 1;
            hence abs(i2'-i1')+abs(j2'-j1') = 1 by UNIFORM1:13;
           end;
          hence F.(k+1) is_sequence_on G by A146,A184,A200,A220,A222,A224,A232,
A233,JORDAN8:9;
         end;
       suppose
A235:       front_left_cell(F.k,(len(F.k))-'1,G) meets C;
        then consider i,j such that
A236:       (F.k)^<*G*(i,j)*> turns_left (len(F.k))-'1,G and
A237:       F.(k+1) = (F.k)^<*G*(i,j)*> by A92,A146,A148,A183,A192;
        thus thesis
         proof set f = (F.k)^<*G*(i,j)*>;
A238:       f/.(len(F.k)-'1) = G*(i1,j1) & f/.len(F.k) = G*(i2,j2)
           by A187,A188,A190,GROUP_5:95;
A239:        f/.(len(F.k)+1) = G*(i,j) by TOPREAL4:1;
          per cases by A186,A187,A188,A191,A236,A238,GOBRD13:def 7;
          suppose that
A240:         i1 = i2 & j1+1 = j2 and
A241:         f/.(len(F.k)+1) = G*(i2-'1,j2);
             now let i1',j1',i2',j2' be Nat;
            assume [i1',j1'] in Indices G & [i2',j2'] in Indices G &
                (F.k)/.len(F.k) = G*(i1',j1') & G*(i2-'1,j2) = G*(i2',j2');
             then A242:           i2 = i1' & j2 = j1' & i2-'1 = i2' & j2 = j2'
                by A188,A201,A235,A240,GOBOARD1:21;
             then i1'-i2' = i2-(i2-1) by A189,SCMFSA_7:3;
             then j2'-j1' = 0 & i1'-i2' = 1 by A242,XCMPLX_1:14,18;
             then abs(j2'-j1') = 0 & abs(i1'-i2') = 1 by ABSVALUE:def 1;
            hence abs(i2'-i1')+abs(j2'-j1') = 1 by UNIFORM1:13;
           end;
          hence F.(k+1) is_sequence_on G by A146,A184,A201,A235,A237,A239,A240,
A241,JORDAN8:9;
          suppose that
A243:          i1+1 = i2 & j1 = j2 and
A244:          f/.(len(F.k)+1) = G*(i2,j2+1);
             now let i1',j1',i2',j2' be Nat;
            assume [i1',j1'] in Indices G & [i2',j2'] in Indices G &
              (F.k)/.len(F.k) = G*(i1',j1') & G*(i2,j2+1) = G*
(i2',j2');
              then i2 = i1' & j2 = j1' & i2 = i2' & j2+1 = j2'
                by A188,A202,A235,A243,GOBOARD1:21;
             then i2'-i1' = 0 & j2'-j1' = 1 by XCMPLX_1:14,26;
             then abs(i2'-i1') = 0 & abs(j2'-j1') = 1 by ABSVALUE:def 1;
            hence abs(i2'-i1')+abs(j2'-j1') = 1;
           end;
          hence F.(k+1) is_sequence_on G by A146,A184,A202,A235,A237,A239,A243,
A244,JORDAN8:9;
          suppose that
A245:          i1 = i2+1 & j1 = j2 and
A246:          f/.(len(F.k)+1) = G*(i2,j2-'1);
             now let i1',j1',i2',j2' be Nat;
            assume [i1',j1'] in Indices G & [i2',j2'] in Indices G &
                (F.k)/.len(F.k) = G*(i1',j1') & G*(i2,j2-'1) = G*(i2',j2');
             then A247:           i2 = i1' & j2 = j1' & i2 = i2' & j2-'1 = j2'
                by A188,A203,A235,A245,GOBOARD1:21;
             then j1'-j2' = j2-(j2-1) by A189,SCMFSA_7:3;
             then i2'-i1' = 0 & j1'-j2' = 1 by A247,XCMPLX_1:14,18;
             then abs(i2'-i1') = 0 & abs(j1'-j2') = 1 by ABSVALUE:def 1;
            hence abs(i2'-i1')+abs(j2'-j1') = 1 by UNIFORM1:13;
           end;
          hence F.(k+1) is_sequence_on G by A146,A184,A203,A235,A237,A239,A245,
A246,JORDAN8:9;
          suppose that
A248:          i1 = i2 & j1 = j2+1 and
A249:          f/.(len(F.k)+1) = G*(i2+1,j2);
             now let i1',j1',i2',j2' be Nat;
            assume [i1',j1'] in Indices G & [i2',j2'] in Indices G &
              (F.k)/.len(F.k) = G*(i1',j1') & G*(i2+1,j2) = G*
(i2',j2');
              then i2 = i1' & j2 = j1' & i2+1 = i2' & j2 = j2'
                by A188,A204,A235,A248,GOBOARD1:21;
             then i2'-i1' = 1 & j2'-j1' = 0 by XCMPLX_1:14,26;
             then abs(i2'-i1') = 1 & abs(j2'-j1') = 0 by ABSVALUE:def 1;
            hence abs(i2'-i1')+abs(j2'-j1') = 1;
           end;
          hence F.(k+1) is_sequence_on G by A146,A184,A204,A235,A237,A239,A248,
A249,JORDAN8:9;
         end;
      end;
     let m such that
A250:    1 <= m & m+1 <= len(F.(k+1));
       now
     per cases;
      suppose m+1 = len(F.(k+1));
then A251:    m = len(F.k) by A148,A149,XCMPLX_1:2;
A252:    (i2-'1)+1 = i2 by A189,AMI_5:4;
A253:    (j2-'1)+1 = j2 by A189,AMI_5:4;
       thus thesis
        proof per cases;
         suppose
A254:        front_left_cell(F.k,(len(F.k))-'1,G) misses C &
            front_right_cell(F.k,(len(F.k))-'1,G) misses C;
          then consider i,j such that
A255:        (F.k)^<*G*(i,j)*> turns_right (len(F.k))-'1,G and
A256:        F.(k+1) = (F.k)^<*G*(i,j)*> by A92,A146,A148,A183,A192;
A257:        (F.(k+1))/.(len(F.k)-'1) = G*(i1,j1) & (F.(k+1))/.len(F.k) = G*
(i2,j2)
            by A187,A188,A190,A256,GROUP_5:95;
            now per cases by A186,A187,A188,A191,A255,A256,A257,GOBRD13:def 6;
           suppose that
A258:           i1 = i2 & j1+1 = j2 and
A259:           (F.(k+1))/.((len(F.k))+1) = G*(i2+1,j2);
               front_right_cell(F.k,(len(F.k))-'1,G) = cell(G,i1,j2)
               by A146,A185,A186,A187,A188,A258,GOBRD13:36;
            hence left_cell(F.(k+1),m,G) misses C
               by A188,A193,A205,A250,A251,A254,A257,A258,A259,GOBRD13:24;
               j2-'1 = j1 & cell(G,i1,j1) meets C
               by A146,A185,A186,A187,A188,A192,A258,BINARITH:39,GOBRD13:23;
            hence right_cell(F.(k+1),m,G) meets C
               by A188,A193,A205,A250,A251,A257,A258,A259,GOBRD13:25;
           suppose that
A260:           i1+1 = i2 & j1 = j2 and
A261:           (F.(k+1))/.((len(F.k))+1) = G*(i2,j2-'1);
               front_right_cell(F.k,(len(F.k))-'1,G) = cell(G,i2,j2-'1)
               by A146,A185,A186,A187,A188,A260,GOBRD13:38;
            hence left_cell(F.(k+1),m,G) misses C
               by A188,A194,A205,A250,A251,A253,A254,A257,A260,A261,GOBRD13:28;
               i2-'1 = i1 & cell(G,i1,j1-'1) meets C
               by A146,A185,A186,A187,A188,A192,A260,BINARITH:39,GOBRD13:25;
            hence right_cell(F.(k+1),m,G) meets C
               by A188,A194,A205,A250,A251,A253,A257,A260,A261,GOBRD13:29;
           suppose that
A262:           i1 = i2+1 & j1 = j2 and
A263:           (F.(k+1))/.((len(F.k))+1) = G*(i2,j2+1);
               front_right_cell(F.k,(len(F.k))-'1,G) = cell(G,i2-'1,j2)
               by A146,A185,A186,A187,A188,A262,GOBRD13:40;
            hence left_cell(F.(k+1),m,G) misses C
               by A188,A195,A205,A250,A251,A254,A257,A262,A263,GOBRD13:22;
               cell(G,i2,j2) meets C
               by A146,A185,A186,A187,A188,A192,A262,GOBRD13:27;
            hence right_cell(F.(k+1),m,G) meets C
               by A188,A195,A205,A250,A251,A257,A262,A263,GOBRD13:23;
           suppose that
A264:           i1 = i2 & j1 = j2+1 and
A265:           (F.(k+1))/.((len(F.k))+1) = G*(i2-'1,j2);
               front_right_cell(F.k,(len(F.k))-'1,G) = cell(G,i2-'1,j2-'1)
               by A146,A185,A186,A187,A188,A264,GOBRD13:42;
            hence left_cell(F.(k+1),m,G) misses C
              by A188,A196,A205,A250,A251,A252,A254,A257,A264,A265,GOBRD13:26;
               cell(G,i2-'1,j2) meets C
              by A146,A185,A186,A187,A188,A192,A264,GOBRD13:29;
            hence right_cell(F.(k+1),m,G) meets C
             by A188,A196,A205,A250,A251,A252,A257,A264,A265,GOBRD13:27;
          end;
          hence thesis;
         suppose
A266:        front_left_cell(F.k,(len(F.k))-'1,G) misses C &
             front_right_cell(F.k,(len(F.k))-'1,G) meets C;
          then consider i,j such that
A267:        (F.k)^<*G*(i,j)*> goes_straight (len(F.k))-'1,G and
A268:        F.(k+1) = (F.k)^<*G*(i,j)*> by A92,A146,A148,A183,A192;
A269:        (F.(k+1))/.(len(F.k)-'1) = G*(i1,j1) & (F.(k+1))/.len(F.k)
                      = G*(i2,j2)
            by A187,A188,A190,A268,GROUP_5:95;
            now per cases by A186,A187,A188,A191,A267,A268,A269,GOBRD13:def 8;
           suppose that
A270:           i1 = i2 & j1+1 = j2 and
A271:           (F.(k+1))/.(len(F.k)+1) = G*(i2,j2+1);
               front_left_cell(F.k,(len(F.k))-'1,G) = cell(G,i1-'1,j2)
               by A146,A185,A186,A187,A188,A270,GOBRD13:35;
             hence left_cell(F.(k+1),m,G) misses C
               by A188,A197,A205,A250,A251,A266,A269,A270,A271,GOBRD13:22;
              front_right_cell(F.k,(len(F.k))-'1,G) = cell(G,i1,j2)
               by A146,A185,A186,A187,A188,A270,GOBRD13:36;
            hence right_cell(F.(k+1),m,G) meets C
               by A188,A197,A205,A250,A251,A266,A269,A270,A271,GOBRD13:23;
           suppose that
A272:           i1+1 = i2 & j1 = j2 and
A273:           (F.(k+1))/.(len(F.k)+1) = G*(i2+1,j2);
               front_left_cell(F.k,(len(F.k))-'1,G) = cell(G,i2,j2)
               by A146,A185,A186,A187,A188,A272,GOBRD13:37;
            hence left_cell(F.(k+1),m,G) misses C
               by A188,A198,A205,A250,A251,A266,A269,A272,A273,GOBRD13:24;
               front_right_cell(F.k,(len(F.k))-'1,G) = cell(G,i2,j2-'1)
               by A146,A185,A186,A187,A188,A272,GOBRD13:38;
            hence right_cell(F.(k+1),m,G) meets C
               by A188,A198,A205,A250,A251,A266,A269,A272,A273,GOBRD13:25;
           suppose that
A274:          i1 = i2+1 & j1 = j2 and
A275:          (F.(k+1))/.(len(F.k)+1) = G*(i2-'1,j2);
               front_left_cell(F.k,(len(F.k))-'1,G) = cell(G,i2-'1,j2-'1)
               by A146,A185,A186,A187,A188,A274,GOBRD13:39;
            hence left_cell(F.(k+1),m,G) misses C
               by A188,A199,A205,A250,A251,A252,A266,A269,A274,A275,GOBRD13:26;
               front_right_cell(F.k,(len(F.k))-'1,G) = cell(G,i2-'1,j2)
               by A146,A185,A186,A187,A188,A274,GOBRD13:40;
            hence right_cell(F.(k+1),m,G) meets C
               by A188,A199,A205,A250,A251,A252,A266,A269,A274,A275,GOBRD13:27;
           suppose that
A276:          i1 = i2 & j1 = j2+1 and
A277:          (F.(k+1))/.(len(F.k)+1) = G*(i2,j2-'1);
               front_left_cell(F.k,(len(F.k))-'1,G) = cell(G,i2,j2-'1)
               by A146,A185,A186,A187,A188,A276,GOBRD13:41;
            hence left_cell(F.(k+1),m,G) misses C
               by A188,A200,A205,A250,A251,A253,A266,A269,A276,A277,GOBRD13:28;
               front_right_cell(F.k,(len(F.k))-'1,G) = cell(G,i2-'1,j2-'1)
               by A146,A185,A186,A187,A188,A276,GOBRD13:42;
            hence right_cell(F.(k+1),m,G) meets C
               by A188,A200,A205,A250,A251,A253,A266,A269,A276,A277,GOBRD13:29;
          end;
          hence thesis;
         suppose
A278:        front_left_cell(F.k,(len(F.k))-'1,G) meets C;
          then consider i,j such that
A279:        (F.k)^<*G*(i,j)*> turns_left (len(F.k))-'1,G and
A280:        F.(k+1) = (F.k)^<*G*(i,j)*> by A92,A146,A148,A183,A192;
A281:        (F.(k+1))/.(len(F.k)-'1) = G*(i1,j1) & (F.(k+1))/.len(F.k) = G*
(i2,j2)
             by A187,A188,A190,A280,GROUP_5:95;
            now per cases by A186,A187,A188,A191,A279,A280,A281,GOBRD13:def 7;
           suppose that
A282:           i1 = i2 & j1+1 = j2 and
A283:           (F.(k+1))/.(len(F.k)+1) = G*(i2-'1,j2);
               j2-'1 = j1 &
             cell(G,i1-'1,j1) misses C
               by A146,A185,A186,A187,A188,A192,A282,BINARITH:39,GOBRD13:22;
            hence left_cell(F.(k+1),m,G) misses C
               by A188,A201,A205,A250,A251,A252,A278,A281,A282,A283,GOBRD13:26;
               front_left_cell(F.k,(len(F.k))-'1,G) = cell(G,i1-'1,j2)
               by A146,A185,A186,A187,A188,A282,GOBRD13:35;
            hence right_cell(F.(k+1),m,G) meets C
               by A188,A201,A205,A250,A251,A252,A278,A281,A282,A283,GOBRD13:27;
           suppose that
A284:           i1+1 = i2 & j1 = j2 and
A285:           (F.(k+1))/.(len(F.k)+1) = G*(i2,j2+1);
               i2-'1 = i1 & cell(G,i1,j1) misses C
               by A146,A185,A186,A187,A188,A192,A284,BINARITH:39,GOBRD13:24;
             hence left_cell(F.(k+1),m,G) misses C
               by A188,A202,A205,A250,A251,A278,A281,A284,A285,GOBRD13:22;
               front_left_cell(F.k,(len(F.k))-'1,G) = cell(G,i2,j2)
               by A146,A185,A186,A187,A188,A284,GOBRD13:37;
            hence right_cell(F.(k+1),m,G) meets C
               by A188,A202,A205,A250,A251,A278,A281,A284,A285,GOBRD13:23;
           suppose that
A286:           i1 = i2+1 & j1 = j2 and
A287:           (F.(k+1))/.(len(F.k)+1) = G*(i2,j2-'1);
               cell(G,i2,j2-'1) misses C
               by A146,A185,A186,A187,A188,A192,A286,GOBRD13:26;
            hence left_cell(F.(k+1),m,G) misses C
               by A188,A203,A205,A250,A251,A253,A278,A281,A286,A287,GOBRD13:28;
               front_left_cell(F.k,(len(F.k))-'1,G) = cell(G,i2-'1,j2-'1)
               by A146,A185,A186,A187,A188,A286,GOBRD13:39;
            hence right_cell(F.(k+1),m,G) meets C
               by A188,A203,A205,A250,A251,A253,A278,A281,A286,A287,GOBRD13:29;
           suppose that
A288:           i1 = i2 & j1 = j2+1 and
A289:           (F.(k+1))/.(len(F.k)+1) = G*(i2+1,j2);
               cell(G,i2,j2) misses C by A146,A185,A186,A187,A188,A192,A288,
GOBRD13:28;
            hence left_cell(F.(k+1),m,G) misses C
               by A188,A204,A205,A250,A251,A278,A281,A288,A289,GOBRD13:24;
               front_left_cell(F.k,(len(F.k))-'1,G) = cell(G,i2,j2-'1)
               by A146,A185,A186,A187,A188,A288,GOBRD13:41;
            hence right_cell(F.(k+1),m,G) meets C
               by A188,A204,A205,A250,A251,A278,A281,A288,A289,GOBRD13:25;
          end;
          hence thesis;
       end;
      suppose m+1 <> len(F.(k+1));
       then m+1 < len(F.(k+1)) by A250,REAL_1:def 5;
       then A290:     m+1 <= len(F.k)by A148,A149,NAT_1:38;
       then consider i1,j1,i2,j2 being Nat such that
A291:     [i1,j1] in Indices G & (F.k)/.m = G*(i1,j1) and
A292:     [i2,j2] in Indices G & (F.k)/.(m+1) = G*(i2,j2) and
A293:      i1 = i2 & j1+1 = j2 or i1+1 = i2 & j1 = j2 or
          i1 = i2+1 & j1 = j2 or i1 = i2 & j1 = j2+1 by A146,A250,JORDAN8:6;
A294:    1 <= m+1 by NAT_1:37;
          m <= len(F.k) by A290,NAT_1:38;
then A295:    m in dom(F.k) & m+1 in dom(F.k) by A250,A290,A294,FINSEQ_3:27;
A296:     left_cell(F.k,m,G) misses C & right_cell(F.k,m,G) meets C
         by A147,A250,A290;
         now per cases;
         suppose
            front_left_cell(F.k,(len(F.k))-'1,G) misses C &
            front_right_cell(F.k,(len(F.k))-'1,G) misses C;
          then consider i,j such that
             (F.k)^<*G*(i,j)*> turns_right (len(F.k))-'1,G and
A297:        F.(k+1) = (F.k)^<*G*(i,j)*> by A92,A146,A148,A183,A192;
          take i,j;
          thus F.(k+1) = (F.k)^<*G*(i,j)*> by A297;
         suppose
             front_left_cell(F.k,(len(F.k))-'1,G) misses C &
             front_right_cell(F.k,(len(F.k))-'1,G) meets C;
          then consider i,j such that
             (F.k)^<*G*(i,j)*> goes_straight (len(F.k))-'1,G and
A298:        F.(k+1) = (F.k)^<*G*(i,j)*> by A92,A146,A148,A183,A192;
          take i,j;
          thus F.(k+1) = (F.k)^<*G*(i,j)*> by A298;
         suppose
            front_left_cell(F.k,(len(F.k))-'1,G) meets C;
          then consider i,j such that
             (F.k)^<*G*(i,j)*> turns_left (len(F.k))-'1,G and
A299:        F.(k+1) = (F.k)^<*G*(i,j)*> by A92,A146,A148,A183,A192;
          take i,j;
          thus F.(k+1) = (F.k)^<*G*(i,j)*> by A299;
       end;
       then consider i,j such that
A300:      F.(k+1) = (F.k)^<*G*(i,j)*>;
A301:      (F.(k+1))/.m = G*(i1,j1) & (F.(k+1))/.(m+1) = G*(i2,j2)
            by A291,A292,A295,A300,GROUP_5:95;
         now per cases by A293;
        suppose
A302:        i1 = i2 & j1+1 = j2;
          then left_cell(F.k,m,G) = cell(G,i1-'1,j1) &
            right_cell(F.k,m,G) = cell(G,i1,j1)
           by A146,A250,A290,A291,A292,GOBRD13:22,23;
         hence thesis by A205,A250,A291,A292,A296,A301,A302,GOBRD13:22,23;
        suppose
A303:        i1+1 = i2 & j1 = j2;
          then left_cell(F.k,m,G) = cell(G,i1,j1) &
            right_cell(F.k,m,G) = cell(G,i1,j1-'1)
           by A146,A250,A290,A291,A292,GOBRD13:24,25;
         hence thesis by A205,A250,A291,A292,A296,A301,A303,GOBRD13:24,25;
        suppose
A304:        i1 = i2+1 & j1 = j2;
          then left_cell(F.k,m,G) = cell(G,i2,j2-'1) &
            right_cell(F.k,m,G) = cell(G,i2,j2)
           by A146,A250,A290,A291,A292,GOBRD13:26,27;
         hence thesis by A205,A250,A291,A292,A296,A301,A304,GOBRD13:26,27;
        suppose
A305:        i1 = i2 & j1 = j2+1;
          then left_cell(F.k,m,G) = cell(G,i2,j2) &
            right_cell(F.k,m,G) = cell(G,i1-'1,j2)
           by A146,A250,A290,A291,A292,GOBRD13:28,29;
         hence thesis by A205,A250,A291,A292,A296,A301,A305,GOBRD13:28,29;
       end;
       hence thesis;
     end;
    hence thesis;
   end;

A306: for k holds P[k] from Ind(A108,A145);
A307: for k,i1,i2,j1,j2 st k > 1 &
        [i1,j1] in Indices G & (F.k)/.((len(F.k)) -'1) = G*(i1,j1) &
        [i2,j2] in Indices G & (F.k)/.len(F.k) = G*(i2,j2)
      holds
       (front_left_cell(F.k,(len(F.k))-'1,G) misses C &
          front_right_cell(F.k,(len(F.k))-'1,G) misses C
         implies F.(k+1) turns_right (len(F.k))-'1,G &
          (i1 = i2 & j1+1 = j2 implies [i2+1,j2] in Indices G &
            F.(k+1) = (F.k)^<*G*(i2+1,j2)*>) &
          (i1+1 = i2 & j1 = j2 implies [i2,j2-'1] in Indices G &
            F.(k+1) = (F.k)^<*G*(i2,j2-'1)*>) &
          (i1 = i2+1 & j1 = j2 implies [i2,j2+1] in Indices G &
            F.(k+1) = (F.k)^<*G*(i2,j2+1)*>)&
          (i1 = i2 & j1 = j2+1 implies [i2-'1,j2] in Indices G &
            F.(k+1) = (F.k)^<*G*(i2-'1,j2)*>)) &
       (front_left_cell(F.k,(len(F.k))-'1,G) misses C &
          front_right_cell(F.k,(len(F.k))-'1,G) meets C
         implies F.(k+1) goes_straight (len(F.k))-'1,G &
          (i1 = i2 & j1+1 = j2 implies [i2,j2+1] in Indices G &
            F.(k+1) = (F.k)^<*G*(i2,j2+1)*>) &
          (i1+1 = i2 & j1 = j2 implies [i2+1,j2] in Indices G &
            F.(k+1) = (F.k)^<*G*(i2+1,j2)*>) &
          (i1 = i2+1 & j1 = j2 implies [i2-'1,j2] in Indices G &
            F.(k+1) = (F.k)^<*G*(i2-'1,j2)*>)&
          (i1 = i2 & j1 = j2+1 implies [i2,j2-'1] in Indices G &
            F.(k+1) = (F.k)^<*G*(i2,j2-'1)*>)) &
       (front_left_cell(F.k,(len(F.k))-'1,G) meets C
         implies F.(k+1) turns_left (len(F.k))-'1,G &
          (i1 = i2 & j1+1 = j2 implies [i2-'1,j2] in Indices G &
            F.(k+1) = (F.k)^<*G*(i2-'1,j2)*>) &
          (i1+1 = i2 & j1 = j2 implies [i2,j2+1] in Indices G &
            F.(k+1) = (F.k)^<*G*(i2,j2+1)*>) &
          (i1 = i2+1 & j1 = j2 implies [i2,j2-'1] in Indices G &
            F.(k+1) = (F.k)^<*G*(i2,j2-'1)*>)&
          (i1 = i2 & j1 = j2+1 implies [i2+1,j2] in Indices G &
            F.(k+1) = (F.k)^<*G*(i2+1,j2)*>))
  proof let k,i1,i2,j1,j2 such that
A308:  k > 1 and
A309:  [i1,j1] in Indices G & (F.k)/.((len(F.k)) -'1) = G*(i1,j1) and
A310:  [i2,j2] in Indices G & (F.k)/.len(F.k) = G*(i2,j2);
A311:  (len(F.k))-'1 <= len(F.k) by GOBOARD9:2;
A312:   F.k is_sequence_on G by A306;
A313:   len(F.k) = k by A107;
then A314:  1 <= (len(F.k))-'1 by A308,JORDAN3:12;
A315:  (len(F.k)) -'1 +1 = len(F.k) by A308,A313,AMI_5:4;
then A316: right_cell(F.k,(len(F.k))-'1,G) meets C by A306,A314;
A317: (len(F.k))-'1 in dom(F.k) & len(F.k) in dom(F.k)
       by A308,A311,A313,A314,FINSEQ_3:27;
A318: (len(F.k))-'1+(1+1) = (len(F.k))+1 by A315,XCMPLX_1:1;
A319: i1+1 > i1 & i2+1 > i2 & j1+1 > j1 & j2+1 > j2 by NAT_1:38;
    hereby assume
        front_left_cell(F.k,(len(F.k))-'1,G) misses C &
        front_right_cell(F.k,(len(F.k))-'1,G) misses C;
     then consider i,j such that
A320:   (F.k)^<*G*(i,j)*> turns_right (len(F.k))-'1,G and
A321:   F.(k+1) = (F.k)^<*G*(i,j)*> by A92,A308,A312,A313,A316;
A322:   (F.(k+1))/.(len(F.k)-'1) = G*(i1,j1) & (F.(k+1))/.len(F.k) = G*
(i2,j2)
        by A309,A310,A317,A321,GROUP_5:95;
A323:   (F.(k+1))/.(len(F.k)+1) = G*(i,j) by A321,TOPREAL4:1;
     thus
      F.(k+1) turns_right (len(F.k))-'1,G by A320,A321;
     thus i1 = i2 & j1+1 = j2 implies
      [i2+1,j2] in Indices G &
       F.(k+1) = (F.k)^<*G*(i2+1,j2)*>
       by A309,A310,A315,A318,A319,A320,A321,A322,A323,GOBRD13:def 6;
    thus i1+1 = i2 & j1 = j2 implies
       [i2,j2-'1] in Indices G &
       F.(k+1) = (F.k)^<*G*(i2,j2-'1)*>
       by A309,A310,A315,A318,A319,A320,A321,A322,A323,GOBRD13:def 6;
     thus i1 = i2+1 & j1 = j2 implies
      [i2,j2+1] in Indices G &
      F.(k+1) = (F.k)^<*G*(i2,j2+1)*>
       by A309,A310,A315,A318,A319,A320,A321,A322,A323,GOBRD13:def 6;
     thus i1 = i2 & j1 = j2+1 implies
      [i2-'1,j2] in Indices G &
      F.(k+1) = (F.k)^<*G*(i2-'1,j2)*>
       by A309,A310,A315,A318,A319,A320,A321,A322,A323,GOBRD13:def 6;
   end;
   hereby assume
        front_left_cell(F.k,(len(F.k))-'1,G) misses C &
        front_right_cell(F.k,(len(F.k))-'1,G) meets C;
    then consider i,j such that
A324:   (F.k)^<*G*(i,j)*> goes_straight (len(F.k))-'1,G and
A325:   F.(k+1) = (F.k)^<*G*(i,j)*> by A92,A308,A312,A313,A316;
A326:   (F.(k+1))/.(len(F.k)-'1) = G*(i1,j1) & (F.(k+1))/.len(F.k) = G*
(i2,j2)
        by A309,A310,A317,A325,GROUP_5:95;
A327:   (F.(k+1))/.(len(F.k)+1) = G*(i,j) by A325,TOPREAL4:1;
    thus
    F.(k+1) goes_straight (len(F.k))-'1,G by A324,A325;
    thus i1 = i2 & j1+1 = j2 implies
      [i2,j2+1] in Indices G & F.(k+1) = (F.k)^<*G*(i2,j2+1)*>
       by A309,A310,A315,A318,A319,A324,A325,A326,A327,GOBRD13:def 8;
    thus i1+1 = i2 & j1 = j2 implies
      [i2+1,j2] in Indices G & F.(k+1) = (F.k)^<*G*(i2+1,j2)*>
       by A309,A310,A315,A318,A319,A324,A325,A326,A327,GOBRD13:def 8;
    thus i1 = i2+1 & j1 = j2 implies
      [i2-'1,j2] in Indices G & F.(k+1) = (F.k)^<*G*(i2-'1,j2)*>
       by A309,A310,A315,A318,A319,A324,A325,A326,A327,GOBRD13:def 8;
    thus i1 = i2 & j1 = j2+1 implies
      [i2,j2-'1] in Indices G & F.(k+1) = (F.k)^<*G*(i2,j2-'1)*>
       by A309,A310,A315,A318,A319,A324,A325,A326,A327,GOBRD13:def 8;
   end;
   assume
        front_left_cell(F.k,(len(F.k))-'1,G) meets C;
   then consider i,j such that
A328:   (F.k)^<*G*(i,j)*> turns_left (len(F.k))-'1,G and
A329:   F.(k+1) = (F.k)^<*G*(i,j)*> by A92,A308,A312,A313,A316;
A330:   (F.(k+1))/.(len(F.k)-'1) = G*(i1,j1) & (F.(k+1))/.len(F.k) = G*
(i2,j2)
        by A309,A310,A317,A329,GROUP_5:95;
A331:   (F.(k+1))/.(len(F.k)+1) = G*(i,j) by A329,TOPREAL4:1;
   thus
    F.(k+1) turns_left (len(F.k))-'1,G by A328,A329;
   thus i1 = i2 & j1+1 = j2 implies
    [i2-'1,j2] in Indices G & F.(k+1) = (F.k)^<*G*(i2-'1,j2)*>
       by A309,A310,A315,A318,A319,A328,A329,A330,A331,GOBRD13:def 7;
   thus i1+1 = i2 & j1 = j2 implies
    [i2,j2+1] in Indices G & F.(k+1) = (F.k)^<*G*(i2,j2+1)*>
       by A309,A310,A315,A318,A319,A328,A329,A330,A331,GOBRD13:def 7;
   thus i1 = i2+1 & j1 = j2 implies
    [i2,j2-'1] in Indices G & F.(k+1) = (F.k)^<*G*(i2,j2-'1)*>
       by A309,A310,A315,A318,A319,A328,A329,A330,A331,GOBRD13:def 7;
   thus i1 = i2 & j1 = j2+1 implies
    [i2+1,j2] in Indices G & F.(k+1) = (F.k)^<*G*(i2+1,j2)*>
       by A309,A310,A315,A318,A319,A328,A329,A330,A331,GOBRD13:def 7;
 end;

A332:
  for k holds ex i,j st [i,j] in Indices G & F.(k+1) = (F.k)^<*G*(i,j)*>
   proof let k;
A333:  F.k is_sequence_on G by A306;
A334:  len(F.k) = k by A107;
A335:  len G = 2|^n+3 & len G = width G by JORDAN8:def 1;
      len G >= 4 by JORDAN8:13;
    then A336:    1 < len G by AXIOMS:22;
    per cases by REAL_1:def 5;
    suppose k < 1;
then A337:   k = 0 by RLVECT_1:98;
     consider i such that
A338:      1 <= i & i+1 <= len G and
          N-min C in cell(G,i,width G-'1) & N-min C <> G*(i,width G-'1) and
A339:    F.(0+1) = <*G*(i,width G)*> by A92;
      take i, j = width G;
        i < len G by A338,NAT_1:38;
      hence [i,j] in Indices G by A335,A336,A338,GOBOARD7:10;
      thus thesis by A91,A337,A339,FINSEQ_1:47;
    suppose
A340:   k = 1;
     consider i such that
A341:    1 <= i & i+1 <= len G and
A342:    N-min C in cell(G,i,width G-'1) & N-min C <> G*(i,width G-'1) and
A343:    F.(0+1) = <*G*(i,width G)*> by A92;
     consider i' being Nat such that
A344:    1 <= i' & i'+1 <= len G &
       N-min C in cell(G,i',width G-'1) & N-min C <> G*(i',width G-'1) and
A345:    F.(1+1) = <*G*(i',width G),G*(i'+1,width G)*> by A92;

A346:    i = i' by A341,A342,A344,Th31;
     take i+1,j = width G;
       1 <= i+1 & i+1 <= len G by A341,NAT_1:38;
     hence [i+1,j] in Indices G by A335,A336,GOBOARD7:10;
     thus thesis by A340,A343,A345,A346,FINSEQ_1:def 9;
    suppose
A347:   k > 1;
then A348:   1 <= (len(F.k))-'1 by A334,JORDAN3:12;
        (len(F.k)) -'1 +1 = len(F.k) by A334,A347,AMI_5:4;
     then consider i1,j1,i2,j2 being Nat such that
A349:   [i1,j1] in Indices G & (F.k)/.((len(F.k)) -'1) = G*(i1,j1) and
A350:   [i2,j2] in Indices G & (F.k)/.len(F.k) = G*(i2,j2) and
A351:   i1 = i2 & j1+1 = j2 or i1+1 = i2 & j1 = j2 or
        i1 = i2+1 & j1 = j2 or i1 = i2 & j1 = j2+1 by A333,A348,JORDAN8:6;
       now per cases;
       suppose
A352:       front_left_cell(F.k,(len(F.k))-'1,G) misses C &
          front_right_cell(F.k,(len(F.k))-'1,G) misses C;
           now per cases by A351;
          suppose i1 = i2 & j1+1 = j2;
           then [i2+1,j2] in Indices G & F.(k+1) = (F.k)^<*G*(i2+1,j2)*>
             by A307,A347,A349,A350,A352;
            hence thesis;
          suppose i1+1 = i2 & j1 = j2;
           then [i2,j2-'1] in Indices G & F.(k+1) = (F.k)^<*G*(i2,j2-'1)*>
             by A307,A347,A349,A350,A352;
            hence thesis;
          suppose i1 = i2+1 & j1 = j2;
           then [i2,j2+1] in Indices G & F.(k+1) = (F.k)^<*G*(i2,j2+1)*>
             by A307,A347,A349,A350,A352;
            hence thesis;
          suppose i1 = i2 & j1 = j2+1;
           then [i2-'1,j2] in Indices G & F.(k+1) = (F.k)^<*G*(i2-'1,j2)*>
             by A307,A347,A349,A350,A352;
            hence thesis;
         end;
        hence thesis;
       suppose
A353:       front_left_cell(F.k,(len(F.k))-'1,G) misses C &
          front_right_cell(F.k,(len(F.k))-'1,G) meets C;
           now per cases by A351;
          suppose i1 = i2 & j1+1 = j2;
           then [i2,j2+1] in Indices G & F.(k+1) = (F.k)^<*G*(i2,j2+1)*>
             by A307,A347,A349,A350,A353;
            hence thesis;
          suppose i1+1 = i2 & j1 = j2;
           then [i2+1,j2] in Indices G & F.(k+1) = (F.k)^<*G*(i2+1,j2)*>
             by A307,A347,A349,A350,A353;
            hence thesis;
          suppose i1 = i2+1 & j1 = j2;
           then [i2-'1,j2] in Indices G & F.(k+1) = (F.k)^<*G*(i2-'1,j2)*>
             by A307,A347,A349,A350,A353;
            hence thesis;
          suppose i1 = i2 & j1 = j2+1;
           then [i2,j2-'1] in Indices G & F.(k+1) = (F.k)^<*G*(i2,j2-'1)*>
             by A307,A347,A349,A350,A353;
            hence thesis;
         end;
        hence thesis;
       suppose
A354:       front_left_cell(F.k,(len(F.k))-'1,G) meets C;
           now per cases by A351;
          suppose i1 = i2 & j1+1 = j2;
           then [i2-'1,j2] in Indices G & F.(k+1) = (F.k)^<*G*(i2-'1,j2)*>
             by A307,A347,A349,A350,A354;
            hence thesis;
          suppose i1+1 = i2 & j1 = j2;
           then [i2,j2+1] in Indices G & F.(k+1) = (F.k)^<*G*(i2,j2+1)*>
             by A307,A347,A349,A350,A354;
            hence thesis;
          suppose i1 = i2+1 & j1 = j2;
           then [i2,j2-'1] in Indices G & F.(k+1) = (F.k)^<*G*(i2,j2-'1)*>
             by A307,A347,A349,A350,A354;
            hence thesis;
          suppose i1 = i2 & j1 = j2+1;
           then [i2+1,j2] in Indices G & F.(k+1) = (F.k)^<*G*(i2+1,j2)*>
             by A307,A347,A349,A350,A354;
            hence thesis;
         end;
        hence thesis;
     end;
     hence thesis;
   end;
defpred P[Nat] means  for m st m <= $1 holds (F.$1)|m = F.m;
A355: P[0]
   proof let m;
    assume
A356:  m <= 0;
    A357: 0 <= m by NAT_1:18;
    then (F.0)|0 = (F.0)|m by A356,AXIOMS:21;
    hence thesis by A91,A356,A357,AXIOMS:21;
   end;
A358: for k st P[k] holds P[k+1]
   proof let k such that
A359:   for m st m <= k holds (F.k)|m = F.m;
    let m such that
A360:   m <= k+1;
    per cases by A360,REAL_1:def 5;
    suppose m < k+1;
then A361:    m <= k by NAT_1:38;
A362:    len(F.k) = k by A107;
     consider i,j such that
         [i,j] in Indices G and
A363:    F.(k+1) = F.k^<*G*(i,j)*> by A332;
       (F.(k+1))|m = (F.k)|m by A361,A362,A363,FINSEQ_5:25;
     hence (F.(k+1))|m = F.m by A359,A361;
    suppose
A364:    m = k+1;
       len(F.(k+1)) = k+1 by A107;
     hence (F.(k+1))|m = F.m by A364,TOPREAL1:2;
   end;

A365: for k holds P[k] from Ind(A355,A358);
A366: for k st k > 1 holds
       (front_left_cell(F.k,k-'1,Gauge(C,n)) misses C &
          front_right_cell(F.k,k-'1,Gauge(C,n)) misses C
         implies F.(k+1) turns_right k-'1,Gauge(C,n)) &
       (front_left_cell(F.k,k-'1,Gauge(C,n)) misses C &
          front_right_cell(F.k,k-'1,Gauge(C,n)) meets C
         implies F.(k+1) goes_straight k-'1,Gauge(C,n)) &
       (front_left_cell(F.k,k-'1,Gauge(C,n)) meets C
         implies F.(k+1) turns_left k-'1,Gauge(C,n))
   proof
     let k such that
A367:  k > 1;
A368:  F.k is_sequence_on G by A306;
A369:  len(F.k) = k by A107;
then A370:  1 <= (len(F.k))-'1 by A367,JORDAN3:12;
      (len(F.k)) -'1 +1 = len(F.k) by A367,A369,AMI_5:4;
    then ex i1,j1,i2,j2 being Nat st
         [i1,j1] in Indices G & (F.k)/.((len(F.k)) -'1) = G*(i1,j1) &
         [i2,j2] in Indices G & (F.k)/.len(F.k) = G*(i2,j2) &
         (i1 = i2 & j1+1 = j2 or i1+1 = i2 & j1 = j2 or
      i1 = i2+1 & j1 = j2 or i1 = i2 & j1 = j2+1) by A368,A370,JORDAN8:6;
    hence thesis by A307,A367,A369;
   end;
   defpred P[Nat] means F.$1 is unfolded;
     len(F.0) = 0 by A91,FINSEQ_1:25;
then A371: P[0] by SPPOL_2:27;
A372: for k st P[k] holds P[k+1]
   proof let k such that
A373:  F.k is unfolded;
A374:  F.k is_sequence_on G by A306;
    per cases;
    suppose k <= 1;
     then k+1 <= 1+1 by AXIOMS:24;
     then len(F.(k+1)) <= 2 by A107;
     hence F.(k+1) is unfolded by SPPOL_2:27;
    suppose
A375:   k > 1;
    set m = k-'1;
A376:   1 <= m by A375,JORDAN3:12;
A377:   m+1 = k by A375,AMI_5:4;
A378:   len(F.k) = k by A107;
     then consider i1,j1,i2,j2 being Nat such that
A379:   [i1,j1] in Indices G & (F.k)/.m = G*(i1,j1) and
A380:   [i2,j2] in Indices G & (F.k)/.len(F.k) = G*(i2,j2) and
A381:   i1 = i2 & j1+1 = j2 or i1+1 = i2 & j1 = j2 or
        i1 = i2+1 & j1 = j2 or i1 = i2 & j1 = j2+1 by A374,A376,A377,JORDAN8:6;
A382:  1 <= i1 & i1 <= len G & 1 <= j1 & j1 <= width G by A379,GOBOARD5:1;
A383:  1 <= i2 & i2 <= len G & 1 <= j2 & j2 <= width G by A380,GOBOARD5:1;
then A384:  (i2-'1)+1 = i2 by AMI_5:4;
A385:  (j2-'1)+1 = j2 by A383,AMI_5:4;
A386:   LSeg(F.k,m) = LSeg(G*(i1,j1),G*(i2,j2)) by A376,A377,A378,A379,A380,
TOPREAL1:def 5;
      now per cases;
     suppose
A387:     front_left_cell(F.k,(len(F.k))-'1,G) misses C &
          front_right_cell(F.k,(len(F.k))-'1,G) misses C;
        now per cases by A381;
       suppose
A388:       i1 = i2 & j1+1 = j2;
then A389:       [i2+1,j2] in Indices G & F.(k+1) = (F.k)^<*G*(i2+1,j2)*>
           by A307,A375,A378,A379,A380,A387;
         then i2+1 <= len G by GOBOARD5:1;
         then LSeg(F.k,m) /\ LSeg((F.k)/.len(F.k),G*
(i2+1,j2)) = {(F.k)/.len(F.k)}
          by A380,A382,A383,A386,A388,GOBOARD7:17;
        hence F.(k+1) is unfolded by A373,A377,A378,A389,SPPOL_2:31;
       suppose
A390:       i1+1 = i2 & j1 = j2;
then A391:       [i2,j2-'1] in Indices G & F.(k+1) = (F.k)^<*G*(i2,j2-'1)*>
           by A307,A375,A378,A379,A380,A387;
         then 1 <= j2-'1 &
          LSeg((F.k)/.len(F.k),G*(i2,j2-'1))=LSeg(G*(i2,j2-'1),
          (F.k)/.len(F.k)) by GOBOARD5:1;
         then LSeg(F.k,m) /\
 LSeg((F.k)/.len(F.k),G*(i2,j2-'1)) = {(F.k)/.len(F.k)}
           by A380,A382,A383,A385,A386,A390,GOBOARD7:18;
        hence F.(k+1) is unfolded by A373,A377,A378,A391,SPPOL_2:31;
       suppose
A392:       i1 = i2+1 & j1 = j2;
then A393:       [i2,j2+1] in Indices G & F.(k+1) = (F.k)^<*G*(i2,j2+1)*>
           by A307,A375,A378,A379,A380,A387;
         then j2+1 <= width G by GOBOARD5:1;
         then LSeg(F.k,m) /\ LSeg((F.k)/.len(F.k),G*
(i2,j2+1)) = {(F.k)/.len(F.k)}
          by A380,A382,A383,A386,A392,GOBOARD7:19;
        hence F.(k+1) is unfolded by A373,A377,A378,A393,SPPOL_2:31;
       suppose
A394:       i1 = i2 & j1 = j2+1;
then A395:       [i2-'1,j2] in Indices G & F.(k+1) = (F.k)^<*G*(i2-'1,j2)*>
           by A307,A375,A378,A379,A380,A387;
         then 1 <= i2-'1 by GOBOARD5:1;
         then LSeg(F.k,m) /\
 LSeg((F.k)/.len(F.k),G*(i2-'1,j2)) = {(F.k)/.len(F.k)}
          by A380,A382,A383,A384,A386,A394,GOBOARD7:20;
        hence F.(k+1) is unfolded by A373,A377,A378,A395,SPPOL_2:31;
      end;
      hence F.(k+1) is unfolded;
     suppose
A396:     front_left_cell(F.k,(len(F.k))-'1,G) misses C &
              front_right_cell(F.k,(len(F.k))-'1,G) meets C;
        now per cases by A381;
       suppose
A397:       i1 = i2 & j1+1 = j2;
then A398:       [i2,j2+1] in Indices G & F.(k+1) = (F.k)^<*G*(i2,j2+1)*>
           by A307,A375,A378,A379,A380,A396;
         then 1 <= j2+1 & j2+1 <= width G & j2+1 = j1+(1+1)
          by A397,GOBOARD5:1,XCMPLX_1:1;
         then LSeg(F.k,m) /\ LSeg((F.k)/.len(F.k),G*
(i2,j2+1)) = {(F.k)/.len(F.k)}
          by A380,A382,A386,A397,GOBOARD7:15;
        hence F.(k+1) is unfolded by A373,A377,A378,A398,SPPOL_2:31;
       suppose
A399:       i1+1 = i2 & j1 = j2;
then A400:       [i2+1,j2] in Indices G & F.(k+1) = (F.k)^<*G*(i2+1,j2)*>
           by A307,A375,A378,A379,A380,A396;
         then 1 <= i2+1 & i2+1 <= len G & i2+1 = i1+(1+1)
          by A399,GOBOARD5:1,XCMPLX_1:1;
         then LSeg(F.k,m) /\ LSeg((F.k)/.len(F.k),G*
(i2+1,j2)) = {(F.k)/.len(F.k)}
          by A380,A382,A386,A399,GOBOARD7:16;
        hence F.(k+1) is unfolded by A373,A377,A378,A400,SPPOL_2:31;
       suppose
A401:       i1 = i2+1 & j1 = j2;
then A402:       [i2-'1,j2] in Indices G & F.(k+1) = (F.k)^<*G*(i2-'1,j2)*>
           by A307,A375,A378,A379,A380,A396;
         then 1 <= i2-'1 & i2-'1+1+1 = i2-'1+(1+1) by GOBOARD5:1,XCMPLX_1:1;
         then LSeg(F.k,m) /\
 LSeg((F.k)/.len(F.k),G*(i2-'1,j2)) = {(F.k)/.len(F.k)}
          by A380,A382,A384,A386,A401,GOBOARD7:16;
        hence F.(k+1) is unfolded by A373,A377,A378,A402,SPPOL_2:31;
       suppose
A403:       i1 = i2 & j1 = j2+1;
then A404:       [i2,j2-'1] in Indices G & F.(k+1) = (F.k)^<*G*(i2,j2-'1)*>
           by A307,A375,A378,A379,A380,A396;
         then 1 <= j2-'1 & j2-'1+1+1 = j2-'1+(1+1)
          by GOBOARD5:1,XCMPLX_1:1;
         then LSeg(F.k,m) /\
 LSeg((F.k)/.len(F.k),G*(i2,j2-'1)) = {(F.k)/.len(F.k)}
          by A380,A382,A385,A386,A403,GOBOARD7:15;
        hence F.(k+1) is unfolded by A373,A377,A378,A404,SPPOL_2:31;
      end;
      hence F.(k+1) is unfolded;
     suppose
A405:     front_left_cell(F.k,(len(F.k))-'1,G) meets C;
        now per cases by A381;
       suppose
A406:       i1 = i2 & j1+1 = j2;
then A407:       [i2-'1,j2] in Indices G & F.(k+1) = (F.k)^<*G*(i2-'1,j2)*>
           by A307,A375,A378,A379,A380,A405;
         then 1 <= i2-'1 &
         LSeg((F.k)/.len(F.k),G*(i2-'1,j2))=LSeg(G*(i2-'1,j2),
         (F.k)/.len(F.k)) by GOBOARD5:1;
         then LSeg(F.k,m) /\
 LSeg((F.k)/.len(F.k),G*(i2-'1,j2)) = {(F.k)/.len(F.k)}
          by A380,A382,A383,A384,A386,A406,GOBOARD7:18;
        hence F.(k+1) is unfolded by A373,A377,A378,A407,SPPOL_2:31;
       suppose
A408:       i1+1 = i2 & j1 = j2;
then A409:       [i2,j2+1] in Indices G & F.(k+1) = (F.k)^<*G*(i2,j2+1)*>
           by A307,A375,A378,A379,A380,A405;
         then 1 <= j2+1 & j2+1 <= width G by GOBOARD5:1;
         then LSeg(F.k,m) /\ LSeg((F.k)/.len(F.k),G*
(i2,j2+1)) = {(F.k)/.len(F.k)}
          by A380,A382,A383,A386,A408,GOBOARD7:20;
        hence F.(k+1) is unfolded by A373,A377,A378,A409,SPPOL_2:31;
       suppose
A410:       i1 = i2+1 & j1 = j2;
then A411:       [i2,j2-'1] in Indices G & F.(k+1) = (F.k)^<*G*(i2,j2-'1)*>
           by A307,A375,A378,A379,A380,A405;
         then 1 <= j2-'1 &
         LSeg((F.k)/.len(F.k),G*(i2,j2-'1))=LSeg(G*(i2,j2-'1),
         (F.k)/.len(F.k)) by GOBOARD5:1;
         then LSeg(F.k,m) /\
 LSeg((F.k)/.len(F.k),G*(i2,j2-'1)) = {(F.k)/.len(F.k)}
          by A380,A382,A383,A385,A386,A410,GOBOARD7:17;
        hence F.(k+1) is unfolded by A373,A377,A378,A411,SPPOL_2:31;
       suppose
A412:       i1 = i2 & j1 = j2+1;
then A413:       [i2+1,j2] in Indices G & F.(k+1) = (F.k)^<*G*(i2+1,j2)*>
           by A307,A375,A378,A379,A380,A405;
         then 1 <= i2+1 & i2+1 <= len G &
         LSeg((F.k)/.len(F.k),G*(i2+1,j2))=LSeg(G*
(i2+1,j2),(F.k)/.len(F.k))
          by GOBOARD5:1;
         then LSeg(F.k,m) /\ LSeg((F.k)/.len(F.k),G*
(i2+1,j2)) = {(F.k)/.len(F.k)}
          by A380,A382,A383,A386,A412,GOBOARD7:19;
        hence F.(k+1) is unfolded by A373,A377,A378,A413,SPPOL_2:31;
      end;
      hence F.(k+1) is unfolded;
    end;
    hence F.(k+1) is unfolded;
   end;

A414: for k holds P[k] from Ind(A371,A372);
defpred Q[Nat] means
$1 >= 1 & ex m st m in dom(F.$1) &
m <> len(F.$1) & (F.$1)/.m = (F.$1)/.len(F.$1);
now assume
A415:   for k st k >= 1 holds
        for m st m in dom(F.k) & m <> len(F.k)
         holds (F.k)/.m <> (F.k)/.len(F.k);
         defpred P[Nat] means F.$1 is one-to-one;
A416:    P[0] by A91;
A417:    for k st P[k] holds P[k+1]
     proof let k; assume
A418:    F.k is one-to-one;
        now let n,m such that
A419:     n in dom(F.(k+1)) & m in dom(F.(k+1)) and
A420:     (F.(k+1))/.n = (F.(k+1))/.m;
       consider i,j such that
          [i,j] in Indices G and
A421:     F.(k+1) = (F.k)^<*G*(i,j)*> by A332;
A422:     1 <= n & n <= len(F.(k+1)) & 1 <= m & m <= len(F.(k+1))
         by A419,FINSEQ_3:27;
A423:     len(F.k) = k & len(F.(k+1)) = k+1 by A107;
A424:     1 <= k+1 by NAT_1:37;
       per cases by A422,A423,NAT_1:26;
        suppose n <= k & m <= k;
then A425:        n in dom(F.k) & m in dom(F.k) by A422,A423,FINSEQ_3:27;
         then (F.(k+1))/.n = (F.k)/.n & (F.(k+1))/.m = (F.k)/.m
          by A421,GROUP_5:95;
         hence n = m by A418,A420,A425,PARTFUN2:17;
        suppose n = k+1 & m <= k;
         hence n = m by A415,A419,A420,A423,A424;
        suppose n <= k & m = k+1;
         hence n = m by A415,A419,A420,A423,A424;
        suppose n = k+1 & m = k+1;
         hence n = m;
      end;
      hence F.(k+1) is one-to-one by PARTFUN2:16;
     end;
A426:   for k holds P[k] from Ind(A416,A417);
A427:   for k holds card rng(F.k) = k
     proof let k;
        F.k is one-to-one by A426;
      hence card rng(F.k) = len(F.k) by FINSEQ_4:77
                         .= k by A107;
     end;
    set k = (len G)*(width G)+1;
A428:   k > (len G)*(width G) by NAT_1:38;
A429:   card Values G <= (len G)*(width G) by GOBRD13:8;
      F.k is_sequence_on G by A306;
    then rng(F.k) c= Values G by GOBRD13:9;
    then card rng(F.k) <= card Values G by CARD_1:80;
    then card rng(F.k) <= (len G)*(width G) by A429,AXIOMS:22;
    hence contradiction by A427,A428;
   end; then A430: ex k st Q[k];
   consider k such that
A431:  Q[k] and
A432:  for l st Q[l] holds k <= l from Min(A430);

   consider m such that
A433:  m in dom(F.k) and
A434:  m <> len(F.k) and
A435:  (F.k)/.m = (F.k)/.len(F.k) by A431;
A436:  len(F.k) = k by A107;
    0<k by A431,AXIOMS:22;
  then reconsider f = F.k as non empty FinSequence of TOP-REAL 2 by A436,
FINSEQ_1:25;
A437: f is_sequence_on G by A306;
A438: 1 <= m & m <= len f by A433,FINSEQ_3:27;
then A439: m < len f by A434,REAL_1:def 5;
   then 1 < len f by A438,AXIOMS:22;
then A440: len f >= 1+1 by NAT_1:38;
  then reconsider f as non constant special unfolded non empty
    FinSequence of TOP-REAL 2 by A414,A437,JORDAN8:7,8;
  set g = f/^(m-'1);
A441: m+1 <= len f by A439,NAT_1:38;
    m-'1 <= m & m < m+1 by JORDAN3:7,NAT_1:38;
  then m-'1 < m+1 by AXIOMS:22;
then A442: m-'1 < len f by A441,AXIOMS:22;
then A443: len g = len f - (m-'1) by RFINSEQ:def 2;
  then (m-'1)-(m-'1) < len g by A442,REAL_1:54;
     then len g > 0 by XCMPLX_1:14;
  then reconsider g as non empty FinSequence of TOP-REAL 2 by FINSEQ_1:25;
    1 in dom g by FINSEQ_5:6;
then A444:  g/.1 = f/.(m-'1+1) by FINSEQ_5:30
           .= f/.m by A438,AMI_5:4;
    len g in dom g by FINSEQ_5:6;
then A445:  g/.len g = f/.(m-'1+len g) by FINSEQ_5:30
                .= f/.len f by A443,XCMPLX_1:27;
A446: g is_sequence_on G by A437,JORDAN8:5;
then A447: g is standard by JORDAN8:7;
A448: for i st 1 <= i & i < len g & 1 <= j & j < len g & g/.i = g/.j
      holds i = j
   proof let i such that
A449:  1 <= i & i < len g and
A450:  1 <= j & j < len g and
A451:  g/.i = g/.j and
A452:  i <> j;
       i in dom g by A449,FINSEQ_3:27;
then A453:  g/.i = f/.(m-'1+i) by FINSEQ_5:30;
       j in dom g by A450,FINSEQ_3:27;
then A454:  g/.j = f/.(m-'1+j) by FINSEQ_5:30;
A455:  0 <= m-'1 by NAT_1:18;
    per cases by A452,REAL_1:def 5;
    suppose
A456:    i < j;
     set l = m-'1+j, m'= m-'1+i;
A457:    m' < l by A456,REAL_1:53;
A458:    l < k by A436,A443,A450,REAL_1:86;
then A459:    f|l = F.l by A365;
       0+j <= l by A455,AXIOMS:24;
then A460:    1 <= l by A450,AXIOMS:22;
A461:    len(F.l) = l by A107;
       0+i <= m' by A455,AXIOMS:24;
     then 1 <= m' by A449,AXIOMS:22;
then A462:    m' in dom(F.l) by A457,A461,FINSEQ_3:27;
then A463:    (F.l)/.m' = f/.m' by A459,TOPREAL1:1;
       l in dom(F.l) by A460,A461,FINSEQ_3:27;
     then (F.l)/.l = f/.l by A459,TOPREAL1:1;
     hence contradiction by A432,A451,A453,A454,A457,A458,A460,A461,A462,A463;
    suppose
A464:    j < i;
     set l = m-'1+i, m'= m-'1+j;
A465:    m' < l by A464,REAL_1:53;
A466:    l < k by A436,A443,A449,REAL_1:86;
then A467:    f|l = F.l by A365;
       0+i <= l by A455,AXIOMS:24;
then A468:    1 <= l by A449,AXIOMS:22;
A469:    len(F.l) = l by A107;
       0+j <= m' by A455,AXIOMS:24;
     then 1 <= m' by A450,AXIOMS:22;
then A470:    m' in dom(F.l) by A465,A469,FINSEQ_3:27;
then A471:    (F.l)/.m' = f/.m' by A467,TOPREAL1:1;
       l in dom(F.l) by A468,A469,FINSEQ_3:27;
     then (F.l)/.l = f/.l by A467,TOPREAL1:1;
     hence contradiction by A432,A451,A453,A454,A465,A466,A468,A469,A470,A471;
   end;
A472: for i st 1 < i & i < j & j <= len g holds g/.i <> g/.j
   proof let i such that
A473:   1 < i and
A474:   i < j and
A475:   j <= len g and
A476:   g/.i = g/.j;
A477:   1 < j by A473,A474,AXIOMS:22;
A478:   i < len g by A474,A475,AXIOMS:22;
then A479:   1 < len g by A473,AXIOMS:22;
    per cases;
    suppose j <> len g;
     then j < len g by A475,REAL_1:def 5;
     hence contradiction by A448,A473,A474,A476,A477,A478;
    suppose j = len g;
     hence contradiction by A435,A444,A445,A448,A473,A474,A476,A479;
   end;
A480: for i st 1 <= i & i < j & j < len g holds g/.i <> g/.j
   proof let i such that
A481:   1 <= i and
A482:   i < j and
A483:   j < len g and
A484:   g/.i = g/.j;
A485:   1 < j by A481,A482,AXIOMS:22;
      i < len g by A482,A483,AXIOMS:22;
    hence contradiction by A448,A481,A482,A483,A484,A485;
   end;
A486: for i st 1 <= i & i+1 <= len f
       holds right_cell(f,i,G) = Cl Int right_cell(f,i,G)
   proof let i such that
A487:    1 <= i & i+1 <= len f;
     consider i1,j1,i2,j2 such that
A488:    [i1,j1] in Indices G and
A489:    f/.i = G*(i1,j1) and
A490:    [i2,j2] in Indices G and
A491:    f/.(i+1) = G*(i2,j2) and
A492:    i1 = i2 & j1+1 = j2 or i1+1 = i2 & j1 = j2 or
        i1 = i2+1 & j1 = j2 or i1 = i2 & j1 = j2+1 by A437,A487,JORDAN8:6;
A493:    1 <= i1 & i1 <= len G & 1 <= j1 & j1 <= width G by A488,GOBOARD5:1;
A494:    1 <= i2 & i2 <= len G & 1 <= j2 & j2 <= width G by A490,GOBOARD5:1;
A495:    i1+1 > i1 & j1+1 > j1 & i2+1 > i2 & j2+1 > j2 by NAT_1:38;
    per cases by A492;
    suppose
        i1 = i2 & j1+1 = j2;
     then right_cell(f,i,G) = cell(G,i1,j1)
      by A437,A487,A488,A489,A490,A491,A495,GOBRD13:def 2;
     hence thesis by A493,GOBRD11:35;
    suppose
A496:   i1+1 = i2 & j1 = j2;
A497:   j1-'1 <= width G by A493,JORDAN3:7;
       right_cell(f,i,G) = cell(G,i1,j1-'1)
      by A437,A487,A488,A489,A490,A491,A495,A496,GOBRD13:def 2;
     hence thesis by A493,A497,GOBRD11:35;
    suppose
        i1 = i2+1 & j1 = j2;
     then right_cell(f,i,G) = cell(G,i2,j2)
      by A437,A487,A488,A489,A490,A491,A495,GOBRD13:def 2;
     hence thesis by A494,GOBRD11:35;
    suppose
A498:   i1 = i2 & j1 = j2+1;
A499:   i1-'1 <= len G by A493,JORDAN3:7;
       right_cell(f,i,G) = cell(G,i1-'1,j2)
      by A437,A487,A488,A489,A490,A491,A495,A498,GOBRD13:def 2;
     hence thesis by A494,A499,GOBRD11:35;
   end;
A500: g is s.c.c.
   proof
    let i,j such that
A501:   i+1 < j and
A502:   i > 1 & j < len g or j+1 < len g;
A503:   j <= j+1 by NAT_1:37;
then A504:  i+1 < j+1 by A501,AXIOMS:22;
A505:   1 <= i+1 by NAT_1:37;
A506:   1 < j by A501,NAT_1:37;
     i < j by A501,NAT_1:38;
then A507:    i < j+1 by A503,AXIOMS:22;
    per cases by A502,RLVECT_1:99;
    suppose
A508:    i > 1 & j < len g;
then A509:    i+1 < len g by A501,AXIOMS:22;
then A510:    LSeg(g,i) = LSeg(g/.i,g/.(i+1)) by A508,TOPREAL1:def 5;
A511:   1 < i+1 by A508,NAT_1:38;
A512:    i < len g by A509,NAT_1:38;
     consider i1,j1,i2,j2 such that
A513:    [i1,j1] in Indices G and
A514:    g/.i = G*(i1,j1) and
A515:    [i2,j2] in Indices G and
A516:    g/.(i+1) = G*(i2,j2) and
A517:    i1 = i2 & j1+1 = j2 or i1+1 = i2 & j1 = j2 or
        i1 = i2+1 & j1 = j2 or i1 = i2 & j1 = j2+1 by A446,A508,A509,JORDAN8:6;
A518:    1 <= i1 & i1 <= len G & 1 <= j1 & j1 <= width G by A513,GOBOARD5:1;
A519:    1 <= i2 & i2 <= len G & 1 <= j2 & j2 <= width G by A515,GOBOARD5:1;
A520:    j+1 <= len g by A508,NAT_1:38;
then A521:    LSeg(g,j) = LSeg(g/.j,g/.(j+1)) by A506,TOPREAL1:def 5;
     consider i1',j1',i2',j2' being Nat such that
A522:    [i1',j1'] in Indices G and
A523:    g/.j = G*(i1',j1') and
A524:    [i2',j2'] in Indices G and
A525:    g/.(j+1) = G*(i2',j2') and
A526:    i1' = i2' & j1'+1 = j2' or i1'+1 = i2' & j1' = j2' or
        i1' = i2'+1 & j1' = j2' or i1' = i2' & j1' = j2'+1
         by A446,A506,A520,JORDAN8:6;
A527:    1 <= i1' & i1' <= len G & 1 <= j1' & j1' <= width G by A522,GOBOARD5:1
;
A528:    1 <= i2' & i2' <= len G & 1 <= j2' & j2' <= width G by A524,GOBOARD5:1
;
     assume
A529:  LSeg(g,i) meets LSeg(g,j);
        now per cases by A517,A526;
       suppose
A530:      i1 = i2 & j1+1 = j2 & i1' = i2' & j1'+1 = j2';
then A531:      i1 = i1' by A510,A514,A516,A518,A519,A521,A523,A525,A527,A528,
A529,GOBOARD7:21;
           j1 = j1' or j1 = j1'+1 or j1+1 = j1'
           by A510,A514,A516,A518,A519,A521,A523,A525,A527,A528,A529,A530,
GOBOARD7:24;
        hence contradiction
           by A448,A472,A501,A503,A505,A506,A507,A508,A509,A512,A514,A516,A520,
A523,A525,A530,A531;
       suppose
A532:      i1 = i2 & j1+1 = j2 & i1'+1 = i2' & j1' = j2';
         then i1 = i1' & j1 = j1' or i1 = i1' & j1+1 = j1' or
              i1 = i1'+1 & j1 = j1' or i1 = i1'+1 & j1+1 = j1'
          by A510,A514,A516,A518,A519,A521,A523,A525,A527,A528,A529,GOBOARD7:23
;
        hence contradiction
            by A448,A472,A501,A504,A506,A507,A508,A511,A512,A514,A516,A520,A523
,A525,A532;
       suppose
A533:      i1 = i2 & j1+1 = j2 & i1' = i2'+1 & j1' = j2';
         then i1 = i2' & j1' = j1 or i1 = i2' & j1+1 = j1' or
          i1 = i2'+1 & j1' = j1 or i1 = i2'+1 & j1+1 = j1'
            by A510,A514,A516,A518,A519,A521,A523,A525,A527,A528,A529,GOBOARD7:
23;
        hence contradiction
         by A448,A472,A501,A504,A506,A507,A508,A511,A512,A514,A516,A520,A523,
A525,A533;
       suppose
A534:      i1 = i2 & j1+1 = j2 & i1' = i2' & j1' = j2'+1;
then A535:      i1 = i1' by A510,A514,A516,A518,A519,A521,A523,A525,A527,A528,
A529,GOBOARD7:21;
          j1 = j2' or j1 = j2'+1 or j1+1 = j2'
              by A510,A514,A516,A518,A519,A521,A523,A525,A527,A528,A529,A534,
GOBOARD7:24;
        hence contradiction by A448,A472,A501,A504,A506,A508,A511,A512,A514,
A516,A520,A523,A525,A534,A535;
       suppose
A536:      i1+1 = i2 & j1 = j2 & i1' = i2' & j1'+1 = j2';
        then i1' = i1 & j1 = j1' or i1' = i1 & j1'+1 = j1 or
         i1' = i1+1 & j1 = j1' or i1' = i1+1 & j1'+1 = j1
            by A510,A514,A516,A518,A519,A521,A523,A525,A527,A528,A529,GOBOARD7:
23;
        hence contradiction
            by A448,A472,A501,A504,A506,A507,A508,A511,A512,A514,A516,A520,A523
,A525,A536;
       suppose
A537:      i1+1 = i2 & j1 = j2 & i1'+1 = i2' & j1' = j2';
then A538:      j1 = j1' by A510,A514,A516,A518,A519,A521,A523,A525,A527,A528,
A529,GOBOARD7:22;
          i1 = i1' or i1 = i1'+1 or i1+1 = i1'
         by A510,A514,A516,A518,A519,A521,A523,A525,A527,A528,A529,A537,
GOBOARD7:25;
        hence contradiction
         by A448,A472,A501,A503,A505,A506,A507,A508,A509,A512,A514,A516,A520,
A523,A525,A537,A538;
       suppose
A539:      i1+1 = i2 & j1 = j2 & i1' = i2'+1 & j1' = j2';
then A540:      j1 = j1' by A510,A514,A516,A518,A519,A521,A523,A525,A527,A528,
A529,GOBOARD7:22;
          i1 = i2' or i1 = i2'+1 or i1+1 = i2'
         by A510,A514,A516,A518,A519,A521,A523,A525,A527,A528,A529,A539,
GOBOARD7:25;
        hence contradiction
         by A448,A472,A504,A506,A507,A508,A511,A512,A514,A516,A520,A523,A525,
A539,A540;
       suppose
A541:      i1+1 = i2 & j1 = j2 & i1' = i2' & j1' = j2'+1;
         then i1' = i1 & j1 = j2' or i1' = i1 & j2'+1 = j1 or
              i1' = i1+1 & j1 = j2' or i1' = i1+1 & j2'+1 = j1
             by A510,A514,A516,A518,A519,A521,A523,A525,A527,A528,A529,GOBOARD7
:23;
        hence contradiction
         by A448,A472,A501,A504,A506,A507,A508,A511,A512,A514,A516,A520,A523,
A525,A541;
       suppose
A542:      i1 = i2+1 & j1 = j2 & i1' = i2' & j1'+1 = j2';
         then i1' = i2 & j1' = j1 or i1' = i2 & j1'+1 = j1 or
              i1' = i2+1 & j1' = j1 or i1' = i2+1 & j1'+1 = j1
          by A510,A514,A516,A518,A519,A521,A523,A525,A527,A528,A529,GOBOARD7:23
;
        hence contradiction
         by A448,A472,A501,A504,A506,A507,A508,A511,A512,A514,A516,A520,A523,
A525,A542;
       suppose
A543:      i1 = i2+1 & j1 = j2 & i1'+1 = i2' & j1' = j2';
then A544:      j1 = j1' by A510,A514,A516,A518,A519,A521,A523,A525,A527,A528,
A529,GOBOARD7:22;
           i2 = i1' or i2 = i1'+1 or i2+1 = i1'
            by A510,A514,A516,A518,A519,A521,A523,A525,A527,A528,A529,A543,
GOBOARD7:25;
        hence contradiction
          by A448,A472,A501,A504,A506,A508,A511,A512,A514,A516,A520,A523,A525,
A543,A544;
       suppose
A545:      i1 = i2+1 & j1 = j2 & i1' = i2'+1 & j1' = j2';
then A546:      j1 = j1' by A510,A514,A516,A518,A519,A521,A523,A525,A527,A528,
A529,GOBOARD7:22;
           i2 = i2' or i2 = i2'+1 or i2+1 = i2'
           by A510,A514,A516,A518,A519,A521,A523,A525,A527,A528,A529,A545,
GOBOARD7:25;
        hence contradiction by A472,A501,A504,A507,A508,A511,A514,A516,A520,
A523,A525,A545,A546;
       suppose
A547:      i1 = i2+1 & j1 = j2 & i1' = i2' & j1' = j2'+1;
         then i1' = i2 & j2' = j1 or i1' = i2 & j2'+1 = j1 or
         i1' = i2+1 & j2' = j1 or i1' = i2+1 & j2'+1 = j1
           by A510,A514,A516,A518,A519,A521,A523,A525,A527,A528,A529,GOBOARD7:
23;
        hence contradiction
          by A448,A472,A501,A504,A506,A507,A508,A511,A512,A514,A516,A520,A523,
A525,A547;
       suppose
A548:      i1 = i2 & j1 = j2+1 & i1' = i2' & j1'+1 = j2';
then A549:      i1 = i1' by A510,A514,A516,A518,A519,A521,A523,A525,A527,A528,
A529,GOBOARD7:21;
           j2 = j1' or j2 = j1'+1 or j2+1 = j1'
           by A510,A514,A516,A518,A519,A521,A523,A525,A527,A528,A529,A548,
GOBOARD7:24;
        hence contradiction
          by A448,A472,A501,A504,A506,A508,A511,A512,A514,A516,A520,A523,A525,
A548,A549;
       suppose
A550:      i1 = i2 & j1 = j2+1 & i1'+1 = i2' & j1' = j2';
         then i1 = i1' & j2 = j1' or i1 = i1' & j2+1 = j1' or
             i1 = i1'+1 & j2 = j1' or i1 = i1'+1 & j2+1 = j1'
          by A510,A514,A516,A518,A519,A521,A523,A525,A527,A528,A529,GOBOARD7:23
;
        hence contradiction
          by A448,A472,A501,A504,A506,A507,A508,A511,A512,A514,A516,A520,A523,
A525,A550;
       suppose
A551:      i1 = i2 & j1 = j2+1 & i1' = i2'+1 & j1' = j2';
         then i1 = i2' & j2 = j1' or i1 = i2' & j2+1 = j1' or
             i1 = i2'+1 & j2 = j1' or i1 = i2'+1 & j2+1 = j1'
          by A510,A514,A516,A518,A519,A521,A523,A525,A527,A528,A529,GOBOARD7:23
;
        hence contradiction
          by A448,A472,A501,A504,A506,A507,A508,A511,A512,A514,A516,A520,A523,
A525,A551;
       suppose
A552:      i1 = i2 & j1 = j2+1 & i1' = i2' & j1' = j2'+1;
then A553:      i1 = i1' by A510,A514,A516,A518,A519,A521,A523,A525,A527,A528,
A529,GOBOARD7:21;
           j2 = j2' or j2 = j2'+1 or j2+1 = j2'
           by A510,A514,A516,A518,A519,A521,A523,A525,A527,A528,A529,A552,
GOBOARD7:24;
        hence contradiction by A472,A501,A504,A507,A508,A511,A514,A516,A520,
A523,A525,A552,A553;
      end;
     hence contradiction;
    suppose i = 0 & j+1 < len g;
     then LSeg(g,i) = {} by TOPREAL1:def 5;
     hence LSeg(g,i) misses LSeg(g,j) by XBOOLE_1:65;
    suppose
A554:   1 <= i & j+1 < len g;
then A555:   i+1 < len g by A504,AXIOMS:22;
then A556:    LSeg(g,i) = LSeg(g/.i,g/.(i+1)) by A554,TOPREAL1:def 5;
A557:   1 < i+1 by A554,NAT_1:38;
A558:    i < len g by A555,NAT_1:38;
A559:   j < len g by A554,NAT_1:37;
     consider i1,j1,i2,j2 such that
A560:    [i1,j1] in Indices G and
A561:    g/.i = G*(i1,j1) and
A562:    [i2,j2] in Indices G and
A563:    g/.(i+1) = G*(i2,j2) and
A564:    i1 = i2 & j1+1 = j2 or i1+1 = i2 & j1 = j2 or
        i1 = i2+1 & j1 = j2 or i1 = i2 & j1 = j2+1 by A446,A554,A555,JORDAN8:6;
A565:    1 <= i1 & i1 <= len G & 1 <= j1 & j1 <= width G by A560,GOBOARD5:1;
A566:    1 <= i2 & i2 <= len G & 1 <= j2 & j2 <= width G by A562,GOBOARD5:1;
A567:    LSeg(g,j) = LSeg(g/.j,g/.(j+1)) by A506,A554,TOPREAL1:def 5;
     consider i1',j1',i2',j2' being Nat such that
A568:    [i1',j1'] in Indices G and
A569:    g/.j = G*(i1',j1') and
A570:    [i2',j2'] in Indices G and
A571:    g/.(j+1) = G*(i2',j2') and
A572:    i1' = i2' & j1'+1 = j2' or i1'+1 = i2' & j1' = j2' or
        i1' = i2'+1 & j1' = j2' or i1' = i2' & j1' = j2'+1
         by A446,A506,A554,JORDAN8:6;
A573:    1 <= i1' & i1' <= len G & 1 <= j1' & j1' <= width G by A568,GOBOARD5:1
;
A574:    1 <= i2' & i2' <= len G & 1 <= j2' & j2' <= width G by A570,GOBOARD5:1
;
     assume
A575: LSeg(g,i) meets LSeg(g,j);
        now per cases by A564,A572;
       suppose
A576:      i1 = i2 & j1+1 = j2 & i1' = i2' & j1'+1 = j2';
then A577:      i1 = i1' by A556,A561,A563,A565,A566,A567,A569,A571,A573,A574,
A575,GOBOARD7:21;
           j1 = j1' or j1 = j1'+1 or j1+1 = j1'
           by A556,A561,A563,A565,A566,A567,A569,A571,A573,A574,A575,A576,
GOBOARD7:24;
        hence contradiction by A480,A501,A504,A505,A507,A554,A559,A561,A563,
A569,A571,A576,A577;
       suppose
A578:      i1 = i2 & j1+1 = j2 & i1'+1 = i2' & j1' = j2';
         then i1 = i1' & j1 = j1' or i1 = i1' & j1+1 = j1' or
            i1 = i1'+1 & j1 = j1' or i1 = i1'+1 & j1+1 = j1'
           by A556,A561,A563,A565,A566,A567,A569,A571,A573,A574,A575,GOBOARD7:
23;
        hence contradiction
           by A448,A480,A501,A504,A506,A507,A554,A557,A558,A559,A561,A563,A569,
A571,A578;
       suppose
A579:      i1 = i2 & j1+1 = j2 & i1' = i2'+1 & j1' = j2';
         then i1 = i2' & j1' = j1 or i1 = i2' & j1+1 = j1' or
              i1 = i2'+1 & j1' = j1 or i1 = i2'+1 & j1+1 = j1'
           by A556,A561,A563,A565,A566,A567,A569,A571,A573,A574,A575,GOBOARD7:
23;
        hence contradiction
           by A448,A480,A501,A504,A506,A507,A554,A557,A558,A559,A561,A563,A569,
A571,A579;
       suppose
A580:      i1 = i2 & j1+1 = j2 & i1' = i2' & j1' = j2'+1;
then A581:      i1 = i1' by A556,A561,A563,A565,A566,A567,A569,A571,A573,A574,
A575,GOBOARD7:21;
           j1 = j2' or j1 = j2'+1 or j1+1 = j2'
           by A556,A561,A563,A565,A566,A567,A569,A571,A573,A574,A575,A580,
GOBOARD7:24;
        hence contradiction by A448,A472,A501,A504,A506,A554,A557,A558,A559,
A561,A563,A569,A571,A580,A581;
       suppose
A582:      i1+1 = i2 & j1 = j2 & i1' = i2' & j1'+1 = j2';
         then i1' = i1 & j1 = j1' or i1' = i1 & j1'+1 = j1 or
              i1' = i1+1 & j1 = j1' or i1' = i1+1 & j1'+1 = j1
           by A556,A561,A563,A565,A566,A567,A569,A571,A573,A574,A575,GOBOARD7:
23;
        hence contradiction by A448,A480,A501,A504,A505,A506,A507,A554,A558,
A559,A561,A563,A569,A571,A582;
       suppose
A583:      i1+1 = i2 & j1 = j2 & i1'+1 = i2' & j1' = j2';
then A584:      j1 = j1' by A556,A561,A563,A565,A566,A567,A569,A571,A573,A574,
A575,GOBOARD7:22;
           i1 = i1' or i1 = i1'+1 or i1+1 = i1'
           by A556,A561,A563,A565,A566,A567,A569,A571,A573,A574,A575,A583,
GOBOARD7:25;
        hence contradiction by A480,A501,A504,A505,A507,A554,A559,A561,A563,
A569,A571,A583,A584;
       suppose
A585:      i1+1 = i2 & j1 = j2 & i1' = i2'+1 & j1' = j2';
then A586:      j1 = j1' by A556,A561,A563,A565,A566,A567,A569,A571,A573,A574,
A575,GOBOARD7:22;
           i1 = i2' or i1 = i2'+1 or i1+1 = i2'
           by A556,A561,A563,A565,A566,A567,A569,A571,A573,A574,A575,A585,
GOBOARD7:25;
        hence contradiction by A448,A480,A504,A506,A507,A554,A557,A558,A559,
A561,A563,A569,A571,A585,A586;
       suppose
A587:      i1+1 = i2 & j1 = j2 & i1' = i2' & j1' = j2'+1;
         then i1' = i1 & j1 = j2' or i1' = i1 & j2'+1 = j1 or
              i1' = i1+1 & j1 = j2' or i1' = i1+1 & j2'+1 = j1
           by A556,A561,A563,A565,A566,A567,A569,A571,A573,A574,A575,GOBOARD7:
23;
        hence contradiction by A448,A480,A501,A504,A506,A507,A554,A557,A558,
A559,A561,A563,A569,A571,A587;
       suppose
A588:      i1 = i2+1 & j1 = j2 & i1' = i2' & j1'+1 = j2';
         then i1' = i2 & j1' = j1 or i1' = i2 & j1'+1 = j1 or
              i1' = i2+1 & j1' = j1 or i1' = i2+1 & j1'+1 = j1
          by A556,A561,A563,A565,A566,A567,A569,A571,A573,A574,A575,GOBOARD7:23
;
        hence contradiction by A448,A480,A501,A504,A506,A507,A554,A557,A558,
A559,A561,A563,A569,A571,A588;
       suppose
A589:      i1 = i2+1 & j1 = j2 & i1'+1 = i2' & j1' = j2';
then A590:      j1 = j1' by A556,A561,A563,A565,A566,A567,A569,A571,A573,A574,
A575,GOBOARD7:22;
           i2 = i1' or i2 = i1'+1 or i2+1 = i1'
            by A556,A561,A563,A565,A566,A567,A569,A571,A573,A574,A575,A589,
GOBOARD7:25;
        hence contradiction by A448,A472,A501,A504,A506,A554,A557,A558,A559,
A561,A563,A569,A571,A589,A590;
       suppose
A591:      i1 = i2+1 & j1 = j2 & i1' = i2'+1 & j1' = j2';
then A592:      j1 = j1' by A556,A561,A563,A565,A566,A567,A569,A571,A573,A574,
A575,GOBOARD7:22;
           i2 = i2' or i2 = i2'+1 or i2+1 = i2'
           by A556,A561,A563,A565,A566,A567,A569,A571,A573,A574,A575,A591,
GOBOARD7:25;
        hence contradiction by A480,A501,A504,A507,A554,A557,A559,A561,A563,
A569,A571,A591,A592;
       suppose
A593:      i1 = i2+1 & j1 = j2 & i1' = i2' & j1' = j2'+1;
         then i1' = i2 & j2' = j1 or i1' = i2 & j2'+1 = j1 or
              i1' = i2+1 & j2' = j1 or i1' = i2+1 & j2'+1 = j1
            by A556,A561,A563,A565,A566,A567,A569,A571,A573,A574,A575,GOBOARD7:
23;
        hence contradiction by A448,A480,A501,A504,A506,A507,A554,A557,A558,
A559,A561,A563,A569,A571,A593;
       suppose
A594:      i1 = i2 & j1 = j2+1 & i1' = i2' & j1'+1 = j2';
then A595:      i1 = i1' by A556,A561,A563,A565,A566,A567,A569,A571,A573,A574,
A575,GOBOARD7:21;
           j2 = j1' or j2 = j1'+1 or j2+1 = j1'
           by A556,A561,A563,A565,A566,A567,A569,A571,A573,A574,A575,A594,
GOBOARD7:24;
        hence contradiction by A448,A472,A501,A504,A506,A554,A557,A558,A559,
A561,A563,A569,A571,A594,A595;
       suppose
A596:      i1 = i2 & j1 = j2+1 & i1'+1 = i2' & j1' = j2';
         then i1 = i1' & j2 = j1' or i1 = i1' & j2+1 = j1' or
              i1 = i1'+1 & j2 = j1' or i1 = i1'+1 & j2+1 = j1'
           by A556,A561,A563,A565,A566,A567,A569,A571,A573,A574,A575,GOBOARD7:
23;
        hence contradiction by A448,A480,A501,A504,A506,A507,A554,A557,A558,
A559,A561,A563,A569,A571,A596;
       suppose
A597:      i1 = i2 & j1 = j2+1 & i1' = i2'+1 & j1' = j2';
         then i1 = i2' & j2 = j1' or i1 = i2' & j2+1 = j1' or
              i1 = i2'+1 & j2 = j1' or i1 = i2'+1 & j2+1 = j1'
           by A556,A561,A563,A565,A566,A567,A569,A571,A573,A574,A575,GOBOARD7:
23;
        hence contradiction by A448,A480,A501,A504,A506,A507,A554,A557,A558,
A559,A561,A563,A569,A571,A597;
       suppose
A598:      i1 = i2 & j1 = j2+1 & i1' = i2' & j1' = j2'+1;
then A599:      i1 = i1' by A556,A561,A563,A565,A566,A567,A569,A571,A573,A574,
A575,GOBOARD7:21;
           j2 = j2' or j2 = j2'+1 or j2+1 = j2'
           by A556,A561,A563,A565,A566,A567,A569,A571,A573,A574,A575,A598,
GOBOARD7:24;
        hence contradiction by A480,A501,A504,A507,A554,A557,A559,A561,A563,
A569,A571,A598,A599;
      end;
     hence contradiction;
   end;
    m+1-(m-'1) <= len g by A441,A443,REAL_1:49;
  then m+1-(m-1) <= len g by A438,SCMFSA_7:3;
  then 1+m-m+1 <= len g by XCMPLX_1:37;
then A600:  1+1 <= len g by XCMPLX_1:26;
    g is non constant
   proof
    take 1,2;
    thus
A601:   1 in dom g by FINSEQ_5:6;
    thus
A602:   2 in dom g by A600,FINSEQ_3:27;
    then g/.1 <> g/.(1+1) by A447,A601,GOBOARD7:31;
    then g.1 <> g/.(1+1) by A601,FINSEQ_4:def 4;
    hence g.1 <> g.2 by A602,FINSEQ_4:def 4;
   end;
  then reconsider g as standard non constant special_circular_sequence
   by A435,A444,A445,A446,A500,FINSEQ_6:def 1,JORDAN8:7;
    now
    assume m <> 1;
     then A603:  1 < m by A438,REAL_1:def 5;
A604:  LeftComp g is_a_component_of (L~g)` &
       RightComp g is_a_component_of (L~g)` by GOBOARD9:def 1,def 2;
    reconsider Lg' = (L~g)` as Subset of TOP-REAL 2;

A605:   C c= Lg'
    proof
     let c be set;
     assume that
A606:   c in C and
A607:   not c in Lg';
     reconsider c as Point of TOP-REAL 2 by A606;
       c in L~g by A607,SUBSET_1:50;
     then consider i such that
A608:   1 <= i & i+1 <= len g and
A609:   c in LSeg(g/.i,g/.(i+1)) by SPPOL_2:14;
       i in dom g & i+1 in dom g by A608,GOBOARD2:3;
then A610:   g/.i = f/.(i+(m-'1)) & g/.(i+1) = f/.(i+1+(m-'1)) by FINSEQ_5:30;
A611:   1 <= i+(m-'1) by A608,NAT_1:37;
A612:   i+1+(m-'1) = i+(m-'1)+1 by XCMPLX_1:1;
     then i+(m-'1)+1 <= (len g)+(m-'1) by A608,AXIOMS:24;
then A613:   i+(m-'1)+1 <= len f by A443,XCMPLX_1:27;
     then c in LSeg(f,i+(m-'1)) by A609,A610,A611,A612,TOPREAL1:def 5;
     then c in left_cell(f,i+(m-'1),G) /\ right_cell(f,i+(m-'1),G)
      by A437,A611,A613,GOBRD13:30;
     then c in left_cell(f,i+(m-'1),G) by XBOOLE_0:def 3;
     then left_cell(f,i+(m-'1),G) meets C by A606,XBOOLE_0:3;
     hence contradiction by A306,A611,A613;
    end;
A614:   TOP-REAL 2 = TopSpaceMetr (Euclid 2) by EUCLID:def 8;
A615:   the carrier of TOP-REAL 2 = the carrier of Euclid 2 by TOPREAL3:13;
      L~g is closed by SPPOL_1:49;
    then (L~g)` is open by TOPS_1:29;
    then Lg' is open;
then A616:  (L~g)` = Int (L~g)` by TOPS_1:55;

A617: for j,k st 1 <= j & j <= k holds (F.k)/.j = (F.j)/.j
  proof
    let j,k;
    assume that
  A618: 1 <= j and
  A619: j <= k;
  A620: (F.k)|j = F.j by A365,A619;
        j <= len(F.k) by A107,A619;
  then len(F.k|j) = j by TOPREAL1:3;
      then j in dom((F.k)|j) by A618,FINSEQ_3:27;
      hence thesis by A620,TOPREAL1:1;
  end;

A621:  L~g c= L~f by JORDAN3:75;

A622: for n st 1 <= n & n <= m-'1 holds not f/.n in L~g
  proof
    let n such that
  A623: 1 <= n and
  A624: n <= m-'1;
  A625: n <= len f by A442,A624,AXIOMS:22;

    set p = f/.n;
    assume p in L~g;
    then consider j such that
  A626: m-'1+1 <= j and
  A627: j+1 <= len f and
  A628: p in LSeg(f,j) by A442,Th9;

  A629: n < m-'1+1 by A624,NAT_1:38;
  then A630: n < j by A626,AXIOMS:22;

  A631: m-'1+1=m by A438,AMI_5:4;
  then A632: 1 < j by A603,A626,AXIOMS:22;
  A633: j < k by A436,A627,NAT_1:38;
 A634: 2 <= len G & 2 <= width G by A2,NAT_1:37;
A635: p in Values G by A437,A623,A625,Th8;
    per cases by A437,A627,A628,A632,A634,A635,Th25;
    suppose
    A636: p = f/.j;
        n <= len(F.j) by A107,A630;
    then A637: n in dom(F.j) by A623,FINSEQ_3:27;
    A638: n <> len(F.j) by A107,A626,A629;
        (F.j)/.n = (F.n)/.n by A617,A623,A630
              .= p by A436,A617,A623,A625
              .= (F.j)/.j by A617,A632,A633,A636
              .= (F.j)/.len(F.j) by A107;
      hence contradiction by A432,A632,A633,A637,A638;

    suppose
    A639: p = f/.(j+1);
        now
        per cases by A436,A627,REAL_1:def 5;
        suppose
        A640: j+1 = k;
            n <= len(F.m) by A107,A629,A631;
        then A641: n in dom(F.m) by A623,FINSEQ_3:27;
        A642: n <> len(F.m) by A107,A629,A631;
            (F.m)/.n = (F.n)/.n by A617,A623,A629,A631
                  .= (F.k)/.k by A436,A617,A623,A625,A639,A640
                  .= (F.m)/.m by A435,A436,A438,A617
                  .= (F.m)/.len(F.m) by A107;
          hence contradiction by A432,A436,A438,A439,A641,A642;
        suppose
        A643: j+1 < k;
          set l = j+1;
        A644: 1 <= l by NAT_1:29;
          A645: n < n+1 & n+1 < l by A630,REAL_1:53,69;
        then A646: n < l by AXIOMS:22;
          then n <= len(F.l) by A107;
        then A647: n in dom(F.l) by A623,FINSEQ_3:27;
        A648: n <> len(F.l) by A107,A645;
            (F.l)/.n = (F.n)/.n by A617,A623,A646
                  .= p by A436,A617,A623,A625
                  .= (F.l)/.l by A617,A639,A643,A644
                  .= (F.l)/.len(F.l) by A107;
          hence contradiction by A432,A643,A644,A647,A648;
      end;
      hence contradiction;
  end;

A649: for h being standard non constant special_circular_sequence st L~h c= L~f
    for Comp being Subset of TOP-REAL 2 st Comp is_a_component_of (L~h)`
    for n st 1 <= n & n+1 <= len f & f/.n in Comp & not f/.n in L~h
      holds
     C meets Comp
  proof
    let h be standard non constant special_circular_sequence such that
  A650: L~h c= L~f;
    let Comp be Subset of TOP-REAL 2 such that
  A651: Comp is_a_component_of (L~h)`;
    let n such that
  A652: 1 <= n & n+1 <= len f and
  A653: f/.n in Comp and
  A654: not f/.n in L~h;

    reconsider rc = right_cell(f,n,G)\L~h as Subset of TOP-REAL 2;
  A655: rc c= right_cell(f,n,G) by XBOOLE_1:36;
  A656: rc = right_cell(f,n,G) /\ (L~h)` by SUBSET_1:32;

  A657: rc meets C
  proof
      right_cell(f,n,G) meets C by A306,A652;
    then consider p being set such that
  A658: p in right_cell(f,n,G) & p in C by XBOOLE_0:3;
    reconsider p as Element of TOP-REAL 2 by A658;
      now take a = p;
      now
      assume p in L~h;
      then consider j such that
    A659: 1 <= j and
    A660: j+1 <= len f and
    A661: p in LSeg(f,j) by A650,SPPOL_2:13;
        p in
 left_cell(f,j,G) /\ right_cell(f,j,G) by A437,A659,A660,A661,GOBRD13:30;
    then A662: p in left_cell(f,j,G) by XBOOLE_0:def 3;
        left_cell(f,j,G) misses C by A306,A659,A660;
      hence contradiction by A658,A662,XBOOLE_0:3;
    end;
    hence a in rc by A658,XBOOLE_0:def 4;
    thus a in C by A658;
  end; hence thesis by XBOOLE_0:3; end;

  A663: Int right_cell(f,n,G) is connected by A437,A652,Th12;
      Int right_cell(f,n,G) misses L~f by A437,A652,Th17;
    then Int right_cell(f,n,G) misses L~h by A650,XBOOLE_1:63;
  then A664: Int right_cell(f,n,G) c= (L~h)` by SUBSET_1:43;
      Int right_cell(f,n,G) c= right_cell(f,n,G) by TOPS_1:44;
  then A665: Int right_cell(f,n,G) c= rc by A656,A664,XBOOLE_1:19;
      rc c= Cl Int right_cell(f,n,G) by A437,A652,A655,Th13;
  then A666: rc is connected by A663,A665,CONNSP_1:19;
      f/.n in right_cell(f,n,G) by A437,A652,Th10;
    then f/.n in rc by A654,XBOOLE_0:def 4;
  then A667: rc meets Comp by A653,XBOOLE_0:3;
      rc c= (L~h)` by A656,XBOOLE_1:17;
    then rc c= Comp by A651,A666,A667,GOBOARD9:6;
    hence C meets Comp by A657,XBOOLE_1:63;
  end;

A668:
 C meets RightComp Rev g
  proof
    set rg=Rev g;

  A669: rg is_sequence_on G by A446,Th7;

  A670: 1+1 <= len rg by A600,FINSEQ_5:def 3;
    then 1+1-'1 <= len rg-'1 by JORDAN3:5;
  then A671: 1 <= len rg -'1 by BINARITH:39;
      1 < len rg by A670,NAT_1:38;
  then A672: len rg -'1+1 = len rg by AMI_5:4;

    set p = rg/.1, q = rg/.2;
    set a = f/.(m-'1);
    set l = (m-'1)+(len g-'1);

  A673: p = f/.m by A435,A445,FINSEQ_5:68;
  A674: m-'1+1 = m by A438,AMI_5:4;
  then A675: 1 <= m-'1 by A603,NAT_1:38;

      1+1-'1 <= len g-'1 by A600,JORDAN3:5;
  then A676: 1 <= len g-'1 by BINARITH:39;
      1+1 - 1 <= len g - 1 by A600,REAL_1:49;
  then A677: 0 <= len g - 1 by AXIOMS:22;
A678: 1 - 1 < m - 1 by A603,REAL_1:54;
      l = m+(len g -'1)-'1 by A438,JORDAN4:3
     .= (len g -'1)+m - 1 by A676,JORDAN4:2
     .= (len g - 1)+m - 1 by A677,BINARITH:def 3
     .= ((k - (m - 1)) - 1)+m - 1 by A436,A443,A678,BINARITH:def 3
     .= ((k - m+1) - 1)+m - 1 by XCMPLX_1:37
     .= (k - m)+m - 1 by XCMPLX_1:26
     .= (k+-m+m) - 1 by XCMPLX_0:def 8
     .= k - 1 by XCMPLX_1:139;
  then A679: k = l+1 by XCMPLX_1:27;
  then A680: l < k by REAL_1:69;
      len g-'1 <= l by NAT_1:29;
  then A681: 1 <= l by A676,AXIOMS:22;
      len g-'1 <= len g by JORDAN3:7;
  then A682: len g-'1 in dom g by A676,FINSEQ_3:27;
      1 <= len g by A600,SPPOL_1:5;
  then A683: len g-'1+2 = len g+1 by Lm1;
  then A684: q = g/.(len g-'1) by A682,FINSEQ_5:69
       .= f/.l by A682,FINSEQ_5:30;
  A685: right_cell(f,m-'1,G) meets C by A306,A438,A674,A675;
  A686: right_cell(f,l,G) meets C by A306,A436,A679,A681;

      not a in L~g by A622,A675;
  then A687: not a in L~rg by SPPOL_2:22;

  A688: m-'1+2 = m+1 by A438,Lm1;

  A689: m-'1 <= l by NAT_1:29;
    then m-'1 <= len(F.l) by A107;
  then A690: m-'1 in dom(F.l) by A675,FINSEQ_3:27;
      (m-'1)+1 <= l by A676,AXIOMS:24;
    then m-'1 < l by NAT_1:38;
  then A691: m-'1 <> len(F.l) by A107;

    consider p1,p2,q1,q2 being Nat such that
  A692: [p1,p2] in Indices G and
  A693: p = G*(p1,p2) and
  A694: [q1,q2] in Indices G and
  A695: q = G*(q1,q2) and
  A696: p1 = q1 & p2+1 = q2 or p1+1 = q1 & p2 = q2 or
        p1 = q1+1 & p2 = q2 or p1 = q1 & p2 = q2+1 by A669,A670,JORDAN8:6;

  A697: 1 <= p1 & p1 <= len G & 1 <= p2 & p2 <= width G by A692,GOBOARD5:1;

    per cases by A696;
    suppose
    A698: p1 = q1 & p2+1 = q2;
      consider a1,a2,p'1,p'2 being Nat such that
    A699: [a1,a2] in Indices G and
    A700: a = G*(a1,a2) and
    A701: [p'1,p'2] in Indices G and
    A702: p = G*(p'1,p'2) and
    A703: a1 = p'1 & a2+1 = p'2 or a1+1 = p'1 & a2 = p'2 or
          a1 = p'1+1 & a2 = p'2 or a1 = p'1 & a2 = p'2+1
            by A437,A438,A673,A674,A675,JORDAN8:6;
    A704: 1 <= a1 & a1 <= len G & 1 <= a2 & a2 <= width G by A699,GOBOARD5:1;

      thus thesis
      proof
        per cases by A703;
        suppose
            a1 = p'1 & a2+1 = p'2;
        then A705: a1 = p1 & a2+1 = p2 by A692,A693,A701,A702,GOBOARD1:21;
        A706: F.m is_sequence_on G by A306;
        A707: m-'1+1 <= len (F.m) by A107,A674;
        A708: m-'1 <= m by A674,NAT_1:29;
        A709: F.k|(m+1)=F.(m+1) by A365,A436,A441;
A710:      p2+1 > a2+1 & a2+1 > a2 by A705,NAT_1:38;
        A711: f/.(m-'1) = (F.(m-'1))/.(m-'1) by A436,A442,A617,A675
                     .= (F.m)/.(m-'1) by A617,A675,A708;
        A712: f/.(m-'1+1) = (F.m)/.m by A436,A438,A617,A674;
            right_cell(f,l,G) = cell(G,p1-'1,p2)
                 by A435,A436,A437,A673,A679,A681,A684,A692,A693,A694,A695,A698
,GOBRD13:29
            .= front_left_cell(F.m,m-'1,G) by A673,A674,A675,A692,A693,A699,
A700,A705,A706,A707,A711,A712,GOBRD13:35;
          then F.(m+1) turns_left m-'1,G by A366,A603,A686;
          then f turns_left m-'1,G by A441,A675,A688,A709,GOBRD13:45;
        then A713: f/.(m+1) = G*(p1-'1,p2) & [p1-'1,p2] in Indices G
            by A673,A674,A688,A692,A693,A699,A700,A710,GOBRD13:def 7;
        then A714: 1 <= p1-'1 & p1-'1 <= len G by GOBOARD5:1;

          set rc = right_cell(rg,len rg-'1,G)\L~rg;
        A715: rc c= RightComp rg by A669,A671,A672,Th29;
        A716: p2-'1 = a2 & p1-'1+1 = p1 by A697,A705,AMI_5:4,BINARITH:39;
        A717: 2 in dom g by A600,FINSEQ_3:27;
             len rg-'1+2 = len g +1 by A683,FINSEQ_5:def 3;
        then A718: rg/.(len rg-'1) = g/.2 by A717,FINSEQ_5:69
                           .= f/.(m+1) by A688,A717,FINSEQ_5:30;
          p = g/.1 by A435,A444,A445,FINSEQ_5:68
              .= rg/.len g by FINSEQ_5:68
              .= rg/.len rg by FINSEQ_5:def 3;
          then right_cell(rg,len rg-'1,G) = cell(G,p1-'1,a2)
            by A669,A671,A672,A692,A693,A713,A716,A718,GOBRD13:25;
          then a in right_cell(rg,len rg-'1,G)
            by A697,A700,A704,A705,A714,A716,Th22;
then A719:       a in rc by A687,XBOOLE_0:def 4;
        A720: L~rg c= L~f by A621,SPPOL_2:22;
            RightComp rg is_a_component_of (L~rg)` by GOBOARD9:def 2;
          hence C meets RightComp Rev g by A438,A649,A674,A675,A687,A715,A719,
A720;

        suppose
          A721: a1+1 = p'1 & a2 = p'2;
        then A722: a1+1 = p1 & a2 = p2 by A692,A693,A701,A702,GOBOARD1:21;
        then q1-'1 = a1 by A698,BINARITH:39;
          then right_cell(f,l,G) = cell(G,a1,a2)
                 by A435,A436,A437,A673,A679,A681,A684,A692,A693,A694,A695,A698
,A722,GOBRD13:29
            .= left_cell(f,m-'1,G)
                 by A437,A438,A673,A674,A675,A699,A700,A701,A702,A721,GOBRD13:
24;
          hence C meets RightComp Rev g by A306,A438,A674,A675,A686;

        suppose
            a1 = p'1+1 & a2 = p'2;
        then a1 = p1+1 & a2 = p2 by A692,A693,A701,A702,GOBOARD1:21;
          then right_cell(f,m-'1,G) = cell(G,p1,p2)
                 by A437,A438,A673,A674,A675,A692,A693,A699,A700,GOBRD13:27
            .= left_cell(f,l,G)
                 by A435,A436,A437,A673,A679,A681,A684,A692,A693,A694,A695,A698
,GOBRD13:28;
          hence C meets RightComp Rev g by A306,A436,A679,A681,A685;

        suppose
            a1 = p'1 & a2 = p'2+1;
        then A723: a1 = q1 & a2 = q2 by A692,A693,A698,A701,A702,GOBOARD1:21;
            (F.l)/.(m-'1) = (F.(m-'1))/.(m-'1) by A617,A675,A689
                     .= q by A436,A442,A617,A675,A695,A700,A723
                     .= (F.l)/.l by A617,A680,A681,A684
                     .= (F.l)/.len(F.l) by A107;
          hence C meets RightComp Rev g by A432,A680,A681,A690,A691;
      end;

    suppose
    A724: p1+1 = q1 & p2 = q2;
      consider a1,a2,p'1,p'2 being Nat such that
    A725: [a1,a2] in Indices G and
    A726: a = G*(a1,a2) and
    A727: [p'1,p'2] in Indices G and
    A728: p = G*(p'1,p'2) and
    A729: a1 = p'1 & a2+1 = p'2 or a1+1 = p'1 & a2 = p'2 or
          a1 = p'1+1 & a2 = p'2 or a1 = p'1 & a2 = p'2+1
            by A437,A438,A673,A674,A675,JORDAN8:6;
    A730: 1 <= a1 & a1 <= len G & 1 <= a2 & a2 <= width G by A725,GOBOARD5:1;

      thus thesis
      proof
        per cases by A729;
        suppose
          A731: a1 = p'1 & a2+1 = p'2;
        then A732: a1 = p1 & a2+1 = p2 by A692,A693,A727,A728,GOBOARD1:21;
        then A733: q2-'1 = a2 by A724,BINARITH:39;
            right_cell(f,m-'1,G) = cell(G,a1,a2)
                 by A437,A438,A673,A674,A675,A725,A726,A727,A728,A731,GOBRD13:
23
            .= left_cell(f,l,G)
                by A435,A436,A437,A673,A679,A681,A684,A692,A693,A694,A695,A724,
A732,A733,GOBRD13:26;
          hence C meets RightComp Rev g by A306,A436,A679,A681,A685;

        suppose
            a1+1 = p'1 & a2 = p'2;
        then A734: a1+1 = p1 & a2 = p2 by A692,A693,A727,A728,GOBOARD1:21;
        A735: F.m is_sequence_on G by A306;
        A736: m-'1+1 <= len (F.m) by A107,A674;
        A737: m-'1 <= m by A674,NAT_1:29;
        A738: F.k|(m+1)=F.(m+1) by A365,A436,A441;
A739:      a1 < a1+1 & p1 < p1+1 by REAL_1:69;
        A740: f/.(m-'1) = (F.(m-'1))/.(m-'1) by A436,A442,A617,A675
                     .= (F.m)/.(m-'1) by A617,A675,A737;
        A741: f/.(m-'1+1) = (F.m)/.m by A436,A438,A617,A674;
            right_cell(f,l,G) = cell(G,p1,p2)
                 by A435,A436,A437,A673,A679,A681,A684,A692,A693,A694,A695,A724
,GOBRD13:27
            .= front_left_cell(F.m,m-'1,G) by A673,A674,A675,A692,A693,A725,
A726,A734,A735,A736,A740,A741,GOBRD13:37;
          then F.(m+1) turns_left m-'1,G by A366,A603,A686;
          then f turns_left m-'1,G by A441,A675,A688,A738,GOBRD13:45;
        then A742: f/.(m+1) = G*(p1,p2+1) & [p1,p2+1] in Indices G
            by A673,A674,A688,A692,A693,A725,A726,A734,A739,GOBRD13:def 7;
        then A743: p2+1 <= width G by GOBOARD5:1;

          set rc = right_cell(rg,len rg-'1,G)\L~rg;
        A744: rc c= RightComp rg by A669,A671,A672,Th29;
        A745: a1 = p1-'1 by A734,BINARITH:39;
        A746: 2 in dom g by A600,FINSEQ_3:27;
             len rg-'1+2 = len g +1 by A683,FINSEQ_5:def 3;
        then A747: rg/.(len rg-'1) = g/.2 by A746,FINSEQ_5:69
                           .= f/.(m+1) by A688,A746,FINSEQ_5:30;
             p = g/.1 by A435,A444,A445,FINSEQ_5:68
              .= rg/.len g by FINSEQ_5:68
              .= rg/.len rg by FINSEQ_5:def 3;
          then right_cell(rg,len rg-'1,G) = cell(G,p1-'1,a2)
            by A669,A671,A672,A692,A693,A734,A742,A747,GOBRD13:29;
          then a in right_cell(rg,len rg-'1,G)
            by A697,A726,A730,A734,A743,A745,Th22;
then A748:       a in rc by A687,XBOOLE_0:def 4;
        A749: L~rg c= L~f by A621,SPPOL_2:22;
            RightComp rg is_a_component_of (L~rg)` by GOBOARD9:def 2;
          hence C meets RightComp Rev g by A438,A649,A674,A675,A687,A744,A748,
A749;

        suppose
            a1 = p'1+1 & a2 = p'2;
        then A750: a1 = q1 & a2 = q2 by A692,A693,A724,A727,A728,GOBOARD1:21;
            (F.l)/.(m-'1) = (F.(m-'1))/.(m-'1) by A617,A675,A689
                     .= q by A436,A442,A617,A675,A695,A726,A750
                     .= (F.l)/.l by A617,A680,A681,A684
                     .= (F.l)/.len(F.l) by A107;
          hence C meets RightComp Rev g by A432,A680,A681,A690,A691;

        suppose
            a1 = p'1 & a2 = p'2+1;
        then A751: a1 = p1 & a2 = p2+1 by A692,A693,A727,A728,GOBOARD1:21;
            right_cell(f,l,G) = cell(G,p1,p2)
                 by A435,A436,A437,A673,A679,A681,A684,A692,A693,A694,A695,A724
,GOBRD13:27
            .= left_cell(f,m-'1,G)
                 by A437,A438,A673,A674,A675,A692,A693,A725,A726,A751,GOBRD13:
28;
          hence C meets RightComp Rev g by A306,A438,A674,A675,A686;
      end;

    suppose
    A752: p1 = q1+1 & p2 = q2;
      consider a1,a2,p'1,p'2 being Nat such that
    A753: [a1,a2] in Indices G and
    A754: a = G*(a1,a2) and
    A755: [p'1,p'2] in Indices G and
    A756: p = G*(p'1,p'2) and
    A757: a1 = p'1 & a2+1 = p'2 or a1+1 = p'1 & a2 = p'2 or
          a1 = p'1+1 & a2 = p'2 or a1 = p'1 & a2 = p'2+1
            by A437,A438,A673,A674,A675,JORDAN8:6;
    A758: 1 <= a1 & a1 <= len G & 1 <= a2 & a2 <= width G by A753,GOBOARD5:1;

      thus thesis
      proof
        per cases by A757;
        suppose
          A759: a1 = p'1 & a2+1 = p'2;
        then A760: a1 = p1 & a2+1 = p2 by A692,A693,A755,A756,GOBOARD1:21;
        then A761: q1 = a1-'1 by A752,BINARITH:39;
        A762: q2-'1 = a2 by A752,A760,BINARITH:39;
            right_cell(f,l,G) = cell(G,q1,q2-'1)
              by A435,A436,A437,A673,A679,A681,A684,A692,A693,A694,A695,A752,
GOBRD13:25
            .= left_cell(f,m-'1,G)
              by A437,A438,A673,A674,A675,A753,A754,A755,A756,A759,A761,A762,
GOBRD13:22;
          hence C meets RightComp Rev g by A306,A438,A674,A675,A686;

        suppose
            a1+1 = p'1 & a2 = p'2;
          then a1+1 = p1 & a2 = p2 by A692,A693,A755,A756,GOBOARD1:21;
        then A763: a1 = q1 & a2 = q2 by A752,XCMPLX_1:2;
            (F.l)/.(m-'1) = (F.(m-'1))/.(m-'1) by A617,A675,A689
                     .= q by A436,A442,A617,A675,A695,A754,A763
                     .= (F.l)/.l by A617,A680,A681,A684
                     .= (F.l)/.len(F.l) by A107;
          hence C meets RightComp Rev g by A432,A680,A681,A690,A691;

        suppose
            a1 = p'1+1 & a2 = p'2;
        then A764: a1 = p1+1 & a2 = p2 by A692,A693,A755,A756,GOBOARD1:21;
        A765: p1-'1 = q1 by A752,BINARITH:39;
        A766: F.m is_sequence_on G by A306;
        A767: m-'1+1 <= len (F.m) by A107,A674;
        A768: m-'1 <= m by A674,NAT_1:29;
        A769: F.k|(m+1)=F.(m+1) by A365,A436,A441;
             p1+1>p1 by REAL_1:69;
        then A770: a2+1 > p2 & a1+1 > p1 & a2 < p2+1 by A764,NAT_1:38;
        A771: f/.(m-'1) = (F.(m-'1))/.(m-'1) by A436,A442,A617,A675
                     .= (F.m)/.(m-'1) by A617,A675,A768;
        A772: f/.(m-'1+1) = (F.m)/.m by A436,A438,A617,A674;
            right_cell(f,l,G) = cell(G,q1,q2-'1)
                 by A435,A436,A437,A673,A679,A681,A684,A692,A693,A694,A695,A752
,GOBRD13:25
            .= front_left_cell(F.m,m-'1,G) by A673,A674,A675,A692,A693,A752,
A753,A754,A764,A765,A766,A767,A771,A772,GOBRD13:39
;
          then F.(m+1) turns_left m-'1,G by A366,A603,A686;
          then f turns_left m-'1,G by A441,A675,A688,A769,GOBRD13:45;
        then A773: f/.(m+1) = G*(p1,p2-'1) & [p1,p2-'1] in Indices G
            by A673,A674,A688,A692,A693,A753,A754,A770,GOBRD13:def 7;
        then A774: 1 <= p2-'1 & p2-'1 <= width G by GOBOARD5:1;

          set rc = right_cell(rg,len rg-'1,G)\L~rg;
        A775: rc c= RightComp rg by A669,A671,A672,Th29;
        A776: p2-'1+1 = p2 by A697,AMI_5:4;
        A777: 2 in dom g by A600,FINSEQ_3:27;
             len rg-'1+2 = len g +1 by A683,FINSEQ_5:def 3;
        then A778: rg/.(len rg-'1) = g/.2 by A777,FINSEQ_5:69
                           .= f/.(m+1) by A688,A777,FINSEQ_5:30;
              p = g/.1 by A435,A444,A445,FINSEQ_5:68
              .= rg/.len g by FINSEQ_5:68
              .= rg/.len rg by FINSEQ_5:def 3;
          then right_cell(rg,len rg-'1,G) = cell(G,p1,p2-'1)
            by A669,A671,A672,A692,A693,A773,A776,A778,GOBRD13:23;
          then a in right_cell(rg,len rg-'1,G)
            by A697,A754,A758,A764,A774,A776,Th22;
then A779:       a in rc by A687,XBOOLE_0:def 4;
        A780: L~rg c= L~f by A621,SPPOL_2:22;
            RightComp rg is_a_component_of (L~rg)` by GOBOARD9:def 2;
          hence C meets RightComp Rev g by A438,A649,A674,A675,A687,A775,A779,
A780;

        suppose
            a1 = p'1 & a2 = p'2+1;
        then A781: a1 = p1 & a2 = p2+1 by A692,A693,A755,A756,GOBOARD1:21;
        then q1 = a1-'1 by A752,BINARITH:39;
          then right_cell(f,m-'1,G) = cell(G,q1,q2)
                 by A437,A438,A673,A674,A675,A692,A693,A752,A753,A754,A781,
GOBRD13:29
            .= left_cell(f,l,G)
                 by A435,A436,A437,A673,A679,A681,A684,A692,A693,A694,A695,A752
,GOBRD13:24;
          hence C meets RightComp Rev g by A306,A436,A679,A681,A685;
      end;

    suppose
    A782: p1 = q1 & p2 = q2+1;
      consider a1,a2,p'1,p'2 being Nat such that
    A783: [a1,a2] in Indices G and
    A784: a = G*(a1,a2) and
    A785: [p'1,p'2] in Indices G and
    A786: p = G*(p'1,p'2) and
    A787: a1 = p'1 & a2+1 = p'2 or a1+1 = p'1 & a2 = p'2 or
          a1 = p'1+1 & a2 = p'2 or a1 = p'1 & a2 = p'2+1
            by A437,A438,A673,A674,A675,JORDAN8:6;
    A788: 1 <= a1 & a1 <= len G & 1 <= a2 & a2 <= width G by A783,GOBOARD5:1;

      thus thesis
      proof
        per cases by A787;
        suppose
            a1 = p'1 & a2+1 = p'2;
          then a1 = p1 & a2+1 = p2 by A692,A693,A785,A786,GOBOARD1:21;
        then A789: a1 = q1 & a2 = q2 by A782,XCMPLX_1:2;
            (F.l)/.(m-'1) = (F.(m-'1))/.(m-'1) by A617,A675,A689
                     .= q by A436,A442,A617,A675,A695,A784,A789
                     .= (F.l)/.l by A617,A680,A681,A684
                     .= (F.l)/.len(F.l) by A107;
          hence C meets RightComp Rev g by A432,A680,A681,A690,A691;

        suppose
          A790: a1+1 = p'1 & a2 = p'2;
        then A791: a1+1 = p1 & a2 = p2 by A692,A693,A785,A786,GOBOARD1:21;
        then A792: a1 = q1-'1 by A782,BINARITH:39;
        A793: a2-'1 = q2 by A782,A791,BINARITH:39;
            right_cell(f,m-'1,G) = cell(G,a1,a2-'1)
                 by A437,A438,A673,A674,A675,A783,A784,A785,A786,A790,GOBRD13:
25
            .= left_cell(f,l,G)
                by A435,A436,A437,A673,A679,A681,A684,A692,A693,A694,A695,A782,
A792,A793,GOBRD13:22;
          hence C meets RightComp Rev g by A306,A436,A679,A681,A685;

        suppose
            a1 = p'1+1 & a2 = p'2;
        then A794: a1 = p1+1 & a2 = p2 by A692,A693,A785,A786,GOBOARD1:21;
        then A795: a2-'1 = q2 by A782,BINARITH:39;
            right_cell(f,l,G) = cell(G,q1,q2)
                 by A435,A436,A437,A673,A679,A681,A684,A692,A693,A694,A695,A782
,GOBRD13:23
            .= left_cell(f,m-'1,G)
                 by A437,A438,A673,A674,A675,A692,A693,A782,A783,A784,A794,A795
,GOBRD13:26;
          hence C meets RightComp Rev g by A306,A438,A674,A675,A686;

        suppose
            a1 = p'1 & a2 = p'2+1;
        then A796: a1 = p1 & a2 = p2+1 by A692,A693,A785,A786,GOBOARD1:21;
        A797: p2-'1 = q2 by A782,BINARITH:39;
        A798: F.m is_sequence_on G by A306;
        A799: m-'1+1 <= len (F.m) by A107,A674;
        A800: m-'1 <= m by A674,NAT_1:29;
        A801: F.k|(m+1)=F.(m+1) by A365,A436,A441;
A802:      a2+1>p2+1 & p2+1>p2 by A796,NAT_1:38;
        A803: f/.(m-'1) = (F.(m-'1))/.(m-'1) by A436,A442,A617,A675
                     .= (F.m)/.(m-'1) by A617,A675,A800;
        A804: f/.(m-'1+1) = (F.m)/.m by A436,A438,A617,A674;
            right_cell(f,l,G) = cell(G,q1,q2)
                 by A435,A436,A437,A673,A679,A681,A684,A692,A693,A694,A695,A782
,GOBRD13:23
            .= front_left_cell(F.m,m-'1,G) by A673,A674,A675,A692,A693,A782,
A783,A784,A796,A797,A798,A799,A803,A804,GOBRD13:41
;
          then F.(m+1) turns_left m-'1,G by A366,A603,A686;
          then f turns_left m-'1,G by A441,A675,A688,A801,GOBRD13:45;
        then A805: f/.(m+1) = G*(p1+1,p2) & [p1+1,p2] in Indices G
            by A673,A674,A688,A692,A693,A783,A784,A802,GOBRD13:def 7;
        then A806: p1+1 <= len G by GOBOARD5:1;

          set rc = right_cell(rg,len rg-'1,G)\L~rg;
        A807: rc c= RightComp rg by A669,A671,A672,Th29;
        A808: 2 in dom g by A600,FINSEQ_3:27;
             len rg-'1+2 = len g +1 by A683,FINSEQ_5:def 3;
        then A809: rg/.(len rg-'1) = g/.2 by A808,FINSEQ_5:69
                           .= f/.(m+1) by A688,A808,FINSEQ_5:30;
          p = g/.1 by A435,A444,A445,FINSEQ_5:68
              .= rg/.len g by FINSEQ_5:68
              .= rg/.len rg by FINSEQ_5:def 3;
          then right_cell(rg,len rg-'1,G) = cell(G,p1,p2)
            by A669,A671,A672,A692,A693,A805,A809,GOBRD13:27;
          then a in right_cell(rg,len rg-'1,G)
            by A697,A784,A788,A796,A806,Th22;
then A810:       a in rc by A687,XBOOLE_0:def 4;
        A811: L~rg c= L~f by A621,SPPOL_2:22;
            RightComp rg is_a_component_of (L~rg)` by GOBOARD9:def 2;
          hence C meets RightComp Rev g by A438,A649,A674,A675,A687,A807,A810,
A811;
      end;
  end;
A812: C meets RightComp g
   proof
      right_cell(f,m,G) meets C by A306,A438,A441;
    then consider p being set such that
A813:  p in right_cell(f,m,G) & p in C by XBOOLE_0:3;
    reconsider p as Element of TOP-REAL 2 by A813;
      now take a = p;
    thus a in C by A813;
    reconsider u = p as Element of Euclid 2 by A615;
    consider r being real number such that
A814:   r > 0 and
A815:   Ball(u,r) c= (L~g)` by A605,A616,A813,GOBOARD6:8;
    reconsider r as Real by XREAL_0:def 1;
A816:   p in Ball(u,r) by A814,GOBOARD6:4;
      Ball(u,r) is Subset of TOP-REAL 2 by A614,TOPMETR:16;
    then reconsider B = Ball(u,r) as non empty Subset of TOP-REAL 2
       by A814,GOBOARD6:4;
A817:   B is connected by SPRECT_3:17;
      right_cell(g,1,G) c= right_cell(g,1) by A446,A600,GOBRD13:34;
then A818:   Int right_cell(g,1,G) c= Int right_cell(g,1) by TOPS_1:48;
      Int right_cell(g,1) c= RightComp g by A600,GOBOARD9:28;
    then Int right_cell(g,1,G) c= RightComp g by A818,XBOOLE_1:1;
    then Int right_cell(f,m-'1+1,G) c= RightComp g by A437,A442,A600,GOBRD13:33
;
    then A819:  Int right_cell(f,m,G) c= RightComp g by A438,AMI_5:4;
A820:   right_cell(f,m,G) = Cl Int right_cell(f,m,G) by A438,A441,A486;
      B is open by GOBOARD6:6;
then A821:   Int right_cell(f,m,G) meets B by A813,A816,A820,TOPS_1:39;
A822:   p in B by A814,GOBOARD6:4;
      B c= RightComp g by A604,A815,A817,A819,A821,GOBOARD9:6;
    hence a in RightComp g by A822;
   end; hence thesis by XBOOLE_0:3; end;
      C meets LeftComp g by A668,GOBOARD9:27;
    then LeftComp g = RightComp g by A1,A604,A605,A812,Th3;
    hence contradiction by SPRECT_4:7;
  end;
then A823: g = f/^0 by GOBOARD9:1
      .= f by FINSEQ_5:31;
A824:
  now
   consider i such that
A825:   1 <= i & i+1 <= len G and
A826:   N-min C in cell(G,i,width G-'1) & N-min C <> G*(i,width G-'1) and
A827:   F.(1+1) = <*G*(i,width G),G*(i+1,width G)*> by A92;
   take i;
   thus 1 <= i & i+1 <= len G by A825;
A828:  f|2 = F.2 by A365,A436,A440;
A829:  len(f|2) = 2 by A440,TOPREAL1:3;
   then 1 in dom(f|2) by FINSEQ_3:27;
   hence f/.1 = (f|2)/.1 by TOPREAL1:1
               .= G*(i,width G) by A827,A828,FINSEQ_4:26;
     2 in dom(f|2) by A829,FINSEQ_3:27;
   hence f/.2 = (f|2)/.2 by TOPREAL1:1
               .= G*(i+1,width G) by A827,A828,FINSEQ_4:26;
   thus N-min C in cell(G,i,width G-'1) & N-min C <> G*
(i,width G-'1) by A826;
  end;
  reconsider f as standard non constant special_circular_sequence by A823;
    f is clockwise_oriented
  proof
      len G >= 3 by A2,NAT_1:37;
then A830:  1 < len G by AXIOMS:22;
   for k st 1 <= k & k+1 <= len f
      holds left_cell(f,k,G) misses C & right_cell(f,k,G) meets C by A306;
    then A831:   N-min L~f = f/.1 by A437,A824,Th32;
      f/.2 in LSeg(f/.1,f/.(1+1)) by TOPREAL1:6;
then A832:  f/.2 in L~f by A440,SPPOL_2:15;
    consider i such that
A833:  1 <= i & i+1 <= len G and
A834:  f/.1 = G*(i,width G) and
A835:  f/.2 = G*(i+1,width G) and
         N-min C in cell(G,i,width G-'1) & N-min C <> G*(i,width G-'1) by A824;
A836:  (NW-corner L~f)`2 = (NE-corner L~f)`2 by PSCOMP_1:82;
A837:  (NE-corner L~f)`2 = N-bound L~f & (N-min L~f)`2 = N-bound L~f
        by PSCOMP_1:77,94;
        i < len G by A833,NAT_1:38;
then A838:  G*(i,width G)`2 = G*(1,width G)`2 by A2,A830,A833,GOBOARD5:2;
        1 <= i+1 by NAT_1:37;
then A839:   G*(i+1,width G)`2 = G*(1,width G)`2 by A2,A830,A833,GOBOARD5:2;
      (NW-corner L~f)`1 = W-bound L~f & (NE-corner L~f)`1 = E-bound L~f
       by PSCOMP_1:74,76;
    then (NW-corner L~f)`1 <= (f/.2)`1 & (f/.2)`1 <= (NE-corner L~f)`1
       by A832,PSCOMP_1:71;
    then f/.2 in LSeg(NW-corner L~f, NE-corner L~f)
     by A831,A834,A835,A836,A837,A838,A839,GOBOARD7:9;
    then f/.2 in LSeg(NW-corner L~f, NE-corner L~f) /\ L~f by A832,XBOOLE_0:def
3;
    then f/.2 in N-most L~f by PSCOMP_1:def 39;
    hence thesis by A831,SPRECT_2:34;
   end;
  then reconsider f as
   clockwise_oriented (standard non constant special_circular_sequence);
  take f;
  thus f is_sequence_on G by A306;
  thus ex i st 1 <= i & i+1 <= len G &
    f/.1 = G*(i,width G) & f/.2 = G*(i+1,width G) &
     N-min C in cell(G,i,width G-'1) & N-min C <> G*(i,width G-'1) by A824;
   let m such that
A840:  1 <= m & m+2 <= len f;
     m+1 < m+2 by REAL_1:53;
then A841:  m+1 <= len f by A840,AXIOMS:22;
then A842:  f|(m+1) = F.(m+1) by A365,A436;
A843:  m = m+1-'1 by BINARITH:39;
A844:  m+1 > 1 by A840,NAT_1:38;
A845:  front_left_cell(F.(m+1),m,G) = front_left_cell(f,m,G) &
       front_right_cell(F.(m+1),m,G) = front_right_cell(f,m,G)
        by A437,A840,A841,A842,GOBRD13:43;
A846:  m+1+1 = m+(1+1) by XCMPLX_1:1;
then A847:  F.(m+1+1) = f|(m+1+1) by A365,A436,A840;
   hereby assume front_left_cell(f,m,G) misses C &
          front_right_cell(f,m,G) misses C;
     then F.(m+1+1) turns_right m,G by A366,A843,A844,A845;
    hence f turns_right m,G by A840,A846,A847,GOBRD13:44;
   end;
   hereby assume front_left_cell(f,m,G) misses C &
          front_right_cell(f,m,G) meets C;
    then F.(m+1+1) goes_straight m,G by A366,A843,A844,A845;
    hence f goes_straight m,G by A840,A846,A847,GOBRD13:46;
   end;
   assume front_left_cell(f,m,G) meets C;
    then F.(m+1+1) turns_left m,G by A366,A843,A844,A845;
   hence f turns_left m,G by A840,A846,A847,GOBRD13:45;
 end;
uniqueness
proof
  let f1,f2 be clockwise_oriented
     (standard non constant special_circular_sequence)
   such that
A848:   f1 is_sequence_on Gauge(C,n);
    given i1 such that
A849:   1 <= i1 & i1+1 <= len Gauge(C,n) and
A850:   f1/.1 = Gauge(C,n)*(i1,width Gauge(C,n)) &
      f1/.2 = Gauge(C,n)*(i1+1,width Gauge(C,n)) and
A851:   N-min C in cell(Gauge(C,n),i1,width Gauge(C,n)-'1) and
A852:   N-min C <> Gauge(C,n)*(i1,width Gauge(C,n)-'1);
   assume that
A853: for k st 1 <= k & k+2 <= len f1 holds
      (front_left_cell(f1,k,Gauge(C,n)) misses C &
         front_right_cell(f1,k,Gauge(C,n)) misses C
       implies f1 turns_right k,Gauge(C,n)) &
      (front_left_cell(f1,k,Gauge(C,n)) misses C &
        front_right_cell(f1,k,Gauge(C,n)) meets C
       implies f1 goes_straight k,Gauge(C,n)) &
      (front_left_cell(f1,k,Gauge(C,n)) meets C
       implies f1 turns_left k,Gauge(C,n)) and
A854:   f2 is_sequence_on Gauge(C,n);
   given i2 such that
A855:   1 <= i2 & i2+1 <= len Gauge(C,n) and
A856:   f2/.1 = Gauge(C,n)*(i2,width Gauge(C,n)) &
      f2/.2 = Gauge(C,n)*(i2+1,width Gauge(C,n)) and
A857:   N-min C in cell(Gauge(C,n),i2,width Gauge(C,n)-'1) and
A858:   N-min C <> Gauge(C,n)*(i2,width Gauge(C,n)-'1);
   assume
A859: for k st 1 <= k & k+2 <= len f2 holds
      (front_left_cell(f2,k,Gauge(C,n)) misses C &
         front_right_cell(f2,k,Gauge(C,n)) misses C
       implies f2 turns_right k,Gauge(C,n)) &
      (front_left_cell(f2,k,Gauge(C,n)) misses C &
        front_right_cell(f2,k,Gauge(C,n)) meets C
       implies f2 goes_straight k,Gauge(C,n)) &
      (front_left_cell(f2,k,Gauge(C,n)) meets C
       implies f2 turns_left k,Gauge(C,n));
       defpred P[Nat] means f1|$1 = f2|$1;
A860:  P[0];
A861: for k st P[k] holds P[k+1]
     proof let k such that
A862:    f1|k = f2|k;
A863:    f1|1 = <*f1/.1*> & f2|1 = <*f2/.1*> by FINSEQ_5:23;
A864:    i1 = i2 by A849,A851,A852,A855,A857,A858,Th31;
A865:    len f1 > 4 & len f2 > 4 by GOBOARD7:36;
      per cases by CQC_THE1:2;
      suppose k = 0;
       hence f1|(k+1) = f2|(k+1) by A849,A850,A851,A852,A855,A856,A857,A858,
A863,Th31;
      suppose
A866:      k = 1;
         len f1 > 2 & len f2 > 2 by A865,AXIOMS:22;
       then f1|2 = <*f1/.1,f1/.2*> & f2|2 = <*f2/.1,f2/.2*>
        by JORDAN8:1;
       hence f1|(k+1) = f2|(k+1) by A850,A856,A864,A866;
      suppose
A867:      k > 1;
A868:      f1/.1 = f1/.len f1 & f2/.1 = f2/.len f2 by FINSEQ_6:def 1;
         now
        per cases;
        suppose
A869:       len f1 > k;
         set m = k-'1;
A870:       1 <= m by A867,JORDAN3:12;
then A871:       m+1 = k by JORDAN3:6;
A872:      now assume
A873:        len f2 <= k;
then A874:        f2 = f2|k by TOPREAL1:2;
          then 1 in dom(f2|k) by FINSEQ_5:6;
then A875:        (f1|k)/.1 = f1/.1 by A862,TOPREAL1:1;
            len f2 in dom(f2|k) by A874,FINSEQ_5:6;
then A876:        (f1|k)/.len f2 = f1/.len f2 by A862,TOPREAL1:1;
          1 < len f2 & len f2 <= len f1 by A862,A865,A874,AXIOMS:22,FINSEQ_5:18
;
          hence contradiction by A862,A868,A869,A873,A874,A875,A876,GOBOARD7:40
;
         end;
then A877:       k+1 <= len f2 by NAT_1:38;
A878:       k+1 <= len f1 by A869,NAT_1:38;
A879:    front_left_cell(f1,m,Gauge(C,n))= front_left_cell(f1|k,m,Gauge(C,n))
              by A848,A869,A870,A871,GOBRD13:43;
A880:       front_right_cell(f1,m,Gauge(C,n))
              = front_right_cell(f1|k,m,Gauge(C,n))
              by A848,A869,A870,A871,GOBRD13:43;
A881:       front_left_cell(f2,m,Gauge(C,n))
              = front_left_cell(f2|k,m,Gauge(C,n))
              by A854,A870,A871,A872,GOBRD13:43;
A882:       front_right_cell(f2,m,Gauge(C,n))
              = front_right_cell(f2|k,m,Gauge(C,n))
              by A854,A870,A871,A872,GOBRD13:43;
A883:      m+(1+1) = k+1 by A871,XCMPLX_1:1;
           now per cases;
          suppose front_left_cell(f1,m,Gauge(C,n)) misses C &
             front_right_cell(f1,m,Gauge(C,n)) misses C;
           then f1 turns_right m,Gauge(C,n) & f2 turns_right m,Gauge(C,n)
            by A853,A859,A862,A870,A877,A878,A879,A880,A881,A882,A883;
          hence f1|(k+1) = f2|(k+1) by A854,A862,A867,A877,A878,GOBRD13:47;
         suppose front_left_cell(f1,m,Gauge(C,n)) misses C &
              front_right_cell(f1,m,Gauge(C,n)) meets C;
            then f1 goes_straight m,Gauge(C,n) & f2 goes_straight m,Gauge(C,n)
             by A853,A859,A862,A870,A877,A878,A879,A880,A881,A882,A883;
            hence f1|(k+1) = f2|(k+1) by A854,A862,A867,A877,A878,GOBRD13:49;
         suppose front_left_cell(f1,m,Gauge(C,n)) meets C;
           then f1 turns_left m,Gauge(C,n) & f2 turns_left m,Gauge(C,n)
             by A853,A859,A862,A870,A877,A878,A879,A881,A883;
           hence f1|(k+1) = f2|(k+1) by A854,A862,A867,A877,A878,GOBRD13:48;
         end;
         hence f1|(k+1) = f2|(k+1);
        suppose
A884:       k >= len f1;
then A885:       f1 = f1|k by TOPREAL1:2;
         then 1 in dom(f1|k) by FINSEQ_5:6;
then A886:       (f2|k)/.1 = f2/.1 by A862,TOPREAL1:1;
           len f1 in dom(f1|k) by A885,FINSEQ_5:6;
then A887:       (f2|k)/.len f1 = f2/.len f1 by A862,TOPREAL1:1;
           1 < len f1 & len f1 <= len f2 by A862,A865,A885,AXIOMS:22,FINSEQ_5:
18
;
then A888:       len f2 = len f1 by A862,A868,A885,A886,A887,GOBOARD7:40;
A889:        k+1 > len f1 by A884,NAT_1:38;
         hence f1|(k+1) = f1 by TOPREAL1:2
                       .= f2 by A862,A884,A885,A888,TOPREAL1:2
                       .= f2|(k+1) by A888,A889,TOPREAL1:2;
        end;
       hence f1|(k+1) = f2|(k+1);
     end;
     for k holds P[k] from Ind(A860,A861);
   hence f1 = f2 by Th4;
 end;
end;

theorem Th33:
  C is connected & 1 <= k & k+1 <= len Cage(C,n)
    implies
   left_cell(Cage(C,n),k,Gauge(C,n)) misses C &
     right_cell(Cage(C,n),k,Gauge(C,n)) meets C
proof
  set G = Gauge(C,n), f = Cage(C,n);
  set W = W-bound C, E = E-bound C, S = S-bound C, N = N-bound C;
  assume
A1:  C is connected;
then A2:  f is_sequence_on G by Def1;
defpred P[Nat] means
for m st 1 <= m & m+1 <= len(f|$1)
       holds left_cell(f|$1,m,G) misses C & right_cell(f|$1,m,G) meets C;
A3:
  P[0]
  proof
   let m; assume
A4:  1 <= m & m+1 <= len(f|0);
     1 <= m+1 by NAT_1:37;
   then 1 <= len(f|0) by A4,AXIOMS:22;
   then 1 <= 0 by FINSEQ_1:25;
   hence left_cell(f|0,m,G) misses C & right_cell(f|0,m,G) meets C;
  end;
A5:   len G = 2|^n+3 & len G = width G by JORDAN8:def 1;
A6: for k st P[k] holds P[k+1]
         proof let k such that
A7:   for m st 1 <= m & m+1 <= len(f|k)
        holds left_cell(f|k,m,G) misses C & right_cell(f|k,m,G) meets C;
    per cases;
    suppose
A8:    k >= len f;
     then k+1 >= len f by NAT_1:37;
     then f|k = f & f|(k+1) = f by A8,TOPREAL1:2;
     hence thesis by A7;
    suppose
A9:   k < len f;
then A10:  k+1 <= len f by NAT_1:38;
A11:   f|k is_sequence_on G by A2,GOBOARD1:38;
A12:  f|(k+1) is_sequence_on G by A2,GOBOARD1:38;
     consider i such that
A13:   1 <= i & i+1 <= len G and
A14:   f/.1 = G*(i,width G) and
A15:   f/.2 = G*(i+1,width G) and
A16:   N-min C in cell(G,i,width G-'1) & N-min C <> G*(i,width G-'1)
         by A1,Def1;
A17:   len(f|k) = k by A9,TOPREAL1:3;
A18:   len(f|(k+1)) = k+1 by A10,TOPREAL1:3;
A19:   1 <= len G by A5,NAT_1:37;
       2|^n >= n+1 & n+1 >= 1 by HEINE:7,NAT_1:37;
then A20:   2|^n > 0 by AXIOMS:22;
     let m such that
A21:  1 <= m & m+1 <= len(f|(k+1));
      now per cases by CQC_THE1:2;
    suppose
A22:   k = 0;
       1 <= m+1 by NAT_1:37;
     then m+1 = 0+1 by A18,A21,A22,AXIOMS:21;
     then m = 0 by XCMPLX_1:2;
     hence left_cell(f|(k+1),m,G) misses C &
              right_cell(f|(k+1),m,G) meets C by A21;
    suppose
A23:   k = 1;
then f|(k+1) = <*G*(i,width G),G*(i+1,width G)*> by A10,A14,A15,JORDAN8:1;
then A24:   (f|(k+1))/.1 = G*(i,width G) & (f|(k+1))/.2 = G*(i+1,width G)
        by FINSEQ_4:26;
A25:  i < len G & 1 <= i+1 & i+1 <= len G by A13,NAT_1:38;
then A26:   [i,len G] in Indices G & [i+1,len G] in Indices G
        by A5,A13,A19,GOBOARD7:10;
       1+1 <= m+1 by A21,AXIOMS:24;
     then m+1 = 1+1 by A18,A21,A23,AXIOMS:21;
then A27:   m = 1 by XCMPLX_1:2;
A28:     i < i+1 & i+1 < (i+1)+1 by NAT_1:38;
then A29:   left_cell(f|(k+1),m,G) = cell(G,i,len G)
        by A5,A12,A21,A24,A26,A27,GOBRD13:def 3;
       now assume left_cell(f|(k+1),m,G) meets C;
      then consider p being set such that
A30:     p in cell(G,i,len G) & p in C by A29,XBOOLE_0:3;
       reconsider p as Element of TOP-REAL 2 by A30;
        cell(G,i,len G) = { |[r,s]|: G*(i,1)`1 <= r & r <= G*(i+1,1)`1 &
                                  G*(1,len G)`2 <= s }
       by A5,A13,A25,GOBRD11:31;
      then consider r,s such that
A31:     p = |[r,s]| and
          G*(i,1)`1 <= r & r <= G*(i+1,1)`1 and
A32:     G*(1,len G)`2 <= s by A30;
        [1,len G] in Indices G by A5,A19,GOBOARD7:10;
then A33:     G*(1,len G) = |[W+((E-W)/(2|^n))*(1-2),S+((N-S)/(2|^n))*((len G)-
2)]|
         by JORDAN8:def 1;
         2|^n+(2+1) = 2|^n+1+2 by XCMPLX_1:1;
       then (len G) - 2 = 2|^n+1 by A5,XCMPLX_1:26;
       then ((N-S)/(2|^n))*((len G)-2)
           = ((N-S)/(2|^n))*(2|^n)+((N-S)/(2|^n))*1 by XCMPLX_1:8
          .= (N-S)+(N-S)/(2|^n) by A20,XCMPLX_1:88;
then A34:     S+((N-S)/(2|^n))*((len G) -2)
          = S+(N-S)+(N-S)/(2|^n) by XCMPLX_1:1
         .= N+(N-S)/(2|^n) by XCMPLX_1:27;
A35:     G*(1,len G)`2 = S+((N-S)/(2|^n))*((len G)-2) by A33,EUCLID:56;
         N > S by JORDAN8:12;
       then N-S > S-S by REAL_1:54;
       then N-S > 0 by XCMPLX_1:14;
       then (N-S)/(2|^n) > 0 by A20,REAL_2:127;
       then N+0 < N+(N-S)/(2|^n) by REAL_1:53;
then A36:     N < s by A32,A34,A35,AXIOMS:22;
         p`2 <= N by A30,PSCOMP_1:71;
      hence contradiction by A31,A36,EUCLID:56;
     end;
     hence left_cell(f|(k+1),m,G) misses C;
A37:   N-min C in C by SPRECT_1:13;
       N-min C in right_cell(f|(k+1),m,G)
      by A5,A12,A16,A21,A24,A26,A27,A28,GOBRD13:def 2;
     hence right_cell(f|(k+1),m,G) meets C by A37,XBOOLE_0:3;
    suppose
A38:   k > 1;
then A39:   1 <= (len(f|k))-'1 by A17,JORDAN3:12;
A40:   (len(f|k)) -'1 +1 = len(f|k) by A17,A38,AMI_5:4;
       now per cases;
      suppose A41: m+1 = len(f|(k+1));
then A42:    m = len(f|k) by A17,A18,XCMPLX_1:2;
A43:     len(f|k) <= len f by FINSEQ_5:18;
          now
         consider i1,j1,i2,j2 being Nat such that
A44:       [i1,j1] in Indices G & f/.((len(f|k)) -'1) = G*(i1,j1) and
A45:       [i2,j2] in Indices G & f/.len(f|k) = G*(i2,j2) and
            i1 = i2 & j1+1 = j2 or i1+1 = i2 & j1 = j2 or
           i1 = i2+1 & j1 = j2 or i1 = i2 & j1 = j2+1
            by A2,A9,A17,A39,A40,JORDAN8:6;
A46:      1 <= i2 & i2 <= len G & 1 <= j2 & j2 <= width G by A45,GOBOARD5:1;
A47:      (len(f|k))-'1+(1+1) = (len(f|k))+1 by A40,XCMPLX_1:1;
A48:      (i2-'1)+1 = i2 by A46,AMI_5:4;
A49:       (j2-'1)+1 = j2 by A46,AMI_5:4;
            left_cell(f|k,(len(f|k))-'1,G) misses C &
            right_cell(f|k,(len(f|k))-'1,G) meets C by A7,A39,A40;
then A50:      left_cell(f,(len(f|k))-'1,G) misses C &
           right_cell(f,(len(f|k))-'1,G) meets C
 by A2,A17,A39,A40,A43,GOBRD13:32;
         per cases;
         suppose
A51:        front_left_cell(f,(len(f|k))-'1,G) misses C &
            front_right_cell(f,(len(f|k))-'1,G) misses C;
          then A52:        f turns_right (len(f|k))-'1,G by A1,A10,A17,A39,A47,
Def1;
            now per cases by A40,A44,A45,A47,A52,GOBRD13:def 6;
           suppose that
A53:           i1 = i2 & j1+1 = j2 and
A54:           [i2+1,j2] in Indices G and
A55:           f/.((len(f|k))+1) = G*(i2+1,j2);
               front_right_cell(f,(len(f|k))-'1,G) = cell(G,i1,j2)
               by A2,A39,A40,A43,A44,A45,A53,GOBRD13:36;
            then left_cell(f,m,G) misses C
               by A2,A10,A17,A21,A42,A45,A51,A53,A54,A55,GOBRD13:24;
            hence left_cell(f|(k+1),m,G) misses C
              by A2,A10,A18,A21,A41,GOBRD13:32;
               j2-'1 = j1 & cell(G,i1,j1) meets C
               by A2,A39,A40,A43,A44,A45,A50,A53,BINARITH:39,GOBRD13:23;
            then right_cell(f,m,G) meets C
               by A2,A10,A17,A21,A42,A45,A53,A54,A55,GOBRD13:25;
            hence right_cell(f|(k+1),m,G) meets C
              by A2,A10,A18,A21,A41,GOBRD13:32;
           suppose that
A56:           i1+1 = i2 & j1 = j2 and
A57:           [i2,j2-'1] in Indices G and
A58:           f/.((len(f|k))+1) = G*(i2,j2-'1);
               front_right_cell(f,(len(f|k))-'1,G) = cell(G,i2,j2-'1)
               by A2,A39,A40,A43,A44,A45,A56,GOBRD13:38;
            then left_cell(f,m,G) misses C
               by A2,A10,A17,A21,A42,A45,A49,A51,A57,A58,GOBRD13:28;
            hence left_cell(f|(k+1),m,G) misses C
               by A2,A10,A18,A21,A41,GOBRD13:32;
               i2-'1 = i1 & cell(G,i1,j1-'1) meets C
               by A2,A39,A40,A43,A44,A45,A50,A56,BINARITH:39,GOBRD13:25;
            then right_cell(f,m,G) meets C
               by A2,A10,A17,A21,A42,A45,A49,A56,A57,A58,GOBRD13:29;
            hence right_cell(f|(k+1),m,G) meets C
               by A2,A10,A18,A21,A41,GOBRD13:32;
           suppose that
A59:           i1 = i2+1 & j1 = j2 and
A60:           [i2,j2+1] in Indices G and
A61:           f/.((len(f|k))+1) = G*(i2,j2+1);
               front_right_cell(f,(len(f|k))-'1,G) = cell(G,i2-'1,j2)
               by A2,A39,A40,A43,A44,A45,A59,GOBRD13:40;
            then left_cell(f,m,G) misses C
               by A2,A10,A17,A21,A42,A45,A51,A60,A61,GOBRD13:22;
            hence left_cell(f|(k+1),m,G) misses C
               by A2,A10,A18,A21,A41,GOBRD13:32;
               cell(G,i2,j2) meets C
               by A2,A39,A40,A43,A44,A45,A50,A59,GOBRD13:27;
            then right_cell(f,m,G) meets C
               by A2,A10,A17,A21,A42,A45,A60,A61,GOBRD13:23;
            hence right_cell(f|(k+1),m,G) meets C
               by A2,A10,A18,A21,A41,GOBRD13:32;
           suppose that
A62:           i1 = i2 & j1 = j2+1 and
A63:           [i2-'1,j2] in Indices G and
A64:           f/.((len(f|k))+1) = G*(i2-'1,j2);
               front_right_cell(f,(len(f|k))-'1,G) = cell(G,i2-'1,j2-'1)
               by A2,A39,A40,A43,A44,A45,A62,GOBRD13:42;
            then left_cell(f,m,G) misses C
              by A2,A10,A17,A21,A42,A45,A48,A51,A63,A64,GOBRD13:26;
            hence left_cell(f|(k+1),m,G) misses C
              by A2,A10,A18,A21,A41,GOBRD13:32;
               cell(G,i2-'1,j2) meets C by A2,A39,A40,A43,A44,A45,A50,A62,
GOBRD13:29;
            then right_cell(f,m,G) meets C
             by A2,A10,A17,A21,A42,A45,A48,A63,A64,GOBRD13:27;
            hence right_cell(f|(k+1),m,G) meets C
             by A2,A10,A18,A21,A41,GOBRD13:32;
          end;
          hence thesis;
         suppose
A65:        front_left_cell(f,(len(f|k))-'1,G) misses C &
             front_right_cell(f,(len(f|k))-'1,G) meets C;
            then A66:        f goes_straight (len(f|k))-'1,G by A1,A10,A17,A39,
A47,Def1;
            now per cases by A40,A44,A45,A47,A66,GOBRD13:def 8;
           suppose that
A67:           i1 = i2 & j1+1 = j2 and
A68:           [i2,j2+1] in Indices G and
A69:           f/.(len(f|k)+1) = G*(i2,j2+1);
               front_left_cell(f,(len(f|k))-'1,G) = cell(G,i1-'1,j2)
               by A2,A39,A40,A43,A44,A45,A67,GOBRD13:35;
             then left_cell(f,m,G) misses C
               by A2,A10,A17,A21,A42,A45,A65,A67,A68,A69,GOBRD13:22;
             hence left_cell(f|(k+1),m,G) misses C
               by A2,A10,A18,A21,A41,GOBRD13:32;
              front_right_cell(f,(len(f|k))-'1,G) = cell(G,i1,j2)
               by A2,A39,A40,A43,A44,A45,A67,GOBRD13:36;
            then right_cell(f,m,G) meets C
               by A2,A10,A17,A21,A42,A45,A65,A67,A68,A69,GOBRD13:23;
            hence right_cell(f|(k+1),m,G) meets C
               by A2,A10,A18,A21,A41,GOBRD13:32;
           suppose that
A70:           i1+1 = i2 & j1 = j2 and
A71:           [i2+1,j2] in Indices G and
A72:           f/.(len(f|k)+1) = G*(i2+1,j2);
               front_left_cell(f,(len(f|k))-'1,G) = cell(G,i2,j2)
               by A2,A39,A40,A43,A44,A45,A70,GOBRD13:37;
            then left_cell(f,m,G) misses C
               by A2,A10,A17,A21,A42,A45,A65,A71,A72,GOBRD13:24;
            hence left_cell(f|(k+1),m,G) misses C
               by A2,A10,A18,A21,A41,GOBRD13:32;
               front_right_cell(f,(len(f|k))-'1,G) = cell(G,i2,j2-'1)
               by A2,A39,A40,A43,A44,A45,A70,GOBRD13:38;
            then right_cell(f,m,G) meets C
               by A2,A10,A17,A21,A42,A45,A65,A71,A72,GOBRD13:25;
            hence right_cell(f|(k+1),m,G) meets C
               by A2,A10,A18,A21,A41,GOBRD13:32;
           suppose that
A73:          i1 = i2+1 & j1 = j2 and
A74:          [i2-'1,j2] in Indices G and
A75:          f/.(len(f|k)+1) = G*(i2-'1,j2);
               front_left_cell(f,(len(f|k))-'1,G) = cell(G,i2-'1,j2-'1)
               by A2,A39,A40,A43,A44,A45,A73,GOBRD13:39;
            then left_cell(f,m,G) misses C
               by A2,A10,A17,A21,A42,A45,A48,A65,A74,A75,GOBRD13:26;
            hence left_cell(f|(k+1),m,G) misses C
               by A2,A10,A18,A21,A41,GOBRD13:32;
               front_right_cell(f,(len(f|k))-'1,G) = cell(G,i2-'1,j2)
               by A2,A39,A40,A43,A44,A45,A73,GOBRD13:40;
            then right_cell(f,m,G) meets C
               by A2,A10,A17,A21,A42,A45,A48,A65,A74,A75,GOBRD13:27;
            hence right_cell(f|(k+1),m,G) meets C
               by A2,A10,A18,A21,A41,GOBRD13:32;
           suppose that
A76:          i1 = i2 & j1 = j2+1 and
A77:          [i2,j2-'1] in Indices G and
A78:          f/.(len(f|k)+1) = G*(i2,j2-'1);
               front_left_cell(f,(len(f|k))-'1,G) = cell(G,i2,j2-'1)
               by A2,A39,A40,A43,A44,A45,A76,GOBRD13:41;
            then left_cell(f,m,G) misses C
               by A2,A10,A17,A21,A42,A45,A49,A65,A77,A78,GOBRD13:28;
            hence left_cell(f|(k+1),m,G) misses C
               by A2,A10,A18,A21,A41,GOBRD13:32;
               front_right_cell(f,(len(f|k))-'1,G) = cell(G,i2-'1,j2-'1)
               by A2,A39,A40,A43,A44,A45,A76,GOBRD13:42;
            then right_cell(f,m,G) meets C
               by A2,A10,A17,A21,A42,A45,A49,A65,A77,A78,GOBRD13:29;
            hence right_cell(f|(k+1),m,G) meets C
               by A2,A10,A18,A21,A41,GOBRD13:32;
          end;
          hence thesis;
         suppose
A79:        front_left_cell(f,(len(f|k))-'1,G) meets C;
then A80:        f turns_left (len(f|k))-'1,G by A1,A10,A17,A39,A47,Def1;
            now per cases by A40,A44,A45,A47,A80,GOBRD13:def 7;
           suppose that
A81:           i1 = i2 & j1+1 = j2 and
A82:           [i2-'1,j2] in Indices G and
A83:           f/.(len(f|k)+1) = G*(i2-'1,j2);
               j2-'1 = j1 &
             cell(G,i1-'1,j1) misses C
               by A2,A39,A40,A43,A44,A45,A50,A81,BINARITH:39,GOBRD13:22;
            then left_cell(f,m,G) misses C
               by A2,A10,A17,A21,A42,A45,A48,A81,A82,A83,GOBRD13:26;
            hence left_cell(f|(k+1),m,G) misses C
               by A2,A10,A18,A21,A41,GOBRD13:32;
               front_left_cell(f,(len(f|k))-'1,G) = cell(G,i1-'1,j2)
               by A2,A39,A40,A43,A44,A45,A81,GOBRD13:35;
            then right_cell(f,m,G) meets C
               by A2,A10,A17,A21,A42,A45,A48,A79,A81,A82,A83,GOBRD13:27;
            hence right_cell(f|(k+1),m,G) meets C
               by A2,A10,A18,A21,A41,GOBRD13:32;
           suppose that
A84:           i1+1 = i2 & j1 = j2 and
A85:           [i2,j2+1] in Indices G and
A86:           f/.(len(f|k)+1) = G*(i2,j2+1);
               i2-'1 = i1 & cell(G,i1,j1) misses C
               by A2,A39,A40,A43,A44,A45,A50,A84,BINARITH:39,GOBRD13:24;
             then left_cell(f,m,G) misses C
               by A2,A10,A17,A21,A42,A45,A84,A85,A86,GOBRD13:22;
             hence left_cell(f|(k+1),m,G) misses C
               by A2,A10,A18,A21,A41,GOBRD13:32;
               front_left_cell(f,(len(f|k))-'1,G) = cell(G,i2,j2)
               by A2,A39,A40,A43,A44,A45,A84,GOBRD13:37;
            then right_cell(f,m,G) meets C
               by A2,A10,A17,A21,A42,A45,A79,A85,A86,GOBRD13:23;
            hence right_cell(f|(k+1),m,G) meets C
               by A2,A10,A18,A21,A41,GOBRD13:32;
           suppose that
A87:           i1 = i2+1 & j1 = j2 and
A88:           [i2,j2-'1] in Indices G and
A89:           f/.(len(f|k)+1) = G*(i2,j2-'1);
               cell(G,i2,j2-'1) misses C
               by A2,A39,A40,A43,A44,A45,A50,A87,GOBRD13:26;
            then left_cell(f,m,G) misses C
               by A2,A10,A17,A21,A42,A45,A49,A88,A89,GOBRD13:28;
            hence left_cell(f|(k+1),m,G) misses C
               by A2,A10,A18,A21,A41,GOBRD13:32;
               front_left_cell(f,(len(f|k))-'1,G) = cell(G,i2-'1,j2-'1)
               by A2,A39,A40,A43,A44,A45,A87,GOBRD13:39;
            then right_cell(f,m,G) meets C
               by A2,A10,A17,A21,A42,A45,A49,A79,A88,A89,GOBRD13:29;
            hence right_cell(f|(k+1),m,G) meets C
               by A2,A10,A18,A21,A41,GOBRD13:32;
           suppose that
A90:           i1 = i2 & j1 = j2+1 and
A91:           [i2+1,j2] in Indices G and
A92:           f/.(len(f|k)+1) = G*(i2+1,j2);
               cell(G,i2,j2) misses C
              by A2,A39,A40,A43,A44,A45,A50,A90,GOBRD13:28;
            then left_cell(f,m,G) misses C
               by A2,A10,A17,A21,A42,A45,A91,A92,GOBRD13:24;
            hence left_cell(f|(k+1),m,G) misses C
               by A2,A10,A18,A21,A41,GOBRD13:32;
               front_left_cell(f,(len(f|k))-'1,G) = cell(G,i2,j2-'1)
               by A2,A39,A40,A43,A44,A45,A90,GOBRD13:41;
            then right_cell(f,m,G) meets C
               by A2,A10,A17,A21,A42,A45,A79,A91,A92,GOBRD13:25;
            hence right_cell(f|(k+1),m,G) meets C
               by A2,A10,A18,A21,A41,GOBRD13:32;
          end;
         hence thesis;
        end;
       hence thesis;
      suppose m+1 <> len(f|(k+1));
       then m+1 < len(f|(k+1)) by A21,REAL_1:def 5;
       then A93:     m+1 <= len(f|k)by A17,A18,NAT_1:38;
       then consider i1,j1,i2,j2 being Nat such that
A94:     [i1,j1] in Indices G & (f|k)/.m = G*(i1,j1) and
A95:     [i2,j2] in Indices G & (f|k)/.(m+1) = G*(i2,j2) and
A96:      i1 = i2 & j1+1 = j2 or i1+1 = i2 & j1 = j2 or
          i1 = i2+1 & j1 = j2 or i1 = i2 & j1 = j2+1 by A11,A21,JORDAN8:6;
A97:    1 <= m+1 by NAT_1:37;
          m <= len(f|k) by A93,NAT_1:38;
then A98:    m in dom(f|k) & m+1 in dom(f|k) by A21,A93,A97,FINSEQ_3:27;
A99:     left_cell(f|k,m,G) misses C & right_cell(f|k,m,G) meets C
         by A7,A21,A93;
       f|(k+1) = (f|k)^<*f/.(k+1)*> by A10,JORDAN8:2;
then A100:      (f|(k+1))/.m = G*(i1,j1) & (f|(k+1))/.(m+1) = G*(i2,j2)
            by A94,A95,A98,GROUP_5:95;
         now per cases by A96;
        suppose
A101:        i1 = i2 & j1+1 = j2;
          then left_cell(f|k,m,G) = cell(G,i1-'1,j1) &
            right_cell(f|k,m,G) = cell(G,i1,j1)
           by A11,A21,A93,A94,A95,GOBRD13:22,23;
         hence thesis by A12,A21,A94,A95,A99,A100,A101,GOBRD13:22,23;
        suppose
A102:        i1+1 = i2 & j1 = j2;
          then left_cell(f|k,m,G) = cell(G,i1,j1) &
            right_cell(f|k,m,G) = cell(G,i1,j1-'1)
           by A11,A21,A93,A94,A95,GOBRD13:24,25;
         hence thesis by A12,A21,A94,A95,A99,A100,A102,GOBRD13:24,25;
        suppose
A103:        i1 = i2+1 & j1 = j2;
          then left_cell(f|k,m,G) = cell(G,i2,j2-'1) &
            right_cell(f|k,m,G) = cell(G,i2,j2)
           by A11,A21,A93,A94,A95,GOBRD13:26,27;
         hence thesis by A12,A21,A94,A95,A99,A100,A103,GOBRD13:26,27;
        suppose
A104:        i1 = i2 & j1 = j2+1;
          then left_cell(f|k,m,G) = cell(G,i2,j2) &
            right_cell(f|k,m,G) = cell(G,i1-'1,j2)
           by A11,A21,A93,A94,A95,GOBRD13:28,29;
         hence thesis by A12,A21,A94,A95,A99,A100,A104,GOBRD13:28,29;
       end;
       hence thesis;
     end;
     hence thesis;
    end;
    hence thesis;
   end;
A105: for k holds P[k]
     from Ind(A3,A6);
    f|len f = f by TOPREAL1:2;
  hence thesis by A105;
end;

theorem
    C is connected implies N-min L~Cage(C,n) = (Cage(C,n))/.1
proof
 set f = Cage(C,n);
 assume
A1: C is connected;
then A2: f is_sequence_on Gauge(C,n) &
 (ex i 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)) by Def1;
    for k 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 by A1,Th33;
  hence thesis by A2,Th32;
end;

Back to top