The Mizar article:

Properties of the Upper and Lower Sequence on the Cage

by
Robert Milewski

Received August 1, 2002

Copyright (c) 2002 Association of Mizar Users

MML identifier: JORDAN15
[ MML identifier index ]


environ

 vocabulary JORDAN8, PSCOMP_1, BOOLE, JORDAN1A, JORDAN9, GOBOARD1, GOBOARD5,
      JORDAN6, SPRECT_2, JORDAN3, FINSEQ_5, JORDAN2C, GROUP_2, REALSET1,
      JORDAN1E, SQUARE_1, FINSEQ_4, TOPREAL2, CONNSP_1, COMPTS_1, TOPREAL1,
      SPPOL_1, FINSEQ_1, TREES_1, EUCLID, RELAT_1, MCART_1, MATRIX_1, GOBOARD9,
      PRE_TOPC, RELAT_2, SEQM_3, SUBSET_1, FUNCT_1, ARYTM_1, FINSEQ_6,
      GOBOARD2, GRAPH_2, ORDINAL2, LATTICES, SEQ_2, FUNCT_5, ARYTM;
 notation TARSKI, XBOOLE_0, ENUMSET1, SUBSET_1, ORDINAL1, NUMBERS, XREAL_0,
      REAL_1, NAT_1, SQUARE_1, FUNCT_1, FUNCT_2, FINSEQ_1, FINSEQ_4, FINSEQ_5,
      SEQ_4, MATRIX_1, REALSET1, FINSEQ_6, GRAPH_2, STRUCT_0, PRE_TOPC,
      RCOMP_1, TOPREAL2, BINARITH, CONNSP_1, COMPTS_1, EUCLID, PSCOMP_1,
      SPRECT_2, GOBOARD1, TOPREAL1, GOBOARD2, GOBOARD5, GOBOARD9, GOBRD13,
      SPPOL_1, JORDAN3, JORDAN8, JORDAN2C, JORDAN6, JORDAN9, JORDAN1A,
      JORDAN1E;
 constructors JORDAN8, REALSET1, GOBOARD9, JORDAN6, REAL_1, CARD_4, PSCOMP_1,
      BINARITH, JORDAN2C, CONNSP_1, JORDAN9, JORDAN1A, SQUARE_1, FINSEQ_4,
      JORDAN1, JORDAN3, TOPREAL2, JORDAN5C, GRAPH_2, SPRECT_1, GOBOARD2,
      JORDAN1E, RCOMP_1, WSIERP_1, JORDAN1H, ENUMSET1, INT_1;
 clusters XREAL_0, SPRECT_1, TOPREAL6, JORDAN8, REVROT_1, INT_1, RELSET_1,
      SPRECT_3, GOBOARD2, SPPOL_2, JORDAN1A, JORDAN1B, JORDAN1E, GOBRD11,
      XBOOLE_0, SPRECT_4, BORSUK_2, PSCOMP_1, JORDAN1J, MEMBERED, NUMBERS,
      ORDINAL2;
 requirements NUMERALS, BOOLE, SUBSET, REAL, ARITHM;
 definitions TARSKI, XBOOLE_0;
 theorems NAT_1, FINSEQ_1, GOBOARD1, FINSEQ_4, EUCLID, FINSEQ_3, AXIOMS,
      REAL_1, SPPOL_2, TARSKI, JORDAN3, SQUARE_1, PSCOMP_1, FINSEQ_5, FINSEQ_6,
      GOBOARD7, TOPREAL1, BINARITH, AMI_5, GOBOARD5, SPRECT_2, SPPOL_1,
      FUNCT_2, GOBOARD9, FINSEQ_2, SUBSET_1, JORDAN4, SPRECT_3, TOPREAL3,
      JORDAN8, PARTFUN2, SPRECT_1, XBOOLE_0, XBOOLE_1, ZFMISC_1, SEQ_4,
      GOBRD14, TOPREAL6, JORDAN2C, PRE_TOPC, JORDAN6, JORDAN9, JORDAN1H,
      JORDAN1A, JORDAN1C, JORDAN1E, JORDAN10, JGRAPH_1, REVROT_1, RCOMP_1,
      COMPTS_1, ENUMSET1, JORDAN1B, JORDAN1F, JORDAN1G, JORDAN1I, JORDAN1J,
      GOBOARD3, TOPREAL8, AMISTD_1, GRAPH_2, SPRECT_5, JORDAN1D, XCMPLX_1;

begin

  reserve n for Nat;

  theorem
     for A,B be Subset of TOP-REAL 2 st A meets B holds
    proj1.:A meets proj1.:B
   proof
    let A,B be Subset of TOP-REAL 2;
    assume A meets B;
    then consider e be set such that
     A1: e in A and
     A2: e in B by XBOOLE_0:3;
    reconsider e as Point of TOP-REAL 2 by A1;
    A3: e`1 = proj1.e by PSCOMP_1:def 28;
    then A4: e`1 in proj1.:A by A1,FUNCT_2:43;
      e`1 in proj1.:B by A2,A3,FUNCT_2:43;
    hence proj1.:A meets proj1.:B by A4,XBOOLE_0:3;
   end;

  theorem
     for A,B be Subset of TOP-REAL 2
   for s be real number st
    A misses B & A c= Horizontal_Line s & B c= Horizontal_Line s holds
     proj1.:A misses proj1.:B
   proof
    let A,B be Subset of TOP-REAL 2;
    let s be real number such that
     A1: A misses B and
     A2: A c= Horizontal_Line s and
     A3: B c= Horizontal_Line s;
    assume proj1.:A meets proj1.:B;
    then consider e be set such that
     A4: e in proj1.:A and
     A5: e in proj1.:B by XBOOLE_0:3;
    reconsider e as Real by A4;
    consider c being Point of TOP-REAL 2 such that
     A6: c in A and
     A7: e = proj1.c by A4,FUNCT_2:116;
    consider d be Point of TOP-REAL 2 such that
     A8: d in B and
     A9: e = proj1.d by A5,FUNCT_2:116;
      c`2 = s & d`2 = s by A2,A3,A6,A8,JORDAN6:35;
    then c = |[c`1,d`2]| by EUCLID:57
       .= |[e,d`2]| by A7,PSCOMP_1:def 28
       .= |[d`1,d`2]| by A9,PSCOMP_1:def 28
       .= d by EUCLID:57;
    hence contradiction by A1,A6,A8,XBOOLE_0:3;
   end;

  theorem Th3:
   for S be closed Subset of TOP-REAL 2 st S is Bounded holds
    proj1.:S is closed
   proof
    let S be closed Subset of TOP-REAL 2;
    assume S is Bounded;
    then Cl(proj1.:S) = proj1.:Cl S by TOPREAL6:92
       .= proj1.:S by PRE_TOPC:52;
    hence proj1.:S is closed;
   end;

  theorem Th4:
   for S be compact Subset of TOP-REAL 2 holds proj1.:S is compact
   proof
    let S being compact Subset of TOP-REAL 2;
    A1: S is Bounded closed by JORDAN2C:73;
    then A2: proj1.:S is closed by Th3;
      proj1.:S is bounded by A1,JORDAN1C:4;
    hence proj1.:S is compact by A2,RCOMP_1:29;
   end;

  theorem Th5:
   for p,q,p1,q1 be Point of TOP-REAL 2 st
    LSeg(p,q) is vertical & LSeg(p1,q1) is vertical & p`1 = p1`1 &
    p`2 <= p1`2 & p1`2 <= q1`2 & q1`2 <= q`2 holds
     LSeg(p1,q1) c= LSeg(p,q)
   proof
    let p,q,p1,q1 be Point of TOP-REAL 2;
    assume that
     A1: LSeg(p,q) is vertical and
     A2: LSeg(p1,q1) is vertical and
     A3: p`1 = p1`1 and
     A4: p`2 <= p1`2 & p1`2 <= q1`2 & q1`2 <= q`2;
    let x be set;
    assume A5: x in LSeg(p1,q1);
    then reconsider r=x as Point of TOP-REAL 2;
    A6: p1`1 = r`1 by A2,A5,SPRECT_3:20;
    A7: p`1 = q`1 by A1,SPPOL_1:37;
      p1`2 <= r`2 & r`2 <= q1`2 by A4,A5,TOPREAL1:10;
    then p`2 <= r`2 & r`2 <= q`2 by A4,AXIOMS:22;
    hence x in LSeg(p,q) by A3,A6,A7,GOBOARD7:8;
   end;

  theorem Th6:
   for p,q,p1,q1 be Point of TOP-REAL 2 st
    LSeg(p,q) is horizontal & LSeg(p1,q1) is horizontal & p`2 = p1`2 &
    p`1 <= p1`1 & p1`1 <= q1`1 & q1`1 <= q`1 holds
     LSeg(p1,q1) c= LSeg(p,q)
   proof
    let p,q,p1,q1 be Point of TOP-REAL 2;
    assume that
     A1: LSeg(p,q) is horizontal and
     A2: LSeg(p1,q1) is horizontal and
     A3: p`2 = p1`2 and
     A4: p`1 <= p1`1 & p1`1 <= q1`1 & q1`1 <= q`1;
    let x be set;
    assume A5: x in LSeg(p1,q1);
    then reconsider r=x as Point of TOP-REAL 2;
    A6: p1`2 = r`2 by A2,A5,SPRECT_3:19;
    A7: p`2 = q`2 by A1,SPPOL_1:36;
      p1`1 <= r`1 & r`1 <= q1`1 by A4,A5,TOPREAL1:9;
    then p`1 <= r`1 & r`1 <= q`1 by A4,AXIOMS:22;
    hence x in LSeg(p,q) by A3,A6,A7,GOBOARD7:9;
   end;

  theorem Th7:
   for G be Go-board
   for i,j,k,j1,k1 be Nat st
    1 <= i & i <= len G &
    1 <= j & j <= j1 & j1 <= k1 & k1 <= k & k <= width G holds
     LSeg(G*(i,j1),G*(i,k1)) c= LSeg(G*(i,j),G*(i,k))
   proof
    let G be Go-board;
    let i,j,k,j1,k1 be Nat;
    assume that
     A1: 1 <= i & i <= len G and
     A2: 1 <= j & j <= j1 & j1 <= k1 & k1 <= k & k <= width G;
      j <= k1 & k1 <= width G & 1 <= j1 & j1 <= k by A2,AXIOMS:22;
    then A3: j <= width G & 1 <= k by AXIOMS:22;
    then G*(i,j)`1 = G*(i,1)`1 by A1,A2,GOBOARD5:3
       .= G*(i,k)`1 by A1,A2,A3,GOBOARD5:3;
    then A4: LSeg(G*(i,j),G*(i,k)) is vertical by SPPOL_1:37;
      j1 <= k & j <= k1 by A2,AXIOMS:22;
    then A5: 1 <= j1 & j1 <= width G & 1 <= k1 & k1 <= width G by A2,AXIOMS:22;
    then G*(i,j1)`1 = G*(i,1)`1 by A1,GOBOARD5:3
       .= G*(i,k1)`1 by A1,A5,GOBOARD5:3;
    then A6: LSeg(G*(i,j1),G*(i,k1)) is vertical by SPPOL_1:37;
    A7: G*(i,j)`1 = G*(i,1)`1 by A1,A2,A3,GOBOARD5:3
       .= G*(i,j1)`1 by A1,A5,GOBOARD5:3;
    A8: G*(i,j)`2 <= G*(i,j1)`2 by A1,A2,A5,SPRECT_3:24;
    A9: G*(i,j1)`2 <= G*(i,k1)`2 by A1,A2,A5,SPRECT_3:24;
      G*(i,k1)`2 <= G*(i,k)`2 by A1,A2,A5,SPRECT_3:24;
    hence LSeg(G*(i,j1),G*(i,k1)) c= LSeg(G*(i,j),G*(i,k))
                                                    by A4,A6,A7,A8,A9,Th5;
   end;

  theorem Th8:
   for G be Go-board
   for i,j,k,j1,k1 be Nat st
    1 <= i & i <= width G &
    1 <= j & j <= j1 & j1 <= k1 & k1 <= k & k <= len G holds
     LSeg(G*(j1,i),G*(k1,i)) c= LSeg(G*(j,i),G*(k,i))
   proof
    let G be Go-board;
    let i,j,k,j1,k1 be Nat;
    assume that
     A1: 1 <= i & i <= width G and
     A2: 1 <= j & j <= j1 & j1 <= k1 & k1 <= k & k <= len G;
      j <= k1 & k1 <= len G & 1 <= j1 & j1 <= k by A2,AXIOMS:22;
    then A3: j <= len G & 1 <= k by AXIOMS:22;
    then G*(j,i)`2 = G*(1,i)`2 by A1,A2,GOBOARD5:2
       .= G*(k,i)`2 by A1,A2,A3,GOBOARD5:2;
    then A4: LSeg(G*(j,i),G*(k,i)) is horizontal by SPPOL_1:36;
      j1 <= k & j <= k1 by A2,AXIOMS:22;
    then A5: 1 <= j1 & j1 <= len G & 1 <= k1 & k1 <= len G by A2,AXIOMS:22;
    then G*(j1,i)`2 = G*(1,i)`2 by A1,GOBOARD5:2
       .= G*(k1,i)`2 by A1,A5,GOBOARD5:2;
    then A6: LSeg(G*(j1,i),G*(k1,i)) is horizontal by SPPOL_1:36;
    A7: G*(j,i)`2 = G*(1,i)`2 by A1,A2,A3,GOBOARD5:2
       .= G*(j1,i)`2 by A1,A5,GOBOARD5:2;
    A8: G*(j,i)`1 <= G*(j1,i)`1 by A1,A2,A5,SPRECT_3:25;
    A9: G*(j1,i)`1 <= G*(k1,i)`1 by A1,A2,A5,SPRECT_3:25;
      G*(k1,i)`1 <= G*(k,i)`1 by A1,A2,A5,SPRECT_3:25;
    hence LSeg(G*(j1,i),G*(k1,i)) c= LSeg(G*(j,i),G*(k,i))
                                                    by A4,A6,A7,A8,A9,Th6;
   end;

  theorem
     for G be Go-board
   for j,k,j1,k1 be Nat st
    1 <= j & j <= j1 & j1 <= k1 & k1 <= k & k <= width G holds
     LSeg(G*(Center G,j1),G*(Center G,k1)) c=
      LSeg(G*(Center G,j),G*(Center G,k))
   proof
    let G be Go-board;
    let j,k,j1,k1 be Nat;
    assume A1: 1 <= j & j <= j1 & j1 <= k1 & k1 <= k & k <= width G;
      1 <= Center G & Center G <= len G by JORDAN1B:12,14;
    hence LSeg(G*(Center G,j1),G*(Center G,k1)) c=
     LSeg(G*(Center G,j),G*(Center G,k)) by A1,Th7;
   end;

  theorem
     for G be Go-board st len G = width G
   for j,k,j1,k1 be Nat st
    1 <= j & j <= j1 & j1 <= k1 & k1 <= k & k <= len G holds
     LSeg(G*(j1,Center G),G*(k1,Center G)) c=
      LSeg(G*(j,Center G),G*(k,Center G))
   proof
    let G be Go-board;
    assume A1: len G = width G;
    let j,k,j1,k1 be Nat;
    assume A2: 1 <= j & j <= j1 & j1 <= k1 & k1 <= k & k <= len G;
      1 <= Center G & Center G <= width G by A1,JORDAN1B:12,14;
    hence LSeg(G*(j1,Center G),G*(k1,Center G)) c=
     LSeg(G*(j,Center G),G*(k,Center G)) by A2,Th8;
   end;

  theorem Th11:
   for C be compact connected non vertical non horizontal Subset of TOP-REAL 2
   for i,j,k be Nat st
    1 <= i & i <= len Gauge(C,n) &
    1 <= j & j <= k & k <= width Gauge(C,n) &
    Gauge(C,n)*(i,j) in L~Lower_Seq(C,n)
   ex j1 be Nat st
    j <= j1 & j1 <= k &
    LSeg(Gauge(C,n)*(i,j1),Gauge(C,n)*(i,k)) /\ L~Lower_Seq(C,n) =
     {Gauge(C,n)*(i,j1)}
   proof
    let C be compact connected non vertical non horizontal
     Subset of TOP-REAL 2;
    let i,j,k be Nat;
    assume that
     A1: 1 <= i & i <= len Gauge(C,n) and
     A2: 1 <= j & j <= k & k <= width Gauge(C,n) and
     A3: Gauge(C,n)*(i,j) in L~Lower_Seq(C,n);
    set G = Gauge(C,n);
    set s = G*(i,1)`1;
    set f = G*(i,j);
    set e = G*(i,k);
    set w2 = sup(proj2.:(LSeg(f,e) /\ L~Lower_Seq(C,n)));
    set q = |[s,w2]|;
    A4: j <= width G by A2,AXIOMS:22;
    then A5: [i,j] in Indices G by A1,A2,GOBOARD7:10;
    A6: k >= 1 by A2,AXIOMS:22;
    then A7: G*(i,k)`1 = s by A1,A2,GOBOARD5:3;
    A8: Lower_Seq(C,n) is_sequence_on G by JORDAN1G:5;
    A9: G*(i,j) in LSeg(G*(i,j),G*(i,k)) by TOPREAL1:6;
    then A10: LSeg(G*(i,j),G*(i,k)) meets L~Lower_Seq(C,n) by A3,XBOOLE_0:3;
      [i,k] in Indices G by A1,A2,A6,GOBOARD7:10;
    then consider j1 be Nat such that
     A11: j <= j1 & j1 <= k and
     A12: G*(i,j1)`2 = w2 by A2,A5,A8,A10,JORDAN1F:2;
    A13: 1 <= j1 & j1 <= width G by A2,A11,AXIOMS:22;
    then A14: G*(i,j1)`1 = s by A1,GOBOARD5:3;
    then A15: q = G*(i,j1) by A12,EUCLID:57;
    take j1;
    thus j <= j1 & j1 <= k by A11;
    A16: q`2 <= e`2 by A1,A2,A11,A13,A15,SPRECT_3:24;
    set X = LSeg(G*(i,j),G*(i,k)) /\ L~Lower_Seq(C,n);
    reconsider X1=X as non empty compact Subset of TOP-REAL 2
                                              by A3,A9,PSCOMP_1:64,XBOOLE_0:def
3;
    consider pp be set such that
     A17: pp in N-most X1 by XBOOLE_0:def 1;
    reconsider pp as Point of TOP-REAL 2 by A17;
      pp in LSeg(NW-corner X, NE-corner X)/\X by A17,PSCOMP_1:def 39;
    then A18: pp in X by XBOOLE_0:def 3;
    then A19: pp in L~Lower_Seq(C,n) by XBOOLE_0:def 3;
    A20: pp in LSeg(G*(i,j),G*(i,k)) by A18,XBOOLE_0:def 3;
      f`1 = e`1 by A1,A2,A4,A7,GOBOARD5:3;
    then LSeg(f,e) is vertical by SPPOL_1:37;
    then A21: pp`1 = q`1 by A7,A14,A15,A20,SPRECT_3:20;
      q`2 = N-bound X by A12,A15,SPRECT_1:50
       .= (N-min X)`2 by PSCOMP_1:94
       .= pp`2 by A17,PSCOMP_1:98;
    then A22: q in L~Lower_Seq(C,n) by A19,A21,TOPREAL3:11;
      for x be set holds
     x in LSeg(e,q) /\ L~Lower_Seq(C,n) iff x = q
    proof
     let x be set;
     thus x in LSeg(e,q) /\ L~Lower_Seq(C,n) implies x = q
     proof
      assume A23: x in LSeg(e,q) /\ L~Lower_Seq(C,n);
      then reconsider pp = x as Point of TOP-REAL 2;
      A24: pp in LSeg(e,q) by A23,XBOOLE_0:def 3;
      then A25: pp`1 = q`1 by A7,A14,A15,GOBOARD7:5;
      A26: pp`2 >= q`2 by A16,A24,TOPREAL1:10;
      reconsider EE = LSeg(f,e) /\ L~Lower_Seq(C,n)
                              as compact Subset of TOP-REAL 2 by PSCOMP_1:64;
      reconsider E0 = proj2.:EE as compact Subset of REAL by JORDAN1A:28;
      A27: f`1 = q`1 by A1,A2,A4,A14,A15,GOBOARD5:3;
        f`2 <= q`2 by A1,A2,A11,A13,A15,SPRECT_3:24;
      then A28: q in LSeg(e,f) by A7,A14,A15,A16,A27,GOBOARD7:8;
        e in LSeg(f,e) by TOPREAL1:6;
      then A29: LSeg(e,q) c= LSeg(f,e) by A28,TOPREAL1:12;
        E0 is bounded by RCOMP_1:28;
      then A30: E0 is bounded_above by SEQ_4:def 3;
        pp in L~Lower_Seq(C,n) by A23,XBOOLE_0:def 3;
      then pp in EE by A24,A29,XBOOLE_0:def 3;
      then proj2.pp in E0 by FUNCT_2:43;
      then pp`2 in E0 by PSCOMP_1:def 29;
      then q`2 >= pp`2 by A12,A15,A30,SEQ_4:def 4;
      then pp`2 = q`2 by A26,AXIOMS:21;
      hence x = q by A25,TOPREAL3:11;
     end;
     assume A31: x = q;
     then x in LSeg(e,q) by TOPREAL1:6;
     hence x in LSeg(e,q) /\ L~Lower_Seq(C,n) by A22,A31,XBOOLE_0:def 3;
    end;
    hence LSeg(Gauge(C,n)*(i,j1),Gauge(C,n)*(i,k)) /\ L~Lower_Seq(C,n) =
     {Gauge(C,n)*(i,j1)} by A15,TARSKI:def 1;
   end;

  theorem
     for C be compact connected non vertical non horizontal Subset of TOP-REAL
2
   for i,j,k be Nat st
    1 <= i & i <= len Gauge(C,n) &
    1 <= j & j <= k & k <= width Gauge(C,n) &
    Gauge(C,n)*(i,k) in L~Upper_Seq(C,n)
   ex k1 be Nat st
    j <= k1 & k1 <= k &
    LSeg(Gauge(C,n)*(i,j),Gauge(C,n)*(i,k1)) /\ L~Upper_Seq(C,n) =
     {Gauge(C,n)*(i,k1)}
   proof
    let C be compact connected non vertical non horizontal
     Subset of TOP-REAL 2;
    let i,j,k be Nat;
    assume that
     A1: 1 <= i & i <= len Gauge(C,n) and
     A2: 1 <= j & j <= k & k <= width Gauge(C,n) and
     A3: Gauge(C,n)*(i,k) in L~Upper_Seq(C,n);
    set G = Gauge(C,n);
    set s = G*(i,1)`1;
    set f = G*(i,j);
    set e = G*(i,k);
    set w1 = inf(proj2.:(LSeg(f,e) /\ L~Upper_Seq(C,n)));
    set p = |[s,w1]|;
    A4: Upper_Seq(C,n) is_sequence_on G by JORDAN1G:4;
    A5: G*(i,k) in LSeg(G*(i,j),G*(i,k)) by TOPREAL1:6;
    then A6: LSeg(G*(i,j),G*(i,k)) meets L~Upper_Seq(C,n) by A3,XBOOLE_0:3;
    A7: j <= width G by A2,AXIOMS:22;
    then A8: [i,j] in Indices G by A1,A2,GOBOARD7:10;
    A9: k >= 1 by A2,AXIOMS:22;
    then [i,k] in Indices G by A1,A2,GOBOARD7:10;
    then consider k1 be Nat such that
     A10: j <= k1 & k1 <= k and
     A11: G*(i,k1)`2 = w1 by A2,A4,A6,A8,JORDAN1F:1;
    A12: 1 <= k1 & k1 <= width G by A2,A10,AXIOMS:22;
    then A13: G*(i,k1)`1 = s by A1,GOBOARD5:3;
    then A14: p = G*(i,k1) by A11,EUCLID:57;
    take k1;
    thus j <= k1 & k1 <= k by A10;
    A15: f`2 <= p`2 by A1,A2,A10,A12,A14,SPRECT_3:24;
    A16: f`1 = p`1 by A1,A2,A7,A13,A14,GOBOARD5:3;
    set X = LSeg(G*(i,j),G*(i,k)) /\ L~Upper_Seq(C,n);
    reconsider X1=X as non empty compact Subset of TOP-REAL 2
                                               by A3,A5,PSCOMP_1:64,XBOOLE_0:
def 3;
    consider pp be set such that
     A17: pp in S-most X1 by XBOOLE_0:def 1;
    reconsider pp as Point of TOP-REAL 2 by A17;
      pp in LSeg(SW-corner X, SE-corner X)/\X by A17,PSCOMP_1:def 41;
    then A18: pp in X by XBOOLE_0:def 3;
    then A19: pp in L~Upper_Seq(C,n) by XBOOLE_0:def 3;
    A20: pp in LSeg(G*(i,j),G*(i,k)) by A18,XBOOLE_0:def 3;
      f`1 = s by A1,A2,A7,GOBOARD5:3
       .= e`1 by A1,A2,A9,GOBOARD5:3;
    then LSeg(f,e) is vertical by SPPOL_1:37;
    then A21: pp`1 = p`1 by A16,A20,SPRECT_3:20;
      p`2 = S-bound X by A11,A14,SPRECT_1:49
       .= (S-min X)`2 by PSCOMP_1:114
       .= pp`2 by A17,PSCOMP_1:118;
    then A22: p in L~Upper_Seq(C,n) by A19,A21,TOPREAL3:11;
      for x be set holds
     x in LSeg(p,f) /\ L~Upper_Seq(C,n) iff x = p
    proof
     let x be set;
     thus x in LSeg(p,f) /\ L~Upper_Seq(C,n) implies x = p
     proof
      assume A23: x in LSeg(p,f) /\ L~Upper_Seq(C,n);
      then reconsider pp = x as Point of TOP-REAL 2;
      A24: pp in LSeg(p,f) by A23,XBOOLE_0:def 3;
      then A25: pp`1 = p`1 by A16,GOBOARD7:5;
      A26: pp`2 <= p`2 by A15,A24,TOPREAL1:10;
      reconsider EE = LSeg(f,e) /\ L~Upper_Seq(C,n)
                              as compact Subset of TOP-REAL 2 by PSCOMP_1:64;
      reconsider E0 = proj2.:EE as compact Subset of REAL by JORDAN1A:28;
      A27: f in LSeg(f,e) by TOPREAL1:6;
      A28: f`1 = p`1 by A1,A2,A7,A13,A14,GOBOARD5:3;
      A29: e`1 = p`1 by A1,A2,A9,A13,A14,GOBOARD5:3;
      A30: f`2 <= p`2 by A1,A2,A10,A12,A14,SPRECT_3:24;
        p`2 <= e`2 by A1,A2,A10,A12,A14,SPRECT_3:24;
      then p in LSeg(f,e) by A28,A29,A30,GOBOARD7:8;
      then A31: LSeg(p,f) c= LSeg(f,e) by A27,TOPREAL1:12;
        E0 is bounded by RCOMP_1:28;
      then A32: E0 is bounded_below by SEQ_4:def 3;
        pp in L~Upper_Seq(C,n) by A23,XBOOLE_0:def 3;
      then pp in EE by A24,A31,XBOOLE_0:def 3;
      then proj2.pp in E0 by FUNCT_2:43;
      then pp`2 in E0 by PSCOMP_1:def 29;
      then p`2 <= pp`2 by A11,A14,A32,SEQ_4:def 5;
      then pp`2 = p`2 by A26,AXIOMS:21;
      hence x = p by A25,TOPREAL3:11;
     end;
     assume A33: x = p;
     then x in LSeg(p,f) by TOPREAL1:6;
     hence x in LSeg(p,f) /\ L~Upper_Seq(C,n) by A22,A33,XBOOLE_0:def 3;
    end;
    hence LSeg(Gauge(C,n)*(i,j),Gauge(C,n)*(i,k1)) /\ L~Upper_Seq(C,n) =
     {Gauge(C,n)*(i,k1)} by A14,TARSKI:def 1;
   end;

  theorem Th13:
   for C be compact connected non vertical non horizontal Subset of TOP-REAL 2
   for i,j,k be Nat st
    1 <= i & i <= len Gauge(C,n) &
    1 <= j & j <= k & k <= width Gauge(C,n) &
    Gauge(C,n)*(i,j) in L~Lower_Seq(C,n) &
    Gauge(C,n)*(i,k) in L~Upper_Seq(C,n)
   ex j1,k1 be Nat st
    j <= j1 & j1 <= k1 & k1 <= k &
    LSeg(Gauge(C,n)*(i,j1),Gauge(C,n)*(i,k1)) /\ L~Lower_Seq(C,n) =
     {Gauge(C,n)*(i,j1)} &
    LSeg(Gauge(C,n)*(i,j1),Gauge(C,n)*(i,k1)) /\ L~Upper_Seq(C,n) =
     {Gauge(C,n)*(i,k1)}
   proof
    let C be compact connected non vertical non horizontal
     Subset of TOP-REAL 2;
    let i,j,k be Nat;
    assume that
     A1: 1 <= i & i <= len Gauge(C,n) and
     A2: 1 <= j & j <= k & k <= width Gauge(C,n) and
     A3: Gauge(C,n)*(i,j) in L~Lower_Seq(C,n) and
     A4: Gauge(C,n)*(i,k) in L~Upper_Seq(C,n);
    set G = Gauge(C,n);
    set s = G*(i,1)`1;
    set f = G*(i,j);
    set e = G*(i,k);
    set w1 = inf(proj2.:(LSeg(f,e) /\ L~Upper_Seq(C,n)));
    set p = |[s,w1]|;
    set w2 = sup(proj2.:(LSeg(f,p) /\ L~Lower_Seq(C,n)));
    set q = |[s,w2]|;
    A5: Upper_Seq(C,n) is_sequence_on G by JORDAN1G:4;
    A6: G*(i,k) in LSeg(G*(i,j),G*(i,k)) by TOPREAL1:6;
    then A7: LSeg(G*(i,j),G*(i,k)) meets L~Upper_Seq(C,n) by A4,XBOOLE_0:3;
    A8: j <= width G by A2,AXIOMS:22;
    then A9: [i,j] in Indices G by A1,A2,GOBOARD7:10;
    A10: k >= 1 by A2,AXIOMS:22;
    then [i,k] in Indices G by A1,A2,GOBOARD7:10;
    then consider k1 be Nat such that
     A11: j <= k1 & k1 <= k and
     A12: G*(i,k1)`2 = w1 by A2,A5,A7,A9,JORDAN1F:1;
    A13: 1 <= k1 & k1 <= width G by A2,A11,AXIOMS:22;
    then A14: G*(i,k1)`1 = s by A1,GOBOARD5:3;
    then A15: p = G*(i,k1) by A12,EUCLID:57;
    A16: Lower_Seq(C,n) is_sequence_on G by JORDAN1G:5;
    A17: G*(i,j) in LSeg(G*(i,j),G*(i,k1)) by TOPREAL1:6;
    then A18: LSeg(G*(i,j),G*(i,k1)) meets L~Lower_Seq(C,n) by A3,XBOOLE_0:3;
      [i,k1] in Indices G by A1,A13,GOBOARD7:10;
    then consider j1 be Nat such that
     A19: j <= j1 & j1 <= k1 and
     A20: G*(i,j1)`2 = w2 by A9,A11,A15,A16,A18,JORDAN1F:2;
    A21: 1 <= j1 & j1 <= width G by A2,A13,A19,AXIOMS:22;
    then A22: G*(i,j1)`1 = s by A1,GOBOARD5:3;
    then A23: q = G*(i,j1) by A20,EUCLID:57;
    take j1,k1;
    thus j <= j1 & j1 <= k1 & k1 <= k by A11,A19;
    A24: q`2 <= p`2 by A1,A13,A15,A19,A21,A23,SPRECT_3:24;
    set X = LSeg(G*(i,j),G*(i,k1)) /\ L~Lower_Seq(C,n);
    reconsider X1=X as non empty compact Subset of TOP-REAL 2
                                              by A3,A17,PSCOMP_1:64,XBOOLE_0:
def 3;
    consider pp be set such that
     A25: pp in N-most X1 by XBOOLE_0:def 1;
    reconsider pp as Point of TOP-REAL 2 by A25;
      pp in LSeg(NW-corner X, NE-corner X)/\X by A25,PSCOMP_1:def 39;
    then A26: pp in X by XBOOLE_0:def 3;
    then A27: pp in L~Lower_Seq(C,n) by XBOOLE_0:def 3;
    A28: pp in LSeg(G*(i,j),G*(i,k1)) by A26,XBOOLE_0:def 3;
    A29: f`1 = p`1 by A1,A2,A8,A14,A15,GOBOARD5:3;
    then LSeg(f,p) is vertical by SPPOL_1:37;
    then A30: pp`1 = q`1 by A14,A15,A22,A23,A28,SPRECT_3:20;
      q`2 = N-bound X by A15,A20,A23,SPRECT_1:50
       .= (N-min X)`2 by PSCOMP_1:94
       .= pp`2 by A25,PSCOMP_1:98;
    then A31: q in L~Lower_Seq(C,n) by A27,A30,TOPREAL3:11;
      for x be set holds
     x in LSeg(p,q) /\ L~Lower_Seq(C,n) iff x = q
    proof
     let x be set;
     thus x in LSeg(p,q) /\ L~Lower_Seq(C,n) implies x = q
     proof
      assume A32: x in LSeg(p,q) /\ L~Lower_Seq(C,n);
      then reconsider pp = x as Point of TOP-REAL 2;
      A33: pp in LSeg(p,q) by A32,XBOOLE_0:def 3;
      then A34: pp`1 = q`1 by A14,A15,A22,A23,GOBOARD7:5;
      A35: pp`2 >= q`2 by A24,A33,TOPREAL1:10;
      reconsider EE = LSeg(f,p) /\ L~Lower_Seq(C,n)
                              as compact Subset of TOP-REAL 2 by PSCOMP_1:64;
      reconsider E0 = proj2.:EE as compact Subset of REAL by JORDAN1A:28;
      A36: f`1 = q`1 by A1,A2,A8,A22,A23,GOBOARD5:3;
        f`2 <= q`2 by A1,A2,A19,A21,A23,SPRECT_3:24;
      then A37: q in LSeg(p,f) by A14,A15,A22,A23,A24,A36,GOBOARD7:8;
        p in LSeg(f,p) by TOPREAL1:6;
      then A38: LSeg(p,q) c= LSeg(f,p) by A37,TOPREAL1:12;
        E0 is bounded by RCOMP_1:28;
      then A39: E0 is bounded_above by SEQ_4:def 3;
        pp in L~Lower_Seq(C,n) by A32,XBOOLE_0:def 3;
      then pp in EE by A33,A38,XBOOLE_0:def 3;
      then proj2.pp in E0 by FUNCT_2:43;
      then pp`2 in E0 by PSCOMP_1:def 29;
      then q`2 >= pp`2 by A20,A23,A39,SEQ_4:def 4;
      then pp`2 = q`2 by A35,AXIOMS:21;
      hence x = q by A34,TOPREAL3:11;
     end;
     assume A40: x = q;
     then x in LSeg(p,q) by TOPREAL1:6;
     hence x in LSeg(p,q) /\ L~Lower_Seq(C,n) by A31,A40,XBOOLE_0:def 3;
    end;
    hence LSeg(Gauge(C,n)*(i,j1),
     Gauge(C,n)*(i,k1)) /\ L~Lower_Seq(C,n) =
      {Gauge(C,n)*(i,j1)} by A15,A23,TARSKI:def 1;
    set X = LSeg(G*(i,j),G*(i,k)) /\ L~Upper_Seq(C,n);
    reconsider X1=X as non empty compact Subset of TOP-REAL 2
                                               by A4,A6,PSCOMP_1:64,XBOOLE_0:
def 3;
    consider pp be set such that
     A41: pp in S-most X1 by XBOOLE_0:def 1;
    reconsider pp as Point of TOP-REAL 2 by A41;
      pp in LSeg(SW-corner X, SE-corner X)/\X by A41,PSCOMP_1:def 41;
    then A42: pp in X by XBOOLE_0:def 3;
    then A43: pp in L~Upper_Seq(C,n) by XBOOLE_0:def 3;
    A44: pp in LSeg(G*(i,j),G*(i,k)) by A42,XBOOLE_0:def 3;
      f`1 = s by A1,A2,A8,GOBOARD5:3
       .= e`1 by A1,A2,A10,GOBOARD5:3;
    then LSeg(f,e) is vertical by SPPOL_1:37;
    then A45: pp`1 = p`1 by A29,A44,SPRECT_3:20;
      p`2 = S-bound X by A12,A15,SPRECT_1:49
       .= (S-min X)`2 by PSCOMP_1:114
       .= pp`2 by A41,PSCOMP_1:118;
    then A46: p in L~Upper_Seq(C,n) by A43,A45,TOPREAL3:11;
      for x be set holds
     x in LSeg(p,q) /\ L~Upper_Seq(C,n) iff x = p
    proof
     let x be set;
     thus x in LSeg(p,q) /\ L~Upper_Seq(C,n) implies x = p
     proof
      assume A47: x in LSeg(p,q) /\ L~Upper_Seq(C,n);
      then reconsider pp = x as Point of TOP-REAL 2;
      A48: pp in LSeg(p,q) by A47,XBOOLE_0:def 3;
      then A49: pp`1 = p`1 by A14,A15,A22,A23,GOBOARD7:5;
      A50: pp`2 <= p`2 by A24,A48,TOPREAL1:10;
      reconsider EE = LSeg(f,e) /\ L~Upper_Seq(C,n)
                              as compact Subset of TOP-REAL 2 by PSCOMP_1:64;
      reconsider E0 = proj2.:EE as compact Subset of REAL by JORDAN1A:28;
      A51: f`1 = q`1 by A1,A2,A8,A22,A23,GOBOARD5:3;
      A52: e`1 = q`1 by A1,A2,A10,A22,A23,GOBOARD5:3;
      A53: f`2 <= q`2 by A1,A2,A19,A21,A23,SPRECT_3:24;
        j1 <= k by A11,A19,AXIOMS:22;
      then q`2 <= e`2 by A1,A2,A21,A23,SPRECT_3:24;
      then A54: q in LSeg(f,e) by A51,A52,A53,GOBOARD7:8;
      A55: f`1 = p`1 by A1,A2,A8,A14,A15,GOBOARD5:3;
      A56: e`1 = p`1 by A1,A2,A10,A14,A15,GOBOARD5:3;
      A57: f`2 <= p`2 by A1,A2,A11,A13,A15,SPRECT_3:24;
        p`2 <= e`2 by A1,A2,A11,A13,A15,SPRECT_3:24;
      then p in LSeg(f,e) by A55,A56,A57,GOBOARD7:8;
      then A58: LSeg(p,q) c= LSeg(f,e) by A54,TOPREAL1:12;
        E0 is bounded by RCOMP_1:28;
      then A59: E0 is bounded_below by SEQ_4:def 3;
        pp in L~Upper_Seq(C,n) by A47,XBOOLE_0:def 3;
      then pp in EE by A48,A58,XBOOLE_0:def 3;
      then proj2.pp in E0 by FUNCT_2:43;
      then pp`2 in E0 by PSCOMP_1:def 29;
      then p`2 <= pp`2 by A12,A15,A59,SEQ_4:def 5;
      then pp`2 = p`2 by A50,AXIOMS:21;
      hence x = p by A49,TOPREAL3:11;
     end;
     assume A60: x = p;
     then x in LSeg(p,q) by TOPREAL1:6;
     hence x in LSeg(p,q) /\ L~Upper_Seq(C,n) by A46,A60,XBOOLE_0:def 3;
    end;
    hence LSeg(Gauge(C,n)*(i,j1),
     Gauge(C,n)*(i,k1)) /\ L~Upper_Seq(C,n) =
      {Gauge(C,n)*(i,k1)} by A15,A23,TARSKI:def 1;
   end;

  theorem
     for C be compact connected non vertical non horizontal Subset of TOP-REAL
2
   for i,j,k be Nat st
    1 <= j & j <= k & k <= len Gauge(C,n) &
    1 <= i & i <= width Gauge(C,n) &
    Gauge(C,n)*(j,i) in L~Lower_Seq(C,n)
   ex j1 be Nat st
    j <= j1 & j1 <= k &
    LSeg(Gauge(C,n)*(j1,i),Gauge(C,n)*(k,i)) /\ L~Lower_Seq(C,n) =
     {Gauge(C,n)*(j1,i)}
   proof
    let C be compact connected non vertical non horizontal
     Subset of TOP-REAL 2;
    let i,j,k be Nat;
    assume that
     A1: 1 <= j & j <= k & k <= len Gauge(C,n) and
     A2: 1 <= i & i <= width Gauge(C,n) and
     A3: Gauge(C,n)*(j,i) in L~Lower_Seq(C,n);
    set G = Gauge(C,n);
    set s = G*(1,i)`2;
    set f = G*(j,i);
    set e = G*(k,i);
    set w2 = sup(proj1.:(LSeg(f,e) /\ L~Lower_Seq(C,n)));
    set q = |[w2,s]|;
    A4: len G = width G by JORDAN8:def 1;
    then A5: j <= width G by A1,AXIOMS:22;
    then A6: [j,i] in Indices G by A1,A2,A4,GOBOARD7:10;
    A7: k >= 1 by A1,AXIOMS:22;
    then A8: G*(k,i)`2 = s by A1,A2,GOBOARD5:2;
    A9: Lower_Seq(C,n) is_sequence_on G by JORDAN1G:5;
    A10: G*(j,i) in LSeg(G*(j,i),G*(k,i)) by TOPREAL1:6;
    then A11: LSeg(G*(j,i),G*(k,i)) meets L~Lower_Seq(C,n) by A3,XBOOLE_0:3;
      [k,i] in Indices G by A1,A2,A7,GOBOARD7:10;
    then consider j1 be Nat such that
     A12: j <= j1 & j1 <= k and
     A13: G*(j1,i)`1 = w2 by A1,A6,A9,A11,JORDAN1F:4;
    A14: 1 <= j1 & j1 <= width G by A1,A4,A12,AXIOMS:22;
    then A15: G*(j1,i)`2 = s by A2,A4,GOBOARD5:2;
    then A16: q = G*(j1,i) by A13,EUCLID:57;
    take j1;
    thus j <= j1 & j1 <= k by A12;
    A17: q`1 <= e`1 by A1,A2,A12,A14,A16,SPRECT_3:25;
    set X = LSeg(G*(j,i),G*(k,i)) /\ L~Lower_Seq(C,n);
    reconsider X1=X as non empty compact Subset of TOP-REAL 2
                                              by A3,A10,PSCOMP_1:64,XBOOLE_0:
def 3;
    consider pp be set such that
     A18: pp in E-most X1 by XBOOLE_0:def 1;
    reconsider pp as Point of TOP-REAL 2 by A18;
      pp in LSeg(SE-corner X, NE-corner X)/\X by A18,PSCOMP_1:def 40;
    then A19: pp in X by XBOOLE_0:def 3;
    then A20: pp in L~Lower_Seq(C,n) by XBOOLE_0:def 3;
    A21: pp in LSeg(G*(j,i),G*(k,i)) by A19,XBOOLE_0:def 3;
      f`2 = e`2 by A1,A2,A4,A5,A8,GOBOARD5:2;
    then LSeg(f,e) is horizontal by SPPOL_1:36;
    then A22: pp`2 = q`2 by A8,A15,A16,A21,SPRECT_3:19;
      q`1 = E-bound X by A13,A16,SPRECT_1:51
       .= (E-min X)`1 by PSCOMP_1:104
       .= pp`1 by A18,PSCOMP_1:108;
    then A23: q in L~Lower_Seq(C,n) by A20,A22,TOPREAL3:11;
      for x be set holds
     x in LSeg(e,q) /\ L~Lower_Seq(C,n) iff x = q
    proof
     let x be set;
     thus x in LSeg(e,q) /\ L~Lower_Seq(C,n) implies x = q
     proof
      assume A24: x in LSeg(e,q) /\ L~Lower_Seq(C,n);
      then reconsider pp = x as Point of TOP-REAL 2;
      A25: pp in LSeg(e,q) by A24,XBOOLE_0:def 3;
      then A26: pp`2 = q`2 by A8,A15,A16,GOBOARD7:6;
      A27: pp`1 >= q`1 by A17,A25,TOPREAL1:9;
      reconsider EE = LSeg(f,e) /\ L~Lower_Seq(C,n)
                              as compact Subset of TOP-REAL 2 by PSCOMP_1:64;
      reconsider E0 = proj1.:EE as compact Subset of REAL by Th4;
      A28: f`2 = q`2 by A1,A2,A4,A5,A15,A16,GOBOARD5:2;
        f`1 <= q`1 by A1,A2,A4,A12,A14,A16,SPRECT_3:25;
      then A29: q in LSeg(e,f) by A8,A15,A16,A17,A28,GOBOARD7:9;
        e in LSeg(f,e) by TOPREAL1:6;
      then A30: LSeg(e,q) c= LSeg(f,e) by A29,TOPREAL1:12;
        E0 is bounded by RCOMP_1:28;
      then A31: E0 is bounded_above by SEQ_4:def 3;
        pp in L~Lower_Seq(C,n) by A24,XBOOLE_0:def 3;
      then pp in EE by A25,A30,XBOOLE_0:def 3;
      then proj1.pp in E0 by FUNCT_2:43;
      then pp`1 in E0 by PSCOMP_1:def 28;
      then q`1 >= pp`1 by A13,A16,A31,SEQ_4:def 4;
      then pp`1 = q`1 by A27,AXIOMS:21;
      hence x = q by A26,TOPREAL3:11;
     end;
     assume A32: x = q;
     then x in LSeg(e,q) by TOPREAL1:6;
     hence x in LSeg(e,q) /\ L~Lower_Seq(C,n) by A23,A32,XBOOLE_0:def 3;
    end;
    hence LSeg(Gauge(C,n)*(j1,i),Gauge(C,n)*(k,i)) /\ L~Lower_Seq(C,n) =
     {Gauge(C,n)*(j1,i)} by A16,TARSKI:def 1;
   end;

  theorem Th15:
   for C be compact connected non vertical non horizontal Subset of TOP-REAL 2
   for i,j,k be Nat st
    1 <= j & j <= k & k <= len Gauge(C,n) &
    1 <= i & i <= width Gauge(C,n) &
    Gauge(C,n)*(k,i) in L~Upper_Seq(C,n)
   ex k1 be Nat st
    j <= k1 & k1 <= k &
    LSeg(Gauge(C,n)*(j,i),Gauge(C,n)*(k1,i)) /\ L~Upper_Seq(C,n) =
     {Gauge(C,n)*(k1,i)}
   proof
    let C be compact connected non vertical non horizontal
     Subset of TOP-REAL 2;
    let i,j,k be Nat;
    assume that
     A1: 1 <= j & j <= k & k <= len Gauge(C,n) and
     A2: 1 <= i & i <= width Gauge(C,n) and
     A3: Gauge(C,n)*(k,i) in L~Upper_Seq(C,n);
    set G = Gauge(C,n);
    set s = G*(1,i)`2;
    set f = G*(j,i);
    set e = G*(k,i);
    set w1 = inf(proj1.:(LSeg(f,e) /\ L~Upper_Seq(C,n)));
    set p = |[w1,s]|;
    A4: len G = width G by JORDAN8:def 1;
    A5: Upper_Seq(C,n) is_sequence_on G by JORDAN1G:4;
    A6: G*(k,i) in LSeg(G*(j,i),G*(k,i)) by TOPREAL1:6;
    then A7: LSeg(G*(j,i),G*(k,i)) meets L~Upper_Seq(C,n) by A3,XBOOLE_0:3;
    A8: j <= width G by A1,A4,AXIOMS:22;
    then A9: [j,i] in Indices G by A1,A2,A4,GOBOARD7:10;
    A10: k >= 1 by A1,AXIOMS:22;
    then [k,i] in Indices G by A1,A2,GOBOARD7:10;
    then consider k1 be Nat such that
     A11: j <= k1 & k1 <= k and
     A12: G*(k1,i)`1 = w1 by A1,A5,A7,A9,JORDAN1F:3;
    A13: 1 <= k1 & k1 <= width G by A1,A4,A11,AXIOMS:22;
    then A14: G*(k1,i)`2 = s by A2,A4,GOBOARD5:2;
    then A15: p = G*(k1,i) by A12,EUCLID:57;
    take k1;
    thus j <= k1 & k1 <= k by A11;
    A16: f`1 <= p`1 by A1,A2,A4,A11,A13,A15,SPRECT_3:25;
    A17: f`2 = p`2 by A1,A2,A4,A8,A14,A15,GOBOARD5:2;
    set X = LSeg(G*(j,i),G*(k,i)) /\ L~Upper_Seq(C,n);
    reconsider X1=X as non empty compact Subset of TOP-REAL 2
                                                by A3,A6,PSCOMP_1:64,XBOOLE_0:
def 3;
    consider pp be set such that
     A18: pp in W-most X1 by XBOOLE_0:def 1;
    reconsider pp as Point of TOP-REAL 2 by A18;
      pp in LSeg(SW-corner X, NW-corner X)/\X by A18,PSCOMP_1:def 38;
    then A19: pp in X by XBOOLE_0:def 3;
    then A20: pp in L~Upper_Seq(C,n) by XBOOLE_0:def 3;
    A21: pp in LSeg(G*(j,i),G*(k,i)) by A19,XBOOLE_0:def 3;
      f`2 = s by A1,A2,A4,A8,GOBOARD5:2
       .= e`2 by A1,A2,A10,GOBOARD5:2;
    then LSeg(f,e) is horizontal by SPPOL_1:36;
    then A22: pp`2 = p`2 by A17,A21,SPRECT_3:19;
      p`1 = W-bound X by A12,A15,SPRECT_1:48
       .= (W-min X)`1 by PSCOMP_1:84
       .= pp`1 by A18,PSCOMP_1:88;
    then A23: p in L~Upper_Seq(C,n) by A20,A22,TOPREAL3:11;
      for x be set holds
     x in LSeg(p,f) /\ L~Upper_Seq(C,n) iff x = p
    proof
     let x be set;
     thus x in LSeg(p,f) /\ L~Upper_Seq(C,n) implies x = p
     proof
      assume A24: x in LSeg(p,f) /\ L~Upper_Seq(C,n);
      then reconsider pp = x as Point of TOP-REAL 2;
      A25: pp in LSeg(p,f) by A24,XBOOLE_0:def 3;
      then A26: pp`2 = p`2 by A17,GOBOARD7:6;
      A27: pp`1 <= p`1 by A16,A25,TOPREAL1:9;
      reconsider EE = LSeg(f,e) /\ L~Upper_Seq(C,n)
                              as compact Subset of TOP-REAL 2 by PSCOMP_1:64;
      reconsider E0 = proj1.:EE as compact Subset of REAL by Th4;
      A28: f in LSeg(f,e) by TOPREAL1:6;
      A29: f`2 = p`2 by A1,A2,A4,A8,A14,A15,GOBOARD5:2;
      A30: e`2 = p`2 by A1,A2,A10,A14,A15,GOBOARD5:2;
      A31: f`1 <= p`1 by A1,A2,A4,A11,A13,A15,SPRECT_3:25;
        p`1 <= e`1 by A1,A2,A11,A13,A15,SPRECT_3:25;
      then p in LSeg(f,e) by A29,A30,A31,GOBOARD7:9;
      then A32: LSeg(p,f) c= LSeg(f,e) by A28,TOPREAL1:12;
        E0 is bounded by RCOMP_1:28;
      then A33: E0 is bounded_below by SEQ_4:def 3;
        pp in L~Upper_Seq(C,n) by A24,XBOOLE_0:def 3;
      then pp in EE by A25,A32,XBOOLE_0:def 3;
      then proj1.pp in E0 by FUNCT_2:43;
      then pp`1 in E0 by PSCOMP_1:def 28;
      then p`1 <= pp`1 by A12,A15,A33,SEQ_4:def 5;
      then pp`1 = p`1 by A27,AXIOMS:21;
      hence x = p by A26,TOPREAL3:11;
     end;
     assume A34: x = p;
     then x in LSeg(p,f) by TOPREAL1:6;
     hence x in LSeg(p,f) /\ L~Upper_Seq(C,n) by A23,A34,XBOOLE_0:def 3;
    end;
    hence LSeg(Gauge(C,n)*(j,i),Gauge(C,n)*(k1,i)) /\ L~Upper_Seq(C,n) =
     {Gauge(C,n)*(k1,i)} by A15,TARSKI:def 1;
   end;

  theorem Th16:
   for C be compact connected non vertical non horizontal Subset of TOP-REAL 2
   for i,j,k be Nat st
    1 <= j & j <= k & k <= len Gauge(C,n) &
    1 <= i & i <= width Gauge(C,n) &
    Gauge(C,n)*(j,i) in L~Lower_Seq(C,n) &
    Gauge(C,n)*(k,i) in L~Upper_Seq(C,n)
   ex j1,k1 be Nat st
    j <= j1 & j1 <= k1 & k1 <= k &
    LSeg(Gauge(C,n)*(j1,i),Gauge(C,n)*(k1,i)) /\ L~Lower_Seq(C,n) =
     {Gauge(C,n)*(j1,i)} &
    LSeg(Gauge(C,n)*(j1,i),Gauge(C,n)*(k1,i)) /\ L~Upper_Seq(C,n) =
     {Gauge(C,n)*(k1,i)}
   proof
    let C be compact connected non vertical non horizontal
     Subset of TOP-REAL 2;
    let i,j,k be Nat;
    assume that
     A1: 1 <= j & j <= k & k <= len Gauge(C,n) and
     A2: 1 <= i & i <= width Gauge(C,n) and
     A3: Gauge(C,n)*(j,i) in L~Lower_Seq(C,n) and
     A4: Gauge(C,n)*(k,i) in L~Upper_Seq(C,n);
    set G = Gauge(C,n);
    set s = G*(1,i)`2;
    set f = G*(j,i);
    set e = G*(k,i);
    set w1 = inf(proj1.:(LSeg(f,e) /\ L~Upper_Seq(C,n)));
    set p = |[w1,s]|;
    set w2 = sup(proj1.:(LSeg(f,p) /\ L~Lower_Seq(C,n)));
    set q = |[w2,s]|;
    A5: len G = width G by JORDAN8:def 1;
    A6: Upper_Seq(C,n) is_sequence_on G by JORDAN1G:4;
    A7: G*(k,i) in LSeg(G*(j,i),G*(k,i)) by TOPREAL1:6;
    then A8: LSeg(G*(j,i),G*(k,i)) meets L~Upper_Seq(C,n) by A4,XBOOLE_0:3;
    A9: j <= width G by A1,A5,AXIOMS:22;
    then A10: [j,i] in Indices G by A1,A2,A5,GOBOARD7:10;
    A11: k >= 1 by A1,AXIOMS:22;
    then [k,i] in Indices G by A1,A2,GOBOARD7:10;
    then consider k1 be Nat such that
     A12: j <= k1 & k1 <= k and
     A13: G*(k1,i)`1 = w1 by A1,A6,A8,A10,JORDAN1F:3;
    A14: 1 <= k1 & k1 <= width G by A1,A5,A12,AXIOMS:22;
    then A15: G*(k1,i)`2 = s by A2,A5,GOBOARD5:2;
    then A16: p = G*(k1,i) by A13,EUCLID:57;
    A17: Lower_Seq(C,n) is_sequence_on G by JORDAN1G:5;
    A18: G*(j,i) in LSeg(G*(j,i),G*(k1,i)) by TOPREAL1:6;
    then A19: LSeg(G*(j,i),G*(k1,i)) meets L~Lower_Seq(C,n) by A3,XBOOLE_0:3;
      [k1,i] in Indices G by A2,A5,A14,GOBOARD7:10;
    then consider j1 be Nat such that
     A20: j <= j1 & j1 <= k1 and
     A21: G*(j1,i)`1 = w2 by A10,A12,A16,A17,A19,JORDAN1F:4;
    A22: 1 <= j1 & j1 <= width G by A1,A14,A20,AXIOMS:22;
    then A23: G*(j1,i)`2 = s by A2,A5,GOBOARD5:2;
    then A24: q = G*(j1,i) by A21,EUCLID:57;
    take j1,k1;
    thus j <= j1 & j1 <= k1 & k1 <= k by A12,A20;
    A25: q`1 <= p`1 by A2,A5,A14,A16,A20,A22,A24,SPRECT_3:25;
    set X = LSeg(G*(j,i),G*(k1,i)) /\ L~Lower_Seq(C,n);
    reconsider X1=X as non empty compact Subset of TOP-REAL 2
                                               by A3,A18,PSCOMP_1:64,XBOOLE_0:
def 3;
    consider pp be set such that
     A26: pp in E-most X1 by XBOOLE_0:def 1;
    reconsider pp as Point of TOP-REAL 2 by A26;
      pp in LSeg(SE-corner X, NE-corner X)/\X by A26,PSCOMP_1:def 40;
    then A27: pp in X by XBOOLE_0:def 3;
    then A28: pp in L~Lower_Seq(C,n) by XBOOLE_0:def 3;
    A29: pp in LSeg(G*(j,i),G*(k1,i)) by A27,XBOOLE_0:def 3;
    A30: f`2 = p`2 by A1,A2,A5,A9,A15,A16,GOBOARD5:2;
    then LSeg(f,p) is horizontal by SPPOL_1:36;
    then A31: pp`2 = q`2 by A15,A16,A23,A24,A29,SPRECT_3:19;
      q`1 = E-bound X by A16,A21,A24,SPRECT_1:51
       .= (E-min X)`1 by PSCOMP_1:104
       .= pp`1 by A26,PSCOMP_1:108;
    then A32: q in L~Lower_Seq(C,n) by A28,A31,TOPREAL3:11;
      for x be set holds
     x in LSeg(p,q) /\ L~Lower_Seq(C,n) iff x = q
    proof
     let x be set;
     thus x in LSeg(p,q) /\ L~Lower_Seq(C,n) implies x = q
     proof
      assume A33: x in LSeg(p,q) /\ L~Lower_Seq(C,n);
      then reconsider pp = x as Point of TOP-REAL 2;
      A34: pp in LSeg(p,q) by A33,XBOOLE_0:def 3;
      then A35: pp`2 = q`2 by A15,A16,A23,A24,GOBOARD7:6;
      A36: pp`1 >= q`1 by A25,A34,TOPREAL1:9;
      reconsider EE = LSeg(f,p) /\ L~Lower_Seq(C,n)
                              as compact Subset of TOP-REAL 2 by PSCOMP_1:64;
      reconsider E0 = proj1.:EE as compact Subset of REAL by Th4;
      A37: f`2 = q`2 by A1,A2,A5,A9,A23,A24,GOBOARD5:2;
        f`1 <= q`1 by A1,A2,A5,A20,A22,A24,SPRECT_3:25;
      then A38: q in LSeg(p,f) by A15,A16,A23,A24,A25,A37,GOBOARD7:9;
        p in LSeg(f,p) by TOPREAL1:6;
      then A39: LSeg(p,q) c= LSeg(f,p) by A38,TOPREAL1:12;
        E0 is bounded by RCOMP_1:28;
      then A40: E0 is bounded_above by SEQ_4:def 3;
        pp in L~Lower_Seq(C,n) by A33,XBOOLE_0:def 3;
      then pp in EE by A34,A39,XBOOLE_0:def 3;
      then proj1.pp in E0 by FUNCT_2:43;
      then pp`1 in E0 by PSCOMP_1:def 28;
      then q`1 >= pp`1 by A21,A24,A40,SEQ_4:def 4;
      then pp`1 = q`1 by A36,AXIOMS:21;
      hence x = q by A35,TOPREAL3:11;
     end;
     assume A41: x = q;
     then x in LSeg(p,q) by TOPREAL1:6;
     hence x in LSeg(p,q) /\ L~Lower_Seq(C,n) by A32,A41,XBOOLE_0:def 3;
    end;
    hence LSeg(Gauge(C,n)*(j1,i),Gauge(C,n)*(k1,i)) /\ L~Lower_Seq(C,n) =
     {Gauge(C,n)*(j1,i)} by A16,A24,TARSKI:def 1;
    set X = LSeg(G*(j,i),G*(k,i)) /\ L~Upper_Seq(C,n);
    reconsider X1=X as non empty compact Subset of TOP-REAL 2
                                              by A4,A7,PSCOMP_1:64,XBOOLE_0:def
3;
    consider pp be set such that
     A42: pp in W-most X1 by XBOOLE_0:def 1;
    reconsider pp as Point of TOP-REAL 2 by A42;
      pp in LSeg(SW-corner X, NW-corner X)/\X by A42,PSCOMP_1:def 38;
    then A43: pp in X by XBOOLE_0:def 3;
    then A44: pp in L~Upper_Seq(C,n) by XBOOLE_0:def 3;
    A45: pp in LSeg(G*(j,i),G*(k,i)) by A43,XBOOLE_0:def 3;
      f`2 = s by A1,A2,A5,A9,GOBOARD5:2
       .= e`2 by A1,A2,A11,GOBOARD5:2;
    then LSeg(f,e) is horizontal by SPPOL_1:36;
    then A46: pp`2 = p`2 by A30,A45,SPRECT_3:19;
      p`1 = W-bound X by A13,A16,SPRECT_1:48
       .= (W-min X)`1 by PSCOMP_1:84
       .= pp`1 by A42,PSCOMP_1:88;
    then A47: p in L~Upper_Seq(C,n) by A44,A46,TOPREAL3:11;
      for x be set holds
     x in LSeg(p,q) /\ L~Upper_Seq(C,n) iff x = p
    proof
     let x be set;
     thus x in LSeg(p,q) /\ L~Upper_Seq(C,n) implies x = p
     proof
      assume A48: x in LSeg(p,q) /\ L~Upper_Seq(C,n);
      then reconsider pp = x as Point of TOP-REAL 2;
      A49: pp in LSeg(p,q) by A48,XBOOLE_0:def 3;
      then A50: pp`2 = p`2 by A15,A16,A23,A24,GOBOARD7:6;
      A51: pp`1 <= p`1 by A25,A49,TOPREAL1:9;
      reconsider EE = LSeg(f,e) /\ L~Upper_Seq(C,n)
                              as compact Subset of TOP-REAL 2 by PSCOMP_1:64;
      reconsider E0 = proj1.:EE as compact Subset of REAL by Th4;
      A52: f`2 = q`2 by A1,A2,A5,A9,A23,A24,GOBOARD5:2;
      A53: e`2 = q`2 by A1,A2,A11,A23,A24,GOBOARD5:2;
      A54: f`1 <= q`1 by A1,A2,A5,A20,A22,A24,SPRECT_3:25;
        j1 <= k by A12,A20,AXIOMS:22;
      then q`1 <= e`1 by A1,A2,A22,A24,SPRECT_3:25;
      then A55: q in LSeg(f,e) by A52,A53,A54,GOBOARD7:9;
      A56: f`2 = p`2 by A1,A2,A5,A9,A15,A16,GOBOARD5:2;
      A57: e`2 = p`2 by A1,A2,A11,A15,A16,GOBOARD5:2;
      A58: f`1 <= p`1 by A1,A2,A5,A12,A14,A16,SPRECT_3:25;
        p`1 <= e`1 by A1,A2,A12,A14,A16,SPRECT_3:25;
      then p in LSeg(f,e) by A56,A57,A58,GOBOARD7:9;
      then A59: LSeg(p,q) c= LSeg(f,e) by A55,TOPREAL1:12;
        E0 is bounded by RCOMP_1:28;
      then A60: E0 is bounded_below by SEQ_4:def 3;
        pp in L~Upper_Seq(C,n) by A48,XBOOLE_0:def 3;
      then pp in EE by A49,A59,XBOOLE_0:def 3;
      then proj1.pp in E0 by FUNCT_2:43;
      then pp`1 in E0 by PSCOMP_1:def 28;
      then p`1 <= pp`1 by A13,A16,A60,SEQ_4:def 5;
      then pp`1 = p`1 by A51,AXIOMS:21;
      hence x = p by A50,TOPREAL3:11;
     end;
     assume A61: x = p;
     then x in LSeg(p,q) by TOPREAL1:6;
     hence x in LSeg(p,q) /\ L~Upper_Seq(C,n) by A47,A61,XBOOLE_0:def 3;
    end;
    hence LSeg(Gauge(C,n)*(j1,i),Gauge(C,n)*(k1,i)) /\ L~Upper_Seq(C,n) =
      {Gauge(C,n)*(k1,i)} by A16,A24,TARSKI:def 1;
   end;

  theorem
     for C be compact connected non vertical non horizontal Subset of TOP-REAL
2
   for i,j,k be Nat st
    1 <= i & i <= len Gauge(C,n) &
    1 <= j & j <= k & k <= width Gauge(C,n) &
    Gauge(C,n)*(i,j) in L~Upper_Seq(C,n)
   ex j1 be Nat st
    j <= j1 & j1 <= k &
    LSeg(Gauge(C,n)*(i,j1),Gauge(C,n)*(i,k)) /\ L~Upper_Seq(C,n) =
     {Gauge(C,n)*(i,j1)}
   proof
    let C be compact connected non vertical non horizontal
     Subset of TOP-REAL 2;
    let i,j,k be Nat;
    assume that
     A1: 1 <= i & i <= len Gauge(C,n) and
     A2: 1 <= j & j <= k & k <= width Gauge(C,n) and
     A3: Gauge(C,n)*(i,j) in L~Upper_Seq(C,n);
    set G = Gauge(C,n);
    set s = G*(i,1)`1;
    set f = G*(i,j);
    set e = G*(i,k);
    set w2 = sup(proj2.:(LSeg(f,e) /\ L~Upper_Seq(C,n)));
    set q = |[s,w2]|;
    A4: j <= width G by A2,AXIOMS:22;
    then A5: [i,j] in Indices G by A1,A2,GOBOARD7:10;
    A6: k >= 1 by A2,AXIOMS:22;
    then A7: G*(i,k)`1 = s by A1,A2,GOBOARD5:3;
    A8: Upper_Seq(C,n) is_sequence_on G by JORDAN1G:4;
    A9: G*(i,j) in LSeg(G*(i,j),G*(i,k)) by TOPREAL1:6;
    then A10: LSeg(G*(i,j),G*(i,k)) meets L~Upper_Seq(C,n) by A3,XBOOLE_0:3;
      [i,k] in Indices G by A1,A2,A6,GOBOARD7:10;
    then consider j1 be Nat such that
     A11: j <= j1 & j1 <= k and
     A12: G*(i,j1)`2 = w2 by A2,A5,A8,A10,JORDAN1F:2;
    A13: 1 <= j1 & j1 <= width G by A2,A11,AXIOMS:22;
    then A14: G*(i,j1)`1 = s by A1,GOBOARD5:3;
    then A15: q = G*(i,j1) by A12,EUCLID:57;
    take j1;
    thus j <= j1 & j1 <= k by A11;
    A16: q`2 <= e`2 by A1,A2,A11,A13,A15,SPRECT_3:24;
    set X = LSeg(G*(i,j),G*(i,k)) /\ L~Upper_Seq(C,n);
    reconsider X1=X as non empty compact Subset of TOP-REAL 2
                                             by A3,A9,PSCOMP_1:64,XBOOLE_0:def
3;
    consider pp be set such that
     A17: pp in N-most X1 by XBOOLE_0:def 1;
    reconsider pp as Point of TOP-REAL 2 by A17;
      pp in LSeg(NW-corner X, NE-corner X)/\X by A17,PSCOMP_1:def 39;
    then A18: pp in X by XBOOLE_0:def 3;
    then A19: pp in L~Upper_Seq(C,n) by XBOOLE_0:def 3;
    A20: pp in LSeg(G*(i,j),G*(i,k)) by A18,XBOOLE_0:def 3;
      f`1 = e`1 by A1,A2,A4,A7,GOBOARD5:3;
    then LSeg(f,e) is vertical by SPPOL_1:37;
    then A21: pp`1 = q`1 by A7,A14,A15,A20,SPRECT_3:20;
      q`2 = N-bound X by A12,A15,SPRECT_1:50
       .= (N-min X)`2 by PSCOMP_1:94
       .= pp`2 by A17,PSCOMP_1:98;
    then A22: q in L~Upper_Seq(C,n) by A19,A21,TOPREAL3:11;
      for x be set holds
     x in LSeg(e,q) /\ L~Upper_Seq(C,n) iff x = q
    proof
     let x be set;
     thus x in LSeg(e,q) /\ L~Upper_Seq(C,n) implies x = q
     proof
      assume A23: x in LSeg(e,q) /\ L~Upper_Seq(C,n);
      then reconsider pp = x as Point of TOP-REAL 2;
      A24: pp in LSeg(e,q) by A23,XBOOLE_0:def 3;
      then A25: pp`1 = q`1 by A7,A14,A15,GOBOARD7:5;
      A26: pp`2 >= q`2 by A16,A24,TOPREAL1:10;
      reconsider EE = LSeg(f,e) /\ L~Upper_Seq(C,n)
                              as compact Subset of TOP-REAL 2 by PSCOMP_1:64;
      reconsider E0 = proj2.:EE as compact Subset of REAL by JORDAN1A:28;
      A27: f`1 = q`1 by A1,A2,A4,A14,A15,GOBOARD5:3;
        f`2 <= q`2 by A1,A2,A11,A13,A15,SPRECT_3:24;
      then A28: q in LSeg(e,f) by A7,A14,A15,A16,A27,GOBOARD7:8;
        e in LSeg(f,e) by TOPREAL1:6;
      then A29: LSeg(e,q) c= LSeg(f,e) by A28,TOPREAL1:12;
        E0 is bounded by RCOMP_1:28;
      then A30: E0 is bounded_above by SEQ_4:def 3;
        pp in L~Upper_Seq(C,n) by A23,XBOOLE_0:def 3;
      then pp in EE by A24,A29,XBOOLE_0:def 3;
      then proj2.pp in E0 by FUNCT_2:43;
      then pp`2 in E0 by PSCOMP_1:def 29;
      then q`2 >= pp`2 by A12,A15,A30,SEQ_4:def 4;
      then pp`2 = q`2 by A26,AXIOMS:21;
      hence x = q by A25,TOPREAL3:11;
     end;
     assume A31: x = q;
     then x in LSeg(e,q) by TOPREAL1:6;
     hence x in LSeg(e,q) /\ L~Upper_Seq(C,n) by A22,A31,XBOOLE_0:def 3;
    end;
    hence LSeg(Gauge(C,n)*(i,j1),Gauge(C,n)*(i,k)) /\ L~Upper_Seq(C,n) =
     {Gauge(C,n)*(i,j1)} by A15,TARSKI:def 1;
   end;

  theorem
     for C be compact connected non vertical non horizontal Subset of TOP-REAL
2
   for i,j,k be Nat st
    1 <= i & i <= len Gauge(C,n) &
    1 <= j & j <= k & k <= width Gauge(C,n) &
    Gauge(C,n)*(i,k) in L~Lower_Seq(C,n)
   ex k1 be Nat st
    j <= k1 & k1 <= k &
    LSeg(Gauge(C,n)*(i,j),Gauge(C,n)*(i,k1)) /\ L~Lower_Seq(C,n) =
     {Gauge(C,n)*(i,k1)}
   proof
    let C be compact connected non vertical non horizontal
     Subset of TOP-REAL 2;
    let i,j,k be Nat;
    assume that
     A1: 1 <= i & i <= len Gauge(C,n) and
     A2: 1 <= j & j <= k & k <= width Gauge(C,n) and
     A3: Gauge(C,n)*(i,k) in L~Lower_Seq(C,n);
    set G = Gauge(C,n);
    set s = G*(i,1)`1;
    set f = G*(i,j);
    set e = G*(i,k);
    set w1 = inf(proj2.:(LSeg(f,e) /\ L~Lower_Seq(C,n)));
    set p = |[s,w1]|;
    A4: Lower_Seq(C,n) is_sequence_on G by JORDAN1G:5;
    A5: G*(i,k) in LSeg(G*(i,j),G*(i,k)) by TOPREAL1:6;
    then A6: LSeg(G*(i,j),G*(i,k)) meets L~Lower_Seq(C,n) by A3,XBOOLE_0:3;
    A7: j <= width G by A2,AXIOMS:22;
    then A8: [i,j] in Indices G by A1,A2,GOBOARD7:10;
    A9: k >= 1 by A2,AXIOMS:22;
    then [i,k] in Indices G by A1,A2,GOBOARD7:10;
    then consider k1 be Nat such that
     A10: j <= k1 & k1 <= k and
     A11: G*(i,k1)`2 = w1 by A2,A4,A6,A8,JORDAN1F:1;
    A12: 1 <= k1 & k1 <= width G by A2,A10,AXIOMS:22;
    then A13: G*(i,k1)`1 = s by A1,GOBOARD5:3;
    then A14: p = G*(i,k1) by A11,EUCLID:57;
    take k1;
    thus j <= k1 & k1 <= k by A10;
    A15: f`2 <= p`2 by A1,A2,A10,A12,A14,SPRECT_3:24;
    A16: f`1 = p`1 by A1,A2,A7,A13,A14,GOBOARD5:3;
    set X = LSeg(G*(i,j),G*(i,k)) /\ L~Lower_Seq(C,n);
    reconsider X1=X as non empty compact Subset of TOP-REAL 2
                                            by A3,A5,PSCOMP_1:64,XBOOLE_0:def 3
;
    consider pp be set such that
     A17: pp in S-most X1 by XBOOLE_0:def 1;
    reconsider pp as Point of TOP-REAL 2 by A17;
      pp in LSeg(SW-corner X, SE-corner X)/\X by A17,PSCOMP_1:def 41;
    then A18: pp in X by XBOOLE_0:def 3;
    then A19: pp in L~Lower_Seq(C,n) by XBOOLE_0:def 3;
    A20: pp in LSeg(G*(i,j),G*(i,k)) by A18,XBOOLE_0:def 3;
      f`1 = s by A1,A2,A7,GOBOARD5:3
       .= e`1 by A1,A2,A9,GOBOARD5:3;
    then LSeg(f,e) is vertical by SPPOL_1:37;
    then A21: pp`1 = p`1 by A16,A20,SPRECT_3:20;
      p`2 = S-bound X by A11,A14,SPRECT_1:49
       .= (S-min X)`2 by PSCOMP_1:114
       .= pp`2 by A17,PSCOMP_1:118;
    then A22: p in L~Lower_Seq(C,n) by A19,A21,TOPREAL3:11;
      for x be set holds
     x in LSeg(p,f) /\ L~Lower_Seq(C,n) iff x = p
    proof
     let x be set;
     thus x in LSeg(p,f) /\ L~Lower_Seq(C,n) implies x = p
     proof
      assume A23: x in LSeg(p,f) /\ L~Lower_Seq(C,n);
      then reconsider pp = x as Point of TOP-REAL 2;
      A24: pp in LSeg(p,f) by A23,XBOOLE_0:def 3;
      then A25: pp`1 = p`1 by A16,GOBOARD7:5;
      A26: pp`2 <= p`2 by A15,A24,TOPREAL1:10;
      reconsider EE = LSeg(f,e) /\ L~Lower_Seq(C,n)
                              as compact Subset of TOP-REAL 2 by PSCOMP_1:64;
      reconsider E0 = proj2.:EE as compact Subset of REAL by JORDAN1A:28;
      A27: f in LSeg(f,e) by TOPREAL1:6;
      A28: f`1 = p`1 by A1,A2,A7,A13,A14,GOBOARD5:3;
      A29: e`1 = p`1 by A1,A2,A9,A13,A14,GOBOARD5:3;
      A30: f`2 <= p`2 by A1,A2,A10,A12,A14,SPRECT_3:24;
        p`2 <= e`2 by A1,A2,A10,A12,A14,SPRECT_3:24;
      then p in LSeg(f,e) by A28,A29,A30,GOBOARD7:8;
      then A31: LSeg(p,f) c= LSeg(f,e) by A27,TOPREAL1:12;
        E0 is bounded by RCOMP_1:28;
      then A32: E0 is bounded_below by SEQ_4:def 3;
        pp in L~Lower_Seq(C,n) by A23,XBOOLE_0:def 3;
      then pp in EE by A24,A31,XBOOLE_0:def 3;
      then proj2.pp in E0 by FUNCT_2:43;
      then pp`2 in E0 by PSCOMP_1:def 29;
      then p`2 <= pp`2 by A11,A14,A32,SEQ_4:def 5;
      then pp`2 = p`2 by A26,AXIOMS:21;
      hence x = p by A25,TOPREAL3:11;
     end;
     assume A33: x = p;
     then x in LSeg(p,f) by TOPREAL1:6;
     hence x in LSeg(p,f) /\ L~Lower_Seq(C,n) by A22,A33,XBOOLE_0:def 3;
    end;
    hence LSeg(Gauge(C,n)*(i,j),Gauge(C,n)*(i,k1)) /\ L~Lower_Seq(C,n) =
     {Gauge(C,n)*(i,k1)} by A14,TARSKI:def 1;
   end;

  theorem
     for C be compact connected non vertical non horizontal Subset of TOP-REAL
2
   for i,j,k be Nat st
    1 <= i & i <= len Gauge(C,n) &
    1 <= j & j <= k & k <= width Gauge(C,n) &
    Gauge(C,n)*(i,j) in L~Upper_Seq(C,n) &
    Gauge(C,n)*(i,k) in L~Lower_Seq(C,n)
   ex j1,k1 be Nat st
    j <= j1 & j1 <= k1 & k1 <= k &
    LSeg(Gauge(C,n)*(i,j1),Gauge(C,n)*(i,k1)) /\ L~Upper_Seq(C,n) =
     {Gauge(C,n)*(i,j1)} &
    LSeg(Gauge(C,n)*(i,j1),Gauge(C,n)*(i,k1)) /\ L~Lower_Seq(C,n) =
     {Gauge(C,n)*(i,k1)}
   proof
    let C be compact connected non vertical non horizontal
     Subset of TOP-REAL 2;
    let i,j,k be Nat;
    assume that
     A1: 1 <= i & i <= len Gauge(C,n) and
     A2: 1 <= j & j <= k & k <= width Gauge(C,n) and
     A3: Gauge(C,n)*(i,j) in L~Upper_Seq(C,n) and
     A4: Gauge(C,n)*(i,k) in L~Lower_Seq(C,n);
    set G = Gauge(C,n);
    set s = G*(i,1)`1;
    set f = G*(i,j);
    set e = G*(i,k);
    set w1 = inf(proj2.:(LSeg(f,e) /\ L~Lower_Seq(C,n)));
    set p = |[s,w1]|;
    set w2 = sup(proj2.:(LSeg(f,p) /\ L~Upper_Seq(C,n)));
    set q = |[s,w2]|;
    A5: Lower_Seq(C,n) is_sequence_on G by JORDAN1G:5;
    A6: G*(i,k) in LSeg(G*(i,j),G*(i,k)) by TOPREAL1:6;
    then A7: LSeg(G*(i,j),G*(i,k)) meets L~Lower_Seq(C,n) by A4,XBOOLE_0:3;
    A8: j <= width G by A2,AXIOMS:22;
    then A9: [i,j] in Indices G by A1,A2,GOBOARD7:10;
    A10: k >= 1 by A2,AXIOMS:22;
    then [i,k] in Indices G by A1,A2,GOBOARD7:10;
    then consider k1 be Nat such that
     A11: j <= k1 & k1 <= k and
     A12: G*(i,k1)`2 = w1 by A2,A5,A7,A9,JORDAN1F:1;
    A13: 1 <= k1 & k1 <= width G by A2,A11,AXIOMS:22;
    then A14: G*(i,k1)`1 = s by A1,GOBOARD5:3;
    then A15: p = G*(i,k1) by A12,EUCLID:57;
    A16: Upper_Seq(C,n) is_sequence_on G by JORDAN1G:4;
    A17: G*(i,j) in LSeg(G*(i,j),G*(i,k1)) by TOPREAL1:6;
    then A18: LSeg(G*(i,j),G*(i,k1)) meets L~Upper_Seq(C,n) by A3,XBOOLE_0:3;
      [i,k1] in Indices G by A1,A13,GOBOARD7:10;
    then consider j1 be Nat such that
     A19: j <= j1 & j1 <= k1 and
     A20: G*(i,j1)`2 = w2 by A9,A11,A15,A16,A18,JORDAN1F:2;
    A21: 1 <= j1 & j1 <= width G by A2,A13,A19,AXIOMS:22;
    then A22: G*(i,j1)`1 = s by A1,GOBOARD5:3;
    then A23: q = G*(i,j1) by A20,EUCLID:57;
    take j1,k1;
    thus j <= j1 & j1 <= k1 & k1 <= k by A11,A19;
    A24: q`2 <= p`2 by A1,A13,A15,A19,A21,A23,SPRECT_3:24;
    set X = LSeg(G*(i,j),G*(i,k1)) /\ L~Upper_Seq(C,n);
    reconsider X1=X as non empty compact Subset of TOP-REAL 2
                                              by A3,A17,PSCOMP_1:64,XBOOLE_0:
def 3;
    consider pp be set such that
     A25: pp in N-most X1 by XBOOLE_0:def 1;
    reconsider pp as Point of TOP-REAL 2 by A25;
      pp in LSeg(NW-corner X, NE-corner X)/\X by A25,PSCOMP_1:def 39;
    then A26: pp in X by XBOOLE_0:def 3;
    then A27: pp in L~Upper_Seq(C,n) by XBOOLE_0:def 3;
    A28: pp in LSeg(G*(i,j),G*(i,k1)) by A26,XBOOLE_0:def 3;
    A29: f`1 = p`1 by A1,A2,A8,A14,A15,GOBOARD5:3;
    then LSeg(f,p) is vertical by SPPOL_1:37;
    then A30: pp`1 = q`1 by A14,A15,A22,A23,A28,SPRECT_3:20;
      q`2 = N-bound X by A15,A20,A23,SPRECT_1:50
       .= (N-min X)`2 by PSCOMP_1:94
       .= pp`2 by A25,PSCOMP_1:98;
    then A31: q in L~Upper_Seq(C,n) by A27,A30,TOPREAL3:11;
      for x be set holds
     x in LSeg(p,q) /\ L~Upper_Seq(C,n) iff x = q
    proof
     let x be set;
     thus x in LSeg(p,q) /\ L~Upper_Seq(C,n) implies x = q
     proof
      assume A32: x in LSeg(p,q) /\ L~Upper_Seq(C,n);
      then reconsider pp = x as Point of TOP-REAL 2;
      A33: pp in LSeg(p,q) by A32,XBOOLE_0:def 3;
      then A34: pp`1 = q`1 by A14,A15,A22,A23,GOBOARD7:5;
      A35: pp`2 >= q`2 by A24,A33,TOPREAL1:10;
      reconsider EE = LSeg(f,p) /\ L~Upper_Seq(C,n)
                              as compact Subset of TOP-REAL 2 by PSCOMP_1:64;
      reconsider E0 = proj2.:EE as compact Subset of REAL by JORDAN1A:28;
      A36: f`1 = q`1 by A1,A2,A8,A22,A23,GOBOARD5:3;
        f`2 <= q`2 by A1,A2,A19,A21,A23,SPRECT_3:24;
      then A37: q in LSeg(p,f) by A14,A15,A22,A23,A24,A36,GOBOARD7:8;
        p in LSeg(f,p) by TOPREAL1:6;
      then A38: LSeg(p,q) c= LSeg(f,p) by A37,TOPREAL1:12;
        E0 is bounded by RCOMP_1:28;
      then A39: E0 is bounded_above by SEQ_4:def 3;
        pp in L~Upper_Seq(C,n) by A32,XBOOLE_0:def 3;
      then pp in EE by A33,A38,XBOOLE_0:def 3;
      then proj2.pp in E0 by FUNCT_2:43;
      then pp`2 in E0 by PSCOMP_1:def 29;
      then q`2 >= pp`2 by A20,A23,A39,SEQ_4:def 4;
      then pp`2 = q`2 by A35,AXIOMS:21;
      hence x = q by A34,TOPREAL3:11;
     end;
     assume A40: x = q;
     then x in LSeg(p,q) by TOPREAL1:6;
     hence x in LSeg(p,q) /\ L~Upper_Seq(C,n) by A31,A40,XBOOLE_0:def 3;
    end;
    hence LSeg(Gauge(C,n)*(i,j1),
     Gauge(C,n)*(i,k1)) /\ L~Upper_Seq(C,n) =
      {Gauge(C,n)*(i,j1)} by A15,A23,TARSKI:def 1;
    set X = LSeg(G*(i,j),G*(i,k)) /\ L~Lower_Seq(C,n);
    reconsider X1=X as non empty compact Subset of TOP-REAL 2
                                            by A4,A6,PSCOMP_1:64,XBOOLE_0:def 3
;
    consider pp be set such that
     A41: pp in S-most X1 by XBOOLE_0:def 1;
    reconsider pp as Point of TOP-REAL 2 by A41;
      pp in LSeg(SW-corner X, SE-corner X)/\X by A41,PSCOMP_1:def 41;
    then A42: pp in X by XBOOLE_0:def 3;
    then A43: pp in L~Lower_Seq(C,n) by XBOOLE_0:def 3;
    A44: pp in LSeg(G*(i,j),G*(i,k)) by A42,XBOOLE_0:def 3;
      f`1 = s by A1,A2,A8,GOBOARD5:3
       .= e`1 by A1,A2,A10,GOBOARD5:3;
    then LSeg(f,e) is vertical by SPPOL_1:37;
    then A45: pp`1 = p`1 by A29,A44,SPRECT_3:20;
      p`2 = S-bound X by A12,A15,SPRECT_1:49
       .= (S-min X)`2 by PSCOMP_1:114
       .= pp`2 by A41,PSCOMP_1:118;
    then A46: p in L~Lower_Seq(C,n) by A43,A45,TOPREAL3:11;
      for x be set holds
     x in LSeg(p,q) /\ L~Lower_Seq(C,n) iff x = p
    proof
     let x be set;
     thus x in LSeg(p,q) /\ L~Lower_Seq(C,n) implies x = p
     proof
      assume A47: x in LSeg(p,q) /\ L~Lower_Seq(C,n);
      then reconsider pp = x as Point of TOP-REAL 2;
      A48: pp in LSeg(p,q) by A47,XBOOLE_0:def 3;
      then A49: pp`1 = p`1 by A14,A15,A22,A23,GOBOARD7:5;
      A50: pp`2 <= p`2 by A24,A48,TOPREAL1:10;
      reconsider EE = LSeg(f,e) /\ L~Lower_Seq(C,n)
                              as compact Subset of TOP-REAL 2 by PSCOMP_1:64;
      reconsider E0 = proj2.:EE as compact Subset of REAL by JORDAN1A:28;
      A51: f`1 = q`1 by A1,A2,A8,A22,A23,GOBOARD5:3;
      A52: e`1 = q`1 by A1,A2,A10,A22,A23,GOBOARD5:3;
      A53: f`2 <= q`2 by A1,A2,A19,A21,A23,SPRECT_3:24;
        j1 <= k by A11,A19,AXIOMS:22;
      then q`2 <= e`2 by A1,A2,A21,A23,SPRECT_3:24;
      then A54: q in LSeg(f,e) by A51,A52,A53,GOBOARD7:8;
      A55: f`1 = p`1 by A1,A2,A8,A14,A15,GOBOARD5:3;
      A56: e`1 = p`1 by A1,A2,A10,A14,A15,GOBOARD5:3;
      A57: f`2 <= p`2 by A1,A2,A11,A13,A15,SPRECT_3:24;
        p`2 <= e`2 by A1,A2,A11,A13,A15,SPRECT_3:24;
      then p in LSeg(f,e) by A55,A56,A57,GOBOARD7:8;
      then A58: LSeg(p,q) c= LSeg(f,e) by A54,TOPREAL1:12;
        E0 is bounded by RCOMP_1:28;
      then A59: E0 is bounded_below by SEQ_4:def 3;
        pp in L~Lower_Seq(C,n) by A47,XBOOLE_0:def 3;
      then pp in EE by A48,A58,XBOOLE_0:def 3;
      then proj2.pp in E0 by FUNCT_2:43;
      then pp`2 in E0 by PSCOMP_1:def 29;
      then p`2 <= pp`2 by A12,A15,A59,SEQ_4:def 5;
      then pp`2 = p`2 by A50,AXIOMS:21;
      hence x = p by A49,TOPREAL3:11;
     end;
     assume A60: x = p;
     then x in LSeg(p,q) by TOPREAL1:6;
     hence x in LSeg(p,q) /\ L~Lower_Seq(C,n) by A46,A60,XBOOLE_0:def 3;
    end;
    hence LSeg(Gauge(C,n)*(i,j1),
     Gauge(C,n)*(i,k1)) /\ L~Lower_Seq(C,n) =
      {Gauge(C,n)*(i,k1)} by A15,A23,TARSKI:def 1;
   end;

  theorem Th20:
   for C be compact connected non vertical non horizontal Subset of TOP-REAL 2
   for i,j,k be Nat st
    1 <= j & j <= k & k <= len Gauge(C,n) &
    1 <= i & i <= width Gauge(C,n) &
    Gauge(C,n)*(j,i) in L~Upper_Seq(C,n)
   ex j1 be Nat st
    j <= j1 & j1 <= k &
    LSeg(Gauge(C,n)*(j1,i),Gauge(C,n)*(k,i)) /\ L~Upper_Seq(C,n) =
     {Gauge(C,n)*(j1,i)}
   proof
    let C be compact connected non vertical non horizontal
     Subset of TOP-REAL 2;
    let i,j,k be Nat;
    assume that
     A1: 1 <= j & j <= k & k <= len Gauge(C,n) and
     A2: 1 <= i & i <= width Gauge(C,n) and
     A3: Gauge(C,n)*(j,i) in L~Upper_Seq(C,n);
    set G = Gauge(C,n);
    set s = G*(1,i)`2;
    set f = G*(j,i);
    set e = G*(k,i);
    set w2 = sup(proj1.:(LSeg(f,e) /\ L~Upper_Seq(C,n)));
    set q = |[w2,s]|;
    A4: len G = width G by JORDAN8:def 1;
    then A5: j <= width G by A1,AXIOMS:22;
    then A6: [j,i] in Indices G by A1,A2,A4,GOBOARD7:10;
    A7: k >= 1 by A1,AXIOMS:22;
    then A8: G*(k,i)`2 = s by A1,A2,GOBOARD5:2;
    A9: Upper_Seq(C,n) is_sequence_on G by JORDAN1G:4;
    A10: G*(j,i) in LSeg(G*(j,i),G*(k,i)) by TOPREAL1:6;
    then A11: LSeg(G*(j,i),G*(k,i)) meets L~Upper_Seq(C,n) by A3,XBOOLE_0:3;
      [k,i] in Indices G by A1,A2,A7,GOBOARD7:10;
    then consider j1 be Nat such that
     A12: j <= j1 & j1 <= k and
     A13: G*(j1,i)`1 = w2 by A1,A6,A9,A11,JORDAN1F:4;
    A14: 1 <= j1 & j1 <= width G by A1,A4,A12,AXIOMS:22;
    then A15: G*(j1,i)`2 = s by A2,A4,GOBOARD5:2;
    then A16: q = G*(j1,i) by A13,EUCLID:57;
    take j1;
    thus j <= j1 & j1 <= k by A12;
    A17: q`1 <= e`1 by A1,A2,A12,A14,A16,SPRECT_3:25;
    set X = LSeg(G*(j,i),G*(k,i)) /\ L~Upper_Seq(C,n);
    reconsider X1=X as non empty compact Subset of TOP-REAL 2
                                               by A3,A10,PSCOMP_1:64,XBOOLE_0:
def 3;
    consider pp be set such that
     A18: pp in E-most X1 by XBOOLE_0:def 1;
    reconsider pp as Point of TOP-REAL 2 by A18;
      pp in LSeg(SE-corner X, NE-corner X)/\X by A18,PSCOMP_1:def 40;
    then A19: pp in X by XBOOLE_0:def 3;
    then A20: pp in L~Upper_Seq(C,n) by XBOOLE_0:def 3;
    A21: pp in LSeg(G*(j,i),G*(k,i)) by A19,XBOOLE_0:def 3;
      f`2 = e`2 by A1,A2,A4,A5,A8,GOBOARD5:2;
    then LSeg(f,e) is horizontal by SPPOL_1:36;
    then A22: pp`2 = q`2 by A8,A15,A16,A21,SPRECT_3:19;
      q`1 = E-bound X by A13,A16,SPRECT_1:51
       .= (E-min X)`1 by PSCOMP_1:104
       .= pp`1 by A18,PSCOMP_1:108;
    then A23: q in L~Upper_Seq(C,n) by A20,A22,TOPREAL3:11;
      for x be set holds
     x in LSeg(e,q) /\ L~Upper_Seq(C,n) iff x = q
    proof
     let x be set;
     thus x in LSeg(e,q) /\ L~Upper_Seq(C,n) implies x = q
     proof
      assume A24: x in LSeg(e,q) /\ L~Upper_Seq(C,n);
      then reconsider pp = x as Point of TOP-REAL 2;
      A25: pp in LSeg(e,q) by A24,XBOOLE_0:def 3;
      then A26: pp`2 = q`2 by A8,A15,A16,GOBOARD7:6;
      A27: pp`1 >= q`1 by A17,A25,TOPREAL1:9;
      reconsider EE = LSeg(f,e) /\ L~Upper_Seq(C,n)
                              as compact Subset of TOP-REAL 2 by PSCOMP_1:64;
      reconsider E0 = proj1.:EE as compact Subset of REAL by Th4;
      A28: f`2 = q`2 by A1,A2,A4,A5,A15,A16,GOBOARD5:2;
        f`1 <= q`1 by A1,A2,A4,A12,A14,A16,SPRECT_3:25;
      then A29: q in LSeg(e,f) by A8,A15,A16,A17,A28,GOBOARD7:9;
        e in LSeg(f,e) by TOPREAL1:6;
      then A30: LSeg(e,q) c= LSeg(f,e) by A29,TOPREAL1:12;
        E0 is bounded by RCOMP_1:28;
      then A31: E0 is bounded_above by SEQ_4:def 3;
        pp in L~Upper_Seq(C,n) by A24,XBOOLE_0:def 3;
      then pp in EE by A25,A30,XBOOLE_0:def 3;
      then proj1.pp in E0 by FUNCT_2:43;
      then pp`1 in E0 by PSCOMP_1:def 28;
      then q`1 >= pp`1 by A13,A16,A31,SEQ_4:def 4;
      then pp`1 = q`1 by A27,AXIOMS:21;
      hence x = q by A26,TOPREAL3:11;
     end;
     assume A32: x = q;
     then x in LSeg(e,q) by TOPREAL1:6;
     hence x in LSeg(e,q) /\ L~Upper_Seq(C,n) by A23,A32,XBOOLE_0:def 3;
    end;
    hence LSeg(Gauge(C,n)*(j1,i),Gauge(C,n)*(k,i)) /\ L~Upper_Seq(C,n) =
     {Gauge(C,n)*(j1,i)} by A16,TARSKI:def 1;
   end;

  theorem
     for C be compact connected non vertical non horizontal Subset of TOP-REAL
2
   for i,j,k be Nat st
    1 <= j & j <= k & k <= len Gauge(C,n) &
    1 <= i & i <= width Gauge(C,n) &
    Gauge(C,n)*(k,i) in L~Lower_Seq(C,n)
   ex k1 be Nat st
    j <= k1 & k1 <= k &
    LSeg(Gauge(C,n)*(j,i),Gauge(C,n)*(k1,i)) /\ L~Lower_Seq(C,n) =
     {Gauge(C,n)*(k1,i)}
   proof
    let C be compact connected non vertical non horizontal
     Subset of TOP-REAL 2;
    let i,j,k be Nat;
    assume that
     A1: 1 <= j & j <= k & k <= len Gauge(C,n) and
     A2: 1 <= i & i <= width Gauge(C,n) and
     A3: Gauge(C,n)*(k,i) in L~Lower_Seq(C,n);
    set G = Gauge(C,n);
    set s = G*(1,i)`2;
    set f = G*(j,i);
    set e = G*(k,i);
    set w1 = inf(proj1.:(LSeg(f,e) /\ L~Lower_Seq(C,n)));
    set p = |[w1,s]|;
    A4: len G = width G by JORDAN8:def 1;
    A5: Lower_Seq(C,n) is_sequence_on G by JORDAN1G:5;
    A6: G*(k,i) in LSeg(G*(j,i),G*(k,i)) by TOPREAL1:6;
    then A7: LSeg(G*(j,i),G*(k,i)) meets L~Lower_Seq(C,n) by A3,XBOOLE_0:3;
    A8: j <= width G by A1,A4,AXIOMS:22;
    then A9: [j,i] in Indices G by A1,A2,A4,GOBOARD7:10;
    A10: k >= 1 by A1,AXIOMS:22;
    then [k,i] in Indices G by A1,A2,GOBOARD7:10;
    then consider k1 be Nat such that
     A11: j <= k1 & k1 <= k and
     A12: G*(k1,i)`1 = w1 by A1,A5,A7,A9,JORDAN1F:3;
    A13: 1 <= k1 & k1 <= width G by A1,A4,A11,AXIOMS:22;
    then A14: G*(k1,i)`2 = s by A2,A4,GOBOARD5:2;
    then A15: p = G*(k1,i) by A12,EUCLID:57;
    take k1;
    thus j <= k1 & k1 <= k by A11;
    A16: f`1 <= p`1 by A1,A2,A4,A11,A13,A15,SPRECT_3:25;
    A17: f`2 = p`2 by A1,A2,A4,A8,A14,A15,GOBOARD5:2;
    set X = LSeg(G*(j,i),G*(k,i)) /\ L~Lower_Seq(C,n);
    reconsider X1=X as non empty compact Subset of TOP-REAL 2
                                              by A3,A6,PSCOMP_1:64,XBOOLE_0:def
3;
    consider pp be set such that
     A18: pp in W-most X1 by XBOOLE_0:def 1;
    reconsider pp as Point of TOP-REAL 2 by A18;
      pp in LSeg(SW-corner X, NW-corner X)/\X by A18,PSCOMP_1:def 38;
    then A19: pp in X by XBOOLE_0:def 3;
    then A20: pp in L~Lower_Seq(C,n) by XBOOLE_0:def 3;
    A21: pp in LSeg(G*(j,i),G*(k,i)) by A19,XBOOLE_0:def 3;
      f`2 = s by A1,A2,A4,A8,GOBOARD5:2
       .= e`2 by A1,A2,A10,GOBOARD5:2;
    then LSeg(f,e) is horizontal by SPPOL_1:36;
    then A22: pp`2 = p`2 by A17,A21,SPRECT_3:19;
      p`1 = W-bound X by A12,A15,SPRECT_1:48
       .= (W-min X)`1 by PSCOMP_1:84
       .= pp`1 by A18,PSCOMP_1:88;
    then A23: p in L~Lower_Seq(C,n) by A20,A22,TOPREAL3:11;
      for x be set holds
     x in LSeg(p,f) /\ L~Lower_Seq(C,n) iff x = p
    proof
     let x be set;
     thus x in LSeg(p,f) /\ L~Lower_Seq(C,n) implies x = p
     proof
      assume A24: x in LSeg(p,f) /\ L~Lower_Seq(C,n);
      then reconsider pp = x as Point of TOP-REAL 2;
      A25: pp in LSeg(p,f) by A24,XBOOLE_0:def 3;
      then A26: pp`2 = p`2 by A17,GOBOARD7:6;
      A27: pp`1 <= p`1 by A16,A25,TOPREAL1:9;
      reconsider EE = LSeg(f,e) /\ L~Lower_Seq(C,n)
                              as compact Subset of TOP-REAL 2 by PSCOMP_1:64;
      reconsider E0 = proj1.:EE as compact Subset of REAL by Th4;
      A28: f in LSeg(f,e) by TOPREAL1:6;
      A29: f`2 = p`2 by A1,A2,A4,A8,A14,A15,GOBOARD5:2;
      A30: e`2 = p`2 by A1,A2,A10,A14,A15,GOBOARD5:2;
      A31: f`1 <= p`1 by A1,A2,A4,A11,A13,A15,SPRECT_3:25;
        p`1 <= e`1 by A1,A2,A11,A13,A15,SPRECT_3:25;
      then p in LSeg(f,e) by A29,A30,A31,GOBOARD7:9;
      then A32: LSeg(p,f) c= LSeg(f,e) by A28,TOPREAL1:12;
        E0 is bounded by RCOMP_1:28;
      then A33: E0 is bounded_below by SEQ_4:def 3;
        pp in L~Lower_Seq(C,n) by A24,XBOOLE_0:def 3;
      then pp in EE by A25,A32,XBOOLE_0:def 3;
      then proj1.pp in E0 by FUNCT_2:43;
      then pp`1 in E0 by PSCOMP_1:def 28;
      then p`1 <= pp`1 by A12,A15,A33,SEQ_4:def 5;
      then pp`1 = p`1 by A27,AXIOMS:21;
      hence x = p by A26,TOPREAL3:11;
     end;
     assume A34: x = p;
     then x in LSeg(p,f) by TOPREAL1:6;
     hence x in LSeg(p,f) /\ L~Lower_Seq(C,n) by A23,A34,XBOOLE_0:def 3;
    end;
    hence LSeg(Gauge(C,n)*(j,i),Gauge(C,n)*(k1,i)) /\ L~Lower_Seq(C,n) =
     {Gauge(C,n)*(k1,i)} by A15,TARSKI:def 1;
   end;

  theorem Th22:
   for C be compact connected non vertical non horizontal Subset of TOP-REAL 2
   for i,j,k be Nat st
    1 <= j & j <= k & k <= len Gauge(C,n) &
    1 <= i & i <= width Gauge(C,n) &
    Gauge(C,n)*(j,i) in L~Upper_Seq(C,n) &
    Gauge(C,n)*(k,i) in L~Lower_Seq(C,n)
   ex j1,k1 be Nat st
    j <= j1 & j1 <= k1 & k1 <= k &
    LSeg(Gauge(C,n)*(j1,i),Gauge(C,n)*(k1,i)) /\ L~Upper_Seq(C,n) =
     {Gauge(C,n)*(j1,i)} &
    LSeg(Gauge(C,n)*(j1,i),Gauge(C,n)*(k1,i)) /\ L~Lower_Seq(C,n) =
     {Gauge(C,n)*(k1,i)}
   proof
    let C be compact connected non vertical non horizontal
     Subset of TOP-REAL 2;
    let i,j,k be Nat;
    assume that
     A1: 1 <= j & j <= k & k <= len Gauge(C,n) and
     A2: 1 <= i & i <= width Gauge(C,n) and
     A3: Gauge(C,n)*(j,i) in L~Upper_Seq(C,n) and
     A4: Gauge(C,n)*(k,i) in L~Lower_Seq(C,n);
    set G = Gauge(C,n);
    set s = G*(1,i)`2;
    set f = G*(j,i);
    set e = G*(k,i);
    set w1 = inf(proj1.:(LSeg(f,e) /\ L~Lower_Seq(C,n)));
    set p = |[w1,s]|;
    set w2 = sup(proj1.:(LSeg(f,p) /\ L~Upper_Seq(C,n)));
    set q = |[w2,s]|;
    A5: len G = width G by JORDAN8:def 1;
    A6: Lower_Seq(C,n) is_sequence_on G by JORDAN1G:5;
    A7: G*(k,i) in LSeg(G*(j,i),G*(k,i)) by TOPREAL1:6;
    then A8: LSeg(G*(j,i),G*(k,i)) meets L~Lower_Seq(C,n) by A4,XBOOLE_0:3;
    A9: j <= width G by A1,A5,AXIOMS:22;
    then A10: [j,i] in Indices G by A1,A2,A5,GOBOARD7:10;
    A11: k >= 1 by A1,AXIOMS:22;
    then [k,i] in Indices G by A1,A2,GOBOARD7:10;
    then consider k1 be Nat such that
     A12: j <= k1 & k1 <= k and
     A13: G*(k1,i)`1 = w1 by A1,A6,A8,A10,JORDAN1F:3;
    A14: 1 <= k1 & k1 <= width G by A1,A5,A12,AXIOMS:22;
    then A15: G*(k1,i)`2 = s by A2,A5,GOBOARD5:2;
    then A16: p = G*(k1,i) by A13,EUCLID:57;
    A17: Upper_Seq(C,n) is_sequence_on G by JORDAN1G:4;
    A18: G*(j,i) in LSeg(G*(j,i),G*(k1,i)) by TOPREAL1:6;
    then A19: LSeg(G*(j,i),G*(k1,i)) meets L~Upper_Seq(C,n) by A3,XBOOLE_0:3;
      [k1,i] in Indices G by A2,A5,A14,GOBOARD7:10;
    then consider j1 be Nat such that
     A20: j <= j1 & j1 <= k1 and
     A21: G*(j1,i)`1 = w2 by A10,A12,A16,A17,A19,JORDAN1F:4;
    A22: 1 <= j1 & j1 <= width G by A1,A14,A20,AXIOMS:22;
    then A23: G*(j1,i)`2 = s by A2,A5,GOBOARD5:2;
    then A24: q = G*(j1,i) by A21,EUCLID:57;
    take j1,k1;
    thus j <= j1 & j1 <= k1 & k1 <= k by A12,A20;
    A25: q`1 <= p`1 by A2,A5,A14,A16,A20,A22,A24,SPRECT_3:25;
    set X = LSeg(G*(j,i),G*(k1,i)) /\ L~Upper_Seq(C,n);
    reconsider X1=X as non empty compact Subset of TOP-REAL 2
                                              by A3,A18,PSCOMP_1:64,XBOOLE_0:
def 3;
    consider pp be set such that
     A26: pp in E-most X1 by XBOOLE_0:def 1;
    reconsider pp as Point of TOP-REAL 2 by A26;
      pp in LSeg(SE-corner X, NE-corner X)/\X by A26,PSCOMP_1:def 40;
    then A27: pp in X by XBOOLE_0:def 3;
    then A28: pp in L~Upper_Seq(C,n) by XBOOLE_0:def 3;
    A29: pp in LSeg(G*(j,i),G*(k1,i)) by A27,XBOOLE_0:def 3;
    A30: f`2 = p`2 by A1,A2,A5,A9,A15,A16,GOBOARD5:2;
    then LSeg(f,p) is horizontal by SPPOL_1:36;
    then A31: pp`2 = q`2 by A15,A16,A23,A24,A29,SPRECT_3:19;
      q`1 = E-bound X by A16,A21,A24,SPRECT_1:51
       .= (E-min X)`1 by PSCOMP_1:104
       .= pp`1 by A26,PSCOMP_1:108;
    then A32: q in L~Upper_Seq(C,n) by A28,A31,TOPREAL3:11;
      for x be set holds
     x in LSeg(p,q) /\ L~Upper_Seq(C,n) iff x = q
    proof
     let x be set;
     thus x in LSeg(p,q) /\ L~Upper_Seq(C,n) implies x = q
     proof
      assume A33: x in LSeg(p,q) /\ L~Upper_Seq(C,n);
      then reconsider pp = x as Point of TOP-REAL 2;
      A34: pp in LSeg(p,q) by A33,XBOOLE_0:def 3;
      then A35: pp`2 = q`2 by A15,A16,A23,A24,GOBOARD7:6;
      A36: pp`1 >= q`1 by A25,A34,TOPREAL1:9;
      reconsider EE = LSeg(f,p) /\ L~Upper_Seq(C,n)
                              as compact Subset of TOP-REAL 2 by PSCOMP_1:64;
      reconsider E0 = proj1.:EE as compact Subset of REAL by Th4;
      A37: f`2 = q`2 by A1,A2,A5,A9,A23,A24,GOBOARD5:2;
        f`1 <= q`1 by A1,A2,A5,A20,A22,A24,SPRECT_3:25;
      then A38: q in LSeg(p,f) by A15,A16,A23,A24,A25,A37,GOBOARD7:9;
        p in LSeg(f,p) by TOPREAL1:6;
      then A39: LSeg(p,q) c= LSeg(f,p) by A38,TOPREAL1:12;
        E0 is bounded by RCOMP_1:28;
      then A40: E0 is bounded_above by SEQ_4:def 3;
        pp in L~Upper_Seq(C,n) by A33,XBOOLE_0:def 3;
      then pp in EE by A34,A39,XBOOLE_0:def 3;
      then proj1.pp in E0 by FUNCT_2:43;
      then pp`1 in E0 by PSCOMP_1:def 28;
      then q`1 >= pp`1 by A21,A24,A40,SEQ_4:def 4;
      then pp`1 = q`1 by A36,AXIOMS:21;
      hence x = q by A35,TOPREAL3:11;
     end;
     assume A41: x = q;
     then x in LSeg(p,q) by TOPREAL1:6;
     hence x in LSeg(p,q) /\ L~Upper_Seq(C,n) by A32,A41,XBOOLE_0:def 3;
    end;
    hence LSeg(Gauge(C,n)*(j1,i),Gauge(C,n)*(k1,i)) /\ L~Upper_Seq(C,n) =
     {Gauge(C,n)*(j1,i)} by A16,A24,TARSKI:def 1;
    set X = LSeg(G*(j,i),G*(k,i)) /\ L~Lower_Seq(C,n);
    reconsider X1=X as non empty compact Subset of TOP-REAL 2
                                              by A4,A7,PSCOMP_1:64,XBOOLE_0:def
3;
    consider pp be set such that
     A42: pp in W-most X1 by XBOOLE_0:def 1;
    reconsider pp as Point of TOP-REAL 2 by A42;
      pp in LSeg(SW-corner X, NW-corner X)/\X by A42,PSCOMP_1:def 38;
    then A43: pp in X by XBOOLE_0:def 3;
    then A44: pp in L~Lower_Seq(C,n) by XBOOLE_0:def 3;
    A45: pp in LSeg(G*(j,i),G*(k,i)) by A43,XBOOLE_0:def 3;
      f`2 = s by A1,A2,A5,A9,GOBOARD5:2
       .= e`2 by A1,A2,A11,GOBOARD5:2;
    then LSeg(f,e) is horizontal by SPPOL_1:36;
    then A46: pp`2 = p`2 by A30,A45,SPRECT_3:19;
      p`1 = W-bound X by A13,A16,SPRECT_1:48
       .= (W-min X)`1 by PSCOMP_1:84
       .= pp`1 by A42,PSCOMP_1:88;
    then A47: p in L~Lower_Seq(C,n) by A44,A46,TOPREAL3:11;
      for x be set holds
     x in LSeg(p,q) /\ L~Lower_Seq(C,n) iff x = p
    proof
     let x be set;
     thus x in LSeg(p,q) /\ L~Lower_Seq(C,n) implies x = p
     proof
      assume A48: x in LSeg(p,q) /\ L~Lower_Seq(C,n);
      then reconsider pp = x as Point of TOP-REAL 2;
      A49: pp in LSeg(p,q) by A48,XBOOLE_0:def 3;
      then A50: pp`2 = p`2 by A15,A16,A23,A24,GOBOARD7:6;
      A51: pp`1 <= p`1 by A25,A49,TOPREAL1:9;
      reconsider EE = LSeg(f,e) /\ L~Lower_Seq(C,n)
                              as compact Subset of TOP-REAL 2 by PSCOMP_1:64;
      reconsider E0 = proj1.:EE as compact Subset of REAL by Th4;
      A52: f`2 = q`2 by A1,A2,A5,A9,A23,A24,GOBOARD5:2;
      A53: e`2 = q`2 by A1,A2,A11,A23,A24,GOBOARD5:2;
      A54: f`1 <= q`1 by A1,A2,A5,A20,A22,A24,SPRECT_3:25;
        j1 <= k by A12,A20,AXIOMS:22;
      then q`1 <= e`1 by A1,A2,A22,A24,SPRECT_3:25;
      then A55: q in LSeg(f,e) by A52,A53,A54,GOBOARD7:9;
      A56: f`2 = p`2 by A1,A2,A5,A9,A15,A16,GOBOARD5:2;
      A57: e`2 = p`2 by A1,A2,A11,A15,A16,GOBOARD5:2;
      A58: f`1 <= p`1 by A1,A2,A5,A12,A14,A16,SPRECT_3:25;
        p`1 <= e`1 by A1,A2,A12,A14,A16,SPRECT_3:25;
      then p in LSeg(f,e) by A56,A57,A58,GOBOARD7:9;
      then A59: LSeg(p,q) c= LSeg(f,e) by A55,TOPREAL1:12;
        E0 is bounded by RCOMP_1:28;
      then A60: E0 is bounded_below by SEQ_4:def 3;
        pp in L~Lower_Seq(C,n) by A48,XBOOLE_0:def 3;
      then pp in EE by A49,A59,XBOOLE_0:def 3;
      then proj1.pp in E0 by FUNCT_2:43;
      then pp`1 in E0 by PSCOMP_1:def 28;
      then p`1 <= pp`1 by A13,A16,A60,SEQ_4:def 5;
      then pp`1 = p`1 by A51,AXIOMS:21;
      hence x = p by A50,TOPREAL3:11;
     end;
     assume A61: x = p;
     then x in LSeg(p,q) by TOPREAL1:6;
     hence x in LSeg(p,q) /\ L~Lower_Seq(C,n) by A47,A61,XBOOLE_0:def 3;
    end;
    hence LSeg(Gauge(C,n)*(j1,i),Gauge(C,n)*(k1,i)) /\ L~Lower_Seq(C,n) =
      {Gauge(C,n)*(k1,i)} by A16,A24,TARSKI:def 1;
   end;

  theorem Th23:
   for C be Simple_closed_curve
   for i,j,k be Nat st
    1 < i & i < len Gauge(C,n) &
    1 <= j & j <= k & k <= width Gauge(C,n) &
    Gauge(C,n)*(i,k) in L~Upper_Seq(C,n) &
    Gauge(C,n)*(i,j) in L~Lower_Seq(C,n) holds
     LSeg(Gauge(C,n)*(i,j),Gauge(C,n)*(i,k)) meets Lower_Arc C
   proof
    let C be Simple_closed_curve;
    let i,j,k be Nat;
    assume that
     A1: 1 < i & i < len Gauge(C,n) and
     A2: 1 <= j & j <= k & k <= width Gauge(C,n) and
     A3: Gauge(C,n)*(i,k) in L~Upper_Seq(C,n) and
     A4: Gauge(C,n)*(i,j) in L~Lower_Seq(C,n);
    consider j1,k1 be Nat such that
     A5: j <= j1 and
     A6: j1 <= k1 and
     A7: k1 <= k and
     A8: LSeg(Gauge(C,n)*(i,j1),Gauge(C,n)*(i,k1)) /\ L~Lower_Seq(C,n) =
      {Gauge(C,n)*(i,j1)} and
     A9: LSeg(Gauge(C,n)*(i,j1),Gauge(C,n)*(i,k1)) /\ L~Upper_Seq(C,n) =
      {Gauge(C,n)*(i,k1)} by A1,A2,A3,A4,Th13;
      1 <= j1 & k1 <= width Gauge(C,n) by A2,A5,A7,AXIOMS:22;
    then A10: LSeg(Gauge(C,n)*(i,j1),Gauge(C,n)*(i,k1)) meets Lower_Arc C
                                               by A1,A6,A8,A9,JORDAN1J:58;
      LSeg(Gauge(C,n)*(i,j1),Gauge(C,n)*(i,k1)) c=
     LSeg(Gauge(C,n)*(i,j),Gauge(C,n)*(i,k)) by A1,A2,A5,A6,A7,Th7;
    hence LSeg(Gauge(C,n)*(i,j),Gauge(C,n)*(i,k)) meets Lower_Arc C
                                                          by A10,XBOOLE_1:63;
   end;

  theorem Th24:
   for C be Simple_closed_curve
   for i,j,k be Nat st
    1 < i & i < len Gauge(C,n) &
    1 <= j & j <= k & k <= width Gauge(C,n) &
    Gauge(C,n)*(i,k) in L~Upper_Seq(C,n) &
    Gauge(C,n)*(i,j) in L~Lower_Seq(C,n) holds
     LSeg(Gauge(C,n)*(i,j),Gauge(C,n)*(i,k)) meets Upper_Arc C
   proof
    let C be Simple_closed_curve;
    let i,j,k be Nat;
    assume that
     A1: 1 < i & i < len Gauge(C,n) and
     A2: 1 <= j & j <= k & k <= width Gauge(C,n) and
     A3: Gauge(C,n)*(i,k) in L~Upper_Seq(C,n) and
     A4: Gauge(C,n)*(i,j) in L~Lower_Seq(C,n);
    consider j1,k1 be Nat such that
     A5: j <= j1 and
     A6: j1 <= k1 and
     A7: k1 <= k and
     A8: LSeg(Gauge(C,n)*(i,j1),Gauge(C,n)*(i,k1)) /\ L~Lower_Seq(C,n) =
      {Gauge(C,n)*(i,j1)} and
     A9: LSeg(Gauge(C,n)*(i,j1),Gauge(C,n)*(i,k1)) /\ L~Upper_Seq(C,n) =
      {Gauge(C,n)*(i,k1)} by A1,A2,A3,A4,Th13;
      1 <= j1 & k1 <= width Gauge(C,n) by A2,A5,A7,AXIOMS:22;
    then A10: LSeg(Gauge(C,n)*(i,j1),Gauge(C,n)*(i,k1)) meets Upper_Arc C
                                               by A1,A6,A8,A9,JORDAN1J:59;
      LSeg(Gauge(C,n)*(i,j1),Gauge(C,n)*(i,k1)) c=
     LSeg(Gauge(C,n)*(i,j),Gauge(C,n)*(i,k)) by A1,A2,A5,A6,A7,Th7;
    hence LSeg(Gauge(C,n)*(i,j),Gauge(C,n)*(i,k)) meets Upper_Arc C
                                                          by A10,XBOOLE_1:63;
   end;

  theorem Th25:
   for C be Simple_closed_curve
   for i,j,k be Nat st
    1 < i & i < len Gauge(C,n) &
    1 <= j & j <= k & k <= width Gauge(C,n) & n > 0 &
    Gauge(C,n)*(i,k) in Upper_Arc L~Cage(C,n) &
    Gauge(C,n)*(i,j) in Lower_Arc L~Cage(C,n) holds
     LSeg(Gauge(C,n)*(i,j),Gauge(C,n)*(i,k)) meets Lower_Arc C
   proof
    let C be Simple_closed_curve;
    let i,j,k be Nat;
    assume that
     A1: 1 < i & i < len Gauge(C,n) and
     A2: 1 <= j & j <= k & k <= width Gauge(C,n) and
     A3: n > 0 and
     A4: Gauge(C,n)*(i,k) in Upper_Arc L~Cage(C,n) and
     A5: Gauge(C,n)*(i,j) in Lower_Arc L~Cage(C,n);
    A6: L~Upper_Seq(C,n) = Upper_Arc L~Cage(C,n) by A3,JORDAN1G:63;
      L~Lower_Seq(C,n) = Lower_Arc L~Cage(C,n) by A3,JORDAN1G:64;
    hence LSeg(Gauge(C,n)*(i,j),Gauge(C,n)*(i,k)) meets Lower_Arc C
                                                    by A1,A2,A4,A5,A6,Th23;
   end;

  theorem Th26:
   for C be Simple_closed_curve
   for i,j,k be Nat st
    1 < i & i < len Gauge(C,n) &
    1 <= j & j <= k & k <= width Gauge(C,n) & n > 0 &
    Gauge(C,n)*(i,k) in Upper_Arc L~Cage(C,n) &
    Gauge(C,n)*(i,j) in Lower_Arc L~Cage(C,n) holds
     LSeg(Gauge(C,n)*(i,j),Gauge(C,n)*(i,k)) meets Upper_Arc C
   proof
    let C be Simple_closed_curve;
    let i,j,k be Nat;
    assume that
     A1: 1 < i & i < len Gauge(C,n) and
     A2: 1 <= j & j <= k & k <= width Gauge(C,n) and
     A3: n > 0 and
     A4: Gauge(C,n)*(i,k) in Upper_Arc L~Cage(C,n) and
     A5: Gauge(C,n)*(i,j) in Lower_Arc L~Cage(C,n);
    A6: L~Upper_Seq(C,n) = Upper_Arc L~Cage(C,n) by A3,JORDAN1G:63;
      L~Lower_Seq(C,n) = Lower_Arc L~Cage(C,n) by A3,JORDAN1G:64;
    hence LSeg(Gauge(C,n)*(i,j),Gauge(C,n)*(i,k)) meets Upper_Arc C
                                                    by A1,A2,A4,A5,A6,Th24;
   end;

  theorem
     for C be Simple_closed_curve
   for j,k be Nat holds
    1 <= j & j <= k & k <= width Gauge(C,n+1) &
    Gauge(C,n+1)*(Center Gauge(C,n+1),k) in Upper_Arc L~Cage(C,n+1) &
    Gauge(C,n+1)*(Center Gauge(C,n+1),j) in Lower_Arc L~Cage(C,n+1) implies
     LSeg(Gauge(C,n+1)*(Center Gauge(C,n+1),j),
          Gauge(C,n+1)*(Center Gauge(C,n+1),k)) meets Lower_Arc C
   proof
    let C be Simple_closed_curve;
    let j,k be Nat;
    assume that
     A1: 1 <= j & j <= k & k <= width Gauge(C,n+1) and
     A2: Gauge(C,n+1)*(Center Gauge(C,n+1),k) in Upper_Arc L~Cage(C,n+1) and
     A3: Gauge(C,n+1)*(Center Gauge(C,n+1),j) in Lower_Arc L~Cage(C,n+1);
      n+1 >= 0+1 by NAT_1:29;
    then A4: n+1 > 0 by NAT_1:38;
    A5: len Gauge(C,n+1) >= 4 by JORDAN8:13;
    then A6: len Gauge(C,n+1) >= 3 by AXIOMS:22;
      len Gauge(C,n+1) >= 2 by A5,AXIOMS:22;
    then A7: 1 < Center Gauge(C,n+1) by JORDAN1B:15;
      Center Gauge(C,n+1) < len Gauge(C,n+1) by A6,JORDAN1B:16;
    hence LSeg(Gauge(C,n+1)*(Center Gauge(C,n+1),j),
          Gauge(C,n+1)*(Center Gauge(C,n+1),k)) meets Lower_Arc C
                                                    by A1,A2,A3,A4,A7,Th25;
   end;

  theorem
     for C be Simple_closed_curve
   for j,k be Nat holds
    1 <= j & j <= k & k <= width Gauge(C,n+1) &
    Gauge(C,n+1)*(Center Gauge(C,n+1),k) in Upper_Arc L~Cage(C,n+1) &
    Gauge(C,n+1)*(Center Gauge(C,n+1),j) in Lower_Arc L~Cage(C,n+1) implies
     LSeg(Gauge(C,n+1)*(Center Gauge(C,n+1),j),
          Gauge(C,n+1)*(Center Gauge(C,n+1),k)) meets Upper_Arc C
   proof
    let C be Simple_closed_curve;
    let j,k be Nat;
    assume that
     A1: 1 <= j & j <= k & k <= width Gauge(C,n+1) and
     A2: Gauge(C,n+1)*(Center Gauge(C,n+1),k) in Upper_Arc L~Cage(C,n+1) and
     A3: Gauge(C,n+1)*(Center Gauge(C,n+1),j) in Lower_Arc L~Cage(C,n+1);
      n+1 >= 0+1 by NAT_1:29;
    then A4: n+1 > 0 by NAT_1:38;
    A5: len Gauge(C,n+1) >= 4 by JORDAN8:13;
    then A6: len Gauge(C,n+1) >= 3 by AXIOMS:22;
      len Gauge(C,n+1) >= 2 by A5,AXIOMS:22;
    then A7: 1 < Center Gauge(C,n+1) by JORDAN1B:15;
      Center Gauge(C,n+1) < len Gauge(C,n+1) by A6,JORDAN1B:16;
    hence LSeg(Gauge(C,n+1)*(Center Gauge(C,n+1),j),
          Gauge(C,n+1)*(Center Gauge(C,n+1),k)) meets Upper_Arc C
                                                   by A1,A2,A3,A4,A7,Th26;
   end;

  theorem Th29:
   for C be compact connected non vertical non horizontal Subset of TOP-REAL 2
   for i,j,k be Nat st
    1 < j & k < len Gauge(C,n) &
    1 <= i & i <= width Gauge(C,n) &
    Gauge(C,n)*(k,i) in L~Upper_Seq(C,n) &
    Gauge(C,n)*(j,i) in L~Lower_Seq(C,n) holds j <> k
   proof
    let C be compact connected non vertical non horizontal
     Subset of TOP-REAL 2;
    let i,j,k be Nat;
    assume that
     A1: 1 < j & k < len Gauge(C,n) and
     A2: 1 <= i & i <= width Gauge(C,n) and
     A3: Gauge(C,n)*(k,i) in L~Upper_Seq(C,n) and
     A4: Gauge(C,n)*(j,i) in L~Lower_Seq(C,n) and
     A5: j = k;
    A6: len Gauge(C,n) = width Gauge(C,n) by JORDAN8:def 1;
      Gauge(C,n)*(k,i) in L~Upper_Seq(C,n) /\ L~Lower_Seq(C,n)
                                                 by A3,A4,A5,XBOOLE_0:def 3;
    then A7: Gauge(C,n)*(k,i) in {W-min L~Cage(C,n),E-max L~Cage(C,n)}
                                                              by JORDAN1E:20;
    A8: [j,i] in Indices Gauge(C,n) by A1,A2,A5,GOBOARD7:10;
      len Gauge(C,n) >= 4 by JORDAN8:13;
    then A9: len Gauge(C,n) >= 1 by AXIOMS:22;
    then A10: [len Gauge(C,n),i] in Indices Gauge(C,n) by A2,GOBOARD7:10;
    A11: [1,i] in Indices Gauge(C,n) by A2,A9,GOBOARD7:10;
    per cases by A7,TARSKI:def 2;
     suppose A12: Gauge(C,n)*(k,i) = W-min L~Cage(C,n);
        Gauge(C,n)*(1,i)`1 = W-bound L~Cage(C,n) by A2,A6,JORDAN1A:94;
      then (W-min L~Cage(C,n))`1 <> W-bound L~Cage(C,n)
                                            by A1,A5,A8,A11,A12,JORDAN1G:7;
      hence contradiction by PSCOMP_1:84;
     suppose A13: Gauge(C,n)*(k,i) = E-max L~Cage(C,n);
        Gauge(C,n)*(len Gauge(C,n),i)`1 = E-bound L~Cage(C,n)
                                                        by A2,A6,JORDAN1A:92;
      then (E-max L~Cage(C,n))`1 <> E-bound L~Cage(C,n)
                                                by A1,A5,A8,A10,A13,JORDAN1G:7;
      hence contradiction by PSCOMP_1:104;
   end;

  theorem Th30:
   for C be Simple_closed_curve
   for i,j,k be Nat st
    1 < j & j <= k & k < len Gauge(C,n) &
    1 <= i & i <= width Gauge(C,n) &
    LSeg(Gauge(C,n)*(j,i),Gauge(C,n)*(k,i)) /\ L~Upper_Seq(C,n) =
     {Gauge(C,n)*(k,i)} &
    LSeg(Gauge(C,n)*(j,i),Gauge(C,n)*(k,i)) /\ L~Lower_Seq(C,n) =
     {Gauge(C,n)*(j,i)} holds
      LSeg(Gauge(C,n)*(j,i),Gauge(C,n)*(k,i)) meets Lower_Arc C
   proof
    let C be Simple_closed_curve;
    let i,j,k be Nat;
    set Ga = Gauge(C,n);
    set US = Upper_Seq(C,n);
    set LS = Lower_Seq(C,n);
    set LA = Lower_Arc C;
    set Wmin = W-min L~Cage(C,n);
    set Emax = E-max L~Cage(C,n);
    set Wbo = W-bound L~Cage(C,n);
    set Ebo = E-bound L~Cage(C,n);
    set Gij = Ga*(j,i);
    set Gik = Ga*(k,i);
    assume that
     A1: 1 < j & j <= k & k < len Ga and
     A2: 1 <= i & i <= width Ga and
     A3: LSeg(Gij,Gik) /\ L~US = {Gik} and
     A4: LSeg(Gij,Gik) /\ L~LS = {Gij} and
     A5: LSeg(Gij,Gik) misses LA;
      Gij in {Gij} by TARSKI:def 1;
    then A6: Gij in L~LS by A4,XBOOLE_0:def 3;
      Gik in {Gik} by TARSKI:def 1;
    then A7: Gik in L~US by A3,XBOOLE_0:def 3;
    A8: len Ga = width Ga by JORDAN8:def 1;
    A9: j <> k by A1,A2,A6,A7,Th29;
    A10: 1 <= j & j <= width Ga by A1,A8,AXIOMS:22;
    A11: 1 <= k & k <= width Ga by A1,AXIOMS:22,JORDAN8:def 1;
    A12: [j,i] in Indices Ga by A2,A8,A10,GOBOARD7:10;
    A13: [k,i] in Indices Ga by A2,A8,A11,GOBOARD7:10;
    A14: US is_sequence_on Ga by JORDAN1G:4;
    A15: LS is_sequence_on Ga by JORDAN1G:5;
    set go = R_Cut(US,Gik);
    set do = L_Cut(LS,Gij);
    A16: len US >= 3 by JORDAN1E:19;
    then len US >= 1 by AXIOMS:22;
    then 1 in dom US by FINSEQ_3:27;
    then A17: US.1 = US/.1 by FINSEQ_4:def 4
       .= Wmin by JORDAN1F:5;
    A18: Wmin`1 = Wbo by PSCOMP_1:84
       .= Ga*(1,k)`1 by A8,A11,JORDAN1A:94;
      len Ga >= 4 by JORDAN8:13;
    then A19: len Ga >= 1 by AXIOMS:22;
    then A20: [1,k] in Indices Ga by A11,GOBOARD7:10;
    then A21: Gik <> US.1 by A1,A13,A17,A18,JORDAN1G:7;
    then reconsider go as being_S-Seq FinSequence of TOP-REAL 2
                                                            by A7,JORDAN3:70;
    A22: [1,j] in Indices Ga by A10,A19,GOBOARD7:10;
    A23: len LS >= 1+2 by JORDAN1E:19;
    then len LS >= 1 by AXIOMS:22;
    then A24: 1 in dom LS & len LS in dom LS by FINSEQ_3:27;
    then A25: LS.len LS = LS/.len LS by FINSEQ_4:def 4
       .= Wmin by JORDAN1F:8;
    A26: Wmin`1 = Wbo by PSCOMP_1:84
       .= Ga*(1,k)`1 by A8,A11,JORDAN1A:94;
    A27: [j,i] in Indices Ga by A2,A8,A10,GOBOARD7:10;
    then A28: Gij <> LS.len LS by A1,A20,A25,A26,JORDAN1G:7;
    then reconsider do as being_S-Seq FinSequence of TOP-REAL 2
                                                            by A6,JORDAN3:69;
    A29: [len Ga,k] in Indices Ga by A11,A19,GOBOARD7:10;
    A30: LS.1 = LS/.1 by A24,FINSEQ_4:def 4
       .= Emax by JORDAN1F:6;
      Emax`1 = Ebo by PSCOMP_1:104
       .= Ga*(len Ga,k)`1 by A8,A11,JORDAN1A:92;
    then A31: Gij <> LS.1 by A1,A27,A29,A30,JORDAN1G:7;
    A32: len go >= 1+1 by TOPREAL1:def 10;
    A33: Gik in rng US by A2,A7,A8,A11,A14,JORDAN1J:40;
    then A34: go is_sequence_on Ga by A14,JORDAN1J:38;
    A35: len do >= 1+1 by TOPREAL1:def 10;
    A36: Gij in rng LS by A2,A6,A8,A10,A15,JORDAN1J:40;
    then A37: do is_sequence_on Ga by A15,JORDAN1J:39;
    reconsider go as non constant s.c.c.
            (being_S-Seq FinSequence of TOP-REAL 2) by A32,A34,JGRAPH_1:16,
JORDAN8:8;
    reconsider do as non constant s.c.c.
            (being_S-Seq FinSequence of TOP-REAL 2) by A35,A37,JGRAPH_1:16,
JORDAN8:8;
    A38: len go > 1 by A32,NAT_1:38;
    then A39: len go in dom go by FINSEQ_3:27;
    then A40: go/.len go = go.len go by FINSEQ_4:def 4
       .= Gik by A7,JORDAN3:59;
      len do >= 1 by A35,AXIOMS:22;
    then 1 in dom do by FINSEQ_3:27;
    then A41: do/.1 = do.1 by FINSEQ_4:def 4
       .= Gij by A6,JORDAN3:58;
    reconsider m = len go - 1 as Nat by A39,FINSEQ_3:28;
    A42: m+1 = len go by XCMPLX_1:27;
    then A43: len go-'1 = m by BINARITH:39;
    A44: LSeg(go,m) c= L~go by TOPREAL3:26;
    A45: L~go c= L~US by A7,JORDAN3:76;
    then LSeg(go,m) c= L~US by A44,XBOOLE_1:1;
    then A46: LSeg(go,m) /\ LSeg(Gik,Gij) c= {Gik} by A3,XBOOLE_1:26;
      m >= 1 by A32,REAL_1:84;
    then A47: LSeg(go,m) = LSeg(go/.m,Gik) by A40,A42,TOPREAL1:def 5;
      {Gik} c= LSeg(go,m) /\ LSeg(Gik,Gij)
    proof
     let x be set;
     assume x in {Gik};
     then A48: x = Gik by TARSKI:def 1;
     A49: Gik in LSeg(go,m) by A47,TOPREAL1:6;
       Gik in LSeg(Gik,Gij) by TOPREAL1:6;
     hence x in LSeg(go,m) /\ LSeg(Gik,Gij) by A48,A49,XBOOLE_0:def 3;
    end;
    then A50: LSeg(go,m) /\ LSeg(Gik,Gij) = {Gik} by A46,XBOOLE_0:def 10;
    A51: LSeg(do,1) c= L~do by TOPREAL3:26;
    A52: L~do c= L~LS by A6,JORDAN3:77;
    then LSeg(do,1) c= L~LS by A51,XBOOLE_1:1;
    then A53: LSeg(do,1) /\ LSeg(Gik,Gij) c= {Gij} by A4,XBOOLE_1:26;
    A54: LSeg(do,1) = LSeg(Gij,do/.(1+1)) by A35,A41,TOPREAL1:def 5;
      {Gij} c= LSeg(do,1) /\ LSeg(Gik,Gij)
    proof
     let x be set;
     assume x in {Gij};
     then A55: x = Gij by TARSKI:def 1;
     A56: Gij in LSeg(do,1) by A54,TOPREAL1:6;
       Gij in LSeg(Gik,Gij) by TOPREAL1:6;
     hence x in LSeg(do,1) /\ LSeg(Gik,Gij) by A55,A56,XBOOLE_0:def 3;
    end;
    then A57: LSeg(Gik,Gij) /\ LSeg(do,1) = {Gij} by A53,XBOOLE_0:def 10;
    A58: go/.1 = US/.1 by A7,SPRECT_3:39
       .= Wmin by JORDAN1F:5;
    then A59: go/.1 = LS/.len LS by JORDAN1F:8
       .= do/.len do by A6,JORDAN1J:35;
    A60: rng go c= L~go & rng do c= L~do by A32,A35,SPPOL_2:18;
    A61: {go/.1} c= L~go /\ L~do
    proof
     let x be set;
     assume x in {go/.1};
     then x = go/.1 by TARSKI:def 1;
     then x in rng go & x in rng do by A59,FINSEQ_6:46,REVROT_1:3;
     hence x in L~go /\ L~do by A60,XBOOLE_0:def 3;
    end;
    A62: LS.1 = LS/.1 by A24,FINSEQ_4:def 4
       .= Emax by JORDAN1F:6;
    A63: [len Ga,j] in Indices Ga by A10,A19,GOBOARD7:10;
      L~go /\ L~do c= {go/.1}
    proof
     let x be set;
     assume x in L~go /\ L~do;
     then A64: x in L~go & x in L~do by XBOOLE_0:def 3;
     then x in L~US /\ L~LS by A45,A52,XBOOLE_0:def 3;
     then x in {Wmin,Emax} by JORDAN1E:20;
     then A65: x = Wmin or x = Emax by TARSKI:def 2;
       now assume x = Emax;
      then A66: Emax = Gij by A6,A62,A64,JORDAN1E:11;
        Ga*(len Ga,j)`1 = Ebo by A8,A10,JORDAN1A:92;
      then Emax`1 <> Ebo by A1,A12,A63,A66,JORDAN1G:7;
      hence contradiction by PSCOMP_1:104;
     end;
     hence x in {go/.1} by A58,A65,TARSKI:def 1;
    end;
    then A67: L~go /\ L~do = {go/.1} by A61,XBOOLE_0:def 10;
    set W2 = go/.2;
    A68: 2 in dom go by A32,FINSEQ_3:27;
    A69: Gik..US >= 1 by A33,FINSEQ_4:31;
    A70: now assume Gij`1 = Wbo;
     then Ga*(1,j)`1 = Ga*(j,i)`1 by A8,A10,JORDAN1A:94;
     hence contradiction by A1,A12,A22,JORDAN1G:7;
    end;
      go = mid(US,1,Gik..US) by A33,JORDAN1G:57
       .= US|(Gik..US) by A69,JORDAN3:25;
    then A71: W2 = US/.2 by A68,TOPREAL1:1;
    A72: Wmin in rng go by A58,FINSEQ_6:46;
    set pion = <*Gik,Gij*>;
    A73: now let n be Nat;
     assume n in dom pion;
     then n in {1,2} by FINSEQ_1:4,FINSEQ_3:29;
     then n = 1 or n = 2 by TARSKI:def 2;
     then pion/.n = Gik or pion/.n = Gij by FINSEQ_4:26;
     hence ex j,i be Nat st [j,i] in Indices Ga & pion/.n = Ga*(j,i)
                                                                  by A12,A13;
    end;
    A74: Gik <> Gij by A9,A12,A13,GOBOARD1:21;
      Gik`2 = Ga*(1,i)`2 by A2,A8,A11,GOBOARD5:2
       .= Gij`2 by A2,A8,A10,GOBOARD5:2;
    then LSeg(Gik,Gij) is horizontal by SPPOL_1:36;
    then pion is_S-Seq by A74,JORDAN1B:9;
    then consider pion1 be FinSequence of TOP-REAL 2 such that
     A75: pion1 is_sequence_on Ga and
     A76: pion1 is_S-Seq and
     A77: L~pion = L~pion1 and
     A78: pion/.1 = pion1/.1 and
     A79: pion/.len pion = pion1/.len pion1 and
     A80: len pion <= len pion1 by A73,GOBOARD3:2;
    reconsider pion1 as being_S-Seq FinSequence of TOP-REAL 2 by A76;
    set godo = go^'pion1^'do;
      len Cage(C,n) > 4 by GOBOARD7:36;
    then A81: 1+1 <= len Cage(C,n) by AXIOMS:22;
    then A82: 1+1 <= len Rotate(Cage(C,n),Wmin) by REVROT_1:14;
      len (go^'pion1) >= len go by TOPREAL8:7;
    then A83: len (go^'pion1) >= 1+1 by A32,AXIOMS:22;
    then A84: len (go^'pion1) > 1+0 by NAT_1:38;
      len godo >= len (go^'pion1) by TOPREAL8:7;
    then A85: 1+1 <= len godo by A83,AXIOMS:22;
    A86: US is_sequence_on Ga by JORDAN1G:4;
    A87: go/.len go = pion1/.1 by A40,A78,FINSEQ_4:26;
    then A88: go^'pion1 is_sequence_on Ga by A34,A75,TOPREAL8:12;
    A89: (go^'pion1)/.len (go^'pion1) = pion/.len pion by A79,AMISTD_1:6
       .= pion/.2 by FINSEQ_1:61
       .= do/.1 by A41,FINSEQ_4:26;
    then A90: godo is_sequence_on Ga by A37,A88,TOPREAL8:12;
      LSeg(pion1,1) c= L~<*Gik,Gij*> by A77,TOPREAL3:26;
    then LSeg(pion1,1) c= LSeg(Gik,Gij) by SPPOL_2:21;
    then A91: LSeg(go,len go-'1) /\ LSeg(pion1,1) c= {Gik}
                                                      by A43,A50,XBOOLE_1:27;
    A92: len pion1 >= 1+1 by A80,FINSEQ_1:61;
      {Gik} c= LSeg(go,m) /\ LSeg(pion1,1)
    proof
     let x be set;
     assume x in {Gik};
     then A93: x = Gik by TARSKI:def 1;
     A94: Gik in LSeg(go,m) by A47,TOPREAL1:6;
       Gik in LSeg(pion1,1) by A40,A87,A92,TOPREAL1:27;
     hence x in LSeg(go,m) /\ LSeg(pion1,1) by A93,A94,XBOOLE_0:def 3;
    end;
    then LSeg(go,len go-'1) /\ LSeg(pion1,1) = { go/.len go }
                                             by A40,A43,A91,XBOOLE_0:def 10;
    then A95: go^'pion1 is unfolded by A87,TOPREAL8:34;
      len pion1 >= 2+0 by A80,FINSEQ_1:61;
    then A96: len pion1-2 >= 0 by REAL_1:84;
    A97: len (go^'pion1)-1 >= 0 by A84,REAL_1:84;
      len (go^'pion1)+1-1 = len go+len pion1-1 by GRAPH_2:13;
    then len (go^'pion1)-1 = len go+len pion1-1-1 by XCMPLX_1:26
       .= len go + len pion1-(1+1) by XCMPLX_1:36
       .= len go + (len pion1-2) by XCMPLX_1:29
       .= len go + (len pion1-'2) by A96,BINARITH:def 3;
    then A98: len (go^'pion1)-'1 = len go + (len pion1-'2)
                                                      by A97,BINARITH:def 3;
    A99: len pion1-1 >= 1 by A92,REAL_1:84;
    then A100: len pion1-1 >= 0 by AXIOMS:22;
    then A101: len pion1-'1 = len pion1-1 by BINARITH:def 3;
    A102: len pion1-'2+1 = len pion1-2+1 by A96,BINARITH:def 3
       .= len pion1-(2-1) by XCMPLX_1:37
       .= len pion1-'1 by A100,BINARITH:def 3;
      len pion1-1+1 <= len pion1 by XCMPLX_1:27;
    then A103: len pion1-'1 < len pion1 by A101,NAT_1:38;
      LSeg(pion1,len pion1-'1) c= L~<*Gik,Gij*> by A77,TOPREAL3:26;
    then LSeg(pion1,len pion1-'1) c= LSeg(Gik,Gij) by SPPOL_2:21;
    then A104: LSeg(pion1,len pion1-'1) /\ LSeg(do,1) c= {Gij}
                                                          by A57,XBOOLE_1:27;
      {Gij} c= LSeg(pion1,len pion1-'1) /\ LSeg(do,1)
    proof
     let x be set;
     assume x in {Gij};
     then A105: x = Gij by TARSKI:def 1;
     A106: Gij in LSeg(do,1) by A54,TOPREAL1:6;
     A107: len pion1-'1+1 = len pion1 by A101,XCMPLX_1:27;
     then pion1/.(len pion1-'1+1) = pion/.2 by A79,FINSEQ_1:61
        .= Gij by FINSEQ_4:26;
     then Gij in LSeg(pion1,len pion1-'1) by A99,A101,A107,TOPREAL1:27;
     hence x in LSeg(pion1,len pion1-'1) /\ LSeg(do,1)
                                                  by A105,A106,XBOOLE_0:def 3;
    end;
    then LSeg(pion1,len pion1-'1) /\ LSeg(do,1) = {Gij} by A104,XBOOLE_0:def 10
;
    then A108: LSeg(go^'pion1,len go+(len pion1-'2)) /\ LSeg(do,1) =
       {(go^'pion1)/.len (go^'pion1)} by A41,A87,A89,A102,A103,TOPREAL8:31;
    A109: (go^'pion1) is non trivial by A83,SPPOL_1:2;
    A110: rng pion1 c= L~pion1 by A92,SPPOL_2:18;
    A111: {pion1/.1} c= L~go /\ L~pion1
    proof
     let x be set;
     assume x in {pion1/.1};
     then x = pion1/.1 by TARSKI:def 1;
     then x in rng go & x in rng pion1 by A87,FINSEQ_6:46,REVROT_1:3;
     hence x in L~go /\ L~pion1 by A60,A110,XBOOLE_0:def 3;
    end;
      L~go /\ L~pion1 c= {pion1/.1}
    proof
     let x be set;
     assume x in L~go /\ L~pion1;
     then x in L~go & x in L~pion1 by XBOOLE_0:def 3;
     then x in L~pion1 /\ L~US by A45,XBOOLE_0:def 3;
     hence x in {pion1/.1} by A3,A40,A77,A87,SPPOL_2:21;
    end;
    then A112: L~go /\ L~pion1 = {pion1/.1} by A111,XBOOLE_0:def 10;
    then A113: (go^'pion1) is s.n.c. by A87,JORDAN1J:54;
      rng go /\ rng pion1 c= {pion1/.1} by A60,A110,A112,XBOOLE_1:27;
    then A114: go^'pion1 is one-to-one by JORDAN1J:55;
    A115: pion/.len pion = pion/.2 by FINSEQ_1:61
       .= do/.1 by A41,FINSEQ_4:26;
    A116: {pion1/.len pion1} c= L~do /\ L~pion1
    proof
     let x be set;
     assume x in {pion1/.len pion1};
     then x = pion1/.len pion1 by TARSKI:def 1;
     then x in rng do & x in rng pion1 by A79,A115,FINSEQ_6:46,REVROT_1:3;
     hence x in L~do /\ L~pion1 by A60,A110,XBOOLE_0:def 3;
    end;
      L~do /\ L~pion1 c= {pion1/.len pion1}
    proof
     let x be set;
     assume x in L~do /\ L~pion1;
     then x in L~do & x in L~pion1 by XBOOLE_0:def 3;
     then x in L~pion1 /\ L~LS by A52,XBOOLE_0:def 3;
     hence x in {pion1/.len pion1} by A4,A41,A77,A79,A115,SPPOL_2:21;
    end;
    then A117: L~do /\ L~pion1 = {pion1/.len pion1} by A116,XBOOLE_0:def 10;
    A118: L~(go^'pion1) /\ L~do = (L~go \/ L~pion1) /\ L~do by A87,TOPREAL8:35
       .= {go/.1} \/ {do/.1} by A67,A79,A115,A117,XBOOLE_1:23
       .= {(go^'pion1)/.1} \/ {do/.1} by AMISTD_1:5
       .= {(go^'pion1)/.1,do/.1} by ENUMSET1:41;
      do/.len do = (go^'pion1)/.1 by A59,AMISTD_1:5;
    then reconsider godo as non constant standard special_circular_sequence
       by A85,A89,A90,A95,A98,A108,A109,A113,A114,A118,JORDAN8:7,8,TOPREAL8:11,
33,34;
    A119: LA is_an_arc_of E-max C,W-min C by JORDAN6:def 9;
    then A120: LA is connected by JORDAN6:11;
    A121: W-min C in LA & E-max C in LA by A119,TOPREAL1:4;
    set ff = Rotate(Cage(C,n),Wmin);
      Wmin in rng Cage(C,n) by SPRECT_2:47;
    then A122: ff/.1 = Wmin by FINSEQ_6:98;
    A123: L~ff = L~Cage(C,n) by REVROT_1:33;
    then A124: (W-max L~ff)..ff > 1 by A122,SPRECT_5:23;
      (W-max L~ff)..ff <= (N-min L~ff)..ff by A122,A123,SPRECT_5:24;
    then A125: (N-min L~ff)..ff > 1 by A124,AXIOMS:22;
      (N-min L~ff)..ff < (N-max L~ff)..ff by A122,A123,SPRECT_5:25;
    then A126: (N-max L~ff)..ff > 1 by A125,AXIOMS:22;
      (N-max L~ff)..ff <= (E-max L~ff)..ff by A122,A123,SPRECT_5:26;
    then A127: Emax..ff > 1 by A123,A126,AXIOMS:22;
    A128: now assume A129: Gik..US <= 1;
       Gik..US >= 1 by A33,FINSEQ_4:31;
     then Gik..US = 1 by A129,AXIOMS:21;
     then Gik = US/.1 by A33,FINSEQ_5:41;
     hence contradiction by A17,A21,JORDAN1F:5;
    end;
    A130: Cage(C,n) is_sequence_on Ga by JORDAN9:def 1;
    then A131: ff is_sequence_on Ga by REVROT_1:34;
    A132: right_cell(godo,1,Ga)\L~godo c= RightComp godo by A85,A90,JORDAN9:29;
    A133: L~godo = L~(go^'pion1) \/ L~do by A89,TOPREAL8:35
       .= L~go \/ L~pion1 \/ L~do by A87,TOPREAL8:35;
      L~Cage(C,n) = L~US \/ L~LS by JORDAN1E:17;
    then A134: L~US c= L~Cage(C,n) & L~LS c= L~Cage(C,n) by XBOOLE_1:7;
    then A135: L~go c= L~Cage(C,n) & L~do c= L~Cage(C,n)
                                                     by A45,A52,XBOOLE_1:1;
    A136: W-min C in C by SPRECT_1:15;
    A137: L~pion = LSeg(Gik,Gij) by SPPOL_2:21;
    A138: now assume W-min C in L~godo;
     then A139: W-min C in L~go \/ L~pion1 or W-min C in L~do by A133,XBOOLE_0:
def 2;
     per cases by A139,XBOOLE_0:def 2;
      suppose W-min C in L~go;
       then C meets L~Cage(C,n) by A135,A136,XBOOLE_0:3;
       hence contradiction by JORDAN10:5;
      suppose W-min C in L~pion1;
       hence contradiction by A5,A77,A121,A137,XBOOLE_0:3;
      suppose W-min C in L~do;
       then C meets L~Cage(C,n) by A135,A136,XBOOLE_0:3;
       hence contradiction by JORDAN10:5;
    end;
      right_cell(Rotate(Cage(C,n),Wmin),1) =
          right_cell(ff,1,GoB ff) by A82,JORDAN1H:29
       .= right_cell(ff,1,GoB Cage(C,n)) by REVROT_1:28
       .= right_cell(ff,1,Ga) by JORDAN1H:52
       .= right_cell(ff-:Emax,1,Ga) by A127,A131,JORDAN1J:53
       .= right_cell(US,1,Ga) by JORDAN1E:def 1
       .= right_cell(R_Cut(US,Gik),1,Ga) by A33,A86,A128,JORDAN1J:52
       .= right_cell(go^'pion1,1,Ga) by A38,A88,JORDAN1J:51
       .= right_cell(godo,1,Ga) by A84,A90,JORDAN1J:51;
    then W-min C in right_cell(godo,1,Ga) by JORDAN1I:8;
    then A140: W-min C in right_cell(godo,1,Ga)\L~godo by A138,XBOOLE_0:def 4;
    A141: godo/.1 = (go^'pion1)/.1 by AMISTD_1:5
       .= Wmin by A58,AMISTD_1:5;
    A142: len US >= 2 by A16,AXIOMS:22;
    A143: godo/.2 = (go^'pion1)/.2 by A83,AMISTD_1:9
       .= US/.2 by A32,A71,AMISTD_1:9
       .= (US^'LS)/.2 by A142,AMISTD_1:9
       .= Rotate(Cage(C,n),Wmin)/.2 by JORDAN1E:15;
    A144: L~go \/ L~do is compact by COMPTS_1:19;
    A145: L~go \/ L~do c= L~Cage(C,n) by A135,XBOOLE_1:8;
      Wmin in L~go \/ L~do by A60,A72,XBOOLE_0:def 2;
    then A146: W-min (L~go \/ L~do) = Wmin by A144,A145,JORDAN1J:21;
    A147: (W-min (L~go \/ L~do))`1 = W-bound (L~go \/ L~do) & Wmin`1 = Wbo
                                                              by PSCOMP_1:84;
    A148: Gij`1 <= Gik`1 by A1,A2,SPRECT_3:25;
    then W-bound LSeg(Gik,Gij) = Gij`1 by SPRECT_1:62;
    then A149: W-bound L~pion1 = Gij`1 by A77,SPPOL_2:21;
      Gij`1 >= Wbo by A6,A134,PSCOMP_1:71;
    then Gij`1 > Wbo by A70,REAL_1:def 5;
    then W-min (L~go\/L~do\/L~pion1) = W-min (L~go \/ L~do)
                                              by A144,A146,A147,A149,JORDAN1J:
33;
    then A150: W-min L~godo = Wmin by A133,A146,XBOOLE_1:4;
    A151: rng godo c= L~godo by A85,SPPOL_2:18;
      2 in dom godo by A85,FINSEQ_3:27;
    then A152: godo/.2 in rng godo by PARTFUN2:4;
      godo/.2 in W-most L~Cage(C,n) by A143,JORDAN1I:27;
    then (godo/.2)`1 = (W-min L~godo)`1 by A150,PSCOMP_1:88
       .= W-bound L~godo by PSCOMP_1:84;
    then godo/.2 in W-most L~godo by A151,A152,SPRECT_2:16;
    then Rotate(godo,W-min L~godo)/.2 in W-most L~godo by A141,A150,FINSEQ_6:95
;
    then reconsider godo as clockwise_oriented non constant standard
                                    special_circular_sequence by JORDAN1I:27;
      len US in dom US by FINSEQ_5:6;
    then A153: US.len US = US/.len US by FINSEQ_4:def 4
       .= Emax by JORDAN1F:7;
    A154: E-max C in E-most C by PSCOMP_1:111;
    A155: east_halfline E-max C misses L~go
    proof
     assume east_halfline E-max C meets L~go;
     then consider p be set such that
      A156: p in east_halfline E-max C and
      A157: p in L~go by XBOOLE_0:3;
     reconsider p as Point of TOP-REAL 2 by A156;
       p in L~US by A45,A157;
     then p in east_halfline E-max C /\ L~Cage(C,n) by A134,A156,XBOOLE_0:def 3
;
     then A158: p`1 = Ebo by A154,JORDAN1A:104;
     then A159: p = Emax by A45,A157,JORDAN1J:46;
     then Emax = Gik by A7,A153,A157,JORDAN1J:43;
     then Gik`1 = Ga*(len Ga,k)`1 by A8,A11,A158,A159,JORDAN1A:92;
     hence contradiction by A1,A13,A29,JORDAN1G:7;
    end;
      now assume east_halfline E-max C meets L~godo;
     then A160: east_halfline E-max C meets (L~go \/ L~pion1) or
               east_halfline E-max C meets L~do by A133,XBOOLE_1:70;
     per cases by A160,XBOOLE_1:70;
      suppose east_halfline E-max C meets L~go;
       hence contradiction by A155;
      suppose east_halfline E-max C meets L~pion1;
       then consider p be set such that
        A161: p in east_halfline E-max C and
        A162: p in L~pion1 by XBOOLE_0:3;
       reconsider p as Point of TOP-REAL 2 by A161;
       A163: p`1 <= Gik`1 by A77,A137,A148,A162,TOPREAL1:9;
         k+1 <= len Ga by A1,NAT_1:38;
       then k+1-1 <= len Ga-1 by REAL_1:49;
       then A164: k <= len Ga-1 by XCMPLX_1:26;
       then len Ga-1 > 0 by A1,AXIOMS:22;
       then A165: k <= len Ga-'1 by A164,BINARITH:def 3;
         len Ga-'1 <= len Ga by GOBOARD9:2;
       then Gik`1 <= Ga*(len Ga-'1,1)`1 by A2,A8,A11,A19,A165,JORDAN1A:39;
       then p`1 <= Ga*(len Ga-'1,1)`1 by A163,AXIOMS:22;
       then p`1 <= E-bound C by A19,JORDAN8:15;
       then A166: p`1 <= (E-max C)`1 by PSCOMP_1:104;
         p`1 >= (E-max C)`1 by A161,JORDAN1A:def 3;
       then A167: p`1 = (E-max C)`1 by A166,AXIOMS:21;
         p`2 = (E-max C)`2 by A161,JORDAN1A:def 3;
       then p = E-max C by A167,TOPREAL3:11;
       hence contradiction by A5,A77,A121,A137,A162,XBOOLE_0:3;
      suppose east_halfline E-max C meets L~do;
       then consider p be set such that
        A168: p in east_halfline E-max C and
        A169: p in L~do by XBOOLE_0:3;
       reconsider p as Point of TOP-REAL 2 by A168;
         p in L~LS by A52,A169;
       then p in east_halfline E-max C /\ L~Cage(C,n)
                              by A134,A168,XBOOLE_0:def 3;
       then A170: p`1 = Ebo by A154,JORDAN1A:104;
       A171: (E-max C)`2 = p`2 by A168,JORDAN1A:def 3;
       set RC = Rotate(Cage(C,n),Emax);
       A172: E-max C in right_cell(RC,1) by JORDAN1I:9;
       A173: 1+1 <= len LS by A23,AXIOMS:22;
         LS = RC-:Wmin by JORDAN1G:26;
       then A174: LSeg(LS,1) = LSeg(RC,1) by A173,SPPOL_2:9;
       A175: L~RC = L~Cage(C,n) by REVROT_1:33;
       A176: len RC = len Cage(C,n) by REVROT_1:14;
       A177: GoB RC = GoB Cage(C,n) by REVROT_1:28
          .= Ga by JORDAN1H:52;
       A178: Emax in rng Cage(C,n) by SPRECT_2:50;
       A179: RC is_sequence_on Ga by A130,REVROT_1:34;
       A180: RC/.1 = E-max L~RC by A175,A178,FINSEQ_6:98;
       then consider ii,jj be Nat such that
        A181: [ii,jj+1] in Indices Ga and
        A182: [ii,jj] in Indices Ga and
        A183: RC/.1 = Ga*(ii,jj+1) and
        A184: RC/.(1+1) = Ga*(ii,jj) by A81,A176,A179,JORDAN1I:25;
       consider jj2 be Nat such that
        A185: 1 <= jj2 & jj2 <= width Ga and
        A186: Emax = Ga*(len Ga,jj2) by JORDAN1D:29;
       A187: len Ga >= 4 by JORDAN8:13;
       then len Ga >= 1 by AXIOMS:22;
       then [len Ga,jj2] in Indices Ga by A185,GOBOARD7:10;
       then A188: ii = len Ga by A175,A180,A181,A183,A186,GOBOARD1:21;
       A189: 1 <= ii & ii <= len Ga & 1 <= jj+1 & jj+1 <= width Ga
                                                           by A181,GOBOARD5:1;
       A190: 1 <= ii & ii <= len Ga & 1 <= jj & jj <= width Ga
                                                           by A182,GOBOARD5:1;
       A191: ii+1 <> ii by NAT_1:38;
         jj+1 > jj by NAT_1:38;
       then jj+1+1 <> jj by NAT_1:38;
       then A192: right_cell(RC,1) = cell(Ga,ii-'1,jj)
                          by A81,A176,A177,A181,A182,A183,A184,A191,GOBOARD5:
def 6;
       A193: ii-'1+1 = ii by A189,AMI_5:4;
       A194: ii-1 >= 4-1 by A187,A188,REAL_1:49;
       then A195: ii-1 >= 1 by AXIOMS:22;
         ii-1 >= 0 by A194,AXIOMS:22;
       then A196: 1 <= ii-'1 by A195,BINARITH:def 3;
       then A197: Ga*(ii-'1,jj)`2 <= p`2 & p`2 <= Ga*(ii-'1,jj+1)`2
                             by A171,A172,A189,A190,A192,A193,JORDAN9:19;
       A198: ii-'1 < len Ga by A189,A193,NAT_1:38;
       then A199: Ga*(ii-'1,jj)`2 = Ga*(1,jj)`2 by A190,A196,GOBOARD5:2
          .= Ga*(ii,jj)`2 by A190,GOBOARD5:2;
       A200: Ga*(ii-'1,jj+1)`2 = Ga*(1,jj+1)`2 by A189,A196,A198,GOBOARD5:2
          .= Ga*(ii,jj+1)`2 by A189,GOBOARD5:2;
         Ga*(len Ga,jj)`1 = Ebo & Ebo = Ga*(len Ga,jj+1)`1
                                                by A8,A189,A190,JORDAN1A:92;
       then p in LSeg(RC/.1,RC/.(1+1))
                              by A170,A183,A184,A188,A197,A199,A200,GOBOARD7:8;
       then A201: p in LSeg(LS,1) by A81,A174,A176,TOPREAL1:def 5;
       A202: p in LSeg(do,Index(p,do)) by A169,JORDAN3:42;
       A203: do = mid(LS,Gij..LS,len LS) by A36,JORDAN1J:37;
       A204: 1<=Gij..LS & Gij..LS<=len LS by A36,FINSEQ_4:31;
         Gij..LS <> len LS by A28,A36,FINSEQ_4:29;
       then A205: Gij..LS < len LS by A204,REAL_1:def 5;
       A206: 1<=Index(p,do) & Index(p,do) < len do by A169,JORDAN3:41;
       A207: Index(Gij,LS)+1 = Gij..LS by A31,A36,JORDAN1J:56;
       consider t be Nat such that
        A208: t in dom LS and
        A209: LS.t = Gij by A36,FINSEQ_2:11;
       A210: 1 <= t & t <= len LS by A208,FINSEQ_3:27;
       then 1 < t by A31,A209,REAL_1:def 5;
       then Index(Gij,LS)+1 = t by A209,A210,JORDAN3:45;
       then A211: len L_Cut(LS,Gij) = len LS-Index(Gij,LS)
                                                  by A6,A28,A209,JORDAN3:61;
       set tt = Index(p,do)+(Gij..LS)-'1;
       A212: 1<=Index(Gij,LS) & 0+Index(Gij,LS) < len LS by A6,JORDAN3:41;
       then A213: len LS-Index(Gij,LS) > 0 by REAL_1:86;
       then Index(p,do) < len LS-'Index(Gij,LS) by A206,A211,BINARITH:def 3;
       then Index(p,do)+1 <= len LS-'Index(Gij,LS) by NAT_1:38;
       then Index(p,do) <= len LS-'Index(Gij,LS)-1 by REAL_1:84;
       then Index(p,do) <= len LS-Index(Gij,LS)-1 by A213,BINARITH:def 3;
       then A214: Index(p,do) <= len LS-Gij..LS by A207,XCMPLX_1:36;
       then len LS-Gij..LS >= 1 by A206,AXIOMS:22;
       then len LS-Gij..LS >= 0 by AXIOMS:22;
       then Index(p,do) <= len LS-'Gij..LS by A214,BINARITH:def 3;
       then Index(p,do) < len LS-'(Gij..LS)+1 by NAT_1:38;
       then A215: LSeg(mid(LS,Gij..LS,len LS),Index(p,do)) =
            LSeg(LS,Index(p,do)+(Gij..LS)-'1) by A204,A205,A206,JORDAN4:31;
       A216: 1+1 <= Gij..LS by A207,A212,REAL_1:55;
       then Index(p,do)+Gij..LS >= 1+1+1 by A206,REAL_1:55;
       then A217: Index(p,do)+Gij..LS-1 >= 1+1+1-1 by REAL_1:49;
       then A218: Index(p,do)+Gij..LS-1 >= 0 by AXIOMS:22;
       then A219: tt >= 1+1 by A217,BINARITH:def 3;
       A220: 2 in dom LS by A173,FINSEQ_3:27;
         now per cases by A219,REAL_1:def 5;
        suppose tt > 1+1;
         then LSeg(LS,1) misses LSeg(LS,tt) by TOPREAL1:def 9;
         hence contradiction by A201,A202,A203,A215,XBOOLE_0:3;
        suppose A221: tt = 1+1;
         then LSeg(LS,1) /\ LSeg(LS,tt) = {LS/.2} by A23,TOPREAL1:def 8;
         then p in {LS/.2} by A201,A202,A203,A215,XBOOLE_0:def 3;
         then A222: p = LS/.2 by TARSKI:def 1;
         then A223: p..LS = 2 by A220,FINSEQ_5:44;
           1+1 = Index(p,do)+(Gij..LS)-1 by A218,A221,BINARITH:def 3;
         then 1+1+1 = Index(p,do)+(Gij..LS) by XCMPLX_1:27;
         then A224: Gij..LS = 2 by A206,A216,JORDAN1E:10;
           p in rng LS by A220,A222,PARTFUN2:4;
         then p = Gij by A36,A223,A224,FINSEQ_5:10;
         then Gij`1 = Ebo by A222,JORDAN1G:40;
         then Gij`1 = Ga*(len Ga,j)`1 by A8,A10,JORDAN1A:92;
         hence contradiction by A1,A12,A63,JORDAN1G:7;
       end;
       hence contradiction;
    end;
    then east_halfline E-max C c= (L~godo)` by SUBSET_1:43;
    then consider W be Subset of TOP-REAL 2 such that
     A225: W is_a_component_of (L~godo)` and
     A226: east_halfline E-max C c= W by GOBOARD9:5;
      east_halfline E-max C is not Bounded by JORDAN1C:9;
    then W is not Bounded by A226,JORDAN2C:16;
    then W is_outside_component_of L~godo by A225,JORDAN2C:def 4;
    then W c= UBD L~godo by JORDAN2C:27;
    then A227: east_halfline E-max C c= UBD L~godo by A226,XBOOLE_1:1;
      E-max C in east_halfline E-max C by JORDAN1C:7;
    then E-max C in UBD L~godo by A227;
    then E-max C in LeftComp godo by GOBRD14:46;
    then LA meets L~godo by A120,A121,A132,A140,JORDAN1J:36;
    then A228: LA meets (L~go \/ L~pion1) or LA meets L~do by A133,XBOOLE_1:70;
    A229: LA c= C by JORDAN1A:16;
    per cases by A228,XBOOLE_1:70;
     suppose LA meets L~go;
      then LA meets L~Cage(C,n) by A135,XBOOLE_1:63;
      then C meets L~Cage(C,n) by A229,XBOOLE_1:63;
      hence contradiction by JORDAN10:5;
     suppose LA meets L~pion1;
      hence contradiction by A5,A77,A137;
     suppose LA meets L~do;
      then LA meets L~Cage(C,n) by A135,XBOOLE_1:63;
      then C meets L~Cage(C,n) by A229,XBOOLE_1:63;
      hence contradiction by JORDAN10:5;
   end;

  theorem Th31:
   for C be Simple_closed_curve
   for i,j,k be Nat st
    1 < j & j <= k & k < len Gauge(C,n) &
    1 <= i & i <= width Gauge(C,n) &
    LSeg(Gauge(C,n)*(j,i),Gauge(C,n)*(k,i)) /\ L~Upper_Seq(C,n) =
     {Gauge(C,n)*(k,i)} &
    LSeg(Gauge(C,n)*(j,i),Gauge(C,n)*(k,i)) /\ L~Lower_Seq(C,n) =
     {Gauge(C,n)*(j,i)} holds
      LSeg(Gauge(C,n)*(j,i),Gauge(C,n)*(k,i)) meets Upper_Arc C
   proof
    let C be Simple_closed_curve;
    let i,j,k be Nat;
    set Ga = Gauge(C,n);
    set US = Upper_Seq(C,n);
    set LS = Lower_Seq(C,n);
    set UA = Upper_Arc C;
    set Wmin = W-min L~Cage(C,n);
    set Emax = E-max L~Cage(C,n);
    set Wbo = W-bound L~Cage(C,n);
    set Ebo = E-bound L~Cage(C,n);
    set Gij = Ga*(j,i);
    set Gik = Ga*(k,i);
    assume that
     A1: 1 < j & j <= k & k < len Ga and
     A2: 1 <= i & i <= width Ga and
     A3: LSeg(Gij,Gik) /\ L~US = {Gik} and
     A4: LSeg(Gij,Gik) /\ L~LS = {Gij} and
     A5: LSeg(Gij,Gik) misses UA;
      Gij in {Gij} by TARSKI:def 1;
    then A6: Gij in L~LS by A4,XBOOLE_0:def 3;
      Gik in {Gik} by TARSKI:def 1;
    then A7: Gik in L~US by A3,XBOOLE_0:def 3;
    A8: len Ga = width Ga by JORDAN8:def 1;
    A9: j <> k by A1,A2,A6,A7,Th29;
    A10: 1 <= j & j <= width Ga by A1,A8,AXIOMS:22;
    A11: 1 <= k & k <= width Ga by A1,AXIOMS:22,JORDAN8:def 1;
    A12: [j,i] in Indices Ga by A2,A8,A10,GOBOARD7:10;
    A13: [k,i] in Indices Ga by A2,A8,A11,GOBOARD7:10;
    A14: US is_sequence_on Ga by JORDAN1G:4;
    A15: LS is_sequence_on Ga by JORDAN1G:5;
    set go = R_Cut(US,Gik);
    set do = L_Cut(LS,Gij);
    A16: len US >= 3 by JORDAN1E:19;
    then len US >= 1 by AXIOMS:22;
    then 1 in dom US by FINSEQ_3:27;
    then A17: US.1 = US/.1 by FINSEQ_4:def 4
       .= Wmin by JORDAN1F:5;
    A18: Wmin`1 = Wbo by PSCOMP_1:84
       .= Ga*(1,k)`1 by A8,A11,JORDAN1A:94;
      len Ga >= 4 by JORDAN8:13;
    then A19: len Ga >= 1 by AXIOMS:22;
    then A20: [1,k] in Indices Ga by A11,GOBOARD7:10;
    then A21: Gik <> US.1 by A1,A13,A17,A18,JORDAN1G:7;
    then reconsider go as being_S-Seq FinSequence of TOP-REAL 2
                                                            by A7,JORDAN3:70;
    A22: [1,j] in Indices Ga by A10,A19,GOBOARD7:10;
    A23: len LS >= 1+2 by JORDAN1E:19;
    then len LS >= 1 by AXIOMS:22;
    then A24: 1 in dom LS & len LS in dom LS by FINSEQ_3:27;
    then A25: LS.len LS = LS/.len LS by FINSEQ_4:def 4
       .= Wmin by JORDAN1F:8;
    A26: Wmin`1 = Wbo by PSCOMP_1:84
       .= Ga*(1,k)`1 by A8,A11,JORDAN1A:94;
    A27: [j,i] in Indices Ga by A2,A8,A10,GOBOARD7:10;
    then A28: Gij <> LS.len LS by A1,A20,A25,A26,JORDAN1G:7;
    then reconsider do as being_S-Seq FinSequence of TOP-REAL 2
                                                            by A6,JORDAN3:69;
    A29: [len Ga,k] in Indices Ga by A11,A19,GOBOARD7:10;
    A30: LS.1 = LS/.1 by A24,FINSEQ_4:def 4
       .= Emax by JORDAN1F:6;
      Emax`1 = Ebo by PSCOMP_1:104
       .= Ga*(len Ga,k)`1 by A8,A11,JORDAN1A:92;
    then A31: Gij <> LS.1 by A1,A27,A29,A30,JORDAN1G:7;
    A32: len go >= 1+1 by TOPREAL1:def 10;
    A33: Gik in rng US by A2,A7,A8,A11,A14,JORDAN1J:40;
    then A34: go is_sequence_on Ga by A14,JORDAN1J:38;
    A35: len do >= 1+1 by TOPREAL1:def 10;
    A36: Gij in rng LS by A2,A6,A8,A10,A15,JORDAN1J:40;
    then A37: do is_sequence_on Ga by A15,JORDAN1J:39;
    reconsider go as non constant s.c.c.
            (being_S-Seq FinSequence of TOP-REAL 2) by A32,A34,JGRAPH_1:16,
JORDAN8:8;
    reconsider do as non constant s.c.c.
            (being_S-Seq FinSequence of TOP-REAL 2) by A35,A37,JGRAPH_1:16,
JORDAN8:8;
    A38: len go > 1 by A32,NAT_1:38;
    then A39: len go in dom go by FINSEQ_3:27;
    then A40: go/.len go = go.len go by FINSEQ_4:def 4
       .= Gik by A7,JORDAN3:59;
      len do >= 1 by A35,AXIOMS:22;
    then 1 in dom do by FINSEQ_3:27;
    then A41: do/.1 = do.1 by FINSEQ_4:def 4
       .= Gij by A6,JORDAN3:58;
    reconsider m = len go - 1 as Nat by A39,FINSEQ_3:28;
    A42: m+1 = len go by XCMPLX_1:27;
    then A43: len go-'1 = m by BINARITH:39;
    A44: LSeg(go,m) c= L~go by TOPREAL3:26;
    A45: L~go c= L~US by A7,JORDAN3:76;
    then LSeg(go,m) c= L~US by A44,XBOOLE_1:1;
    then A46: LSeg(go,m) /\ LSeg(Gik,Gij) c= {Gik} by A3,XBOOLE_1:26;
      m >= 1 by A32,REAL_1:84;
    then A47: LSeg(go,m) = LSeg(go/.m,Gik) by A40,A42,TOPREAL1:def 5;
      {Gik} c= LSeg(go,m) /\ LSeg(Gik,Gij)
    proof
     let x be set;
     assume x in {Gik};
     then A48: x = Gik by TARSKI:def 1;
     A49: Gik in LSeg(go,m) by A47,TOPREAL1:6;
       Gik in LSeg(Gik,Gij) by TOPREAL1:6;
     hence x in LSeg(go,m) /\ LSeg(Gik,Gij) by A48,A49,XBOOLE_0:def 3;
    end;
    then A50: LSeg(go,m) /\ LSeg(Gik,Gij) = {Gik} by A46,XBOOLE_0:def 10;
    A51: LSeg(do,1) c= L~do by TOPREAL3:26;
    A52: L~do c= L~LS by A6,JORDAN3:77;
    then LSeg(do,1) c= L~LS by A51,XBOOLE_1:1;
    then A53: LSeg(do,1) /\ LSeg(Gik,Gij) c= {Gij} by A4,XBOOLE_1:26;
    A54: LSeg(do,1) = LSeg(Gij,do/.(1+1)) by A35,A41,TOPREAL1:def 5;
      {Gij} c= LSeg(do,1) /\ LSeg(Gik,Gij)
    proof
     let x be set;
     assume x in {Gij};
     then A55: x = Gij by TARSKI:def 1;
     A56: Gij in LSeg(do,1) by A54,TOPREAL1:6;
       Gij in LSeg(Gik,Gij) by TOPREAL1:6;
     hence x in LSeg(do,1) /\ LSeg(Gik,Gij) by A55,A56,XBOOLE_0:def 3;
    end;
    then A57: LSeg(Gik,Gij) /\ LSeg(do,1) = {Gij} by A53,XBOOLE_0:def 10;
    A58: go/.1 = US/.1 by A7,SPRECT_3:39
       .= Wmin by JORDAN1F:5;
    then A59: go/.1 = LS/.len LS by JORDAN1F:8
       .= do/.len do by A6,JORDAN1J:35;
    A60: rng go c= L~go & rng do c= L~do by A32,A35,SPPOL_2:18;
    A61: {go/.1} c= L~go /\ L~do
    proof
     let x be set;
     assume x in {go/.1};
     then x = go/.1 by TARSKI:def 1;
     then x in rng go & x in rng do by A59,FINSEQ_6:46,REVROT_1:3;
     hence x in L~go /\ L~do by A60,XBOOLE_0:def 3;
    end;
    A62: LS.1 = LS/.1 by A24,FINSEQ_4:def 4
       .= Emax by JORDAN1F:6;
    A63: [len Ga,j] in Indices Ga by A10,A19,GOBOARD7:10;
      L~go /\ L~do c= {go/.1}
    proof
     let x be set;
     assume x in L~go /\ L~do;
     then A64: x in L~go & x in L~do by XBOOLE_0:def 3;
     then x in L~US /\ L~LS by A45,A52,XBOOLE_0:def 3;
     then x in {Wmin,Emax} by JORDAN1E:20;
     then A65: x = Wmin or x = Emax by TARSKI:def 2;
       now assume x = Emax;
      then A66: Emax = Gij by A6,A62,A64,JORDAN1E:11;
        Ga*(len Ga,j)`1 = Ebo by A8,A10,JORDAN1A:92;
      then Emax`1 <> Ebo by A1,A12,A63,A66,JORDAN1G:7;
      hence contradiction by PSCOMP_1:104;
     end;
     hence x in {go/.1} by A58,A65,TARSKI:def 1;
    end;
    then A67: L~go /\ L~do = {go/.1} by A61,XBOOLE_0:def 10;
    set W2 = go/.2;
    A68: 2 in dom go by A32,FINSEQ_3:27;
    A69: Gik..US >= 1 by A33,FINSEQ_4:31;
    A70: now assume Gij`1 = Wbo;
     then Ga*(1,j)`1 = Ga*(j,i)`1 by A8,A10,JORDAN1A:94;
     hence contradiction by A1,A12,A22,JORDAN1G:7;
    end;
      go = mid(US,1,Gik..US) by A33,JORDAN1G:57
       .= US|(Gik..US) by A69,JORDAN3:25;
    then A71: W2 = US/.2 by A68,TOPREAL1:1;
    A72: Wmin in rng go by A58,FINSEQ_6:46;
    set pion = <*Gik,Gij*>;
    A73: now let n be Nat;
     assume n in dom pion;
     then n in {1,2} by FINSEQ_1:4,FINSEQ_3:29;
     then n = 1 or n = 2 by TARSKI:def 2;
     then pion/.n = Gik or pion/.n = Gij by FINSEQ_4:26;
     hence ex j,i be Nat st [j,i] in Indices Ga & pion/.n = Ga*(j,i)
                                                                  by A12,A13;
    end;
    A74: Gik <> Gij by A9,A12,A13,GOBOARD1:21;
      Gik`2 = Ga*(1,i)`2 by A2,A8,A11,GOBOARD5:2
       .= Gij`2 by A2,A8,A10,GOBOARD5:2;
    then LSeg(Gik,Gij) is horizontal by SPPOL_1:36;
    then pion is_S-Seq by A74,JORDAN1B:9;
    then consider pion1 be FinSequence of TOP-REAL 2 such that
     A75: pion1 is_sequence_on Ga and
     A76: pion1 is_S-Seq and
     A77: L~pion = L~pion1 and
     A78: pion/.1 = pion1/.1 and
     A79: pion/.len pion = pion1/.len pion1 and
     A80: len pion <= len pion1 by A73,GOBOARD3:2;
    reconsider pion1 as being_S-Seq FinSequence of TOP-REAL 2 by A76;
    set godo = go^'pion1^'do;
      len Cage(C,n) > 4 by GOBOARD7:36;
    then A81: 1+1 <= len Cage(C,n) by AXIOMS:22;
    then A82: 1+1 <= len Rotate(Cage(C,n),Wmin) by REVROT_1:14;
      len (go^'pion1) >= len go by TOPREAL8:7;
    then A83: len (go^'pion1) >= 1+1 by A32,AXIOMS:22;
    then A84: len (go^'pion1) > 1+0 by NAT_1:38;
      len godo >= len (go^'pion1) by TOPREAL8:7;
    then A85: 1+1 <= len godo by A83,AXIOMS:22;
    A86: US is_sequence_on Ga by JORDAN1G:4;
    A87: go/.len go = pion1/.1 by A40,A78,FINSEQ_4:26;
    then A88: go^'pion1 is_sequence_on Ga by A34,A75,TOPREAL8:12;
    A89: (go^'pion1)/.len (go^'pion1) = pion/.len pion by A79,AMISTD_1:6
       .= pion/.2 by FINSEQ_1:61
       .= do/.1 by A41,FINSEQ_4:26;
    then A90: godo is_sequence_on Ga by A37,A88,TOPREAL8:12;
      LSeg(pion1,1) c= L~<*Gik,Gij*> by A77,TOPREAL3:26;
    then LSeg(pion1,1) c= LSeg(Gik,Gij) by SPPOL_2:21;
    then A91: LSeg(go,len go-'1) /\ LSeg(pion1,1) c= {Gik}
                                                      by A43,A50,XBOOLE_1:27;
    A92: len pion1 >= 1+1 by A80,FINSEQ_1:61;
      {Gik} c= LSeg(go,m) /\ LSeg(pion1,1)
    proof
     let x be set;
     assume x in {Gik};
     then A93: x = Gik by TARSKI:def 1;
     A94: Gik in LSeg(go,m) by A47,TOPREAL1:6;
       Gik in LSeg(pion1,1) by A40,A87,A92,TOPREAL1:27;
     hence x in LSeg(go,m) /\ LSeg(pion1,1) by A93,A94,XBOOLE_0:def 3;
    end;
    then LSeg(go,len go-'1) /\ LSeg(pion1,1) = { go/.len go }
                                             by A40,A43,A91,XBOOLE_0:def 10;
    then A95: go^'pion1 is unfolded by A87,TOPREAL8:34;
      len pion1 >= 2+0 by A80,FINSEQ_1:61;
    then A96: len pion1-2 >= 0 by REAL_1:84;
    A97: len (go^'pion1)-1 >= 0 by A84,REAL_1:84;
      len (go^'pion1)+1-1 = len go+len pion1-1 by GRAPH_2:13;
    then len (go^'pion1)-1 = len go+len pion1-1-1 by XCMPLX_1:26
       .= len go + len pion1-(1+1) by XCMPLX_1:36
       .= len go + (len pion1-2) by XCMPLX_1:29
       .= len go + (len pion1-'2) by A96,BINARITH:def 3;
    then A98: len (go^'pion1)-'1 = len go + (len pion1-'2)
                                                      by A97,BINARITH:def 3;
    A99: len pion1-1 >= 1 by A92,REAL_1:84;
    then A100: len pion1-1 >= 0 by AXIOMS:22;
    then A101: len pion1-'1 = len pion1-1 by BINARITH:def 3;
    A102: len pion1-'2+1 = len pion1-2+1 by A96,BINARITH:def 3
       .= len pion1-(2-1) by XCMPLX_1:37
       .= len pion1-'1 by A100,BINARITH:def 3;
      len pion1-1+1 <= len pion1 by XCMPLX_1:27;
    then A103: len pion1-'1 < len pion1 by A101,NAT_1:38;
      LSeg(pion1,len pion1-'1) c= L~<*Gik,Gij*> by A77,TOPREAL3:26;
    then LSeg(pion1,len pion1-'1) c= LSeg(Gik,Gij) by SPPOL_2:21;
    then A104: LSeg(pion1,len pion1-'1) /\ LSeg(do,1) c= {Gij}
                                                          by A57,XBOOLE_1:27;
      {Gij} c= LSeg(pion1,len pion1-'1) /\ LSeg(do,1)
    proof
     let x be set;
     assume x in {Gij};
     then A105: x = Gij by TARSKI:def 1;
     A106: Gij in LSeg(do,1) by A54,TOPREAL1:6;
     A107: len pion1-'1+1 = len pion1 by A101,XCMPLX_1:27;
     then pion1/.(len pion1-'1+1) = pion/.2 by A79,FINSEQ_1:61
        .= Gij by FINSEQ_4:26;
     then Gij in LSeg(pion1,len pion1-'1) by A99,A101,A107,TOPREAL1:27;
     hence x in LSeg(pion1,len pion1-'1) /\ LSeg(do,1)
                                                  by A105,A106,XBOOLE_0:def 3;
    end;
    then LSeg(pion1,len pion1-'1) /\ LSeg(do,1) = {Gij} by A104,XBOOLE_0:def 10
;
    then A108: LSeg(go^'pion1,len go+(len pion1-'2)) /\ LSeg(do,1) =
       {(go^'pion1)/.len (go^'pion1)} by A41,A87,A89,A102,A103,TOPREAL8:31;
    A109: (go^'pion1) is non trivial by A83,SPPOL_1:2;
    A110: rng pion1 c= L~pion1 by A92,SPPOL_2:18;
    A111: {pion1/.1} c= L~go /\ L~pion1
    proof
     let x be set;
     assume x in {pion1/.1};
     then x = pion1/.1 by TARSKI:def 1;
     then x in rng go & x in rng pion1 by A87,FINSEQ_6:46,REVROT_1:3;
     hence x in L~go /\ L~pion1 by A60,A110,XBOOLE_0:def 3;
    end;
      L~go /\ L~pion1 c= {pion1/.1}
    proof
     let x be set;
     assume x in L~go /\ L~pion1;
     then x in L~go & x in L~pion1 by XBOOLE_0:def 3;
     then x in L~pion1 /\ L~US by A45,XBOOLE_0:def 3;
     hence x in {pion1/.1} by A3,A40,A77,A87,SPPOL_2:21;
    end;
    then A112: L~go /\ L~pion1 = {pion1/.1} by A111,XBOOLE_0:def 10;
    then A113: (go^'pion1) is s.n.c. by A87,JORDAN1J:54;
      rng go /\ rng pion1 c= {pion1/.1} by A60,A110,A112,XBOOLE_1:27;
    then A114: go^'pion1 is one-to-one by JORDAN1J:55;
    A115: pion/.len pion = pion/.2 by FINSEQ_1:61
       .= do/.1 by A41,FINSEQ_4:26;
    A116: {pion1/.len pion1} c= L~do /\ L~pion1
    proof
     let x be set;
     assume x in {pion1/.len pion1};
     then x = pion1/.len pion1 by TARSKI:def 1;
     then x in rng do & x in rng pion1 by A79,A115,FINSEQ_6:46,REVROT_1:3;
     hence x in L~do /\ L~pion1 by A60,A110,XBOOLE_0:def 3;
    end;
      L~do /\ L~pion1 c= {pion1/.len pion1}
    proof
     let x be set;
     assume x in L~do /\ L~pion1;
     then x in L~do & x in L~pion1 by XBOOLE_0:def 3;
     then x in L~pion1 /\ L~LS by A52,XBOOLE_0:def 3;
     hence x in {pion1/.len pion1} by A4,A41,A77,A79,A115,SPPOL_2:21;
    end;
    then A117: L~do /\ L~pion1 = {pion1/.len pion1} by A116,XBOOLE_0:def 10;
    A118: L~(go^'pion1) /\ L~do = (L~go \/ L~pion1) /\ L~do by A87,TOPREAL8:35
       .= {go/.1} \/ {do/.1} by A67,A79,A115,A117,XBOOLE_1:23
       .= {(go^'pion1)/.1} \/ {do/.1} by AMISTD_1:5
       .= {(go^'pion1)/.1,do/.1} by ENUMSET1:41;
      do/.len do = (go^'pion1)/.1 by A59,AMISTD_1:5;
    then reconsider godo as non constant standard special_circular_sequence
           by A85,A89,A90,A95,A98,A108,A109,A113,A114,A118,JORDAN8:7,8,TOPREAL8
:11,33,34;
    A119: UA is_an_arc_of W-min C,E-max C by JORDAN6:def 8;
    then A120: UA is connected by JORDAN6:11;
    A121: W-min C in UA & E-max C in UA by A119,TOPREAL1:4;
    set ff = Rotate(Cage(C,n),Wmin);
      Wmin in rng Cage(C,n) by SPRECT_2:47;
    then A122: ff/.1 = Wmin by FINSEQ_6:98;
    A123: L~ff = L~Cage(C,n) by REVROT_1:33;
    then A124: (W-max L~ff)..ff > 1 by A122,SPRECT_5:23;
      (W-max L~ff)..ff <= (N-min L~ff)..ff by A122,A123,SPRECT_5:24;
    then A125: (N-min L~ff)..ff > 1 by A124,AXIOMS:22;
      (N-min L~ff)..ff < (N-max L~ff)..ff by A122,A123,SPRECT_5:25;
    then A126: (N-max L~ff)..ff > 1 by A125,AXIOMS:22;
      (N-max L~ff)..ff <= (E-max L~ff)..ff by A122,A123,SPRECT_5:26;
    then A127: Emax..ff > 1 by A123,A126,AXIOMS:22;
    A128: now assume A129: Gik..US <= 1;
       Gik..US >= 1 by A33,FINSEQ_4:31;
     then Gik..US = 1 by A129,AXIOMS:21;
     then Gik = US/.1 by A33,FINSEQ_5:41;
     hence contradiction by A17,A21,JORDAN1F:5;
    end;
    A130: Cage(C,n) is_sequence_on Ga by JORDAN9:def 1;
    then A131: ff is_sequence_on Ga by REVROT_1:34;
    A132: right_cell(godo,1,Ga)\L~godo c= RightComp godo by A85,A90,JORDAN9:29;
    A133: L~godo = L~(go^'pion1) \/ L~do by A89,TOPREAL8:35
       .= L~go \/ L~pion1 \/ L~do by A87,TOPREAL8:35;
      L~Cage(C,n) = L~US \/ L~LS by JORDAN1E:17;
    then A134: L~US c= L~Cage(C,n) & L~LS c= L~Cage(C,n) by XBOOLE_1:7;
    then A135: L~go c= L~Cage(C,n) & L~do c= L~Cage(C,n)
                                                     by A45,A52,XBOOLE_1:1;
    A136: W-min C in C by SPRECT_1:15;
    A137: L~pion = LSeg(Gik,Gij) by SPPOL_2:21;
    A138: now assume W-min C in L~godo;
     then A139: W-min C in L~go \/ L~pion1 or W-min C in L~do by A133,XBOOLE_0:
def 2;
     per cases by A139,XBOOLE_0:def 2;
      suppose W-min C in L~go;
       then C meets L~Cage(C,n) by A135,A136,XBOOLE_0:3;
       hence contradiction by JORDAN10:5;
      suppose W-min C in L~pion1;
       hence contradiction by A5,A77,A121,A137,XBOOLE_0:3;
      suppose W-min C in L~do;
       then C meets L~Cage(C,n) by A135,A136,XBOOLE_0:3;
       hence contradiction by JORDAN10:5;
    end;
      right_cell(Rotate(Cage(C,n),Wmin),1) =
          right_cell(ff,1,GoB ff) by A82,JORDAN1H:29
       .= right_cell(ff,1,GoB Cage(C,n)) by REVROT_1:28
       .= right_cell(ff,1,Ga) by JORDAN1H:52
       .= right_cell(ff-:Emax,1,Ga) by A127,A131,JORDAN1J:53
       .= right_cell(US,1,Ga) by JORDAN1E:def 1
       .= right_cell(R_Cut(US,Gik),1,Ga) by A33,A86,A128,JORDAN1J:52
       .= right_cell(go^'pion1,1,Ga) by A38,A88,JORDAN1J:51
       .= right_cell(godo,1,Ga) by A84,A90,JORDAN1J:51;
    then W-min C in right_cell(godo,1,Ga) by JORDAN1I:8;
    then A140: W-min C in right_cell(godo,1,Ga)\L~godo by A138,XBOOLE_0:def 4;
    A141: godo/.1 = (go^'pion1)/.1 by AMISTD_1:5
       .= Wmin by A58,AMISTD_1:5;
    A142: len US >= 2 by A16,AXIOMS:22;
    A143: godo/.2 = (go^'pion1)/.2 by A83,AMISTD_1:9
       .= US/.2 by A32,A71,AMISTD_1:9
       .= (US^'LS)/.2 by A142,AMISTD_1:9
       .= Rotate(Cage(C,n),Wmin)/.2 by JORDAN1E:15;
    A144: L~go \/ L~do is compact by COMPTS_1:19;
    A145: L~go \/ L~do c= L~Cage(C,n) by A135,XBOOLE_1:8;
      Wmin in L~go \/ L~do by A60,A72,XBOOLE_0:def 2;
    then A146: W-min (L~go \/ L~do) = Wmin by A144,A145,JORDAN1J:21;
    A147: (W-min (L~go \/ L~do))`1 = W-bound (L~go \/ L~do) & Wmin`1 = Wbo
                                                              by PSCOMP_1:84;
    A148: Gij`1 <= Gik`1 by A1,A2,SPRECT_3:25;
    then W-bound LSeg(Gik,Gij) = Gij`1 by SPRECT_1:62;
    then A149: W-bound L~pion1 = Gij`1 by A77,SPPOL_2:21;
      Gij`1 >= Wbo by A6,A134,PSCOMP_1:71;
    then Gij`1 > Wbo by A70,REAL_1:def 5;
    then W-min (L~go\/L~do\/L~pion1) = W-min (L~go \/ L~do)
                                              by A144,A146,A147,A149,JORDAN1J:
33;
    then A150: W-min L~godo = Wmin by A133,A146,XBOOLE_1:4;
    A151: rng godo c= L~godo by A85,SPPOL_2:18;
      2 in dom godo by A85,FINSEQ_3:27;
    then A152: godo/.2 in rng godo by PARTFUN2:4;
      godo/.2 in W-most L~Cage(C,n) by A143,JORDAN1I:27;
    then (godo/.2)`1 = (W-min L~godo)`1 by A150,PSCOMP_1:88
       .= W-bound L~godo by PSCOMP_1:84;
    then godo/.2 in W-most L~godo by A151,A152,SPRECT_2:16;
    then Rotate(godo,W-min L~godo)/.2 in W-most L~godo by A141,A150,FINSEQ_6:95
;
    then reconsider godo as clockwise_oriented non constant standard
                                    special_circular_sequence by JORDAN1I:27;
      len US in dom US by FINSEQ_5:6;
    then A153: US.len US = US/.len US by FINSEQ_4:def 4
       .= Emax by JORDAN1F:7;
    A154: E-max C in E-most C by PSCOMP_1:111;
    A155: east_halfline E-max C misses L~go
    proof
     assume east_halfline E-max C meets L~go;
     then consider p be set such that
      A156: p in east_halfline E-max C and
      A157: p in L~go by XBOOLE_0:3;
     reconsider p as Point of TOP-REAL 2 by A156;
       p in L~US by A45,A157;
     then p in east_halfline E-max C /\ L~Cage(C,n) by A134,A156,XBOOLE_0:def 3
;
     then A158: p`1 = Ebo by A154,JORDAN1A:104;
     then A159: p = Emax by A45,A157,JORDAN1J:46;
     then Emax = Gik by A7,A153,A157,JORDAN1J:43;
     then Gik`1 = Ga*(len Ga,k)`1 by A8,A11,A158,A159,JORDAN1A:92;
     hence contradiction by A1,A13,A29,JORDAN1G:7;
    end;
      now assume east_halfline E-max C meets L~godo;
     then A160: east_halfline E-max C meets (L~go \/ L~pion1) or
               east_halfline E-max C meets L~do by A133,XBOOLE_1:70;
     per cases by A160,XBOOLE_1:70;
      suppose east_halfline E-max C meets L~go;
       hence contradiction by A155;
      suppose east_halfline E-max C meets L~pion1;
       then consider p be set such that
        A161: p in east_halfline E-max C and
        A162: p in L~pion1 by XBOOLE_0:3;
       reconsider p as Point of TOP-REAL 2 by A161;
       A163: p`1 <= Gik`1 by A77,A137,A148,A162,TOPREAL1:9;
         k+1 <= len Ga by A1,NAT_1:38;
       then k+1-1 <= len Ga-1 by REAL_1:49;
       then A164: k <= len Ga-1 by XCMPLX_1:26;
       then len Ga-1 > 0 by A1,AXIOMS:22;
       then A165: k <= len Ga-'1 by A164,BINARITH:def 3;
         len Ga-'1 <= len Ga by GOBOARD9:2;
       then Gik`1 <= Ga*(len Ga-'1,1)`1 by A2,A8,A11,A19,A165,JORDAN1A:39;
       then p`1 <= Ga*(len Ga-'1,1)`1 by A163,AXIOMS:22;
       then p`1 <= E-bound C by A19,JORDAN8:15;
       then A166: p`1 <= (E-max C)`1 by PSCOMP_1:104;
         p`1 >= (E-max C)`1 by A161,JORDAN1A:def 3;
       then A167: p`1 = (E-max C)`1 by A166,AXIOMS:21;
         p`2 = (E-max C)`2 by A161,JORDAN1A:def 3;
       then p = E-max C by A167,TOPREAL3:11;
       hence contradiction by A5,A77,A121,A137,A162,XBOOLE_0:3;
      suppose east_halfline E-max C meets L~do;
       then consider p be set such that
        A168: p in east_halfline E-max C and
        A169: p in L~do by XBOOLE_0:3;
       reconsider p as Point of TOP-REAL 2 by A168;
         p in L~LS by A52,A169;
       then p in east_halfline E-max C /\ L~Cage(C,n)
                 by A134,A168,XBOOLE_0:def 3;
       then A170: p`1 = Ebo by A154,JORDAN1A:104;
       A171: (E-max C)`2 = p`2 by A168,JORDAN1A:def 3;
       set RC = Rotate(Cage(C,n),Emax);
       A172: E-max C in right_cell(RC,1) by JORDAN1I:9;
       A173: 1+1 <= len LS by A23,AXIOMS:22;
         LS = RC-:Wmin by JORDAN1G:26;
       then A174: LSeg(LS,1) = LSeg(RC,1) by A173,SPPOL_2:9;
       A175: L~RC = L~Cage(C,n) by REVROT_1:33;
       A176: len RC = len Cage(C,n) by REVROT_1:14;
       A177: GoB RC = GoB Cage(C,n) by REVROT_1:28
          .= Ga by JORDAN1H:52;
       A178: Emax in rng Cage(C,n) by SPRECT_2:50;
       A179: RC is_sequence_on Ga by A130,REVROT_1:34;
       A180: RC/.1 = E-max L~RC by A175,A178,FINSEQ_6:98;
       then consider ii,jj be Nat such that
        A181: [ii,jj+1] in Indices Ga and
        A182: [ii,jj] in Indices Ga and
        A183: RC/.1 = Ga*(ii,jj+1) and
        A184: RC/.(1+1) = Ga*(ii,jj) by A81,A176,A179,JORDAN1I:25;
       consider jj2 be Nat such that
        A185: 1 <= jj2 & jj2 <= width Ga and
        A186: Emax = Ga*(len Ga,jj2) by JORDAN1D:29;
       A187: len Ga >= 4 by JORDAN8:13;
       then len Ga >= 1 by AXIOMS:22;
       then [len Ga,jj2] in Indices Ga by A185,GOBOARD7:10;
       then A188: ii = len Ga by A175,A180,A181,A183,A186,GOBOARD1:21;
       A189: 1 <= ii & ii <= len Ga & 1 <= jj+1 & jj+1 <= width Ga
                                                           by A181,GOBOARD5:1;
       A190: 1 <= ii & ii <= len Ga & 1 <= jj & jj <= width Ga
                                                           by A182,GOBOARD5:1;
       A191: ii+1 <> ii by NAT_1:38;
         jj+1 > jj by NAT_1:38;
       then jj+1+1 <> jj by NAT_1:38;
       then A192: right_cell(RC,1) = cell(Ga,ii-'1,jj)
                          by A81,A176,A177,A181,A182,A183,A184,A191,GOBOARD5:
def 6;
       A193: ii-'1+1 = ii by A189,AMI_5:4;
       A194: ii-1 >= 4-1 by A187,A188,REAL_1:49;
       then A195: ii-1 >= 1 by AXIOMS:22;
         ii-1 >= 0 by A194,AXIOMS:22;
       then A196: 1 <= ii-'1 by A195,BINARITH:def 3;
       then A197: Ga*(ii-'1,jj)`2 <= p`2 & p`2 <= Ga*(ii-'1,jj+1)`2
                by A171,A172,A189,A190,A192,A193,JORDAN9:19;
       A198: ii-'1 < len Ga by A189,A193,NAT_1:38;
       then A199: Ga*(ii-'1,jj)`2 = Ga*(1,jj)`2 by A190,A196,GOBOARD5:2
          .= Ga*(ii,jj)`2 by A190,GOBOARD5:2;
       A200: Ga*(ii-'1,jj+1)`2 = Ga*(1,jj+1)`2 by A189,A196,A198,GOBOARD5:2
          .= Ga*(ii,jj+1)`2 by A189,GOBOARD5:2;
         Ga*(len Ga,jj)`1 = Ebo & Ebo = Ga*(len Ga,jj+1)`1
                                                by A8,A189,A190,JORDAN1A:92;
       then p in LSeg(RC/.1,RC/.(1+1))
                              by A170,A183,A184,A188,A197,A199,A200,GOBOARD7:8;
       then A201: p in LSeg(LS,1) by A81,A174,A176,TOPREAL1:def 5;
       A202: p in LSeg(do,Index(p,do)) by A169,JORDAN3:42;
       A203: do = mid(LS,Gij..LS,len LS) by A36,JORDAN1J:37;
       A204: 1<=Gij..LS & Gij..LS<=len LS by A36,FINSEQ_4:31;
         Gij..LS <> len LS by A28,A36,FINSEQ_4:29;
       then A205: Gij..LS < len LS by A204,REAL_1:def 5;
       A206: 1<=Index(p,do) & Index(p,do) < len do by A169,JORDAN3:41;
       A207: Index(Gij,LS)+1 = Gij..LS by A31,A36,JORDAN1J:56;
       consider t be Nat such that
        A208: t in dom LS and
        A209: LS.t = Gij by A36,FINSEQ_2:11;
       A210: 1 <= t & t <= len LS by A208,FINSEQ_3:27;
       then 1 < t by A31,A209,REAL_1:def 5;
       then Index(Gij,LS)+1 = t by A209,A210,JORDAN3:45;
       then A211: len L_Cut(LS,Gij) = len LS-Index(Gij,LS)
                                                  by A6,A28,A209,JORDAN3:61;
       set tt = Index(p,do)+(Gij..LS)-'1;
       A212: 1<=Index(Gij,LS) & 0+Index(Gij,LS) < len LS by A6,JORDAN3:41;
       then A213: len LS-Index(Gij,LS) > 0 by REAL_1:86;
       then Index(p,do) < len LS-'Index(Gij,LS) by A206,A211,BINARITH:def 3;
       then Index(p,do)+1 <= len LS-'Index(Gij,LS) by NAT_1:38;
       then Index(p,do) <= len LS-'Index(Gij,LS)-1 by REAL_1:84;
       then Index(p,do) <= len LS-Index(Gij,LS)-1 by A213,BINARITH:def 3;
       then A214: Index(p,do) <= len LS-Gij..LS by A207,XCMPLX_1:36;
       then len LS-Gij..LS >= 1 by A206,AXIOMS:22;
       then len LS-Gij..LS >= 0 by AXIOMS:22;
       then Index(p,do) <= len LS-'Gij..LS by A214,BINARITH:def 3;
       then Index(p,do) < len LS-'(Gij..LS)+1 by NAT_1:38;
       then A215: LSeg(mid(LS,Gij..LS,len LS),Index(p,do)) =
            LSeg(LS,Index(p,do)+(Gij..LS)-'1) by A204,A205,A206,JORDAN4:31;
       A216: 1+1 <= Gij..LS by A207,A212,REAL_1:55;
       then Index(p,do)+Gij..LS >= 1+1+1 by A206,REAL_1:55;
       then A217: Index(p,do)+Gij..LS-1 >= 1+1+1-1 by REAL_1:49;
       then A218: Index(p,do)+Gij..LS-1 >= 0 by AXIOMS:22;
       then A219: tt >= 1+1 by A217,BINARITH:def 3;
       A220: 2 in dom LS by A173,FINSEQ_3:27;
         now per cases by A219,REAL_1:def 5;
        suppose tt > 1+1;
         then LSeg(LS,1) misses LSeg(LS,tt) by TOPREAL1:def 9;
         hence contradiction by A201,A202,A203,A215,XBOOLE_0:3;
        suppose A221: tt = 1+1;
         then LSeg(LS,1) /\ LSeg(LS,tt) = {LS/.2} by A23,TOPREAL1:def 8;
         then p in {LS/.2} by A201,A202,A203,A215,XBOOLE_0:def 3;
         then A222: p = LS/.2 by TARSKI:def 1;
         then A223: p..LS = 2 by A220,FINSEQ_5:44;
           1+1 = Index(p,do)+(Gij..LS)-1 by A218,A221,BINARITH:def 3;
         then 1+1+1 = Index(p,do)+(Gij..LS) by XCMPLX_1:27;
         then A224: Gij..LS = 2 by A206,A216,JORDAN1E:10;
           p in rng LS by A220,A222,PARTFUN2:4;
         then p = Gij by A36,A223,A224,FINSEQ_5:10;
         then Gij`1 = Ebo by A222,JORDAN1G:40;
         then Gij`1 = Ga*(len Ga,j)`1 by A8,A10,JORDAN1A:92;
         hence contradiction by A1,A12,A63,JORDAN1G:7;
       end;
       hence contradiction;
    end;
    then east_halfline E-max C c= (L~godo)` by SUBSET_1:43;
    then consider W be Subset of TOP-REAL 2 such that
     A225: W is_a_component_of (L~godo)` and
     A226: east_halfline E-max C c= W by GOBOARD9:5;
      east_halfline E-max C is not Bounded by JORDAN1C:9;
    then W is not Bounded by A226,JORDAN2C:16;
    then W is_outside_component_of L~godo by A225,JORDAN2C:def 4;
    then W c= UBD L~godo by JORDAN2C:27;
    then A227: east_halfline E-max C c= UBD L~godo by A226,XBOOLE_1:1;
      E-max C in east_halfline E-max C by JORDAN1C:7;
    then E-max C in UBD L~godo by A227;
    then E-max C in LeftComp godo by GOBRD14:46;
    then UA meets L~godo by A120,A121,A132,A140,JORDAN1J:36;
    then A228: UA meets (L~go \/ L~pion1) or UA meets L~do by A133,XBOOLE_1:70;
    A229: UA c= C by JORDAN1A:16;
    per cases by A228,XBOOLE_1:70;
     suppose UA meets L~go;
      then UA meets L~Cage(C,n) by A135,XBOOLE_1:63;
      then C meets L~Cage(C,n) by A229,XBOOLE_1:63;
      hence contradiction by JORDAN10:5;
     suppose UA meets L~pion1;
      hence contradiction by A5,A77,A137;
     suppose UA meets L~do;
      then UA meets L~Cage(C,n) by A135,XBOOLE_1:63;
      then C meets L~Cage(C,n) by A229,XBOOLE_1:63;
      hence contradiction by JORDAN10:5;
   end;

  theorem Th32:
   for C be Simple_closed_curve
   for i,j,k be Nat st
    1 < j & j <= k & k < len Gauge(C,n) &
    1 <= i & i <= width Gauge(C,n) &
    Gauge(C,n)*(k,i) in L~Upper_Seq(C,n) &
    Gauge(C,n)*(j,i) in L~Lower_Seq(C,n) holds
     LSeg(Gauge(C,n)*(j,i),Gauge(C,n)*(k,i)) meets Lower_Arc C
   proof
    let C be Simple_closed_curve;
    let i,j,k be Nat;
    assume that
     A1: 1 < j & j <= k & k < len Gauge(C,n) and
     A2: 1 <= i & i <= width Gauge(C,n) and
     A3: Gauge(C,n)*(k,i) in L~Upper_Seq(C,n) and
     A4: Gauge(C,n)*(j,i) in L~Lower_Seq(C,n);
    consider j1,k1 be Nat such that
     A5: j <= j1 and
     A6: j1 <= k1 and
     A7: k1 <= k and
     A8: LSeg(Gauge(C,n)*(j1,i),Gauge(C,n)*(k1,i)) /\ L~Lower_Seq(C,n) =
      {Gauge(C,n)*(j1,i)} and
     A9: LSeg(Gauge(C,n)*(j1,i),Gauge(C,n)*(k1,i)) /\ L~Upper_Seq(C,n) =
      {Gauge(C,n)*(k1,i)} by A1,A2,A3,A4,Th16;
      1 < j1 & k1 < len Gauge(C,n) by A1,A5,A7,AXIOMS:22;
    then A10: LSeg(Gauge(C,n)*(j1,i),Gauge(C,n)*(k1,i)) meets Lower_Arc C
                                                   by A2,A6,A8,A9,Th30;
      LSeg(Gauge(C,n)*(j1,i),Gauge(C,n)*(k1,i)) c=
     LSeg(Gauge(C,n)*(j,i),Gauge(C,n)*(k,i)) by A1,A2,A5,A6,A7,Th8;
    hence LSeg(Gauge(C,n)*(j,i),Gauge(C,n)*(k,i)) meets Lower_Arc C
                                                          by A10,XBOOLE_1:63;
   end;

  theorem Th33:
   for C be Simple_closed_curve
   for i,j,k be Nat st
    1 < j & j <= k & k < len Gauge(C,n) &
    1 <= i & i <= width Gauge(C,n) &
    Gauge(C,n)*(k,i) in L~Upper_Seq(C,n) &
    Gauge(C,n)*(j,i) in L~Lower_Seq(C,n) holds
     LSeg(Gauge(C,n)*(j,i),Gauge(C,n)*(k,i)) meets Upper_Arc C
   proof
    let C be Simple_closed_curve;
    let i,j,k be Nat;
    assume that
     A1: 1 < j & j <= k & k < len Gauge(C,n) and
     A2: 1 <= i & i <= width Gauge(C,n) and
     A3: Gauge(C,n)*(k,i) in L~Upper_Seq(C,n) and
     A4: Gauge(C,n)*(j,i) in L~Lower_Seq(C,n);
    consider j1,k1 be Nat such that
     A5: j <= j1 and
     A6: j1 <= k1 and
     A7: k1 <= k and
     A8: LSeg(Gauge(C,n)*(j1,i),Gauge(C,n)*(k1,i)) /\ L~Lower_Seq(C,n) =
      {Gauge(C,n)*(j1,i)} and
     A9: LSeg(Gauge(C,n)*(j1,i),Gauge(C,n)*(k1,i)) /\ L~Upper_Seq(C,n) =
      {Gauge(C,n)*(k1,i)} by A1,A2,A3,A4,Th16;
      1 < j1 & k1 < len Gauge(C,n) by A1,A5,A7,AXIOMS:22;
    then A10: LSeg(Gauge(C,n)*(j1,i),Gauge(C,n)*(k1,i)) meets Upper_Arc C
                                                   by A2,A6,A8,A9,Th31;
      LSeg(Gauge(C,n)*(j1,i),Gauge(C,n)*(k1,i)) c=
     LSeg(Gauge(C,n)*(j,i),Gauge(C,n)*(k,i)) by A1,A2,A5,A6,A7,Th8;
    hence LSeg(Gauge(C,n)*(j,i),Gauge(C,n)*(k,i)) meets Upper_Arc C
                                                          by A10,XBOOLE_1:63;
   end;

  theorem Th34:
   for C be Simple_closed_curve
   for i,j,k be Nat st
    1 < j & j <= k & k < len Gauge(C,n) &
    1 <= i & i <= width Gauge(C,n) & n > 0 &
    Gauge(C,n)*(k,i) in Upper_Arc L~Cage(C,n) &
    Gauge(C,n)*(j,i) in Lower_Arc L~Cage(C,n) holds
     LSeg(Gauge(C,n)*(j,i),Gauge(C,n)*(k,i)) meets Lower_Arc C
   proof
    let C be Simple_closed_curve;
    let i,j,k be Nat;
    assume that
     A1: 1 < j & j <= k & k < len Gauge(C,n) and
     A2: 1 <= i & i <= width Gauge(C,n) and
     A3: n > 0 and
     A4: Gauge(C,n)*(k,i) in Upper_Arc L~Cage(C,n) and
     A5: Gauge(C,n)*(j,i) in Lower_Arc L~Cage(C,n);
    A6: L~Upper_Seq(C,n) = Upper_Arc L~Cage(C,n) by A3,JORDAN1G:63;
      L~Lower_Seq(C,n) = Lower_Arc L~Cage(C,n) by A3,JORDAN1G:64;
    hence LSeg(Gauge(C,n)*(j,i),Gauge(C,n)*(k,i)) meets Lower_Arc C
                                                   by A1,A2,A4,A5,A6,Th32;
   end;

  theorem Th35:
   for C be Simple_closed_curve
   for i,j,k be Nat st
    1 < j & j <= k & k < len Gauge(C,n) &
    1 <= i & i <= width Gauge(C,n) & n > 0 &
    Gauge(C,n)*(k,i) in Upper_Arc L~Cage(C,n) &
    Gauge(C,n)*(j,i) in Lower_Arc L~Cage(C,n) holds
     LSeg(Gauge(C,n)*(j,i),Gauge(C,n)*(k,i)) meets Upper_Arc C
   proof
    let C be Simple_closed_curve;
    let i,j,k be Nat;
    assume that
     A1: 1 < j & j <= k & k < len Gauge(C,n) and
     A2: 1 <= i & i <= width Gauge(C,n) and
     A3: n > 0 and
     A4: Gauge(C,n)*(k,i) in Upper_Arc L~Cage(C,n) and
     A5: Gauge(C,n)*(j,i) in Lower_Arc L~Cage(C,n);
    A6: L~Upper_Seq(C,n) = Upper_Arc L~Cage(C,n) by A3,JORDAN1G:63;
      L~Lower_Seq(C,n) = Lower_Arc L~Cage(C,n) by A3,JORDAN1G:64;
    hence LSeg(Gauge(C,n)*(j,i),Gauge(C,n)*(k,i)) meets Upper_Arc C
                                                   by A1,A2,A4,A5,A6,Th33;
   end;

  theorem
     for C be Simple_closed_curve
   for j,k be Nat holds
    1 < j & j <= k & k < len Gauge(C,n+1) &
    Gauge(C,n+1)*(k,Center Gauge(C,n+1)) in Upper_Arc L~Cage(C,n+1) &
    Gauge(C,n+1)*(j,Center Gauge(C,n+1)) in Lower_Arc L~Cage(C,n+1) implies
     LSeg(Gauge(C,n+1)*(j,Center Gauge(C,n+1)),
          Gauge(C,n+1)*(k,Center Gauge(C,n+1))) meets Lower_Arc C
   proof
    let C be Simple_closed_curve;
    let j,k be Nat;
    assume that
     A1: 1 < j & j <= k & k < len Gauge(C,n+1) and
     A2: Gauge(C,n+1)*(k,Center Gauge(C,n+1)) in Upper_Arc L~Cage(C,n+1) and
     A3: Gauge(C,n+1)*(j,Center Gauge(C,n+1)) in Lower_Arc L~Cage(C,n+1);
      n+1 >= 0+1 by NAT_1:29;
    then A4: n+1 > 0 by NAT_1:38;
    A5: len Gauge(C,n+1) >= 4 by JORDAN8:13;
    then A6: len Gauge(C,n+1) >= 3 by AXIOMS:22;
      len Gauge(C,n+1) >= 2 by A5,AXIOMS:22;
    then A7: 1 < Center Gauge(C,n+1) by JORDAN1B:15;
      Center Gauge(C,n+1) < len Gauge(C,n+1) by A6,JORDAN1B:16;
    then Center Gauge(C,n+1) < width Gauge(C,n+1) by JORDAN8:def 1;
    hence LSeg(Gauge(C,n+1)*(j,Center Gauge(C,n+1)),
          Gauge(C,n+1)*(k,Center Gauge(C,n+1))) meets Lower_Arc C
                                                   by A1,A2,A3,A4,A7,Th34;
   end;

  theorem
     for C be Simple_closed_curve
   for j,k be Nat holds
    1 < j & j <= k & k < len Gauge(C,n+1) &
    Gauge(C,n+1)*(k,Center Gauge(C,n+1)) in Upper_Arc L~Cage(C,n+1) &
    Gauge(C,n+1)*(j,Center Gauge(C,n+1)) in Lower_Arc L~Cage(C,n+1) implies
     LSeg(Gauge(C,n+1)*(j,Center Gauge(C,n+1)),
          Gauge(C,n+1)*(k,Center Gauge(C,n+1))) meets Upper_Arc C
   proof
    let C be Simple_closed_curve;
    let j,k be Nat;
    assume that
     A1: 1 < j & j <= k & k < len Gauge(C,n+1) and
     A2: Gauge(C,n+1)*(k,Center Gauge(C,n+1)) in Upper_Arc L~Cage(C,n+1) and
     A3: Gauge(C,n+1)*(j,Center Gauge(C,n+1)) in Lower_Arc L~Cage(C,n+1);
      n+1 >= 0+1 by NAT_1:29;
    then A4: n+1 > 0 by NAT_1:38;
    A5: len Gauge(C,n+1) >= 4 by JORDAN8:13;
    then A6: len Gauge(C,n+1) >= 3 by AXIOMS:22;
      len Gauge(C,n+1) >= 2 by A5,AXIOMS:22;
    then A7: 1 < Center Gauge(C,n+1) by JORDAN1B:15;
      Center Gauge(C,n+1) < len Gauge(C,n+1) by A6,JORDAN1B:16;
    then Center Gauge(C,n+1) < width Gauge(C,n+1) by JORDAN8:def 1;
    hence LSeg(Gauge(C,n+1)*(j,Center Gauge(C,n+1)),
          Gauge(C,n+1)*(k,Center Gauge(C,n+1))) meets Upper_Arc C
                                                   by A1,A2,A3,A4,A7,Th35;
   end;

  theorem Th38:
   for C be Simple_closed_curve
   for i,j,k be Nat st
    1 < j & j <= k & k < len Gauge(C,n) &
    1 <= i & i <= width Gauge(C,n) &
    LSeg(Gauge(C,n)*(j,i),Gauge(C,n)*(k,i)) /\ L~Upper_Seq(C,n) =
     {Gauge(C,n)*(j,i)} &
    LSeg(Gauge(C,n)*(j,i),Gauge(C,n)*(k,i)) /\ L~Lower_Seq(C,n) =
     {Gauge(C,n)*(k,i)} holds
      LSeg(Gauge(C,n)*(j,i),Gauge(C,n)*(k,i)) meets Lower_Arc C
   proof
    let C be Simple_closed_curve;
    let i,j,k be Nat;
    set Ga = Gauge(C,n);
    set US = Upper_Seq(C,n);
    set LS = Lower_Seq(C,n);
    set LA = Lower_Arc C;
    set Wmin = W-min L~Cage(C,n);
    set Emax = E-max L~Cage(C,n);
    set Wbo = W-bound L~Cage(C,n);
    set Ebo = E-bound L~Cage(C,n);
    set Gij = Ga*(j,i);
    set Gik = Ga*(k,i);
    assume that
     A1: 1 < j & j <= k & k < len Ga and
     A2: 1 <= i & i <= width Ga and
     A3: LSeg(Gij,Gik) /\ L~US = {Gij} and
     A4: LSeg(Gij,Gik) /\ L~LS = {Gik} and
     A5: LSeg(Gij,Gik) misses LA;
      Gik in {Gik} by TARSKI:def 1;
    then A6: Gik in L~LS by A4,XBOOLE_0:def 3;
      Gij in {Gij} by TARSKI:def 1;
    then A7: Gij in L~US by A3,XBOOLE_0:def 3;
    A8: len Ga = width Ga by JORDAN8:def 1;
    A9: j <> k by A1,A2,A6,A7,Th29;
    A10: 1 < j & j < width Ga by A1,A8,AXIOMS:22;
    A11: 1 < k & k < width Ga by A1,AXIOMS:22,JORDAN8:def 1;
    A12: [j,i] in Indices Ga by A2,A8,A10,GOBOARD7:10;
    A13: [k,i] in Indices Ga by A2,A8,A11,GOBOARD7:10;
    A14: US is_sequence_on Ga by JORDAN1G:4;
    A15: LS is_sequence_on Ga by JORDAN1G:5;
    set go = R_Cut(US,Gij);
    set do = L_Cut(LS,Gik);
    A16: len US >= 3 by JORDAN1E:19;
    then len US >= 1 by AXIOMS:22;
    then 1 in dom US by FINSEQ_3:27;
    then A17: US.1 = US/.1 by FINSEQ_4:def 4
       .= Wmin by JORDAN1F:5;
    A18: Wmin`1 = Wbo by PSCOMP_1:84
       .= Ga*(1,i)`1 by A2,A8,JORDAN1A:94;
      len Ga >= 4 by JORDAN8:13;
    then A19: len Ga >= 1 by AXIOMS:22;
    then A20: [1,k] in Indices Ga by A11,GOBOARD7:10;
    A21: [1,i] in Indices Ga by A2,A19,GOBOARD7:10;
    then A22: Gij <> US.1 by A1,A12,A17,A18,JORDAN1G:7;
    then reconsider go as being_S-Seq FinSequence of TOP-REAL 2
                                                            by A7,JORDAN3:70;
    A23: len LS >= 1+2 by JORDAN1E:19;
    then len LS >= 1 by AXIOMS:22;
    then A24: 1 in dom LS & len LS in dom LS by FINSEQ_3:27;
    then A25: LS.len LS = LS/.len LS by FINSEQ_4:def 4
       .= Wmin by JORDAN1F:8;
      Wmin`1 = Wbo by PSCOMP_1:84
       .= Ga*(1,i)`1 by A2,A8,JORDAN1A:94;
    then A26: Gik <> LS.len LS by A1,A13,A21,A25,JORDAN1G:7;
    then reconsider do as being_S-Seq FinSequence of TOP-REAL 2
                                                            by A6,JORDAN3:69;
    A27: [len Ga,k] in Indices Ga by A11,A19,GOBOARD7:10;
    A28: LS.1 = LS/.1 by A24,FINSEQ_4:def 4
       .= Emax by JORDAN1F:6;
      Emax`1 = Ebo by PSCOMP_1:104
       .= Ga*(len Ga,k)`1 by A8,A11,JORDAN1A:92;
    then A29: Gik <> LS.1 by A1,A13,A27,A28,JORDAN1G:7;
    A30: len go >= 1+1 by TOPREAL1:def 10;
    A31: Gij in rng US by A2,A7,A8,A10,A14,JORDAN1J:40;
    then A32: go is_sequence_on Ga by A14,JORDAN1J:38;
    A33: len do >= 1+1 by TOPREAL1:def 10;
    A34: Gik in rng LS by A2,A6,A8,A11,A15,JORDAN1J:40;
    then A35: do is_sequence_on Ga by A15,JORDAN1J:39;
    reconsider go as non constant s.c.c.
            (being_S-Seq FinSequence of TOP-REAL 2) by A30,A32,JGRAPH_1:16,
JORDAN8:8;
    reconsider do as non constant s.c.c.
            (being_S-Seq FinSequence of TOP-REAL 2) by A33,A35,JGRAPH_1:16,
JORDAN8:8;
    A36: len go > 1 by A30,NAT_1:38;
    then A37: len go in dom go by FINSEQ_3:27;
    then A38: go/.len go = go.len go by FINSEQ_4:def 4
       .= Gij by A7,JORDAN3:59;
      len do >= 1 by A33,AXIOMS:22;
    then 1 in dom do by FINSEQ_3:27;
    then A39: do/.1 = do.1 by FINSEQ_4:def 4
       .= Gik by A6,JORDAN3:58;
    reconsider m = len go - 1 as Nat by A37,FINSEQ_3:28;
    A40: m+1 = len go by XCMPLX_1:27;
    then A41: len go-'1 = m by BINARITH:39;
    A42: LSeg(go,m) c= L~go by TOPREAL3:26;
    A43: L~go c= L~US by A7,JORDAN3:76;
    then LSeg(go,m) c= L~US by A42,XBOOLE_1:1;
    then A44: LSeg(go,m) /\ LSeg(Gik,Gij) c= {Gij} by A3,XBOOLE_1:26;
      m >= 1 by A30,REAL_1:84;
    then A45: LSeg(go,m) = LSeg(go/.m,Gij) by A38,A40,TOPREAL1:def 5;
      {Gij} c= LSeg(go,m) /\ LSeg(Gik,Gij)
    proof
     let x be set;
     assume x in {Gij};
     then A46: x = Gij by TARSKI:def 1;
     A47: Gij in LSeg(go,m) by A45,TOPREAL1:6;
       Gij in LSeg(Gik,Gij) by TOPREAL1:6;
     hence x in LSeg(go,m) /\ LSeg(Gik,Gij) by A46,A47,XBOOLE_0:def 3;
    end;
    then A48: LSeg(go,m) /\ LSeg(Gik,Gij) = {Gij} by A44,XBOOLE_0:def 10;
    A49: LSeg(do,1) c= L~do by TOPREAL3:26;
    A50: L~do c= L~LS by A6,JORDAN3:77;
    then LSeg(do,1) c= L~LS by A49,XBOOLE_1:1;
    then A51: LSeg(do,1) /\ LSeg(Gik,Gij) c= {Gik} by A4,XBOOLE_1:26;
    A52: LSeg(do,1) = LSeg(Gik,do/.(1+1)) by A33,A39,TOPREAL1:def 5;
      {Gik} c= LSeg(do,1) /\ LSeg(Gik,Gij)
    proof
     let x be set;
     assume x in {Gik};
     then A53: x = Gik by TARSKI:def 1;
     A54: Gik in LSeg(do,1) by A52,TOPREAL1:6;
       Gik in LSeg(Gik,Gij) by TOPREAL1:6;
     hence x in LSeg(do,1) /\ LSeg(Gik,Gij) by A53,A54,XBOOLE_0:def 3;
    end;
    then A55: LSeg(Gik,Gij) /\ LSeg(do,1) = {Gik} by A51,XBOOLE_0:def 10;
    A56: go/.1 = US/.1 by A7,SPRECT_3:39
       .= Wmin by JORDAN1F:5;
    then A57: go/.1 = LS/.len LS by JORDAN1F:8
       .= do/.len do by A6,JORDAN1J:35;
    A58: rng go c= L~go & rng do c= L~do by A30,A33,SPPOL_2:18;
    A59: {go/.1} c= L~go /\ L~do
    proof
     let x be set;
     assume x in {go/.1};
     then x = go/.1 by TARSKI:def 1;
     then x in rng go & x in rng do by A57,FINSEQ_6:46,REVROT_1:3;
     hence x in L~go /\ L~do by A58,XBOOLE_0:def 3;
    end;
    A60: LS.1 = LS/.1 by A24,FINSEQ_4:def 4
       .= Emax by JORDAN1F:6;
    A61: [len Ga,j] in Indices Ga by A10,A19,GOBOARD7:10;
      L~go /\ L~do c= {go/.1}
    proof
     let x be set;
     assume x in L~go /\ L~do;
     then A62: x in L~go & x in L~do by XBOOLE_0:def 3;
     then x in L~US /\ L~LS by A43,A50,XBOOLE_0:def 3;
     then x in {Wmin,Emax} by JORDAN1E:20;
     then A63: x = Wmin or x = Emax by TARSKI:def 2;
       now assume x = Emax;
      then A64: Emax = Gik by A6,A60,A62,JORDAN1E:11;
        Ga*(len Ga,j)`1 = Ebo by A8,A10,JORDAN1A:92;
      then Emax`1 <> Ebo by A1,A13,A61,A64,JORDAN1G:7;
      hence contradiction by PSCOMP_1:104;
     end;
     hence x in {go/.1} by A56,A63,TARSKI:def 1;
    end;
    then A65: L~go /\ L~do = {go/.1} by A59,XBOOLE_0:def 10;
    set W2 = go/.2;
    A66: 2 in dom go by A30,FINSEQ_3:27;
    A67: Gij..US >= 1 by A31,FINSEQ_4:31;
    A68: now assume Gij`1 = Wbo;
     then Ga*(1,k)`1 = Ga*(j,i)`1 by A8,A11,JORDAN1A:94;
     hence contradiction by A1,A12,A20,JORDAN1G:7;
    end;
      go = mid(US,1,Gij..US) by A31,JORDAN1G:57
       .= US|(Gij..US) by A67,JORDAN3:25;
    then A69: W2 = US/.2 by A66,TOPREAL1:1;
    A70: Wmin in rng go by A56,FINSEQ_6:46;
    set pion = <*Gij,Gik*>;
    A71: now let n be Nat;
     assume n in dom pion;
     then n in {1,2} by FINSEQ_1:4,FINSEQ_3:29;
     then n = 1 or n = 2 by TARSKI:def 2;
     then pion/.n = Gik or pion/.n = Gij by FINSEQ_4:26;
     hence ex j,i be Nat st [j,i] in Indices Ga & pion/.n = Ga*(j,i)
                                                                  by A12,A13;
    end;
    A72: Gik <> Gij by A9,A12,A13,GOBOARD1:21;
      Gik`2 = Ga*(1,i)`2 by A2,A8,A11,GOBOARD5:2
       .= Gij`2 by A2,A8,A10,GOBOARD5:2;
    then LSeg(Gik,Gij) is horizontal by SPPOL_1:36;
    then pion is_S-Seq by A72,JORDAN1B:9;
    then consider pion1 be FinSequence of TOP-REAL 2 such that
     A73: pion1 is_sequence_on Ga and
     A74: pion1 is_S-Seq and
     A75: L~pion = L~pion1 and
     A76: pion/.1 = pion1/.1 and
     A77: pion/.len pion = pion1/.len pion1 and
     A78: len pion <= len pion1 by A71,GOBOARD3:2;
    reconsider pion1 as being_S-Seq FinSequence of TOP-REAL 2 by A74;
    set godo = go^'pion1^'do;
      len Cage(C,n) > 4 by GOBOARD7:36;
    then A79: 1+1 <= len Cage(C,n) by AXIOMS:22;
    then A80: 1+1 <= len Rotate(Cage(C,n),Wmin) by REVROT_1:14;
      len (go^'pion1) >= len go by TOPREAL8:7;
    then A81: len (go^'pion1) >= 1+1 by A30,AXIOMS:22;
    then A82: len (go^'pion1) > 1+0 by NAT_1:38;
      len godo >= len (go^'pion1) by TOPREAL8:7;
    then A83: 1+1 <= len godo by A81,AXIOMS:22;
    A84: US is_sequence_on Ga by JORDAN1G:4;
    A85: go/.len go = pion1/.1 by A38,A76,FINSEQ_4:26;
    then A86: go^'pion1 is_sequence_on Ga by A32,A73,TOPREAL8:12;
    A87: (go^'pion1)/.len (go^'pion1) = pion/.len pion by A77,AMISTD_1:6
       .= pion/.2 by FINSEQ_1:61
       .= do/.1 by A39,FINSEQ_4:26;
    then A88: godo is_sequence_on Ga by A35,A86,TOPREAL8:12;
      LSeg(pion1,1) c= L~<*Gij,Gik*> by A75,TOPREAL3:26;
    then LSeg(pion1,1) c= LSeg(Gij,Gik) by SPPOL_2:21;
    then A89: LSeg(go,len go-'1) /\ LSeg(pion1,1) c= {Gij}
                                                      by A41,A48,XBOOLE_1:27;
    A90: len pion1 >= 1+1 by A78,FINSEQ_1:61;
      {Gij} c= LSeg(go,m) /\ LSeg(pion1,1)
    proof
     let x be set;
     assume x in {Gij};
     then A91: x = Gij by TARSKI:def 1;
     A92: Gij in LSeg(go,m) by A45,TOPREAL1:6;
       Gij in LSeg(pion1,1) by A38,A85,A90,TOPREAL1:27;
     hence x in LSeg(go,m) /\ LSeg(pion1,1) by A91,A92,XBOOLE_0:def 3;
    end;
    then LSeg(go,len go-'1) /\ LSeg(pion1,1) = { go/.len go }
                                             by A38,A41,A89,XBOOLE_0:def 10;
    then A93: go^'pion1 is unfolded by A85,TOPREAL8:34;
      len pion1 >= 2+0 by A78,FINSEQ_1:61;
    then A94: len pion1-2 >= 0 by REAL_1:84;
    A95: len (go^'pion1)-1 >= 0 by A82,REAL_1:84;
      len (go^'pion1)+1-1 = len go+len pion1-1 by GRAPH_2:13;
    then len (go^'pion1)-1 = len go+len pion1-1-1 by XCMPLX_1:26
       .= len go + len pion1-(1+1) by XCMPLX_1:36
       .= len go + (len pion1-2) by XCMPLX_1:29
       .= len go + (len pion1-'2) by A94,BINARITH:def 3;
    then A96: len (go^'pion1)-'1 = len go + (len pion1-'2)
                                                      by A95,BINARITH:def 3;
    A97: len pion1-1 >= 1 by A90,REAL_1:84;
    then A98: len pion1-1 >= 0 by AXIOMS:22;
    then A99: len pion1-'1 = len pion1-1 by BINARITH:def 3;
    A100: len pion1-'2+1 = len pion1-2+1 by A94,BINARITH:def 3
       .= len pion1-(2-1) by XCMPLX_1:37
       .= len pion1-'1 by A98,BINARITH:def 3;
      len pion1-1+1 <= len pion1 by XCMPLX_1:27;
    then A101: len pion1-'1 < len pion1 by A99,NAT_1:38;
      LSeg(pion1,len pion1-'1) c= L~<*Gij,Gik*> by A75,TOPREAL3:26;
    then LSeg(pion1,len pion1-'1) c= LSeg(Gij,Gik) by SPPOL_2:21;
    then A102: LSeg(pion1,len pion1-'1) /\ LSeg(do,1) c= {Gik}
                                                          by A55,XBOOLE_1:27;
      {Gik} c= LSeg(pion1,len pion1-'1) /\ LSeg(do,1)
    proof
     let x be set;
     assume x in {Gik};
     then A103: x = Gik by TARSKI:def 1;
     A104: Gik in LSeg(do,1) by A52,TOPREAL1:6;
     A105: len pion1-'1+1 = len pion1 by A99,XCMPLX_1:27;
     then pion1/.(len pion1-'1+1) = pion/.2 by A77,FINSEQ_1:61
        .= Gik by FINSEQ_4:26;
     then Gik in LSeg(pion1,len pion1-'1) by A97,A99,A105,TOPREAL1:27;
     hence x in LSeg(pion1,len pion1-'1) /\ LSeg(do,1)
                                                  by A103,A104,XBOOLE_0:def 3;
    end;
    then LSeg(pion1,len pion1-'1) /\ LSeg(do,1) = {Gik} by A102,XBOOLE_0:def 10
;
    then A106: LSeg(go^'pion1,len go+(len pion1-'2)) /\ LSeg(do,1) =
       {(go^'pion1)/.len (go^'pion1)} by A39,A85,A87,A100,A101,TOPREAL8:31;
    A107: (go^'pion1) is non trivial by A81,SPPOL_1:2;
    A108: rng pion1 c= L~pion1 by A90,SPPOL_2:18;
    A109: {pion1/.1} c= L~go /\ L~pion1
    proof
     let x be set;
     assume x in {pion1/.1};
     then x = pion1/.1 by TARSKI:def 1;
     then x in rng go & x in rng pion1 by A85,FINSEQ_6:46,REVROT_1:3;
     hence x in L~go /\ L~pion1 by A58,A108,XBOOLE_0:def 3;
    end;
      L~go /\ L~pion1 c= {pion1/.1}
    proof
     let x be set;
     assume x in L~go /\ L~pion1;
     then x in L~go & x in L~pion1 by XBOOLE_0:def 3;
     then x in L~pion1 /\ L~US by A43,XBOOLE_0:def 3;
     hence x in {pion1/.1} by A3,A38,A75,A85,SPPOL_2:21;
    end;
    then A110: L~go /\ L~pion1 = {pion1/.1} by A109,XBOOLE_0:def 10;
    then A111: (go^'pion1) is s.n.c. by A85,JORDAN1J:54;
      rng go /\ rng pion1 c= {pion1/.1} by A58,A108,A110,XBOOLE_1:27;
    then A112: go^'pion1 is one-to-one by JORDAN1J:55;
    A113: pion/.len pion = pion/.2 by FINSEQ_1:61
       .= do/.1 by A39,FINSEQ_4:26;
    A114: {pion1/.len pion1} c= L~do /\ L~pion1
    proof
     let x be set;
     assume x in {pion1/.len pion1};
     then x = pion1/.len pion1 by TARSKI:def 1;
     then x in rng do & x in rng pion1 by A77,A113,FINSEQ_6:46,REVROT_1:3;
     hence x in L~do /\ L~pion1 by A58,A108,XBOOLE_0:def 3;
    end;
      L~do /\ L~pion1 c= {pion1/.len pion1}
    proof
     let x be set;
     assume x in L~do /\ L~pion1;
     then x in L~do & x in L~pion1 by XBOOLE_0:def 3;
     then x in L~pion1 /\ L~LS by A50,XBOOLE_0:def 3;
     hence x in {pion1/.len pion1} by A4,A39,A75,A77,A113,SPPOL_2:21;
    end;
    then A115: L~do /\ L~pion1 = {pion1/.len pion1} by A114,XBOOLE_0:def 10;
    A116: L~(go^'pion1) /\ L~do = (L~go \/ L~pion1) /\ L~do by A85,TOPREAL8:35
       .= {go/.1} \/ {do/.1} by A65,A77,A113,A115,XBOOLE_1:23
       .= {(go^'pion1)/.1} \/ {do/.1} by AMISTD_1:5
       .= {(go^'pion1)/.1,do/.1} by ENUMSET1:41;
      do/.len do = (go^'pion1)/.1 by A57,AMISTD_1:5;
    then reconsider godo as non constant standard special_circular_sequence
          by A83,A87,A88,A93,A96,A106,A107,A111,A112,A116,JORDAN8:7,8,TOPREAL8:
11,33,34;
    A117: LA is_an_arc_of E-max C,W-min C by JORDAN6:def 9;
    then A118: LA is connected by JORDAN6:11;
    A119: W-min C in LA & E-max C in LA by A117,TOPREAL1:4;
    set ff = Rotate(Cage(C,n),Wmin);
      Wmin in rng Cage(C,n) by SPRECT_2:47;
    then A120: ff/.1 = Wmin by FINSEQ_6:98;
    A121: L~ff = L~Cage(C,n) by REVROT_1:33;
    then A122: (W-max L~ff)..ff > 1 by A120,SPRECT_5:23;
      (W-max L~ff)..ff <= (N-min L~ff)..ff by A120,A121,SPRECT_5:24;
    then A123: (N-min L~ff)..ff > 1 by A122,AXIOMS:22;
      (N-min L~ff)..ff < (N-max L~ff)..ff by A120,A121,SPRECT_5:25;
    then A124: (N-max L~ff)..ff > 1 by A123,AXIOMS:22;
      (N-max L~ff)..ff <= (E-max L~ff)..ff by A120,A121,SPRECT_5:26;
    then A125: Emax..ff > 1 by A121,A124,AXIOMS:22;
    A126: now assume A127: Gij..US <= 1;
       Gij..US >= 1 by A31,FINSEQ_4:31;
     then Gij..US = 1 by A127,AXIOMS:21;
     then Gij = US/.1 by A31,FINSEQ_5:41;
     hence contradiction by A17,A22,JORDAN1F:5;
    end;
    A128: Cage(C,n) is_sequence_on Ga by JORDAN9:def 1;
    then A129: ff is_sequence_on Ga by REVROT_1:34;
    A130: right_cell(godo,1,Ga)\L~godo c= RightComp godo by A83,A88,JORDAN9:29;
    A131: L~godo = L~(go^'pion1) \/ L~do by A87,TOPREAL8:35
       .= L~go \/ L~pion1 \/ L~do by A85,TOPREAL8:35;
      L~Cage(C,n) = L~US \/ L~LS by JORDAN1E:17;
    then A132: L~US c= L~Cage(C,n) & L~LS c= L~Cage(C,n) by XBOOLE_1:7;
    then A133: L~go c= L~Cage(C,n) & L~do c= L~Cage(C,n)
                                                     by A43,A50,XBOOLE_1:1;
    A134: W-min C in C by SPRECT_1:15;
    A135: L~pion = LSeg(Gik,Gij) by SPPOL_2:21;
    A136: now assume W-min C in L~godo;
     then A137: W-min C in L~go \/ L~pion1 or W-min C in L~do by A131,XBOOLE_0:
def 2;
     per cases by A137,XBOOLE_0:def 2;
      suppose W-min C in L~go;
       then C meets L~Cage(C,n) by A133,A134,XBOOLE_0:3;
       hence contradiction by JORDAN10:5;
      suppose W-min C in L~pion1;
       hence contradiction by A5,A75,A119,A135,XBOOLE_0:3;
      suppose W-min C in L~do;
       then C meets L~Cage(C,n) by A133,A134,XBOOLE_0:3;
       hence contradiction by JORDAN10:5;
    end;
      right_cell(Rotate(Cage(C,n),Wmin),1) =
          right_cell(ff,1,GoB ff) by A80,JORDAN1H:29
       .= right_cell(ff,1,GoB Cage(C,n)) by REVROT_1:28
       .= right_cell(ff,1,Ga) by JORDAN1H:52
       .= right_cell(ff-:Emax,1,Ga) by A125,A129,JORDAN1J:53
       .= right_cell(US,1,Ga) by JORDAN1E:def 1
       .= right_cell(R_Cut(US,Gij),1,Ga) by A31,A84,A126,JORDAN1J:52
       .= right_cell(go^'pion1,1,Ga) by A36,A86,JORDAN1J:51
       .= right_cell(godo,1,Ga) by A82,A88,JORDAN1J:51;
    then W-min C in right_cell(godo,1,Ga) by JORDAN1I:8;
    then A138: W-min C in right_cell(godo,1,Ga)\L~godo by A136,XBOOLE_0:def 4;
    A139: godo/.1 = (go^'pion1)/.1 by AMISTD_1:5
       .= Wmin by A56,AMISTD_1:5;
    A140: len US >= 2 by A16,AXIOMS:22;
    A141: godo/.2 = (go^'pion1)/.2 by A81,AMISTD_1:9
       .= US/.2 by A30,A69,AMISTD_1:9
       .= (US^'LS)/.2 by A140,AMISTD_1:9
       .= Rotate(Cage(C,n),Wmin)/.2 by JORDAN1E:15;
    A142: L~go \/ L~do is compact by COMPTS_1:19;
    A143: L~go \/ L~do c= L~Cage(C,n) by A133,XBOOLE_1:8;
      Wmin in L~go \/ L~do by A58,A70,XBOOLE_0:def 2;
    then A144: W-min (L~go \/ L~do) = Wmin by A142,A143,JORDAN1J:21;
    A145: (W-min (L~go \/ L~do))`1 = W-bound (L~go \/ L~do) & Wmin`1 = Wbo
                                                              by PSCOMP_1:84;
    A146: Gij`1 <= Gik`1 by A1,A2,SPRECT_3:25;
    then W-bound LSeg(Gik,Gij) = Gij`1 by SPRECT_1:62;
    then A147: W-bound L~pion1 = Gij`1 by A75,SPPOL_2:21;
      Gij`1 >= Wbo by A7,A132,PSCOMP_1:71;
    then Gij`1 > Wbo by A68,REAL_1:def 5;
    then W-min (L~go\/L~do\/L~pion1) = W-min (L~go \/ L~do)
                                              by A142,A144,A145,A147,JORDAN1J:
33;
    then A148: W-min L~godo = Wmin by A131,A144,XBOOLE_1:4;
    A149: rng godo c= L~godo by A83,SPPOL_2:18;
      2 in dom godo by A83,FINSEQ_3:27;
    then A150: godo/.2 in rng godo by PARTFUN2:4;
      godo/.2 in W-most L~Cage(C,n) by A141,JORDAN1I:27;
    then (godo/.2)`1 = (W-min L~godo)`1 by A148,PSCOMP_1:88
       .= W-bound L~godo by PSCOMP_1:84;
    then godo/.2 in W-most L~godo by A149,A150,SPRECT_2:16;
    then Rotate(godo,W-min L~godo)/.2 in W-most L~godo by A139,A148,FINSEQ_6:95
;
    then reconsider godo as clockwise_oriented non constant standard
                                    special_circular_sequence by JORDAN1I:27;
      len US in dom US by FINSEQ_5:6;
    then A151: US.len US = US/.len US by FINSEQ_4:def 4
       .= Emax by JORDAN1F:7;
    A152: E-max C in E-most C by PSCOMP_1:111;
    A153: east_halfline E-max C misses L~go
    proof
     assume east_halfline E-max C meets L~go;
     then consider p be set such that
      A154: p in east_halfline E-max C and
      A155: p in L~go by XBOOLE_0:3;
     reconsider p as Point of TOP-REAL 2 by A154;
       p in L~US by A43,A155;
     then p in east_halfline E-max C /\ L~Cage(C,n) by A132,A154,XBOOLE_0:def 3
;
     then A156: p`1 = Ebo by A152,JORDAN1A:104;
     then A157: p = Emax by A43,A155,JORDAN1J:46;
     then Emax = Gij by A7,A151,A155,JORDAN1J:43;
     then Gij`1 = Ga*(len Ga,k)`1 by A8,A11,A156,A157,JORDAN1A:92;
     hence contradiction by A1,A12,A27,JORDAN1G:7;
    end;
      now assume east_halfline E-max C meets L~godo;
     then A158: east_halfline E-max C meets (L~go \/ L~pion1) or
               east_halfline E-max C meets L~do by A131,XBOOLE_1:70;
     per cases by A158,XBOOLE_1:70;
      suppose east_halfline E-max C meets L~go;
       hence contradiction by A153;
      suppose east_halfline E-max C meets L~pion1;
       then consider p be set such that
        A159: p in east_halfline E-max C and
        A160: p in L~pion1 by XBOOLE_0:3;
       reconsider p as Point of TOP-REAL 2 by A159;
       A161: p`1 <= Gik`1 by A75,A135,A146,A160,TOPREAL1:9;
         k+1 <= len Ga by A1,NAT_1:38;
       then k+1-1 <= len Ga-1 by REAL_1:49;
       then A162: k <= len Ga-1 by XCMPLX_1:26;
       then len Ga-1 > 0 by A1,AXIOMS:22;
       then A163: k <= len Ga-'1 by A162,BINARITH:def 3;
         len Ga-'1 <= len Ga by GOBOARD9:2;
       then Gik`1 <= Ga*(len Ga-'1,1)`1 by A2,A8,A11,A19,A163,JORDAN1A:39;
       then p`1 <= Ga*(len Ga-'1,1)`1 by A161,AXIOMS:22;
       then p`1 <= E-bound C by A19,JORDAN8:15;
       then A164: p`1 <= (E-max C)`1 by PSCOMP_1:104;
         p`1 >= (E-max C)`1 by A159,JORDAN1A:def 3;
       then A165: p`1 = (E-max C)`1 by A164,AXIOMS:21;
         p`2 = (E-max C)`2 by A159,JORDAN1A:def 3;
       then p = E-max C by A165,TOPREAL3:11;
       hence contradiction by A5,A75,A119,A135,A160,XBOOLE_0:3;
      suppose east_halfline E-max C meets L~do;
       then consider p be set such that
        A166: p in east_halfline E-max C and
        A167: p in L~do by XBOOLE_0:3;
       reconsider p as Point of TOP-REAL 2 by A166;
         p in L~LS by A50,A167;
       then p in east_halfline E-max C /\ L~Cage(C,n)
                   by A132,A166,XBOOLE_0:def 3;
       then A168: p`1 = Ebo by A152,JORDAN1A:104;
       A169: (E-max C)`2 = p`2 by A166,JORDAN1A:def 3;
       set RC = Rotate(Cage(C,n),Emax);
       A170: E-max C in right_cell(RC,1) by JORDAN1I:9;
       A171: 1+1 <= len LS by A23,AXIOMS:22;
         LS = RC-:Wmin by JORDAN1G:26;
       then A172: LSeg(LS,1) = LSeg(RC,1) by A171,SPPOL_2:9;
       A173: L~RC = L~Cage(C,n) by REVROT_1:33;
       A174: len RC = len Cage(C,n) by REVROT_1:14;
       A175: GoB RC = GoB Cage(C,n) by REVROT_1:28
          .= Ga by JORDAN1H:52;
       A176: Emax in rng Cage(C,n) by SPRECT_2:50;
       A177: RC is_sequence_on Ga by A128,REVROT_1:34;
       A178: RC/.1 = E-max L~RC by A173,A176,FINSEQ_6:98;
       then consider ii,g be Nat such that
        A179: [ii,g+1] in Indices Ga and
        A180: [ii,g] in Indices Ga and
        A181: RC/.1 = Ga*(ii,g+1) and
        A182: RC/.(1+1) = Ga*(ii,g) by A79,A174,A177,JORDAN1I:25;
       consider g2 be Nat such that
        A183: 1 <= g2 & g2 <= width Ga and
        A184: Emax = Ga*(len Ga,g2) by JORDAN1D:29;
       A185: len Ga >= 4 by JORDAN8:13;
       then len Ga >= 1 by AXIOMS:22;
       then [len Ga,g2] in Indices Ga by A183,GOBOARD7:10;
       then A186: ii = len Ga by A173,A178,A179,A181,A184,GOBOARD1:21;
       A187: 1 <= ii & ii <= len Ga & 1 <= g+1 & g+1 <= width Ga
                                                           by A179,GOBOARD5:1;
       A188: 1 <= ii & ii <= len Ga & 1 <= g & g <= width Ga
                                                           by A180,GOBOARD5:1;
       A189: ii+1 <> ii by NAT_1:38;
         g+1 > g by NAT_1:38;
       then g+1+1 <> g by NAT_1:38;
       then A190: right_cell(RC,1) = cell(Ga,ii-'1,g)
                          by A79,A174,A175,A179,A180,A181,A182,A189,GOBOARD5:
def 6;
       A191: ii-'1+1 = ii by A187,AMI_5:4;
       A192: ii-1 >= 4-1 by A185,A186,REAL_1:49;
       then A193: ii-1 >= 1 by AXIOMS:22;
         ii-1 >= 0 by A192,AXIOMS:22;
       then A194: 1 <= ii-'1 by A193,BINARITH:def 3;
       then A195: Ga*(ii-'1,g)`2 <= p`2 & p`2 <= Ga*(ii-'1,g+1)`2
               by A169,A170,A187,A188,A190,A191,JORDAN9:19;
       A196: ii-'1 < len Ga by A187,A191,NAT_1:38;
       then A197: Ga*(ii-'1,g)`2 = Ga*(1,g)`2 by A188,A194,GOBOARD5:2
          .= Ga*(ii,g)`2 by A188,GOBOARD5:2;
       A198: Ga*(ii-'1,g+1)`2 = Ga*(1,g+1)`2 by A187,A194,A196,GOBOARD5:2
          .= Ga*(ii,g+1)`2 by A187,GOBOARD5:2;
         Ga*(len Ga,g)`1 = Ebo & Ebo = Ga*(len Ga,g+1)`1
                                                by A8,A187,A188,JORDAN1A:92;
       then p in LSeg(RC/.1,RC/.(1+1))
                              by A168,A181,A182,A186,A195,A197,A198,GOBOARD7:8;
       then A199: p in LSeg(LS,1) by A79,A172,A174,TOPREAL1:def 5;
       A200: p in LSeg(do,Index(p,do)) by A167,JORDAN3:42;
       A201: do = mid(LS,Gik..LS,len LS) by A34,JORDAN1J:37;
       A202: 1<=Gik..LS & Gik..LS<=len LS by A34,FINSEQ_4:31;
         Gik..LS <> len LS by A26,A34,FINSEQ_4:29;
       then A203: Gik..LS < len LS by A202,REAL_1:def 5;
       A204: 1<=Index(p,do) & Index(p,do) < len do by A167,JORDAN3:41;
       A205: Index(Gik,LS)+1 = Gik..LS by A29,A34,JORDAN1J:56;
       consider t be Nat such that
        A206: t in dom LS and
        A207: LS.t = Gik by A34,FINSEQ_2:11;
       A208: 1 <= t & t <= len LS by A206,FINSEQ_3:27;
       then 1 < t by A29,A207,REAL_1:def 5;
       then Index(Gik,LS)+1 = t by A207,A208,JORDAN3:45;
       then A209: len L_Cut(LS,Gik) = len LS-Index(Gik,LS)
                                                  by A6,A26,A207,JORDAN3:61;
       set tt = Index(p,do)+(Gik..LS)-'1;
       A210: 1<=Index(Gik,LS) & 0+Index(Gik,LS) < len LS by A6,JORDAN3:41;
       then A211: len LS-Index(Gik,LS) > 0 by REAL_1:86;
       then Index(p,do) < len LS-'Index(Gik,LS) by A204,A209,BINARITH:def 3;
       then Index(p,do)+1 <= len LS-'Index(Gik,LS) by NAT_1:38;
       then Index(p,do) <= len LS-'Index(Gik,LS)-1 by REAL_1:84;
       then Index(p,do) <= len LS-Index(Gik,LS)-1 by A211,BINARITH:def 3;
       then A212: Index(p,do) <= len LS-Gik..LS by A205,XCMPLX_1:36;
       then len LS-Gik..LS >= 1 by A204,AXIOMS:22;
       then len LS-Gik..LS >= 0 by AXIOMS:22;
       then Index(p,do) <= len LS-'Gik..LS by A212,BINARITH:def 3;
       then Index(p,do) < len LS-'(Gik..LS)+1 by NAT_1:38;
       then A213: LSeg(mid(LS,Gik..LS,len LS),Index(p,do)) =
            LSeg(LS,Index(p,do)+(Gik..LS)-'1) by A202,A203,A204,JORDAN4:31;
       A214: 1+1 <= Gik..LS by A205,A210,REAL_1:55;
       then Index(p,do)+Gik..LS >= 1+1+1 by A204,REAL_1:55;
       then A215: Index(p,do)+Gik..LS-1 >= 1+1+1-1 by REAL_1:49;
       then A216: Index(p,do)+Gik..LS-1 >= 0 by AXIOMS:22;
       then A217: tt >= 1+1 by A215,BINARITH:def 3;
       A218: 2 in dom LS by A171,FINSEQ_3:27;
         now per cases by A217,REAL_1:def 5;
        suppose tt > 1+1;
         then LSeg(LS,1) misses LSeg(LS,tt) by TOPREAL1:def 9;
         hence contradiction by A199,A200,A201,A213,XBOOLE_0:3;
        suppose A219: tt = 1+1;
         then LSeg(LS,1) /\ LSeg(LS,tt) = {LS/.2} by A23,TOPREAL1:def 8;
         then p in {LS/.2} by A199,A200,A201,A213,XBOOLE_0:def 3;
         then A220: p = LS/.2 by TARSKI:def 1;
         then A221: p..LS = 2 by A218,FINSEQ_5:44;
           1+1 = Index(p,do)+(Gik..LS)-1 by A216,A219,BINARITH:def 3;
         then 1+1+1 = Index(p,do)+(Gik..LS) by XCMPLX_1:27;
         then A222: Gik..LS = 2 by A204,A214,JORDAN1E:10;
           p in rng LS by A218,A220,PARTFUN2:4;
         then p = Gik by A34,A221,A222,FINSEQ_5:10;
         then Gik`1 = Ebo by A220,JORDAN1G:40;
         then Gik`1 = Ga*(len Ga,j)`1 by A8,A10,JORDAN1A:92;
         hence contradiction by A1,A13,A61,JORDAN1G:7;
       end;
       hence contradiction;
    end;
    then east_halfline E-max C c= (L~godo)` by SUBSET_1:43;
    then consider W be Subset of TOP-REAL 2 such that
     A223: W is_a_component_of (L~godo)` and
     A224: east_halfline E-max C c= W by GOBOARD9:5;
      east_halfline E-max C is not Bounded by JORDAN1C:9;
    then W is not Bounded by A224,JORDAN2C:16;
    then W is_outside_component_of L~godo by A223,JORDAN2C:def 4;
    then W c= UBD L~godo by JORDAN2C:27;
    then A225: east_halfline E-max C c= UBD L~godo by A224,XBOOLE_1:1;
      E-max C in east_halfline E-max C by JORDAN1C:7;
    then E-max C in UBD L~godo by A225;
    then E-max C in LeftComp godo by GOBRD14:46;
    then LA meets L~godo by A118,A119,A130,A138,JORDAN1J:36;
    then A226: LA meets (L~go \/ L~pion1) or LA meets L~do by A131,XBOOLE_1:70;
    A227: LA c= C by JORDAN1A:16;
    per cases by A226,XBOOLE_1:70;
     suppose LA meets L~go;
      then LA meets L~Cage(C,n) by A133,XBOOLE_1:63;
      then C meets L~Cage(C,n) by A227,XBOOLE_1:63;
      hence contradiction by JORDAN10:5;
     suppose LA meets L~pion1;
      hence contradiction by A5,A75,A135;
     suppose LA meets L~do;
      then LA meets L~Cage(C,n) by A133,XBOOLE_1:63;
      then C meets L~Cage(C,n) by A227,XBOOLE_1:63;
      hence contradiction by JORDAN10:5;
   end;

  theorem Th39:
   for C be Simple_closed_curve
   for i,j,k be Nat st
    1 < j & j <= k & k < len Gauge(C,n) &
    1 <= i & i <= width Gauge(C,n) &
    LSeg(Gauge(C,n)*(j,i),Gauge(C,n)*(k,i)) /\ L~Upper_Seq(C,n) =
     {Gauge(C,n)*(j,i)} &
    LSeg(Gauge(C,n)*(j,i),Gauge(C,n)*(k,i)) /\ L~Lower_Seq(C,n) =
     {Gauge(C,n)*(k,i)} holds
      LSeg(Gauge(C,n)*(j,i),Gauge(C,n)*(k,i)) meets Upper_Arc C
   proof
    let C be Simple_closed_curve;
    let i,j,k be Nat;
    set Ga = Gauge(C,n);
    set US = Upper_Seq(C,n);
    set LS = Lower_Seq(C,n);
    set UA = Upper_Arc C;
    set Wmin = W-min L~Cage(C,n);
    set Emax = E-max L~Cage(C,n);
    set Wbo = W-bound L~Cage(C,n);
    set Ebo = E-bound L~Cage(C,n);
    set Gij = Ga*(j,i);
    set Gik = Ga*(k,i);
    assume that
     A1: 1 < j & j <= k & k < len Ga and
     A2: 1 <= i & i <= width Ga and
     A3: LSeg(Gij,Gik) /\ L~US = {Gij} and
     A4: LSeg(Gij,Gik) /\ L~LS = {Gik} and
     A5: LSeg(Gij,Gik) misses UA;
      Gik in {Gik} by TARSKI:def 1;
    then A6: Gik in L~LS by A4,XBOOLE_0:def 3;
      Gij in {Gij} by TARSKI:def 1;
    then A7: Gij in L~US by A3,XBOOLE_0:def 3;
    A8: len Ga = width Ga by JORDAN8:def 1;
    A9: j <> k by A1,A2,A6,A7,Th29;
    A10: 1 < j & j < width Ga by A1,A8,AXIOMS:22;
    A11: 1 < k & k < width Ga by A1,AXIOMS:22,JORDAN8:def 1;
    A12: [j,i] in Indices Ga by A2,A8,A10,GOBOARD7:10;
    A13: [k,i] in Indices Ga by A2,A8,A11,GOBOARD7:10;
    A14: US is_sequence_on Ga by JORDAN1G:4;
    A15: LS is_sequence_on Ga by JORDAN1G:5;
    set go = R_Cut(US,Gij);
    set do = L_Cut(LS,Gik);
    A16: len US >= 3 by JORDAN1E:19;
    then len US >= 1 by AXIOMS:22;
    then 1 in dom US by FINSEQ_3:27;
    then A17: US.1 = US/.1 by FINSEQ_4:def 4
       .= Wmin by JORDAN1F:5;
    A18: Wmin`1 = Wbo by PSCOMP_1:84
       .= Ga*(1,i)`1 by A2,A8,JORDAN1A:94;
      len Ga >= 4 by JORDAN8:13;
    then A19: len Ga >= 1 by AXIOMS:22;
    then A20: [1,k] in Indices Ga by A11,GOBOARD7:10;
    A21: [1,i] in Indices Ga by A2,A19,GOBOARD7:10;
    then A22: Gij <> US.1 by A1,A12,A17,A18,JORDAN1G:7;
    then reconsider go as being_S-Seq FinSequence of TOP-REAL 2
                                                            by A7,JORDAN3:70;
    A23: len LS >= 1+2 by JORDAN1E:19;
    then len LS >= 1 by AXIOMS:22;
    then A24: 1 in dom LS & len LS in dom LS by FINSEQ_3:27;
    then A25: LS.len LS = LS/.len LS by FINSEQ_4:def 4
       .= Wmin by JORDAN1F:8;
      Wmin`1 = Wbo by PSCOMP_1:84
       .= Ga*(1,i)`1 by A2,A8,JORDAN1A:94;
    then A26: Gik <> LS.len LS by A1,A13,A21,A25,JORDAN1G:7;
    then reconsider do as being_S-Seq FinSequence of TOP-REAL 2
                                                            by A6,JORDAN3:69;
    A27: [len Ga,k] in Indices Ga by A11,A19,GOBOARD7:10;
    A28: LS.1 = LS/.1 by A24,FINSEQ_4:def 4
       .= Emax by JORDAN1F:6;
      Emax`1 = Ebo by PSCOMP_1:104
       .= Ga*(len Ga,k)`1 by A8,A11,JORDAN1A:92;
    then A29: Gik <> LS.1 by A1,A13,A27,A28,JORDAN1G:7;
    A30: len go >= 1+1 by TOPREAL1:def 10;
    A31: Gij in rng US by A2,A7,A8,A10,A14,JORDAN1J:40;
    then A32: go is_sequence_on Ga by A14,JORDAN1J:38;
    A33: len do >= 1+1 by TOPREAL1:def 10;
    A34: Gik in rng LS by A2,A6,A8,A11,A15,JORDAN1J:40;
    then A35: do is_sequence_on Ga by A15,JORDAN1J:39;
    reconsider go as non constant s.c.c.
            (being_S-Seq FinSequence of TOP-REAL 2) by A30,A32,JGRAPH_1:16,
JORDAN8:8;
    reconsider do as non constant s.c.c.
            (being_S-Seq FinSequence of TOP-REAL 2) by A33,A35,JGRAPH_1:16,
JORDAN8:8;
    A36: len go > 1 by A30,NAT_1:38;
    then A37: len go in dom go by FINSEQ_3:27;
    then A38: go/.len go = go.len go by FINSEQ_4:def 4
       .= Gij by A7,JORDAN3:59;
      len do >= 1 by A33,AXIOMS:22;
    then 1 in dom do by FINSEQ_3:27;
    then A39: do/.1 = do.1 by FINSEQ_4:def 4
       .= Gik by A6,JORDAN3:58;
    reconsider m = len go - 1 as Nat by A37,FINSEQ_3:28;
    A40: m+1 = len go by XCMPLX_1:27;
    then A41: len go-'1 = m by BINARITH:39;
    A42: LSeg(go,m) c= L~go by TOPREAL3:26;
    A43: L~go c= L~US by A7,JORDAN3:76;
    then LSeg(go,m) c= L~US by A42,XBOOLE_1:1;
    then A44: LSeg(go,m) /\ LSeg(Gik,Gij) c= {Gij} by A3,XBOOLE_1:26;
      m >= 1 by A30,REAL_1:84;
    then A45: LSeg(go,m) = LSeg(go/.m,Gij) by A38,A40,TOPREAL1:def 5;
      {Gij} c= LSeg(go,m) /\ LSeg(Gik,Gij)
    proof
     let x be set;
     assume x in {Gij};
     then A46: x = Gij by TARSKI:def 1;
     A47: Gij in LSeg(go,m) by A45,TOPREAL1:6;
       Gij in LSeg(Gik,Gij) by TOPREAL1:6;
     hence x in LSeg(go,m) /\ LSeg(Gik,Gij) by A46,A47,XBOOLE_0:def 3;
    end;
    then A48: LSeg(go,m) /\ LSeg(Gik,Gij) = {Gij} by A44,XBOOLE_0:def 10;
    A49: LSeg(do,1) c= L~do by TOPREAL3:26;
    A50: L~do c= L~LS by A6,JORDAN3:77;
    then LSeg(do,1) c= L~LS by A49,XBOOLE_1:1;
    then A51: LSeg(do,1) /\ LSeg(Gik,Gij) c= {Gik} by A4,XBOOLE_1:26;
    A52: LSeg(do,1) = LSeg(Gik,do/.(1+1)) by A33,A39,TOPREAL1:def 5;
      {Gik} c= LSeg(do,1) /\ LSeg(Gik,Gij)
    proof
     let x be set;
     assume x in {Gik};
     then A53: x = Gik by TARSKI:def 1;
     A54: Gik in LSeg(do,1) by A52,TOPREAL1:6;
       Gik in LSeg(Gik,Gij) by TOPREAL1:6;
     hence x in LSeg(do,1) /\ LSeg(Gik,Gij) by A53,A54,XBOOLE_0:def 3;
    end;
    then A55: LSeg(Gik,Gij) /\ LSeg(do,1) = {Gik} by A51,XBOOLE_0:def 10;
    A56: go/.1 = US/.1 by A7,SPRECT_3:39
       .= Wmin by JORDAN1F:5;
    then A57: go/.1 = LS/.len LS by JORDAN1F:8
       .= do/.len do by A6,JORDAN1J:35;
    A58: rng go c= L~go & rng do c= L~do by A30,A33,SPPOL_2:18;
    A59: {go/.1} c= L~go /\ L~do
    proof
     let x be set;
     assume x in {go/.1};
     then x = go/.1 by TARSKI:def 1;
     then x in rng go & x in rng do by A57,FINSEQ_6:46,REVROT_1:3;
     hence x in L~go /\ L~do by A58,XBOOLE_0:def 3;
    end;
    A60: LS.1 = LS/.1 by A24,FINSEQ_4:def 4
       .= Emax by JORDAN1F:6;
    A61: [len Ga,j] in Indices Ga by A10,A19,GOBOARD7:10;
      L~go /\ L~do c= {go/.1}
    proof
     let x be set;
     assume x in L~go /\ L~do;
     then A62: x in L~go & x in L~do by XBOOLE_0:def 3;
     then x in L~US /\ L~LS by A43,A50,XBOOLE_0:def 3;
     then x in {Wmin,Emax} by JORDAN1E:20;
     then A63: x = Wmin or x = Emax by TARSKI:def 2;
       now assume x = Emax;
      then A64: Emax = Gik by A6,A60,A62,JORDAN1E:11;
        Ga*(len Ga,j)`1 = Ebo by A8,A10,JORDAN1A:92;
      then Emax`1 <> Ebo by A1,A13,A61,A64,JORDAN1G:7;
      hence contradiction by PSCOMP_1:104;
     end;
     hence x in {go/.1} by A56,A63,TARSKI:def 1;
    end;
    then A65: L~go /\ L~do = {go/.1} by A59,XBOOLE_0:def 10;
    set W2 = go/.2;
    A66: 2 in dom go by A30,FINSEQ_3:27;
    A67: Gij..US >= 1 by A31,FINSEQ_4:31;
    A68: now assume Gij`1 = Wbo;
     then Ga*(1,k)`1 = Ga*(j,i)`1 by A8,A11,JORDAN1A:94;
     hence contradiction by A1,A12,A20,JORDAN1G:7;
    end;
      go = mid(US,1,Gij..US) by A31,JORDAN1G:57
       .= US|(Gij..US) by A67,JORDAN3:25;
    then A69: W2 = US/.2 by A66,TOPREAL1:1;
    A70: Wmin in rng go by A56,FINSEQ_6:46;
    set pion = <*Gij,Gik*>;
    A71: now let n be Nat;
     assume n in dom pion;
     then n in {1,2} by FINSEQ_1:4,FINSEQ_3:29;
     then n = 1 or n = 2 by TARSKI:def 2;
     then pion/.n = Gik or pion/.n = Gij by FINSEQ_4:26;
     hence ex j,i be Nat st [j,i] in Indices Ga & pion/.n = Ga*(j,i)
                                                                  by A12,A13;
    end;
    A72: Gik <> Gij by A9,A12,A13,GOBOARD1:21;
      Gik`2 = Ga*(1,i)`2 by A2,A8,A11,GOBOARD5:2
       .= Gij`2 by A2,A8,A10,GOBOARD5:2;
    then LSeg(Gik,Gij) is horizontal by SPPOL_1:36;
    then pion is_S-Seq by A72,JORDAN1B:9;
    then consider pion1 be FinSequence of TOP-REAL 2 such that
     A73: pion1 is_sequence_on Ga and
     A74: pion1 is_S-Seq and
     A75: L~pion = L~pion1 and
     A76: pion/.1 = pion1/.1 and
     A77: pion/.len pion = pion1/.len pion1 and
     A78: len pion <= len pion1 by A71,GOBOARD3:2;
    reconsider pion1 as being_S-Seq FinSequence of TOP-REAL 2 by A74;
    set godo = go^'pion1^'do;
      len Cage(C,n) > 4 by GOBOARD7:36;
    then A79: 1+1 <= len Cage(C,n) by AXIOMS:22;
    then A80: 1+1 <= len Rotate(Cage(C,n),Wmin) by REVROT_1:14;
      len (go^'pion1) >= len go by TOPREAL8:7;
    then A81: len (go^'pion1) >= 1+1 by A30,AXIOMS:22;
    then A82: len (go^'pion1) > 1+0 by NAT_1:38;
      len godo >= len (go^'pion1) by TOPREAL8:7;
    then A83: 1+1 <= len godo by A81,AXIOMS:22;
    A84: US is_sequence_on Ga by JORDAN1G:4;
    A85: go/.len go = pion1/.1 by A38,A76,FINSEQ_4:26;
    then A86: go^'pion1 is_sequence_on Ga by A32,A73,TOPREAL8:12;
    A87: (go^'pion1)/.len (go^'pion1) = pion/.len pion by A77,AMISTD_1:6
       .= pion/.2 by FINSEQ_1:61
       .= do/.1 by A39,FINSEQ_4:26;
    then A88: godo is_sequence_on Ga by A35,A86,TOPREAL8:12;
      LSeg(pion1,1) c= L~<*Gij,Gik*> by A75,TOPREAL3:26;
    then LSeg(pion1,1) c= LSeg(Gij,Gik) by SPPOL_2:21;
    then A89: LSeg(go,len go-'1) /\ LSeg(pion1,1) c= {Gij}
                                                      by A41,A48,XBOOLE_1:27;
    A90: len pion1 >= 1+1 by A78,FINSEQ_1:61;
      {Gij} c= LSeg(go,m) /\ LSeg(pion1,1)
    proof
     let x be set;
     assume x in {Gij};
     then A91: x = Gij by TARSKI:def 1;
     A92: Gij in LSeg(go,m) by A45,TOPREAL1:6;
       Gij in LSeg(pion1,1) by A38,A85,A90,TOPREAL1:27;
     hence x in LSeg(go,m) /\ LSeg(pion1,1) by A91,A92,XBOOLE_0:def 3;
    end;
    then LSeg(go,len go-'1) /\ LSeg(pion1,1) = { go/.len go }
                                             by A38,A41,A89,XBOOLE_0:def 10;
    then A93: go^'pion1 is unfolded by A85,TOPREAL8:34;
      len pion1 >= 2+0 by A78,FINSEQ_1:61;
    then A94: len pion1-2 >= 0 by REAL_1:84;
    A95: len (go^'pion1)-1 >= 0 by A82,REAL_1:84;
      len (go^'pion1)+1-1 = len go+len pion1-1 by GRAPH_2:13;
    then len (go^'pion1)-1 = len go+len pion1-1-1 by XCMPLX_1:26
       .= len go + len pion1-(1+1) by XCMPLX_1:36
       .= len go + (len pion1-2) by XCMPLX_1:29
       .= len go + (len pion1-'2) by A94,BINARITH:def 3;
    then A96: len (go^'pion1)-'1 = len go + (len pion1-'2)
                                                      by A95,BINARITH:def 3;
    A97: len pion1-1 >= 1 by A90,REAL_1:84;
    then A98: len pion1-1 >= 0 by AXIOMS:22;
    then A99: len pion1-'1 = len pion1-1 by BINARITH:def 3;
    A100: len pion1-'2+1 = len pion1-2+1 by A94,BINARITH:def 3
       .= len pion1-(2-1) by XCMPLX_1:37
       .= len pion1-'1 by A98,BINARITH:def 3;
      len pion1-1+1 <= len pion1 by XCMPLX_1:27;
    then A101: len pion1-'1 < len pion1 by A99,NAT_1:38;
      LSeg(pion1,len pion1-'1) c= L~<*Gij,Gik*> by A75,TOPREAL3:26;
    then LSeg(pion1,len pion1-'1) c= LSeg(Gij,Gik) by SPPOL_2:21;
    then A102: LSeg(pion1,len pion1-'1) /\ LSeg(do,1) c= {Gik}
                                                          by A55,XBOOLE_1:27;
      {Gik} c= LSeg(pion1,len pion1-'1) /\ LSeg(do,1)
    proof
     let x be set;
     assume x in {Gik};
     then A103: x = Gik by TARSKI:def 1;
     A104: Gik in LSeg(do,1) by A52,TOPREAL1:6;
     A105: len pion1-'1+1 = len pion1 by A99,XCMPLX_1:27;
     then pion1/.(len pion1-'1+1) = pion/.2 by A77,FINSEQ_1:61
        .= Gik by FINSEQ_4:26;
     then Gik in LSeg(pion1,len pion1-'1) by A97,A99,A105,TOPREAL1:27;
     hence x in LSeg(pion1,len pion1-'1) /\ LSeg(do,1)
                                                  by A103,A104,XBOOLE_0:def 3;
    end;
    then LSeg(pion1,len pion1-'1) /\ LSeg(do,1) = {Gik} by A102,XBOOLE_0:def 10
;
    then A106: LSeg(go^'pion1,len go+(len pion1-'2)) /\ LSeg(do,1) =
       {(go^'pion1)/.len (go^'pion1)} by A39,A85,A87,A100,A101,TOPREAL8:31;
    A107: (go^'pion1) is non trivial by A81,SPPOL_1:2;
    A108: rng pion1 c= L~pion1 by A90,SPPOL_2:18;
    A109: {pion1/.1} c= L~go /\ L~pion1
    proof
     let x be set;
     assume x in {pion1/.1};
     then x = pion1/.1 by TARSKI:def 1;
     then x in rng go & x in rng pion1 by A85,FINSEQ_6:46,REVROT_1:3;
     hence x in L~go /\ L~pion1 by A58,A108,XBOOLE_0:def 3;
    end;
      L~go /\ L~pion1 c= {pion1/.1}
    proof
     let x be set;
     assume x in L~go /\ L~pion1;
     then x in L~go & x in L~pion1 by XBOOLE_0:def 3;
     then x in L~pion1 /\ L~US by A43,XBOOLE_0:def 3;
     hence x in {pion1/.1} by A3,A38,A75,A85,SPPOL_2:21;
    end;
    then A110: L~go /\ L~pion1 = {pion1/.1} by A109,XBOOLE_0:def 10;
    then A111: (go^'pion1) is s.n.c. by A85,JORDAN1J:54;
      rng go /\ rng pion1 c= {pion1/.1} by A58,A108,A110,XBOOLE_1:27;
    then A112: go^'pion1 is one-to-one by JORDAN1J:55;
    A113: pion/.len pion = pion/.2 by FINSEQ_1:61
       .= do/.1 by A39,FINSEQ_4:26;
    A114: {pion1/.len pion1} c= L~do /\ L~pion1
    proof
     let x be set;
     assume x in {pion1/.len pion1};
     then x = pion1/.len pion1 by TARSKI:def 1;
     then x in rng do & x in rng pion1 by A77,A113,FINSEQ_6:46,REVROT_1:3;
     hence x in L~do /\ L~pion1 by A58,A108,XBOOLE_0:def 3;
    end;
      L~do /\ L~pion1 c= {pion1/.len pion1}
    proof
     let x be set;
     assume x in L~do /\ L~pion1;
     then x in L~do & x in L~pion1 by XBOOLE_0:def 3;
     then x in L~pion1 /\ L~LS by A50,XBOOLE_0:def 3;
     hence x in {pion1/.len pion1} by A4,A39,A75,A77,A113,SPPOL_2:21;
    end;
    then A115: L~do /\ L~pion1 = {pion1/.len pion1} by A114,XBOOLE_0:def 10;
    A116: L~(go^'pion1) /\ L~do = (L~go \/ L~pion1) /\ L~do by A85,TOPREAL8:35
       .= {go/.1} \/ {do/.1} by A65,A77,A113,A115,XBOOLE_1:23
       .= {(go^'pion1)/.1} \/ {do/.1} by AMISTD_1:5
       .= {(go^'pion1)/.1,do/.1} by ENUMSET1:41;
      do/.len do = (go^'pion1)/.1 by A57,AMISTD_1:5;
    then reconsider godo as non constant standard special_circular_sequence
        by A83,A87,A88,A93,A96,A106,A107,A111,A112,A116,JORDAN8:7,8,TOPREAL8:11
,33,34;
    A117: UA is_an_arc_of W-min C,E-max C by JORDAN6:def 8;
    then A118: UA is connected by JORDAN6:11;
    A119: W-min C in UA & E-max C in UA by A117,TOPREAL1:4;
    set ff = Rotate(Cage(C,n),Wmin);
      Wmin in rng Cage(C,n) by SPRECT_2:47;
    then A120: ff/.1 = Wmin by FINSEQ_6:98;
    A121: L~ff = L~Cage(C,n) by REVROT_1:33;
    then A122: (W-max L~ff)..ff > 1 by A120,SPRECT_5:23;
      (W-max L~ff)..ff <= (N-min L~ff)..ff by A120,A121,SPRECT_5:24;
    then A123: (N-min L~ff)..ff > 1 by A122,AXIOMS:22;
      (N-min L~ff)..ff < (N-max L~ff)..ff by A120,A121,SPRECT_5:25;
    then A124: (N-max L~ff)..ff > 1 by A123,AXIOMS:22;
      (N-max L~ff)..ff <= (E-max L~ff)..ff by A120,A121,SPRECT_5:26;
    then A125: Emax..ff > 1 by A121,A124,AXIOMS:22;
    A126: now assume A127: Gij..US <= 1;
       Gij..US >= 1 by A31,FINSEQ_4:31;
     then Gij..US = 1 by A127,AXIOMS:21;
     then Gij = US/.1 by A31,FINSEQ_5:41;
     hence contradiction by A17,A22,JORDAN1F:5;
    end;
    A128: Cage(C,n) is_sequence_on Ga by JORDAN9:def 1;
    then A129: ff is_sequence_on Ga by REVROT_1:34;
    A130: right_cell(godo,1,Ga)\L~godo c= RightComp godo by A83,A88,JORDAN9:29;
    A131: L~godo = L~(go^'pion1) \/ L~do by A87,TOPREAL8:35
       .= L~go \/ L~pion1 \/ L~do by A85,TOPREAL8:35;
      L~Cage(C,n) = L~US \/ L~LS by JORDAN1E:17;
    then A132: L~US c= L~Cage(C,n) & L~LS c= L~Cage(C,n) by XBOOLE_1:7;
    then A133: L~go c= L~Cage(C,n) & L~do c= L~Cage(C,n)
                                                     by A43,A50,XBOOLE_1:1;
    A134: W-min C in C by SPRECT_1:15;
    A135: L~pion = LSeg(Gik,Gij) by SPPOL_2:21;
    A136: now assume W-min C in L~godo;
     then A137: W-min C in L~go \/ L~pion1 or W-min C in L~do by A131,XBOOLE_0:
def 2;
     per cases by A137,XBOOLE_0:def 2;
      suppose W-min C in L~go;
       then C meets L~Cage(C,n) by A133,A134,XBOOLE_0:3;
       hence contradiction by JORDAN10:5;
      suppose W-min C in L~pion1;
       hence contradiction by A5,A75,A119,A135,XBOOLE_0:3;
      suppose W-min C in L~do;
       then C meets L~Cage(C,n) by A133,A134,XBOOLE_0:3;
       hence contradiction by JORDAN10:5;
    end;
      right_cell(Rotate(Cage(C,n),Wmin),1) =
          right_cell(ff,1,GoB ff) by A80,JORDAN1H:29
       .= right_cell(ff,1,GoB Cage(C,n)) by REVROT_1:28
       .= right_cell(ff,1,Ga) by JORDAN1H:52
       .= right_cell(ff-:Emax,1,Ga) by A125,A129,JORDAN1J:53
       .= right_cell(US,1,Ga) by JORDAN1E:def 1
       .= right_cell(R_Cut(US,Gij),1,Ga) by A31,A84,A126,JORDAN1J:52
       .= right_cell(go^'pion1,1,Ga) by A36,A86,JORDAN1J:51
       .= right_cell(godo,1,Ga) by A82,A88,JORDAN1J:51;
    then W-min C in right_cell(godo,1,Ga) by JORDAN1I:8;
    then A138: W-min C in right_cell(godo,1,Ga)\L~godo by A136,XBOOLE_0:def 4;
    A139: godo/.1 = (go^'pion1)/.1 by AMISTD_1:5
       .= Wmin by A56,AMISTD_1:5;
    A140: len US >= 2 by A16,AXIOMS:22;
    A141: godo/.2 = (go^'pion1)/.2 by A81,AMISTD_1:9
       .= US/.2 by A30,A69,AMISTD_1:9
       .= (US^'LS)/.2 by A140,AMISTD_1:9
       .= Rotate(Cage(C,n),Wmin)/.2 by JORDAN1E:15;
    A142: L~go \/ L~do is compact by COMPTS_1:19;
    A143: L~go \/ L~do c= L~Cage(C,n) by A133,XBOOLE_1:8;
      Wmin in L~go \/ L~do by A58,A70,XBOOLE_0:def 2;
    then A144: W-min (L~go \/ L~do) = Wmin by A142,A143,JORDAN1J:21;
    A145: (W-min (L~go \/ L~do))`1 = W-bound (L~go \/ L~do) & Wmin`1 = Wbo
                                                              by PSCOMP_1:84;
    A146: Gij`1 <= Gik`1 by A1,A2,SPRECT_3:25;
    then W-bound LSeg(Gik,Gij) = Gij`1 by SPRECT_1:62;
    then A147: W-bound L~pion1 = Gij`1 by A75,SPPOL_2:21;
      Gij`1 >= Wbo by A7,A132,PSCOMP_1:71;
    then Gij`1 > Wbo by A68,REAL_1:def 5;
    then W-min (L~go\/L~do\/L~pion1) = W-min (L~go \/ L~do)
                                              by A142,A144,A145,A147,JORDAN1J:
33;
    then A148: W-min L~godo = Wmin by A131,A144,XBOOLE_1:4;
    A149: rng godo c= L~godo by A83,SPPOL_2:18;
      2 in dom godo by A83,FINSEQ_3:27;
    then A150: godo/.2 in rng godo by PARTFUN2:4;
      godo/.2 in W-most L~Cage(C,n) by A141,JORDAN1I:27;
    then (godo/.2)`1 = (W-min L~godo)`1 by A148,PSCOMP_1:88
       .= W-bound L~godo by PSCOMP_1:84;
    then godo/.2 in W-most L~godo by A149,A150,SPRECT_2:16;
    then Rotate(godo,W-min L~godo)/.2 in W-most L~godo by A139,A148,FINSEQ_6:95
;
    then reconsider godo as clockwise_oriented non constant standard
                                    special_circular_sequence by JORDAN1I:27;
      len US in dom US by FINSEQ_5:6;
    then A151: US.len US = US/.len US by FINSEQ_4:def 4
       .= Emax by JORDAN1F:7;
    A152: E-max C in E-most C by PSCOMP_1:111;
    A153: east_halfline E-max C misses L~go
    proof
     assume east_halfline E-max C meets L~go;
     then consider p be set such that
      A154: p in east_halfline E-max C and
      A155: p in L~go by XBOOLE_0:3;
     reconsider p as Point of TOP-REAL 2 by A154;
       p in L~US by A43,A155;
     then p in east_halfline E-max C /\ L~Cage(C,n) by A132,A154,XBOOLE_0:def 3
;
     then A156: p`1 = Ebo by A152,JORDAN1A:104;
     then A157: p = Emax by A43,A155,JORDAN1J:46;
     then Emax = Gij by A7,A151,A155,JORDAN1J:43;
     then Gij`1 = Ga*(len Ga,k)`1 by A8,A11,A156,A157,JORDAN1A:92;
     hence contradiction by A1,A12,A27,JORDAN1G:7;
    end;
      now assume east_halfline E-max C meets L~godo;
     then A158: east_halfline E-max C meets (L~go \/ L~pion1) or
               east_halfline E-max C meets L~do by A131,XBOOLE_1:70;
     per cases by A158,XBOOLE_1:70;
      suppose east_halfline E-max C meets L~go;
       hence contradiction by A153;
      suppose east_halfline E-max C meets L~pion1;
       then consider p be set such that
        A159: p in east_halfline E-max C and
        A160: p in L~pion1 by XBOOLE_0:3;
       reconsider p as Point of TOP-REAL 2 by A159;
       A161: p`1 <= Gik`1 by A75,A135,A146,A160,TOPREAL1:9;
         k+1 <= len Ga by A1,NAT_1:38;
       then k+1-1 <= len Ga-1 by REAL_1:49;
       then A162: k <= len Ga-1 by XCMPLX_1:26;
       then len Ga-1 > 0 by A1,AXIOMS:22;
       then A163: k <= len Ga-'1 by A162,BINARITH:def 3;
         len Ga-'1 <= len Ga by GOBOARD9:2;
       then Gik`1 <= Ga*(len Ga-'1,1)`1 by A2,A8,A11,A19,A163,JORDAN1A:39;
       then p`1 <= Ga*(len Ga-'1,1)`1 by A161,AXIOMS:22;
       then p`1 <= E-bound C by A19,JORDAN8:15;
       then A164: p`1 <= (E-max C)`1 by PSCOMP_1:104;
         p`1 >= (E-max C)`1 by A159,JORDAN1A:def 3;
       then A165: p`1 = (E-max C)`1 by A164,AXIOMS:21;
         p`2 = (E-max C)`2 by A159,JORDAN1A:def 3;
       then p = E-max C by A165,TOPREAL3:11;
       hence contradiction by A5,A75,A119,A135,A160,XBOOLE_0:3;
      suppose east_halfline E-max C meets L~do;
       then consider p be set such that
        A166: p in east_halfline E-max C and
        A167: p in L~do by XBOOLE_0:3;
       reconsider p as Point of TOP-REAL 2 by A166;
         p in L~LS by A50,A167;
       then p in east_halfline E-max C /\ L~Cage(C,n)
                   by A132,A166,XBOOLE_0:def 3;
       then A168: p`1 = Ebo by A152,JORDAN1A:104;
       A169: (E-max C)`2 = p`2 by A166,JORDAN1A:def 3;
       set RC = Rotate(Cage(C,n),Emax);
       A170: E-max C in right_cell(RC,1) by JORDAN1I:9;
       A171: 1+1 <= len LS by A23,AXIOMS:22;
         LS = RC-:Wmin by JORDAN1G:26;
       then A172: LSeg(LS,1) = LSeg(RC,1) by A171,SPPOL_2:9;
       A173: L~RC = L~Cage(C,n) by REVROT_1:33;
       A174: len RC = len Cage(C,n) by REVROT_1:14;
       A175: GoB RC = GoB Cage(C,n) by REVROT_1:28
          .= Ga by JORDAN1H:52;
       A176: Emax in rng Cage(C,n) by SPRECT_2:50;
       A177: RC is_sequence_on Ga by A128,REVROT_1:34;
       A178: RC/.1 = E-max L~RC by A173,A176,FINSEQ_6:98;
       then consider ii,g be Nat such that
        A179: [ii,g+1] in Indices Ga and
        A180: [ii,g] in Indices Ga and
        A181: RC/.1 = Ga*(ii,g+1) and
        A182: RC/.(1+1) = Ga*(ii,g) by A79,A174,A177,JORDAN1I:25;
       consider g2 be Nat such that
        A183: 1 <= g2 & g2 <= width Ga and
        A184: Emax = Ga*(len Ga,g2) by JORDAN1D:29;
       A185: len Ga >= 4 by JORDAN8:13;
       then len Ga >= 1 by AXIOMS:22;
       then [len Ga,g2] in Indices Ga by A183,GOBOARD7:10;
       then A186: ii = len Ga by A173,A178,A179,A181,A184,GOBOARD1:21;
       A187: 1 <= ii & ii <= len Ga & 1 <= g+1 & g+1 <= width Ga
                                                           by A179,GOBOARD5:1;
       A188: 1 <= ii & ii <= len Ga & 1 <= g & g <= width Ga
                                                           by A180,GOBOARD5:1;
       A189: ii+1 <> ii by NAT_1:38;
         g+1 > g by NAT_1:38;
       then g+1+1 <> g by NAT_1:38;
       then A190: right_cell(RC,1) = cell(Ga,ii-'1,g)
                          by A79,A174,A175,A179,A180,A181,A182,A189,GOBOARD5:
def 6;
       A191: ii-'1+1 = ii by A187,AMI_5:4;
       A192: ii-1 >= 4-1 by A185,A186,REAL_1:49;
       then A193: ii-1 >= 1 by AXIOMS:22;
         ii-1 >= 0 by A192,AXIOMS:22;
       then A194: 1 <= ii-'1 by A193,BINARITH:def 3;
       then A195: Ga*(ii-'1,g)`2 <= p`2 & p`2 <= Ga*(ii-'1,g+1)`2
              by A169,A170,A187,A188,A190,A191,JORDAN9:19;
       A196: ii-'1 < len Ga by A187,A191,NAT_1:38;
       then A197: Ga*(ii-'1,g)`2 = Ga*(1,g)`2 by A188,A194,GOBOARD5:2
          .= Ga*(ii,g)`2 by A188,GOBOARD5:2;
       A198: Ga*(ii-'1,g+1)`2 = Ga*(1,g+1)`2 by A187,A194,A196,GOBOARD5:2
          .= Ga*(ii,g+1)`2 by A187,GOBOARD5:2;
         Ga*(len Ga,g)`1 = Ebo & Ebo = Ga*(len Ga,g+1)`1
                                                by A8,A187,A188,JORDAN1A:92;
       then p in LSeg(RC/.1,RC/.(1+1))
                              by A168,A181,A182,A186,A195,A197,A198,GOBOARD7:8;
       then A199: p in LSeg(LS,1) by A79,A172,A174,TOPREAL1:def 5;
       A200: p in LSeg(do,Index(p,do)) by A167,JORDAN3:42;
       A201: do = mid(LS,Gik..LS,len LS) by A34,JORDAN1J:37;
       A202: 1<=Gik..LS & Gik..LS<=len LS by A34,FINSEQ_4:31;
         Gik..LS <> len LS by A26,A34,FINSEQ_4:29;
       then A203: Gik..LS < len LS by A202,REAL_1:def 5;
       A204: 1<=Index(p,do) & Index(p,do) < len do by A167,JORDAN3:41;
       A205: Index(Gik,LS)+1 = Gik..LS by A29,A34,JORDAN1J:56;
       consider t be Nat such that
        A206: t in dom LS and
        A207: LS.t = Gik by A34,FINSEQ_2:11;
       A208: 1 <= t & t <= len LS by A206,FINSEQ_3:27;
       then 1 < t by A29,A207,REAL_1:def 5;
       then Index(Gik,LS)+1 = t by A207,A208,JORDAN3:45;
       then A209: len L_Cut(LS,Gik) = len LS-Index(Gik,LS)
                                                  by A6,A26,A207,JORDAN3:61;
       set tt = Index(p,do)+(Gik..LS)-'1;
       A210: 1<=Index(Gik,LS) & 0+Index(Gik,LS) < len LS by A6,JORDAN3:41;
       then A211: len LS-Index(Gik,LS) > 0 by REAL_1:86;
       then Index(p,do) < len LS-'Index(Gik,LS) by A204,A209,BINARITH:def 3;
       then Index(p,do)+1 <= len LS-'Index(Gik,LS) by NAT_1:38;
       then Index(p,do) <= len LS-'Index(Gik,LS)-1 by REAL_1:84;
       then Index(p,do) <= len LS-Index(Gik,LS)-1 by A211,BINARITH:def 3;
       then A212: Index(p,do) <= len LS-Gik..LS by A205,XCMPLX_1:36;
       then len LS-Gik..LS >= 1 by A204,AXIOMS:22;
       then len LS-Gik..LS >= 0 by AXIOMS:22;
       then Index(p,do) <= len LS-'Gik..LS by A212,BINARITH:def 3;
       then Index(p,do) < len LS-'(Gik..LS)+1 by NAT_1:38;
       then A213: LSeg(mid(LS,Gik..LS,len LS),Index(p,do)) =
            LSeg(LS,Index(p,do)+(Gik..LS)-'1) by A202,A203,A204,JORDAN4:31;
       A214: 1+1 <= Gik..LS by A205,A210,REAL_1:55;
       then Index(p,do)+Gik..LS >= 1+1+1 by A204,REAL_1:55;
       then A215: Index(p,do)+Gik..LS-1 >= 1+1+1-1 by REAL_1:49;
       then A216: Index(p,do)+Gik..LS-1 >= 0 by AXIOMS:22;
       then A217: tt >= 1+1 by A215,BINARITH:def 3;
       A218: 2 in dom LS by A171,FINSEQ_3:27;
         now per cases by A217,REAL_1:def 5;
        suppose tt > 1+1;
         then LSeg(LS,1) misses LSeg(LS,tt) by TOPREAL1:def 9;
         hence contradiction by A199,A200,A201,A213,XBOOLE_0:3;
        suppose A219: tt = 1+1;
         then LSeg(LS,1) /\ LSeg(LS,tt) = {LS/.2} by A23,TOPREAL1:def 8;
         then p in {LS/.2} by A199,A200,A201,A213,XBOOLE_0:def 3;
         then A220: p = LS/.2 by TARSKI:def 1;
         then A221: p..LS = 2 by A218,FINSEQ_5:44;
           1+1 = Index(p,do)+(Gik..LS)-1 by A216,A219,BINARITH:def 3;
         then 1+1+1 = Index(p,do)+(Gik..LS) by XCMPLX_1:27;
         then A222: Gik..LS = 2 by A204,A214,JORDAN1E:10;
           p in rng LS by A218,A220,PARTFUN2:4;
         then p = Gik by A34,A221,A222,FINSEQ_5:10;
         then Gik`1 = Ebo by A220,JORDAN1G:40;
         then Gik`1 = Ga*(len Ga,j)`1 by A8,A10,JORDAN1A:92;
         hence contradiction by A1,A13,A61,JORDAN1G:7;
       end;
       hence contradiction;
    end;
    then east_halfline E-max C c= (L~godo)` by SUBSET_1:43;
    then consider W be Subset of TOP-REAL 2 such that
     A223: W is_a_component_of (L~godo)` and
     A224: east_halfline E-max C c= W by GOBOARD9:5;
      east_halfline E-max C is not Bounded by JORDAN1C:9;
    then W is not Bounded by A224,JORDAN2C:16;
    then W is_outside_component_of L~godo by A223,JORDAN2C:def 4;
    then W c= UBD L~godo by JORDAN2C:27;
    then A225: east_halfline E-max C c= UBD L~godo by A224,XBOOLE_1:1;
      E-max C in east_halfline E-max C by JORDAN1C:7;
    then E-max C in UBD L~godo by A225;
    then E-max C in LeftComp godo by GOBRD14:46;
    then UA meets L~godo by A118,A119,A130,A138,JORDAN1J:36;
    then A226: UA meets (L~go \/ L~pion1) or UA meets L~do by A131,XBOOLE_1:70;
    A227: UA c= C by JORDAN1A:16;
    per cases by A226,XBOOLE_1:70;
     suppose UA meets L~go;
      then UA meets L~Cage(C,n) by A133,XBOOLE_1:63;
      then C meets L~Cage(C,n) by A227,XBOOLE_1:63;
      hence contradiction by JORDAN10:5;
     suppose UA meets L~pion1;
      hence contradiction by A5,A75,A135;
     suppose UA meets L~do;
      then UA meets L~Cage(C,n) by A133,XBOOLE_1:63;
      then C meets L~Cage(C,n) by A227,XBOOLE_1:63;
      hence contradiction by JORDAN10:5;
   end;

  theorem Th40:
   for C be Simple_closed_curve
   for i,j,k be Nat st
    1 < j & j <= k & k < len Gauge(C,n) &
    1 <= i & i <= width Gauge(C,n) &
    Gauge(C,n)*(j,i) in L~Upper_Seq(C,n) &
    Gauge(C,n)*(k,i) in L~Lower_Seq(C,n) holds
     LSeg(Gauge(C,n)*(j,i),Gauge(C,n)*(k,i)) meets Lower_Arc C
   proof
    let C be Simple_closed_curve;
    let i,j,k be Nat;
    assume that
     A1: 1 < j & j <= k & k < len Gauge(C,n) and
     A2: 1 <= i & i <= width Gauge(C,n) and
     A3: Gauge(C,n)*(j,i) in L~Upper_Seq(C,n) and
     A4: Gauge(C,n)*(k,i) in L~Lower_Seq(C,n);
    consider j1,k1 be Nat such that
     A5: j <= j1 and
     A6: j1 <= k1 and
     A7: k1 <= k and
     A8: LSeg(Gauge(C,n)*(j1,i),Gauge(C,n)*(k1,i)) /\ L~Upper_Seq(C,n) =
      {Gauge(C,n)*(j1,i)} and
     A9: LSeg(Gauge(C,n)*(j1,i),Gauge(C,n)*(k1,i)) /\ L~Lower_Seq(C,n) =
      {Gauge(C,n)*(k1,i)} by A1,A2,A3,A4,Th22;
      1 < j1 & k1 < len Gauge(C,n) by A1,A5,A7,AXIOMS:22;
    then A10: LSeg(Gauge(C,n)*(j1,i),Gauge(C,n)*(k1,i)) meets Lower_Arc C
                                                   by A2,A6,A8,A9,Th38;
      LSeg(Gauge(C,n)*(j1,i),Gauge(C,n)*(k1,i)) c=
     LSeg(Gauge(C,n)*(j,i),Gauge(C,n)*(k,i)) by A1,A2,A5,A6,A7,Th8;
    hence LSeg(Gauge(C,n)*(j,i),Gauge(C,n)*(k,i)) meets Lower_Arc C
                                                          by A10,XBOOLE_1:63;
   end;

  theorem Th41:
   for C be Simple_closed_curve
   for i,j,k be Nat st
    1 < j & j <= k & k < len Gauge(C,n) &
    1 <= i & i <= width Gauge(C,n) &
    Gauge(C,n)*(j,i) in L~Upper_Seq(C,n) &
    Gauge(C,n)*(k,i) in L~Lower_Seq(C,n) holds
     LSeg(Gauge(C,n)*(j,i),Gauge(C,n)*(k,i)) meets Upper_Arc C
   proof
    let C be Simple_closed_curve;
    let i,j,k be Nat;
    assume that
     A1: 1 < j & j <= k & k < len Gauge(C,n) and
     A2: 1 <= i & i <= width Gauge(C,n) and
     A3: Gauge(C,n)*(j,i) in L~Upper_Seq(C,n) and
     A4: Gauge(C,n)*(k,i) in L~Lower_Seq(C,n);
    consider j1,k1 be Nat such that
     A5: j <= j1 and
     A6: j1 <= k1 and
     A7: k1 <= k and
     A8: LSeg(Gauge(C,n)*(j1,i),Gauge(C,n)*(k1,i)) /\ L~Upper_Seq(C,n) =
      {Gauge(C,n)*(j1,i)} and
     A9: LSeg(Gauge(C,n)*(j1,i),Gauge(C,n)*(k1,i)) /\ L~Lower_Seq(C,n) =
      {Gauge(C,n)*(k1,i)} by A1,A2,A3,A4,Th22;
      1 < j1 & k1 < len Gauge(C,n) by A1,A5,A7,AXIOMS:22;
    then A10: LSeg(Gauge(C,n)*(j1,i),Gauge(C,n)*(k1,i)) meets Upper_Arc C
                                                   by A2,A6,A8,A9,Th39;
      LSeg(Gauge(C,n)*(j1,i),Gauge(C,n)*(k1,i)) c=
     LSeg(Gauge(C,n)*(j,i),Gauge(C,n)*(k,i)) by A1,A2,A5,A6,A7,Th8;
    hence LSeg(Gauge(C,n)*(j,i),Gauge(C,n)*(k,i)) meets Upper_Arc C
                                                          by A10,XBOOLE_1:63;
   end;

  theorem Th42:
   for C be Simple_closed_curve
   for i,j,k be Nat st
    1 < j & j <= k & k < len Gauge(C,n) &
    1 <= i & i <= width Gauge(C,n) & n > 0 &
    Gauge(C,n)*(j,i) in Upper_Arc L~Cage(C,n) &
    Gauge(C,n)*(k,i) in Lower_Arc L~Cage(C,n) holds
     LSeg(Gauge(C,n)*(j,i),Gauge(C,n)*(k,i)) meets Lower_Arc C
   proof
    let C be Simple_closed_curve;
    let i,j,k be Nat;
    assume that
     A1: 1 < j & j <= k & k < len Gauge(C,n) and
     A2: 1 <= i & i <= width Gauge(C,n) and
     A3: n > 0 and
     A4: Gauge(C,n)*(j,i) in Upper_Arc L~Cage(C,n) and
     A5: Gauge(C,n)*(k,i) in Lower_Arc L~Cage(C,n);
    A6: L~Upper_Seq(C,n) = Upper_Arc L~Cage(C,n) by A3,JORDAN1G:63;
      L~Lower_Seq(C,n) = Lower_Arc L~Cage(C,n) by A3,JORDAN1G:64;
    hence LSeg(Gauge(C,n)*(j,i),Gauge(C,n)*(k,i)) meets Lower_Arc C
                                                   by A1,A2,A4,A5,A6,Th40;
   end;

  theorem Th43:
   for C be Simple_closed_curve
   for i,j,k be Nat st
    1 < j & j <= k & k < len Gauge(C,n) &
    1 <= i & i <= width Gauge(C,n) & n > 0 &
    Gauge(C,n)*(j,i) in Upper_Arc L~Cage(C,n) &
    Gauge(C,n)*(k,i) in Lower_Arc L~Cage(C,n) holds
     LSeg(Gauge(C,n)*(j,i),Gauge(C,n)*(k,i)) meets Upper_Arc C
   proof
    let C be Simple_closed_curve;
    let i,j,k be Nat;
    assume that
     A1: 1 < j & j <= k & k < len Gauge(C,n) and
     A2: 1 <= i & i <= width Gauge(C,n) and
     A3: n > 0 and
     A4: Gauge(C,n)*(j,i) in Upper_Arc L~Cage(C,n) and
     A5: Gauge(C,n)*(k,i) in Lower_Arc L~Cage(C,n);
    A6: L~Upper_Seq(C,n) = Upper_Arc L~Cage(C,n) by A3,JORDAN1G:63;
      L~Lower_Seq(C,n) = Lower_Arc L~Cage(C,n) by A3,JORDAN1G:64;
    hence LSeg(Gauge(C,n)*(j,i),Gauge(C,n)*(k,i)) meets Upper_Arc C
                                                   by A1,A2,A4,A5,A6,Th41;
   end;

  theorem
     for C be Simple_closed_curve
   for j,k be Nat holds
    1 < j & j <= k & k < len Gauge(C,n+1) &
    Gauge(C,n+1)*(j,Center Gauge(C,n+1)) in Upper_Arc L~Cage(C,n+1) &
    Gauge(C,n+1)*(k,Center Gauge(C,n+1)) in Lower_Arc L~Cage(C,n+1) implies
     LSeg(Gauge(C,n+1)*(j,Center Gauge(C,n+1)),
          Gauge(C,n+1)*(k,Center Gauge(C,n+1))) meets Lower_Arc C
   proof
    let C be Simple_closed_curve;
    let j,k be Nat;
    assume that
     A1: 1 < j & j <= k & k < len Gauge(C,n+1) and
     A2: Gauge(C,n+1)*(j,Center Gauge(C,n+1)) in Upper_Arc L~Cage(C,n+1) and
     A3: Gauge(C,n+1)*(k,Center Gauge(C,n+1)) in Lower_Arc L~Cage(C,n+1);
      n+1 >= 0+1 by NAT_1:29;
    then A4: n+1 > 0 by NAT_1:38;
    A5: len Gauge(C,n+1) >= 4 by JORDAN8:13;
    then A6: len Gauge(C,n+1) >= 3 by AXIOMS:22;
      len Gauge(C,n+1) >= 2 by A5,AXIOMS:22;
    then A7: 1 < Center Gauge(C,n+1) by JORDAN1B:15;
      Center Gauge(C,n+1) < len Gauge(C,n+1) by A6,JORDAN1B:16;
    then Center Gauge(C,n+1) < width Gauge(C,n+1) by JORDAN8:def 1;
    hence LSeg(Gauge(C,n+1)*(j,Center Gauge(C,n+1)),
          Gauge(C,n+1)*(k,Center Gauge(C,n+1))) meets Lower_Arc C
                                                   by A1,A2,A3,A4,A7,Th42;
   end;

  theorem
     for C be Simple_closed_curve
   for j,k be Nat holds
    1 < j & j <= k & k < len Gauge(C,n+1) &
    Gauge(C,n+1)*(j,Center Gauge(C,n+1)) in Upper_Arc L~Cage(C,n+1) &
    Gauge(C,n+1)*(k,Center Gauge(C,n+1)) in Lower_Arc L~Cage(C,n+1) implies
     LSeg(Gauge(C,n+1)*(j,Center Gauge(C,n+1)),
          Gauge(C,n+1)*(k,Center Gauge(C,n+1))) meets Upper_Arc C
   proof
    let C be Simple_closed_curve;
    let j,k be Nat;
    assume that
     A1: 1 < j & j <= k & k < len Gauge(C,n+1) and
     A2: Gauge(C,n+1)*(j,Center Gauge(C,n+1)) in Upper_Arc L~Cage(C,n+1) and
     A3: Gauge(C,n+1)*(k,Center Gauge(C,n+1)) in Lower_Arc L~Cage(C,n+1);
      n+1 >= 0+1 by NAT_1:29;
    then A4: n+1 > 0 by NAT_1:38;
    A5: len Gauge(C,n+1) >= 4 by JORDAN8:13;
    then A6: len Gauge(C,n+1) >= 3 by AXIOMS:22;
      len Gauge(C,n+1) >= 2 by A5,AXIOMS:22;
    then A7: 1 < Center Gauge(C,n+1) by JORDAN1B:15;
      Center Gauge(C,n+1) < len Gauge(C,n+1) by A6,JORDAN1B:16;
    then Center Gauge(C,n+1) < width Gauge(C,n+1) by JORDAN8:def 1;
    hence LSeg(Gauge(C,n+1)*(j,Center Gauge(C,n+1)),
          Gauge(C,n+1)*(k,Center Gauge(C,n+1))) meets Upper_Arc C
                                                   by A1,A2,A3,A4,A7,Th43;
   end;

  theorem Th46:
   for C be Simple_closed_curve
   for i1,i2,j,k be Nat st
    1 < i1 & i1 <= i2 & i2 < len Gauge(C,n) &
    1 <= j & j <= k & k <= width Gauge(C,n) &
    (LSeg(Gauge(C,n)*(i1,j),Gauge(C,n)*(i1,k)) \/
     LSeg(Gauge(C,n)*(i1,k),Gauge(C,n)*(i2,k))) /\ L~Upper_Seq(C,n) =
     {Gauge(C,n)*(i2,k)} &
    (LSeg(Gauge(C,n)*(i1,j),Gauge(C,n)*(i1,k)) \/
     LSeg(Gauge(C,n)*(i1,k),Gauge(C,n)*(i2,k))) /\ L~Lower_Seq(C,n) =
     {Gauge(C,n)*(i1,j)} holds
      (LSeg(Gauge(C,n)*(i1,j),Gauge(C,n)*(i1,k)) \/
       LSeg(Gauge(C,n)*(i1,k),Gauge(C,n)*(i2,k))) meets Upper_Arc C
   proof
    let C be Simple_closed_curve;
    let i1,i2,j,k be Nat;
    set G = Gauge(C,n);
    set pio = LSeg(G*(i1,j),G*(i1,k));
    set poz = LSeg(G*(i1,k),G*(i2,k));
    set US = Upper_Seq(C,n);
    set LS = Lower_Seq(C,n);
    assume that
     A1: 1 < i1 & i1 <= i2 & i2 < len G and
     A2: 1 <= j & j <= k & k <= width G and
     A3: (pio \/ poz) /\ L~US = {G*(i2,k)} and
     A4: (pio \/ poz) /\ L~LS = {G*(i1,j)} and
     A5: (pio \/ poz) misses Upper_Arc C;
    set UA = Upper_Arc C;
    set Wmin = W-min L~Cage(C,n);
    set Emax = E-max L~Cage(C,n);
    set Wbo = W-bound L~Cage(C,n);
    set Ebo = E-bound L~Cage(C,n);
    set Gij = G*(i1,j);
    set Gik = G*(i2,k);
    set Gi1k = G*(i1,k);
    A6: 1 < i1 & i1 < len G & 1 < i2 & i2 < len G by A1,AXIOMS:22;
    A7: L~<*Gik,Gi1k,Gij*> = poz \/ pio by TOPREAL3:23;
      Gij in {Gij} by TARSKI:def 1;
    then A8: Gij in L~LS by A4,XBOOLE_0:def 3;
      Gik in {Gik} by TARSKI:def 1;
    then A9: Gik in L~US by A3,XBOOLE_0:def 3;
    A10: 1 <= j & j <= width G by A2,AXIOMS:22;
    A11: 1 <= k & k <= width G by A2,AXIOMS:22;
    A12: [i1,j] in Indices G by A6,A10,GOBOARD7:10;
    A13: [i2,k] in Indices G by A6,A11,GOBOARD7:10;
    A14: [i1,k] in Indices G by A6,A11,GOBOARD7:10;
    A15: US is_sequence_on G by JORDAN1G:4;
    A16: LS is_sequence_on G by JORDAN1G:5;
    set go = R_Cut(US,Gik);
    set do = L_Cut(LS,Gij);
    A17: len G = width G by JORDAN8:def 1;
    A18: len US >= 3 by JORDAN1E:19;
    then len US >= 1 by AXIOMS:22;
    then 1 in dom US by FINSEQ_3:27;
    then A19: US.1 = US/.1 by FINSEQ_4:def 4
       .= Wmin by JORDAN1F:5;
    A20: Wmin`1 = Wbo by PSCOMP_1:84
       .= G*(1,k)`1 by A11,A17,JORDAN1A:94;
      len G >= 4 by JORDAN8:13;
    then A21: len G >= 1 by AXIOMS:22;
    then A22: [1,k] in Indices G by A11,GOBOARD7:10;
    then A23: Gik <> US.1 by A1,A13,A19,A20,JORDAN1G:7;
    then reconsider go as being_S-Seq FinSequence of TOP-REAL 2
                                                            by A9,JORDAN3:70;
    A24: [1,j] in Indices G by A10,A21,GOBOARD7:10;
    A25: len LS >= 1+2 by JORDAN1E:19;
    then len LS >= 1 by AXIOMS:22;
    then A26: 1 in dom LS & len LS in dom LS by FINSEQ_3:27;
    then A27: LS.len LS = LS/.len LS by FINSEQ_4:def 4
       .= Wmin by JORDAN1F:8;
    A28: Wmin`1 = Wbo by PSCOMP_1:84
       .= G*(1,k)`1 by A11,A17,JORDAN1A:94;
    A29: [i1,j] in Indices G by A6,A10,GOBOARD7:10;
    then A30: Gij <> LS.len LS by A1,A22,A27,A28,JORDAN1G:7;
    then reconsider do as being_S-Seq FinSequence of TOP-REAL 2
                                                            by A8,JORDAN3:69;
    A31: [len G,k] in Indices G by A11,A21,GOBOARD7:10;
    A32: LS.1 = LS/.1 by A26,FINSEQ_4:def 4
       .= Emax by JORDAN1F:6;
      Emax`1 = Ebo by PSCOMP_1:104
       .= G*(len G,k)`1 by A11,A17,JORDAN1A:92;
    then A33: Gij <> LS.1 by A1,A29,A31,A32,JORDAN1G:7;
    A34: len go >= 1+1 by TOPREAL1:def 10;
    A35: Gik in rng US by A6,A9,A11,A15,JORDAN1J:40;
    then A36: go is_sequence_on G by A15,JORDAN1J:38;
    A37: len do >= 1+1 by TOPREAL1:def 10;
    A38: Gij in rng LS by A6,A8,A10,A16,JORDAN1J:40;
    then A39: do is_sequence_on G by A16,JORDAN1J:39;
    reconsider go as non constant s.c.c.
            (being_S-Seq FinSequence of TOP-REAL 2) by A34,A36,JGRAPH_1:16,
JORDAN8:8;
    reconsider do as non constant s.c.c.
            (being_S-Seq FinSequence of TOP-REAL 2) by A37,A39,JGRAPH_1:16,
JORDAN8:8;
    A40: len go > 1 by A34,NAT_1:38;
    then A41: len go in dom go by FINSEQ_3:27;
    then A42: go/.len go = go.len go by FINSEQ_4:def 4
       .= Gik by A9,JORDAN3:59;
      len do >= 1 by A37,AXIOMS:22;
    then 1 in dom do by FINSEQ_3:27;
    then A43: do/.1 = do.1 by FINSEQ_4:def 4
       .= Gij by A8,JORDAN3:58;
    reconsider m = len go - 1 as Nat by A41,FINSEQ_3:28;
    A44: m+1 = len go by XCMPLX_1:27;
    then A45: len go-'1 = m by BINARITH:39;
    A46: LSeg(go,m) c= L~go by TOPREAL3:26;
    A47: L~go c= L~US by A9,JORDAN3:76;
    then LSeg(go,m) c= L~US by A46,XBOOLE_1:1;
    then A48: LSeg(go,m) /\ L~<*Gik,Gi1k,Gij*> c= {Gik} by A3,A7,XBOOLE_1:26;
      m >= 1 by A34,REAL_1:84;
    then A49: LSeg(go,m) = LSeg(go/.m,Gik) by A42,A44,TOPREAL1:def 5;
      {Gik} c= LSeg(go,m) /\ L~<*Gik,Gi1k,Gij*>
    proof
     let x be set;
     assume x in {Gik};
     then A50: x = Gik by TARSKI:def 1;
     A51: Gik in LSeg(go,m) by A49,TOPREAL1:6;
       Gik in LSeg(Gik,Gi1k) by TOPREAL1:6;
     then Gik in LSeg(Gik,Gi1k) \/ LSeg(Gi1k,Gij) by XBOOLE_0:def 2;
     then Gik in L~<*Gik,Gi1k,Gij*> by SPRECT_1:10;
     hence x in LSeg(go,m) /\ L~<*Gik,Gi1k,Gij*> by A50,A51,XBOOLE_0:def 3;
    end;
    then A52: LSeg(go,m) /\ L~<*Gik,Gi1k,Gij*> = {Gik} by A48,XBOOLE_0:def 10;
    A53: LSeg(do,1) c= L~do by TOPREAL3:26;
    A54: L~do c= L~LS by A8,JORDAN3:77;
    then LSeg(do,1) c= L~LS by A53,XBOOLE_1:1;
    then A55: LSeg(do,1) /\ L~<*Gik,Gi1k,Gij*> c= {Gij} by A4,A7,XBOOLE_1:26;
    A56: LSeg(do,1) = LSeg(Gij,do/.(1+1)) by A37,A43,TOPREAL1:def 5;
      {Gij} c= LSeg(do,1) /\ L~<*Gik,Gi1k,Gij*>
    proof
     let x be set;
     assume x in {Gij};
     then A57: x = Gij by TARSKI:def 1;
     A58: Gij in LSeg(do,1) by A56,TOPREAL1:6;
       Gij in LSeg(Gi1k,Gij) by TOPREAL1:6;
     then Gij in LSeg(Gik,Gi1k) \/ LSeg(Gi1k,Gij) by XBOOLE_0:def 2;
     then Gij in L~<*Gik,Gi1k,Gij*> by SPRECT_1:10;
     hence x in LSeg(do,1) /\ L~<*Gik,Gi1k,Gij*> by A57,A58,XBOOLE_0:def 3;
    end;
    then A59: L~<*Gik,Gi1k,Gij*> /\ LSeg(do,1) = {Gij} by A55,XBOOLE_0:def 10;
    A60: go/.1 = US/.1 by A9,SPRECT_3:39
       .= Wmin by JORDAN1F:5;
    then A61: go/.1 = LS/.len LS by JORDAN1F:8
       .= do/.len do by A8,JORDAN1J:35;
    A62: rng go c= L~go & rng do c= L~do by A34,A37,SPPOL_2:18;
    A63: {go/.1} c= L~go /\ L~do
    proof
     let x be set;
     assume x in {go/.1};
     then x = go/.1 by TARSKI:def 1;
     then x in rng go & x in rng do by A61,FINSEQ_6:46,REVROT_1:3;
     hence x in L~go /\ L~do by A62,XBOOLE_0:def 3;
    end;
    A64: LS.1 = LS/.1 by A26,FINSEQ_4:def 4
       .= Emax by JORDAN1F:6;
    A65: [len G,j] in Indices G by A10,A21,GOBOARD7:10;
      L~go /\ L~do c= {go/.1}
    proof
     let x be set;
     assume x in L~go /\ L~do;
     then A66: x in L~go & x in L~do by XBOOLE_0:def 3;
     then x in L~US /\ L~LS by A47,A54,XBOOLE_0:def 3;
     then x in {Wmin,Emax} by JORDAN1E:20;
     then A67: x = Wmin or x = Emax by TARSKI:def 2;
       now assume x = Emax;
      then A68: Emax = Gij by A8,A64,A66,JORDAN1E:11;
        G*(len G,j)`1 = Ebo by A10,A17,JORDAN1A:92;
      then Emax`1 <> Ebo by A1,A12,A65,A68,JORDAN1G:7;
      hence contradiction by PSCOMP_1:104;
     end;
     hence x in {go/.1} by A60,A67,TARSKI:def 1;
    end;
    then A69: L~go /\ L~do = {go/.1} by A63,XBOOLE_0:def 10;
    set W2 = go/.2;
    A70: 2 in dom go by A34,FINSEQ_3:27;
    A71: Gik..US >= 1 by A35,FINSEQ_4:31;
    A72: now assume Gij`1 = Wbo;
     then G*(1,j)`1 = G*(i1,j)`1 by A10,A17,JORDAN1A:94;
     hence contradiction by A1,A12,A24,JORDAN1G:7;
    end;
      go = mid(US,1,Gik..US) by A35,JORDAN1G:57
       .= US|(Gik..US) by A71,JORDAN3:25;
    then A73: W2 = US/.2 by A70,TOPREAL1:1;
    set pion = <*Gik,Gi1k,Gij*>;
    A74: now let n be Nat;
     assume n in dom pion;
     then n in {1,2,3} by FINSEQ_3:1,30;
     then n = 1 or n = 2 or n = 3 by ENUMSET1:def 1;
     then pion/.n = Gik or pion/.n = Gi1k or pion/.n = Gij by FINSEQ_4:27;
     hence ex i,j be Nat st [i,j] in Indices G & pion/.n = G*(i,j)
                                                             by A12,A13,A14;
    end;
    A75: Gi1k`1 = G*(i1,1)`1 by A6,A11,GOBOARD5:3
       .= Gij`1 by A6,A10,GOBOARD5:3;
      Gi1k`2 = G*(1,k)`2 by A6,A11,GOBOARD5:2
       .= Gik`2 by A6,A11,GOBOARD5:2;
    then A76: Gi1k = |[Gij`1,Gik`2]| by A75,EUCLID:57;
    A77: Gi1k in pio by TOPREAL1:6;
    A78: Gi1k in poz by TOPREAL1:6;
      now per cases;
     suppose Gij`1 <> Gik`1 & Gij`2 <> Gik`2;
      then pion is_S-Seq by A76,TOPREAL3:42;
      then consider pion1 be FinSequence of TOP-REAL 2 such that
       A79: pion1 is_sequence_on G and
       A80: pion1 is_S-Seq and
       A81: L~pion = L~pion1 and
       A82: pion/.1 = pion1/.1 and
       A83: pion/.len pion = pion1/.len pion1 and
       A84: len pion <= len pion1 by A74,GOBOARD3:2;
      reconsider pion1 as being_S-Seq FinSequence of TOP-REAL 2 by A80;
      set godo = go^'pion1^'do;
      A85: Gi1k`1 = G*(i1,1)`1 by A6,A11,GOBOARD5:3
         .= Gij`1 by A6,A10,GOBOARD5:3;
      A86: Gi1k`1 <= Gik`1 by A1,A11,JORDAN1A:39;
      then A87: W-bound poz = Gi1k`1 by SPRECT_1:62;
      A88: W-bound pio = Gij`1 by A85,SPRECT_1:62;
        W-bound (poz \/ pio) = min(W-bound poz, W-bound pio) by SPRECT_1:54
         .= Gij`1 by A85,A87,A88;
      then A89: W-bound L~pion1 = Gij`1 by A81,TOPREAL3:23;
        len Cage(C,n) > 4 by GOBOARD7:36;
      then A90: 1+1 <= len Cage(C,n) by AXIOMS:22;
      then A91: 1+1 <= len Rotate(Cage(C,n),Wmin) by REVROT_1:14;
        len (go^'pion1) >= len go by TOPREAL8:7;
      then A92: len (go^'pion1) >= 1+1 by A34,AXIOMS:22;
      then A93: len (go^'pion1) > 1+0 by NAT_1:38;
        len godo >= len (go^'pion1) by TOPREAL8:7;
      then A94: 1+1 <= len godo by A92,AXIOMS:22;
      A95: US is_sequence_on G by JORDAN1G:4;
      A96: go/.len go = pion1/.1 by A42,A82,FINSEQ_4:27;
      then A97: go^'pion1 is_sequence_on G by A36,A79,TOPREAL8:12;
      A98: (go^'pion1)/.len (go^'pion1) = pion/.len pion by A83,AMISTD_1:6
         .= pion/.3 by FINSEQ_1:62
         .= do/.1 by A43,FINSEQ_4:27;
      then A99: godo is_sequence_on G by A39,A97,TOPREAL8:12;
        LSeg(pion1,1) c= L~pion by A81,TOPREAL3:26;
      then A100: LSeg(go,len go-'1) /\ LSeg(pion1,1) c={Gik}
                                                      by A45,A52,XBOOLE_1:27;
        len pion1 >= 2+1 by A84,FINSEQ_1:62;
      then A101: len pion1 > 1+1 by NAT_1:38;
        {Gik} c= LSeg(go,m) /\ LSeg(pion1,1)
      proof
       let x be set;
       assume x in {Gik};
       then A102: x = Gik by TARSKI:def 1;
       A103: Gik in LSeg(go,m) by A49,TOPREAL1:6;
         Gik in LSeg(pion1,1) by A42,A96,A101,TOPREAL1:27;
       hence x in LSeg(go,m) /\ LSeg(pion1,1) by A102,A103,XBOOLE_0:def 3;
      end;
      then LSeg(go,len go-'1) /\ LSeg(pion1,1) = { go/.len go }
                                              by A42,A45,A100,XBOOLE_0:def 10;
      then A104: go^'pion1 is unfolded by A96,TOPREAL8:34;
        len pion1 >= 2+1 by A84,FINSEQ_1:62;
      then len pion1 > 2+0 by NAT_1:38;
      then A105: len pion1-2 >= 0 by REAL_1:84;
      A106: len (go^'pion1)-1 >= 0 by A93,REAL_1:84;
        len (go^'pion1)+1-1 = len go+len pion1-1 by GRAPH_2:13;
      then len (go^'pion1)-1 = len go+len pion1-1-1 by XCMPLX_1:26
         .= len go + len pion1-(1+1) by XCMPLX_1:36
         .= len go + (len pion1-2) by XCMPLX_1:29
         .= len go + (len pion1-'2) by A105,BINARITH:def 3;
      then A107: len (go^'pion1)-'1 = len go + (len pion1-'2)
                                                      by A106,BINARITH:def 3;
      A108: len pion1-1 >= 1 by A101,REAL_1:84;
      then A109: len pion1-1 >= 0 by AXIOMS:22;
      then A110: len pion1-'1 = len pion1-1 by BINARITH:def 3;
      A111: len pion1-'2+1 = len pion1-2+1 by A105,BINARITH:def 3
         .= len pion1-(2-1) by XCMPLX_1:37
         .= len pion1-'1 by A109,BINARITH:def 3;
        len pion1-1+1 <= len pion1 by XCMPLX_1:27;
      then A112: len pion1-'1 < len pion1 by A110,NAT_1:38;
        LSeg(pion1,len pion1-'1) c= L~pion by A81,TOPREAL3:26;
      then A113: LSeg(pion1,len pion1-'1) /\ LSeg(do,1) c= {Gij}
                                                          by A59,XBOOLE_1:27;
        {Gij} c= LSeg(pion1,len pion1-'1) /\ LSeg(do,1)
      proof
       let x be set;
       assume x in {Gij};
       then A114: x = Gij by TARSKI:def 1;
       A115: Gij in LSeg(do,1) by A56,TOPREAL1:6;
       A116: len pion1-'1+1 = len pion1 by A110,XCMPLX_1:27;
       then pion1/.(len pion1-'1+1) = pion/.3 by A83,FINSEQ_1:62
          .= Gij by FINSEQ_4:27;
       then Gij in LSeg(pion1,len pion1-'1) by A108,A110,A116,TOPREAL1:27;
       hence x in LSeg(pion1,len pion1-'1) /\ LSeg(do,1)
                                                   by A114,A115,XBOOLE_0:def 3;
      end;
      then LSeg(pion1,len pion1-'1) /\ LSeg(do,1) = {Gij}
                                                     by A113,XBOOLE_0:def 10;
      then A117: LSeg(go^'pion1,len go+(len pion1-'2)) /\ LSeg(do,1) =
         {(go^'pion1)/.len (go^'pion1)} by A43,A96,A98,A111,A112,TOPREAL8:31;
      A118: (go^'pion1) is non trivial by A92,SPPOL_1:2;
      A119: rng pion1 c= L~pion1 by A101,SPPOL_2:18;
      A120: {pion1/.1} c= L~go /\ L~pion1
      proof
       let x be set;
       assume x in {pion1/.1};
       then x = pion1/.1 by TARSKI:def 1;
       then x in rng go & x in rng pion1 by A96,FINSEQ_6:46,REVROT_1:3;
       hence x in L~go /\ L~pion1 by A62,A119,XBOOLE_0:def 3;
      end;
        L~go /\ L~pion1 c= {pion1/.1}
      proof
       let x be set;
       assume x in L~go /\ L~pion1;
       then x in L~go & x in L~pion1 by XBOOLE_0:def 3;
       hence x in {pion1/.1} by A3,A7,A42,A47,A81,A96,XBOOLE_0:def 3;
      end;
      then A121: L~go /\ L~pion1 = {pion1/.1} by A120,XBOOLE_0:def 10;
      then A122: (go^'pion1) is s.n.c. by A96,JORDAN1J:54;
        rng go /\ rng pion1 c= {pion1/.1} by A62,A119,A121,XBOOLE_1:27;
      then A123: go^'pion1 is one-to-one by JORDAN1J:55;
      A124: pion/.len pion = pion/.3 by FINSEQ_1:62
         .= do/.1 by A43,FINSEQ_4:27;
      A125: {pion1/.len pion1} c= L~do /\ L~pion1
      proof
       let x be set;
       assume x in {pion1/.len pion1};
       then x = pion1/.len pion1 by TARSKI:def 1;
       then x in rng do & x in rng pion1 by A83,A124,FINSEQ_6:46,REVROT_1:3;
       hence x in L~do /\ L~pion1 by A62,A119,XBOOLE_0:def 3;
      end;
        L~do /\ L~pion1 c= {pion1/.len pion1}
      proof
       let x be set;
       assume x in L~do /\ L~pion1;
       then x in L~do & x in L~pion1 by XBOOLE_0:def 3;
       hence x in {pion1/.len pion1}
          by A4,A7,A43,A54,A81,A83,A124,XBOOLE_0:def 3;
      end;
      then A126: L~do /\ L~pion1 = {pion1/.len pion1} by A125,XBOOLE_0:def 10;
      A127: L~(go^'pion1) /\ L~do = (L~go \/ L~pion1) /\ L~do by A96,TOPREAL8:
35
         .= {go/.1} \/ {do/.1} by A69,A83,A124,A126,XBOOLE_1:23
         .= {(go^'pion1)/.1} \/ {do/.1} by AMISTD_1:5
         .= {(go^'pion1)/.1,do/.1} by ENUMSET1:41;
        do/.len do = (go^'pion1)/.1 by A61,AMISTD_1:5;
      then reconsider godo as non constant standard special_circular_sequence
          by A94,A98,A99,A104,A107,A117,A118,A122,A123,A127,JORDAN8:7,8,
TOPREAL8:11,33,34;
      A128: UA is_an_arc_of W-min C,E-max C by JORDAN6:def 8;
      then A129: UA is connected by JORDAN6:11;
      A130: W-min C in UA & E-max C in UA by A128,TOPREAL1:4;
      set ff = Rotate(Cage(C,n),Wmin);
        Wmin in rng Cage(C,n) by SPRECT_2:47;
      then A131: ff/.1 = Wmin by FINSEQ_6:98;
      A132: L~ff = L~Cage(C,n) by REVROT_1:33;
      then A133: (W-max L~ff)..ff > 1 by A131,SPRECT_5:23;
        (W-max L~ff)..ff <= (N-min L~ff)..ff by A131,A132,SPRECT_5:24;
      then A134: (N-min L~ff)..ff > 1 by A133,AXIOMS:22;
        (N-min L~ff)..ff < (N-max L~ff)..ff by A131,A132,SPRECT_5:25;
      then A135: (N-max L~ff)..ff > 1 by A134,AXIOMS:22;
        (N-max L~ff)..ff <= (E-max L~ff)..ff by A131,A132,SPRECT_5:26;
      then A136: Emax..ff > 1 by A132,A135,AXIOMS:22;
      A137: now assume A138: Gik..US <= 1;
         Gik..US >= 1 by A35,FINSEQ_4:31;
       then Gik..US = 1 by A138,AXIOMS:21;
       then Gik = US/.1 by A35,FINSEQ_5:41;
       hence contradiction by A19,A23,JORDAN1F:5;
      end;
      A139: Cage(C,n) is_sequence_on G by JORDAN9:def 1;
      then A140: ff is_sequence_on G by REVROT_1:34;
      A141: right_cell(godo,1,G)\L~godo c= RightComp godo by A94,A99,JORDAN9:29
;
      A142: L~godo = L~(go^'pion1) \/ L~do by A98,TOPREAL8:35
         .= L~go \/ L~pion1 \/ L~do by A96,TOPREAL8:35;
        L~Cage(C,n) = L~US \/ L~LS by JORDAN1E:17;
      then A143: L~US c= L~Cage(C,n) & L~LS c= L~Cage(C,n) by XBOOLE_1:7;
      then A144: L~go c=L~Cage(C,n) & L~do c=L~Cage(C,n)
                                                     by A47,A54,XBOOLE_1:1;
      A145: W-min C in C by SPRECT_1:15;
      A146: now assume W-min C in L~godo;
       then A147: W-min C in L~go \/ L~pion1 or W-min C in L~do
                                                       by A142,XBOOLE_0:def 2;
       per cases by A147,XBOOLE_0:def 2;
        suppose W-min C in L~go;
         then C meets L~Cage(C,n) by A144,A145,XBOOLE_0:3;
         hence contradiction by JORDAN10:5;
        suppose W-min C in L~pion1;
         hence contradiction by A5,A7,A81,A130,XBOOLE_0:3;
        suppose W-min C in L~do;
         then C meets L~Cage(C,n) by A144,A145,XBOOLE_0:3;
         hence contradiction by JORDAN10:5;
      end;
        right_cell(Rotate(Cage(C,n),Wmin),1) =
            right_cell(ff,1,GoB ff) by A91,JORDAN1H:29
         .= right_cell(ff,1,GoB Cage(C,n)) by REVROT_1:28
         .= right_cell(ff,1,G) by JORDAN1H:52
         .= right_cell(ff-:Emax,1,G) by A136,A140,JORDAN1J:53
         .= right_cell(US,1,G) by JORDAN1E:def 1
         .= right_cell(R_Cut(US,Gik),1,G) by A35,A95,A137,JORDAN1J:52
         .= right_cell(go^'pion1,1,G) by A40,A97,JORDAN1J:51
         .= right_cell(godo,1,G) by A93,A99,JORDAN1J:51;
      then W-min C in right_cell(godo,1,G) by JORDAN1I:8;
      then A148: W-min C in right_cell(godo,1,G)\L~godo by A146,XBOOLE_0:def 4;
      A149: godo/.1 = (go^'pion1)/.1 by AMISTD_1:5
         .= Wmin by A60,AMISTD_1:5;
      A150: len US >= 2 by A18,AXIOMS:22;
      A151: godo/.2 = (go^'pion1)/.2 by A92,AMISTD_1:9
         .= US/.2 by A34,A73,AMISTD_1:9
         .= (US^'LS)/.2 by A150,AMISTD_1:9
         .= Rotate(Cage(C,n),Wmin)/.2 by JORDAN1E:15;
      A152: L~go \/ L~do is compact by COMPTS_1:19;
      A153: L~go \/ L~do c= L~Cage(C,n) by A144,XBOOLE_1:8;
        Wmin in rng go by A60,FINSEQ_6:46;
      then Wmin in L~go \/ L~do by A62,XBOOLE_0:def 2;
      then A154: W-min (L~go \/ L~do) = Wmin by A152,A153,JORDAN1J:21;
      A155: (W-min (L~go \/ L~do))`1 = W-bound (L~go \/ L~do) & Wmin`1 = Wbo
                                                              by PSCOMP_1:84;
        Gij`1 >= Wbo by A8,A143,PSCOMP_1:71;
      then Gij`1 > Wbo by A72,REAL_1:def 5;
      then W-min (L~go\/L~do\/L~pion1) = W-min (L~go \/ L~do)
                                              by A89,A152,A154,A155,JORDAN1J:33
;
      then A156: W-min L~godo = Wmin by A142,A154,XBOOLE_1:4;
      A157: rng godo c= L~godo by A94,SPPOL_2:18;
        2 in dom godo by A94,FINSEQ_3:27;
      then A158: godo/.2 in rng godo by PARTFUN2:4;
        godo/.2 in W-most L~Cage(C,n) by A151,JORDAN1I:27;
      then (godo/.2)`1 = (W-min L~godo)`1 by A156,PSCOMP_1:88
         .= W-bound L~godo by PSCOMP_1:84;
      then godo/.2 in W-most L~godo by A157,A158,SPRECT_2:16;
      then Rotate(godo,W-min L~godo)/.2 in W-most L~godo
                                                      by A149,A156,FINSEQ_6:95;
      then reconsider godo as clockwise_oriented non constant standard
                                    special_circular_sequence by JORDAN1I:27;
        len US in dom US by FINSEQ_5:6;
      then A159: US.len US = US/.len US by FINSEQ_4:def 4
         .= Emax by JORDAN1F:7;
      A160: E-max C in E-most C by PSCOMP_1:111;
      A161: east_halfline E-max C misses L~go
      proof
       assume east_halfline E-max C meets L~go;
       then consider p be set such that
        A162: p in east_halfline E-max C and
        A163: p in L~go by XBOOLE_0:3;
       reconsider p as Point of TOP-REAL 2 by A162;
         p in L~US by A47,A163;
       then p in east_halfline E-max C /\ L~Cage(C,n)
                    by A143,A162,XBOOLE_0:def 3;
       then A164: p`1 = Ebo by A160,JORDAN1A:104;
       then A165: p = Emax by A47,A163,JORDAN1J:46;
       then Emax = Gik by A9,A159,A163,JORDAN1J:43;
       then Gik`1 = G*(len G,k)`1 by A11,A17,A164,A165,JORDAN1A:92;
       hence contradiction by A1,A13,A31,JORDAN1G:7;
      end;
        now assume east_halfline E-max C meets L~godo;
       then A166: east_halfline E-max C meets (L~go \/ L~pion1) or
                 east_halfline E-max C meets L~do by A142,XBOOLE_1:70;
       per cases by A166,XBOOLE_1:70;
        suppose east_halfline E-max C meets L~go;
         hence contradiction by A161;
        suppose east_halfline E-max C meets L~pion1;
         then consider p be set such that
          A167: p in east_halfline E-max C and
          A168: p in L~pion1 by XBOOLE_0:3;
         reconsider p as Point of TOP-REAL 2 by A167;
         A169: now per cases by A7,A81,A168,XBOOLE_0:def 2;
          suppose p in poz;
           hence p`1 <= Gik`1 by A86,TOPREAL1:9;
          suppose p in pio;
           hence p`1 <= Gik`1 by A85,A86,GOBOARD7:5;
         end;
           i2+1 <= len G by A1,NAT_1:38;
         then i2+1-1 <= len G-1 by REAL_1:49;
         then A170: i2 <= len G-1 by XCMPLX_1:26;
         then len G-1 > 0 by A1,AXIOMS:22;
         then A171: i2 <= len G-'1 by A170,BINARITH:def 3;
           len G-'1 <= len G by GOBOARD9:2;
         then Gik`1 <= G*(len G-'1,1)`1 by A6,A11,A17,A21,A171,JORDAN1A:39;
         then p`1 <= G*(len G-'1,1)`1 by A169,AXIOMS:22;
         then p`1 <= E-bound C by A21,JORDAN8:15;
         then A172: p`1 <= (E-max C)`1 by PSCOMP_1:104;
           p`1 >= (E-max C)`1 by A167,JORDAN1A:def 3;
         then A173: p`1 = (E-max C)`1 by A172,AXIOMS:21;
           p`2 = (E-max C)`2 by A167,JORDAN1A:def 3;
         then p = E-max C by A173,TOPREAL3:11;
         hence contradiction by A5,A7,A81,A130,A168,XBOOLE_0:3;
        suppose east_halfline E-max C meets L~do;
         then consider p be set such that
          A174: p in east_halfline E-max C and
          A175: p in L~do by XBOOLE_0:3;
         reconsider p as Point of TOP-REAL 2 by A174;
           p in L~LS by A54,A175;
         then p in east_halfline E-max C /\ L~Cage(C,n)
                      by A143,A174,XBOOLE_0:def 3;
         then A176: p`1 = Ebo by A160,JORDAN1A:104;
         A177: (E-max C)`2 = p`2 by A174,JORDAN1A:def 3;
         set RC = Rotate(Cage(C,n),Emax);
         A178: E-max C in right_cell(RC,1) by JORDAN1I:9;
         A179: 1+1 <= len LS by A25,AXIOMS:22;
           LS = RC-:Wmin by JORDAN1G:26;
         then A180: LSeg(LS,1) = LSeg(RC,1) by A179,SPPOL_2:9;
         A181: L~RC = L~Cage(C,n) by REVROT_1:33;
         A182: len RC = len Cage(C,n) by REVROT_1:14;
         A183: GoB RC = GoB Cage(C,n) by REVROT_1:28
            .= G by JORDAN1H:52;
         A184: Emax in rng Cage(C,n) by SPRECT_2:50;
         A185: RC is_sequence_on G by A139,REVROT_1:34;
         A186: RC/.1 = E-max L~RC by A181,A184,FINSEQ_6:98;
         then consider ii,jj be Nat such that
          A187: [ii,jj+1] in Indices G and
          A188: [ii,jj] in Indices G and
          A189: RC/.1 = G*(ii,jj+1) and
          A190: RC/.(1+1) = G*(ii,jj) by A90,A182,A185,JORDAN1I:25;
         consider jj2 be Nat such that
          A191: 1 <= jj2 & jj2 <= width G and
          A192: Emax = G*(len G,jj2) by JORDAN1D:29;
         A193: len G >= 4 by JORDAN8:13;
         then len G >= 1 by AXIOMS:22;
         then [len G,jj2] in Indices G by A191,GOBOARD7:10;
         then A194: ii = len G by A181,A186,A187,A189,A192,GOBOARD1:21;
         A195: 1 <= ii & ii <= len G & 1 <= jj+1 & jj+1 <= width G
                                                           by A187,GOBOARD5:1;
         A196: 1 <= ii & ii <= len G & 1 <= jj & jj <= width G
                                                           by A188,GOBOARD5:1;
         A197: ii+1 <> ii by NAT_1:38;
           jj+1 > jj by NAT_1:38;
         then jj+1+1 <> jj by NAT_1:38;
         then A198: right_cell(RC,1) = cell(G,ii-'1,jj)
                           by A90,A182,A183,A187,A188,A189,A190,A197,GOBOARD5:
def 6;
         A199: ii-'1+1 = ii by A195,AMI_5:4;
         A200: ii-1 >= 4-1 by A193,A194,REAL_1:49;
         then A201: ii-1 >= 1 by AXIOMS:22;
           ii-1 >= 0 by A200,AXIOMS:22;
         then A202: 1 <= ii-'1 by A201,BINARITH:def 3;
         then A203: G*(ii-'1,jj)`2 <= p`2 & p`2 <= G*(ii-'1,jj+1)`2
             by A177,A178,A195,A196,A198,A199,JORDAN9:19;
         A204: ii-'1 < len G by A195,A199,NAT_1:38;
         then A205: G*(ii-'1,jj)`2 = G*(1,jj)`2 by A196,A202,GOBOARD5:2
            .= G*(ii,jj)`2 by A196,GOBOARD5:2;
         A206: G*(ii-'1,jj+1)`2 = G*(1,jj+1)`2 by A195,A202,A204,GOBOARD5:2
            .= G*(ii,jj+1)`2 by A195,GOBOARD5:2;
           G*(len G,jj)`1 = Ebo & Ebo = G*(len G,jj+1)`1
                                                by A17,A195,A196,JORDAN1A:92;
         then p in LSeg(RC/.1,RC/.(1+1))
                              by A176,A189,A190,A194,A203,A205,A206,GOBOARD7:8;
         then A207: p in LSeg(LS,1) by A90,A180,A182,TOPREAL1:def 5;
         A208: p in LSeg(do,Index(p,do)) by A175,JORDAN3:42;
         A209: do = mid(LS,Gij..LS,len LS) by A38,JORDAN1J:37;
         A210: 1<=Gij..LS & Gij..LS<=len LS by A38,FINSEQ_4:31;
           Gij..LS <> len LS by A30,A38,FINSEQ_4:29;
         then A211: Gij..LS < len LS by A210,REAL_1:def 5;
         A212: 1<=Index(p,do) & Index(p,do) < len do by A175,JORDAN3:41;
         A213: Index(Gij,LS)+1 = Gij..LS by A33,A38,JORDAN1J:56;
         consider t be Nat such that
          A214: t in dom LS and
          A215: LS.t = Gij by A38,FINSEQ_2:11;
         A216: 1 <= t & t <= len LS by A214,FINSEQ_3:27;
         then 1 < t by A33,A215,REAL_1:def 5;
         then Index(Gij,LS)+1 = t by A215,A216,JORDAN3:45;
         then A217: len L_Cut(LS,Gij) = len LS-Index(Gij,LS)
                                                  by A8,A30,A215,JORDAN3:61;
         set tt = Index(p,do)+(Gij..LS)-'1;
         A218: 1<=Index(Gij,LS) & 0+Index(Gij,LS) < len LS by A8,JORDAN3:41;
         then A219: len LS-Index(Gij,LS) > 0 by REAL_1:86;
         then Index(p,do) < len LS-'Index(Gij,LS) by A212,A217,BINARITH:def 3;
         then Index(p,do)+1 <= len LS-'Index(Gij,LS) by NAT_1:38;
         then Index(p,do) <= len LS-'Index(Gij,LS)-1 by REAL_1:84;
         then Index(p,do) <= len LS-Index(Gij,LS)-1 by A219,BINARITH:def 3;
         then A220: Index(p,do) <= len LS-Gij..LS by A213,XCMPLX_1:36;
         then len LS-Gij..LS >= 1 by A212,AXIOMS:22;
         then len LS-Gij..LS >= 0 by AXIOMS:22;
         then Index(p,do) <= len LS-'Gij..LS by A220,BINARITH:def 3;
         then Index(p,do) < len LS-'(Gij..LS)+1 by NAT_1:38;
         then A221: LSeg(mid(LS,Gij..LS,len LS),Index(p,do)) =
              LSeg(LS,Index(p,do)+(Gij..LS)-'1) by A210,A211,A212,JORDAN4:31;
         A222: 1+1 <= Gij..LS by A213,A218,REAL_1:55;
         then Index(p,do)+Gij..LS >= 1+1+1 by A212,REAL_1:55;
         then A223: Index(p,do)+Gij..LS-1 >= 1+1+1-1 by REAL_1:49;
         then A224: Index(p,do)+Gij..LS-1 >= 0 by AXIOMS:22;
         then A225: tt >= 1+1 by A223,BINARITH:def 3;
         A226: 2 in dom LS by A179,FINSEQ_3:27;
           now per cases by A225,REAL_1:def 5;
          suppose tt > 1+1;
           then LSeg(LS,1) misses LSeg(LS,tt) by TOPREAL1:def 9;
           hence contradiction by A207,A208,A209,A221,XBOOLE_0:3;
          suppose A227: tt = 1+1;
           then LSeg(LS,1) /\ LSeg(LS,tt) = {LS/.2} by A25,TOPREAL1:def 8;
           then p in {LS/.2} by A207,A208,A209,A221,XBOOLE_0:def 3;
           then A228: p = LS/.2 by TARSKI:def 1;
           then A229: p..LS = 2 by A226,FINSEQ_5:44;
             1+1 = Index(p,do)+(Gij..LS)-1 by A224,A227,BINARITH:def 3;
           then 1+1+1 = Index(p,do)+(Gij..LS) by XCMPLX_1:27;
           then A230: Gij..LS = 2 by A212,A222,JORDAN1E:10;
             p in rng LS by A226,A228,PARTFUN2:4;
           then p = Gij by A38,A229,A230,FINSEQ_5:10;
           then Gij`1 = Ebo by A228,JORDAN1G:40;
           then Gij`1 = G*(len G,j)`1 by A10,A17,JORDAN1A:92;
           hence contradiction by A1,A12,A65,JORDAN1G:7;
         end;
         hence contradiction;
      end;
      then east_halfline E-max C c= (L~godo)` by SUBSET_1:43;
      then consider W be Subset of TOP-REAL 2 such that
       A231: W is_a_component_of (L~godo)` and
       A232: east_halfline E-max C c= W by GOBOARD9:5;
        east_halfline E-max C is not Bounded by JORDAN1C:9;
      then W is not Bounded by A232,JORDAN2C:16;
      then W is_outside_component_of L~godo by A231,JORDAN2C:def 4;
      then W c= UBD L~godo by JORDAN2C:27;
      then A233: east_halfline E-max C c= UBD L~godo by A232,XBOOLE_1:1;
        E-max C in east_halfline E-max C by JORDAN1C:7;
      then E-max C in UBD L~godo by A233;
      then E-max C in LeftComp godo by GOBRD14:46;
      then UA meets L~godo by A129,A130,A141,A148,JORDAN1J:36;
      then A234: UA meets (L~go \/ L~pion1) or UA meets L~do by A142,XBOOLE_1:
70;
      A235: UA c= C by JORDAN1A:16;
        now per cases by A234,XBOOLE_1:70;
       suppose UA meets L~go;
        then UA meets L~Cage(C,n) by A144,XBOOLE_1:63;
        then C meets L~Cage(C,n) by A235,XBOOLE_1:63;
        hence contradiction by JORDAN10:5;
       suppose UA meets L~pion1;
        hence contradiction by A5,A7,A81;
       suppose UA meets L~do;
        then UA meets L~Cage(C,n) by A144,XBOOLE_1:63;
        then C meets L~Cage(C,n) by A235,XBOOLE_1:63;
        hence contradiction by JORDAN10:5;
      end;
      hence contradiction;
    suppose Gij`1 = Gik`1;
      then A236: i1 = i2 by A12,A13,JORDAN1G:7;
      then poz = {Gi1k} by TOPREAL1:7;
      then poz c= pio by A77,ZFMISC_1:37;
      then pio \/ poz = pio by XBOOLE_1:12;
      hence contradiction by A1,A2,A3,A4,A5,A236,JORDAN1J:59;
     suppose Gij`2 = Gik`2;
      then A237: j = k by A12,A13,JORDAN1G:6;
      then pio = {Gi1k} by TOPREAL1:7;
      then pio c= poz by A78,ZFMISC_1:37;
      then pio \/ poz = poz by XBOOLE_1:12;
      hence contradiction by A1,A2,A3,A4,A5,A237,Th31;
    end;
    hence contradiction;
   end;

  theorem Th47:
   for C be Simple_closed_curve
   for i1,i2,j,k be Nat st
    1 < i1 & i1 <= i2 & i2 < len Gauge(C,n) &
    1 <= j & j <= k & k <= width Gauge(C,n) &
    (LSeg(Gauge(C,n)*(i1,j),Gauge(C,n)*(i1,k)) \/
     LSeg(Gauge(C,n)*(i1,k),Gauge(C,n)*(i2,k))) /\ L~Upper_Seq(C,n) =
     {Gauge(C,n)*(i2,k)} &
    (LSeg(Gauge(C,n)*(i1,j),Gauge(C,n)*(i1,k)) \/
     LSeg(Gauge(C,n)*(i1,k),Gauge(C,n)*(i2,k))) /\ L~Lower_Seq(C,n) =
     {Gauge(C,n)*(i1,j)} holds
      (LSeg(Gauge(C,n)*(i1,j),Gauge(C,n)*(i1,k)) \/
       LSeg(Gauge(C,n)*(i1,k),Gauge(C,n)*(i2,k))) meets Lower_Arc C
   proof
    let C be Simple_closed_curve;
    let i1,i2,j,k be Nat;
    set G = Gauge(C,n);
    set pio = LSeg(G*(i1,j),G*(i1,k));
    set poz = LSeg(G*(i1,k),G*(i2,k));
    set US = Upper_Seq(C,n);
    set LS = Lower_Seq(C,n);
    assume that
     A1: 1 < i1 & i1 <= i2 & i2 < len G and
     A2: 1 <= j & j <= k & k <= width G and
     A3: (pio \/ poz) /\ L~US = {G*(i2,k)} and
     A4: (pio \/ poz) /\ L~LS = {G*(i1,j)} and
     A5: (pio \/ poz) misses Lower_Arc C;
    set LA = Lower_Arc C;
    set Wmin = W-min L~Cage(C,n);
    set Emax = E-max L~Cage(C,n);
    set Wbo = W-bound L~Cage(C,n);
    set Ebo = E-bound L~Cage(C,n);
    set Gij = G*(i1,j);
    set Gik = G*(i2,k);
    set Gi1k = G*(i1,k);
    A6: 1 < i1 & i1 < len G & 1 < i2 & i2 < len G by A1,AXIOMS:22;
    A7: L~<*Gik,Gi1k,Gij*> = poz \/ pio by TOPREAL3:23;
      Gij in {Gij} by TARSKI:def 1;
    then A8: Gij in L~LS by A4,XBOOLE_0:def 3;
      Gik in {Gik} by TARSKI:def 1;
    then A9: Gik in L~US by A3,XBOOLE_0:def 3;
    A10: 1 <= j & j <= width G by A2,AXIOMS:22;
    A11: 1 <= k & k <= width G by A2,AXIOMS:22;
    A12: [i1,j] in Indices G by A6,A10,GOBOARD7:10;
    A13: [i2,k] in Indices G by A6,A11,GOBOARD7:10;
    A14: [i1,k] in Indices G by A6,A11,GOBOARD7:10;
    A15: US is_sequence_on G by JORDAN1G:4;
    A16: LS is_sequence_on G by JORDAN1G:5;
    set go = R_Cut(US,Gik);
    set do = L_Cut(LS,Gij);
    A17: len G = width G by JORDAN8:def 1;
    A18: len US >= 3 by JORDAN1E:19;
    then len US >= 1 by AXIOMS:22;
    then 1 in dom US by FINSEQ_3:27;
    then A19: US.1 = US/.1 by FINSEQ_4:def 4
       .= Wmin by JORDAN1F:5;
    A20: Wmin`1 = Wbo by PSCOMP_1:84
       .= G*(1,k)`1 by A11,A17,JORDAN1A:94;
      len G >= 4 by JORDAN8:13;
    then A21: len G >= 1 by AXIOMS:22;
    then A22: [1,k] in Indices G by A11,GOBOARD7:10;
    then A23: Gik <> US.1 by A1,A13,A19,A20,JORDAN1G:7;
    then reconsider go as being_S-Seq FinSequence of TOP-REAL 2
                                                            by A9,JORDAN3:70;
    A24: [1,j] in Indices G by A10,A21,GOBOARD7:10;
    A25: len LS >= 1+2 by JORDAN1E:19;
    then len LS >= 1 by AXIOMS:22;
    then A26: 1 in dom LS & len LS in dom LS by FINSEQ_3:27;
    then A27: LS.len LS = LS/.len LS by FINSEQ_4:def 4
       .= Wmin by JORDAN1F:8;
    A28: Wmin`1 = Wbo by PSCOMP_1:84
       .= G*(1,k)`1 by A11,A17,JORDAN1A:94;
    A29: [i1,j] in Indices G by A6,A10,GOBOARD7:10;
    then A30: Gij <> LS.len LS by A1,A22,A27,A28,JORDAN1G:7;
    then reconsider do as being_S-Seq FinSequence of TOP-REAL 2
                                                            by A8,JORDAN3:69;
    A31: [len G,k] in Indices G by A11,A21,GOBOARD7:10;
    A32: LS.1 = LS/.1 by A26,FINSEQ_4:def 4
       .= Emax by JORDAN1F:6;
      Emax`1 = Ebo by PSCOMP_1:104
       .= G*(len G,k)`1 by A11,A17,JORDAN1A:92;
    then A33: Gij <> LS.1 by A1,A29,A31,A32,JORDAN1G:7;
    A34: len go >= 1+1 by TOPREAL1:def 10;
    A35: Gik in rng US by A6,A9,A11,A15,JORDAN1J:40;
    then A36: go is_sequence_on G by A15,JORDAN1J:38;
    A37: len do >= 1+1 by TOPREAL1:def 10;
    A38: Gij in rng LS by A6,A8,A10,A16,JORDAN1J:40;
    then A39: do is_sequence_on G by A16,JORDAN1J:39;
    reconsider go as non constant s.c.c.
            (being_S-Seq FinSequence of TOP-REAL 2) by A34,A36,JGRAPH_1:16,
JORDAN8:8;
    reconsider do as non constant s.c.c.
            (being_S-Seq FinSequence of TOP-REAL 2) by A37,A39,JGRAPH_1:16,
JORDAN8:8;
    A40: len go > 1 by A34,NAT_1:38;
    then A41: len go in dom go by FINSEQ_3:27;
    then A42: go/.len go = go.len go by FINSEQ_4:def 4
       .= Gik by A9,JORDAN3:59;
      len do >= 1 by A37,AXIOMS:22;
    then 1 in dom do by FINSEQ_3:27;
    then A43: do/.1 = do.1 by FINSEQ_4:def 4
       .= Gij by A8,JORDAN3:58;
    reconsider m = len go - 1 as Nat by A41,FINSEQ_3:28;
    A44: m+1 = len go by XCMPLX_1:27;
    then A45: len go-'1 = m by BINARITH:39;
    A46: LSeg(go,m) c= L~go by TOPREAL3:26;
    A47: L~go c= L~US by A9,JORDAN3:76;
    then LSeg(go,m) c= L~US by A46,XBOOLE_1:1;
    then A48: LSeg(go,m) /\ L~<*Gik,Gi1k,Gij*> c= {Gik} by A3,A7,XBOOLE_1:26;
      m >= 1 by A34,REAL_1:84;
    then A49: LSeg(go,m) = LSeg(go/.m,Gik) by A42,A44,TOPREAL1:def 5;
      {Gik} c= LSeg(go,m) /\ L~<*Gik,Gi1k,Gij*>
    proof
     let x be set;
     assume x in {Gik};
     then A50: x = Gik by TARSKI:def 1;
     A51: Gik in LSeg(go,m) by A49,TOPREAL1:6;
       Gik in LSeg(Gik,Gi1k) by TOPREAL1:6;
     then Gik in LSeg(Gik,Gi1k) \/ LSeg(Gi1k,Gij) by XBOOLE_0:def 2;
     then Gik in L~<*Gik,Gi1k,Gij*> by SPRECT_1:10;
     hence x in LSeg(go,m) /\ L~<*Gik,Gi1k,Gij*> by A50,A51,XBOOLE_0:def 3;
    end;
    then A52: LSeg(go,m) /\ L~<*Gik,Gi1k,Gij*> = {Gik} by A48,XBOOLE_0:def 10;
    A53: LSeg(do,1) c= L~do by TOPREAL3:26;
    A54: L~do c= L~LS by A8,JORDAN3:77;
    then LSeg(do,1) c= L~LS by A53,XBOOLE_1:1;
    then A55: LSeg(do,1) /\ L~<*Gik,Gi1k,Gij*> c= {Gij} by A4,A7,XBOOLE_1:26;
    A56: LSeg(do,1) = LSeg(Gij,do/.(1+1)) by A37,A43,TOPREAL1:def 5;
      {Gij} c= LSeg(do,1) /\ L~<*Gik,Gi1k,Gij*>
    proof
     let x be set;
     assume x in {Gij};
     then A57: x = Gij by TARSKI:def 1;
     A58: Gij in LSeg(do,1) by A56,TOPREAL1:6;
       Gij in LSeg(Gi1k,Gij) by TOPREAL1:6;
     then Gij in LSeg(Gik,Gi1k) \/ LSeg(Gi1k,Gij) by XBOOLE_0:def 2;
     then Gij in L~<*Gik,Gi1k,Gij*> by SPRECT_1:10;
     hence x in LSeg(do,1) /\ L~<*Gik,Gi1k,Gij*> by A57,A58,XBOOLE_0:def 3;
    end;
    then A59: L~<*Gik,Gi1k,Gij*> /\ LSeg(do,1) = {Gij} by A55,XBOOLE_0:def 10;
    A60: go/.1 = US/.1 by A9,SPRECT_3:39
       .= Wmin by JORDAN1F:5;
    then A61: go/.1 = LS/.len LS by JORDAN1F:8
       .= do/.len do by A8,JORDAN1J:35;
    A62: rng go c= L~go & rng do c= L~do by A34,A37,SPPOL_2:18;
    A63: {go/.1} c= L~go /\ L~do
    proof
     let x be set;
     assume x in {go/.1};
     then x = go/.1 by TARSKI:def 1;
     then x in rng go & x in rng do by A61,FINSEQ_6:46,REVROT_1:3;
     hence x in L~go /\ L~do by A62,XBOOLE_0:def 3;
    end;
    A64: LS.1 = LS/.1 by A26,FINSEQ_4:def 4
       .= Emax by JORDAN1F:6;
    A65: [len G,j] in Indices G by A10,A21,GOBOARD7:10;
      L~go /\ L~do c= {go/.1}
    proof
     let x be set;
     assume x in L~go /\ L~do;
     then A66: x in L~go & x in L~do by XBOOLE_0:def 3;
     then x in L~US /\ L~LS by A47,A54,XBOOLE_0:def 3;
     then x in {Wmin,Emax} by JORDAN1E:20;
     then A67: x = Wmin or x = Emax by TARSKI:def 2;
       now assume x = Emax;
      then A68: Emax = Gij by A8,A64,A66,JORDAN1E:11;
        G*(len G,j)`1 = Ebo by A10,A17,JORDAN1A:92;
      then Emax`1 <> Ebo by A1,A12,A65,A68,JORDAN1G:7;
      hence contradiction by PSCOMP_1:104;
     end;
     hence x in {go/.1} by A60,A67,TARSKI:def 1;
    end;
    then A69: L~go /\ L~do = {go/.1} by A63,XBOOLE_0:def 10;
    set W2 = go/.2;
    A70: 2 in dom go by A34,FINSEQ_3:27;
    A71: Gik..US >= 1 by A35,FINSEQ_4:31;
    A72: now assume Gij`1 = Wbo;
     then G*(1,j)`1 = G*(i1,j)`1 by A10,A17,JORDAN1A:94;
     hence contradiction by A1,A12,A24,JORDAN1G:7;
    end;
      go = mid(US,1,Gik..US) by A35,JORDAN1G:57
       .= US|(Gik..US) by A71,JORDAN3:25;
    then A73: W2 = US/.2 by A70,TOPREAL1:1;
    A74: Wmin in rng go by A60,FINSEQ_6:46;
    set pion = <*Gik,Gi1k,Gij*>;
    A75: now let n be Nat;
     assume n in dom pion;
     then n in {1,2,3} by FINSEQ_3:1,30;
     then n = 1 or n = 2 or n = 3 by ENUMSET1:def 1;
     then pion/.n = Gik or pion/.n = Gi1k or pion/.n = Gij by FINSEQ_4:27;
     hence ex i,j be Nat st [i,j] in Indices G & pion/.n = G*(i,j)
                                                             by A12,A13,A14;
    end;
    A76: Gi1k`1 = G*(i1,1)`1 by A6,A11,GOBOARD5:3
       .= Gij`1 by A6,A10,GOBOARD5:3;
      Gi1k`2 = G*(1,k)`2 by A6,A11,GOBOARD5:2
       .= Gik`2 by A6,A11,GOBOARD5:2;
    then A77: Gi1k = |[Gij`1,Gik`2]| by A76,EUCLID:57;
    A78: Gi1k in pio by TOPREAL1:6;
    A79: Gi1k in poz by TOPREAL1:6;
      now per cases;
     suppose Gij`1 <> Gik`1 & Gij`2 <> Gik`2;
      then pion is_S-Seq by A77,TOPREAL3:42;
      then consider pion1 be FinSequence of TOP-REAL 2 such that
       A80: pion1 is_sequence_on G and
       A81: pion1 is_S-Seq and
       A82: L~pion = L~pion1 and
       A83: pion/.1 = pion1/.1 and
       A84: pion/.len pion = pion1/.len pion1 and
       A85: len pion <= len pion1 by A75,GOBOARD3:2;
      reconsider pion1 as being_S-Seq FinSequence of TOP-REAL 2 by A81;
      set godo = go^'pion1^'do;
      A86: Gi1k`1 = G*(i1,1)`1 by A6,A11,GOBOARD5:3
         .= Gij`1 by A6,A10,GOBOARD5:3;
      A87: Gi1k`1 <= Gik`1 by A1,A11,JORDAN1A:39;
      then A88: W-bound poz = Gi1k`1 by SPRECT_1:62;
      A89: W-bound pio = Gij`1 by A86,SPRECT_1:62;
        W-bound (poz \/ pio) = min(W-bound poz, W-bound pio) by SPRECT_1:54
         .= Gij`1 by A86,A88,A89;
      then A90: W-bound L~pion1 = Gij`1 by A82,TOPREAL3:23;
        len Cage(C,n) > 4 by GOBOARD7:36;
      then A91: 1+1 <= len Cage(C,n) by AXIOMS:22;
      then A92: 1+1 <= len Rotate(Cage(C,n),Wmin) by REVROT_1:14;
        len (go^'pion1) >= len go by TOPREAL8:7;
      then A93: len (go^'pion1) >= 1+1 by A34,AXIOMS:22;
      then A94: len (go^'pion1) > 1+0 by NAT_1:38;
        len godo >= len (go^'pion1) by TOPREAL8:7;
      then A95: 1+1 <= len godo by A93,AXIOMS:22;
      A96: US is_sequence_on G by JORDAN1G:4;
      A97: go/.len go = pion1/.1 by A42,A83,FINSEQ_4:27;
      then A98: go^'pion1 is_sequence_on G by A36,A80,TOPREAL8:12;
      A99: (go^'pion1)/.len (go^'pion1) = pion/.len pion by A84,AMISTD_1:6
         .= pion/.3 by FINSEQ_1:62
         .= do/.1 by A43,FINSEQ_4:27;
      then A100: godo is_sequence_on G by A39,A98,TOPREAL8:12;
        LSeg(pion1,1) c= L~pion by A82,TOPREAL3:26;
      then A101: LSeg(go,len go-'1) /\ LSeg(pion1,1) c={Gik}
                                                      by A45,A52,XBOOLE_1:27;
        len pion1 >= 2+1 by A85,FINSEQ_1:62;
      then A102: len pion1 > 1+1 by NAT_1:38;
        {Gik} c= LSeg(go,m) /\ LSeg(pion1,1)
      proof
       let x be set;
       assume x in {Gik};
       then A103: x = Gik by TARSKI:def 1;
       A104: Gik in LSeg(go,m) by A49,TOPREAL1:6;
         Gik in LSeg(pion1,1) by A42,A97,A102,TOPREAL1:27;
       hence x in LSeg(go,m) /\ LSeg(pion1,1) by A103,A104,XBOOLE_0:def 3;
      end;
      then LSeg(go,len go-'1) /\ LSeg(pion1,1) = { go/.len go }
                                              by A42,A45,A101,XBOOLE_0:def 10;
      then A105: go^'pion1 is unfolded by A97,TOPREAL8:34;
        len pion1 >= 2+1 by A85,FINSEQ_1:62;
      then len pion1 > 2+0 by NAT_1:38;
      then A106: len pion1-2 >= 0 by REAL_1:84;
      A107: len (go^'pion1)-1 >= 0 by A94,REAL_1:84;
        len (go^'pion1)+1-1 = len go+len pion1-1 by GRAPH_2:13;
      then len (go^'pion1)-1 = len go+len pion1-1-1 by XCMPLX_1:26
         .= len go + len pion1-(1+1) by XCMPLX_1:36
         .= len go + (len pion1-2) by XCMPLX_1:29
         .= len go + (len pion1-'2) by A106,BINARITH:def 3;
      then A108: len (go^'pion1)-'1 = len go + (len pion1-'2)
                                                      by A107,BINARITH:def 3;
      A109: len pion1-1 >= 1 by A102,REAL_1:84;
      then A110: len pion1-1 >= 0 by AXIOMS:22;
      then A111: len pion1-'1 = len pion1-1 by BINARITH:def 3;
      A112: len pion1-'2+1 = len pion1-2+1 by A106,BINARITH:def 3
         .= len pion1-(2-1) by XCMPLX_1:37
         .= len pion1-'1 by A110,BINARITH:def 3;
        len pion1-1+1 <= len pion1 by XCMPLX_1:27;
      then A113: len pion1-'1 < len pion1 by A111,NAT_1:38;
        LSeg(pion1,len pion1-'1) c= L~pion by A82,TOPREAL3:26;
      then A114: LSeg(pion1,len pion1-'1) /\ LSeg(do,1) c= {Gij}
                                                          by A59,XBOOLE_1:27;
        {Gij} c= LSeg(pion1,len pion1-'1) /\ LSeg(do,1)
      proof
       let x be set;
       assume x in {Gij};
       then A115: x = Gij by TARSKI:def 1;
       A116: Gij in LSeg(do,1) by A56,TOPREAL1:6;
       A117: len pion1-'1+1 = len pion1 by A111,XCMPLX_1:27;
       then pion1/.(len pion1-'1+1) = pion/.3 by A84,FINSEQ_1:62
          .= Gij by FINSEQ_4:27;
       then Gij in LSeg(pion1,len pion1-'1) by A109,A111,A117,TOPREAL1:27;
       hence x in LSeg(pion1,len pion1-'1) /\ LSeg(do,1)
                                                   by A115,A116,XBOOLE_0:def 3;
      end;
      then LSeg(pion1,len pion1-'1) /\ LSeg(do,1) = {Gij}
                                                     by A114,XBOOLE_0:def 10;
      then A118: LSeg(go^'pion1,len go+(len pion1-'2)) /\ LSeg(do,1) =
         {(go^'pion1)/.len (go^'pion1)} by A43,A97,A99,A112,A113,TOPREAL8:31;
      A119: (go^'pion1) is non trivial by A93,SPPOL_1:2;
      A120: rng pion1 c= L~pion1 by A102,SPPOL_2:18;
      A121: {pion1/.1} c= L~go /\ L~pion1
      proof
       let x be set;
       assume x in {pion1/.1};
       then x = pion1/.1 by TARSKI:def 1;
       then x in rng go & x in rng pion1 by A97,FINSEQ_6:46,REVROT_1:3;
       hence x in L~go /\ L~pion1 by A62,A120,XBOOLE_0:def 3;
      end;
        L~go /\ L~pion1 c= {pion1/.1}
      proof
       let x be set;
       assume x in L~go /\ L~pion1;
       then x in L~go & x in L~pion1 by XBOOLE_0:def 3;
       hence x in {pion1/.1} by A3,A7,A42,A47,A82,A97,XBOOLE_0:def 3;
      end;
      then A122: L~go /\ L~pion1 = {pion1/.1} by A121,XBOOLE_0:def 10;
      then A123: (go^'pion1) is s.n.c. by A97,JORDAN1J:54;
        rng go /\ rng pion1 c= {pion1/.1} by A62,A120,A122,XBOOLE_1:27;
      then A124: go^'pion1 is one-to-one by JORDAN1J:55;
      A125: pion/.len pion = pion/.3 by FINSEQ_1:62
         .= do/.1 by A43,FINSEQ_4:27;
      A126: {pion1/.len pion1} c= L~do /\ L~pion1
      proof
       let x be set;
       assume x in {pion1/.len pion1};
       then x = pion1/.len pion1 by TARSKI:def 1;
       then x in rng do & x in rng pion1 by A84,A125,FINSEQ_6:46,REVROT_1:3;
       hence x in L~do /\ L~pion1 by A62,A120,XBOOLE_0:def 3;
      end;
        L~do /\ L~pion1 c= {pion1/.len pion1}
      proof
       let x be set;
       assume x in L~do /\ L~pion1;
       then x in L~do & x in L~pion1 by XBOOLE_0:def 3;
       hence x in {pion1/.len pion1}
         by A4,A7,A43,A54,A82,A84,A125,XBOOLE_0:def 3;
      end;
      then A127: L~do /\ L~pion1 = {pion1/.len pion1} by A126,XBOOLE_0:def 10;
      A128: L~(go^'pion1) /\ L~do = (L~go \/ L~pion1) /\ L~do by A97,TOPREAL8:
35
         .= {go/.1} \/ {do/.1} by A69,A84,A125,A127,XBOOLE_1:23
         .= {(go^'pion1)/.1} \/ {do/.1} by AMISTD_1:5
         .= {(go^'pion1)/.1,do/.1} by ENUMSET1:41;
        do/.len do = (go^'pion1)/.1 by A61,AMISTD_1:5;
      then reconsider godo as non constant standard special_circular_sequence
          by A95,A99,A100,A105,A108,A118,A119,A123,A124,A128,JORDAN8:7,8,
TOPREAL8:11,33,34;
      A129: LA is_an_arc_of E-max C,W-min C by JORDAN6:def 9;
      then A130: LA is connected by JORDAN6:11;
      A131: W-min C in LA & E-max C in LA by A129,TOPREAL1:4;
      set ff = Rotate(Cage(C,n),Wmin);
        Wmin in rng Cage(C,n) by SPRECT_2:47;
      then A132: ff/.1 = Wmin by FINSEQ_6:98;
      A133: L~ff = L~Cage(C,n) by REVROT_1:33;
      then A134: (W-max L~ff)..ff > 1 by A132,SPRECT_5:23;
        (W-max L~ff)..ff <= (N-min L~ff)..ff by A132,A133,SPRECT_5:24;
      then A135: (N-min L~ff)..ff > 1 by A134,AXIOMS:22;
        (N-min L~ff)..ff < (N-max L~ff)..ff by A132,A133,SPRECT_5:25;
      then A136: (N-max L~ff)..ff > 1 by A135,AXIOMS:22;
        (N-max L~ff)..ff <= (E-max L~ff)..ff by A132,A133,SPRECT_5:26;
      then A137: Emax..ff > 1 by A133,A136,AXIOMS:22;
      A138: now assume A139: Gik..US <= 1;
         Gik..US >= 1 by A35,FINSEQ_4:31;
       then Gik..US = 1 by A139,AXIOMS:21;
       then Gik = US/.1 by A35,FINSEQ_5:41;
       hence contradiction by A19,A23,JORDAN1F:5;
      end;
      A140: Cage(C,n) is_sequence_on G by JORDAN9:def 1;
      then A141: ff is_sequence_on G by REVROT_1:34;
      A142: right_cell(godo,1,G)\L~godo c= RightComp godo by A95,A100,JORDAN9:
29;
      A143: L~godo = L~(go^'pion1) \/ L~do by A99,TOPREAL8:35
         .= L~go \/ L~pion1 \/ L~do by A97,TOPREAL8:35;
        L~Cage(C,n) = L~US \/ L~LS by JORDAN1E:17;
      then A144: L~US c= L~Cage(C,n) & L~LS c= L~Cage(C,n) by XBOOLE_1:7;
      then A145: L~go c=L~Cage(C,n) & L~do c=L~Cage(C,n)
                                                     by A47,A54,XBOOLE_1:1;
      A146: W-min C in C by SPRECT_1:15;
      A147: now assume W-min C in L~godo;
       then A148: W-min C in L~go \/ L~pion1 or W-min C in L~do
                                                       by A143,XBOOLE_0:def 2;
       per cases by A148,XBOOLE_0:def 2;
        suppose W-min C in L~go;
         then C meets L~Cage(C,n) by A145,A146,XBOOLE_0:3;
         hence contradiction by JORDAN10:5;
        suppose W-min C in L~pion1;
         hence contradiction by A5,A7,A82,A131,XBOOLE_0:3;
        suppose W-min C in L~do;
         then C meets L~Cage(C,n) by A145,A146,XBOOLE_0:3;
         hence contradiction by JORDAN10:5;
      end;
        right_cell(Rotate(Cage(C,n),Wmin),1) =
            right_cell(ff,1,GoB ff) by A92,JORDAN1H:29
         .= right_cell(ff,1,GoB Cage(C,n)) by REVROT_1:28
         .= right_cell(ff,1,G) by JORDAN1H:52
         .= right_cell(ff-:Emax,1,G) by A137,A141,JORDAN1J:53
         .= right_cell(US,1,G) by JORDAN1E:def 1
         .= right_cell(R_Cut(US,Gik),1,G) by A35,A96,A138,JORDAN1J:52
         .= right_cell(go^'pion1,1,G) by A40,A98,JORDAN1J:51
         .= right_cell(godo,1,G) by A94,A100,JORDAN1J:51;
      then W-min C in right_cell(godo,1,G) by JORDAN1I:8;
      then A149: W-min C in right_cell(godo,1,G)\L~godo by A147,XBOOLE_0:def 4;
      A150: godo/.1 = (go^'pion1)/.1 by AMISTD_1:5
         .= Wmin by A60,AMISTD_1:5;
      A151: len US >= 2 by A18,AXIOMS:22;
      A152: godo/.2 = (go^'pion1)/.2 by A93,AMISTD_1:9
         .= US/.2 by A34,A73,AMISTD_1:9
         .= (US^'LS)/.2 by A151,AMISTD_1:9
         .= Rotate(Cage(C,n),Wmin)/.2 by JORDAN1E:15;
      A153: L~go \/ L~do is compact by COMPTS_1:19;
      A154: L~go \/ L~do c= L~Cage(C,n) by A145,XBOOLE_1:8;
        Wmin in L~go \/ L~do by A62,A74,XBOOLE_0:def 2;
      then A155: W-min (L~go \/ L~do) = Wmin by A153,A154,JORDAN1J:21;
      A156: (W-min (L~go \/ L~do))`1 = W-bound (L~go \/ L~do) & Wmin`1 = Wbo
                                                              by PSCOMP_1:84;
        Gij`1 >= Wbo by A8,A144,PSCOMP_1:71;
      then Gij`1 > Wbo by A72,REAL_1:def 5;
      then W-min (L~go\/L~do\/L~pion1) = W-min (L~go \/ L~do)
                                              by A90,A153,A155,A156,JORDAN1J:33
;
      then A157: W-min L~godo = Wmin by A143,A155,XBOOLE_1:4;
      A158: rng godo c= L~godo by A95,SPPOL_2:18;
        2 in dom godo by A95,FINSEQ_3:27;
      then A159: godo/.2 in rng godo by PARTFUN2:4;
        godo/.2 in W-most L~Cage(C,n) by A152,JORDAN1I:27;
      then (godo/.2)`1 = (W-min L~godo)`1 by A157,PSCOMP_1:88
         .= W-bound L~godo by PSCOMP_1:84;
      then godo/.2 in W-most L~godo by A158,A159,SPRECT_2:16;
      then Rotate(godo,W-min L~godo)/.2 in W-most L~godo
                                                      by A150,A157,FINSEQ_6:95;
      then reconsider godo as clockwise_oriented non constant standard
                                    special_circular_sequence by JORDAN1I:27;
        len US in dom US by FINSEQ_5:6;
      then A160: US.len US = US/.len US by FINSEQ_4:def 4
         .= Emax by JORDAN1F:7;
      A161: E-max C in E-most C by PSCOMP_1:111;
      A162: east_halfline E-max C misses L~go
      proof
       assume east_halfline E-max C meets L~go;
       then consider p be set such that
        A163: p in east_halfline E-max C and
        A164: p in L~go by XBOOLE_0:3;
       reconsider p as Point of TOP-REAL 2 by A163;
         p in L~US by A47,A164;
       then p in east_halfline E-max C /\ L~Cage(C,n)
                    by A144,A163,XBOOLE_0:def 3;
       then A165: p`1 = Ebo by A161,JORDAN1A:104;
       then A166: p = Emax by A47,A164,JORDAN1J:46;
       then Emax = Gik by A9,A160,A164,JORDAN1J:43;
       then Gik`1 = G*(len G,k)`1 by A11,A17,A165,A166,JORDAN1A:92;
       hence contradiction by A1,A13,A31,JORDAN1G:7;
      end;
        now assume east_halfline E-max C meets L~godo;
       then A167: east_halfline E-max C meets (L~go \/ L~pion1) or
                 east_halfline E-max C meets L~do by A143,XBOOLE_1:70;
       per cases by A167,XBOOLE_1:70;
        suppose east_halfline E-max C meets L~go;
         hence contradiction by A162;
        suppose east_halfline E-max C meets L~pion1;
         then consider p be set such that
          A168: p in east_halfline E-max C and
          A169: p in L~pion1 by XBOOLE_0:3;
         reconsider p as Point of TOP-REAL 2 by A168;
         A170: now per cases by A7,A82,A169,XBOOLE_0:def 2;
          suppose p in poz;
           hence p`1 <= Gik`1 by A87,TOPREAL1:9;
          suppose p in pio;
           hence p`1 <= Gik`1 by A86,A87,GOBOARD7:5;
         end;
           i2+1 <= len G by A1,NAT_1:38;
         then i2+1-1 <= len G-1 by REAL_1:49;
         then A171: i2 <= len G-1 by XCMPLX_1:26;
         then len G-1 > 0 by A1,AXIOMS:22;
         then A172: i2 <= len G-'1 by A171,BINARITH:def 3;
           len G-'1 <= len G by GOBOARD9:2;
         then Gik`1 <= G*(len G-'1,1)`1 by A6,A11,A17,A21,A172,JORDAN1A:39;
         then p`1 <= G*(len G-'1,1)`1 by A170,AXIOMS:22;
         then p`1 <= E-bound C by A21,JORDAN8:15;
         then A173: p`1 <= (E-max C)`1 by PSCOMP_1:104;
           p`1 >= (E-max C)`1 by A168,JORDAN1A:def 3;
         then A174: p`1 = (E-max C)`1 by A173,AXIOMS:21;
           p`2 = (E-max C)`2 by A168,JORDAN1A:def 3;
         then p = E-max C by A174,TOPREAL3:11;
         hence contradiction by A5,A7,A82,A131,A169,XBOOLE_0:3;
        suppose east_halfline E-max C meets L~do;
         then consider p be set such that
          A175: p in east_halfline E-max C and
          A176: p in L~do by XBOOLE_0:3;
         reconsider p as Point of TOP-REAL 2 by A175;
           p in L~LS by A54,A176;
         then p in east_halfline E-max C /\ L~Cage(C,n)
                   by A144,A175,XBOOLE_0:def 3;
         then A177: p`1 = Ebo by A161,JORDAN1A:104;
         A178: (E-max C)`2 = p`2 by A175,JORDAN1A:def 3;
         set RC = Rotate(Cage(C,n),Emax);
         A179: E-max C in right_cell(RC,1) by JORDAN1I:9;
         A180: 1+1 <= len LS by A25,AXIOMS:22;
           LS = RC-:Wmin by JORDAN1G:26;
         then A181: LSeg(LS,1) = LSeg(RC,1) by A180,SPPOL_2:9;
         A182: L~RC = L~Cage(C,n) by REVROT_1:33;
         A183: len RC = len Cage(C,n) by REVROT_1:14;
         A184: GoB RC = GoB Cage(C,n) by REVROT_1:28
            .= G by JORDAN1H:52;
         A185: Emax in rng Cage(C,n) by SPRECT_2:50;
         A186: RC is_sequence_on G by A140,REVROT_1:34;
         A187: RC/.1 = E-max L~RC by A182,A185,FINSEQ_6:98;
         then consider ii,jj be Nat such that
          A188: [ii,jj+1] in Indices G and
          A189: [ii,jj] in Indices G and
          A190: RC/.1 = G*(ii,jj+1) and
          A191: RC/.(1+1) = G*(ii,jj) by A91,A183,A186,JORDAN1I:25;
         consider jj2 be Nat such that
          A192: 1 <= jj2 & jj2 <= width G and
          A193: Emax = G*(len G,jj2) by JORDAN1D:29;
         A194: len G >= 4 by JORDAN8:13;
         then len G >= 1 by AXIOMS:22;
         then [len G,jj2] in Indices G by A192,GOBOARD7:10;
         then A195: ii = len G by A182,A187,A188,A190,A193,GOBOARD1:21;
         A196: 1 <= ii & ii <= len G & 1 <= jj+1 & jj+1 <= width G
                                                           by A188,GOBOARD5:1;
         A197: 1 <= ii & ii <= len G & 1 <= jj & jj <= width G
                                                           by A189,GOBOARD5:1;
         A198: ii+1 <> ii by NAT_1:38;
           jj+1 > jj by NAT_1:38;
         then jj+1+1 <> jj by NAT_1:38;
         then A199: right_cell(RC,1) = cell(G,ii-'1,jj)
                           by A91,A183,A184,A188,A189,A190,A191,A198,GOBOARD5:
def 6;
         A200: ii-'1+1 = ii by A196,AMI_5:4;
         A201: ii-1 >= 4-1 by A194,A195,REAL_1:49;
         then A202: ii-1 >= 1 by AXIOMS:22;
           ii-1 >= 0 by A201,AXIOMS:22;
         then A203: 1 <= ii-'1 by A202,BINARITH:def 3;
         then A204: G*(ii-'1,jj)`2 <= p`2 & p`2 <= G*(ii-'1,jj+1)`2
                by A178,A179,A196,A197,A199,A200,JORDAN9:19;
         A205: ii-'1 < len G by A196,A200,NAT_1:38;
         then A206: G*(ii-'1,jj)`2 = G*(1,jj)`2 by A197,A203,GOBOARD5:2
            .= G*(ii,jj)`2 by A197,GOBOARD5:2;
         A207: G*(ii-'1,jj+1)`2 = G*(1,jj+1)`2 by A196,A203,A205,GOBOARD5:2
            .= G*(ii,jj+1)`2 by A196,GOBOARD5:2;
           G*(len G,jj)`1 = Ebo & Ebo = G*(len G,jj+1)`1
                                                by A17,A196,A197,JORDAN1A:92;
         then p in LSeg(RC/.1,RC/.(1+1))
                              by A177,A190,A191,A195,A204,A206,A207,GOBOARD7:8;
         then A208: p in LSeg(LS,1) by A91,A181,A183,TOPREAL1:def 5;
         A209: p in LSeg(do,Index(p,do)) by A176,JORDAN3:42;
         A210: do = mid(LS,Gij..LS,len LS) by A38,JORDAN1J:37;
         A211: 1<=Gij..LS & Gij..LS<=len LS by A38,FINSEQ_4:31;
           Gij..LS <> len LS by A30,A38,FINSEQ_4:29;
         then A212: Gij..LS < len LS by A211,REAL_1:def 5;
         A213: 1<=Index(p,do) & Index(p,do) < len do by A176,JORDAN3:41;
         A214: Index(Gij,LS)+1 = Gij..LS by A33,A38,JORDAN1J:56;
         consider t be Nat such that
          A215: t in dom LS and
          A216: LS.t = Gij by A38,FINSEQ_2:11;
         A217: 1 <= t & t <= len LS by A215,FINSEQ_3:27;
         then 1 < t by A33,A216,REAL_1:def 5;
         then Index(Gij,LS)+1 = t by A216,A217,JORDAN3:45;
         then A218: len L_Cut(LS,Gij) = len LS-Index(Gij,LS)
                                                  by A8,A30,A216,JORDAN3:61;
         set tt = Index(p,do)+(Gij..LS)-'1;
         A219: 1<=Index(Gij,LS) & 0+Index(Gij,LS) < len LS by A8,JORDAN3:41;
         then A220: len LS-Index(Gij,LS) > 0 by REAL_1:86;
         then Index(p,do) < len LS-'Index(Gij,LS) by A213,A218,BINARITH:def 3;
         then Index(p,do)+1 <= len LS-'Index(Gij,LS) by NAT_1:38;
         then Index(p,do) <= len LS-'Index(Gij,LS)-1 by REAL_1:84;
         then Index(p,do) <= len LS-Index(Gij,LS)-1 by A220,BINARITH:def 3;
         then A221: Index(p,do) <= len LS-Gij..LS by A214,XCMPLX_1:36;
         then len LS-Gij..LS >= 1 by A213,AXIOMS:22;
         then len LS-Gij..LS >= 0 by AXIOMS:22;
         then Index(p,do) <= len LS-'Gij..LS by A221,BINARITH:def 3;
         then Index(p,do) < len LS-'(Gij..LS)+1 by NAT_1:38;
         then A222: LSeg(mid(LS,Gij..LS,len LS),Index(p,do)) =
              LSeg(LS,Index(p,do)+(Gij..LS)-'1) by A211,A212,A213,JORDAN4:31;
         A223: 1+1 <= Gij..LS by A214,A219,REAL_1:55;
         then Index(p,do)+Gij..LS >= 1+1+1 by A213,REAL_1:55;
         then A224: Index(p,do)+Gij..LS-1 >= 1+1+1-1 by REAL_1:49;
         then A225: Index(p,do)+Gij..LS-1 >= 0 by AXIOMS:22;
         then A226: tt >= 1+1 by A224,BINARITH:def 3;
         A227: 2 in dom LS by A180,FINSEQ_3:27;
           now per cases by A226,REAL_1:def 5;
          suppose tt > 1+1;
           then LSeg(LS,1) misses LSeg(LS,tt) by TOPREAL1:def 9;
           hence contradiction by A208,A209,A210,A222,XBOOLE_0:3;
          suppose A228: tt = 1+1;
           then LSeg(LS,1) /\ LSeg(LS,tt) = {LS/.2} by A25,TOPREAL1:def 8;
           then p in {LS/.2} by A208,A209,A210,A222,XBOOLE_0:def 3;
           then A229: p = LS/.2 by TARSKI:def 1;
           then A230: p..LS = 2 by A227,FINSEQ_5:44;
             1+1 = Index(p,do)+(Gij..LS)-1 by A225,A228,BINARITH:def 3;
           then 1+1+1 = Index(p,do)+(Gij..LS) by XCMPLX_1:27;
           then A231: Gij..LS = 2 by A213,A223,JORDAN1E:10;
             p in rng LS by A227,A229,PARTFUN2:4;
           then p = Gij by A38,A230,A231,FINSEQ_5:10;
           then Gij`1 = Ebo by A229,JORDAN1G:40;
           then Gij`1 = G*(len G,j)`1 by A10,A17,JORDAN1A:92;
           hence contradiction by A1,A12,A65,JORDAN1G:7;
         end;
         hence contradiction;
      end;
      then east_halfline E-max C c= (L~godo)` by SUBSET_1:43;
      then consider W be Subset of TOP-REAL 2 such that
       A232: W is_a_component_of (L~godo)` and
       A233: east_halfline E-max C c= W by GOBOARD9:5;
        east_halfline E-max C is not Bounded by JORDAN1C:9;
      then W is not Bounded by A233,JORDAN2C:16;
      then W is_outside_component_of L~godo by A232,JORDAN2C:def 4;
      then W c= UBD L~godo by JORDAN2C:27;
      then A234: east_halfline E-max C c= UBD L~godo by A233,XBOOLE_1:1;
        E-max C in east_halfline E-max C by JORDAN1C:7;
      then E-max C in UBD L~godo by A234;
      then E-max C in LeftComp godo by GOBRD14:46;
      then LA meets L~godo by A130,A131,A142,A149,JORDAN1J:36;
      then A235: LA meets (L~go \/ L~pion1) or LA meets L~do by A143,XBOOLE_1:
70;
      A236: LA c= C by JORDAN1A:16;
        now per cases by A235,XBOOLE_1:70;
       suppose LA meets L~go;
        then LA meets L~Cage(C,n) by A145,XBOOLE_1:63;
        then C meets L~Cage(C,n) by A236,XBOOLE_1:63;
        hence contradiction by JORDAN10:5;
       suppose LA meets L~pion1;
        hence contradiction by A5,A7,A82;
       suppose LA meets L~do;
        then LA meets L~Cage(C,n) by A145,XBOOLE_1:63;
        then C meets L~Cage(C,n) by A236,XBOOLE_1:63;
        hence contradiction by JORDAN10:5;
      end;
      hence contradiction;
    suppose Gij`1 = Gik`1;
      then A237: i1 = i2 by A12,A13,JORDAN1G:7;
      then poz = {Gi1k} by TOPREAL1:7;
      then poz c= pio by A78,ZFMISC_1:37;
      then pio \/ poz = pio by XBOOLE_1:12;
      hence contradiction by A1,A2,A3,A4,A5,A237,JORDAN1J:58;
     suppose Gij`2 = Gik`2;
      then A238: j = k by A12,A13,JORDAN1G:6;
      then pio = {Gi1k} by TOPREAL1:7;
      then pio c= poz by A79,ZFMISC_1:37;
      then pio \/ poz = poz by XBOOLE_1:12;
      hence contradiction by A1,A2,A3,A4,A5,A238,Th30;
    end;
    hence contradiction;
   end;

  theorem Th48:
   for C be Simple_closed_curve
   for i1,i2,j,k be Nat st
    1 < i2 & i2 <= i1 & i1 < len Gauge(C,n) &
    1 <= j & j <= k & k <= width Gauge(C,n) &
    (LSeg(Gauge(C,n)*(i1,j),Gauge(C,n)*(i1,k)) \/
     LSeg(Gauge(C,n)*(i1,k),Gauge(C,n)*(i2,k))) /\ L~Upper_Seq(C,n) =
     {Gauge(C,n)*(i2,k)} &
    (LSeg(Gauge(C,n)*(i1,j),Gauge(C,n)*(i1,k)) \/
     LSeg(Gauge(C,n)*(i1,k),Gauge(C,n)*(i2,k))) /\ L~Lower_Seq(C,n) =
     {Gauge(C,n)*(i1,j)} holds
      (LSeg(Gauge(C,n)*(i1,j),Gauge(C,n)*(i1,k)) \/
       LSeg(Gauge(C,n)*(i1,k),Gauge(C,n)*(i2,k))) meets Upper_Arc C
   proof
    let C be Simple_closed_curve;
    let i1,i2,j,k be Nat;
    set G = Gauge(C,n);
    set pio = LSeg(G*(i1,j),G*(i1,k));
    set poz = LSeg(G*(i1,k),G*(i2,k));
    set US = Upper_Seq(C,n);
    set LS = Lower_Seq(C,n);
    assume that
     A1: 1 < i2 & i2 <= i1 & i1 < len G and
     A2: 1 <= j & j <= k & k <= width G and
     A3: (pio \/ poz) /\ L~US = {G*(i2,k)} and
     A4: (pio \/ poz) /\ L~LS = {G*(i1,j)} and
     A5: (pio \/ poz) misses Upper_Arc C;
    set UA = Upper_Arc C;
    set Wmin = W-min L~Cage(C,n);
    set Emax = E-max L~Cage(C,n);
    set Wbo = W-bound L~Cage(C,n);
    set Ebo = E-bound L~Cage(C,n);
    set Gij = G*(i1,j);
    set Gik = G*(i2,k);
    set Gi1k = G*(i1,k);
    A6: 1 < i1 & i1 < len G & 1 < i2 & i2 < len G by A1,AXIOMS:22;
    A7: L~<*Gik,Gi1k,Gij*> = poz \/ pio by TOPREAL3:23;
      Gij in {Gij} by TARSKI:def 1;
    then A8: Gij in L~LS by A4,XBOOLE_0:def 3;
      Gik in {Gik} by TARSKI:def 1;
    then A9: Gik in L~US by A3,XBOOLE_0:def 3;
    A10: 1 <= j & j <= width G by A2,AXIOMS:22;
    A11: 1 <= k & k <= width G by A2,AXIOMS:22;
    A12: [i1,j] in Indices G by A6,A10,GOBOARD7:10;
    A13: [i2,k] in Indices G by A6,A11,GOBOARD7:10;
    A14: [i1,k] in Indices G by A6,A11,GOBOARD7:10;
    A15: US is_sequence_on G by JORDAN1G:4;
    A16: LS is_sequence_on G by JORDAN1G:5;
    set go = R_Cut(US,Gik);
    set do = L_Cut(LS,Gij);
    A17: len G = width G by JORDAN8:def 1;
    A18: len US >= 3 by JORDAN1E:19;
    then len US >= 1 by AXIOMS:22;
    then 1 in dom US by FINSEQ_3:27;
    then A19: US.1 = US/.1 by FINSEQ_4:def 4
       .= Wmin by JORDAN1F:5;
    A20: Wmin`1 = Wbo by PSCOMP_1:84
       .= G*(1,k)`1 by A11,A17,JORDAN1A:94;
      len G >= 4 by JORDAN8:13;
    then A21: len G >= 1 by AXIOMS:22;
    then A22: [1,k] in Indices G by A11,GOBOARD7:10;
    then A23: Gik <> US.1 by A1,A13,A19,A20,JORDAN1G:7;
    then reconsider go as being_S-Seq FinSequence of TOP-REAL 2
                                                            by A9,JORDAN3:70;
    A24: len LS >= 1+2 by JORDAN1E:19;
    then len LS >= 1 by AXIOMS:22;
    then A25: 1 in dom LS & len LS in dom LS by FINSEQ_3:27;
    then A26: LS.len LS = LS/.len LS by FINSEQ_4:def 4
       .= Wmin by JORDAN1F:8;
    A27: Wmin`1 = Wbo by PSCOMP_1:84
       .= G*(1,k)`1 by A11,A17,JORDAN1A:94;
    A28: [i1,j] in Indices G by A6,A10,GOBOARD7:10;
    then A29: Gij <> LS.len LS by A1,A22,A26,A27,JORDAN1G:7;
    then reconsider do as being_S-Seq FinSequence of TOP-REAL 2
                                                            by A8,JORDAN3:69;
    A30: [len G,k] in Indices G by A11,A21,GOBOARD7:10;
    A31: LS.1 = LS/.1 by A25,FINSEQ_4:def 4
       .= Emax by JORDAN1F:6;
      Emax`1 = Ebo by PSCOMP_1:104
       .= G*(len G,k)`1 by A11,A17,JORDAN1A:92;
    then A32: Gij <> LS.1 by A1,A28,A30,A31,JORDAN1G:7;
    A33: len go >= 1+1 by TOPREAL1:def 10;
    A34: Gik in rng US by A6,A9,A11,A15,JORDAN1J:40;
    then A35: go is_sequence_on G by A15,JORDAN1J:38;
    A36: len do >= 1+1 by TOPREAL1:def 10;
    A37: Gij in rng LS by A6,A8,A10,A16,JORDAN1J:40;
    then A38: do is_sequence_on G by A16,JORDAN1J:39;
    reconsider go as non constant s.c.c.
            (being_S-Seq FinSequence of TOP-REAL 2) by A33,A35,JGRAPH_1:16,
JORDAN8:8;
    reconsider do as non constant s.c.c.
            (being_S-Seq FinSequence of TOP-REAL 2) by A36,A38,JGRAPH_1:16,
JORDAN8:8;
    A39: len go > 1 by A33,NAT_1:38;
    then A40: len go in dom go by FINSEQ_3:27;
    then A41: go/.len go = go.len go by FINSEQ_4:def 4
       .= Gik by A9,JORDAN3:59;
      len do >= 1 by A36,AXIOMS:22;
    then 1 in dom do by FINSEQ_3:27;
    then A42: do/.1 = do.1 by FINSEQ_4:def 4
       .= Gij by A8,JORDAN3:58;
    reconsider m = len go - 1 as Nat by A40,FINSEQ_3:28;
    A43: m+1 = len go by XCMPLX_1:27;
    then A44: len go-'1 = m by BINARITH:39;
    A45: LSeg(go,m) c= L~go by TOPREAL3:26;
    A46: L~go c= L~US by A9,JORDAN3:76;
    then LSeg(go,m) c= L~US by A45,XBOOLE_1:1;
    then A47: LSeg(go,m) /\ L~<*Gik,Gi1k,Gij*> c= {Gik} by A3,A7,XBOOLE_1:26;
      m >= 1 by A33,REAL_1:84;
    then A48: LSeg(go,m) = LSeg(go/.m,Gik) by A41,A43,TOPREAL1:def 5;
      {Gik} c= LSeg(go,m) /\ L~<*Gik,Gi1k,Gij*>
    proof
     let x be set;
     assume x in {Gik};
     then A49: x = Gik by TARSKI:def 1;
     A50: Gik in LSeg(go,m) by A48,TOPREAL1:6;
       Gik in LSeg(Gik,Gi1k) by TOPREAL1:6;
     then Gik in LSeg(Gik,Gi1k) \/ LSeg(Gi1k,Gij) by XBOOLE_0:def 2;
     then Gik in L~<*Gik,Gi1k,Gij*> by SPRECT_1:10;
     hence x in LSeg(go,m) /\ L~<*Gik,Gi1k,Gij*> by A49,A50,XBOOLE_0:def 3;
    end;
    then A51: LSeg(go,m) /\ L~<*Gik,Gi1k,Gij*> = {Gik} by A47,XBOOLE_0:def 10;
    A52: LSeg(do,1) c= L~do by TOPREAL3:26;
    A53: L~do c= L~LS by A8,JORDAN3:77;
    then LSeg(do,1) c= L~LS by A52,XBOOLE_1:1;
    then A54: LSeg(do,1) /\ L~<*Gik,Gi1k,Gij*> c= {Gij} by A4,A7,XBOOLE_1:26;
    A55: LSeg(do,1) = LSeg(Gij,do/.(1+1)) by A36,A42,TOPREAL1:def 5;
      {Gij} c= LSeg(do,1) /\ L~<*Gik,Gi1k,Gij*>
    proof
     let x be set;
     assume x in {Gij};
     then A56: x = Gij by TARSKI:def 1;
     A57: Gij in LSeg(do,1) by A55,TOPREAL1:6;
       Gij in LSeg(Gi1k,Gij) by TOPREAL1:6;
     then Gij in LSeg(Gik,Gi1k) \/ LSeg(Gi1k,Gij) by XBOOLE_0:def 2;
     then Gij in L~<*Gik,Gi1k,Gij*> by SPRECT_1:10;
     hence x in LSeg(do,1) /\ L~<*Gik,Gi1k,Gij*> by A56,A57,XBOOLE_0:def 3;
    end;
    then A58: L~<*Gik,Gi1k,Gij*> /\ LSeg(do,1) = {Gij} by A54,XBOOLE_0:def 10;
    A59: go/.1 = US/.1 by A9,SPRECT_3:39
       .= Wmin by JORDAN1F:5;
    then A60: go/.1 = LS/.len LS by JORDAN1F:8
       .= do/.len do by A8,JORDAN1J:35;
    A61: rng go c= L~go & rng do c= L~do by A33,A36,SPPOL_2:18;
    A62: {go/.1} c= L~go /\ L~do
    proof
     let x be set;
     assume x in {go/.1};
     then x = go/.1 by TARSKI:def 1;
     then x in rng go & x in rng do by A60,FINSEQ_6:46,REVROT_1:3;
     hence x in L~go /\ L~do by A61,XBOOLE_0:def 3;
    end;
    A63: LS.1 = LS/.1 by A25,FINSEQ_4:def 4
       .= Emax by JORDAN1F:6;
    A64: [len G,j] in Indices G by A10,A21,GOBOARD7:10;
      L~go /\ L~do c= {go/.1}
    proof
     let x be set;
     assume x in L~go /\ L~do;
     then A65: x in L~go & x in L~do by XBOOLE_0:def 3;
     then x in L~US /\ L~LS by A46,A53,XBOOLE_0:def 3;
     then x in {Wmin,Emax} by JORDAN1E:20;
     then A66: x = Wmin or x = Emax by TARSKI:def 2;
       now assume x = Emax;
      then A67: Emax = Gij by A8,A63,A65,JORDAN1E:11;
        G*(len G,j)`1 = Ebo by A10,A17,JORDAN1A:92;
      then Emax`1 <> Ebo by A1,A12,A64,A67,JORDAN1G:7;
      hence contradiction by PSCOMP_1:104;
     end;
     hence x in {go/.1} by A59,A66,TARSKI:def 1;
    end;
    then A68: L~go /\ L~do = {go/.1} by A62,XBOOLE_0:def 10;
    set W2 = go/.2;
    A69: 2 in dom go by A33,FINSEQ_3:27;
    A70: Gik..US >= 1 by A34,FINSEQ_4:31;
    A71: now assume Gik`1 = Wbo;
     then G*(1,k)`1 = G*(i2,k)`1 by A11,A17,JORDAN1A:94;
     hence contradiction by A1,A13,A22,JORDAN1G:7;
    end;
      go = mid(US,1,Gik..US) by A34,JORDAN1G:57
       .= US|(Gik..US) by A70,JORDAN3:25;
    then A72: W2 = US/.2 by A69,TOPREAL1:1;
    A73: Wmin in rng go by A59,FINSEQ_6:46;
    set pion = <*Gik,Gi1k,Gij*>;
    A74: now let n be Nat;
     assume n in dom pion;
     then n in {1,2,3} by FINSEQ_3:1,30;
     then n = 1 or n = 2 or n = 3 by ENUMSET1:def 1;
     then pion/.n = Gik or pion/.n = Gi1k or pion/.n = Gij by FINSEQ_4:27;
     hence ex i,j be Nat st [i,j] in Indices G & pion/.n = G*(i,j)
                                                             by A12,A13,A14;
    end;
    A75: Gi1k`1 = G*(i1,1)`1 by A6,A11,GOBOARD5:3
       .= Gij`1 by A6,A10,GOBOARD5:3;
      Gi1k`2 = G*(1,k)`2 by A6,A11,GOBOARD5:2
       .= Gik`2 by A6,A11,GOBOARD5:2;
    then A76: Gi1k = |[Gij`1,Gik`2]| by A75,EUCLID:57;
    A77: Gi1k in pio by TOPREAL1:6;
    A78: Gi1k in poz by TOPREAL1:6;
      now per cases;
     suppose Gij`1 <> Gik`1 & Gij`2 <> Gik`2;
      then pion is_S-Seq by A76,TOPREAL3:42;
      then consider pion1 be FinSequence of TOP-REAL 2 such that
       A79: pion1 is_sequence_on G and
       A80: pion1 is_S-Seq and
       A81: L~pion = L~pion1 and
       A82: pion/.1 = pion1/.1 and
       A83: pion/.len pion = pion1/.len pion1 and
       A84: len pion <= len pion1 by A74,GOBOARD3:2;
      reconsider pion1 as being_S-Seq FinSequence of TOP-REAL 2 by A80;
      set godo = go^'pion1^'do;
      A85: Gi1k`1 = G*(i1,1)`1 by A6,A11,GOBOARD5:3
         .= Gij`1 by A6,A10,GOBOARD5:3;
      A86: Gik`1 <= Gi1k`1 by A1,A11,JORDAN1A:39;
      then A87: W-bound poz = Gik`1 by SPRECT_1:62;
      A88: W-bound pio = Gij`1 by A85,SPRECT_1:62;
        W-bound (poz \/ pio) = min(W-bound poz, W-bound pio) by SPRECT_1:54
         .= Gik`1 by A85,A86,A87,A88,SQUARE_1:def 1;
      then A89: W-bound L~pion1 = Gik`1 by A81,TOPREAL3:23;
        len Cage(C,n) > 4 by GOBOARD7:36;
      then A90: 1+1 <= len Cage(C,n) by AXIOMS:22;
      then A91: 1+1 <= len Rotate(Cage(C,n),Wmin) by REVROT_1:14;
        len (go^'pion1) >= len go by TOPREAL8:7;
      then A92: len (go^'pion1) >= 1+1 by A33,AXIOMS:22;
      then A93: len (go^'pion1) > 1+0 by NAT_1:38;
        len godo >= len (go^'pion1) by TOPREAL8:7;
      then A94: 1+1 <= len godo by A92,AXIOMS:22;
      A95: US is_sequence_on G by JORDAN1G:4;
      A96: go/.len go = pion1/.1 by A41,A82,FINSEQ_4:27;
      then A97: go^'pion1 is_sequence_on G by A35,A79,TOPREAL8:12;
      A98: (go^'pion1)/.len (go^'pion1) = pion/.len pion by A83,AMISTD_1:6
         .= pion/.3 by FINSEQ_1:62
         .= do/.1 by A42,FINSEQ_4:27;
      then A99: godo is_sequence_on G by A38,A97,TOPREAL8:12;
        LSeg(pion1,1) c= L~pion by A81,TOPREAL3:26;
      then A100: LSeg(go,len go-'1) /\ LSeg(pion1,1) c={Gik}
                                                      by A44,A51,XBOOLE_1:27;
        len pion1 >= 2+1 by A84,FINSEQ_1:62;
      then A101: len pion1 > 1+1 by NAT_1:38;
        {Gik} c= LSeg(go,m) /\ LSeg(pion1,1)
      proof
       let x be set;
       assume x in {Gik};
       then A102: x = Gik by TARSKI:def 1;
       A103: Gik in LSeg(go,m) by A48,TOPREAL1:6;
         Gik in LSeg(pion1,1) by A41,A96,A101,TOPREAL1:27;
       hence x in LSeg(go,m) /\ LSeg(pion1,1) by A102,A103,XBOOLE_0:def 3;
      end;
      then LSeg(go,len go-'1) /\ LSeg(pion1,1) = { go/.len go }
                                              by A41,A44,A100,XBOOLE_0:def 10;
      then A104: go^'pion1 is unfolded by A96,TOPREAL8:34;
        len pion1 >= 2+1 by A84,FINSEQ_1:62;
      then len pion1 > 2+0 by NAT_1:38;
      then A105: len pion1-2 >= 0 by REAL_1:84;
      A106: len (go^'pion1)-1 >= 0 by A93,REAL_1:84;
        len (go^'pion1)+1-1 = len go+len pion1-1 by GRAPH_2:13;
      then len (go^'pion1)-1 = len go+len pion1-1-1 by XCMPLX_1:26
         .= len go + len pion1-(1+1) by XCMPLX_1:36
         .= len go + (len pion1-2) by XCMPLX_1:29
         .= len go + (len pion1-'2) by A105,BINARITH:def 3;
      then A107: len (go^'pion1)-'1 = len go + (len pion1-'2)
                                                      by A106,BINARITH:def 3;
      A108: len pion1-1 >= 1 by A101,REAL_1:84;
      then A109: len pion1-1 >= 0 by AXIOMS:22;
      then A110: len pion1-'1 = len pion1-1 by BINARITH:def 3;
      A111: len pion1-'2+1 = len pion1-2+1 by A105,BINARITH:def 3
         .= len pion1-(2-1) by XCMPLX_1:37
         .= len pion1-'1 by A109,BINARITH:def 3;
        len pion1-1+1 <= len pion1 by XCMPLX_1:27;
      then A112: len pion1-'1 < len pion1 by A110,NAT_1:38;
        LSeg(pion1,len pion1-'1) c= L~pion by A81,TOPREAL3:26;
      then A113: LSeg(pion1,len pion1-'1) /\ LSeg(do,1) c= {Gij}
                                                          by A58,XBOOLE_1:27;
        {Gij} c= LSeg(pion1,len pion1-'1) /\ LSeg(do,1)
      proof
       let x be set;
       assume x in {Gij};
       then A114: x = Gij by TARSKI:def 1;
       A115: Gij in LSeg(do,1) by A55,TOPREAL1:6;
       A116: len pion1-'1+1 = len pion1 by A110,XCMPLX_1:27;
       then pion1/.(len pion1-'1+1) = pion/.3 by A83,FINSEQ_1:62
          .= Gij by FINSEQ_4:27;
       then Gij in LSeg(pion1,len pion1-'1) by A108,A110,A116,TOPREAL1:27;
       hence x in LSeg(pion1,len pion1-'1) /\ LSeg(do,1)
                                                   by A114,A115,XBOOLE_0:def 3;
      end;
      then LSeg(pion1,len pion1-'1) /\ LSeg(do,1) = {Gij}
                                                     by A113,XBOOLE_0:def 10;
      then A117: LSeg(go^'pion1,len go+(len pion1-'2)) /\ LSeg(do,1) =
         {(go^'pion1)/.len (go^'pion1)} by A42,A96,A98,A111,A112,TOPREAL8:31;
      A118: (go^'pion1) is non trivial by A92,SPPOL_1:2;
      A119: rng pion1 c= L~pion1 by A101,SPPOL_2:18;
      A120: {pion1/.1} c= L~go /\ L~pion1
      proof
       let x be set;
       assume x in {pion1/.1};
       then x = pion1/.1 by TARSKI:def 1;
       then x in rng go & x in rng pion1 by A96,FINSEQ_6:46,REVROT_1:3;
       hence x in L~go /\ L~pion1 by A61,A119,XBOOLE_0:def 3;
      end;
        L~go /\ L~pion1 c= {pion1/.1}
      proof
       let x be set;
       assume x in L~go /\ L~pion1;
       then x in L~go & x in L~pion1 by XBOOLE_0:def 3;
       hence x in {pion1/.1} by A3,A7,A41,A46,A81,A96,XBOOLE_0:def 3;
      end;
      then A121: L~go /\ L~pion1 = {pion1/.1} by A120,XBOOLE_0:def 10;
      then A122: (go^'pion1) is s.n.c. by A96,JORDAN1J:54;
        rng go /\ rng pion1 c= {pion1/.1} by A61,A119,A121,XBOOLE_1:27;
      then A123: go^'pion1 is one-to-one by JORDAN1J:55;
      A124: pion/.len pion = pion/.3 by FINSEQ_1:62
         .= do/.1 by A42,FINSEQ_4:27;
      A125: {pion1/.len pion1} c= L~do /\ L~pion1
      proof
       let x be set;
       assume x in {pion1/.len pion1};
       then x = pion1/.len pion1 by TARSKI:def 1;
       then x in rng do & x in rng pion1 by A83,A124,FINSEQ_6:46,REVROT_1:3;
       hence x in L~do /\ L~pion1 by A61,A119,XBOOLE_0:def 3;
      end;
        L~do /\ L~pion1 c= {pion1/.len pion1}
      proof
       let x be set;
       assume x in L~do /\ L~pion1;
       then x in L~do & x in L~pion1 by XBOOLE_0:def 3;
       hence x in {pion1/.len pion1}
              by A4,A7,A42,A53,A81,A83,A124,XBOOLE_0:def 3;
      end;
      then A126: L~do /\ L~pion1 = {pion1/.len pion1} by A125,XBOOLE_0:def 10;
      A127: L~(go^'pion1) /\ L~do = (L~go \/ L~pion1) /\ L~do by A96,TOPREAL8:
35
         .= {go/.1} \/ {do/.1} by A68,A83,A124,A126,XBOOLE_1:23
         .= {(go^'pion1)/.1} \/ {do/.1} by AMISTD_1:5
         .= {(go^'pion1)/.1,do/.1} by ENUMSET1:41;
        do/.len do = (go^'pion1)/.1 by A60,AMISTD_1:5;
      then reconsider godo as non constant standard special_circular_sequence
          by A94,A98,A99,A104,A107,A117,A118,A122,A123,A127,JORDAN8:7,8,
TOPREAL8:11,33,34;
      A128: UA is_an_arc_of W-min C,E-max C by JORDAN6:def 8;
      then A129: UA is connected by JORDAN6:11;
      A130: W-min C in UA & E-max C in UA by A128,TOPREAL1:4;
      set ff = Rotate(Cage(C,n),Wmin);
        Wmin in rng Cage(C,n) by SPRECT_2:47;
      then A131: ff/.1 = Wmin by FINSEQ_6:98;
      A132: L~ff = L~Cage(C,n) by REVROT_1:33;
      then A133: (W-max L~ff)..ff > 1 by A131,SPRECT_5:23;
        (W-max L~ff)..ff <= (N-min L~ff)..ff by A131,A132,SPRECT_5:24;
      then A134: (N-min L~ff)..ff > 1 by A133,AXIOMS:22;
        (N-min L~ff)..ff < (N-max L~ff)..ff by A131,A132,SPRECT_5:25;
      then A135: (N-max L~ff)..ff > 1 by A134,AXIOMS:22;
        (N-max L~ff)..ff <= (E-max L~ff)..ff by A131,A132,SPRECT_5:26;
      then A136: Emax..ff > 1 by A132,A135,AXIOMS:22;
      A137: now assume A138: Gik..US <= 1;
         Gik..US >= 1 by A34,FINSEQ_4:31;
       then Gik..US = 1 by A138,AXIOMS:21;
       then Gik = US/.1 by A34,FINSEQ_5:41;
       hence contradiction by A19,A23,JORDAN1F:5;
      end;
      A139: Cage(C,n) is_sequence_on G by JORDAN9:def 1;
      then A140: ff is_sequence_on G by REVROT_1:34;
      A141: right_cell(godo,1,G)\L~godo c= RightComp godo by A94,A99,JORDAN9:29
;
      A142: L~godo = L~(go^'pion1) \/ L~do by A98,TOPREAL8:35
         .= L~go \/ L~pion1 \/ L~do by A96,TOPREAL8:35;
        L~Cage(C,n) = L~US \/ L~LS by JORDAN1E:17;
      then A143: L~US c= L~Cage(C,n) & L~LS c= L~Cage(C,n) by XBOOLE_1:7;
      then A144: L~go c=L~Cage(C,n) & L~do c=L~Cage(C,n)
                                                     by A46,A53,XBOOLE_1:1;
      A145: W-min C in C by SPRECT_1:15;
      A146: now assume W-min C in L~godo;
       then A147: W-min C in L~go \/ L~pion1 or W-min C in L~do
                                                       by A142,XBOOLE_0:def 2;
       per cases by A147,XBOOLE_0:def 2;
        suppose W-min C in L~go;
         then C meets L~Cage(C,n) by A144,A145,XBOOLE_0:3;
         hence contradiction by JORDAN10:5;
        suppose W-min C in L~pion1;
         hence contradiction by A5,A7,A81,A130,XBOOLE_0:3;
        suppose W-min C in L~do;
         then C meets L~Cage(C,n) by A144,A145,XBOOLE_0:3;
         hence contradiction by JORDAN10:5;
      end;
        right_cell(Rotate(Cage(C,n),Wmin),1) =
            right_cell(ff,1,GoB ff) by A91,JORDAN1H:29
         .= right_cell(ff,1,GoB Cage(C,n)) by REVROT_1:28
         .= right_cell(ff,1,G) by JORDAN1H:52
         .= right_cell(ff-:Emax,1,G) by A136,A140,JORDAN1J:53
         .= right_cell(US,1,G) by JORDAN1E:def 1
         .= right_cell(R_Cut(US,Gik),1,G) by A34,A95,A137,JORDAN1J:52
         .= right_cell(go^'pion1,1,G) by A39,A97,JORDAN1J:51
         .= right_cell(godo,1,G) by A93,A99,JORDAN1J:51;
      then W-min C in right_cell(godo,1,G) by JORDAN1I:8;
      then A148: W-min C in right_cell(godo,1,G)\L~godo by A146,XBOOLE_0:def 4;
      A149: godo/.1 = (go^'pion1)/.1 by AMISTD_1:5
         .= Wmin by A59,AMISTD_1:5;
      A150: len US >= 2 by A18,AXIOMS:22;
      A151: godo/.2 = (go^'pion1)/.2 by A92,AMISTD_1:9
         .= US/.2 by A33,A72,AMISTD_1:9
         .= (US^'LS)/.2 by A150,AMISTD_1:9
         .= Rotate(Cage(C,n),Wmin)/.2 by JORDAN1E:15;
      A152: L~go \/ L~do is compact by COMPTS_1:19;
      A153: L~go \/ L~do c= L~Cage(C,n) by A144,XBOOLE_1:8;
        Wmin in L~go \/ L~do by A61,A73,XBOOLE_0:def 2;
      then A154: W-min (L~go \/ L~do) = Wmin by A152,A153,JORDAN1J:21;
      A155: (W-min (L~go \/ L~do))`1 = W-bound (L~go \/ L~do) & Wmin`1 = Wbo
                                                              by PSCOMP_1:84;
        Gik`1 >= Wbo by A9,A143,PSCOMP_1:71;
      then Gik`1 > Wbo by A71,REAL_1:def 5;
      then W-min (L~go\/L~do\/L~pion1) = W-min (L~go \/ L~do)
                                              by A89,A152,A154,A155,JORDAN1J:33
;
      then A156: W-min L~godo = Wmin by A142,A154,XBOOLE_1:4;
      A157: rng godo c= L~godo by A94,SPPOL_2:18;
        2 in dom godo by A94,FINSEQ_3:27;
      then A158: godo/.2 in rng godo by PARTFUN2:4;
        godo/.2 in W-most L~Cage(C,n) by A151,JORDAN1I:27;
      then (godo/.2)`1 = (W-min L~godo)`1 by A156,PSCOMP_1:88
         .= W-bound L~godo by PSCOMP_1:84;
      then godo/.2 in W-most L~godo by A157,A158,SPRECT_2:16;
      then Rotate(godo,W-min L~godo)/.2 in W-most L~godo
                                                      by A149,A156,FINSEQ_6:95;
      then reconsider godo as clockwise_oriented non constant standard
                                    special_circular_sequence by JORDAN1I:27;
        len US in dom US by FINSEQ_5:6;
      then A159: US.len US = US/.len US by FINSEQ_4:def 4
         .= Emax by JORDAN1F:7;
      A160: E-max C in E-most C by PSCOMP_1:111;
      A161: east_halfline E-max C misses L~go
      proof
       assume east_halfline E-max C meets L~go;
       then consider p be set such that
        A162: p in east_halfline E-max C and
        A163: p in L~go by XBOOLE_0:3;
       reconsider p as Point of TOP-REAL 2 by A162;
         p in L~US by A46,A163;
       then p in east_halfline E-max C /\ L~Cage(C,n)
              by A143,A162,XBOOLE_0:def 3;
       then A164: p`1 = Ebo by A160,JORDAN1A:104;
       then A165: p = Emax by A46,A163,JORDAN1J:46;
       then Emax = Gik by A9,A159,A163,JORDAN1J:43;
       then Gik`1 = G*(len G,k)`1 by A11,A17,A164,A165,JORDAN1A:92;
       hence contradiction by A1,A13,A30,JORDAN1G:7;
      end;
        now assume east_halfline E-max C meets L~godo;
       then A166: east_halfline E-max C meets (L~go \/ L~pion1) or
                 east_halfline E-max C meets L~do by A142,XBOOLE_1:70;
       per cases by A166,XBOOLE_1:70;
        suppose east_halfline E-max C meets L~go;
         hence contradiction by A161;
        suppose east_halfline E-max C meets L~pion1;
         then consider p be set such that
          A167: p in east_halfline E-max C and
          A168: p in L~pion1 by XBOOLE_0:3;
         reconsider p as Point of TOP-REAL 2 by A167;
         A169: now per cases by A7,A81,A168,XBOOLE_0:def 2;
          suppose p in poz;
           hence p`1 <= Gi1k`1 by A86,TOPREAL1:9;
          suppose p in pio;
           hence p`1 <= Gi1k`1 by A85,GOBOARD7:5;
         end;
           i1+1 <= len G by A1,NAT_1:38;
         then i1+1-1 <= len G-1 by REAL_1:49;
         then A170: i1 <= len G-1 by XCMPLX_1:26;
         then len G-1 > 0 by A1,AXIOMS:22;
         then A171: i1 <= len G-'1 by A170,BINARITH:def 3;
           len G-'1 <= len G by GOBOARD9:2;
         then Gi1k`1 <= G*(len G-'1,1)`1 by A6,A11,A17,A21,A171,JORDAN1A:39;
         then p`1 <= G*(len G-'1,1)`1 by A169,AXIOMS:22;
         then p`1 <= E-bound C by A21,JORDAN8:15;
         then A172: p`1 <= (E-max C)`1 by PSCOMP_1:104;
           p`1 >= (E-max C)`1 by A167,JORDAN1A:def 3;
         then A173: p`1 = (E-max C)`1 by A172,AXIOMS:21;
           p`2 = (E-max C)`2 by A167,JORDAN1A:def 3;
         then p = E-max C by A173,TOPREAL3:11;
         hence contradiction by A5,A7,A81,A130,A168,XBOOLE_0:3;
        suppose east_halfline E-max C meets L~do;
         then consider p be set such that
          A174: p in east_halfline E-max C and
          A175: p in L~do by XBOOLE_0:3;
         reconsider p as Point of TOP-REAL 2 by A174;
           p in L~LS by A53,A175;
         then p in east_halfline E-max C /\ L~Cage(C,n)
                     by A143,A174,XBOOLE_0:def 3;
         then A176: p`1 = Ebo by A160,JORDAN1A:104;
         A177: (E-max C)`2 = p`2 by A174,JORDAN1A:def 3;
         set RC = Rotate(Cage(C,n),Emax);
         A178: E-max C in right_cell(RC,1) by JORDAN1I:9;
         A179: 1+1 <= len LS by A24,AXIOMS:22;
           LS = RC-:Wmin by JORDAN1G:26;
         then A180: LSeg(LS,1) = LSeg(RC,1) by A179,SPPOL_2:9;
         A181: L~RC = L~Cage(C,n) by REVROT_1:33;
         A182: len RC = len Cage(C,n) by REVROT_1:14;
         A183: GoB RC = GoB Cage(C,n) by REVROT_1:28
            .= G by JORDAN1H:52;
         A184: Emax in rng Cage(C,n) by SPRECT_2:50;
         A185: RC is_sequence_on G by A139,REVROT_1:34;
         A186: RC/.1 = E-max L~RC by A181,A184,FINSEQ_6:98;
         then consider ii,jj be Nat such that
          A187: [ii,jj+1] in Indices G and
          A188: [ii,jj] in Indices G and
          A189: RC/.1 = G*(ii,jj+1) and
          A190: RC/.(1+1) = G*(ii,jj) by A90,A182,A185,JORDAN1I:25;
         consider jj2 be Nat such that
          A191: 1 <= jj2 & jj2 <= width G and
          A192: Emax = G*(len G,jj2) by JORDAN1D:29;
         A193: len G >= 4 by JORDAN8:13;
         then len G >= 1 by AXIOMS:22;
         then [len G,jj2] in Indices G by A191,GOBOARD7:10;
         then A194: ii = len G by A181,A186,A187,A189,A192,GOBOARD1:21;
         A195: 1 <= ii & ii <= len G & 1 <= jj+1 & jj+1 <= width G
                                                           by A187,GOBOARD5:1;
         A196: 1 <= ii & ii <= len G & 1 <= jj & jj <= width G
                                                           by A188,GOBOARD5:1;
         A197: ii+1 <> ii by NAT_1:38;
           jj+1 > jj by NAT_1:38;
         then jj+1+1 <> jj by NAT_1:38;
         then A198: right_cell(RC,1) = cell(G,ii-'1,jj)
                           by A90,A182,A183,A187,A188,A189,A190,A197,GOBOARD5:
def 6;
         A199: ii-'1+1 = ii by A195,AMI_5:4;
         A200: ii-1 >= 4-1 by A193,A194,REAL_1:49;
         then A201: ii-1 >= 1 by AXIOMS:22;
           ii-1 >= 0 by A200,AXIOMS:22;
         then A202: 1 <= ii-'1 by A201,BINARITH:def 3;
         then A203: G*(ii-'1,jj)`2 <= p`2 & p`2 <= G*(ii-'1,jj+1)`2
             by A177,A178,A195,A196,A198,A199,JORDAN9:19;
         A204: ii-'1 < len G by A195,A199,NAT_1:38;
         then A205: G*(ii-'1,jj)`2 = G*(1,jj)`2 by A196,A202,GOBOARD5:2
            .= G*(ii,jj)`2 by A196,GOBOARD5:2;
         A206: G*(ii-'1,jj+1)`2 = G*(1,jj+1)`2 by A195,A202,A204,GOBOARD5:2
            .= G*(ii,jj+1)`2 by A195,GOBOARD5:2;
           G*(len G,jj)`1 = Ebo & Ebo = G*(len G,jj+1)`1
                                                by A17,A195,A196,JORDAN1A:92;
         then p in LSeg(RC/.1,RC/.(1+1))
                              by A176,A189,A190,A194,A203,A205,A206,GOBOARD7:8;
         then A207: p in LSeg(LS,1) by A90,A180,A182,TOPREAL1:def 5;
         A208: p in LSeg(do,Index(p,do)) by A175,JORDAN3:42;
         A209: do = mid(LS,Gij..LS,len LS) by A37,JORDAN1J:37;
         A210: 1<=Gij..LS & Gij..LS<=len LS by A37,FINSEQ_4:31;
           Gij..LS <> len LS by A29,A37,FINSEQ_4:29;
         then A211: Gij..LS < len LS by A210,REAL_1:def 5;
         A212: 1<=Index(p,do) & Index(p,do) < len do by A175,JORDAN3:41;
         A213: Index(Gij,LS)+1 = Gij..LS by A32,A37,JORDAN1J:56;
         consider t be Nat such that
          A214: t in dom LS and
          A215: LS.t = Gij by A37,FINSEQ_2:11;
         A216: 1 <= t & t <= len LS by A214,FINSEQ_3:27;
         then 1 < t by A32,A215,REAL_1:def 5;
         then Index(Gij,LS)+1 = t by A215,A216,JORDAN3:45;
         then A217: len L_Cut(LS,Gij) = len LS-Index(Gij,LS)
                                                  by A8,A29,A215,JORDAN3:61;
         set tt = Index(p,do)+(Gij..LS)-'1;
         A218: 1<=Index(Gij,LS) & 0+Index(Gij,LS) < len LS by A8,JORDAN3:41;
         then A219: len LS-Index(Gij,LS) > 0 by REAL_1:86;
         then Index(p,do) < len LS-'Index(Gij,LS) by A212,A217,BINARITH:def 3;
         then Index(p,do)+1 <= len LS-'Index(Gij,LS) by NAT_1:38;
         then Index(p,do) <= len LS-'Index(Gij,LS)-1 by REAL_1:84;
         then Index(p,do) <= len LS-Index(Gij,LS)-1 by A219,BINARITH:def 3;
         then A220: Index(p,do) <= len LS-Gij..LS by A213,XCMPLX_1:36;
         then len LS-Gij..LS >= 1 by A212,AXIOMS:22;
         then len LS-Gij..LS >= 0 by AXIOMS:22;
         then Index(p,do) <= len LS-'Gij..LS by A220,BINARITH:def 3;
         then Index(p,do) < len LS-'(Gij..LS)+1 by NAT_1:38;
         then A221: LSeg(mid(LS,Gij..LS,len LS),Index(p,do)) =
              LSeg(LS,Index(p,do)+(Gij..LS)-'1) by A210,A211,A212,JORDAN4:31;
         A222: 1+1 <= Gij..LS by A213,A218,REAL_1:55;
         then Index(p,do)+Gij..LS >= 1+1+1 by A212,REAL_1:55;
         then A223: Index(p,do)+Gij..LS-1 >= 1+1+1-1 by REAL_1:49;
         then A224: Index(p,do)+Gij..LS-1 >= 0 by AXIOMS:22;
         then A225: tt >= 1+1 by A223,BINARITH:def 3;
         A226: 2 in dom LS by A179,FINSEQ_3:27;
           now per cases by A225,REAL_1:def 5;
          suppose tt > 1+1;
           then LSeg(LS,1) misses LSeg(LS,tt) by TOPREAL1:def 9;
           hence contradiction by A207,A208,A209,A221,XBOOLE_0:3;
          suppose A227: tt = 1+1;
           then LSeg(LS,1) /\ LSeg(LS,tt) = {LS/.2} by A24,TOPREAL1:def 8;
           then p in {LS/.2} by A207,A208,A209,A221,XBOOLE_0:def 3;
           then A228: p = LS/.2 by TARSKI:def 1;
           then A229: p..LS = 2 by A226,FINSEQ_5:44;
             1+1 = Index(p,do)+(Gij..LS)-1 by A224,A227,BINARITH:def 3;
           then 1+1+1 = Index(p,do)+(Gij..LS) by XCMPLX_1:27;
           then A230: Gij..LS = 2 by A212,A222,JORDAN1E:10;
             p in rng LS by A226,A228,PARTFUN2:4;
           then p = Gij by A37,A229,A230,FINSEQ_5:10;
           then Gij`1 = Ebo by A228,JORDAN1G:40;
           then Gij`1 = G*(len G,j)`1 by A10,A17,JORDAN1A:92;
           hence contradiction by A1,A12,A64,JORDAN1G:7;
         end;
         hence contradiction;
      end;
      then east_halfline E-max C c= (L~godo)` by SUBSET_1:43;
      then consider W be Subset of TOP-REAL 2 such that
       A231: W is_a_component_of (L~godo)` and
       A232: east_halfline E-max C c= W by GOBOARD9:5;
        east_halfline E-max C is not Bounded by JORDAN1C:9;
      then W is not Bounded by A232,JORDAN2C:16;
      then W is_outside_component_of L~godo by A231,JORDAN2C:def 4;
      then W c= UBD L~godo by JORDAN2C:27;
      then A233: east_halfline E-max C c= UBD L~godo by A232,XBOOLE_1:1;
        E-max C in east_halfline E-max C by JORDAN1C:7;
      then E-max C in UBD L~godo by A233;
      then E-max C in LeftComp godo by GOBRD14:46;
      then UA meets L~godo by A129,A130,A141,A148,JORDAN1J:36;
      then A234: UA meets (L~go \/ L~pion1) or UA meets L~do by A142,XBOOLE_1:
70;
      A235: UA c= C by JORDAN1A:16;
        now per cases by A234,XBOOLE_1:70;
       suppose UA meets L~go;
        then UA meets L~Cage(C,n) by A144,XBOOLE_1:63;
        then C meets L~Cage(C,n) by A235,XBOOLE_1:63;
        hence contradiction by JORDAN10:5;
       suppose UA meets L~pion1;
        hence contradiction by A5,A7,A81;
       suppose UA meets L~do;
        then UA meets L~Cage(C,n) by A144,XBOOLE_1:63;
        then C meets L~Cage(C,n) by A235,XBOOLE_1:63;
        hence contradiction by JORDAN10:5;
      end;
      hence contradiction;
    suppose Gij`1 = Gik`1;
      then A236: i1 = i2 by A12,A13,JORDAN1G:7;
      then poz = {Gi1k} by TOPREAL1:7;
      then poz c= pio by A77,ZFMISC_1:37;
      then pio \/ poz = pio by XBOOLE_1:12;
      hence contradiction by A1,A2,A3,A4,A5,A236,JORDAN1J:59;
     suppose Gij`2 = Gik`2;
      then A237: j = k by A12,A13,JORDAN1G:6;
      then pio = {Gi1k} by TOPREAL1:7;
      then pio c= poz by A78,ZFMISC_1:37;
      then pio \/ poz = poz by XBOOLE_1:12;
      hence contradiction by A1,A2,A3,A4,A5,A237,Th39;
    end;
    hence contradiction;
   end;

  theorem Th49:
   for C be Simple_closed_curve
   for i1,i2,j,k be Nat st
    1 < i2 & i2 <= i1 & i1 < len Gauge(C,n) &
    1 <= j & j <= k & k <= width Gauge(C,n) &
    (LSeg(Gauge(C,n)*(i1,j),Gauge(C,n)*(i1,k)) \/
     LSeg(Gauge(C,n)*(i1,k),Gauge(C,n)*(i2,k))) /\ L~Upper_Seq(C,n) =
     {Gauge(C,n)*(i2,k)} &
    (LSeg(Gauge(C,n)*(i1,j),Gauge(C,n)*(i1,k)) \/
     LSeg(Gauge(C,n)*(i1,k),Gauge(C,n)*(i2,k))) /\ L~Lower_Seq(C,n) =
     {Gauge(C,n)*(i1,j)} holds
      (LSeg(Gauge(C,n)*(i1,j),Gauge(C,n)*(i1,k)) \/
       LSeg(Gauge(C,n)*(i1,k),Gauge(C,n)*(i2,k))) meets Lower_Arc C
   proof
    let C be Simple_closed_curve;
    let i1,i2,j,k be Nat;
    set G = Gauge(C,n);
    set pio = LSeg(G*(i1,j),G*(i1,k));
    set poz = LSeg(G*(i1,k),G*(i2,k));
    set US = Upper_Seq(C,n);
    set LS = Lower_Seq(C,n);
    assume that
     A1: 1 < i2 & i2 <= i1 & i1 < len G and
     A2: 1 <= j & j <= k & k <= width G and
     A3: (pio \/ poz) /\ L~US = {G*(i2,k)} and
     A4: (pio \/ poz) /\ L~LS = {G*(i1,j)} and
     A5: (pio \/ poz) misses Lower_Arc C;
    set LA = Lower_Arc C;
    set Wmin = W-min L~Cage(C,n);
    set Emax = E-max L~Cage(C,n);
    set Wbo = W-bound L~Cage(C,n);
    set Ebo = E-bound L~Cage(C,n);
    set Gij = G*(i1,j);
    set Gik = G*(i2,k);
    set Gi1k = G*(i1,k);
    A6: 1 < i1 & i1 < len G & 1 < i2 & i2 < len G by A1,AXIOMS:22;
    A7: L~<*Gik,Gi1k,Gij*> = poz \/ pio by TOPREAL3:23;
      Gij in {Gij} by TARSKI:def 1;
    then A8: Gij in L~LS by A4,XBOOLE_0:def 3;
      Gik in {Gik} by TARSKI:def 1;
    then A9: Gik in L~US by A3,XBOOLE_0:def 3;
    A10: 1 <= j & j <= width G by A2,AXIOMS:22;
    A11: 1 <= k & k <= width G by A2,AXIOMS:22;
    A12: [i1,j] in Indices G by A6,A10,GOBOARD7:10;
    A13: [i2,k] in Indices G by A6,A11,GOBOARD7:10;
    A14: [i1,k] in Indices G by A6,A11,GOBOARD7:10;
    A15: US is_sequence_on G by JORDAN1G:4;
    A16: LS is_sequence_on G by JORDAN1G:5;
    set go = R_Cut(US,Gik);
    set do = L_Cut(LS,Gij);
    A17: len G = width G by JORDAN8:def 1;
    A18: len US >= 3 by JORDAN1E:19;
    then len US >= 1 by AXIOMS:22;
    then 1 in dom US by FINSEQ_3:27;
    then A19: US.1 = US/.1 by FINSEQ_4:def 4
       .= Wmin by JORDAN1F:5;
    A20: Wmin`1 = Wbo by PSCOMP_1:84
       .= G*(1,k)`1 by A11,A17,JORDAN1A:94;
      len G >= 4 by JORDAN8:13;
    then A21: len G >= 1 by AXIOMS:22;
    then A22: [1,k] in Indices G by A11,GOBOARD7:10;
    then A23: Gik <> US.1 by A1,A13,A19,A20,JORDAN1G:7;
    then reconsider go as being_S-Seq FinSequence of TOP-REAL 2
                                                            by A9,JORDAN3:70;
    A24: len LS >= 1+2 by JORDAN1E:19;
    then len LS >= 1 by AXIOMS:22;
    then A25: 1 in dom LS & len LS in dom LS by FINSEQ_3:27;
    then A26: LS.len LS = LS/.len LS by FINSEQ_4:def 4
       .= Wmin by JORDAN1F:8;
    A27: Wmin`1 = Wbo by PSCOMP_1:84
       .= G*(1,k)`1 by A11,A17,JORDAN1A:94;
    A28: [i1,j] in Indices G by A6,A10,GOBOARD7:10;
    then A29: Gij <> LS.len LS by A1,A22,A26,A27,JORDAN1G:7;
    then reconsider do as being_S-Seq FinSequence of TOP-REAL 2
                                                            by A8,JORDAN3:69;
    A30: [len G,k] in Indices G by A11,A21,GOBOARD7:10;
    A31: LS.1 = LS/.1 by A25,FINSEQ_4:def 4
       .= Emax by JORDAN1F:6;
      Emax`1 = Ebo by PSCOMP_1:104
       .= G*(len G,k)`1 by A11,A17,JORDAN1A:92;
    then A32: Gij <> LS.1 by A1,A28,A30,A31,JORDAN1G:7;
    A33: len go >= 1+1 by TOPREAL1:def 10;
    A34: Gik in rng US by A6,A9,A11,A15,JORDAN1J:40;
    then A35: go is_sequence_on G by A15,JORDAN1J:38;
    A36: len do >= 1+1 by TOPREAL1:def 10;
    A37: Gij in rng LS by A6,A8,A10,A16,JORDAN1J:40;
    then A38: do is_sequence_on G by A16,JORDAN1J:39;
    reconsider go as non constant s.c.c.
            (being_S-Seq FinSequence of TOP-REAL 2) by A33,A35,JGRAPH_1:16,
JORDAN8:8;
    reconsider do as non constant s.c.c.
            (being_S-Seq FinSequence of TOP-REAL 2) by A36,A38,JGRAPH_1:16,
JORDAN8:8;
    A39: len go > 1 by A33,NAT_1:38;
    then A40: len go in dom go by FINSEQ_3:27;
    then A41: go/.len go = go.len go by FINSEQ_4:def 4
       .= Gik by A9,JORDAN3:59;
      len do >= 1 by A36,AXIOMS:22;
    then 1 in dom do by FINSEQ_3:27;
    then A42: do/.1 = do.1 by FINSEQ_4:def 4
       .= Gij by A8,JORDAN3:58;
    reconsider m = len go - 1 as Nat by A40,FINSEQ_3:28;
    A43: m+1 = len go by XCMPLX_1:27;
    then A44: len go-'1 = m by BINARITH:39;
    A45: LSeg(go,m) c= L~go by TOPREAL3:26;
    A46: L~go c= L~US by A9,JORDAN3:76;
    then LSeg(go,m) c= L~US by A45,XBOOLE_1:1;
    then A47: LSeg(go,m) /\ L~<*Gik,Gi1k,Gij*> c= {Gik} by A3,A7,XBOOLE_1:26;
      m >= 1 by A33,REAL_1:84;
    then A48: LSeg(go,m) = LSeg(go/.m,Gik) by A41,A43,TOPREAL1:def 5;
      {Gik} c= LSeg(go,m) /\ L~<*Gik,Gi1k,Gij*>
    proof
     let x be set;
     assume x in {Gik};
     then A49: x = Gik by TARSKI:def 1;
     A50: Gik in LSeg(go,m) by A48,TOPREAL1:6;
       Gik in LSeg(Gik,Gi1k) by TOPREAL1:6;
     then Gik in LSeg(Gik,Gi1k) \/ LSeg(Gi1k,Gij) by XBOOLE_0:def 2;
     then Gik in L~<*Gik,Gi1k,Gij*> by SPRECT_1:10;
     hence x in LSeg(go,m) /\ L~<*Gik,Gi1k,Gij*> by A49,A50,XBOOLE_0:def 3;
    end;
    then A51: LSeg(go,m) /\ L~<*Gik,Gi1k,Gij*> = {Gik} by A47,XBOOLE_0:def 10;
    A52: LSeg(do,1) c= L~do by TOPREAL3:26;
    A53: L~do c= L~LS by A8,JORDAN3:77;
    then LSeg(do,1) c= L~LS by A52,XBOOLE_1:1;
    then A54: LSeg(do,1) /\ L~<*Gik,Gi1k,Gij*> c= {Gij} by A4,A7,XBOOLE_1:26;
    A55: LSeg(do,1) = LSeg(Gij,do/.(1+1)) by A36,A42,TOPREAL1:def 5;
      {Gij} c= LSeg(do,1) /\ L~<*Gik,Gi1k,Gij*>
    proof
     let x be set;
     assume x in {Gij};
     then A56: x = Gij by TARSKI:def 1;
     A57: Gij in LSeg(do,1) by A55,TOPREAL1:6;
       Gij in LSeg(Gi1k,Gij) by TOPREAL1:6;
     then Gij in LSeg(Gik,Gi1k) \/ LSeg(Gi1k,Gij) by XBOOLE_0:def 2;
     then Gij in L~<*Gik,Gi1k,Gij*> by SPRECT_1:10;
     hence x in LSeg(do,1) /\ L~<*Gik,Gi1k,Gij*> by A56,A57,XBOOLE_0:def 3;
    end;
    then A58: L~<*Gik,Gi1k,Gij*> /\ LSeg(do,1) = {Gij} by A54,XBOOLE_0:def 10;
    A59: go/.1 = US/.1 by A9,SPRECT_3:39
       .= Wmin by JORDAN1F:5;
    then A60: go/.1 = LS/.len LS by JORDAN1F:8
       .= do/.len do by A8,JORDAN1J:35;
    A61: rng go c= L~go & rng do c= L~do by A33,A36,SPPOL_2:18;
    A62: {go/.1} c= L~go /\ L~do
    proof
     let x be set;
     assume x in {go/.1};
     then x = go/.1 by TARSKI:def 1;
     then x in rng go & x in rng do by A60,FINSEQ_6:46,REVROT_1:3;
     hence x in L~go /\ L~do by A61,XBOOLE_0:def 3;
    end;
    A63: LS.1 = LS/.1 by A25,FINSEQ_4:def 4
       .= Emax by JORDAN1F:6;
    A64: [len G,j] in Indices G by A10,A21,GOBOARD7:10;
      L~go /\ L~do c= {go/.1}
    proof
     let x be set;
     assume x in L~go /\ L~do;
     then A65: x in L~go & x in L~do by XBOOLE_0:def 3;
     then x in L~US /\ L~LS by A46,A53,XBOOLE_0:def 3;
     then x in {Wmin,Emax} by JORDAN1E:20;
     then A66: x = Wmin or x = Emax by TARSKI:def 2;
       now assume x = Emax;
      then A67: Emax = Gij by A8,A63,A65,JORDAN1E:11;
        G*(len G,j)`1 = Ebo by A10,A17,JORDAN1A:92;
      then Emax`1 <> Ebo by A1,A12,A64,A67,JORDAN1G:7;
      hence contradiction by PSCOMP_1:104;
     end;
     hence x in {go/.1} by A59,A66,TARSKI:def 1;
    end;
    then A68: L~go /\ L~do = {go/.1} by A62,XBOOLE_0:def 10;
    set W2 = go/.2;
    A69: 2 in dom go by A33,FINSEQ_3:27;
    A70: Gik..US >= 1 by A34,FINSEQ_4:31;
    A71: now assume Gik`1 = Wbo;
     then G*(1,k)`1 = G*(i2,k)`1 by A11,A17,JORDAN1A:94;
     hence contradiction by A1,A13,A22,JORDAN1G:7;
    end;
      go = mid(US,1,Gik..US) by A34,JORDAN1G:57
       .= US|(Gik..US) by A70,JORDAN3:25;
    then A72: W2 = US/.2 by A69,TOPREAL1:1;
    A73: Wmin in rng go by A59,FINSEQ_6:46;
    set pion = <*Gik,Gi1k,Gij*>;
    A74: now let n be Nat;
     assume n in dom pion;
     then n in {1,2,3} by FINSEQ_3:1,30;
     then n = 1 or n = 2 or n = 3 by ENUMSET1:def 1;
     then pion/.n = Gik or pion/.n = Gi1k or pion/.n = Gij by FINSEQ_4:27;
     hence ex i,j be Nat st [i,j] in Indices G & pion/.n = G*(i,j)
                                                             by A12,A13,A14;
    end;
    A75: Gi1k`1 = G*(i1,1)`1 by A6,A11,GOBOARD5:3
       .= Gij`1 by A6,A10,GOBOARD5:3;
      Gi1k`2 = G*(1,k)`2 by A6,A11,GOBOARD5:2
       .= Gik`2 by A6,A11,GOBOARD5:2;
    then A76: Gi1k = |[Gij`1,Gik`2]| by A75,EUCLID:57;
    A77: Gi1k in pio by TOPREAL1:6;
    A78: Gi1k in poz by TOPREAL1:6;
      now per cases;
     suppose Gij`1 <> Gik`1 & Gij`2 <> Gik`2;
      then pion is_S-Seq by A76,TOPREAL3:42;
      then consider pion1 be FinSequence of TOP-REAL 2 such that
       A79: pion1 is_sequence_on G and
       A80: pion1 is_S-Seq and
       A81: L~pion = L~pion1 and
       A82: pion/.1 = pion1/.1 and
       A83: pion/.len pion = pion1/.len pion1 and
       A84: len pion <= len pion1 by A74,GOBOARD3:2;
      reconsider pion1 as being_S-Seq FinSequence of TOP-REAL 2 by A80;
      set godo = go^'pion1^'do;
      A85: Gi1k`1 = G*(i1,1)`1 by A6,A11,GOBOARD5:3
         .= Gij`1 by A6,A10,GOBOARD5:3;
      A86: Gik`1 <= Gi1k`1 by A1,A11,JORDAN1A:39;
      then A87: W-bound poz = Gik`1 by SPRECT_1:62;
      A88: W-bound pio = Gij`1 by A85,SPRECT_1:62;
        W-bound (poz \/ pio) = min(W-bound poz, W-bound pio) by SPRECT_1:54
         .= Gik`1 by A85,A86,A87,A88,SQUARE_1:def 1;
      then A89: W-bound L~pion1 = Gik`1 by A81,TOPREAL3:23;
        len Cage(C,n) > 4 by GOBOARD7:36;
      then A90: 1+1 <= len Cage(C,n) by AXIOMS:22;
      then A91: 1+1 <= len Rotate(Cage(C,n),Wmin) by REVROT_1:14;
        len (go^'pion1) >= len go by TOPREAL8:7;
      then A92: len (go^'pion1) >= 1+1 by A33,AXIOMS:22;
      then A93: len (go^'pion1) > 1+0 by NAT_1:38;
        len godo >= len (go^'pion1) by TOPREAL8:7;
      then A94: 1+1 <= len godo by A92,AXIOMS:22;
      A95: US is_sequence_on G by JORDAN1G:4;
      A96: go/.len go = pion1/.1 by A41,A82,FINSEQ_4:27;
      then A97: go^'pion1 is_sequence_on G by A35,A79,TOPREAL8:12;
      A98: (go^'pion1)/.len (go^'pion1) = pion/.len pion by A83,AMISTD_1:6
         .= pion/.3 by FINSEQ_1:62
         .= do/.1 by A42,FINSEQ_4:27;
      then A99: godo is_sequence_on G by A38,A97,TOPREAL8:12;
        LSeg(pion1,1) c= L~pion by A81,TOPREAL3:26;
      then A100: LSeg(go,len go-'1) /\ LSeg(pion1,1) c={Gik}
                                                      by A44,A51,XBOOLE_1:27;
        len pion1 >= 2+1 by A84,FINSEQ_1:62;
      then A101: len pion1 > 1+1 by NAT_1:38;
        {Gik} c= LSeg(go,m) /\ LSeg(pion1,1)
      proof
       let x be set;
       assume x in {Gik};
       then A102: x = Gik by TARSKI:def 1;
       A103: Gik in LSeg(go,m) by A48,TOPREAL1:6;
         Gik in LSeg(pion1,1) by A41,A96,A101,TOPREAL1:27;
       hence x in LSeg(go,m) /\ LSeg(pion1,1) by A102,A103,XBOOLE_0:def 3;
      end;
      then LSeg(go,len go-'1) /\ LSeg(pion1,1) = { go/.len go }
                                              by A41,A44,A100,XBOOLE_0:def 10;
      then A104: go^'pion1 is unfolded by A96,TOPREAL8:34;
        len pion1 >= 2+1 by A84,FINSEQ_1:62;
      then len pion1 > 2+0 by NAT_1:38;
      then A105: len pion1-2 >= 0 by REAL_1:84;
      A106: len (go^'pion1)-1 >= 0 by A93,REAL_1:84;
        len (go^'pion1)+1-1 = len go+len pion1-1 by GRAPH_2:13;
      then len (go^'pion1)-1 = len go+len pion1-1-1 by XCMPLX_1:26
         .= len go + len pion1-(1+1) by XCMPLX_1:36
         .= len go + (len pion1-2) by XCMPLX_1:29
         .= len go + (len pion1-'2) by A105,BINARITH:def 3;
      then A107: len (go^'pion1)-'1 = len go + (len pion1-'2)
                                                      by A106,BINARITH:def 3;
      A108: len pion1-1 >= 1 by A101,REAL_1:84;
      then A109: len pion1-1 >= 0 by AXIOMS:22;
      then A110: len pion1-'1 = len pion1-1 by BINARITH:def 3;
      A111: len pion1-'2+1 = len pion1-2+1 by A105,BINARITH:def 3
         .= len pion1-(2-1) by XCMPLX_1:37
         .= len pion1-'1 by A109,BINARITH:def 3;
        len pion1-1+1 <= len pion1 by XCMPLX_1:27;
      then A112: len pion1-'1 < len pion1 by A110,NAT_1:38;
        LSeg(pion1,len pion1-'1) c= L~pion by A81,TOPREAL3:26;
      then A113: LSeg(pion1,len pion1-'1) /\ LSeg(do,1) c= {Gij}
                                                          by A58,XBOOLE_1:27;
        {Gij} c= LSeg(pion1,len pion1-'1) /\ LSeg(do,1)
      proof
       let x be set;
       assume x in {Gij};
       then A114: x = Gij by TARSKI:def 1;
       A115: Gij in LSeg(do,1) by A55,TOPREAL1:6;
       A116: len pion1-'1+1 = len pion1 by A110,XCMPLX_1:27;
       then pion1/.(len pion1-'1+1) = pion/.3 by A83,FINSEQ_1:62
          .= Gij by FINSEQ_4:27;
       then Gij in LSeg(pion1,len pion1-'1) by A108,A110,A116,TOPREAL1:27;
       hence x in LSeg(pion1,len pion1-'1) /\ LSeg(do,1)
                                                   by A114,A115,XBOOLE_0:def 3;
      end;
      then LSeg(pion1,len pion1-'1) /\ LSeg(do,1) = {Gij}
                                                     by A113,XBOOLE_0:def 10;
      then A117: LSeg(go^'pion1,len go+(len pion1-'2)) /\ LSeg(do,1) =
         {(go^'pion1)/.len (go^'pion1)} by A42,A96,A98,A111,A112,TOPREAL8:31;
      A118: (go^'pion1) is non trivial by A92,SPPOL_1:2;
      A119: rng pion1 c= L~pion1 by A101,SPPOL_2:18;
      A120: {pion1/.1} c= L~go /\ L~pion1
      proof
       let x be set;
       assume x in {pion1/.1};
       then x = pion1/.1 by TARSKI:def 1;
       then x in rng go & x in rng pion1 by A96,FINSEQ_6:46,REVROT_1:3;
       hence x in L~go /\ L~pion1 by A61,A119,XBOOLE_0:def 3;
      end;
        L~go /\ L~pion1 c= {pion1/.1}
      proof
       let x be set;
       assume x in L~go /\ L~pion1;
       then x in L~go & x in L~pion1 by XBOOLE_0:def 3;
       hence x in {pion1/.1} by A3,A7,A41,A46,A81,A96,XBOOLE_0:def 3;
      end;
      then A121: L~go /\ L~pion1 = {pion1/.1} by A120,XBOOLE_0:def 10;
      then A122: (go^'pion1) is s.n.c. by A96,JORDAN1J:54;
        rng go /\ rng pion1 c= {pion1/.1} by A61,A119,A121,XBOOLE_1:27;
      then A123: go^'pion1 is one-to-one by JORDAN1J:55;
      A124: pion/.len pion = pion/.3 by FINSEQ_1:62
         .= do/.1 by A42,FINSEQ_4:27;
      A125: {pion1/.len pion1} c= L~do /\ L~pion1
      proof
       let x be set;
       assume x in {pion1/.len pion1};
       then x = pion1/.len pion1 by TARSKI:def 1;
       then x in rng do & x in rng pion1 by A83,A124,FINSEQ_6:46,REVROT_1:3;
       hence x in L~do /\ L~pion1 by A61,A119,XBOOLE_0:def 3;
      end;
        L~do /\ L~pion1 c= {pion1/.len pion1}
      proof
       let x be set;
       assume x in L~do /\ L~pion1;
       then x in L~do & x in L~pion1 by XBOOLE_0:def 3;
       hence x in {pion1/.len pion1}
                by A4,A7,A42,A53,A81,A83,A124,XBOOLE_0:def 3;
      end;
      then A126: L~do /\ L~pion1 = {pion1/.len pion1} by A125,XBOOLE_0:def 10;
      A127: L~(go^'pion1) /\ L~do = (L~go \/ L~pion1) /\ L~do by A96,TOPREAL8:
35
         .= {go/.1} \/ {do/.1} by A68,A83,A124,A126,XBOOLE_1:23
         .= {(go^'pion1)/.1} \/ {do/.1} by AMISTD_1:5
         .= {(go^'pion1)/.1,do/.1} by ENUMSET1:41;
        do/.len do = (go^'pion1)/.1 by A60,AMISTD_1:5;
      then reconsider godo as non constant standard special_circular_sequence
           by A94,A98,A99,A104,A107,A117,A118,A122,A123,A127,JORDAN8:7,8,
TOPREAL8:11,33,34;
      A128: LA is_an_arc_of E-max C,W-min C by JORDAN6:def 9;
      then A129: LA is connected by JORDAN6:11;
      A130: W-min C in LA & E-max C in LA by A128,TOPREAL1:4;
      set ff = Rotate(Cage(C,n),Wmin);
        Wmin in rng Cage(C,n) by SPRECT_2:47;
      then A131: ff/.1 = Wmin by FINSEQ_6:98;
      A132: L~ff = L~Cage(C,n) by REVROT_1:33;
      then A133: (W-max L~ff)..ff > 1 by A131,SPRECT_5:23;
        (W-max L~ff)..ff <= (N-min L~ff)..ff by A131,A132,SPRECT_5:24;
      then A134: (N-min L~ff)..ff > 1 by A133,AXIOMS:22;
        (N-min L~ff)..ff < (N-max L~ff)..ff by A131,A132,SPRECT_5:25;
      then A135: (N-max L~ff)..ff > 1 by A134,AXIOMS:22;
        (N-max L~ff)..ff <= (E-max L~ff)..ff by A131,A132,SPRECT_5:26;
      then A136: Emax..ff > 1 by A132,A135,AXIOMS:22;
      A137: now assume A138: Gik..US <= 1;
         Gik..US >= 1 by A34,FINSEQ_4:31;
       then Gik..US = 1 by A138,AXIOMS:21;
       then Gik = US/.1 by A34,FINSEQ_5:41;
       hence contradiction by A19,A23,JORDAN1F:5;
      end;
      A139: Cage(C,n) is_sequence_on G by JORDAN9:def 1;
      then A140: ff is_sequence_on G by REVROT_1:34;
      A141: right_cell(godo,1,G)\L~godo c= RightComp godo by A94,A99,JORDAN9:29
;
      A142: L~godo = L~(go^'pion1) \/ L~do by A98,TOPREAL8:35
         .= L~go \/ L~pion1 \/ L~do by A96,TOPREAL8:35;
        L~Cage(C,n) = L~US \/ L~LS by JORDAN1E:17;
      then A143: L~US c= L~Cage(C,n) & L~LS c= L~Cage(C,n) by XBOOLE_1:7;
      then A144: L~go c=L~Cage(C,n) & L~do c=L~Cage(C,n)
                                                     by A46,A53,XBOOLE_1:1;
      A145: W-min C in C by SPRECT_1:15;
      A146: now assume W-min C in L~godo;
       then A147: W-min C in L~go \/ L~pion1 or W-min C in L~do
                                                       by A142,XBOOLE_0:def 2;
       per cases by A147,XBOOLE_0:def 2;
        suppose W-min C in L~go;
         then C meets L~Cage(C,n) by A144,A145,XBOOLE_0:3;
         hence contradiction by JORDAN10:5;
        suppose W-min C in L~pion1;
         hence contradiction by A5,A7,A81,A130,XBOOLE_0:3;
        suppose W-min C in L~do;
         then C meets L~Cage(C,n) by A144,A145,XBOOLE_0:3;
         hence contradiction by JORDAN10:5;
      end;
        right_cell(Rotate(Cage(C,n),Wmin),1) =
            right_cell(ff,1,GoB ff) by A91,JORDAN1H:29
         .= right_cell(ff,1,GoB Cage(C,n)) by REVROT_1:28
         .= right_cell(ff,1,G) by JORDAN1H:52
         .= right_cell(ff-:Emax,1,G) by A136,A140,JORDAN1J:53
         .= right_cell(US,1,G) by JORDAN1E:def 1
         .= right_cell(R_Cut(US,Gik),1,G) by A34,A95,A137,JORDAN1J:52
         .= right_cell(go^'pion1,1,G) by A39,A97,JORDAN1J:51
         .= right_cell(godo,1,G) by A93,A99,JORDAN1J:51;
      then W-min C in right_cell(godo,1,G) by JORDAN1I:8;
      then A148: W-min C in right_cell(godo,1,G)\L~godo by A146,XBOOLE_0:def 4;
      A149: godo/.1 = (go^'pion1)/.1 by AMISTD_1:5
         .= Wmin by A59,AMISTD_1:5;
      A150: len US >= 2 by A18,AXIOMS:22;
      A151: godo/.2 = (go^'pion1)/.2 by A92,AMISTD_1:9
         .= US/.2 by A33,A72,AMISTD_1:9
         .= (US^'LS)/.2 by A150,AMISTD_1:9
         .= Rotate(Cage(C,n),Wmin)/.2 by JORDAN1E:15;
      A152: L~go \/ L~do is compact by COMPTS_1:19;
      A153: L~go \/ L~do c= L~Cage(C,n) by A144,XBOOLE_1:8;
        Wmin in L~go \/ L~do by A61,A73,XBOOLE_0:def 2;
      then A154: W-min (L~go \/ L~do) = Wmin by A152,A153,JORDAN1J:21;
      A155: (W-min (L~go \/ L~do))`1 = W-bound (L~go \/ L~do) & Wmin`1 = Wbo
                                                              by PSCOMP_1:84;
        Gik`1 >= Wbo by A9,A143,PSCOMP_1:71;
      then Gik`1 > Wbo by A71,REAL_1:def 5;
      then W-min (L~go\/L~do\/L~pion1) = W-min (L~go \/ L~do)
                                              by A89,A152,A154,A155,JORDAN1J:33
;
      then A156: W-min L~godo = Wmin by A142,A154,XBOOLE_1:4;
      A157: rng godo c= L~godo by A94,SPPOL_2:18;
        2 in dom godo by A94,FINSEQ_3:27;
      then A158: godo/.2 in rng godo by PARTFUN2:4;
        godo/.2 in W-most L~Cage(C,n) by A151,JORDAN1I:27;
      then (godo/.2)`1 = (W-min L~godo)`1 by A156,PSCOMP_1:88
         .= W-bound L~godo by PSCOMP_1:84;
      then godo/.2 in W-most L~godo by A157,A158,SPRECT_2:16;
      then Rotate(godo,W-min L~godo)/.2 in W-most L~godo
                                                      by A149,A156,FINSEQ_6:95;
      then reconsider godo as clockwise_oriented non constant standard
                                    special_circular_sequence by JORDAN1I:27;
        len US in dom US by FINSEQ_5:6;
      then A159: US.len US = US/.len US by FINSEQ_4:def 4
         .= Emax by JORDAN1F:7;
      A160: E-max C in E-most C by PSCOMP_1:111;
      A161: east_halfline E-max C misses L~go
      proof
       assume east_halfline E-max C meets L~go;
       then consider p be set such that
        A162: p in east_halfline E-max C and
        A163: p in L~go by XBOOLE_0:3;
       reconsider p as Point of TOP-REAL 2 by A162;
         p in L~US by A46,A163;
       then p in east_halfline E-max C /\ L~Cage(C,n)
                       by A143,A162,XBOOLE_0:def 3;
       then A164: p`1 = Ebo by A160,JORDAN1A:104;
       then A165: p = Emax by A46,A163,JORDAN1J:46;
       then Emax = Gik by A9,A159,A163,JORDAN1J:43;
       then Gik`1 = G*(len G,k)`1 by A11,A17,A164,A165,JORDAN1A:92;
       hence contradiction by A1,A13,A30,JORDAN1G:7;
      end;
        now assume east_halfline E-max C meets L~godo;
       then A166: east_halfline E-max C meets (L~go \/ L~pion1) or
                 east_halfline E-max C meets L~do by A142,XBOOLE_1:70;
       per cases by A166,XBOOLE_1:70;
        suppose east_halfline E-max C meets L~go;
         hence contradiction by A161;
        suppose east_halfline E-max C meets L~pion1;
         then consider p be set such that
          A167: p in east_halfline E-max C and
          A168: p in L~pion1 by XBOOLE_0:3;
         reconsider p as Point of TOP-REAL 2 by A167;
         A169: now per cases by A7,A81,A168,XBOOLE_0:def 2;
          suppose p in poz;
           hence p`1 <= Gi1k`1 by A86,TOPREAL1:9;
          suppose p in pio;
           hence p`1 <= Gi1k`1 by A85,GOBOARD7:5;
         end;
           i1+1 <= len G by A1,NAT_1:38;
         then i1+1-1 <= len G-1 by REAL_1:49;
         then A170: i1 <= len G-1 by XCMPLX_1:26;
         then len G-1 > 0 by A1,AXIOMS:22;
         then A171: i1 <= len G-'1 by A170,BINARITH:def 3;
           len G-'1 <= len G by GOBOARD9:2;
         then Gi1k`1 <= G*(len G-'1,1)`1 by A6,A11,A17,A21,A171,JORDAN1A:39;
         then p`1 <= G*(len G-'1,1)`1 by A169,AXIOMS:22;
         then p`1 <= E-bound C by A21,JORDAN8:15;
         then A172: p`1 <= (E-max C)`1 by PSCOMP_1:104;
           p`1 >= (E-max C)`1 by A167,JORDAN1A:def 3;
         then A173: p`1 = (E-max C)`1 by A172,AXIOMS:21;
           p`2 = (E-max C)`2 by A167,JORDAN1A:def 3;
         then p = E-max C by A173,TOPREAL3:11;
         hence contradiction by A5,A7,A81,A130,A168,XBOOLE_0:3;
        suppose east_halfline E-max C meets L~do;
         then consider p be set such that
          A174: p in east_halfline E-max C and
          A175: p in L~do by XBOOLE_0:3;
         reconsider p as Point of TOP-REAL 2 by A174;
           p in L~LS by A53,A175;
         then p in east_halfline E-max C /\ L~Cage(C,n)
                    by A143,A174,XBOOLE_0:def 3;
         then A176: p`1 = Ebo by A160,JORDAN1A:104;
         A177: (E-max C)`2 = p`2 by A174,JORDAN1A:def 3;
         set RC = Rotate(Cage(C,n),Emax);
         A178: E-max C in right_cell(RC,1) by JORDAN1I:9;
         A179: 1+1 <= len LS by A24,AXIOMS:22;
           LS = RC-:Wmin by JORDAN1G:26;
         then A180: LSeg(LS,1) = LSeg(RC,1) by A179,SPPOL_2:9;
         A181: L~RC = L~Cage(C,n) by REVROT_1:33;
         A182: len RC = len Cage(C,n) by REVROT_1:14;
         A183: GoB RC = GoB Cage(C,n) by REVROT_1:28
            .= G by JORDAN1H:52;
         A184: Emax in rng Cage(C,n) by SPRECT_2:50;
         A185: RC is_sequence_on G by A139,REVROT_1:34;
         A186: RC/.1 = E-max L~RC by A181,A184,FINSEQ_6:98;
         then consider ii,jj be Nat such that
          A187: [ii,jj+1] in Indices G and
          A188: [ii,jj] in Indices G and
          A189: RC/.1 = G*(ii,jj+1) and
          A190: RC/.(1+1) = G*(ii,jj) by A90,A182,A185,JORDAN1I:25;
         consider jj2 be Nat such that
          A191: 1 <= jj2 & jj2 <= width G and
          A192: Emax = G*(len G,jj2) by JORDAN1D:29;
         A193: len G >= 4 by JORDAN8:13;
         then len G >= 1 by AXIOMS:22;
         then [len G,jj2] in Indices G by A191,GOBOARD7:10;
         then A194: ii = len G by A181,A186,A187,A189,A192,GOBOARD1:21;
         A195: 1 <= ii & ii <= len G & 1 <= jj+1 & jj+1 <= width G
                                                           by A187,GOBOARD5:1;
         A196: 1 <= ii & ii <= len G & 1 <= jj & jj <= width G
                                                           by A188,GOBOARD5:1;
         A197: ii+1 <> ii by NAT_1:38;
           jj+1 > jj by NAT_1:38;
         then jj+1+1 <> jj by NAT_1:38;
         then A198: right_cell(RC,1) = cell(G,ii-'1,jj)
                           by A90,A182,A183,A187,A188,A189,A190,A197,GOBOARD5:
def 6;
         A199: ii-'1+1 = ii by A195,AMI_5:4;
         A200: ii-1 >= 4-1 by A193,A194,REAL_1:49;
         then A201: ii-1 >= 1 by AXIOMS:22;
           ii-1 >= 0 by A200,AXIOMS:22;
         then A202: 1 <= ii-'1 by A201,BINARITH:def 3;
         then A203: G*(ii-'1,jj)`2 <= p`2 & p`2 <= G*(ii-'1,jj+1)`2
                 by A177,A178,A195,A196,A198,A199,JORDAN9:19;
         A204: ii-'1 < len G by A195,A199,NAT_1:38;
         then A205: G*(ii-'1,jj)`2 = G*(1,jj)`2 by A196,A202,GOBOARD5:2
            .= G*(ii,jj)`2 by A196,GOBOARD5:2;
         A206: G*(ii-'1,jj+1)`2 = G*(1,jj+1)`2 by A195,A202,A204,GOBOARD5:2
            .= G*(ii,jj+1)`2 by A195,GOBOARD5:2;
           G*(len G,jj)`1 = Ebo & Ebo = G*(len G,jj+1)`1
                                                by A17,A195,A196,JORDAN1A:92;
         then p in LSeg(RC/.1,RC/.(1+1))
                              by A176,A189,A190,A194,A203,A205,A206,GOBOARD7:8;
         then A207: p in LSeg(LS,1) by A90,A180,A182,TOPREAL1:def 5;
         A208: p in LSeg(do,Index(p,do)) by A175,JORDAN3:42;
         A209: do = mid(LS,Gij..LS,len LS) by A37,JORDAN1J:37;
         A210: 1<=Gij..LS & Gij..LS<=len LS by A37,FINSEQ_4:31;
           Gij..LS <> len LS by A29,A37,FINSEQ_4:29;
         then A211: Gij..LS < len LS by A210,REAL_1:def 5;
         A212: 1<=Index(p,do) & Index(p,do) < len do by A175,JORDAN3:41;
         A213: Index(Gij,LS)+1 = Gij..LS by A32,A37,JORDAN1J:56;
         consider t be Nat such that
          A214: t in dom LS and
          A215: LS.t = Gij by A37,FINSEQ_2:11;
         A216: 1 <= t & t <= len LS by A214,FINSEQ_3:27;
         then 1 < t by A32,A215,REAL_1:def 5;
         then Index(Gij,LS)+1 = t by A215,A216,JORDAN3:45;
         then A217: len L_Cut(LS,Gij) = len LS-Index(Gij,LS)
                                                  by A8,A29,A215,JORDAN3:61;
         set tt = Index(p,do)+(Gij..LS)-'1;
         A218: 1<=Index(Gij,LS) & 0+Index(Gij,LS) < len LS by A8,JORDAN3:41;
         then A219: len LS-Index(Gij,LS) > 0 by REAL_1:86;
         then Index(p,do) < len LS-'Index(Gij,LS) by A212,A217,BINARITH:def 3;
         then Index(p,do)+1 <= len LS-'Index(Gij,LS) by NAT_1:38;
         then Index(p,do) <= len LS-'Index(Gij,LS)-1 by REAL_1:84;
         then Index(p,do) <= len LS-Index(Gij,LS)-1 by A219,BINARITH:def 3;
         then A220: Index(p,do) <= len LS-Gij..LS by A213,XCMPLX_1:36;
         then len LS-Gij..LS >= 1 by A212,AXIOMS:22;
         then len LS-Gij..LS >= 0 by AXIOMS:22;
         then Index(p,do) <= len LS-'Gij..LS by A220,BINARITH:def 3;
         then Index(p,do) < len LS-'(Gij..LS)+1 by NAT_1:38;
         then A221: LSeg(mid(LS,Gij..LS,len LS),Index(p,do)) =
              LSeg(LS,Index(p,do)+(Gij..LS)-'1) by A210,A211,A212,JORDAN4:31;
         A222: 1+1 <= Gij..LS by A213,A218,REAL_1:55;
         then Index(p,do)+Gij..LS >= 1+1+1 by A212,REAL_1:55;
         then A223: Index(p,do)+Gij..LS-1 >= 1+1+1-1 by REAL_1:49;
         then A224: Index(p,do)+Gij..LS-1 >= 0 by AXIOMS:22;
         then A225: tt >= 1+1 by A223,BINARITH:def 3;
         A226: 2 in dom LS by A179,FINSEQ_3:27;
           now per cases by A225,REAL_1:def 5;
          suppose tt > 1+1;
           then LSeg(LS,1) misses LSeg(LS,tt) by TOPREAL1:def 9;
           hence contradiction by A207,A208,A209,A221,XBOOLE_0:3;
          suppose A227: tt = 1+1;
           then LSeg(LS,1) /\ LSeg(LS,tt) = {LS/.2} by A24,TOPREAL1:def 8;
           then p in {LS/.2} by A207,A208,A209,A221,XBOOLE_0:def 3;
           then A228: p = LS/.2 by TARSKI:def 1;
           then A229: p..LS = 2 by A226,FINSEQ_5:44;
             1+1 = Index(p,do)+(Gij..LS)-1 by A224,A227,BINARITH:def 3;
           then 1+1+1 = Index(p,do)+(Gij..LS) by XCMPLX_1:27;
           then A230: Gij..LS = 2 by A212,A222,JORDAN1E:10;
             p in rng LS by A226,A228,PARTFUN2:4;
           then p = Gij by A37,A229,A230,FINSEQ_5:10;
           then Gij`1 = Ebo by A228,JORDAN1G:40;
           then Gij`1 = G*(len G,j)`1 by A10,A17,JORDAN1A:92;
           hence contradiction by A1,A12,A64,JORDAN1G:7;
         end;
         hence contradiction;
      end;
      then east_halfline E-max C c= (L~godo)` by SUBSET_1:43;
      then consider W be Subset of TOP-REAL 2 such that
       A231: W is_a_component_of (L~godo)` and
       A232: east_halfline E-max C c= W by GOBOARD9:5;
        east_halfline E-max C is not Bounded by JORDAN1C:9;
      then W is not Bounded by A232,JORDAN2C:16;
      then W is_outside_component_of L~godo by A231,JORDAN2C:def 4;
      then W c= UBD L~godo by JORDAN2C:27;
      then A233: east_halfline E-max C c= UBD L~godo by A232,XBOOLE_1:1;
        E-max C in east_halfline E-max C by JORDAN1C:7;
      then E-max C in UBD L~godo by A233;
      then E-max C in LeftComp godo by GOBRD14:46;
      then LA meets L~godo by A129,A130,A141,A148,JORDAN1J:36;
      then A234: LA meets (L~go \/ L~pion1) or LA meets L~do by A142,XBOOLE_1:
70;
      A235: LA c= C by JORDAN1A:16;
        now per cases by A234,XBOOLE_1:70;
       suppose LA meets L~go;
        then LA meets L~Cage(C,n) by A144,XBOOLE_1:63;
        then C meets L~Cage(C,n) by A235,XBOOLE_1:63;
        hence contradiction by JORDAN10:5;
       suppose LA meets L~pion1;
        hence contradiction by A5,A7,A81;
       suppose LA meets L~do;
        then LA meets L~Cage(C,n) by A144,XBOOLE_1:63;
        then C meets L~Cage(C,n) by A235,XBOOLE_1:63;
        hence contradiction by JORDAN10:5;
      end;
      hence contradiction;
    suppose Gij`1 = Gik`1;
      then A236: i1 = i2 by A12,A13,JORDAN1G:7;
      then poz = {Gi1k} by TOPREAL1:7;
      then poz c= pio by A77,ZFMISC_1:37;
      then pio \/ poz = pio by XBOOLE_1:12;
      hence contradiction by A1,A2,A3,A4,A5,A236,JORDAN1J:58;
     suppose Gij`2 = Gik`2;
      then A237: j = k by A12,A13,JORDAN1G:6;
      then pio = {Gi1k} by TOPREAL1:7;
      then pio c= poz by A78,ZFMISC_1:37;
      then pio \/ poz = poz by XBOOLE_1:12;
      hence contradiction by A1,A2,A3,A4,A5,A237,Th38;
    end;
    hence contradiction;
   end;

  theorem Th50:
   for C be Simple_closed_curve
   for i1,i2,j,k be Nat holds
    1 < i1 & i1 < len Gauge(C,n+1) & 1 < i2 & i2 < len Gauge(C,n+1) &
    1 <= j & j <= k & k <= width Gauge(C,n+1) &
    Gauge(C,n+1)*(i1,k) in Upper_Arc L~Cage(C,n+1) &
    Gauge(C,n+1)*(i2,j) in Lower_Arc L~Cage(C,n+1) implies
     LSeg(Gauge(C,n+1)*(i2,j),Gauge(C,n+1)*(i2,k)) \/
     LSeg(Gauge(C,n+1)*(i2,k),Gauge(C,n+1)*(i1,k)) meets Upper_Arc C
   proof
    let C be Simple_closed_curve;
    let i1,i2,j,k be Nat;
    set G=Gauge(C,n+1);
    assume that
     A1: 1 < i1 & i1 < len G and
     A2: 1 < i2 & i2 < len G and
     A3: 1 <= j & j <= k & k <= width G and
     A4: G*(i1,k) in Upper_Arc L~Cage(C,n+1) and
     A5: G*(i2,j) in Lower_Arc L~Cage(C,n+1);
      n+1 >= 0+1 by NAT_1:29;
    then A6: n+1 > 0 by NAT_1:38;
    then A7: Upper_Arc L~Cage(C,n+1) = L~Upper_Seq(C,n+1) by JORDAN1G:63;
    A8: Lower_Arc L~Cage(C,n+1) = L~Lower_Seq(C,n+1) by A6,JORDAN1G:64;
    A9: Upper_Seq(C,n+1) is_sequence_on G by JORDAN1G:4;
    A10: Lower_Seq(C,n+1) is_sequence_on G by JORDAN1G:5;
    A11: 1 <= j & j <= width G by A3,AXIOMS:22;
    then A12: [i2,j] in Indices G by A2,GOBOARD7:10;
    A13: 1 <= k & k <= width G by A3,AXIOMS:22;
    then A14: [i2,k] in Indices G by A2,GOBOARD7:10;
      G*(i2,j)`1 = G*(i2,1)`1 by A2,A11,GOBOARD5:3
       .= G*(i2,k)`1 by A2,A13,GOBOARD5:3;
    then A15: LSeg(G*(i2,j),G*(i2,k)) is vertical by SPPOL_1:37;
      G*(i2,k)`2 = G*(1,k)`2 by A2,A13,GOBOARD5:2
       .= G*(i1,k)`2 by A1,A13,GOBOARD5:2;
    then A16: LSeg(G*(i2,k),G*(i1,k)) is horizontal by SPPOL_1:36;
    A17: [i2,k] in Indices G by A2,A13,GOBOARD7:10;
    A18: [i1,k] in Indices G by A1,A13,GOBOARD7:10;
      now per cases;
     suppose A19: LSeg(G*(i2,j),G*(i2,k)) meets
       Upper_Arc L~Cage(C,n+1);
      then consider m be Nat such that
       A20: j <= m & m <= k and
       A21: G*(i2,m)`2 = inf(proj2.:(LSeg(G*(i2,j),G*(i2,k))
            /\ L~Upper_Seq(C,n+1))) by A3,A7,A9,A12,A14,JORDAN1F:1;
      A22: 1 <= m & m <= width G by A3,A20,AXIOMS:22;
      set X = LSeg(G*(i2,j),G*(i2,k)) /\ L~Upper_Seq(C,n+1);
      A23: G*(i2,m)`1 = G*(i2,1)`1 by A2,A22,GOBOARD5:3;
      then A24: |[G*(i2,1)`1,inf(proj2.:X)]| = G*(i2,m) by A21,EUCLID:57;
      then A25: G*(i2,j)`1 = |[G*(i2,1)`1,inf(proj2.:X)]|`1
                                               by A2,A11,A23,GOBOARD5:3;
        ex x be set st x in LSeg(G*(i2,j),G*(i2,k)) &
       x in L~Upper_Seq(C,n+1) by A7,A19,XBOOLE_0:3;
      then reconsider X1=X as non empty compact Subset of TOP-REAL 2
                                              by PSCOMP_1:64,XBOOLE_0:def 3;
      consider pp be set such that
       A26: pp in S-most X1 by XBOOLE_0:def 1;
      reconsider pp as Point of TOP-REAL 2 by A26;
        pp in LSeg(SW-corner X, SE-corner X)/\X by A26,PSCOMP_1:def 41;
      then A27: pp in X by XBOOLE_0:def 3;
      then A28: pp in L~Upper_Seq(C,n+1) by XBOOLE_0:def 3;
        pp in LSeg(G*(i2,j),G*(i2,k)) by A27,XBOOLE_0:def 3;
      then A29: pp`1 = |[G*(i2,1)`1,inf(proj2.:X)]|`1
                            by A15,A25,SPRECT_3:20;
        |[G*(i2,1)`1,inf(proj2.:X)]|`2 = S-bound X by A21,A24,SPRECT_1:49
         .= (S-min X)`2 by PSCOMP_1:114
         .= pp`2 by A26,PSCOMP_1:118;
      then G*(i2,m) in Upper_Arc L~Cage(C,n+1) by A7,A24,A28,A29,TOPREAL3:11;
      then A30: LSeg(G*(i2,j),G*(i2,m)) meets Upper_Arc C
                                             by A2,A3,A5,A6,A20,A22,Th26;
        LSeg(G*(i2,j),G*(i2,m)) c= LSeg(G*(i2,j),G*(i2,k)) by A2,A3,A20,Th7;
      then LSeg(G*(i2,j),G*(i2,k)) meets Upper_Arc C by A30,XBOOLE_1:63;
      hence LSeg(G*(i2,j),G*(i2,k)) \/
            LSeg(G*(i2,k),G*(i1,k)) meets Upper_Arc C by XBOOLE_1:70;
     suppose A31: LSeg(G*(i2,k),G*(i1,k)) meets Lower_Arc L~Cage(C,n+1) &
       i2 <= i1;
      then consider m be Nat such that
       A32: i2 <= m & m <= i1 and
       A33: G*(m,k)`1 = sup(proj1.:(LSeg(G*(i2,k),G*(i1,k))
            /\ L~Lower_Seq(C,n+1))) by A8,A10,A17,A18,JORDAN1F:4;
      A34: 1 < m & m < len G by A1,A2,A32,AXIOMS:22;
      set X = LSeg(G*(i2,k),G*(i1,k)) /\ L~Lower_Seq(C,n+1);
      A35: G*(m,k)`2 = G*(1,k)`2 by A13,A34,GOBOARD5:2;
      then A36: |[sup(proj1.:X),G*(1,k)`2]| = G*(m,k) by A33,EUCLID:57;
      then A37: G*(i2,k)`2 = |[sup(proj1.:X),G*(1,k)`2]|`2
                                               by A2,A13,A35,GOBOARD5:2;
        ex x be set st x in LSeg(G*(i2,k),G*(i1,k)) &
       x in L~Lower_Seq(C,n+1) by A8,A31,XBOOLE_0:3;
      then reconsider X1=X as non empty compact Subset of TOP-REAL 2
                                             by PSCOMP_1:64,XBOOLE_0:def 3;
      consider pp be set such that
       A38: pp in E-most X1 by XBOOLE_0:def 1;
      reconsider pp as Point of TOP-REAL 2 by A38;
        pp in LSeg(SE-corner X, NE-corner X)/\X by A38,PSCOMP_1:def 40;
      then A39: pp in X by XBOOLE_0:def 3;
      then A40: pp in L~Lower_Seq(C,n+1) by XBOOLE_0:def 3;
        pp in LSeg(G*(i2,k),G*(i1,k)) by A39,XBOOLE_0:def 3;
      then A41: pp`2 = |[sup(proj1.:X),G*(1,k)`2]|`2 by A16,A37,SPRECT_3:19;
        |[sup(proj1.:X),G*(1,k)`2]|`1 = E-bound X by A33,A36,SPRECT_1:51
         .= (E-min X)`1 by PSCOMP_1:104
         .= pp`1 by A38,PSCOMP_1:108;
      then G*(m,k) in Lower_Arc L~Cage(C,n+1) by A8,A36,A40,A41,TOPREAL3:11;
      then A42: LSeg(G*(m,k),G*(i1,k)) meets Upper_Arc C
                                             by A1,A4,A6,A13,A32,A34,Th35;
        LSeg(G*(m,k),G*(i1,k)) c= LSeg(G*(i2,k),G*(i1,k))
                                                    by A1,A2,A13,A32,Th8;
      then LSeg(G*(i2,k),G*(i1,k)) meets Upper_Arc C by A42,XBOOLE_1:63;
      hence LSeg(G*(i2,j),G*(i2,k)) \/
            LSeg(G*(i2,k),G*(i1,k)) meets Upper_Arc C by XBOOLE_1:70;
     suppose A43: LSeg(G*(i2,k),G*(i1,k)) meets Lower_Arc L~Cage(C,n+1) &
      i1 < i2;
      then consider m be Nat such that
       A44: i1 <= m & m <= i2 and
       A45: G*(m,k)`1 = inf(proj1.:(LSeg(G*(i1,k),G*(i2,k))
            /\ L~Lower_Seq(C,n+1))) by A8,A10,A17,A18,JORDAN1F:3;
      A46: 1 < m & m < len G by A1,A2,A44,AXIOMS:22;
      set X = LSeg(G*(i1,k),G*(i2,k)) /\ L~Lower_Seq(C,n+1);
      A47: G*(m,k)`2 = G*(1,k)`2 by A13,A46,GOBOARD5:2;
      then A48: |[inf(proj1.:X),G*(1,k)`2]| = G*(m,k) by A45,EUCLID:57;
      then A49: G*(i1,k)`2 = |[inf(proj1.:X),G*(1,k)`2]|`2
                                                by A1,A13,A47,GOBOARD5:2;
        ex x be set st x in LSeg(G*(i1,k),G*(i2,k)) &
       x in L~Lower_Seq(C,n+1) by A8,A43,XBOOLE_0:3;
      then reconsider X1=X as non empty compact Subset of TOP-REAL 2
                                             by PSCOMP_1:64,XBOOLE_0:def 3;
      consider pp be set such that
       A50: pp in W-most X1 by XBOOLE_0:def 1;
      reconsider pp as Point of TOP-REAL 2 by A50;
        pp in LSeg(SW-corner X, NW-corner X)/\X by A50,PSCOMP_1:def 38;
      then A51: pp in X by XBOOLE_0:def 3;
      then A52: pp in L~Lower_Seq(C,n+1) by XBOOLE_0:def 3;
        pp in LSeg(G*(i1,k),G*(i2,k)) by A51,XBOOLE_0:def 3;
      then A53: pp`2 = |[inf(proj1.:X),G*(1,k)`2]|`2 by A16,A49,SPRECT_3:19;
        |[inf(proj1.:X),G*(1,k)`2]|`1 = W-bound X by A45,A48,SPRECT_1:48
         .= (W-min X)`1 by PSCOMP_1:84
         .= pp`1 by A50,PSCOMP_1:88;
      then G*(m,k) in Lower_Arc L~Cage(C,n+1) by A8,A48,A52,A53,TOPREAL3:11;
      then A54: LSeg(G*(i1,k),G*(m,k)) meets Upper_Arc C
                                             by A1,A4,A6,A13,A44,A46,Th43;
        LSeg(G*(i1,k),G*(m,k)) c= LSeg(G*(i1,k),G*(i2,k))
                                                    by A1,A2,A13,A44,Th8;
      then LSeg(G*(i1,k),G*(i2,k)) meets Upper_Arc C by A54,XBOOLE_1:63;
      hence LSeg(G*(i2,j),G*(i2,k)) \/
            LSeg(G*(i2,k),G*(i1,k)) meets Upper_Arc C by XBOOLE_1:70;
     suppose A55: LSeg(G*(i2,j),G*(i2,k)) misses
        Upper_Arc L~Cage(C,n+1) &
       LSeg(Gauge(C,n+1)*(i2,k),Gauge(C,n+1)*(i1,k))
        misses Lower_Arc L~Cage(C,n+1);
      consider j1 be Nat such that
       A56: j <= j1 & j1 <= k and
       A57: LSeg(G*(i2,j1),G*(i2,k)) /\ L~Lower_Seq(C,n+1) =
        {G*(i2,j1)} by A2,A3,A5,A8,Th11;
        G*(i2,j1) in LSeg(G*(i2,j1),G*(i2,k)) /\
       L~Lower_Seq(C,n+1) by A57,TARSKI:def 1;
      then A58: G*(i2,j1) in L~Lower_Seq(C,n+1) by XBOOLE_0:def 3;
      A59: 1 <= j1 & j1 <= width G by A3,A56,AXIOMS:22;
        now per cases;
       suppose i2 <= i1;
        then consider i3 be Nat such that
         A60: i2 <= i3 & i3 <= i1 and
         A61: LSeg(G*(i2,k),G*(i3,k)) /\ L~Upper_Seq(C,n+1) = {G*(i3,k)}
                                                 by A1,A2,A4,A7,A13,Th15;
        A62: 1 < i3 & i3 < len G by A1,A2,A60,AXIOMS:22;
          G*(i3,k) in LSeg(G*(i2,k),G*(i3,k)) /\ L~Upper_Seq(C,n+1)
                                                        by A61,TARSKI:def 1;
        then A63: G*(i3,k) in L~Upper_Seq(C,n+1) by XBOOLE_0:def 3;
        A64: LSeg(G*(i2,j1),G*(i2,k)) c=
         LSeg(G*(i2,j),G*(i2,k)) by A2,A3,A56,Th7;
        A65: LSeg(G*(i2,k),G*(i3,k)) c= LSeg(G*(i2,k),G*(i1,k))
                                                    by A1,A2,A13,A60,Th8;
        then A66: LSeg(G*(i2,j1),G*(i2,k)) \/ LSeg(G*(i2,k),G*(i3,k)) c=
         LSeg(G*(i2,j),G*(i2,k)) \/ LSeg(G*(i2,k),G*(i1,k)) by A64,XBOOLE_1:13;
        A67: (LSeg(G*(i2,j1),G*(i2,k)) \/
         LSeg(G*(i2,k),G*(i3,k))) /\ L~Upper_Seq(C,n+1) = {G*(i3,k)}
        proof
         thus (LSeg(G*(i2,j1),G*(i2,k)) \/
          LSeg(G*(i2,k),G*(i3,k))) /\ L~Upper_Seq(C,n+1) c= {G*(i3,k)}
         proof
          let x be set;
          assume x in (LSeg(G*(i2,j1),G*(i2,k)) \/
           LSeg(G*(i2,k),G*(i3,k))) /\ L~Upper_Seq(C,n+1);
          then x in LSeg(G*(i2,j1),G*(i2,k)) \/
                    LSeg(G*(i2,k),G*(i3,k)) &
               x in L~Upper_Seq(C,n+1) by XBOOLE_0:def 3;
          then (x in LSeg(G*(i2,j1),G*(i2,k)) or
                x in LSeg(G*(i2,k),G*(i3,k))) &
                x in L~Upper_Seq(C,n+1) by XBOOLE_0:def 2;
          hence x in {G*(i3,k)} by A7,A55,A61,A64,XBOOLE_0:3,def 3;
         end;
         let x be set;
         assume x in {G*(i3,k)};
         then A68: x = G*(i3,k) by TARSKI:def 1;
           G*(i3,k) in LSeg(G*(i2,k),G*(i3,k)) by TOPREAL1:6;
         then G*(i3,k) in LSeg(G*(i2,j1),G*(i2,k)) \/
                          LSeg(G*(i2,k),G*(i3,k)) by XBOOLE_0:def 2;
         hence x in (LSeg(G*(i2,j1),G*(i2,k)) \/
          LSeg(G*(i2,k),G*(i3,k))) /\ L~Upper_Seq(C,n+1)
                                                  by A63,A68,XBOOLE_0:def 3;
        end;
          (LSeg(G*(i2,j1),G*(i2,k)) \/ LSeg(G*(i2,k),G*(i3,k))) /\
         L~Lower_Seq(C,n+1) = {G*(i2,j1)}
        proof
         thus (LSeg(G*(i2,j1),G*(i2,k)) \/ LSeg(G*(i2,k),G*(i3,k))) /\
          L~Lower_Seq(C,n+1) c= {G*(i2,j1)}
         proof
          let x be set;
          assume x in (LSeg(G*(i2,j1),G*(i2,k)) \/
                       LSeg(G*(i2,k),G*(i3,k))) /\ L~Lower_Seq(C,n+1);
          then x in LSeg(G*(i2,j1),G*(i2,k)) \/ LSeg(G*(i2,k),G*(i3,k)) &
               x in L~Lower_Seq(C,n+1) by XBOOLE_0:def 3;
          then (x in LSeg(G*(i2,j1),G*(i2,k)) or
                x in LSeg(G*(i2,k),G*(i3,k))) &
                x in L~Lower_Seq(C,n+1) by XBOOLE_0:def 2;
          hence x in {G*(i2,j1)} by A8,A55,A57,A65,XBOOLE_0:3,def 3;
         end;
         let x be set;
         assume x in {G*(i2,j1)};
         then A69: x = G*(i2,j1) by TARSKI:def 1;
           G*(i2,j1) in LSeg(G*(i2,j1),G*(i2,k)) by TOPREAL1:6;
         then G*(i2,j1) in LSeg(G*(i2,j1),G*(i2,k)) \/
                           LSeg(G*(i2,k),G*(i3,k)) by XBOOLE_0:def 2;
         hence x in (LSeg(G*(i2,j1),G*(i2,k)) \/
                     LSeg(G*(i2,k),G*(i3,k))) /\ L~Lower_Seq(C,n+1)
                                                  by A58,A69,XBOOLE_0:def 3;
        end;
        then LSeg(G*(i2,j1),G*(i2,k)) \/
             LSeg(G*(i2,k),G*(i3,k)) meets Upper_Arc C
                                       by A2,A3,A56,A59,A60,A62,A67,Th46;
        hence LSeg(G*(i2,j),G*(i2,k)) \/
         LSeg(G*(i2,k),G*(i1,k)) meets Upper_Arc C by A66,XBOOLE_1:63;
       suppose i1 < i2;
        then consider i3 be Nat such that
         A70: i1 <= i3 & i3 <= i2 and
         A71: LSeg(G*(i3,k),G*(i2,k)) /\ L~Upper_Seq(C,n+1) = {G*(i3,k)}
                                                 by A1,A2,A4,A7,A13,Th20;
        A72: 1 < i3 & i3 < len G by A1,A2,A70,AXIOMS:22;
          G*(i3,k) in LSeg(G*(i2,k),G*(i3,k)) /\ L~Upper_Seq(C,n+1)
                                                        by A71,TARSKI:def 1;
        then A73: G*(i3,k) in L~Upper_Seq(C,n+1) by XBOOLE_0:def 3;
        A74: LSeg(G*(i2,j1),G*(i2,k)) c=
         LSeg(G*(i2,j),G*(i2,k)) by A2,A3,A56,Th7;
        A75: LSeg(G*(i2,k),G*(i3,k)) c= LSeg(G*(i2,k),G*(i1,k))
                                                    by A1,A2,A13,A70,Th8;
        then A76: LSeg(G*(i2,j1),G*(i2,k)) \/ LSeg(G*(i2,k),G*(i3,k)) c=
         LSeg(G*(i2,j),G*(i2,k)) \/ LSeg(G*(i2,k),G*(i1,k)) by A74,XBOOLE_1:13;
        A77: (LSeg(G*(i2,j1),G*(i2,k)) \/
         LSeg(G*(i2,k),G*(i3,k))) /\ L~Upper_Seq(C,n+1) = {G*(i3,k)}
        proof
         thus (LSeg(G*(i2,j1),G*(i2,k)) \/
          LSeg(G*(i2,k),G*(i3,k))) /\ L~Upper_Seq(C,n+1) c= {G*(i3,k)}
         proof
          let x be set;
          assume x in (LSeg(G*(i2,j1),G*(i2,k)) \/
           LSeg(G*(i2,k),G*(i3,k))) /\ L~Upper_Seq(C,n+1);
          then x in LSeg(G*(i2,j1),G*(i2,k)) \/ LSeg(G*(i2,k),G*(i3,k)) &
               x in L~Upper_Seq(C,n+1) by XBOOLE_0:def 3;
          then (x in LSeg(G*(i2,j1),G*(i2,k)) or x in LSeg(G*(i2,k),G*(i3,k)))
           & x in L~Upper_Seq(C,n+1) by XBOOLE_0:def 2;
          hence x in {G*(i3,k)} by A7,A55,A71,A74,XBOOLE_0:3,def 3;
         end;
         let x be set;
         assume x in {G*(i3,k)};
         then A78: x = G*(i3,k) by TARSKI:def 1;
           G*(i3,k) in LSeg(G*(i2,k),G*(i3,k)) by TOPREAL1:6;
         then G*(i3,k) in LSeg(G*(i2,j1),G*(i2,k)) \/
                          LSeg(G*(i2,k),G*(i3,k)) by XBOOLE_0:def 2;
         hence x in (LSeg(G*(i2,j1),G*(i2,k)) \/
          LSeg(G*(i2,k),G*(i3,k))) /\ L~Upper_Seq(C,n+1)
                                                  by A73,A78,XBOOLE_0:def 3;
        end;
          (LSeg(G*(i2,j1),G*(i2,k)) \/ LSeg(G*(i2,k),G*(i3,k))) /\
         L~Lower_Seq(C,n+1) = {G*(i2,j1)}
        proof
         thus (LSeg(G*(i2,j1),G*(i2,k)) \/ LSeg(G*(i2,k),G*(i3,k))) /\
          L~Lower_Seq(C,n+1) c= {G*(i2,j1)}
         proof
          let x be set;
          assume x in (LSeg(G*(i2,j1),G*(i2,k)) \/ LSeg(G*(i2,k),G*(i3,k))) /\
           L~Lower_Seq(C,n+1);
          then x in LSeg(G*(i2,j1),G*(i2,k)) \/ LSeg(G*(i2,k),G*(i3,k)) &
               x in L~Lower_Seq(C,n+1) by XBOOLE_0:def 3;
          then (x in LSeg(G*(i2,j1),G*(i2,k)) or x in LSeg(G*(i2,k),G*(i3,k)))
           & x in L~Lower_Seq(C,n+1) by XBOOLE_0:def 2;
          hence x in {G*(i2,j1)} by A8,A55,A57,A75,XBOOLE_0:3,def 3;
         end;
         let x be set;
         assume x in {G*(i2,j1)};
         then A79: x = G*(i2,j1) by TARSKI:def 1;
           G*(i2,j1) in LSeg(G*(i2,j1),G*(i2,k)) by TOPREAL1:6;
         then G*(i2,j1) in LSeg(G*(i2,j1),G*(i2,k)) \/
                           LSeg(G*(i2,k),G*(i3,k)) by XBOOLE_0:def 2;
         hence x in (LSeg(G*(i2,j1),G*(i2,k)) \/
                     LSeg(G*(i2,k),G*(i3,k))) /\ L~Lower_Seq(C,n+1)
                                                  by A58,A79,XBOOLE_0:def 3;
        end;
        then LSeg(G*(i2,j1),G*(i2,k)) \/ LSeg(G*(i2,k),G*(i3,k))
         meets Upper_Arc C by A2,A3,A56,A59,A70,A72,A77,Th48;
        hence LSeg(G*(i2,j),G*(i2,k)) \/
         LSeg(G*(i2,k),G*(i1,k)) meets Upper_Arc C by A76,XBOOLE_1:63;
      end;
      hence LSeg(G*(i2,j),G*(i2,k)) \/ LSeg(G*(i2,k),G*(i1,k))
       meets Upper_Arc C;
    end;
    hence LSeg(G*(i2,j),G*(i2,k)) \/ LSeg(G*(i2,k),G*(i1,k)) meets Upper_Arc C
;
   end;

  theorem Th51:
   for C be Simple_closed_curve
   for i1,i2,j,k be Nat holds
    1 < i1 & i1 < len Gauge(C,n+1) & 1 < i2 & i2 < len Gauge(C,n+1) &
    1 <= j & j <= k & k <= width Gauge(C,n+1) &
    Gauge(C,n+1)*(i1,k) in Upper_Arc L~Cage(C,n+1) &
    Gauge(C,n+1)*(i2,j) in Lower_Arc L~Cage(C,n+1) implies
     LSeg(Gauge(C,n+1)*(i2,j),Gauge(C,n+1)*(i2,k)) \/
     LSeg(Gauge(C,n+1)*(i2,k),Gauge(C,n+1)*(i1,k)) meets Lower_Arc C
   proof
    let C be Simple_closed_curve;
    let i1,i2,j,k be Nat;
    set G=Gauge(C,n+1);
    assume that
     A1: 1 < i1 & i1 < len G and
     A2: 1 < i2 & i2 < len G and
     A3: 1 <= j & j <= k & k <= width G and
     A4: G*(i1,k) in Upper_Arc L~Cage(C,n+1) and
     A5: G*(i2,j) in Lower_Arc L~Cage(C,n+1);
      n+1 >= 0+1 by NAT_1:29;
    then A6: n+1 > 0 by NAT_1:38;
    then A7: Upper_Arc L~Cage(C,n+1) = L~Upper_Seq(C,n+1) by JORDAN1G:63;
    A8: Lower_Arc L~Cage(C,n+1) = L~Lower_Seq(C,n+1) by A6,JORDAN1G:64;
    A9: Upper_Seq(C,n+1) is_sequence_on G by JORDAN1G:4;
    A10: Lower_Seq(C,n+1) is_sequence_on G by JORDAN1G:5;
    A11: 1 <= j & j <= width G by A3,AXIOMS:22;
    then A12: [i2,j] in Indices G by A2,GOBOARD7:10;
    A13: 1 <= k & k <= width G by A3,AXIOMS:22;
    then A14: [i2,k] in Indices G by A2,GOBOARD7:10;
      G*(i2,j)`1 = G*(i2,1)`1 by A2,A11,GOBOARD5:3
       .= G*(i2,k)`1 by A2,A13,GOBOARD5:3;
    then A15: LSeg(G*(i2,j),G*(i2,k)) is vertical by SPPOL_1:37;
      G*(i2,k)`2 = G*(1,k)`2 by A2,A13,GOBOARD5:2
       .= G*(i1,k)`2 by A1,A13,GOBOARD5:2;
    then A16: LSeg(G*(i2,k),G*(i1,k)) is horizontal by SPPOL_1:36;
    A17: [i2,k] in Indices G by A2,A13,GOBOARD7:10;
    A18: [i1,k] in Indices G by A1,A13,GOBOARD7:10;
      now per cases;
     suppose A19: LSeg(G*(i2,j),G*(i2,k)) meets
       Upper_Arc L~Cage(C,n+1);
      then consider m be Nat such that
       A20: j <= m & m <= k and
       A21: G*(i2,m)`2 = inf(proj2.:(LSeg(G*(i2,j),G*(i2,k))
            /\ L~Upper_Seq(C,n+1))) by A3,A7,A9,A12,A14,JORDAN1F:1;
      A22: 1 <= m & m <= width G by A3,A20,AXIOMS:22;
      set X = LSeg(G*(i2,j),G*(i2,k)) /\ L~Upper_Seq(C,n+1);
      A23: G*(i2,m)`1 = G*(i2,1)`1 by A2,A22,GOBOARD5:3;
      then A24: |[G*(i2,1)`1,inf(proj2.:X)]| = G*(i2,m) by A21,EUCLID:57;
      then A25: G*(i2,j)`1 = |[G*(i2,1)`1,inf(proj2.:X)]|`1
                                               by A2,A11,A23,GOBOARD5:3;
        ex x be set st x in LSeg(G*(i2,j),G*(i2,k)) &
       x in L~Upper_Seq(C,n+1) by A7,A19,XBOOLE_0:3;
      then reconsider X1=X as non empty compact Subset of TOP-REAL 2
                                            by PSCOMP_1:64,XBOOLE_0:def 3;
      consider pp be set such that
       A26: pp in S-most X1 by XBOOLE_0:def 1;
      reconsider pp as Point of TOP-REAL 2 by A26;
        pp in LSeg(SW-corner X, SE-corner X)/\X by A26,PSCOMP_1:def 41;
      then A27: pp in X by XBOOLE_0:def 3;
      then A28: pp in L~Upper_Seq(C,n+1) by XBOOLE_0:def 3;
        pp in LSeg(G*(i2,j),G*(i2,k)) by A27,XBOOLE_0:def 3;
      then A29: pp`1 = |[G*(i2,1)`1,inf(proj2.:X)]|`1
               by A15,A25,SPRECT_3:20;
        |[G*(i2,1)`1,inf(proj2.:X)]|`2 = S-bound X by A21,A24,SPRECT_1:49
         .= (S-min X)`2 by PSCOMP_1:114
         .= pp`2 by A26,PSCOMP_1:118;
      then G*(i2,m) in Upper_Arc L~Cage(C,n+1) by A7,A24,A28,A29,TOPREAL3:11;
      then A30: LSeg(G*(i2,j),G*(i2,m)) meets Lower_Arc C
                                              by A2,A3,A5,A6,A20,A22,Th25;
        LSeg(G*(i2,j),G*(i2,m)) c= LSeg(G*(i2,j),G*(i2,k)) by A2,A3,A20,Th7;
      then LSeg(G*(i2,j),G*(i2,k)) meets Lower_Arc C by A30,XBOOLE_1:63;
      hence LSeg(G*(i2,j),G*(i2,k)) \/
            LSeg(G*(i2,k),G*(i1,k)) meets Lower_Arc C by XBOOLE_1:70;
     suppose A31: LSeg(G*(i2,k),G*(i1,k)) meets Lower_Arc L~Cage(C,n+1) &
       i2 <= i1;
      then consider m be Nat such that
       A32: i2 <= m & m <= i1 and
       A33: G*(m,k)`1 = sup(proj1.:(LSeg(G*(i2,k),G*(i1,k))
            /\ L~Lower_Seq(C,n+1))) by A8,A10,A17,A18,JORDAN1F:4;
      A34: 1 < m & m < len G by A1,A2,A32,AXIOMS:22;
      set X = LSeg(G*(i2,k),G*(i1,k)) /\ L~Lower_Seq(C,n+1);
      A35: G*(m,k)`2 = G*(1,k)`2 by A13,A34,GOBOARD5:2;
      then A36: |[sup(proj1.:X),G*(1,k)`2]| = G*(m,k) by A33,EUCLID:57;
      then A37: G*(i2,k)`2 = |[sup(proj1.:X),G*(1,k)`2]|`2
                                               by A2,A13,A35,GOBOARD5:2;
        ex x be set st x in LSeg(G*(i2,k),G*(i1,k)) &
       x in L~Lower_Seq(C,n+1) by A8,A31,XBOOLE_0:3;
      then reconsider X1=X as non empty compact Subset of TOP-REAL 2
                                          by PSCOMP_1:64,XBOOLE_0:def 3;
      consider pp be set such that
       A38: pp in E-most X1 by XBOOLE_0:def 1;
      reconsider pp as Point of TOP-REAL 2 by A38;
        pp in LSeg(SE-corner X, NE-corner X)/\X by A38,PSCOMP_1:def 40;
      then A39: pp in X by XBOOLE_0:def 3;
      then A40: pp in L~Lower_Seq(C,n+1) by XBOOLE_0:def 3;
        pp in LSeg(G*(i2,k),G*(i1,k)) by A39,XBOOLE_0:def 3;
      then A41: pp`2 = |[sup(proj1.:X),G*(1,k)`2]|`2 by A16,A37,SPRECT_3:19;
        |[sup(proj1.:X),G*(1,k)`2]|`1 = E-bound X by A33,A36,SPRECT_1:51
         .= (E-min X)`1 by PSCOMP_1:104
         .= pp`1 by A38,PSCOMP_1:108;
      then G*(m,k) in Lower_Arc L~Cage(C,n+1) by A8,A36,A40,A41,TOPREAL3:11;
      then A42: LSeg(G*(m,k),G*(i1,k)) meets Lower_Arc C
                                             by A1,A4,A6,A13,A32,A34,Th34;
        LSeg(G*(m,k),G*(i1,k)) c= LSeg(G*(i2,k),G*(i1,k))
                                                    by A1,A2,A13,A32,Th8;
      then LSeg(G*(i2,k),G*(i1,k)) meets Lower_Arc C by A42,XBOOLE_1:63;
      hence LSeg(G*(i2,j),G*(i2,k)) \/
            LSeg(G*(i2,k),G*(i1,k)) meets Lower_Arc C by XBOOLE_1:70;
     suppose A43: LSeg(G*(i2,k),G*(i1,k)) meets Lower_Arc L~Cage(C,n+1) &
      i1 < i2;
      then consider m be Nat such that
       A44: i1 <= m & m <= i2 and
       A45: G*(m,k)`1 = inf(proj1.:(LSeg(G*(i1,k),G*(i2,k))
            /\ L~Lower_Seq(C,n+1))) by A8,A10,A17,A18,JORDAN1F:3;
      A46: 1 < m & m < len G by A1,A2,A44,AXIOMS:22;
      set X = LSeg(G*(i1,k),G*(i2,k)) /\ L~Lower_Seq(C,n+1);
      A47: G*(m,k)`2 = G*(1,k)`2 by A13,A46,GOBOARD5:2;
      then A48: |[inf(proj1.:X),G*(1,k)`2]| = G*(m,k) by A45,EUCLID:57;
      then A49: G*(i1,k)`2 = |[inf(proj1.:X),G*(1,k)`2]|`2
                                                by A1,A13,A47,GOBOARD5:2;
        ex x be set st x in LSeg(G*(i1,k),G*(i2,k)) &
       x in L~Lower_Seq(C,n+1) by A8,A43,XBOOLE_0:3;
      then reconsider X1=X as non empty compact Subset of TOP-REAL 2
                                          by PSCOMP_1:64,XBOOLE_0:def 3;
      consider pp be set such that
       A50: pp in W-most X1 by XBOOLE_0:def 1;
      reconsider pp as Point of TOP-REAL 2 by A50;
        pp in LSeg(SW-corner X, NW-corner X)/\X by A50,PSCOMP_1:def 38;
      then A51: pp in X by XBOOLE_0:def 3;
      then A52: pp in L~Lower_Seq(C,n+1) by XBOOLE_0:def 3;
        pp in LSeg(G*(i1,k),G*(i2,k)) by A51,XBOOLE_0:def 3;
      then A53: pp`2 = |[inf(proj1.:X),G*(1,k)`2]|`2 by A16,A49,SPRECT_3:19;
        |[inf(proj1.:X),G*(1,k)`2]|`1 = W-bound X by A45,A48,SPRECT_1:48
         .= (W-min X)`1 by PSCOMP_1:84
         .= pp`1 by A50,PSCOMP_1:88;
      then G*(m,k) in Lower_Arc L~Cage(C,n+1) by A8,A48,A52,A53,TOPREAL3:11;
      then A54: LSeg(G*(i1,k),G*(m,k)) meets Lower_Arc C
                                             by A1,A4,A6,A13,A44,A46,Th42;
        LSeg(G*(i1,k),G*(m,k)) c= LSeg(G*(i1,k),G*(i2,k))
                                                    by A1,A2,A13,A44,Th8;
      then LSeg(G*(i1,k),G*(i2,k)) meets Lower_Arc C by A54,XBOOLE_1:63;
      hence LSeg(G*(i2,j),G*(i2,k)) \/
            LSeg(G*(i2,k),G*(i1,k)) meets Lower_Arc C by XBOOLE_1:70;
     suppose A55: LSeg(G*(i2,j),G*(i2,k)) misses
        Upper_Arc L~Cage(C,n+1) &
       LSeg(Gauge(C,n+1)*(i2,k),Gauge(C,n+1)*(i1,k))
        misses Lower_Arc L~Cage(C,n+1);
      consider j1 be Nat such that
       A56: j <= j1 & j1 <= k and
       A57: LSeg(G*(i2,j1),G*(i2,k)) /\ L~Lower_Seq(C,n+1) =
        {G*(i2,j1)} by A2,A3,A5,A8,Th11;
        G*(i2,j1) in LSeg(G*(i2,j1),G*(i2,k)) /\
       L~Lower_Seq(C,n+1) by A57,TARSKI:def 1;
      then A58: G*(i2,j1) in L~Lower_Seq(C,n+1) by XBOOLE_0:def 3;
      A59: 1 <= j1 & j1 <= width G by A3,A56,AXIOMS:22;
        now per cases;
       suppose i2 <= i1;
        then consider i3 be Nat such that
         A60: i2 <= i3 & i3 <= i1 and
         A61: LSeg(G*(i2,k),G*(i3,k)) /\ L~Upper_Seq(C,n+1) = {G*(i3,k)}
                                                 by A1,A2,A4,A7,A13,Th15;
        A62: 1 < i3 & i3 < len G by A1,A2,A60,AXIOMS:22;
          G*(i3,k) in LSeg(G*(i2,k),G*(i3,k)) /\ L~Upper_Seq(C,n+1)
                                                        by A61,TARSKI:def 1;
        then A63: G*(i3,k) in L~Upper_Seq(C,n+1) by XBOOLE_0:def 3;
        A64: LSeg(G*(i2,j1),G*(i2,k)) c=
         LSeg(G*(i2,j),G*(i2,k)) by A2,A3,A56,Th7;
        A65: LSeg(G*(i2,k),G*(i3,k)) c= LSeg(G*(i2,k),G*(i1,k))
                                                    by A1,A2,A13,A60,Th8;
        then A66: LSeg(G*(i2,j1),G*(i2,k)) \/ LSeg(G*(i2,k),G*(i3,k)) c=
         LSeg(G*(i2,j),G*(i2,k)) \/ LSeg(G*(i2,k),G*(i1,k)) by A64,XBOOLE_1:13;
        A67: (LSeg(G*(i2,j1),G*(i2,k)) \/
         LSeg(G*(i2,k),G*(i3,k))) /\ L~Upper_Seq(C,n+1) = {G*(i3,k)}
        proof
         thus (LSeg(G*(i2,j1),G*(i2,k)) \/
          LSeg(G*(i2,k),G*(i3,k))) /\ L~Upper_Seq(C,n+1) c= {G*(i3,k)}
         proof
          let x be set;
          assume x in (LSeg(G*(i2,j1),G*(i2,k)) \/
           LSeg(G*(i2,k),G*(i3,k))) /\ L~Upper_Seq(C,n+1);
          then x in LSeg(G*(i2,j1),G*(i2,k)) \/
                    LSeg(G*(i2,k),G*(i3,k)) &
               x in L~Upper_Seq(C,n+1) by XBOOLE_0:def 3;
          then (x in LSeg(G*(i2,j1),G*(i2,k)) or
                x in LSeg(G*(i2,k),G*(i3,k))) &
                x in L~Upper_Seq(C,n+1) by XBOOLE_0:def 2;
          hence x in {G*(i3,k)} by A7,A55,A61,A64,XBOOLE_0:3,def 3;
         end;
         let x be set;
         assume x in {G*(i3,k)};
         then A68: x = G*(i3,k) by TARSKI:def 1;
           G*(i3,k) in LSeg(G*(i2,k),G*(i3,k)) by TOPREAL1:6;
         then G*(i3,k) in LSeg(G*(i2,j1),G*(i2,k)) \/
                          LSeg(G*(i2,k),G*(i3,k)) by XBOOLE_0:def 2;
         hence x in (LSeg(G*(i2,j1),G*(i2,k)) \/
          LSeg(G*(i2,k),G*(i3,k))) /\ L~Upper_Seq(C,n+1)
                                                  by A63,A68,XBOOLE_0:def 3;
        end;
          (LSeg(G*(i2,j1),G*(i2,k)) \/ LSeg(G*(i2,k),G*(i3,k))) /\
         L~Lower_Seq(C,n+1) = {G*(i2,j1)}
        proof
         thus (LSeg(G*(i2,j1),G*(i2,k)) \/ LSeg(G*(i2,k),G*(i3,k))) /\
          L~Lower_Seq(C,n+1) c= {G*(i2,j1)}
         proof
          let x be set;
          assume x in (LSeg(G*(i2,j1),G*(i2,k)) \/
                       LSeg(G*(i2,k),G*(i3,k))) /\ L~Lower_Seq(C,n+1);
          then x in LSeg(G*(i2,j1),G*(i2,k)) \/ LSeg(G*(i2,k),G*(i3,k)) &
               x in L~Lower_Seq(C,n+1) by XBOOLE_0:def 3;
          then (x in LSeg(G*(i2,j1),G*(i2,k)) or
                x in LSeg(G*(i2,k),G*(i3,k))) &
                x in L~Lower_Seq(C,n+1) by XBOOLE_0:def 2;
          hence x in {G*(i2,j1)} by A8,A55,A57,A65,XBOOLE_0:3,def 3;
         end;
         let x be set;
         assume x in {G*(i2,j1)};
         then A69: x = G*(i2,j1) by TARSKI:def 1;
           G*(i2,j1) in LSeg(G*(i2,j1),G*(i2,k)) by TOPREAL1:6;
         then G*(i2,j1) in LSeg(G*(i2,j1),G*(i2,k)) \/
                           LSeg(G*(i2,k),G*(i3,k)) by XBOOLE_0:def 2;
         hence x in (LSeg(G*(i2,j1),G*(i2,k)) \/
                     LSeg(G*(i2,k),G*(i3,k))) /\ L~Lower_Seq(C,n+1)
                                                  by A58,A69,XBOOLE_0:def 3;
        end;
        then LSeg(G*(i2,j1),G*(i2,k)) \/
             LSeg(G*(i2,k),G*(i3,k)) meets Lower_Arc C
                                       by A2,A3,A56,A59,A60,A62,A67,Th47;
        hence LSeg(G*(i2,j),G*(i2,k)) \/
         LSeg(G*(i2,k),G*(i1,k)) meets Lower_Arc C by A66,XBOOLE_1:63;
       suppose i1 < i2;
        then consider i3 be Nat such that
         A70: i1 <= i3 & i3 <= i2 and
         A71: LSeg(G*(i3,k),G*(i2,k)) /\ L~Upper_Seq(C,n+1) = {G*(i3,k)}
                                                 by A1,A2,A4,A7,A13,Th20;
        A72: 1 < i3 & i3 < len G by A1,A2,A70,AXIOMS:22;
          G*(i3,k) in LSeg(G*(i2,k),G*(i3,k)) /\ L~Upper_Seq(C,n+1)
                                                        by A71,TARSKI:def 1;
        then A73: G*(i3,k) in L~Upper_Seq(C,n+1) by XBOOLE_0:def 3;
        A74: LSeg(G*(i2,j1),G*(i2,k)) c=
         LSeg(G*(i2,j),G*(i2,k)) by A2,A3,A56,Th7;
        A75: LSeg(G*(i2,k),G*(i3,k)) c= LSeg(G*(i2,k),G*(i1,k))
                                                    by A1,A2,A13,A70,Th8;
        then A76: LSeg(G*(i2,j1),G*(i2,k)) \/ LSeg(G*(i2,k),G*(i3,k)) c=
         LSeg(G*(i2,j),G*(i2,k)) \/ LSeg(G*(i2,k),G*(i1,k)) by A74,XBOOLE_1:13;
        A77: (LSeg(G*(i2,j1),G*(i2,k)) \/
         LSeg(G*(i2,k),G*(i3,k))) /\ L~Upper_Seq(C,n+1) = {G*(i3,k)}
        proof
         thus (LSeg(G*(i2,j1),G*(i2,k)) \/
          LSeg(G*(i2,k),G*(i3,k))) /\ L~Upper_Seq(C,n+1) c= {G*(i3,k)}
         proof
          let x be set;
          assume x in (LSeg(G*(i2,j1),G*(i2,k)) \/
           LSeg(G*(i2,k),G*(i3,k))) /\ L~Upper_Seq(C,n+1);
          then x in LSeg(G*(i2,j1),G*(i2,k)) \/ LSeg(G*(i2,k),G*(i3,k)) &
               x in L~Upper_Seq(C,n+1) by XBOOLE_0:def 3;
          then (x in LSeg(G*(i2,j1),G*(i2,k)) or x in LSeg(G*(i2,k),G*(i3,k)))
           & x in L~Upper_Seq(C,n+1) by XBOOLE_0:def 2;
          hence x in {G*(i3,k)} by A7,A55,A71,A74,XBOOLE_0:3,def 3;
         end;
         let x be set;
         assume x in {G*(i3,k)};
         then A78: x = G*(i3,k) by TARSKI:def 1;
           G*(i3,k) in LSeg(G*(i2,k),G*(i3,k)) by TOPREAL1:6;
         then G*(i3,k) in LSeg(G*(i2,j1),G*(i2,k)) \/
                          LSeg(G*(i2,k),G*(i3,k)) by XBOOLE_0:def 2;
         hence x in (LSeg(G*(i2,j1),G*(i2,k)) \/
          LSeg(G*(i2,k),G*(i3,k))) /\ L~Upper_Seq(C,n+1)
                                                  by A73,A78,XBOOLE_0:def 3;
        end;
          (LSeg(G*(i2,j1),G*(i2,k)) \/ LSeg(G*(i2,k),G*(i3,k))) /\
         L~Lower_Seq(C,n+1) = {G*(i2,j1)}
        proof
         thus (LSeg(G*(i2,j1),G*(i2,k)) \/ LSeg(G*(i2,k),G*(i3,k))) /\
          L~Lower_Seq(C,n+1) c= {G*(i2,j1)}
         proof
          let x be set;
          assume x in (LSeg(G*(i2,j1),G*(i2,k)) \/ LSeg(G*(i2,k),G*(i3,k))) /\
           L~Lower_Seq(C,n+1);
          then x in LSeg(G*(i2,j1),G*(i2,k)) \/ LSeg(G*(i2,k),G*(i3,k)) &
               x in L~Lower_Seq(C,n+1) by XBOOLE_0:def 3;
          then (x in LSeg(G*(i2,j1),G*(i2,k)) or x in LSeg(G*(i2,k),G*(i3,k)))
           & x in L~Lower_Seq(C,n+1) by XBOOLE_0:def 2;
          hence x in {G*(i2,j1)} by A8,A55,A57,A75,XBOOLE_0:3,def 3;
         end;
         let x be set;
         assume x in {G*(i2,j1)};
         then A79: x = G*(i2,j1) by TARSKI:def 1;
           G*(i2,j1) in LSeg(G*(i2,j1),G*(i2,k)) by TOPREAL1:6;
         then G*(i2,j1) in LSeg(G*(i2,j1),G*(i2,k)) \/
                           LSeg(G*(i2,k),G*(i3,k)) by XBOOLE_0:def 2;
         hence x in (LSeg(G*(i2,j1),G*(i2,k)) \/
                     LSeg(G*(i2,k),G*(i3,k))) /\ L~Lower_Seq(C,n+1)
                                                  by A58,A79,XBOOLE_0:def 3;
        end;
        then LSeg(G*(i2,j1),G*(i2,k)) \/ LSeg(G*(i2,k),G*(i3,k))
         meets Lower_Arc C by A2,A3,A56,A59,A70,A72,A77,Th49;
        hence LSeg(G*(i2,j),G*(i2,k)) \/
         LSeg(G*(i2,k),G*(i1,k)) meets Lower_Arc C by A76,XBOOLE_1:63;
      end;
      hence LSeg(G*(i2,j),G*(i2,k)) \/ LSeg(G*(i2,k),G*(i1,k))
       meets Lower_Arc C;
    end;
    hence LSeg(G*(i2,j),G*(i2,k)) \/ LSeg(G*(i2,k),G*(i1,k)) meets Lower_Arc C
;
   end;

  theorem
     for C be Simple_closed_curve
   for i,j,k be Nat holds
    1 < i & i < len Gauge(C,n+1) &
    1 <= j & j <= k & k <= width Gauge(C,n+1) &
    Gauge(C,n+1)*(i,k) in Upper_Arc L~Cage(C,n+1) &
    Gauge(C,n+1)*(Center Gauge(C,n+1),j) in Lower_Arc L~Cage(C,n+1) implies
     LSeg(Gauge(C,n+1)*(Center Gauge(C,n+1),j),
          Gauge(C,n+1)*(Center Gauge(C,n+1),k)) \/
     LSeg(Gauge(C,n+1)*(Center Gauge(C,n+1),k),Gauge(C,n+1)*(i,k)) meets
      Upper_Arc C
   proof
    let C be Simple_closed_curve;
    let i,j,k be Nat;
    assume that
     A1: 1 < i & i < len Gauge(C,n+1) and
     A2: 1 <= j & j <= k & k <= width Gauge(C,n+1) and
     A3: Gauge(C,n+1)*(i,k) in Upper_Arc L~Cage(C,n+1) and
     A4: Gauge(C,n+1)*(Center Gauge(C,n+1),j) in Lower_Arc L~Cage(C,n+1);
      len Gauge(C,n+1) >= 4 by JORDAN8:13;
    then len Gauge(C,n+1) >= 2 & len Gauge(C,n+1) >= 3 by AXIOMS:22;
    then 1 < Center Gauge(C,n+1) & Center Gauge(C,n+1) < len Gauge(C,n+1)
                                                           by JORDAN1B:15,16;
    hence thesis by A1,A2,A3,A4,Th50;
   end;

  theorem
     for C be Simple_closed_curve
   for i,j,k be Nat holds
    1 < i & i < len Gauge(C,n+1) &
    1 <= j & j <= k & k <= width Gauge(C,n+1) &
    Gauge(C,n+1)*(i,k) in Upper_Arc L~Cage(C,n+1) &
    Gauge(C,n+1)*(Center Gauge(C,n+1),j) in Lower_Arc L~Cage(C,n+1) implies
     LSeg(Gauge(C,n+1)*(Center Gauge(C,n+1),j),
          Gauge(C,n+1)*(Center Gauge(C,n+1),k)) \/
     LSeg(Gauge(C,n+1)*(Center Gauge(C,n+1),k),Gauge(C,n+1)*(i,k)) meets
      Lower_Arc C
   proof
    let C be Simple_closed_curve;
    let i,j,k be Nat;
    assume that
     A1: 1 < i & i < len Gauge(C,n+1) and
     A2: 1 <= j & j <= k & k <= width Gauge(C,n+1) and
     A3: Gauge(C,n+1)*(i,k) in Upper_Arc L~Cage(C,n+1) and
     A4: Gauge(C,n+1)*(Center Gauge(C,n+1),j) in Lower_Arc L~Cage(C,n+1);
      len Gauge(C,n+1) >= 4 by JORDAN8:13;
    then len Gauge(C,n+1) >= 2 & len Gauge(C,n+1) >= 3 by AXIOMS:22;
    then 1 < Center Gauge(C,n+1) & Center Gauge(C,n+1) < len Gauge(C,n+1)
                                                           by JORDAN1B:15,16;
    hence thesis by A1,A2,A3,A4,Th51;
   end;

Back to top