The Mizar article:

General Fashoda Meet Theorem for Unit Circle and Square

by
Yatsuka Nakamura

Received February 25, 2003

Copyright (c) 2003 Association of Mizar Users

MML identifier: JGRAPH_6
[ MML identifier index ]


environ

 vocabulary FUNCT_1, BOOLE, ABSVALUE, EUCLID, PRE_TOPC, SQUARE_1, RELAT_1,
      SUBSET_1, ARYTM_3, METRIC_1, RCOMP_1, FUNCT_5, TOPMETR, COMPTS_1,
      ORDINAL2, TOPS_2, ARYTM_1, ARYTM, COMPLEX1, MCART_1, PCOMPS_1, JGRAPH_3,
      BORSUK_1, TOPREAL1, TOPREAL2, JORDAN3, PSCOMP_1, REALSET1, JORDAN5C,
      JORDAN6, JGRAPH_2, JGRAPH_6, RELAT_2, FINSEQ_1, FINSEQ_4, TARSKI, TOPS_1,
      FUNCT_4, JORDAN17;
 notation XBOOLE_0, ORDINAL1, ABSVALUE, EUCLID, TARSKI, RELAT_1, TOPS_2,
      FUNCT_1, FUNCT_2, SUBSET_1, FINSEQ_1, STRUCT_0, TOPMETR, PCOMPS_1,
      COMPTS_1, METRIC_1, SQUARE_1, RCOMP_1, PRE_TOPC, FUNCT_4, JGRAPH_3,
      JORDAN5C, JORDAN6, TOPREAL2, CONNSP_1, FINSEQ_4, TOPS_1, JORDAN17,
      JGRAPH_2, NUMBERS, XCMPLX_0, XREAL_0, TOPRNS_1, TOPREAL1, PSCOMP_1,
      REAL_1, NAT_1;
 constructors TOPS_1, RCOMP_1, JGRAPH_2, TOPREAL2, TOPGRP_1, CONNSP_1,
      WELLFND1, JGRAPH_3, JORDAN5C, JORDAN6, TOPRNS_1, TREAL_1, FINSEQ_4,
      TOPS_2, JORDAN17, TOPREAL1, PSCOMP_1, REAL_1, NAT_1;
 clusters STRUCT_0, RELSET_1, FUNCT_1, EUCLID, PRE_TOPC, TOPMETR, SQUARE_1,
      PSCOMP_1, BORSUK_1, TOPREAL2, BORSUK_2, TOPREAL1, JGRAPH_3, TOPS_1,
      JGRAPH_2, TOPREAL6, XREAL_0, NAT_1, MEMBERED, XBOOLE_0;
 requirements REAL, BOOLE, SUBSET, NUMERALS, ARITHM;
 definitions TARSKI, XBOOLE_0, TOPREAL1;
 theorems TARSKI, XBOOLE_0, XBOOLE_1, AXIOMS, RELAT_1, FUNCT_1, FUNCT_2,
      TOPS_1, TOPS_2, PRE_TOPC, TOPMETR, JORDAN6, EUCLID, REAL_1, REAL_2,
      JGRAPH_1, SQUARE_1, PSCOMP_1, METRIC_1, JGRAPH_2, RCOMP_1, COMPTS_1,
      RFUNCT_2, BORSUK_1, TOPREAL1, TOPREAL3, TOPREAL5, JGRAPH_3, ABSVALUE,
      JORDAN7, HEINE, JGRAPH_4, JORDAN5C, JGRAPH_5, GOBOARD6, JORDAN2C,
      TOPREAL2, TREAL_1, SEQ_2, FINSEQ_1, FINSEQ_3, FINSEQ_4, ZFMISC_1, NAT_1,
      ENUMSET1, UNIFORM1, BORSUK_4, TSEP_1, FUNCT_4, JORDAN1A, JORDAN17,
      JORDAN1, GOBOARD7, XREAL_0, TOPRNS_1, COMPLEX1, XCMPLX_0, XCMPLX_1;
 schemes FUNCT_2, JGRAPH_2, TOPREAL1, FRAENKEL, FUNCT_1, DOMAIN_1;

begin :: Preliminaries

canceled;

theorem Th2: for a,b,r being real number st 0<=r & r<=1 & a<=b holds
 a<=(1-r)*a+r*b & (1-r)*a+r*b<=b
proof let a,b,r be real number;
 assume A1: 0<=r & r<=1 & a<=b;
  then A2: 1-r>=0 by SQUARE_1:12;
    r*a<=r*b by A1,AXIOMS:25;
  then (1-r)*a+r*a<=(1-r)*a+r*b by REAL_1:55;
  then ((1-r)+r)*a<=(1-r)*a+r*b by XCMPLX_1:8;
  then 1*a<=(1-r)*a+r*b by XCMPLX_1:27;
 hence a<=(1-r)*a+r*b;
    (1-r)*a<=(1-r)*b by A1,A2,AXIOMS:25;
  then (1-r)*a+r*b<=(1-r)*b+r*b by REAL_1:55;
  then ((1-r)+r)*b>=(1-r)*a+r*b by XCMPLX_1:8;
  then 1*b>=(1-r)*a+r*b by XCMPLX_1:27;
 hence (1-r)*a+r*b<=b;
end;

theorem Th3:
  for a,b being real number st a>=0 & b>0 or a>0 & b>=0 holds a+b>0
proof let a,b be real number;
 assume A1: a>=0 & b>0 or a>0 & b>=0;
    now per cases by A1;
  case a>=0 & b>0;
   hence a+b>0 by REAL_1:69;
  case a>0 & b>=0;
   hence a+b>0 by REAL_1:69;
  end;
 hence a+b>0;
end;

theorem Th4: for a,b being real number st -1<=a & a<=1 & -1<=b & b<=1
holds a^2*b^2<=1
proof let a,b be real number;
 assume A1: -1<=a & a<=1 & -1<=b & b<=1;
  then A2: a^2<=1 by JGRAPH_2:7,SQUARE_1:59;
  A3: b^2<=1 by A1,JGRAPH_2:7,SQUARE_1:59;
    0<=b^2 by SQUARE_1:72;
  then a^2*b^2 <= 1*b^2 by A2,AXIOMS:25;
 hence a^2*b^2<=1 by A3,AXIOMS:22;
end;

theorem Th5: for a,b being real number st a>=0 & b>=0
holds a*sqrt(b)=sqrt(a^2*b)
proof let a,b be real number;
 assume A1: a>=0 & b>=0;
  then A2: sqrt(a^2)=a by SQUARE_1:89;
    a^2>=0 by SQUARE_1:72;
 hence a*sqrt(b)=sqrt(a^2*b) by A1,A2,SQUARE_1:97;
end;

Lm1: for a,b being real number st -1<=a & a<=1 & -1<=b & b<=1
holds (1+a^2)*b^2<=1+b^2
proof let a,b be real number;
  assume -1<=a & a<=1 & -1<=b & b<=1;
  then a^2*b^2<=1 by Th4;
  then 1*b^2+a^2*b^2<=1+b^2 by REAL_1:55;
 hence (1+a^2)*b^2<=1+b^2 by XCMPLX_1:8;
end;

theorem Th6: for a,b being real number st -1<=a & a<=1 & -1<=b & b<=1
holds (-b)*sqrt(1+a^2)<=sqrt(1+b^2) & -sqrt(1+b^2)<= b*sqrt(1+a^2)
proof let a,b be real number;
  assume A1: -1<=a & a<=1 & -1<=b & b<=1;
    a^2>=0 by SQUARE_1:72;
  then 1+a^2>=1+0 by REAL_1:55;
  then A2: 1+a^2>=0 by AXIOMS:22;
  then A3: sqrt(1+a^2)>=0 by SQUARE_1:def 4;
    b^2>=0 by SQUARE_1:72;
  then 1+b^2>=1+0 by REAL_1:55;
  then 1+b^2>=0 by AXIOMS:22;
  then A4: sqrt(1+b^2)>=0 by SQUARE_1:def 4;
  A5: now per cases;
  case b>=0;
    then -b<=0 by REAL_1:66;
   hence (-b)*sqrt(1+a^2)<=sqrt(1+b^2) by A3,A4,SQUARE_1:23;
  case b<0; then -b>0 by REAL_1:66;
    then A6: (-b)*sqrt(1+a^2)=sqrt((-b)^2*(1+a^2)) by A2,Th5;
      (1+a^2)*b^2<=1+b^2 by A1,Lm1;
    then A7: (-b)^2*(1+a^2)<=1+b^2 by SQUARE_1:61;
      (-b)^2>=0 by SQUARE_1:72;
    then (-b)^2*(1+a^2)>=0 by A2,SQUARE_1:19;
   hence (-b)*sqrt(1+a^2)<=sqrt(1+b^2) by A6,A7,SQUARE_1:94;
  end;
  then -((-b)*sqrt(1+a^2)) >= - sqrt(1+b^2) by REAL_1:50;
  then ((--b)*sqrt(1+a^2)) >= - sqrt(1+b^2) by XCMPLX_1:175;
 hence (-b)*sqrt(1+a^2)<=sqrt(1+b^2) & -sqrt(1+b^2)<= b*sqrt(1+a^2) by A5;
end;

theorem Th7: for a,b being real number st -1<=a & a<=1 & -1<=b & b<=1
holds b*sqrt(1+a^2)<=sqrt(1+b^2)
proof let a,b be real number;
  assume A1: -1<=a & a<=1 & -1<=b & b<=1;
   then A2: --1>=-b by REAL_1:50;
     -1<=-b by A1,REAL_1:50;
   then (--b)*sqrt(1+a^2)<=sqrt(1+(-b)^2) by A1,A2,Th6;
 hence b*sqrt(1+a^2)<=sqrt(1+b^2) by SQUARE_1:61;
end;

Lm2: for a,b being real number st b<=0 & a<=b
holds a*sqrt(1+b^2)<= b*sqrt(1+a^2)
proof let a,b be real number;
 assume A1: b<=0 & a<=b;
    a^2>=0 by SQUARE_1:72;
  then 1+a^2>=1+0 by REAL_1:55;
  then A2: 1+a^2>=0 by AXIOMS:22;
    b^2>=0 by SQUARE_1:72;
  then 1+b^2>=1+0 by REAL_1:55;
  then A3: 1+b^2>=0 by AXIOMS:22;
    --b<=0 by A1;
  then -b>=0 by REAL_1:66;
  then A4: (-b)*sqrt(1+a^2)=sqrt((-b)^2*(1+a^2)) by A2,Th5;
    --a<=0 by A1;
  then -a>=0 by REAL_1:66;
  then A5: (-a)*sqrt(1+b^2)=sqrt((-a)^2*(1+b^2)) by A3,Th5;
  A6: (-a)^2=a^2 by SQUARE_1:61;
  A7: (-b)^2=b^2 by SQUARE_1:61;
    a<b or a=b by A1,REAL_1:def 5;
  then b^2<a^2 or a=b by A1,JGRAPH_5:2;
  then b^2*1+b^2*a^2<=a^2*1+a^2*b^2 by REAL_1:55;
  then b^2*1+b^2*a^2<=a^2*(1+b^2) by XCMPLX_1:8;
  then A8: b^2*(1+a^2)<=a^2*(1+b^2) by XCMPLX_1:8;
    (b)^2>=0 by SQUARE_1:72;
  then (b)^2*(1+a^2)>=0 by A2,SQUARE_1:19;
  then sqrt((a)^2*(1+b^2))>=sqrt(b^2*(1+a^2)) by A8,SQUARE_1:94;
  then -(a*sqrt(1+b^2))>=(-b)*sqrt(1+a^2) by A4,A5,A6,A7,XCMPLX_1:175;
  then -(a*sqrt(1+b^2))>= -(b*sqrt(1+a^2)) by XCMPLX_1:175;
   hence (a)*sqrt(1+b^2)<=b*sqrt(1+a^2) by REAL_1:50;
end;

Lm3: for a,b being real number st a<=0 & a<=b
holds a*sqrt(1+b^2)<= b*sqrt(1+a^2)
proof let a,b be real number;
 assume A1: a<=0 & a<=b;
   now per cases;
 case b<=0;
  hence b*sqrt(1+a^2)>= a*sqrt(1+b^2) by A1,Lm2;
 case A2: b>0;
      A3: (b)^2 >=0 by SQUARE_1:72;
        1+(b)^2 >(b)^2 by REAL_1:69;
      then sqrt(1+(b)^2)>0 by A3,SQUARE_1:93;
      then A4: a*sqrt(1+b^2)<=0 by A1,SQUARE_1:23;
      A5: (a)^2 >=0 by SQUARE_1:72;
        1+(a)^2 >(a)^2 by REAL_1:69;
      then sqrt(1+(a)^2)>0 by A5,SQUARE_1:93;
  hence b*sqrt(1+a^2)>= a*sqrt(1+b^2) by A2,A4,REAL_2:121;
 end;
 hence thesis;
end;

Lm4: for a,b being real number st a>=0 & a>=b
holds a*sqrt(1+b^2)>= b*sqrt(1+a^2)
proof let a,b be real number;
 assume A1: a>=0 & a>=b;
  then A2: -a<=0 by REAL_1:66;
    -a <= -b by A1,REAL_1:50;
  then (-a)*sqrt(1+(-b)^2)<= (-b)*sqrt(1+(-a)^2) by A2,Lm3;
  then (-a)*sqrt(1+(b)^2)<= (-b)*sqrt(1+(-a)^2) by SQUARE_1:61;
  then (-a)*sqrt(1+(b)^2)<= (-b)*sqrt(1+(a)^2) by SQUARE_1:61;
  then -(a*sqrt(1+(b)^2))<= (-b)*sqrt(1+(a)^2) by XCMPLX_1:175;
  then -(a*sqrt(1+(b)^2))<= -(b*sqrt(1+(a)^2)) by XCMPLX_1:175;
  hence a*sqrt(1+b^2)>= b*sqrt(1+a^2) by REAL_1:50;
end;

theorem Th8: for a,b being real number st a>=b
holds a*sqrt(1+b^2)>= b*sqrt(1+a^2)
proof let a,b be real number;
 assume A1: a>=b;
  per cases;
  suppose a>=0;
   hence thesis by A1,Lm4;
  suppose a<0;
   hence thesis by A1,Lm2;
end;

theorem Th9: for a,c,d being real number,p being Point of TOP-REAL 2
st c <=d & p in LSeg(|[a,c]|,|[a,d]|) holds p`1=a & c <=p`2 & p`2<=d
proof let a,c,d be real number,p be Point of TOP-REAL 2;
 assume A1: c <=d & p in LSeg(|[a,c]|,|[a,d]|);
  reconsider a2=a,c2=c,d2=d as Real by XREAL_0:def 1;
    p in LSeg(|[a2,c2]|,|[a2,d2]|) by A1;
 hence p`1=a by TOPREAL3:17;
  A2: (|[a,c]|)`2=c by EUCLID:56;
    (|[a,d]|)`2=d by EUCLID:56;
 hence c <=p`2 & p`2<=d by A1,A2,TOPREAL1:10;
end;

theorem Th10: for a,c,d being real number,p being Point of TOP-REAL 2
st c <d & p`1=a & c <=p`2 & p`2<=d holds p in LSeg(|[a,c]|,|[a,d]|)
proof let a,c,d be real number,p be Point of TOP-REAL 2;
 assume A1: c <d & p`1=a & c <=p`2 & p`2<=d;
  then A2: d-c>0 by SQUARE_1:11;
  reconsider w=(p`2-c)/(d-c) as Real by XREAL_0:def 1;
  A3: (1-w)*(|[a,c]|)+w*(|[a,d]|)
    =|[(1-w)*a,(1-w)*c]|+w*(|[a,d]|) by EUCLID:62
   .=|[(1-w)*a,(1-w)*c]|+(|[w*a,w*d]|) by EUCLID:62
   .=|[(1-w)*a+w*a,(1-w)*c+w*d]| by EUCLID:60
   .=|[((1-w)+w)*a,(1-w)*c+w*d]| by XCMPLX_1:8
   .=|[(1)*a,(1-w)*c+w*d]| by XCMPLX_1:27
   .=|[a,(1*c-w*c)+w*d]| by XCMPLX_1:40
   .= |[a,c+w*d-w*c]| by XCMPLX_1:29
   .= |[a,c+(w*d-w*c)]| by XCMPLX_1:29
   .= |[a,c+w*(d-c)]| by XCMPLX_1:40
   .= |[a,c+(p`2-c)]| by A2,XCMPLX_1:88
   .= |[a,p`2]| by XCMPLX_1:27
   .= p by A1,EUCLID:57;
  A4: p`2-c>=0 by A1,SQUARE_1:12;
    p`2-c <=d-c by A1,REAL_1:49;
  then w<=(d-c)/(d-c) by A2,REAL_1:73;
  then 0<=w & w<=1 by A2,A4,REAL_2:125,XCMPLX_1:60;
  then p in { (1-lambda)*(|[a,c]|) + lambda*(|[a,d]|) where lambda is Real
             : 0 <= lambda & lambda <= 1 } by A3;
 hence thesis by TOPREAL1:def 4;
end;

theorem Th11: for a,b,d being real number,p being Point of TOP-REAL 2
st a <=b & p in LSeg(|[a,d]|,|[b,d]|) holds p`2=d & a <=p`1 & p`1<=b
proof let a,b,d be real number,p be Point of TOP-REAL 2;
 assume A1: a <=b & p in LSeg(|[a,d]|,|[b,d]|);
  reconsider a2=a,b2=b,d2=d as Real by XREAL_0:def 1;
    p in LSeg(|[a2,d2]|,|[b2,d2]|) by A1;
 hence p`2=d by TOPREAL3:18;
  A2: (|[a,d]|)`1=a by EUCLID:56;
    (|[b,d]|)`1=b by EUCLID:56;
 hence a <=p`1 & p`1<=b by A1,A2,TOPREAL1:9;
end;

theorem Th12: :: BORSUK_4:48
  for a,b being real number,B being Subset of I[01] st
   B=[.a,b.] holds B is closed
proof let a,b be real number,B be Subset of I[01];
 assume A1: B=[.a,b.];
    B c= the carrier of R^1 by BORSUK_1:83,TOPMETR:24,XBOOLE_1:1;
  then reconsider B2=B as Subset of R^1;
  A2: B2 is closed by A1,TREAL_1:4;
  reconsider I1=[.0,1.] as Subset of R^1 by TOPMETR:24;
  A3: [#](R^1|I1)=the carrier of I[01] by BORSUK_1:83,PRE_TOPC:def 10;
  A4: I[01]=R^1|I1 by TOPMETR:26,27;
    B=B2 /\ [#](R^1|I1) by A3,XBOOLE_1:28;
 hence B is closed by A2,A4,PRE_TOPC:43;
end;

theorem Th13: for X being TopStruct,Y,Z being non empty TopStruct,
f being map of X,Y, g being map of X,Z holds dom f=dom g
& dom f=the carrier of X & dom f=[#]X
proof let X be TopStruct,Y,Z be non empty TopStruct,
f be map of X,Y, g be map of X,Z;
    dom f=the carrier of X by FUNCT_2:def 1;
 hence dom f=dom g
  & dom f=the carrier of X & dom f=[#]X by FUNCT_2:def 1,PRE_TOPC:12;
end;

theorem Th14:
for X being non empty TopSpace,B being non empty Subset of X
ex f being map of X|B,X st (for p being Point of X|B holds f.p=p) &
  f is continuous
proof let X be non empty TopSpace,B be non empty Subset of X;
  defpred P[set,set] means (for p being Point of X|B holds $2=$1);
  A1: the carrier of X|B = [#](X|B) by PRE_TOPC:12;
  A2: [#](X|B)= B by PRE_TOPC:def 10;
  A3: for x being Element of X|B
    ex y being Element of X st P[x,y]
    proof let x be Element of X|B;
        x in B by A1,A2;
      then reconsider px=x as Point of X;
      set y0=px;
        P[x,y0];
     hence ex y being Element of X st P[x,y];
    end;
    ex g being Function of the carrier of X|B,the carrier of X st
        for x being Element of X|B holds P[x,g.x]
                   from FuncExD(A3);
  then consider g being Function of the carrier of X|B,the carrier of X such
that
  A4: for x being Element of X|B holds P[x,g.x];
  A5: for p being Point of X|B holds g.p=p by A4;
  A6: for r0 being Point of X|B,V being Subset of X
  st g.r0 in V & V is open holds
  ex W being Subset of X|B st r0 in W & W is open & g.:W c= V
  proof let r0 be Point of X|B,V be Subset of X;
    assume A7: g.r0 in V & V is open;
      V /\ [#](X|B) c= [#](X|B) by XBOOLE_1:17;
    then reconsider W2=V /\ [#](X|B) as Subset of X|B by A1;
      r0 in the carrier of (X|B);
    then A8: r0 in [#](X|B) by PRE_TOPC:12;
      g.r0=r0 by A4;
    then A9: r0 in W2 by A7,A8,XBOOLE_0:def 3;
    A10: W2 is open by A7,TOPS_2:32;
      g.:W2 c= V
    proof let y be set;assume y in g.:W2;
      then consider x being set such that
      A11: x in dom g & x in W2 & y=g.x by FUNCT_1:def 12;
      reconsider px=x as Point of X|B by A11;
        g.px=px by A4;
     hence y in V by A11,XBOOLE_0:def 3;
    end;
   hence ex W being Subset of X|B st r0 in W & W is open & g.:W c= V
                           by A9,A10;
  end;
  reconsider g1=g as map of X|B,X ;
    g1 is continuous by A6,JGRAPH_2:20;
 hence thesis by A5;
end;

theorem Th15:
for X being non empty TopSpace,
f1 being map of X,R^1,a being real number st f1 is continuous
holds ex g being map of X,R^1
st (for p being Point of X,r1 being real number st
f1.p=r1 holds g.p=r1-a) & g is continuous
proof let X be non empty TopSpace,
f1 be map of X,R^1,a be real number;
 assume f1 is continuous;
  then consider g1 being map of X,R^1 such that
  A1: (for p being Point of X,r1 being real number st
   f1.p=r1 holds g1.p=r1+-a) & g1 is continuous by JGRAPH_2:34;
    for p being Point of X,r1 being real number st
   f1.p=r1 holds g1.p=r1-a
   proof let p be Point of X,r1 be real number;
     assume f1.p=r1;
      then g1.p=r1+-a by A1;
     hence g1.p=r1-a by XCMPLX_0:def 8;
   end;
 hence thesis by A1;
end;

theorem Th16:
for X being non empty TopSpace,
f1 being map of X,R^1,a being real number st f1 is continuous
holds ex g being map of X,R^1
st (for p being Point of X,r1 being real number st
f1.p=r1 holds g.p=a-r1) & g is continuous
proof let X be non empty TopSpace,
f1 be map of X,R^1,a be real number;
 assume f1 is continuous;
   then consider g1 being map of X,R^1 such that
   A1: (for p being Point of X,r1 being real number st
     f1.p=r1 holds g1.p=r1-a) & g1 is continuous by Th15;
   consider g2 being map of X,R^1 such that
   A2: (for p being Point of X,r1 being real number st
     g1.p=r1 holds g2.p= -r1) & g2 is continuous by A1,JGRAPH_4:13;
     for p being Point of X,r1 being real number st
   f1.p=r1 holds g2.p=a-r1
   proof let p be Point of X,r1 be real number;
    assume f1.p=r1; then g1.p=r1-a by A1;
      then g2.p=-(r1-a) by A2 .=0-(r1-a) by XCMPLX_1:150
      .=0-r1+a by XCMPLX_1:37
      .=a+-r1 by XCMPLX_1:150 .=a-r1 by XCMPLX_0:def 8;
    hence g2.p=a-r1;
   end;
 hence thesis by A2;
end;

theorem Th17: for X being non empty TopSpace, n being Nat,
p being Point of TOP-REAL n,
f being map of X,R^1 st f is continuous ex g being map of X,TOP-REAL n
st (for r being Point of X holds g.r=(f.r)*p) &
g is continuous
proof let X be non empty TopSpace, n be Nat,
p be Point of TOP-REAL n,
f be map of X,R^1;
 assume A1: f is continuous;
   defpred P[set,set] means $2=(f.$1)*p;
  A2: for x being Element of X
    ex y being Element of TOP-REAL n st P[x,y];
    ex g being Function of the carrier of X,the carrier of TOP-REAL n st
        for x being Element of X holds P[x,g.x]
                   from FuncExD(A2);
  then consider g being Function of the carrier of X,the carrier of TOP-REAL n
    such that
  A3: for x being Element of X holds P[x,g.x];
  reconsider g as map of X,TOP-REAL n ;
    for r0 being Point of X,V being Subset of TOP-REAL n
  st g.r0 in V & V is open holds
  ex W being Subset of X st r0 in W & W is open & g.:W c= V
  proof let r0 be Point of X,V be Subset of TOP-REAL n;
    assume A4: g.r0 in V & V is open;
    then A5: g.r0 in Int V by TOPS_1:55;
    reconsider u=g.r0 as Point of Euclid n by TOPREAL3:13;
    consider s being real number such that
    A6: s>0 & Ball(u,s) c= V by A5,GOBOARD6:8;
      now per cases;
    case p <> 0.REAL n;
    then A7: |.p.| <> 0 by TOPRNS_1:25;
     A8: |.p.| >=0 by TOPRNS_1:26;
    set r2=s/|.p.|;
    reconsider G1=].f.r0-r2,f.r0+r2.[ as Subset of R^1
                             by TOPMETR:24;
      r2>0 by A6,A7,A8,SEQ_2:6;
    then A9: f.r0<f.r0+r2 by REAL_1:69;
    then f.r0-r2<f.r0 by REAL_1:84;
    then A10:f.r0 in G1 by A9,JORDAN6:45;
      G1 is open by JORDAN6:46;
    then consider W2 being Subset of X such that
    A11: r0 in W2 & W2 is open & f.:W2 c= G1 by A1,A10,JGRAPH_2:20;
      g.:W2 c= V
    proof let y be set;assume y in g.:W2;
      then consider r being set such that
      A12: r in dom g & r in W2 & y=g.r by FUNCT_1:def 12;
      reconsider r as Point of X by A12;
        dom f=the carrier of X by FUNCT_2:def 1;
      then f.r in f.:W2 by A12,FUNCT_1:def 12;
      then A13: abs(f.r - f.r0) <r2 by A11,RCOMP_1:8;
      reconsider t=f.r,t0=f.r0 as Real by XREAL_0:def 1;
      A14: abs(t0 - t)=abs(t-t0) by UNIFORM1:13;
      reconsider v=g.r as Point of Euclid n by TOPREAL3:13;
        g.r0=(f.r0)*p by A3;
      then A15: |.g.r0 -g.r.| = |.(f.r0)*p -(f.r)*p.| by A3
      .= |.((f.r0)-(f.r))*p.| by EUCLID:54
      .=abs(t0-t)*|.p.| by TOPRNS_1:8;
        abs(f.r - f.r0)*|.p.| < r2*|.p.| by A7,A8,A13,REAL_1:70;
      then |.g.r0 -g.r .|<s by A7,A14,A15,XCMPLX_1:88;
      then dist(u,v)<s by JGRAPH_1:45;
      then g.r in Ball(u,s) by METRIC_1:12;
     hence y in V by A6,A12;
    end;
   hence ex W being Subset of X st r0 in W & W is open
      & g.:W c= V by A11;
   case A16: p =0.REAL n;
    A17: for r being Point of X holds g.r=0.REAL n
    proof let r be Point of X;
     thus g.r=(f.r)*p by A3 .=0.REAL n by A16,EUCLID:32;
    end;
    then A18: 0.REAL n in V by A4;
    set W2=[#]X;
      r0 in the carrier of X;
    then A19: r0 in W2 by PRE_TOPC:12;
      g.:W2 c= V
    proof let y be set;assume y in g.:W2;
      then consider x being set such that
      A20: x in dom g & x in W2 & y=g.x by FUNCT_1:def 12;
     thus y in V by A17,A18,A20;
    end;
   hence ex W being Subset of X st r0 in W & W is open
      & g.:W c= V by A19;
   end;
   hence ex W being Subset of X st r0 in W & W is open
      & g.:W c= V;
  end;
  then g is continuous by JGRAPH_2:20;
 hence thesis by A3;
end;

theorem Th18:  Sq_Circ.(|[-1,0]|)= |[-1,0]|
proof
  set p= |[-1,0]|;
  A1: p`1=-1 & p`2=0 by EUCLID:56;
  then A2: p<>0.REAL 2 by EUCLID:56,58;
    p`2>=p`1 & p`2<=-p`1 by A1;
  then Sq_Circ.p= |[p`1/sqrt(1+(p`2/p`1)^2),p`2/sqrt(1+(p`2/p`1)^2)]|
                            by A2,JGRAPH_3:def 1;
 hence thesis by A1,SQUARE_1:60,83;
end;

theorem for P being compact non empty Subset of TOP-REAL 2
 st P={p where p is Point of TOP-REAL 2: |.p.|=1}
 holds Sq_Circ.(|[-1,0]|)=W-min(P) by Th18,JGRAPH_5:32;

theorem Th20: for X being non empty TopSpace, n being Nat,
g1,g2 being map of X,TOP-REAL n st
g1 is continuous & g2 is continuous
ex g being map of X,TOP-REAL n st
(for r being Point of X holds g.r=g1.r + g2.r) &
g is continuous
proof let X being non empty TopSpace,n be Nat, g1,g2 be map of X,TOP-REAL n;
assume A1: g1 is continuous & g2 is continuous;
defpred P[set,set] means
(for r1,r2 being Element of TOP-REAL n
 st g1.$1=r1 & g2.$1=r2 holds $2=r1+r2);

A2:for x being Element of X
  ex y being Element of TOP-REAL n
  st P[x,y]
  proof let x be Element of X;
     set rr1=g1.x;
     set rr2=g2.x;
     set r3=rr1+rr2;
      for s1,s2 being Point of TOP-REAL n st g1.x=s1 & g2.x=s2
       holds r3=s1+s2;
   hence thesis;
  end;
    ex f being Function of the carrier of X,the carrier of TOP-REAL n
  st for x being Element of X holds P[x,f.x]
               from FuncExD(A2);
  then consider f being Function of the carrier of X,the carrier of TOP-REAL n
  such that A3:  for x being Element of X holds
  (for r1,r2 being Element of TOP-REAL n
         st g1.x=r1 & g2.x=r2 holds f.x=r1+r2);
  reconsider g0=f as map of X,TOP-REAL n ;
  A4: for r being Point of X holds g0.r=g1.r + g2.r by A3;
    for p being Point of X,V being Subset of TOP-REAL n
   st g0.p in V & V is open holds
   ex W being Subset of X st p in W & W is open & g0.:W c= V
    proof let p be Point of X,V be Subset of TOP-REAL n;
     assume g0.p in V & V is open;
      then A5: g0.p in Int V by TOPS_1:55;
      reconsider r=g0.p as Point of Euclid n by TOPREAL3:13;
      consider r0 being real number such that
      A6: r0>0 & Ball(r,r0) c= V by A5,GOBOARD6:8;
      reconsider r01=g1.p as Point of Euclid n by TOPREAL3:13;
      A7: the carrier of Euclid n=the carrier of TOP-REAL n
                           by TOPREAL3:13;
      then reconsider G1=Ball(r01,r0/2) as Subset of TOP-REAL n
                         ;
      A8: r0/2>0 by A6,SEQ_2:3;
      then A9:g1.p in G1 by GOBOARD6:4;
        TOP-REAL n=TopSpaceMetr(Euclid n) by EUCLID:def 8;
      then G1 is open by TOPMETR:21;
      then consider W1 being Subset of X such that
      A10: p in W1 & W1 is open & g1.:W1 c= G1 by A1,A9,JGRAPH_2:20;
      reconsider r02=g2.p as Point of Euclid n by TOPREAL3:13;
      reconsider G2=Ball(r02,r0/2) as Subset of TOP-REAL n
                          by A7;
      A11:g2.p in G2 by A8,GOBOARD6:4;
        TOP-REAL n=TopSpaceMetr(Euclid n) by EUCLID:def 8;
      then G2 is open by TOPMETR:21;
      then consider W2 being Subset of X such that
      A12: p in W2 & W2 is open & g2.:W2 c= G2 by A1,A11,JGRAPH_2:20;
      set W=W1 /\ W2;
      A13:W is open by A10,A12,TOPS_1:38;
      A14:p in W by A10,A12,XBOOLE_0:def 3;
        g0.:W c= Ball(r,r0)
       proof let x be set;assume x in g0.:W;
         then consider z being set such that
         A15: z in dom g0 & z in W & g0.z=x by FUNCT_1:def 12;
         A16:z in W1 by A15,XBOOLE_0:def 3;
         reconsider pz=z as Point of X by A15;
           dom g1=the carrier of X by FUNCT_2:def 1;
         then A17: g1.pz in g1.:W1 by A16,FUNCT_1:def 12;
         reconsider aa1=g1.pz as Point of TOP-REAL n;
         reconsider bb1=aa1 as Point of Euclid n by TOPREAL3:13;
           dist(r01,bb1)<r0/2 by A10,A17,METRIC_1:12;
         then A18: |.g1.p - g1.pz .|<r0/2 by JGRAPH_1:45;
         A19:z in W2 by A15,XBOOLE_0:def 3;
           dom g2=the carrier of X by FUNCT_2:def 1;
         then A20: g2.pz in g2.:W2 by A19,FUNCT_1:def 12;
         reconsider aa2=g2.pz as Point of TOP-REAL n;
         reconsider bb2=aa2 as Point of Euclid n by TOPREAL3:13;
           dist(r02,bb2)<r0/2 by A12,A20,METRIC_1:12;
         then A21: |.g2.p - g2.pz .|<r0/2 by JGRAPH_1:45;
         A22: aa1+aa2=x by A3,A15;
         reconsider bb0=aa1+aa2 as Point of Euclid n by TOPREAL3:13;
         A23: g0.pz= g1.pz+g2.pz by A3;
           (g1.p +g2.p)-(g1.pz+g2.pz)=g1.p+g2.p-g1.pz-g2.pz by EUCLID:50
         .= g1.p+g2.p+-g1.pz-g2.pz by EUCLID:45
         .= g1.p+g2.p+-g1.pz+-g2.pz by EUCLID:45
         .= g1.p+-g1.pz+g2.p+-g2.pz by EUCLID:30
         .= g1.p+-g1.pz+(g2.p+-g2.pz) by EUCLID:30
         .= g1.p-g1.pz+(g2.p+-g2.pz) by EUCLID:45
         .= g1.p-g1.pz+(g2.p-g2.pz) by EUCLID:45;
         then A24: |.(g1.p +g2.p)-(g1.pz+g2.pz).|
         <= |.g1.p-g1.pz.| + |.g2.p-g2.pz.| by TOPRNS_1:30;
           |.g1.p-g1.pz.| + |.g2.p-g2.pz.| < r0/2 +r0/2
                           by A18,A21,REAL_1:67;
         then |.g1.p-g1.pz.| + |.g2.p-g2.pz.| < r0 by XCMPLX_1:66;
         then |.(g1.p +g2.p)-(g1.pz+g2.pz).|<r0 by A24,AXIOMS:22;
         then |.g0.p - g0.pz .|<r0 by A3,A23;
         then dist(r,bb0)<r0 by A15,A22,JGRAPH_1:45;
        hence x in Ball(r,r0) by A22,METRIC_1:12;
       end;
      then g0.:W c= V by A6,XBOOLE_1:1;
     hence ex W being Subset of X st p in W & W is open & g0.:W c= V
                          by A13,A14;
    end;
   then g0 is continuous by JGRAPH_2:20;
hence thesis by A4;
end;

theorem Th21: for X being non empty TopSpace, n being Nat,
p1,p2 being Point of TOP-REAL n,
f1,f2 being map of X,R^1 st
f1 is continuous & f2 is continuous
ex g being map of X,TOP-REAL n st
(for r being Point of X holds g.r=(f1.r)*p1+(f2.r)*p2) &
g is continuous
proof let X be non empty TopSpace, n be Nat,
p1,p2 be Point of TOP-REAL n,
f1,f2 be map of X,R^1;
 assume A1: f1 is continuous & f2 is continuous;
  then consider g1 being map of X,TOP-REAL n such that
  A2: (for r being Point of X holds g1.r=(f1.r)*p1) &
    g1 is continuous by Th17;
  consider g2 being map of X,TOP-REAL n such that
  A3: (for r being Point of X holds g2.r=(f2.r)*p2) &
    g2 is continuous by A1,Th17;
  consider g being map of X,TOP-REAL n such that
  A4: (for r being Point of X holds g.r=g1.r + g2.r) &
    g is continuous by A2,A3,Th20;
   (for r being Point of X holds g.r=(f1.r)*p1+(f2.r)*p2)
  proof let r be Point of X;
      g.r=g1.r+g2.r by A4;
    then g.r=g1.r+(f2.r)*p2 by A3;
   hence g.r=(f1.r)*p1+(f2.r)*p2 by A2;
  end;
 hence thesis by A4;
end;

theorem Th22:
  for f being Function,A being set st f is one-to-one & A c= dom f
   holds f".:(f.:A)=A
proof let f be Function,A be set;
 set B = f.:A;
 assume A1: f is one-to-one & A c= dom f;
  A2: f".:B c= A
   proof let y be set;assume y in f".:B;
     then consider x being set such that
     A3: x in dom (f") & x in B & y=f".x by FUNCT_1:def 12;
     consider y2 being set such that
     A4: y2 in dom f & y2 in A & x=f.y2 by A3,FUNCT_1:def 12;
    thus y in A by A1,A3,A4,FUNCT_1:54;
   end;
    A c= f".:B
  proof let x be set;assume
    A5: x in A;
    set y0=f.x;
    A6: f".y0=x by A1,A5,FUNCT_1:56;
      y0 in rng f by A1,A5,FUNCT_1:12;
    then A7: y0 in dom (f") by A1,FUNCT_1:55;
      y0 in B by A1,A5,FUNCT_1:def 12;
   hence x in f".:B by A6,A7,FUNCT_1:def 12;
  end;
 hence f".:B=A by A2,XBOOLE_0:def 10;
end;

begin :: General Fashoda Theorem for Unit Circle

Lm5: (|[-1,0]|)`1 =-1 & (|[-1,0]|)`2=0 & (|[1,0]|)`1 =1 & (|[1,0]|)`2=0 &
      (|[0,-1]|)`1 =0 & (|[0,-1]|)`2=-1 & (|[0,1]|)`1 =0 & (|[0,1]|)`2=1
               by EUCLID:56;
Lm6: now thus |.(|[-1,0]|).|=sqrt((-1)^2+0^2) by Lm5,JGRAPH_3:10
       .=1 by SQUARE_1:59,60,61,83;
       thus |.(|[1,0]|).|=sqrt(1^2+0^2) by Lm5,JGRAPH_3:10
                        .=1 by SQUARE_1:59,60,83;
       thus |.(|[0,-1]|).|=sqrt(0^2+(-1)^2) by Lm5,JGRAPH_3:10
                         .=1 by SQUARE_1:59,60,61,83;
       thus |.(|[0,1]|).|=sqrt(0^2+1^2) by Lm5,JGRAPH_3:10
                        .=1 by SQUARE_1:59,60,83;
      end;
Lm7: 0 in [.0,1.] by TOPREAL5:1;
Lm8: 1 in [.0,1.] by TOPREAL5:1;

reserve p,p1,p2,p3,q,q1,q2 for Point of TOP-REAL 2,i,j for Nat,
        r,lambda for Real;

theorem Th23: for f,g being map of I[01],TOP-REAL 2,
  C0,KXP,KXN,KYP,KYN being Subset of TOP-REAL 2,
  O,I being Point of I[01] st O=0 & I=1 &
  f is continuous one-to-one &
  g is continuous one-to-one &
  C0={p: |.p.|<=1}&
  KXP={q1 where q1 is Point of TOP-REAL 2:
           |.q1.|=1 & q1`2<=q1`1 & q1`2>=-q1`1} &
  KXN={q2 where q2 is Point of TOP-REAL 2:
           |.q2.|=1 & q2`2>=q2`1 & q2`2<=-q2`1} &
  KYP={q3 where q3 is Point of TOP-REAL 2:
           |.q3.|=1 & q3`2>=q3`1 & q3`2>=-q3`1} &
  KYN={q4 where q4 is Point of TOP-REAL 2:
           |.q4.|=1 & q4`2<=q4`1 & q4`2<=-q4`1} &
  f.O in KXP & f.I in KXN &
  g.O in KYP & g.I in KYN &
  rng f c= C0 & rng g c= C0
   holds rng f meets rng g
proof let f,g be map of I[01],TOP-REAL 2,
  C0,KXP,KXN,KYP,KYN be Subset of TOP-REAL 2,
  O,I be Point of I[01];
  assume A1: O=0 & I=1 &
  f is continuous & f is one-to-one &
  g is continuous & g is one-to-one &
  C0={p: |.p.|<=1}&
  KXP={q1 where q1 is Point of TOP-REAL 2:
           |.q1.|=1 & q1`2<=q1`1 & q1`2>=-q1`1} &
  KXN={q2 where q2 is Point of TOP-REAL 2:
           |.q2.|=1 & q2`2>=q2`1 & q2`2<=-q2`1} &
  KYP={q3 where q3 is Point of TOP-REAL 2:
           |.q3.|=1 & q3`2>=q3`1 & q3`2>=-q3`1} &
  KYN={q4 where q4 is Point of TOP-REAL 2:
           |.q4.|=1 & q4`2<=q4`1 & q4`2<=-q4`1} &
  f.O in KXP & f.I in KXN &
  g.O in KYP & g.I in KYN &
  rng f c= C0 & rng g c= C0;
  then consider f2 being map of I[01],TOP-REAL 2 such that
  A2: f2.0=f.1 & f2.1=f.0 &
  rng f2=rng f & f2 is continuous & f2 is one-to-one by JGRAPH_5:15;
 thus rng f meets rng g by A1,A2,JGRAPH_5:16;
end;

theorem Th24: for f,g being map of I[01],TOP-REAL 2,
  C0,KXP,KXN,KYP,KYN being Subset of TOP-REAL 2,
  O,I being Point of I[01] st O=0 & I=1 &
  f is continuous & f is one-to-one &
  g is continuous & g is one-to-one &
  C0={p: |.p.|<=1}&
  KXP={q1 where q1 is Point of TOP-REAL 2:
           |.q1.|=1 & q1`2<=q1`1 & q1`2>=-q1`1} &
  KXN={q2 where q2 is Point of TOP-REAL 2:
           |.q2.|=1 & q2`2>=q2`1 & q2`2<=-q2`1} &
  KYP={q3 where q3 is Point of TOP-REAL 2:
           |.q3.|=1 & q3`2>=q3`1 & q3`2>=-q3`1} &
  KYN={q4 where q4 is Point of TOP-REAL 2:
           |.q4.|=1 & q4`2<=q4`1 & q4`2<=-q4`1} &
  f.O in KXP & f.I in KXN &
  g.O in KYN & g.I in KYP &
  rng f c= C0 & rng g c= C0
   holds rng f meets rng g
proof let f,g be map of I[01],TOP-REAL 2,
  C0,KXP,KXN,KYP,KYN be Subset of TOP-REAL 2,
  O,I be Point of I[01];
  assume A1: O=0 & I=1 &
  f is continuous & f is one-to-one &
  g is continuous & g is one-to-one &
  C0={p: |.p.|<=1}&
  KXP={q1 where q1 is Point of TOP-REAL 2:
           |.q1.|=1 & q1`2<=q1`1 & q1`2>=-q1`1} &
  KXN={q2 where q2 is Point of TOP-REAL 2:
           |.q2.|=1 & q2`2>=q2`1 & q2`2<=-q2`1} &
  KYP={q3 where q3 is Point of TOP-REAL 2:
           |.q3.|=1 & q3`2>=q3`1 & q3`2>=-q3`1} &
  KYN={q4 where q4 is Point of TOP-REAL 2:
           |.q4.|=1 & q4`2<=q4`1 & q4`2<=-q4`1} &
  f.O in KXP & f.I in KXN &
  g.O in KYN & g.I in KYP &
  rng f c= C0 & rng g c= C0;
  then consider g2 being map of I[01],TOP-REAL 2 such that
  A2: g2.0=g.1 & g2.1=g.0 &
  rng g2=rng g & g2 is continuous & g2 is one-to-one by JGRAPH_5:15;
 thus rng f meets rng g by A1,A2,Th23;
end;

theorem Th25:
for p1,p2,p3,p4 being Point of TOP-REAL 2,
P being compact non empty Subset of TOP-REAL 2,C0 being Subset of TOP-REAL 2
 st P={p where p is Point of TOP-REAL 2: |.p.|=1}
 & LE p1,p2,P & LE p2,p3,P & LE p3,p4,P
 holds
 (for f,g being map of I[01],TOP-REAL 2 st
  f is continuous one-to-one & g is continuous one-to-one &
  C0={p8 where p8 is Point of TOP-REAL 2: |.p8.|<=1}&
  f.0=p3 & f.1=p1 & g.0=p2 & g.1=p4 &
  rng f c= C0 & rng g c= C0 holds rng f meets rng g)
 proof let p1,p2,p3,p4 be Point of TOP-REAL 2,
   P be compact non empty Subset of TOP-REAL 2,C0 be Subset of TOP-REAL 2;
   assume A1: P={p where p is Point of TOP-REAL 2: |.p.|=1}
     & LE p1,p2,P & LE p2,p3,P & LE p3,p4,P;
  let f,g be map of I[01],TOP-REAL 2;
    assume A2:f is continuous one-to-one & g is continuous one-to-one &
    C0={p8 where p8 is Point of TOP-REAL 2: |.p8.|<=1}&
    f.0=p3 & f.1=p1 & g.0=p2 & g.1=p4 &
    rng f c= C0 & rng g c= C0;
  A3: dom f=the carrier of I[01] by FUNCT_2:def 1;
  A4: dom g=the carrier of I[01] by FUNCT_2:def 1;
  per cases;
  suppose A5: not (p1<>p2 & p2<>p3 & p3<>p4);
     now per cases by A5;
   case A6: p1=p2;
    thus rng f meets rng g
       proof
        A7: p1 in rng f by A2,A3,Lm8,BORSUK_1:83,FUNCT_1:def 5;
          p2 in rng g by A2,A4,Lm7,BORSUK_1:83,FUNCT_1:def 5;
        hence thesis by A6,A7,XBOOLE_0:3;
       end;
   case A8: p2=p3;
    thus rng f meets rng g
       proof
        A9: p3 in rng f by A2,A3,Lm7,BORSUK_1:83,FUNCT_1:def 5;
          p2 in rng g by A2,A4,Lm7,BORSUK_1:83,FUNCT_1:def 5;
        hence thesis by A8,A9,XBOOLE_0:3;
       end;
   case A10: p3=p4;
    thus rng f meets rng g
       proof
        A11: p3 in rng f by A2,A3,Lm7,BORSUK_1:83,FUNCT_1:def 5;
          p4 in rng g by A2,A4,Lm8,BORSUK_1:83,FUNCT_1:def 5;
        hence thesis by A10,A11,XBOOLE_0:3;
       end;
   end;
 hence thesis;
  suppose p1<>p2 & p2<>p3 & p3<>p4;
    then consider h being map of TOP-REAL 2,TOP-REAL 2 such that
    A12: h is_homeomorphism &
      (for q being Point of TOP-REAL 2 holds |.(h.q).|=|.q.|)&
       |[-1,0]|=h.p1 & |[0,1]|=h.p2 & |[1,0]|=h.p3 & |[0,-1]|=h.p4
                                   by A1,JGRAPH_5:70;
    A13: h is one-to-one by A12,TOPS_2:def 5;
    reconsider f2=h*f as map of I[01],TOP-REAL 2;
    reconsider g2=h*g as map of I[01],TOP-REAL 2;
    A14: dom f2=the carrier of I[01] by FUNCT_2:def 1;
    A15: dom g2=the carrier of I[01] by FUNCT_2:def 1;
    A16: f2.1= |[-1,0]| by A2,A12,A14,Lm8,BORSUK_1:83,FUNCT_1:22;
    A17: g2.1= |[0,-1]| by A2,A12,A15,Lm8,BORSUK_1:83,FUNCT_1:22;
    A18: f2.0= |[1,0]| by A2,A12,A14,Lm7,BORSUK_1:83,FUNCT_1:22;
    A19: g2.0= |[0,1]| by A2,A12,A15,Lm7,BORSUK_1:83,FUNCT_1:22;
    A20: f2 is continuous one-to-one &
     g2 is continuous one-to-one &
     f2.1=|[-1,0]| & f2.0=|[1,0]| & g2.1=|[0,-1]| & g2.0= |[0,1]|
                        by A2,A12,A14,A15,Lm7,Lm8,BORSUK_1:83,FUNCT_1:22,
JGRAPH_5:8,9;
    A21: rng f2 c= C0
     proof let y be set;assume y in rng f2;
       then consider x being set such that
       A22: x in dom f2 & y=f2.x by FUNCT_1:def 5;
       A23: f2.x=h.(f.x) by A22,FUNCT_1:22;
       A24: f.x in rng f by A3,A14,A22,FUNCT_1:def 5;
       then A25: f.x in C0 by A2;
       reconsider qf=f.x as Point of TOP-REAL 2 by A24;
       consider q5 being Point of TOP-REAL 2 such that
       A26: q5=f.x & |.q5.|<=1 by A2,A25;
         |.(h.qf).|=|.qf.| by A12;
      hence y in C0 by A2,A22,A23,A26;
     end;
    A27: rng g2 c= C0
     proof let y be set;assume y in rng g2;
       then consider x being set such that
       A28: x in dom g2 & y=g2.x by FUNCT_1:def 5;
       A29: g2.x=h.(g.x) by A28,FUNCT_1:22;
       A30: g.x in rng g by A4,A15,A28,FUNCT_1:def 5;
       then A31: g.x in C0 by A2;
       reconsider qg=g.x as Point of TOP-REAL 2 by A30;
       consider q5 being Point of TOP-REAL 2 such that
       A32: q5=g.x & |.q5.|<=1 by A2,A31;
         |.(h.qg).|=|.qg.| by A12;
      hence y in C0 by A2,A28,A29,A32;
     end;
     defpred Q[Point of TOP-REAL 2] means
        |.$1.|=1 & $1`2<=$1`1 & $1`2>=-$1`1;
      {q1 where q1 is Point of TOP-REAL 2:Q[q1]} is
        Subset of TOP-REAL 2 from TopSubset;
    then reconsider KXP={q1 where q1 is Point of TOP-REAL 2:
        |.q1.|=1 & q1`2<=q1`1 & q1`2>=-q1`1} as Subset of TOP-REAL 2;
        defpred Q[Point of TOP-REAL 2] means
        |.$1.|=1 & $1`2>=$1`1 & $1`2<=-$1`1;
      {q2 where q2 is Point of TOP-REAL 2:Q[q2]} is
         Subset of TOP-REAL 2 from TopSubset;
    then reconsider KXN={q2 where q2 is Point of TOP-REAL 2:
        |.q2.|=1 & q2`2>=q2`1 & q2`2<=-q2`1} as Subset of TOP-REAL 2;
        defpred Q[Point of TOP-REAL 2] means
        |.$1.|=1 & $1`2>=$1`1 & $1`2>=-$1`1;
      {q3 where q3 is Point of TOP-REAL 2:Q[q3]} is
         Subset of TOP-REAL 2 from TopSubset;
    then reconsider KYP={q3 where q3 is Point of TOP-REAL 2:
        |.q3.|=1 & q3`2>=q3`1 & q3`2>=-q3`1} as Subset of TOP-REAL 2;
        defpred Q[Point of TOP-REAL 2] means
        |.$1.|=1 & $1`2<=$1`1 & $1`2<=-$1`1;
      {q4 where q4 is Point of TOP-REAL 2:Q[q4]} is
         Subset of TOP-REAL 2 from TopSubset;
    then reconsider KYN={q4 where q4 is Point of TOP-REAL 2:
        |.q4.|=1 & q4`2<=q4`1 & q4`2<=-q4`1} as Subset of TOP-REAL 2;
    reconsider O=0,I=1 as Point of I[01] by BORSUK_1:83,TOPREAL5:1;
       -(|[-1,0]|)`1=1 by Lm5;
    then A33: f2.I in KXN by A16,Lm5,Lm6;
    A34: f2.O in KXP by A18,Lm5,Lm6;
       -(|[0,-1]|)`1= 0 by Lm5;
    then A35: g2.I in KYN by A17,Lm5,Lm6;
       -(|[0,1]|)`1= 0 by Lm5;
    then g2.O in KYP by A19,Lm5,Lm6;
    then rng f2 meets rng g2 by A2,A20,A21,A27,A33,A34,A35,Th23;
    then consider x2 being set such that
    A36: x2 in rng f2 & x2 in rng g2 by XBOOLE_0:3;
    consider z2 being set such that
    A37: z2 in dom f2 & x2=f2.z2 by A36,FUNCT_1:def 5;
    consider z3 being set such that
    A38: z3 in dom g2 & x2=g2.z3 by A36,FUNCT_1:def 5;
    A39: dom h=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
    A40: g.z3 in rng g by A4,A15,A38,FUNCT_1:def 5;
    A41: f.z2 in rng f by A3,A14,A37,FUNCT_1:def 5;
    reconsider h1=h as Function;
    A42: h1".x2=h1".(h.(f.z2)) by A37,FUNCT_1:22 .=f.z2
                     by A13,A39,A41,FUNCT_1:56;
      h1".x2=h1".(h.(g.z3)) by A38,FUNCT_1:22 .=g.z3
                     by A13,A39,A40,FUNCT_1:56;
    then h1".x2 in rng f & h1".x2 in rng g by A3,A4,A14,A15,A37,A38,A42,FUNCT_1
:def 5;
   hence thesis by XBOOLE_0:3;
 end;

theorem Th26:
for p1,p2,p3,p4 being Point of TOP-REAL 2,
P being compact non empty Subset of TOP-REAL 2,C0 being Subset of TOP-REAL 2
 st P={p where p is Point of TOP-REAL 2: |.p.|=1}
 & LE p1,p2,P & LE p2,p3,P & LE p3,p4,P
 holds
 (for f,g being map of I[01],TOP-REAL 2 st
  f is continuous one-to-one & g is continuous one-to-one &
  C0={p8 where p8 is Point of TOP-REAL 2: |.p8.|<=1}&
  f.0=p3 & f.1=p1 & g.0=p4 & g.1=p2 &
  rng f c= C0 & rng g c= C0 holds rng f meets rng g)
 proof let p1,p2,p3,p4 be Point of TOP-REAL 2,
   P be compact non empty Subset of TOP-REAL 2,C0 be Subset of TOP-REAL 2;
   assume A1: P={p where p is Point of TOP-REAL 2: |.p.|=1}
     & LE p1,p2,P & LE p2,p3,P & LE p3,p4,P;
  let f,g be map of I[01],TOP-REAL 2;
    assume A2:f is continuous one-to-one & g is continuous one-to-one &
    C0={p8 where p8 is Point of TOP-REAL 2: |.p8.|<=1}&
    f.0=p3 & f.1=p1 & g.0=p4 & g.1=p2 &
    rng f c= C0 & rng g c= C0;
  A3: dom f=the carrier of I[01] by FUNCT_2:def 1;
  A4: dom g=the carrier of I[01] by FUNCT_2:def 1;
  per cases;
  suppose A5: not (p1<>p2 & p2<>p3 & p3<>p4);
     now per cases by A5;
   case A6: p1=p2;
    thus rng f meets rng g
       proof
        A7: p1 in rng f by A2,A3,Lm8,BORSUK_1:83,FUNCT_1:def 5;
          p2 in rng g by A2,A4,Lm8,BORSUK_1:83,FUNCT_1:def 5;
        hence thesis by A6,A7,XBOOLE_0:3;
       end;
   case A8: p2=p3;
    thus rng f meets rng g
       proof
        A9: p3 in rng f by A2,A3,Lm7,BORSUK_1:83,FUNCT_1:def 5;
          p2 in rng g by A2,A4,Lm8,BORSUK_1:83,FUNCT_1:def 5;
        hence thesis by A8,A9,XBOOLE_0:3;
       end;
   case A10: p3=p4;
    thus rng f meets rng g
       proof
        A11: p3 in rng f by A2,A3,Lm7,BORSUK_1:83,FUNCT_1:def 5;
          p4 in rng g by A2,A4,Lm7,BORSUK_1:83,FUNCT_1:def 5;
        hence thesis by A10,A11,XBOOLE_0:3;
       end;
   end;
 hence thesis;
  suppose p1<>p2 & p2<>p3 & p3<>p4;
    then consider h being map of TOP-REAL 2,TOP-REAL 2 such that
    A12: h is_homeomorphism &
      (for q being Point of TOP-REAL 2 holds |.(h.q).|=|.q.|)&
       |[-1,0]|=h.p1 & |[0,1]|=h.p2 & |[1,0]|=h.p3 & |[0,-1]|=h.p4
                                   by A1,JGRAPH_5:70;
    A13: h is one-to-one by A12,TOPS_2:def 5;
    reconsider f2=h*f as map of I[01],TOP-REAL 2;
    reconsider g2=h*g as map of I[01],TOP-REAL 2;
    A14: dom f2=the carrier of I[01] by FUNCT_2:def 1;
    A15: dom g2=the carrier of I[01] by FUNCT_2:def 1;
    A16: f2.1= |[-1,0]| by A2,A12,A14,Lm8,BORSUK_1:83,FUNCT_1:22;
    A17: g2.1= |[0,1]| by A2,A12,A15,Lm8,BORSUK_1:83,FUNCT_1:22;
    A18: f2.0= |[1,0]| by A2,A12,A14,Lm7,BORSUK_1:83,FUNCT_1:22;
    A19: g2.0= |[0,-1]| by A2,A12,A15,Lm7,BORSUK_1:83,FUNCT_1:22;
    A20: f2 is continuous one-to-one &
     g2 is continuous one-to-one &
     f2.1=|[-1,0]| & f2.0=|[1,0]| & g2.1=|[0,1]| & g2.0= |[0,-1]|
                        by A2,A12,A14,A15,Lm7,Lm8,BORSUK_1:83,FUNCT_1:22,
JGRAPH_5:8,9;
    A21: rng f2 c= C0
     proof let y be set;assume y in rng f2;
       then consider x being set such that
       A22: x in dom f2 & y=f2.x by FUNCT_1:def 5;
       A23: f2.x=h.(f.x) by A22,FUNCT_1:22;
       A24: f.x in rng f by A3,A14,A22,FUNCT_1:def 5;
       then A25: f.x in C0 by A2;
       reconsider qf=f.x as Point of TOP-REAL 2 by A24;
       consider q5 being Point of TOP-REAL 2 such that
       A26: q5=f.x & |.q5.|<=1 by A2,A25;
         |.(h.qf).|=|.qf.| by A12;
      hence y in C0 by A2,A22,A23,A26;
     end;
    A27: rng g2 c= C0
     proof let y be set;assume y in rng g2;
       then consider x being set such that
       A28: x in dom g2 & y=g2.x by FUNCT_1:def 5;
       A29: g2.x=h.(g.x) by A28,FUNCT_1:22;
       A30: g.x in rng g by A4,A15,A28,FUNCT_1:def 5;
       then A31: g.x in C0 by A2;
       reconsider qg=g.x as Point of TOP-REAL 2 by A30;
       consider q5 being Point of TOP-REAL 2 such that
       A32: q5=g.x & |.q5.|<=1 by A2,A31;
         |.(h.qg).|=|.qg.| by A12;
      hence y in C0 by A2,A28,A29,A32;
     end;
     defpred P[Point of TOP-REAL 2] means
        |.$1.|=1 & $1`2<=$1`1 & $1`2>=-$1`1;
      {q1 where q1 is Point of TOP-REAL 2:P[q1]} is
        Subset of TOP-REAL 2 from TopSubset;
    then reconsider KXP={q1 where q1 is Point of TOP-REAL 2:
        |.q1.|=1 & q1`2<=q1`1 & q1`2>=-q1`1} as Subset of TOP-REAL 2;
        defpred P[Point of TOP-REAL 2] means
        |.$1.|=1 & $1`2>=$1`1 & $1`2<=-$1`1;
      {q2 where q2 is Point of TOP-REAL 2:P[q2]} is
         Subset of TOP-REAL 2 from TopSubset;
    then reconsider KXN={q2 where q2 is Point of TOP-REAL 2:
        |.q2.|=1 & q2`2>=q2`1 & q2`2<=-q2`1} as Subset of TOP-REAL 2;
        defpred P[Point of TOP-REAL 2] means
        |.$1.|=1 & $1`2>=$1`1 & $1`2>=-$1`1;
      {q3 where q3 is Point of TOP-REAL 2:P[q3]} is
         Subset of TOP-REAL 2 from TopSubset;
    then reconsider KYP={q3 where q3 is Point of TOP-REAL 2:
        |.q3.|=1 & q3`2>=q3`1 & q3`2>=-q3`1} as Subset of TOP-REAL 2;
        defpred P[Point of TOP-REAL 2] means
        |.$1.|=1 & $1`2<=$1`1 & $1`2<=-$1`1;
      {q4 where q4 is Point of TOP-REAL 2:P[q4]} is
         Subset of TOP-REAL 2 from TopSubset;
    then reconsider KYN={q4 where q4 is Point of TOP-REAL 2:
        |.q4.|=1 & q4`2<=q4`1 & q4`2<=-q4`1} as Subset of TOP-REAL 2;
    reconsider O=0,I=1 as Point of I[01] by BORSUK_1:83,TOPREAL5:1;
       -(|[-1,0]|)`1=1 by Lm5;
    then A33: f2.I in KXN by A16,Lm5,Lm6;
    A34: f2.O in KXP by A18,Lm5,Lm6;
       -(|[0,-1]|)`1= 0 by Lm5;
    then A35: g2.I in KYP by A17,Lm5,Lm6;
       -(|[0,1]|)`1= 0 by Lm5;
    then g2.O in KYN by A19,Lm5,Lm6;
    then rng f2 meets rng g2 by A2,A20,A21,A27,A33,A34,A35,Th24;
    then consider x2 being set such that
    A36: x2 in rng f2 & x2 in rng g2 by XBOOLE_0:3;
    consider z2 being set such that
    A37: z2 in dom f2 & x2=f2.z2 by A36,FUNCT_1:def 5;
    consider z3 being set such that
    A38: z3 in dom g2 & x2=g2.z3 by A36,FUNCT_1:def 5;
    A39: dom h=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
    A40: g.z3 in rng g by A4,A15,A38,FUNCT_1:def 5;
    A41: f.z2 in rng f by A3,A14,A37,FUNCT_1:def 5;
    reconsider h1=h as Function;
    A42: h1".x2=h1".(h.(f.z2)) by A37,FUNCT_1:22 .=f.z2
                     by A13,A39,A41,FUNCT_1:56;
      h1".x2=h1".(h.(g.z3)) by A38,FUNCT_1:22 .=g.z3
                     by A13,A39,A40,FUNCT_1:56;
    then h1".x2 in rng f & h1".x2 in rng g by A3,A4,A14,A15,A37,A38,A42,FUNCT_1
:def 5;
   hence thesis by XBOOLE_0:3;
 end;

theorem Th27:
for p1,p2,p3,p4 being Point of TOP-REAL 2,
    P being compact non empty Subset of TOP-REAL 2,
    C0 being Subset of TOP-REAL 2 st
  P={p where p is Point of TOP-REAL 2: |.p.|=1} &
   p1,p2,p3,p4 are_in_this_order_on P holds
 (for f,g being map of I[01],TOP-REAL 2 st
  f is continuous one-to-one & g is continuous one-to-one &
  C0={p8 where p8 is Point of TOP-REAL 2: |.p8.|<=1}&
  f.0=p1 & f.1=p3 & g.0=p2 & g.1=p4 &
  rng f c= C0 & rng g c= C0 holds rng f meets rng g)
proof let p1,p2,p3,p4 be Point of TOP-REAL 2,
  P be compact non empty Subset of TOP-REAL 2,C0 be Subset of TOP-REAL 2;
 assume A1: P={p where p is Point of TOP-REAL 2: |.p.|=1}
   & p1,p2,p3,p4 are_in_this_order_on P;
  per cases by A1,JORDAN17:def 1;
  suppose LE p1,p2,P & LE p2,p3,P & LE p3,p4,P;
   hence thesis by A1,JGRAPH_5:71;
  suppose LE p2,p3,P & LE p3,p4,P & LE p4,p1,P;
   hence thesis by A1,JGRAPH_5:72;
  suppose LE p3,p4,P & LE p4,p1,P & LE p1,p2,P;
   hence thesis by A1,Th26;
  suppose LE p4,p1,P & LE p1,p2,P & LE p2,p3,P;
   hence thesis by A1,Th25;
end;

begin :: General Rectangles and Circles

definition let a,b,c,d be real number;
  func rectangle(a,b,c,d) -> Subset of TOP-REAL 2 equals :Def1:
   {p: p`1=a & c <=p`2 & p`2<=d or p`2=d & a<=p`1 & p`1<=b or
      p`1=b & c <=p`2 & p`2<=d or p`2=c & a<=p`1 & p`1<=b};
  coherence
   proof
     defpred P[Point of TOP-REAL 2] means
     $1`1=a & c <=$1`2 & $1`2<=d or $1`2=d & a<=$1`1 & $1`1<=b or
      $1`1=b & c <=$1`2 & $1`2<=d or $1`2=c & a<=$1`1 & $1`1<=b;
     {p: P[p]}
       c= the carrier of TOP-REAL 2 from Fr_Set0;
    hence thesis;
   end;
end;

theorem Th28: for a,b,c,d being real number, p being Point of TOP-REAL 2
 st a<=b & c <=d & p in rectangle(a,b,c,d) holds
  a<=p`1 & p`1<=b & c <=p`2 & p`2<=d
proof let a,b,c,d be real number, p be Point of TOP-REAL 2;
 assume A1: a<=b & c <=d & p in rectangle(a,b,c,d);
  then p in {p2: p2`1=a & c <=p2`2 & p2`2<=d or p2`2=d & a<=p2`1 & p2`1<=b or
      p2`1=b & c <=p2`2 & p2`2<=d or p2`2=c & a<=p2`1 & p2`1<=b} by Def1;
  then consider p2 such that
  A2: p2=p &
    (p2`1=a & c <=p2`2 & p2`2<=d or p2`2=d & a<=p2`1 & p2`1<=b or
      p2`1=b & c <=p2`2 & p2`2<=d or p2`2=c & a<=p2`1 & p2`1<=b);
  per cases by A2;
  suppose p`1=a & c <=p`2 & p`2<=d;
   hence thesis by A1;
  suppose p`2=d & a<=p`1 & p`1<=b;
   hence thesis by A1;
  suppose p`1=b & c <=p`2 & p`2<=d;
   hence thesis by A1;
  suppose p`2=c & a<=p`1 & p`1<=b;
   hence thesis by A1;
end;

 definition let a,b,c,d be real number;
  func inside_of_rectangle(a,b,c,d) -> Subset of TOP-REAL 2 equals :Def2:
   {p: a <p`1 & p`1< b & c <p`2 & p`2< d};
  coherence
   proof
     defpred P[Point of TOP-REAL 2] means
     a <$1`1 & $1`1< b & c <$1`2 & $1`2< d;
     {p: P[p]}
       c= the carrier of TOP-REAL 2 from Fr_Set0;
    hence thesis;
   end;
 end;

 definition let a,b,c,d be real number;
  func closed_inside_of_rectangle(a,b,c,d) -> Subset of TOP-REAL 2 equals
   :Def3:
   {p: a <=p`1 & p`1<= b & c <=p`2 & p`2<= d};
  coherence
   proof
     defpred P[Point of TOP-REAL 2] means
     a <=$1`1 & $1`1<= b & c <=$1`2 & $1`2<= d;
     {p: P[p]}
       c= the carrier of TOP-REAL 2 from Fr_Set0;
    hence thesis;
   end;
 end;

 definition let a,b,c,d be real number;
  func outside_of_rectangle(a,b,c,d) -> Subset of TOP-REAL 2 equals
   :Def4:
   {p: not(a <=p`1 & p`1<= b & c <=p`2 & p`2<= d)};
  coherence
   proof
     defpred P[Point of TOP-REAL 2] means
     not(a <=$1`1 & $1`1<= b & c <=$1`2 & $1`2<= d);
     {p: P[p] }
       c= the carrier of TOP-REAL 2 from Fr_Set0;
    hence thesis;
   end;
 end;

 definition let a,b,c,d be real number;
  func closed_outside_of_rectangle(a,b,c,d) -> Subset of TOP-REAL 2 equals
   :Def5:
   {p: not(a <p`1 & p`1< b & c <p`2 & p`2< d)};
  coherence
   proof
     defpred P[Point of TOP-REAL 2] means
     not(a <$1`1 & $1`1< b & c <$1`2 & $1`2< d);
     {p: P[p] }
       c= the carrier of TOP-REAL 2 from Fr_Set0;
    hence thesis;
   end;
 end;

theorem Th29:
for a,b,r being real number,Kb,Cb being Subset of TOP-REAL 2 st
 r>=0 & Kb={q: |.q.|=1}&
 Cb={p2 where p2 is Point of TOP-REAL 2: |.(p2- |[a,b]|).|=r} holds
 AffineMap(r,a,r,b).:Kb=Cb
proof let a,b,r be real number,Kb,Cb be Subset of TOP-REAL 2;
 assume A1: r>=0 & Kb={q: |.q.|=1}&
  Cb={p2 where p2 is Point of TOP-REAL 2: |.(p2- |[a,b]|).|=r};
  reconsider rr=r as Real by XREAL_0:def 1;
  A2: AffineMap(r,a,r,b).:Kb c= Cb
  proof let y be set;assume y in AffineMap(r,a,r,b).:Kb;
    then consider x being set such that
    A3: x in dom (AffineMap(r,a,r,b)) & x in Kb &
        y=(AffineMap(r,a,r,b)).x by FUNCT_1:def 12;
    consider p being Point of TOP-REAL 2 such that
    A4: x=p & |.p.|=1 by A1,A3;
    A5: (AffineMap(r,a,r,b)).p=|[r*(p`1)+a,r*(p`2)+b]| by JGRAPH_2:def 2;
    then reconsider q=y as Point of TOP-REAL 2 by A3,A4;
    A6: q- |[a,b]|= |[r*(p`1)+a-a,r*(p`2)+b-b]| by A3,A4,A5,EUCLID:66
              .= |[r*(p`1),r*(p`2)+b-b]| by XCMPLX_1:26
              .= |[r*(p`1),r*(p`2)]| by XCMPLX_1:26
              .= r*(|[(p`1),(p`2)]|) by EUCLID:62
              .= r*p by EUCLID:57;
      |.r*p.|=abs(rr)*(|.p.|) by TOPRNS_1:8.=r by A1,A4,ABSVALUE:def 1;
   hence y in Cb by A1,A6;
  end;
    Cb c= AffineMap(r,a,r,b).:Kb
  proof let y be set;assume y in Cb;
    then consider p2 being Point of TOP-REAL 2 such that
    A7: y=p2 & |.(p2- |[a,b]|).|=r by A1;
      now per cases by A1;
    case A8: r>0; then r" >0 by REAL_1:72;
    then A9: 1/r >0 by XCMPLX_1:217;
    set p1=(1/r)*(p2- |[a,b]|);
      |.p1.|=abs(1/rr)*|.(p2- |[a,b]|).| by TOPRNS_1:8
         .=(1/r)*r by A7,A9,ABSVALUE:def 1 .= 1 by A8,XCMPLX_1:88;
    then A10: p1 in Kb by A1;
      p1=|[(1/r)*(p2- |[a,b]|)`1,(1/r)*(p2- |[a,b]|)`2]| by EUCLID:61;
    then A11: p1`1=(1/r)*(p2- |[a,b]|)`1 &
       p1`2=(1/r)*(p2- |[a,b]|)`2 by EUCLID:56;
    then A12: r*p1`1=r*(1/r)*(p2- |[a,b]|)`1 by XCMPLX_1:4
         .=1*(p2- |[a,b]|)`1 by A8,XCMPLX_1:88
         .=p2`1 -(|[a,b]|)`1 by TOPREAL3:8
         .=p2`1 - a by EUCLID:56;
    A13: r*p1`2=r*(1/r)*(p2- |[a,b]|)`2 by A11,XCMPLX_1:4
         .=1*(p2- |[a,b]|)`2 by A8,XCMPLX_1:88
         .=p2`2 -(|[a,b]|)`2 by TOPREAL3:8
         .=p2`2 - b by EUCLID:56;
    A14: (AffineMap(r,a,r,b)).p1= |[r*(p1`1)+a,r*(p1`2)+b]| by JGRAPH_2:def 2
                          .= |[p2`1,p2`2 -b+b]| by A12,A13,XCMPLX_1:27
                          .= |[p2`1,p2`2]| by XCMPLX_1:27
                          .=p2 by EUCLID:57;
      dom (AffineMap(r,a,r,b))=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
   hence y in AffineMap(r,a,r,b).:Kb by A7,A10,A14,FUNCT_1:def 12;
   case A15: r=0;
    set p1= |[1,0]|;
      p1`1=1 & p1`2=0 by EUCLID:56;
    then |.p1.|=sqrt(1^2+0) by JGRAPH_3:10,SQUARE_1:60
   .=1 by SQUARE_1:89;
    then A16: p1 in Kb by A1;
    A17: (AffineMap(r,a,r,b)).p1= |[(0)*(p1`1)+a,(0)*(p1`2)+b]|
                                 by A15,JGRAPH_2:def 2
                          .=p2 by A7,A15,TOPRNS_1:29;
      dom (AffineMap(r,a,r,b))=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
   hence y in AffineMap(r,a,r,b).:Kb by A7,A16,A17,FUNCT_1:def 12;
   end;
   hence y in AffineMap(r,a,r,b).:Kb;

  end;
 hence AffineMap(r,a,r,b).:Kb=Cb by A2,XBOOLE_0:def 10;
end;

theorem Th30: for P,Q being Subset of TOP-REAL 2 st
(ex f being map of (TOP-REAL 2)|P,(TOP-REAL 2)|Q st
f is_homeomorphism) & P is_simple_closed_curve holds
Q is_simple_closed_curve
proof let P,Q be Subset of TOP-REAL 2;
 assume A1: (ex f being map of (TOP-REAL 2)|P,(TOP-REAL 2)|Q st
   f is_homeomorphism) & P is_simple_closed_curve;
   then consider f being map of (TOP-REAL 2)|P,(TOP-REAL 2)|Q such that
   A2: f is_homeomorphism;
   consider g being map of (TOP-REAL 2)|R^2-unit_square,
      (TOP-REAL 2)|P such that
   A3: g is_homeomorphism by A1,TOPREAL2:def 1;
     (|[1,0]|)`1=1 & (|[1,0]|)`2=0 by EUCLID:56;
   then A4: |[1,0]| in R^2-unit_square by TOPREAL1:def 3;
   A5: dom g=[#]((TOP-REAL 2)|R^2-unit_square) &
    rng g=[#]((TOP-REAL 2)|P) by A3,TOPS_2:def 5;
   then dom g= R^2-unit_square by PRE_TOPC:def 10;
   then A6: g.(|[1,0]|) in rng g by A4,FUNCT_1:12;
   then A7: g.(|[1,0]|) in P by A5,PRE_TOPC:def 10;
   reconsider P1=P as non empty Subset of TOP-REAL 2 by A5,A6,PRE_TOPC:def 10;
   A8: dom f=[#]((TOP-REAL 2)|P) &
    rng f=[#]((TOP-REAL 2)|Q) by A2,TOPS_2:def 5;
   then dom f= P by PRE_TOPC:def 10;
   then f.(g.(|[1,0]|)) in rng f by A7,FUNCT_1:12;
   then reconsider Q1=Q as non empty Subset of TOP-REAL 2 by A8,PRE_TOPC:def 10
;
   reconsider f1=f as map of (TOP-REAL 2)|P1,(TOP-REAL 2)|Q1;
   reconsider g1=g as map of (TOP-REAL 2)|R^2-unit_square,(TOP-REAL 2)|P1;
   reconsider h=f1*g1 as map of (TOP-REAL 2)|R^2-unit_square,
   (TOP-REAL 2)|Q1;
     h is_homeomorphism by A2,A3,TOPS_2:71;
 hence Q is_simple_closed_curve by TOPREAL2:def 1;
end;

theorem Th31: for P being Subset of TOP-REAL 2 st
  P is being_simple_closed_curve holds P is compact
  proof
    let P be Subset of TOP-REAL 2;
    assume P is being_simple_closed_curve;
    then reconsider P'=P as being_simple_closed_curve Subset of TOP-REAL 2;
      P' is compact;
    hence thesis;
  end;

theorem Th32:
for a,b,r be real number, Cb being Subset of TOP-REAL 2 st r>0 &
  Cb={p where p is Point of TOP-REAL 2: |.(p- |[a,b]|).|=r}
 holds Cb is_simple_closed_curve
 proof let a,b,r be real number,
 Cb be Subset of TOP-REAL 2;
  assume A1: r>0 & Cb={p where p is Point of TOP-REAL 2:
      |.(p- |[a,b]|).|=r};
   A2:(|[r,0]|)`1=r & (|[r,0]|)`2=0 by EUCLID:56;
     |.(|[r+a,b]| - |[a,b]|).|=|.(|[r+a,0+b]| - |[a,b]|).|
   .=|.(|[r,0]|+|[a,b]|- |[a,b]|).| by EUCLID:60
   .= |.(|[r,0]|+(|[a,b]|- |[a,b]|)).| by EUCLID:49
   .=|.|[r,0]|+(0.REAL 2).| by EUCLID:46
   .=|.|[r,0]|.| by EUCLID:31 .=sqrt(r^2+0) by A2,JGRAPH_3:10,SQUARE_1:60
   .=r by A1,SQUARE_1:89;
    then |[r+a,b]| in Cb by A1;
   then reconsider Cbb=Cb as non empty Subset of TOP-REAL 2;
   set v= |[1,0]|;
     v`1=1 & v`2=0 by EUCLID:56;
   then |.v.|=sqrt(1^2+0) by JGRAPH_3:10,SQUARE_1:60
   .=1 by SQUARE_1:89;
then A3:  |[1,0]| in {q: |.q.|=1};
defpred P[Point of TOP-REAL 2] means |.$1.|=1;
     {q where q is Element of TOP-REAL 2:P[q]}
       is Subset of TOP-REAL 2 from SubsetD;
   then reconsider Kb= {q: |.q.|=1}
      as non empty Subset of TOP-REAL 2 by A3;
    A4:the carrier of (TOP-REAL 2)|Kb=Kb by JORDAN1:1;
   set SC= AffineMap(r,a,r,b);
   A5: SC is one-to-one by A1,JGRAPH_2:54;
   A6:dom SC = the carrier of TOP-REAL 2 by FUNCT_2:def 1;
   A7:dom (SC|Kb)=(dom SC)/\ Kb by FUNCT_1:68
   .=the carrier of ((TOP-REAL 2)|Kb) by A4,A6,XBOOLE_1:28;
   A8:rng (SC|Kb) c= rng (SC) by FUNCT_1:76;
   A9:rng (SC|Kb) c= (SC|Kb).:(the carrier of ((TOP-REAL 2)|Kb))
    proof let u be set;assume u in rng (SC|Kb);
      then consider z being set such that
      A10: z in dom ((SC|Kb)) & u=(SC|Kb).z by FUNCT_1:def 5;
     thus u in (SC|Kb).:(the carrier of ((TOP-REAL 2)|Kb))
                       by A7,A10,FUNCT_1:def 12;
    end;
     (SC|Kb).: (the carrier of ((TOP-REAL 2)|Kb))
        = SC.:Kb by A4,RFUNCT_2:5
        .= Cb by A1,Th29
        .=the carrier of (TOP-REAL 2)|Cbb by JORDAN1:1;
   then SC|Kb is Function of the carrier of (TOP-REAL 2)|Kb,
      the carrier of (TOP-REAL 2)|Cbb by A7,A9,FUNCT_2:4;
   then reconsider f0=SC|Kb
         as map of (TOP-REAL 2)|Kb, (TOP-REAL 2)|Cbb ;
     rng (SC|Kb) c= the carrier of (TOP-REAL 2) by A8,XBOOLE_1:1;
   then f0 is Function of the carrier of (TOP-REAL 2)|Kb,
     the carrier of TOP-REAL 2 by A7,FUNCT_2:4;
   then reconsider f00=f0 as map of (TOP-REAL 2)|Kb,TOP-REAL 2;
   A11:rng f0 = (SC|Kb).:
          (the carrier of ((TOP-REAL 2)|Kb)) by FUNCT_2:45
        .= SC.:Kb by A4,RFUNCT_2:5
        .= Cb by A1,Th29;
   A12:f0 is one-to-one by A5,FUNCT_1:84;
   A13:f00 is continuous by TOPMETR:10;
   A14: Kb is being_simple_closed_curve by JGRAPH_3:36;
   then Kb is compact by Th31;
   then (TOP-REAL 2)|Kb is compact by COMPTS_1:12;
   then ex f1 being map of (TOP-REAL 2)|Kb,(TOP-REAL 2)|Cbb st
   f00=f1 & f1 is_homeomorphism by A11,A12,A13,JGRAPH_1:64;
  hence thesis by A14,Th30;
 end;

 definition let a,b,r be real number;
  assume A1: r>0;
  func circle(a,b,r) -> compact non empty Subset of TOP-REAL 2 equals:Def6:
   {p where p is Point of TOP-REAL 2: |.p- |[a,b]| .|=r};
  coherence
   proof
     defpred P[Point of TOP-REAL 2] means |.$1- |[a,b]| .|=r;
     {p where p is Point of TOP-REAL 2:P[p]}
       c= the carrier of TOP-REAL 2 from Fr_Set0;
   then reconsider Cb={p where p is Point of TOP-REAL 2: |.p- |[a,b]| .|=r}
                as Subset of TOP-REAL 2;
   A2:(|[r,0]|)`1=r & (|[r,0]|)`2=0 by EUCLID:56;
     |.(|[r+a,b]| - |[a,b]|).|=|.(|[r+a,0+b]| - |[a,b]|).|
   .=|.(|[r,0]|+|[a,b]|- |[a,b]|).| by EUCLID:60
   .= |.(|[r,0]|+(|[a,b]|- |[a,b]|)).| by EUCLID:49
   .=|.|[r,0]|+(0.REAL 2).| by EUCLID:46
   .=|.|[r,0]|.| by EUCLID:31 .=sqrt(r^2+0^2) by A2,JGRAPH_3:10
   .=r by A1,SQUARE_1:60,89;
    then A3: |[r+a,b]| in Cb;
      Cb is_simple_closed_curve by A1,Th32;
    hence thesis by A3,Th31;
   end;
 end;

definition let a,b,r be real number;
  func inside_of_circle(a,b,r) -> Subset of TOP-REAL 2 equals :Def7:
   {p where p is Point of TOP-REAL 2: |.p- |[a,b]| .|<r};
  coherence
   proof
     defpred P[Point of TOP-REAL 2] means |.$1- |[a,b]| .|<r;
     {p where p is Point of TOP-REAL 2: P[p]}
       c= the carrier of TOP-REAL 2 from Fr_Set0;
    hence thesis;
   end;
end;

definition let a,b,r be real number;
  func closed_inside_of_circle(a,b,r) -> Subset of TOP-REAL 2 equals :Def8:
   {p where p is Point of TOP-REAL 2: |.p- |[a,b]| .|<=r};
  coherence
   proof
     defpred P[Point of TOP-REAL 2] means |.$1- |[a,b]| .|<=r;
     {p where p is Point of TOP-REAL 2: P[p]}
       c= the carrier of TOP-REAL 2 from Fr_Set0;
    hence thesis;
   end;
end;

definition let a,b,r be real number;
  func outside_of_circle(a,b,r) -> Subset of TOP-REAL 2 equals :Def9:
   {p where p is Point of TOP-REAL 2: |.p- |[a,b]| .|>r};
  coherence
   proof
     defpred P[Point of TOP-REAL 2] means |.$1- |[a,b]| .|>r;
     {p where p is Point of TOP-REAL 2: P[p]}
       c= the carrier of TOP-REAL 2 from Fr_Set0;
    hence thesis;
   end;
end;

definition let a,b,r be real number;
  func closed_outside_of_circle(a,b,r) -> Subset of TOP-REAL 2 equals :Def10:
   {p where p is Point of TOP-REAL 2: |.p- |[a,b]| .|>=r};
  coherence
   proof
     defpred P[Point of TOP-REAL 2] means |.$1- |[a,b]| .|>=r;
     {p where p is Point of TOP-REAL 2: P[p]}
       c= the carrier of TOP-REAL 2 from Fr_Set0;
    hence thesis;
   end;
end;

theorem Th33: for r being real number holds
 inside_of_circle(0,0,r)={p : |.p.|<r} &
 (r>0 implies circle(0,0,r)={p2 : |.p2.|=r}) &
 outside_of_circle(0,0,r)={p3 : |.p3.|>r} &
 closed_inside_of_circle(0,0,r)={q : |.q.|<=r} &
 closed_outside_of_circle(0,0,r)={q2 : |.q2.|>=r}
proof let r be real number;
  defpred P[Point of TOP-REAL 2] means |.$1- |[0,0]| .|<r;
  defpred Q[Point of TOP-REAL 2] means |.$1.|<r;
  deffunc F(set)=$1;
  A1: for p holds P[p] iff Q[p] by EUCLID:58,JORDAN2C:13;
    inside_of_circle(0,0,r)
   = {F(p) where p is Point of TOP-REAL 2: P[p]} by Def7
  .= {F(p2) where p2 is Point of TOP-REAL 2: Q[p2]} from Fraenkel6'(A1);
 hence inside_of_circle(0,0,r)={p : |.p.|<r};
  defpred P[Point of TOP-REAL 2] means |.$1- |[0,0]| .|=r;
  defpred Q[Point of TOP-REAL 2] means |.$1.|=r;
  A2: for p holds P[p] iff Q[p] by EUCLID:58,JORDAN2C:13;
  hereby assume r>0;
  then circle(0,0,r)= {F(p): P[p]} by Def6
  .= {F(p2) where p2 is Point of TOP-REAL 2: Q[p2]} from Fraenkel6'(A2);
  hence circle(0,0,r)={p2 : |.p2.|=r};
  end;
  defpred P[Point of TOP-REAL 2] means |.$1- |[0,0]| .|>r;
  defpred Q[Point of TOP-REAL 2] means |.$1.|>r;
  A3: for p holds P[p] iff Q[p] by EUCLID:58,JORDAN2C:13;
    outside_of_circle(0,0,r)= {F(p): P[p]} by Def9
  .= {F(p2) where p2 is Point of TOP-REAL 2: Q[p2]} from Fraenkel6'(A3);
 hence outside_of_circle(0,0,r)={p3 : |.p3.|>r};
  defpred P[Point of TOP-REAL 2] means |.$1- |[0,0]| .|<=r;
  defpred Q[Point of TOP-REAL 2] means |.$1.|<=r;
  A4: for p holds P[p] iff Q[p] by EUCLID:58,JORDAN2C:13;
    closed_inside_of_circle(0,0,r)= {F(p): P[p]} by Def8
  .= {F(p2) where p2 is Point of TOP-REAL 2: Q[p2]} from Fraenkel6'(A4);
 hence closed_inside_of_circle(0,0,r)={p : |.p.|<=r};
  defpred P[Point of TOP-REAL 2] means |.$1- |[0,0]| .|>=r;
  defpred Q[Point of TOP-REAL 2] means |.$1.|>=r;
  A5: for p holds P[p] iff Q[p] by EUCLID:58,JORDAN2C:13;
    closed_outside_of_circle(0,0,r)= {F(p): P[p]} by Def10
  .= {F(p2) where p2 is Point of TOP-REAL 2: Q[p2]} from Fraenkel6'(A5);
 hence closed_outside_of_circle(0,0,r)={p3 : |.p3.|>=r};
end;

theorem Th34: for Kb,Cb being Subset of TOP-REAL 2 st
 Kb={p: -1 <p`1 & p`1< 1 & -1 <p`2 & p`2< 1}&
 Cb={p2 where p2 is Point of TOP-REAL 2: |.p2.|<1} holds Sq_Circ.:Kb=Cb
 proof let Kb,Cb be Subset of TOP-REAL 2;
  assume
   A1:Kb={p: -1 <p`1 & p`1< 1 & -1 <p`2 & p`2< 1}&
      Cb={p2 where p2 is Point of TOP-REAL 2: |.p2.|<1};
   thus Sq_Circ.:Kb c= Cb
    proof let y be set;assume y in Sq_Circ.:Kb;
      then consider x being set such that
      A2: x in dom Sq_Circ & x in Kb & y=Sq_Circ.x by FUNCT_1:def 12;
      consider q being Point of TOP-REAL 2 such that
      A3:q=x & -1 <q`1 & q`1< 1 & -1 <q`2 & q`2< 1 by A1,A2;
        now per cases;
      case A4: q=0.REAL 2;
        then A5: Sq_Circ.q=q by JGRAPH_3:def 1;
          |.q.|=0 by A4,TOPRNS_1:24;
       hence ex p2 being Point of TOP-REAL 2 st p2=y & |.p2.|<1 by A2,A3,A5;
      case A6:q<>0.REAL 2 &
          (q`2<=q`1 & -q`1<=q`2 or q`2>=q`1 & q`2<=-q`1);
        then A7:Sq_Circ.q=|[q`1/sqrt(1+(q`2/q`1)^2),q`2/sqrt(1+(q`2/q`1)^2)]|
                                 by JGRAPH_3:def 1;
        A8: (|[q`1/sqrt(1+(q`2/q`1)^2),q`2/sqrt(1+(q`2/q`1)^2)]|)`1
         = q`1/sqrt(1+(q`2/q`1)^2) &
         (|[q`1/sqrt(1+(q`2/q`1)^2),q`2/sqrt(1+(q`2/q`1)^2)]|)`2
         = q`2/sqrt(1+(q`2/q`1)^2) by EUCLID:56;
          0<=(q`2/q`1)^2 by SQUARE_1:72;
        then A9:1+(q`2/q`1)^2>0 by Th3;
        A10: now assume A11: q`1=0;
          then q`2=0 by A6;
         hence contradiction by A6,A11,EUCLID:57,58;
        end;
        then A12: (q`1)^2>0 by SQUARE_1:74;
        A13: (q`1)^2+(q`2)^2<>0 by A10,COMPLEX1:2;
            (q`1)^2<1^2 by A3,JGRAPH_2:8;
          then A14: sqrt((q`1)^2)<1 by A12,SQUARE_1:59,83,95;
           A15: |.(|[q`1/sqrt(1+(q`2/q`1)^2),q`2/sqrt(1+(q`2/q`1)^2)]|).|^2
          =((q`1)/sqrt(1+(q`2/(q`1))^2))^2+(q`2/sqrt(1+(q`2/(q`1))^2))^2
                        by A8,JGRAPH_3:10
          .=(q`1)^2/(sqrt(1+(q`2/(q`1))^2))^2+(q`2/sqrt(1+(q`2/(q`1))^2))^2
                        by SQUARE_1:69
          .=(q`1)^2/(sqrt(1+(q`2/(q`1))^2))^2
                  +(q`2)^2/(sqrt(1+(q`2/(q`1))^2))^2 by SQUARE_1:69
          .=(q`1)^2/(1+(q`2/q`1)^2)+(q`2)^2/(sqrt(1+(q`2/q`1)^2))^2
                               by A9,SQUARE_1:def 4
          .=(q`1)^2/(1+(q`2/q`1)^2)+(q`2)^2/(1+(q`2/q`1)^2)
                           by A9,SQUARE_1:def 4
          .=((q`1)^2+(q`2)^2)/(1+(q`2/q`1)^2) by XCMPLX_1:63
          .=((q`1)^2+(q`2)^2)/(1+(q`2)^2/(q`1)^2) by SQUARE_1:69
          .=((q`1)^2+(q`2)^2)/((q`1)^2/(q`1)^2+(q`2)^2/(q`1)^2)
                                by A12,XCMPLX_1:60
          .=((q`1)^2+(q`2)^2)/(((q`1)^2+(q`2)^2)/(q`1)^2)
                                by XCMPLX_1:63
          .=(q`1)^2*(((q`1)^2+(q`2)^2)/((q`1)^2+(q`2)^2))
                                by XCMPLX_1:82
          .=(q`1)^2*1 by A13,XCMPLX_1:60
          .=(q`1)^2;
            0<=|.(|[q`1/sqrt(1+(q`2/q`1)^2),q`2/sqrt(1+(q`2/q`1)^2)]|).|
                               by TOPRNS_1:26;
          then |.(|[q`1/sqrt(1+(q`2/q`1)^2),q`2/sqrt(1+(q`2/q`1)^2)]|).|<1
                               by A14,A15,SQUARE_1:89;
         hence ex p2 being Point of TOP-REAL 2 st p2=y & |.p2.|<1
                                     by A2,A3,A7;
      case A16:q<>0.REAL 2 &
        not(q`2<=q`1 & -q`1<=q`2 or q`2>=q`1 & q`2<=-q`1);
        then A17:Sq_Circ.q=|[q`1/sqrt(1+(q`1/q`2)^2),q`2/sqrt(1+(q`1/q`2)^2)]|
                                 by JGRAPH_3:def 1;
        A18: (|[q`1/sqrt(1+(q`1/q`2)^2),q`2/sqrt(1+(q`1/q`2)^2)]|)`1
         = q`1/sqrt(1+(q`1/q`2)^2) &
         (|[q`1/sqrt(1+(q`1/q`2)^2),q`2/sqrt(1+(q`1/q`2)^2)]|)`2
         = q`2/sqrt(1+(q`1/q`2)^2) by EUCLID:56;

          0<=(q`1/q`2)^2 by SQUARE_1:72;
        then A19:1+(q`1/q`2)^2>0 by Th3;
        A20: q`2 <> 0 by A16,REAL_1:66;
        then A21: (q`2)^2>0 by SQUARE_1:74;
        A22: (q`1)^2+(q`2)^2 <>0 by A20,COMPLEX1:2;
            (q`2)^2<1^2 by A3,JGRAPH_2:8;
          then A23: sqrt((q`2)^2)<1 by A21,SQUARE_1:59,83,95;
           A24: |.(|[q`1/sqrt(1+(q`1/q`2)^2),q`2/sqrt(1+(q`1/q`2)^2)]|).|^2
          =((q`1)/sqrt(1+(q`1/(q`2))^2))^2+(q`2/sqrt(1+(q`1/(q`2))^2))^2
                        by A18,JGRAPH_3:10
          .=(q`1)^2/(sqrt(1+(q`1/(q`2))^2))^2+(q`2/sqrt(1+(q`1/(q`2))^2))^2
                        by SQUARE_1:69
          .=(q`1)^2/(sqrt(1+(q`1/(q`2))^2))^2
                  +(q`2)^2/(sqrt(1+(q`1/(q`2))^2))^2 by SQUARE_1:69
          .=(q`1)^2/(1+(q`1/q`2)^2)+(q`2)^2/(sqrt(1+(q`1/q`2)^2))^2
                               by A19,SQUARE_1:def 4
          .=(q`1)^2/(1+(q`1/q`2)^2)+(q`2)^2/(1+(q`1/q`2)^2)
                           by A19,SQUARE_1:def 4
          .=((q`1)^2+(q`2)^2)/(1+(q`1/q`2)^2) by XCMPLX_1:63
          .=((q`1)^2+(q`2)^2)/(1+(q`1)^2/(q`2)^2) by SQUARE_1:69
          .=((q`1)^2+(q`2)^2)/((q`1)^2/(q`2)^2+(q`2)^2/(q`2)^2)
                                by A21,XCMPLX_1:60
          .=((q`1)^2+(q`2)^2)/(((q`1)^2+(q`2)^2)/(q`2)^2)
                                by XCMPLX_1:63
          .=(q`2)^2*(((q`1)^2+(q`2)^2)/((q`1)^2+(q`2)^2))
                                by XCMPLX_1:82
          .=(q`2)^2*1 by A22,XCMPLX_1:60
          .=(q`2)^2;
            0<=|.(|[q`1/sqrt(1+(q`1/q`2)^2),q`2/sqrt(1+(q`1/q`2)^2)]|).|
                               by TOPRNS_1:26;
          then |.(|[q`1/sqrt(1+(q`1/q`2)^2),q`2/sqrt(1+(q`1/q`2)^2)]|).|<1
                               by A23,A24,SQUARE_1:89;
         hence ex p2 being Point of TOP-REAL 2 st p2=y & |.p2.|<1
                                     by A2,A3,A17;
      end;
     hence y in Cb by A1;
    end;
   let y be set;assume y in Cb;
     then consider p2 being Point of TOP-REAL 2 such that
     A25: p2=y & |.p2.|<1 by A1;
     set q=p2;
       now per cases;
     case A26: q=0.REAL 2;
        then q`1=0 & q`2=0 by EUCLID:56,58;
        then A27: y in Kb by A1,A25;
        A28: Sq_Circ".y=y by A25,A26,JGRAPH_3:38;
        A29: dom Sq_Circ=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
          y=Sq_Circ.y by A25,A28,FUNCT_1:57,JGRAPH_3:54;
      hence ex x being set st x in dom Sq_Circ & x in Kb & y=Sq_Circ.x
                              by A27,A29;
     case A30:q<>0.REAL 2 &
          (q`2<=q`1 & -q`1<=q`2 or q`2>=q`1 & q`2<=-q`1);
          set px=|[q`1*sqrt(1+(q`2/q`1)^2),q`2*sqrt(1+(q`2/q`1)^2)]|;
        A31: px`1 = q`1*sqrt(1+(q`2/q`1)^2) &
         px`2 = q`2*sqrt(1+(q`2/q`1)^2) by EUCLID:56;
          0<=(q`2/q`1)^2 by SQUARE_1:72;
        then 1+(q`2/q`1)^2>0 by Th3;
        then A32:sqrt(1+(q`2/q`1)^2)>0 by SQUARE_1:93;
          0<=(px`2/px`1)^2 by SQUARE_1:72;
        then A33:1+(px`2/px`1)^2>0 by Th3;
        A34:px`2/px`1=q`2/q`1 by A31,A32,XCMPLX_1:92;
        A35: q`1=q`1*sqrt(1+(q`2/q`1)^2)/(sqrt(1+(q`2/q`1)^2))by A32,XCMPLX_1:
90
          .=px`1/(sqrt(1+(q`2/q`1)^2)) by EUCLID:56;
        A36: q`2=q`2*sqrt(1+(q`2/q`1)^2)/(sqrt(1+(q`2/q`1)^2))by A32,XCMPLX_1:
90
          .=px`2/(sqrt(1+(q`2/q`1)^2)) by EUCLID:56;
        A37: |.q.|^2=q`1^2+q`2^2 by JGRAPH_3:10;
          |.q.| >=0 by TOPRNS_1:26;
        then A38: |.q.|^2<1 by A25,SQUARE_1:59,78;
        A39:now assume px`1=0 & px`2=0;
          then A40:q`1*sqrt(1+(q`2/q`1)^2)=0 & q`2*sqrt(1+(q`2/q`1)^2)=0
                                    by EUCLID:56;
          then A41:q`1=0 by A32,XCMPLX_1:6;
            q`2=0 by A32,A40,XCMPLX_1:6;
         hence contradiction by A30,A41,EUCLID:57,58;
        end;
          q`2<=q`1 & -q`1<=q`2 or q`2>=q`1 &
        q`2*sqrt(1+(q`2/q`1)^2) <= (-q`1)*sqrt(1+(q`2/q`1)^2)
                             by A30,A32,AXIOMS:25;
        then q`2<=q`1 & (-q`1)*sqrt(1+(q`2/q`1)^2) <= q`2*sqrt(1+(q`2/q`1)^2)
           or px`2>=px`1 & px`2<=-px`1 by A31,A32,AXIOMS:25,XCMPLX_1:175;
        then A42:px`2<=px`1 & -px`1<=px`2 or px`2>=px`1 & px`2<=-px`1
                                by A31,A32,AXIOMS:25,XCMPLX_1:175;
    then A43:Sq_Circ.px=|[px`1/sqrt(1+(px`2/px`1)^2),px`2/sqrt(1+(px`2/px`1)^2)
]|
                                 by A39,JGRAPH_2:11,JGRAPH_3:def 1;
          px`2<=px`1 & --px`1>=-px`2 or px`2>=px`1 & px`2<=-px`1
                                by A42,REAL_1:50;
        then A44:px`2<=px`1 & px`1>=-px`2 or px`2>=px`1 & -px`2>=--px`1
                                by REAL_1:50;
        A45:px`1/sqrt(1+(px`2/px`1)^2)=q`1 by A31,A32,A34,XCMPLX_1:90;
         px`2/sqrt(1+(px`2/px`1)^2)=q`2 by A31,A32,A34,XCMPLX_1:90;
        then A46:q=Sq_Circ.px by A43,A45,EUCLID:57;
        A47: dom Sq_Circ=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
          not px`1=0 by A39,A42;
        then A48: (px`1)^2>0 by SQUARE_1:74;
        A49: (px`2)^2>=0 by SQUARE_1:72;
           (px`1)^2/(sqrt(1+(px`2/px`1)^2))^2+(px`2/sqrt(1+(px`2/px`1)^2))^2<1
                            by A34,A35,A36,A37,A38,SQUARE_1:69;
      then (px`1)^2/(sqrt(1+(px`2/px`1)^2))^2+(px`2)^2/(sqrt(1+(px`2/px`1)^2))
^2<1
                            by SQUARE_1:69;
         then (px`1)^2/(1+(px`2/px`1)^2)+(px`2)^2/(sqrt(1+(px`2/px`1)^2))^2<1
                            by A33,SQUARE_1:def 4;
         then (px`1)^2/(1+(px`2/px`1)^2)+(px`2)^2/(1+(px`2/px`1)^2)<1
                            by A33,SQUARE_1:def 4;
         then ((px`1)^2/(1+(px`2/px`1)^2)
          +(px`2)^2/(1+(px`2/px`1)^2))*(1+(px`2/px`1)^2)<1 *(1+(px`2/px`1)^2)
                            by A33,REAL_1:70;
         then (px`1)^2/(1+(px`2/px`1)^2)*(1+(px`2/px`1)^2)
           +(px`2)^2/(1+(px`2/px`1)^2)*(1+(px`2/px`1)^2)<1 *(1+(px`2/px`1)^2)
                            by XCMPLX_1:8;
    then (px`1)^2+(px`2)^2/(1+(px`2/px`1)^2)*(1+(px`2/px`1)^2)<1 *(1+(px`2/px
`1)^2)
                            by A33,XCMPLX_1:88;
         then A50: (px`1)^2+(px`2)^2<1 *(1+(px`2/px`1)^2) by A33,XCMPLX_1:88;
           1 *(1+(px`2/px`1)^2)
         =1+(px`2)^2/(px`1)^2 by SQUARE_1:69;
         then (px`1)^2+(px`2)^2-1<1+(px`2)^2/(px`1)^2-1 by A50,REAL_1:54;
         then (px`1)^2+(px`2)^2-1<(px`2)^2/(px`1)^2 by XCMPLX_1:26;

         then ((px`1)^2+(px`2)^2-1)*(px`1)^2<((px`2)^2/(px`1)^2)*(px`1)^2
                            by A48,REAL_1:70;
         then ((px`1)^2+(px`2)^2-1)*(px`1)^2<(px`2)^2 by A48,XCMPLX_1:88;
         then ((px`1)^2+((px`2)^2-1))*(px`1)^2<(px`2)^2 by XCMPLX_1:29;
         then (px`1)^2*(px`1)^2+((px`2)^2-1)*(px`1)^2<(px`2)^2 by XCMPLX_1:8;
         then (px`1)^2*(px`1)^2+(px`1)^2*((px`2)^2-1)-(px`2)^2<0
                             by REAL_2:105;
         then A51: 0>(px`1)^2*(px`1)^2+((px`1)^2*(px`2)^2-(px`1)^2*1)-(px`2)^2
                                   by XCMPLX_1:40;
            (px`1)^2*(px`1)^2+((px`1)^2*(px`2)^2-(px`1)^2*1)-(px`2)^2
         = (px`1)^2*(px`1)^2+(px`1)^2*(px`2)^2-(px`1)^2*1-(px`2)^2
                                   by XCMPLX_1:29
         .= (px`1)^2*(px`1)^2-(px`1)^2*1+(px`1)^2*(px`2)^2-1 *(px`2)^2
                                   by XCMPLX_1:29
         .= (px`1)^2*((px`1)^2-1)+(px`1)^2*(px`2)^2-1 *(px`2)^2 by XCMPLX_1:40
         .= (px`1)^2*((px`1)^2-1)+((px`1)^2*(px`2)^2-1 *(px`2)^2) by XCMPLX_1:
29
         .= (px`1)^2*((px`1)^2-1)+((px`1)^2-1)*(px`2)^2 by XCMPLX_1:40
         .= ((px`1)^2-1)*((px`1)^2+(px`2)^2) by XCMPLX_1:8;
         then ((px`1)^2-1<0 or ((px`1)^2+(px`2)^2)<0)
         &((px`1)^2-1>0 or (px`1)^2+(px`2)^2>0) by A51,REAL_2:121;

         then (px`1)^2-1+1<0+1 & ((px`1)^2+(px`2)^2)>0 by A48,A49,Th3,REAL_1:53
;
         then (px`1)^2<1 by XCMPLX_1:27;
          then A52:px`1<1 & px`1>-1 by JGRAPH_2:6,SQUARE_1:59;
          then px`2<1 & 1>-px`2 or px`2>=px`1 & -px`2>=px`1
                                by A44,AXIOMS:22;
          then px`2<1 & -1< --px`2 or px`2>-1 & -px`2> -1
                                by A52,AXIOMS:22,REAL_1:50;
          then px`2<1 & -1<px`2 or px`2>-1 & --px`2< --1
                                by REAL_1:50;
        then px in Kb by A1,A52;
      hence ex x being set st x in dom Sq_Circ & x in Kb & y=Sq_Circ.x
                               by A25,A46,A47;
     case A53:q<>0.REAL 2 & not(q`2<=q`1 & -q`1<=q`2 or q`2>=q`1 & q`2<=-q`1);
          set px=|[q`1*sqrt(1+(q`1/q`2)^2),q`2*sqrt(1+(q`1/q`2)^2)]|;
        A54:q<>0.REAL 2 & (q`1<=q`2 & -q`2<=q`1 or q`1>=q`2 & q`1<=-q`2)
                          by A53,JGRAPH_2:23;
        A55: px`2 = q`2*sqrt(1+(q`1/q`2)^2) &
         px`1 = q`1*sqrt(1+(q`1/q`2)^2) by EUCLID:56;
          0<=(q`1/q`2)^2 by SQUARE_1:72;
        then 1+(q`1/q`2)^2>0 by Th3;
        then A56:sqrt(1+(q`1/q`2)^2)>0 by SQUARE_1:93;
          0<=(px`1/px`2)^2 by SQUARE_1:72;
        then A57:1+(px`1/px`2)^2>0 by Th3;
        A58:px`1/px`2=q`1/q`2 by A55,A56,XCMPLX_1:92;
        A59: q`2=q`2*sqrt(1+(q`1/q`2)^2)/(sqrt(1+(q`1/q`2)^2))by A56,XCMPLX_1:
90
          .=px`2/(sqrt(1+(q`1/q`2)^2)) by EUCLID:56;
        A60: q`1=q`1*sqrt(1+(q`1/q`2)^2)/(sqrt(1+(q`1/q`2)^2))by A56,XCMPLX_1:
90
          .=px`1/(sqrt(1+(q`1/q`2)^2)) by EUCLID:56;
        A61: |.q.|^2=q`2^2+q`1^2 by JGRAPH_3:10;
          |.q.| >=0 by TOPRNS_1:26;
        then A62: |.q.|^2<1 by A25,SQUARE_1:59,78;
        A63:now assume px`2=0 & px`1=0;
          then A64:q`2*sqrt(1+(q`1/q`2)^2)=0 & q`1*sqrt(1+(q`1/q`2)^2)=0
                                    by EUCLID:56;
          then q`1=0 by A56,XCMPLX_1:6;
         hence contradiction by A53,A56,A64,XCMPLX_1:6;
        end;
          q`1<=q`2 & -q`2<=q`1 or q`1>=q`2 &
        q`1*sqrt(1+(q`1/q`2)^2) <= (-q`2)*sqrt(1+(q`1/q`2)^2)
                             by A54,A56,AXIOMS:25;
        then q`1<=q`2 & (-q`2)*sqrt(1+(q`1/q`2)^2) <= q`1*sqrt(1+(q`1/q`2)^2)
           or px`1>=px`2 & px`1<=-px`2 by A55,A56,AXIOMS:25,XCMPLX_1:175;
        then A65:px`1<=px`2 & -px`2<=px`1 or px`1>=px`2 & px`1<=-px`2
                                by A55,A56,AXIOMS:25,XCMPLX_1:175;
    then A66:Sq_Circ.px=|[px`1/sqrt(1+(px`1/px`2)^2),px`2/sqrt(1+(px`1/px`2)^2)
]|
                                 by A63,JGRAPH_2:11,JGRAPH_3:14;
          px`1<=px`2 & --px`2>=-px`1 or px`1>=px`2 & px`1<=-px`2
                                by A65,REAL_1:50;
        then A67:px`1<=px`2 & px`2>=-px`1 or px`1>=px`2 & -px`1>=--px`2
                                by REAL_1:50;
        A68:px`2/sqrt(1+(px`1/px`2)^2)=q`2 by A55,A56,A58,XCMPLX_1:90;
         px`1/sqrt(1+(px`1/px`2)^2)=q`1 by A55,A56,A58,XCMPLX_1:90;
        then A69:q=Sq_Circ.px by A66,A68,EUCLID:57;
        A70: dom Sq_Circ=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
          not px`2=0 by A63,A65;
        then A71: (px`2)^2>0 by SQUARE_1:74;
        A72: (px`1)^2>=0 by SQUARE_1:72;
           (px`2)^2/(sqrt(1+(px`1/px`2)^2))^2+(px`1/sqrt(1+(px`1/px`2)^2))^2<1
                            by A58,A59,A60,A61,A62,SQUARE_1:69;
      then (px`2)^2/(sqrt(1+(px`1/px`2)^2))^2+(px`1)^2/(sqrt(1+(px`1/px`2)^2))
^2<1
                            by SQUARE_1:69;
         then (px`2)^2/(1+(px`1/px`2)^2)+(px`1)^2/(sqrt(1+(px`1/px`2)^2))^2<1
                            by A57,SQUARE_1:def 4;
         then (px`2)^2/(1+(px`1/px`2)^2)+(px`1)^2/(1+(px`1/px`2)^2)<1
                            by A57,SQUARE_1:def 4;
         then ((px`2)^2/(1+(px`1/px`2)^2)
          +(px`1)^2/(1+(px`1/px`2)^2))*(1+(px`1/px`2)^2)<1 *(1+(px`1/px`2)^2)
                            by A57,REAL_1:70;
         then (px`2)^2/(1+(px`1/px`2)^2)*(1+(px`1/px`2)^2)
           +(px`1)^2/(1+(px`1/px`2)^2)*(1+(px`1/px`2)^2)<1 *(1+(px`1/px`2)^2)
                            by XCMPLX_1:8;
    then (px`2)^2+(px`1)^2/(1+(px`1/px`2)^2)*(1+(px`1/px`2)^2)<1 *(1+(px`1/px
`2)^2)
                            by A57,XCMPLX_1:88;
         then A73: (px`2)^2+(px`1)^2<1 *(1+(px`1/px`2)^2) by A57,XCMPLX_1:88;
           1 *(1+(px`1/px`2)^2)
         =1+(px`1)^2/(px`2)^2 by SQUARE_1:69;
         then (px`2)^2+(px`1)^2-1<1+(px`1)^2/(px`2)^2-1 by A73,REAL_1:54;
         then (px`2)^2+(px`1)^2-1<(px`1)^2/(px`2)^2 by XCMPLX_1:26;

         then ((px`2)^2+(px`1)^2-1)*(px`2)^2<((px`1)^2/(px`2)^2)*(px`2)^2
                            by A71,REAL_1:70;
         then ((px`2)^2+(px`1)^2-1)*(px`2)^2<(px`1)^2 by A71,XCMPLX_1:88;
         then ((px`2)^2+((px`1)^2-1))*(px`2)^2<(px`1)^2 by XCMPLX_1:29;
         then (px`2)^2*(px`2)^2+((px`1)^2-1)*(px`2)^2<(px`1)^2 by XCMPLX_1:8;
         then (px`2)^2*(px`2)^2+(px`2)^2*((px`1)^2-1)-(px`1)^2<0
                             by REAL_2:105;
         then A74: 0>(px`2)^2*(px`2)^2+((px`2)^2*(px`1)^2-(px`2)^2*1)-(px`1)^2
                                   by XCMPLX_1:40;
            (px`2)^2*(px`2)^2+((px`2)^2*(px`1)^2-(px`2)^2*1)-(px`1)^2
         = (px`2)^2*(px`2)^2+(px`2)^2*(px`1)^2-(px`2)^2*1-(px`1)^2
                                   by XCMPLX_1:29
         .= (px`2)^2*(px`2)^2-(px`2)^2*1+(px`2)^2*(px`1)^2-1 *(px`1)^2
                                   by XCMPLX_1:29
         .= (px`2)^2*((px`2)^2-1)+(px`2)^2*(px`1)^2-1 *(px`1)^2 by XCMPLX_1:40
         .= (px`2)^2*((px`2)^2-1)+((px`2)^2*(px`1)^2-1 *(px`1)^2) by XCMPLX_1:
29
         .= (px`2)^2*((px`2)^2-1)+((px`2)^2-1)*(px`1)^2 by XCMPLX_1:40
         .= ((px`2)^2-1)*((px`2)^2+(px`1)^2) by XCMPLX_1:8;
         then ((px`2)^2-1<0 or ((px`2)^2+(px`1)^2)<0)
         &((px`2)^2-1>0 or (px`2)^2+(px`1)^2>0) by A74,REAL_2:121;

         then (px`2)^2-1+1<0+1 & ((px`2)^2+(px`1)^2)>0 by A71,A72,Th3,REAL_1:53
;
         then (px`2)^2<1 by XCMPLX_1:27;
          then A75:px`2<1 & px`2>-1 by JGRAPH_2:6,SQUARE_1:59;
          then px`1<1 & 1>-px`1 or px`1>=px`2 & -px`1>=px`2
                                by A67,AXIOMS:22;
          then px`1<1 & -1< --px`1 or px`1>-1 & -px`1> -1
                                by A75,AXIOMS:22,REAL_1:50;
          then px`1<1 & -1<px`1 or px`1>-1 & --px`1< --1
                                by REAL_1:50;
        then px in Kb by A1,A75;
      hence ex x being set st x in dom Sq_Circ & x in Kb & y=Sq_Circ.x
                               by A25,A69,A70;
     end;
    hence thesis by FUNCT_1:def 12;
 end;

theorem Th35: for Kb,Cb being Subset of TOP-REAL 2 st
 Kb={p: not(-1 <=p`1 & p`1<= 1 & -1 <=p`2 & p`2<= 1)}&
 Cb={p2 where p2 is Point of TOP-REAL 2: |.p2.|>1} holds Sq_Circ.:Kb=Cb
 proof let Kb,Cb be Subset of TOP-REAL 2;
  assume
A1:Kb={p: not(-1 <=p`1 & p`1<= 1 & -1 <=p`2 & p`2<= 1)}&
   Cb={p2 where p2 is Point of TOP-REAL 2: |.p2.|>1};
   thus Sq_Circ.:Kb c= Cb
    proof let y be set;assume y in Sq_Circ.:Kb;
      then consider x being set such that
      A2: x in dom Sq_Circ & x in Kb & y=Sq_Circ.x
                   by FUNCT_1:def 12;
      consider q being Point of TOP-REAL 2 such that
      A3:q=x &( not(-1 <=q`1 & q`1<= 1 & -1 <=q`2 & q`2<= 1) )
                        by A1,A2;
        now per cases;
      case q=0.REAL 2;
        then q`1=0 & q`2=0 by EUCLID:56,58;
       hence contradiction by A3;
      case A4:q<>0.REAL 2 &
          (q`2<=q`1 & -q`1<=q`2 or q`2>=q`1 & q`2<=-q`1);
        then A5:Sq_Circ.q=|[q`1/sqrt(1+(q`2/q`1)^2),q`2/sqrt(1+(q`2/q`1)^2)]|
                                 by JGRAPH_3:def 1;
        A6: not(-1 <=q`2 & q`2<= 1) implies -1>q`1 or q`1>1
        proof assume A7: not(-1 <=q`2 & q`2<= 1);
            now per cases by A7;
          case A8: -1>q`2;
            then -q`1< -1 or q`2>=q`1 & q`2<= -q`1 by A4,AXIOMS:22;
           hence -1>q`1 or q`1>1 by A8,AXIOMS:22,REAL_1:50;
          case q`2>1;
            then 1<q`1 or 1< -q`1 by A4,AXIOMS:22;
            then 1<q`1 or --q`1< -1 by REAL_1:50;
           hence -1>q`1 or q`1>1;
          end;
         hence -1>q`1 or q`1>1;
        end;
        A9: (|[q`1/sqrt(1+(q`2/q`1)^2),q`2/sqrt(1+(q`2/q`1)^2)]|)`1
         = q`1/sqrt(1+(q`2/q`1)^2) &
         (|[q`1/sqrt(1+(q`2/q`1)^2),q`2/sqrt(1+(q`2/q`1)^2)]|)`2
         = q`2/sqrt(1+(q`2/q`1)^2) by EUCLID:56;
          0<=(q`2/q`1)^2 by SQUARE_1:72;
        then A10:1+(q`2/q`1)^2>0 by Th3;
        A11: now assume A12: q`1=0;
          then q`2=0 by A4;
         hence contradiction by A4,A12,EUCLID:57,58;
        end;
        then A13: (q`1)^2>0 by SQUARE_1:74;
        A14: (q`1)^2+(q`2)^2 <>0 by A11,COMPLEX1:2;
           (q`1)^2>1^2 by A3,A6,JGRAPH_2:5;
          then A15: sqrt((q`1)^2)>1 by SQUARE_1:59,83,95;
           A16: |.(|[q`1/sqrt(1+(q`2/q`1)^2),q`2/sqrt(1+(q`2/q`1)^2)]|).|^2
          =((q`1)/sqrt(1+(q`2/(q`1))^2))^2+(q`2/sqrt(1+(q`2/(q`1))^2))^2
                        by A9,JGRAPH_3:10
          .=(q`1)^2/(sqrt(1+(q`2/(q`1))^2))^2+(q`2/sqrt(1+(q`2/(q`1))^2))^2
                        by SQUARE_1:69
          .=(q`1)^2/(sqrt(1+(q`2/(q`1))^2))^2
                  +(q`2)^2/(sqrt(1+(q`2/(q`1))^2))^2 by SQUARE_1:69
          .=(q`1)^2/(1+(q`2/q`1)^2)+(q`2)^2/(sqrt(1+(q`2/q`1)^2))^2
                               by A10,SQUARE_1:def 4
          .=(q`1)^2/(1+(q`2/q`1)^2)+(q`2)^2/(1+(q`2/q`1)^2)
                           by A10,SQUARE_1:def 4
          .=((q`1)^2+(q`2)^2)/(1+(q`2/q`1)^2) by XCMPLX_1:63
          .=((q`1)^2+(q`2)^2)/(1+(q`2)^2/(q`1)^2) by SQUARE_1:69
          .=((q`1)^2+(q`2)^2)/((q`1)^2/(q`1)^2+(q`2)^2/(q`1)^2)
                                by A13,XCMPLX_1:60
          .=((q`1)^2+(q`2)^2)/(((q`1)^2+(q`2)^2)/(q`1)^2) by XCMPLX_1:63
          .=(q`1)^2*(((q`1)^2+(q`2)^2)/((q`1)^2+(q`2)^2)) by XCMPLX_1:82
          .=(q`1)^2*1 by A14,XCMPLX_1:60
          .=(q`1)^2;
            0<=|.(|[q`1/sqrt(1+(q`2/q`1)^2),q`2/sqrt(1+(q`2/q`1)^2)]|).|
                               by TOPRNS_1:26;
          then |.(|[q`1/sqrt(1+(q`2/q`1)^2),q`2/sqrt(1+(q`2/q`1)^2)]|).|>1
                               by A15,A16,SQUARE_1:89;
         hence ex p2 being Point of TOP-REAL 2 st p2=y & |.p2.|>1
                                     by A2,A3,A5;
      case A17:q<>0.REAL 2 &
        not(q`2<=q`1 & -q`1<=q`2 or q`2>=q`1 & q`2<=-q`1);
        then A18:Sq_Circ.q=|[q`1/sqrt(1+(q`1/q`2)^2),q`2/sqrt(1+(q`1/q`2)^2)]|
                                 by JGRAPH_3:def 1;
        A19: (|[q`1/sqrt(1+(q`1/q`2)^2),q`2/sqrt(1+(q`1/q`2)^2)]|)`1
         = q`1/sqrt(1+(q`1/q`2)^2) &
         (|[q`1/sqrt(1+(q`1/q`2)^2),q`2/sqrt(1+(q`1/q`2)^2)]|)`2
         = q`2/sqrt(1+(q`1/q`2)^2) by EUCLID:56;
          0<=(q`1/q`2)^2 by SQUARE_1:72;
        then A20:1+(q`1/q`2)^2>0 by Th3;
        A21: q`2 <> 0 by A17,REAL_1:66;
        then A22: (q`2)^2>0 by SQUARE_1:74;
        A23: (q`1)^2+(q`2)^2 <>0 by A21,COMPLEX1:2;
          not(-1 <=q`1 & q`1<= 1) implies -1>q`2 or q`2>1
        proof assume A24: not(-1 <=q`1 & q`1<= 1);
            now per cases by A24;
          case A25: -1>q`1;
           then q`2< -1 or q`1<q`2 & -q`2< --q`1 by A17,AXIOMS:22,REAL_1:50;
            then -q`2< -1 or -1>q`2 by A25,AXIOMS:22;
           hence -1>q`2 or q`2>1 by REAL_1:50;
          case A26: q`1>1;
             --q`1< -q`2 & q`2<q`1 or q`2>q`1 & q`2> -q`1 by A17,REAL_1:50;
            then 1< -q`2 or q`2>q`1 & q`2> -q`1 by A26,AXIOMS:22;
            then -1> --q`2 or 1<q`2 by A26,AXIOMS:22,REAL_1:50;
           hence -1>q`2 or q`2>1;
          end;
         hence -1>q`2 or q`2>1;
        end;
          then (q`2)^2>1^2 by A3,JGRAPH_2:5;
          then A27: sqrt((q`2)^2)>1 by SQUARE_1:59,83,95;
           A28: |.(|[q`1/sqrt(1+(q`1/q`2)^2),q`2/sqrt(1+(q`1/q`2)^2)]|).|^2
          =((q`1)/sqrt(1+(q`1/(q`2))^2))^2+(q`2/sqrt(1+(q`1/(q`2))^2))^2
                        by A19,JGRAPH_3:10
          .=(q`1)^2/(sqrt(1+(q`1/(q`2))^2))^2+(q`2/sqrt(1+(q`1/(q`2))^2))^2
                        by SQUARE_1:69
          .=(q`1)^2/(sqrt(1+(q`1/(q`2))^2))^2
                  +(q`2)^2/(sqrt(1+(q`1/(q`2))^2))^2 by SQUARE_1:69
          .=(q`1)^2/(1+(q`1/q`2)^2)+(q`2)^2/(sqrt(1+(q`1/q`2)^2))^2
                               by A20,SQUARE_1:def 4
          .=(q`1)^2/(1+(q`1/q`2)^2)+(q`2)^2/(1+(q`1/q`2)^2)
                           by A20,SQUARE_1:def 4
          .=((q`1)^2+(q`2)^2)/(1+(q`1/q`2)^2) by XCMPLX_1:63
          .=((q`1)^2+(q`2)^2)/(1+(q`1)^2/(q`2)^2) by SQUARE_1:69
          .=((q`1)^2+(q`2)^2)/((q`1)^2/(q`2)^2+(q`2)^2/(q`2)^2)
                                by A22,XCMPLX_1:60
          .=((q`1)^2+(q`2)^2)/(((q`1)^2+(q`2)^2)/(q`2)^2)
                                by XCMPLX_1:63
          .=(q`2)^2*(((q`1)^2+(q`2)^2)/((q`1)^2+(q`2)^2))
                                by XCMPLX_1:82
          .=(q`2)^2*1 by A23,XCMPLX_1:60
          .=(q`2)^2;
            0<=|.(|[q`1/sqrt(1+(q`1/q`2)^2),q`2/sqrt(1+(q`1/q`2)^2)]|).|
                               by TOPRNS_1:26;
          then |.(|[q`1/sqrt(1+(q`1/q`2)^2),q`2/sqrt(1+(q`1/q`2)^2)]|).|>1
                               by A27,A28,SQUARE_1:89;
         hence ex p2 being Point of TOP-REAL 2 st p2=y & |.p2.|>1
                                     by A2,A3,A18;
      end;
     hence y in Cb by A1;
    end;
   let y be set;assume y in Cb;
     then consider p2 being Point of TOP-REAL 2 such that
     A29: p2=y & |.p2.|>1 by A1;
     set q=p2;
       now per cases;
     case q=0.REAL 2;
       then |.q.|=0 by TOPRNS_1:24;
      hence contradiction by A29;
     case A30:q<>0.REAL 2 &
          (q`2<=q`1 & -q`1<=q`2 or q`2>=q`1 & q`2<=-q`1);
          set px=|[q`1*sqrt(1+(q`2/q`1)^2),q`2*sqrt(1+(q`2/q`1)^2)]|;
        A31: px`1 = q`1*sqrt(1+(q`2/q`1)^2) &
         px`2 = q`2*sqrt(1+(q`2/q`1)^2) by EUCLID:56;
          0<=(q`2/q`1)^2 by SQUARE_1:72;
        then 1+(q`2/q`1)^2>0 by Th3;
        then A32:sqrt(1+(q`2/q`1)^2)>0 by SQUARE_1:93;
          0<=(px`2/px`1)^2 by SQUARE_1:72;
        then A33:1+(px`2/px`1)^2>0 by Th3;
        A34:px`2/px`1=q`2/q`1 by A31,A32,XCMPLX_1:92;
        A35: q`1=q`1*sqrt(1+(q`2/q`1)^2)/(sqrt(1+(q`2/q`1)^2))by A32,XCMPLX_1:
90
          .=px`1/(sqrt(1+(q`2/q`1)^2)) by EUCLID:56;
        A36: q`2=q`2*sqrt(1+(q`2/q`1)^2)/(sqrt(1+(q`2/q`1)^2))by A32,XCMPLX_1:
90
          .=px`2/(sqrt(1+(q`2/q`1)^2)) by EUCLID:56;
        A37: |.q.|^2=q`1^2+q`2^2 by JGRAPH_3:10;
        A38: |.q.|^2>1 by A29,SQUARE_1:59,78;
        A39:now assume px`1=0 & px`2=0;
          then A40:q`1*sqrt(1+(q`2/q`1)^2)=0 & q`2*sqrt(1+(q`2/q`1)^2)=0
                                    by EUCLID:56;
          then A41:q`1=0 by A32,XCMPLX_1:6;
            q`2=0 by A32,A40,XCMPLX_1:6;
         hence contradiction by A30,A41,EUCLID:57,58;
        end;
          q`2<=q`1 & -q`1<=q`2 or q`2>=q`1 &
        q`2*sqrt(1+(q`2/q`1)^2) <= (-q`1)*sqrt(1+(q`2/q`1)^2)
                             by A30,A32,AXIOMS:25;
        then q`2<=q`1 & (-q`1)*sqrt(1+(q`2/q`1)^2) <= q`2*sqrt(1+(q`2/q`1)^2)
           or px`2>=px`1 & px`2<=-px`1 by A31,A32,AXIOMS:25,XCMPLX_1:175;
        then A42:px`2<=px`1 & -px`1<=px`2 or px`2>=px`1 & px`2<=-px`1
                                by A31,A32,AXIOMS:25,XCMPLX_1:175;
    then A43:Sq_Circ.px=|[px`1/sqrt(1+(px`2/px`1)^2),px`2/sqrt(1+(px`2/px`1)^2)
]|
                                 by A39,JGRAPH_2:11,JGRAPH_3:def 1;
        A44:px`1/sqrt(1+(px`2/px`1)^2)=q`1 by A31,A32,A34,XCMPLX_1:90;
         px`2/sqrt(1+(px`2/px`1)^2)=q`2 by A31,A32,A34,XCMPLX_1:90;
        then A45:q=Sq_Circ.px by A43,A44,EUCLID:57;
        A46: dom Sq_Circ=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
          not px`1=0 by A39,A42;
        then A47: (px`1)^2>0 by SQUARE_1:74;
        A48: (px`2)^2>=0 by SQUARE_1:72;
           (px`1)^2/(sqrt(1+(px`2/px`1)^2))^2+(px`2/sqrt(1+(px`2/px`1)^2))^2>1
                            by A34,A35,A36,A37,A38,SQUARE_1:69;
      then (px`1)^2/(sqrt(1+(px`2/px`1)^2))^2+(px`2)^2/(sqrt(1+(px`2/px`1)^2))
^2>1
                            by SQUARE_1:69;
         then (px`1)^2/(1+(px`2/px`1)^2)+(px`2)^2/(sqrt(1+(px`2/px`1)^2))^2>1
                            by A33,SQUARE_1:def 4;
         then (px`1)^2/(1+(px`2/px`1)^2)+(px`2)^2/(1+(px`2/px`1)^2)>1
                            by A33,SQUARE_1:def 4;
         then ((px`1)^2/(1+(px`2/px`1)^2)
          +(px`2)^2/(1+(px`2/px`1)^2))*(1+(px`2/px`1)^2)>1 *(1+(px`2/px`1)^2)
                            by A33,REAL_1:70;
         then (px`1)^2/(1+(px`2/px`1)^2)*(1+(px`2/px`1)^2)
           +(px`2)^2/(1+(px`2/px`1)^2)*(1+(px`2/px`1)^2)>1 *(1+(px`2/px`1)^2)
                            by XCMPLX_1:8;
    then (px`1)^2+(px`2)^2/(1+(px`2/px`1)^2)*(1+(px`2/px`1)^2)>1 *(1+(px`2/px
`1)^2)
                            by A33,XCMPLX_1:88;
         then A49: (px`1)^2+(px`2)^2>1 *(1+(px`2/px`1)^2) by A33,XCMPLX_1:88;
           1 *(1+(px`2/px`1)^2)
         =1+(px`2)^2/(px`1)^2 by SQUARE_1:69;
         then (px`1)^2+(px`2)^2-1>1+(px`2)^2/(px`1)^2-1 by A49,REAL_1:54;
         then (px`1)^2+(px`2)^2-1>(px`2)^2/(px`1)^2 by XCMPLX_1:26;

         then ((px`1)^2+(px`2)^2-1)*(px`1)^2>((px`2)^2/(px`1)^2)*(px`1)^2
                            by A47,REAL_1:70;
         then ((px`1)^2+(px`2)^2-1)*(px`1)^2>(px`2)^2 by A47,XCMPLX_1:88;
         then ((px`1)^2+((px`2)^2-1))*(px`1)^2>(px`2)^2 by XCMPLX_1:29;
         then (px`1)^2*(px`1)^2+((px`2)^2-1)*(px`1)^2>(px`2)^2 by XCMPLX_1:8;
         then (px`1)^2*(px`1)^2+(px`1)^2*((px`2)^2-1)-(px`2)^2>0
                             by SQUARE_1:11;
         then A50: 0<(px`1)^2*(px`1)^2+((px`1)^2*(px`2)^2-(px`1)^2*1)-(px`2)^2
                                   by XCMPLX_1:40;
            (px`1)^2*(px`1)^2+((px`1)^2*(px`2)^2-(px`1)^2*1)-(px`2)^2
         = (px`1)^2*(px`1)^2+(px`1)^2*(px`2)^2-(px`1)^2*1-(px`2)^2
                                   by XCMPLX_1:29
         .= (px`1)^2*(px`1)^2-(px`1)^2*1+(px`1)^2*(px`2)^2-1 *(px`2)^2
                                   by XCMPLX_1:29
         .= (px`1)^2*((px`1)^2-1)+(px`1)^2*(px`2)^2-1 *(px`2)^2 by XCMPLX_1:40
         .= (px`1)^2*((px`1)^2-1)+((px`1)^2*(px`2)^2-1 *(px`2)^2) by XCMPLX_1:
29
         .= (px`1)^2*((px`1)^2-1)+((px`1)^2-1)*(px`2)^2 by XCMPLX_1:40
         .= ((px`1)^2-1)*((px`1)^2+(px`2)^2) by XCMPLX_1:8;
         then ( (px`1)^2-1<0 or (px`1)^2+(px`2)^2>0)&
          ((px`1)^2-1>0 or (px`1)^2+(px`2)^2<0) by A50,REAL_2:123;
         then (px`1)^2-1+1>0+1 & ((px`1)^2+(px`2)^2)>0 by A47,A48,Th3,REAL_1:53
;
         then (px`1)^2>1 by XCMPLX_1:27;
          then px`1>1 or px`1< -1 by JGRAPH_2:7,SQUARE_1:59;
        then px in Kb by A1;
      hence ex x being set st x in dom Sq_Circ & x in Kb & y=Sq_Circ.x
                               by A29,A45,A46;
     case A51:q<>0.REAL 2 & not(q`2<=q`1 & -q`1<=q`2 or q`2>=q`1 & q`2<=-q`1);
          set px=|[q`1*sqrt(1+(q`1/q`2)^2),q`2*sqrt(1+(q`1/q`2)^2)]|;
        A52:q<>0.REAL 2 & (q`1<=q`2 & -q`2<=q`1 or q`1>=q`2 & q`1<=-q`2)
                          by A51,JGRAPH_2:23;
        A53: px`2 = q`2*sqrt(1+(q`1/q`2)^2) &
         px`1 = q`1*sqrt(1+(q`1/q`2)^2) by EUCLID:56;
          0<=(q`1/q`2)^2 by SQUARE_1:72;
        then 1+(q`1/q`2)^2>0 by Th3;
        then A54:sqrt(1+(q`1/q`2)^2)>0 by SQUARE_1:93;
          0<=(px`1/px`2)^2 by SQUARE_1:72;
        then A55:1+(px`1/px`2)^2>0 by Th3;
        A56:px`1/px`2=q`1/q`2 by A53,A54,XCMPLX_1:92;
        A57: q`2=q`2*sqrt(1+(q`1/q`2)^2)/(sqrt(1+(q`1/q`2)^2))by A54,XCMPLX_1:
90
          .=px`2/(sqrt(1+(q`1/q`2)^2)) by EUCLID:56;
        A58: q`1=q`1*sqrt(1+(q`1/q`2)^2)/(sqrt(1+(q`1/q`2)^2))by A54,XCMPLX_1:
90
          .=px`1/(sqrt(1+(q`1/q`2)^2)) by EUCLID:56;
        A59: |.q.|^2=q`2^2+q`1^2 by JGRAPH_3:10;
        A60: |.q.|^2>1 by A29,SQUARE_1:59,78;
        A61:now assume px`2=0 & px`1=0;
          then A62:q`2*sqrt(1+(q`1/q`2)^2)=0 & q`1*sqrt(1+(q`1/q`2)^2)=0
                                    by EUCLID:56;
          then q`1=0 by A54,XCMPLX_1:6;
         hence contradiction by A51,A54,A62,XCMPLX_1:6;
        end;
          q`1<=q`2 & -q`2<=q`1 or q`1>=q`2 &
        q`1*sqrt(1+(q`1/q`2)^2) <= (-q`2)*sqrt(1+(q`1/q`2)^2)
                             by A52,A54,AXIOMS:25;
        then q`1<=q`2 & (-q`2)*sqrt(1+(q`1/q`2)^2) <= q`1*sqrt(1+(q`1/q`2)^2)
           or px`1>=px`2 & px`1<=-px`2 by A53,A54,AXIOMS:25,XCMPLX_1:175;
        then A63:px`1<=px`2 & -px`2<=px`1 or px`1>=px`2 & px`1<=-px`2
                                by A53,A54,AXIOMS:25,XCMPLX_1:175;
    then A64:Sq_Circ.px=|[px`1/sqrt(1+(px`1/px`2)^2),px`2/sqrt(1+(px`1/px`2)^2)
]|
                                 by A61,JGRAPH_2:11,JGRAPH_3:14;
        A65:px`2/sqrt(1+(px`1/px`2)^2)=q`2 by A53,A54,A56,XCMPLX_1:90;
         px`1/sqrt(1+(px`1/px`2)^2)=q`1 by A53,A54,A56,XCMPLX_1:90;
        then A66:q=Sq_Circ.px by A64,A65,EUCLID:57;
        A67: dom Sq_Circ=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
          not px`2=0 by A61,A63;
        then A68: (px`2)^2>0 by SQUARE_1:74;
        A69: (px`1)^2>=0 by SQUARE_1:72;
           (px`2)^2/(sqrt(1+(px`1/px`2)^2))^2+(px`1/sqrt(1+(px`1/px`2)^2))^2>1
                            by A56,A57,A58,A59,A60,SQUARE_1:69;
      then (px`2)^2/(sqrt(1+(px`1/px`2)^2))^2+(px`1)^2/(sqrt(1+(px`1/px`2)^2))
^2>1
                            by SQUARE_1:69;
         then (px`2)^2/(1+(px`1/px`2)^2)+(px`1)^2/(sqrt(1+(px`1/px`2)^2))^2>1
                            by A55,SQUARE_1:def 4;
         then (px`2)^2/(1+(px`1/px`2)^2)+(px`1)^2/(1+(px`1/px`2)^2)>1
                            by A55,SQUARE_1:def 4;
         then ((px`2)^2/(1+(px`1/px`2)^2)
          +(px`1)^2/(1+(px`1/px`2)^2))*(1+(px`1/px`2)^2)>1 *(1+(px`1/px`2)^2)
                            by A55,REAL_1:70;
         then (px`2)^2/(1+(px`1/px`2)^2)*(1+(px`1/px`2)^2)
           +(px`1)^2/(1+(px`1/px`2)^2)*(1+(px`1/px`2)^2)>1 *(1+(px`1/px`2)^2)
                            by XCMPLX_1:8;
    then (px`2)^2+(px`1)^2/(1+(px`1/px`2)^2)*(1+(px`1/px`2)^2)>1 *(1+(px`1/px
`2)^2)
                            by A55,XCMPLX_1:88;
         then A70: (px`2)^2+(px`1)^2>1 *(1+(px`1/px`2)^2) by A55,XCMPLX_1:88;
           1 *(1+(px`1/px`2)^2)
         =1+(px`1)^2/(px`2)^2 by SQUARE_1:69;
         then (px`2)^2+(px`1)^2-1>1+(px`1)^2/(px`2)^2-1 by A70,REAL_1:54;
         then (px`2)^2+(px`1)^2-1>(px`1)^2/(px`2)^2 by XCMPLX_1:26;
         then ((px`2)^2+(px`1)^2-1)*(px`2)^2>((px`1)^2/(px`2)^2)*(px`2)^2
                            by A68,REAL_1:70;
         then ((px`2)^2+(px`1)^2-1)*(px`2)^2>(px`1)^2 by A68,XCMPLX_1:88;
         then ((px`2)^2+((px`1)^2-1))*(px`2)^2>(px`1)^2 by XCMPLX_1:29;
         then (px`2)^2*(px`2)^2+((px`1)^2-1)*(px`2)^2>(px`1)^2 by XCMPLX_1:8;
         then (px`2)^2*(px`2)^2+(px`2)^2*((px`1)^2-1)-(px`1)^2>0
                             by SQUARE_1:11;
         then A71: 0<(px`2)^2*(px`2)^2+((px`2)^2*(px`1)^2-(px`2)^2*1)-(px`1)^2
                                   by XCMPLX_1:40;
            (px`2)^2*(px`2)^2+((px`2)^2*(px`1)^2-(px`2)^2*1)-(px`1)^2
         = (px`2)^2*(px`2)^2+(px`2)^2*(px`1)^2-(px`2)^2*1-(px`1)^2
                                   by XCMPLX_1:29
         .= (px`2)^2*(px`2)^2-(px`2)^2*1+(px`2)^2*(px`1)^2-1 *(px`1)^2
                                   by XCMPLX_1:29
         .= (px`2)^2*((px`2)^2-1)+(px`2)^2*(px`1)^2-1 *(px`1)^2 by XCMPLX_1:40
         .= (px`2)^2*((px`2)^2-1)+((px`2)^2*(px`1)^2-1 *(px`1)^2) by XCMPLX_1:
29
         .= (px`2)^2*((px`2)^2-1)+((px`2)^2-1)*(px`1)^2 by XCMPLX_1:40
         .= ((px`2)^2-1)*((px`2)^2+(px`1)^2) by XCMPLX_1:8;
         then ( (px`2)^2-1<0 or (px`1)^2+(px`2)^2>0)&
          ((px`2)^2-1>0 or (px`1)^2+(px`2)^2<0) by A71,REAL_2:123;
         then (px`2)^2-1+1>0+1 & ((px`1)^2+(px`2)^2)>0 by A68,A69,Th3,REAL_1:53
;
         then (px`2)^2>1 by XCMPLX_1:27;
          then px`2>1 or px`2< -1 by JGRAPH_2:7,SQUARE_1:59;
        then px in Kb by A1;
      hence ex x being set st x in dom Sq_Circ & x in Kb & y=Sq_Circ.x
                               by A29,A66,A67;
     end;
    hence thesis by FUNCT_1:def 12;
 end;

theorem Th36: for Kb,Cb being Subset of TOP-REAL 2 st
 Kb={p: -1 <=p`1 & p`1<= 1 & -1 <=p`2 & p`2<= 1}&
 Cb={p2 where p2 is Point of TOP-REAL 2: |.p2.|<=1} holds Sq_Circ.:Kb=Cb
 proof let Kb,Cb be Subset of TOP-REAL 2;
  assume
   A1:Kb={p: -1 <=p`1 & p`1<= 1 & -1 <=p`2 & p`2<= 1}&
      Cb={p2 where p2 is Point of TOP-REAL 2: |.p2.|<=1};
   thus Sq_Circ.:Kb c= Cb
    proof let y be set;assume y in Sq_Circ.:Kb;
      then consider x being set such that
      A2: x in dom Sq_Circ & x in Kb & y=Sq_Circ.x
                   by FUNCT_1:def 12;
      consider q being Point of TOP-REAL 2 such that
      A3:q=x &( -1 <=q`1 & q`1<= 1 & -1 <=q`2 & q`2<= 1 ) by A1,A2;
        now per cases;
      case A4: q=0.REAL 2;
        then A5: Sq_Circ.q=q by JGRAPH_3:def 1;
          |.q.|=0 by A4,TOPRNS_1:24;
       hence ex p2 being Point of TOP-REAL 2 st p2=y & |.p2.|<=1 by A2,A3,A5;
      case A6:q<>0.REAL 2 &
          (q`2<=q`1 & -q`1<=q`2 or q`2>=q`1 & q`2<=-q`1);
        then A7:Sq_Circ.q=|[q`1/sqrt(1+(q`2/q`1)^2),q`2/sqrt(1+(q`2/q`1)^2)]|
                                 by JGRAPH_3:def 1;
        A8: (|[q`1/sqrt(1+(q`2/q`1)^2),q`2/sqrt(1+(q`2/q`1)^2)]|)`1
         = q`1/sqrt(1+(q`2/q`1)^2) &
         (|[q`1/sqrt(1+(q`2/q`1)^2),q`2/sqrt(1+(q`2/q`1)^2)]|)`2
         = q`2/sqrt(1+(q`2/q`1)^2) by EUCLID:56;
          0<=(q`2/q`1)^2 by SQUARE_1:72;
        then A9:1+(q`2/q`1)^2>0 by Th3;
        A10: now assume A11: q`1=0;
          then q`2=0 by A6;
         hence contradiction by A6,A11,EUCLID:57,58;
        end;
        then A12: (q`1)^2>0 by SQUARE_1:74;
        A13: (q`1)^2+(q`2)^2 <>0 by A10,COMPLEX1:2;
            (q`1)^2<=1^2 by A3,JGRAPH_2:7;
          then A14: sqrt((q`1)^2)<=1 by A12,SQUARE_1:59,83,94;
           A15: |.(|[q`1/sqrt(1+(q`2/q`1)^2),q`2/sqrt(1+(q`2/q`1)^2)]|).|^2
          =((q`1)/sqrt(1+(q`2/(q`1))^2))^2+(q`2/sqrt(1+(q`2/(q`1))^2))^2
                        by A8,JGRAPH_3:10
          .=(q`1)^2/(sqrt(1+(q`2/(q`1))^2))^2+(q`2/sqrt(1+(q`2/(q`1))^2))^2
                        by SQUARE_1:69
          .=(q`1)^2/(sqrt(1+(q`2/(q`1))^2))^2
                  +(q`2)^2/(sqrt(1+(q`2/(q`1))^2))^2 by SQUARE_1:69
          .=(q`1)^2/(1+(q`2/q`1)^2)+(q`2)^2/(sqrt(1+(q`2/q`1)^2))^2
                               by A9,SQUARE_1:def 4
          .=(q`1)^2/(1+(q`2/q`1)^2)+(q`2)^2/(1+(q`2/q`1)^2)
                           by A9,SQUARE_1:def 4
          .=((q`1)^2+(q`2)^2)/(1+(q`2/q`1)^2) by XCMPLX_1:63
          .=((q`1)^2+(q`2)^2)/(1+(q`2)^2/(q`1)^2) by SQUARE_1:69
          .=((q`1)^2+(q`2)^2)/((q`1)^2/(q`1)^2+(q`2)^2/(q`1)^2)
                                by A12,XCMPLX_1:60
          .=((q`1)^2+(q`2)^2)/(((q`1)^2+(q`2)^2)/(q`1)^2)
                                by XCMPLX_1:63
          .=(q`1)^2*(((q`1)^2+(q`2)^2)/((q`1)^2+(q`2)^2))
                                by XCMPLX_1:82
          .=(q`1)^2*1 by A13,XCMPLX_1:60
          .=(q`1)^2;
            0<=|.(|[q`1/sqrt(1+(q`2/q`1)^2),q`2/sqrt(1+(q`2/q`1)^2)]|).|
                               by TOPRNS_1:26;
          then |.(|[q`1/sqrt(1+(q`2/q`1)^2),q`2/sqrt(1+(q`2/q`1)^2)]|).|<=1
                               by A14,A15,SQUARE_1:89;
         hence ex p2 being Point of TOP-REAL 2 st p2=y & |.p2.|<=1
                                     by A2,A3,A7;
      case A16:q<>0.REAL 2 &
        not(q`2<=q`1 & -q`1<=q`2 or q`2>=q`1 & q`2<=-q`1);
        then A17:Sq_Circ.q=|[q`1/sqrt(1+(q`1/q`2)^2),q`2/sqrt(1+(q`1/q`2)^2)]|
                                 by JGRAPH_3:def 1;
        A18: (|[q`1/sqrt(1+(q`1/q`2)^2),q`2/sqrt(1+(q`1/q`2)^2)]|)`1
         = q`1/sqrt(1+(q`1/q`2)^2) &
         (|[q`1/sqrt(1+(q`1/q`2)^2),q`2/sqrt(1+(q`1/q`2)^2)]|)`2
         = q`2/sqrt(1+(q`1/q`2)^2) by EUCLID:56;
          0<=(q`1/q`2)^2 by SQUARE_1:72;
        then A19:1+(q`1/q`2)^2>0 by Th3;
        A20: q`2 <>0 by A16,REAL_1:66;
        then A21: (q`2)^2>0 by SQUARE_1:74;
        A22: (q`1)^2+(q`2)^2 <>0 by A20,COMPLEX1:2;
            (q`2)^2<=1 by A3,JGRAPH_2:7,SQUARE_1:59;
          then A23: sqrt((q`2)^2)<=1 by A21,SQUARE_1:83,94;
           A24: |.(|[q`1/sqrt(1+(q`1/q`2)^2),q`2/sqrt(1+(q`1/q`2)^2)]|).|^2
          =((q`1)/sqrt(1+(q`1/(q`2))^2))^2+(q`2/sqrt(1+(q`1/(q`2))^2))^2
                        by A18,JGRAPH_3:10
          .=(q`1)^2/(sqrt(1+(q`1/(q`2))^2))^2+(q`2/sqrt(1+(q`1/(q`2))^2))^2
                        by SQUARE_1:69
          .=(q`1)^2/(sqrt(1+(q`1/(q`2))^2))^2
                  +(q`2)^2/(sqrt(1+(q`1/(q`2))^2))^2 by SQUARE_1:69
          .=(q`1)^2/(1+(q`1/q`2)^2)+(q`2)^2/(sqrt(1+(q`1/q`2)^2))^2
                               by A19,SQUARE_1:def 4
          .=(q`1)^2/(1+(q`1/q`2)^2)+(q`2)^2/(1+(q`1/q`2)^2)
                           by A19,SQUARE_1:def 4
          .=((q`1)^2+(q`2)^2)/(1+(q`1/q`2)^2) by XCMPLX_1:63
          .=((q`1)^2+(q`2)^2)/(1+(q`1)^2/(q`2)^2) by SQUARE_1:69
          .=((q`1)^2+(q`2)^2)/((q`1)^2/(q`2)^2+(q`2)^2/(q`2)^2)
                                by A21,XCMPLX_1:60
          .=((q`1)^2+(q`2)^2)/(((q`1)^2+(q`2)^2)/(q`2)^2)
                                by XCMPLX_1:63
          .=(q`2)^2*(((q`1)^2+(q`2)^2)/((q`1)^2+(q`2)^2)) by XCMPLX_1:82
          .=(q`2)^2*1 by A22,XCMPLX_1:60
          .=(q`2)^2;
            0<=|.(|[q`1/sqrt(1+(q`1/q`2)^2),q`2/sqrt(1+(q`1/q`2)^2)]|).|
                               by TOPRNS_1:26;
          then |.(|[q`1/sqrt(1+(q`1/q`2)^2),q`2/sqrt(1+(q`1/q`2)^2)]|).|<=1
                               by A23,A24,SQUARE_1:89;
         hence ex p2 being Point of TOP-REAL 2 st p2=y & |.p2.|<=1
                                     by A2,A3,A17;
      end;
     hence y in Cb by A1;
    end;
   let y be set;assume y in Cb;
     then consider p2 being Point of TOP-REAL 2 such that
     A25: p2=y & |.p2.|<=1 by A1;
     set q=p2;
       now per cases;
     case A26: q=0.REAL 2;
        then q`1=0 & q`2=0 by EUCLID:56,58;
        then A27: y in Kb by A1,A25;
        A28: Sq_Circ".y=y by A25,A26,JGRAPH_3:38;
        A29: dom Sq_Circ=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
          y=Sq_Circ.y by A25,A28,FUNCT_1:57,JGRAPH_3:54;
      hence ex x being set st x in dom Sq_Circ & x in Kb & y=Sq_Circ.x
                              by A27,A29;
     case A30:q<>0.REAL 2 &
          (q`2<=q`1 & -q`1<=q`2 or q`2>=q`1 & q`2<=-q`1);
          set px=|[q`1*sqrt(1+(q`2/q`1)^2),q`2*sqrt(1+(q`2/q`1)^2)]|;
        A31: px`1 = q`1*sqrt(1+(q`2/q`1)^2) &
         px`2 = q`2*sqrt(1+(q`2/q`1)^2) by EUCLID:56;
          0<=(q`2/q`1)^2 by SQUARE_1:72;
        then 1+(q`2/q`1)^2>0 by Th3;
        then A32:sqrt(1+(q`2/q`1)^2)>0 by SQUARE_1:93;
          0<=(px`2/px`1)^2 by SQUARE_1:72;
        then A33:1+(px`2/px`1)^2>0 by Th3;
        A34:px`2/px`1=q`2/q`1 by A31,A32,XCMPLX_1:92;
        A35: q`1=q`1*sqrt(1+(q`2/q`1)^2)/(sqrt(1+(q`2/q`1)^2))by A32,XCMPLX_1:
90
          .=px`1/(sqrt(1+(q`2/q`1)^2)) by EUCLID:56;
        A36: q`2=q`2*sqrt(1+(q`2/q`1)^2)/(sqrt(1+(q`2/q`1)^2))by A32,XCMPLX_1:
90
          .=px`2/(sqrt(1+(q`2/q`1)^2)) by EUCLID:56;
        A37: |.q.|^2=q`1^2+q`2^2 by JGRAPH_3:10;
          |.q.| >=0 by TOPRNS_1:26;
        then A38: |.q.|^2<=1 by A25,SQUARE_1:59,77;
        A39:now assume px`1=0 & px`2=0;
          then A40:q`1*sqrt(1+(q`2/q`1)^2)=0 & q`2*sqrt(1+(q`2/q`1)^2)=0
                                    by EUCLID:56;
          then A41:q`1=0 by A32,XCMPLX_1:6;
            q`2=0 by A32,A40,XCMPLX_1:6;
         hence contradiction by A30,A41,EUCLID:57,58;
        end;
          q`2<=q`1 & -q`1<=q`2 or q`2>=q`1 &
        q`2*sqrt(1+(q`2/q`1)^2) <= (-q`1)*sqrt(1+(q`2/q`1)^2)
                             by A30,A32,AXIOMS:25;
        then q`2<=q`1 & (-q`1)*sqrt(1+(q`2/q`1)^2) <= q`2*sqrt(1+(q`2/q`1)^2)
           or px`2>=px`1 & px`2<=-px`1 by A31,A32,AXIOMS:25,XCMPLX_1:175;
        then A42:px`2<=px`1 & -px`1<=px`2 or px`2>=px`1 & px`2<=-px`1
                                by A31,A32,AXIOMS:25,XCMPLX_1:175;
    then A43:Sq_Circ.px=|[px`1/sqrt(1+(px`2/px`1)^2),px`2/sqrt(1+(px`2/px`1)^2)
]|
                                 by A39,JGRAPH_2:11,JGRAPH_3:def 1;
          px`2<=px`1 & --px`1>=-px`2 or px`2>=px`1 & px`2<=-px`1
                                by A42,REAL_1:50;
        then A44:px`2<=px`1 & px`1>=-px`2 or px`2>=px`1 & -px`2>=--px`1
                                by REAL_1:50;
        A45:px`1/sqrt(1+(px`2/px`1)^2)=q`1 by A31,A32,A34,XCMPLX_1:90;
         px`2/sqrt(1+(px`2/px`1)^2)=q`2 by A31,A32,A34,XCMPLX_1:90;
        then A46:q=Sq_Circ.px by A43,A45,EUCLID:57;
        A47: dom Sq_Circ=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
          not px`1=0 by A39,A42;
        then A48: (px`1)^2>0 by SQUARE_1:74;
        A49: (px`2)^2>=0 by SQUARE_1:72;
           (px`1)^2/(sqrt(1+(px`2/px`1)^2))^2+(px`2/sqrt(1+(px`2/px`1)^2))^2
          <=1
                            by A34,A35,A36,A37,A38,SQUARE_1:69;
      then (px`1)^2/(sqrt(1+(px`2/px`1)^2))^2+(px`2)^2/(sqrt(1+(px`2/px`1)^2))
^2<=1
                            by SQUARE_1:69;
         then (px`1)^2/(1+(px`2/px`1)^2)+(px`2)^2/(sqrt(1+(px`2/px`1)^2))^2<=1
                            by A33,SQUARE_1:def 4;
         then (px`1)^2/(1+(px`2/px`1)^2)+(px`2)^2/(1+(px`2/px`1)^2)<=1
                            by A33,SQUARE_1:def 4;
         then ((px`1)^2/(1+(px`2/px`1)^2)
          +(px`2)^2/(1+(px`2/px`1)^2))*(1+(px`2/px`1)^2)<=1 *(1+(px`2/px`1)^2)
                            by A33,AXIOMS:25;
         then (px`1)^2/(1+(px`2/px`1)^2)*(1+(px`2/px`1)^2)
           +(px`2)^2/(1+(px`2/px`1)^2)*(1+(px`2/px`1)^2)<=1 *(1+(px`2/px`1)^2)
                            by XCMPLX_1:8;
    then (px`1)^2+(px`2)^2/(1+(px`2/px`1)^2)*(1+(px`2/px`1)^2)
              <=1 *(1+(px`2/px`1)^2)
                            by A33,XCMPLX_1:88;
         then A50: (px`1)^2+(px`2)^2<=1 *(1+(px`2/px`1)^2) by A33,XCMPLX_1:88;
           1 *(1+(px`2/px`1)^2)
         =1+(px`2)^2/(px`1)^2 by SQUARE_1:69;
         then (px`1)^2+(px`2)^2-1<=1+(px`2)^2/(px`1)^2-1 by A50,REAL_1:49;
         then (px`1)^2+(px`2)^2-1<=(px`2)^2/(px`1)^2 by XCMPLX_1:26;
         then ((px`1)^2+(px`2)^2-1)*(px`1)^2<=((px`2)^2/(px`1)^2)*(px`1)^2
                            by A48,AXIOMS:25;
         then ((px`1)^2+(px`2)^2-1)*(px`1)^2<=(px`2)^2 by A48,XCMPLX_1:88;
         then ((px`1)^2+((px`2)^2-1))*(px`1)^2<=(px`2)^2 by XCMPLX_1:29;
         then (px`1)^2*(px`1)^2+((px`2)^2-1)*(px`1)^2<=(px`2)^2 by XCMPLX_1:8;
         then (px`1)^2*(px`1)^2+(px`1)^2*((px`2)^2-1)-(px`2)^2<=0
                             by REAL_2:106;
         then A51: 0>=(px`1)^2*(px`1)^2+((px`1)^2*(px`2)^2-(px`1)^2*1)-(px`2)^2
                                   by XCMPLX_1:40;
            (px`1)^2*(px`1)^2+((px`1)^2*(px`2)^2-(px`1)^2*1)-(px`2)^2
         = (px`1)^2*(px`1)^2+(px`1)^2*(px`2)^2-(px`1)^2*1-(px`2)^2
                                   by XCMPLX_1:29
         .= (px`1)^2*(px`1)^2-(px`1)^2*1+(px`1)^2*(px`2)^2-1 *(px`2)^2
                                   by XCMPLX_1:29
         .= (px`1)^2*((px`1)^2-1)+(px`1)^2*(px`2)^2-1 *(px`2)^2 by XCMPLX_1:40
         .= (px`1)^2*((px`1)^2-1)+((px`1)^2*(px`2)^2-1 *(px`2)^2) by XCMPLX_1:
29
         .= (px`1)^2*((px`1)^2-1)+((px`1)^2-1)*(px`2)^2 by XCMPLX_1:40
         .= ((px`1)^2-1)*((px`1)^2+(px`2)^2) by XCMPLX_1:8;
         then (px`1)^2-1<=0 & (px`1)^2-1>=0 or
         (px`1)^2-1<=0 & (px`1)^2+(px`2)^2>=0 or
         (px`1)^2+(px`2)^2<=0 & (px`1)^2-1>=0 or
         (px`1)^2+(px`2)^2<=0 & (px`1)^2+(px`2)^2>=0 by A51,REAL_2:122;
         then (px`1)^2-1+1<=0+1 & ((px`1)^2+(px`2)^2)>=0 by A48,A49,Th3,REAL_1:
55;
         then (px`1)^2<=1 by XCMPLX_1:27;
          then A52:px`1<=1 & px`1>= -1 by JGRAPH_2:5,SQUARE_1:59;
          then px`2<=1 & 1>= -px`2 or px`2>= -1 & -px`2>=px`1
                                by A44,AXIOMS:22;
          then px`2<=1 & -1<= --px`2 or px`2>= -1 & -px`2>= -1
                                by A52,AXIOMS:22,REAL_1:50;
          then px`2<=1 & -1<=px`2 or px`2>= -1 & --px`2<= --1
                                by REAL_1:50;
        then px in Kb by A1,A52;
      hence ex x being set st x in dom Sq_Circ & x in Kb & y=Sq_Circ.x
                               by A25,A46,A47;
     case A53:q<>0.REAL 2 & not(q`2<=q`1 & -q`1<=q`2 or q`2>=q`1 & q`2<=-q`1);
          set px=|[q`1*sqrt(1+(q`1/q`2)^2),q`2*sqrt(1+(q`1/q`2)^2)]|;
        A54:q<>0.REAL 2 & (q`1<=q`2 & -q`2<=q`1 or q`1>=q`2 & q`1<=-q`2)
                          by A53,JGRAPH_2:23;
        A55: px`2 = q`2*sqrt(1+(q`1/q`2)^2) &
         px`1 = q`1*sqrt(1+(q`1/q`2)^2) by EUCLID:56;
          0<=(q`1/q`2)^2 by SQUARE_1:72;
        then 1+(q`1/q`2)^2>0 by Th3;
        then A56:sqrt(1+(q`1/q`2)^2)>0 by SQUARE_1:93;
          0<=(px`1/px`2)^2 by SQUARE_1:72;
        then A57:1+(px`1/px`2)^2>0 by Th3;
        A58:px`1/px`2=q`1/q`2 by A55,A56,XCMPLX_1:92;
        A59: q`2=q`2*sqrt(1+(q`1/q`2)^2)/(sqrt(1+(q`1/q`2)^2))by A56,XCMPLX_1:
90
          .=px`2/(sqrt(1+(q`1/q`2)^2)) by EUCLID:56;
        A60: q`1=q`1*sqrt(1+(q`1/q`2)^2)/(sqrt(1+(q`1/q`2)^2))by A56,XCMPLX_1:
90
          .=px`1/(sqrt(1+(q`1/q`2)^2)) by EUCLID:56;
        A61: |.q.|^2=q`2^2+q`1^2 by JGRAPH_3:10;
          |.q.| >=0 by TOPRNS_1:26;
        then A62: |.q.|^2<=1 by A25,SQUARE_1:59,77;
        A63:now assume px`2=0 & px`1=0;
          then A64:q`2*sqrt(1+(q`1/q`2)^2)=0 & q`1*sqrt(1+(q`1/q`2)^2)=0
                                    by EUCLID:56;
          then q`1=0 by A56,XCMPLX_1:6;
         hence contradiction by A53,A56,A64,XCMPLX_1:6;
        end;
          q`1<=q`2 & -q`2<=q`1 or q`1>=q`2 &
        q`1*sqrt(1+(q`1/q`2)^2) <= (-q`2)*sqrt(1+(q`1/q`2)^2)
                             by A54,A56,AXIOMS:25;
        then q`1<=q`2 & (-q`2)*sqrt(1+(q`1/q`2)^2) <= q`1*sqrt(1+(q`1/q`2)^2)
           or px`1>=px`2 & px`1<=-px`2 by A55,A56,AXIOMS:25,XCMPLX_1:175;
        then A65:px`1<=px`2 & -px`2<=px`1 or px`1>=px`2 & px`1<=-px`2
                                by A55,A56,AXIOMS:25,XCMPLX_1:175;
    then A66:Sq_Circ.px=|[px`1/sqrt(1+(px`1/px`2)^2),px`2/sqrt(1+(px`1/px`2)^2)
]|
                                 by A63,JGRAPH_2:11,JGRAPH_3:14;
          px`1<=px`2 & --px`2>=-px`1 or px`1>=px`2 & px`1<=-px`2
                                by A65,REAL_1:50;
        then A67:px`1<=px`2 & px`2>=-px`1 or px`1>=px`2 & -px`1>=--px`2
                                by REAL_1:50;
        A68:px`2/sqrt(1+(px`1/px`2)^2)=q`2 by A55,A56,A58,XCMPLX_1:90;
         px`1/sqrt(1+(px`1/px`2)^2)=q`1 by A55,A56,A58,XCMPLX_1:90;
        then A69:q=Sq_Circ.px by A66,A68,EUCLID:57;
        A70: dom Sq_Circ=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
          not px`2=0 by A63,A65;
        then A71: (px`2)^2>0 by SQUARE_1:74;
        A72: (px`1)^2>=0 by SQUARE_1:72;
           (px`2)^2/(sqrt(1+(px`1/px`2)^2))^2+(px`1/sqrt(1+(px`1/px`2)^2))^2<=1
                            by A58,A59,A60,A61,A62,SQUARE_1:69;
      then (px`2)^2/(sqrt(1+(px`1/px`2)^2))^2+(px`1)^2/(sqrt(1+(px`1/px`2)^2))
^2<=1
                            by SQUARE_1:69;
         then (px`2)^2/(1+(px`1/px`2)^2)+(px`1)^2/(sqrt(1+(px`1/px`2)^2))^2<=1
                            by A57,SQUARE_1:def 4;
         then (px`2)^2/(1+(px`1/px`2)^2)+(px`1)^2/(1+(px`1/px`2)^2)<=1
                            by A57,SQUARE_1:def 4;
         then ((px`2)^2/(1+(px`1/px`2)^2)
          +(px`1)^2/(1+(px`1/px`2)^2))*(1+(px`1/px`2)^2)<=1 *(1+(px`1/px`2)^2)
                            by A57,AXIOMS:25;
         then (px`2)^2/(1+(px`1/px`2)^2)*(1+(px`1/px`2)^2)
           +(px`1)^2/(1+(px`1/px`2)^2)*(1+(px`1/px`2)^2)<=1 *(1+(px`1/px`2)^2)
                            by XCMPLX_1:8;
    then (px`2)^2+(px`1)^2/(1+(px`1/px`2)^2)*(1+(px`1/px`2)^2)
                      <=1 *(1+(px`1/px`2)^2)
                            by A57,XCMPLX_1:88;
         then A73: (px`2)^2+(px`1)^2<=1 *(1+(px`1/px`2)^2) by A57,XCMPLX_1:88;
           1 *(1+(px`1/px`2)^2)
         =1+(px`1)^2/(px`2)^2 by SQUARE_1:69;
         then (px`2)^2+(px`1)^2-1<=1+(px`1)^2/(px`2)^2-1 by A73,REAL_1:49;
         then (px`2)^2+(px`1)^2-1<=(px`1)^2/(px`2)^2 by XCMPLX_1:26;

         then ((px`2)^2+(px`1)^2-1)*(px`2)^2<=((px`1)^2/(px`2)^2)*(px`2)^2
                            by A71,AXIOMS:25;
         then ((px`2)^2+(px`1)^2-1)*(px`2)^2<=(px`1)^2 by A71,XCMPLX_1:88;
         then ((px`2)^2+((px`1)^2-1))*(px`2)^2<=(px`1)^2 by XCMPLX_1:29;
         then (px`2)^2*(px`2)^2+((px`1)^2-1)*(px`2)^2<=(px`1)^2 by XCMPLX_1:8;
         then (px`2)^2*(px`2)^2+(px`2)^2*((px`1)^2-1)-(px`1)^2<=0
                             by REAL_2:106;
         then A74: 0>=(px`2)^2*(px`2)^2+((px`2)^2*(px`1)^2-(px`2)^2*1)-(px`1)^2
                                   by XCMPLX_1:40;
            (px`2)^2*(px`2)^2+((px`2)^2*(px`1)^2-(px`2)^2*1)-(px`1)^2
         = (px`2)^2*(px`2)^2+(px`2)^2*(px`1)^2-(px`2)^2*1-(px`1)^2
                                   by XCMPLX_1:29
         .= (px`2)^2*(px`2)^2-(px`2)^2*1+(px`2)^2*(px`1)^2-1 *(px`1)^2
                                   by XCMPLX_1:29
         .= (px`2)^2*((px`2)^2-1)+(px`2)^2*(px`1)^2-1 *(px`1)^2 by XCMPLX_1:40
         .= (px`2)^2*((px`2)^2-1)+((px`2)^2*(px`1)^2-1 *(px`1)^2) by XCMPLX_1:
29
         .= (px`2)^2*((px`2)^2-1)+((px`2)^2-1)*(px`1)^2 by XCMPLX_1:40
         .= ((px`2)^2-1)*((px`2)^2+(px`1)^2) by XCMPLX_1:8;
         then ((px`2)^2-1<=0 or ((px`2)^2+(px`1)^2)<=0)
         &((px`2)^2-1>=0 or (px`2)^2+(px`1)^2>=0) by A74,REAL_2:122;

         then (px`2)^2-1+1<=0+1 & ((px`2)^2+(px`1)^2)>=0 by A71,A72,Th3,REAL_1:
55;
         then (px`2)^2<=1 by XCMPLX_1:27;
          then A75:px`2<=1 & px`2>= -1 by JGRAPH_2:5,SQUARE_1:59;
          then px`1<=1 & 1>= -px`1 or px`1>= -1 & -px`1>=px`2
                                by A67,AXIOMS:22;
          then px`1<=1 & -1<= --px`1 or px`1>= -1 & -px`1>= -1
                                by A75,AXIOMS:22,REAL_1:50;
          then px`1<=1 & -1<=px`1 or px`1>= -1 & --px`1<= --1
                                by REAL_1:50;
        then px in Kb by A1,A75;
      hence ex x being set st x in dom Sq_Circ & x in Kb & y=Sq_Circ.x
                               by A25,A69,A70;
     end;
    hence thesis by FUNCT_1:def 12;
 end;

theorem Th37: for Kb,Cb being Subset of TOP-REAL 2 st
 Kb={p: not(-1 <p`1 & p`1< 1 & -1 <p`2 & p`2< 1)}&
 Cb={p2 where p2 is Point of TOP-REAL 2: |.p2.|>=1} holds Sq_Circ.:Kb=Cb
 proof let Kb,Cb be Subset of TOP-REAL 2;
  assume
   A1:
 Kb={p: not(-1 <p`1 & p`1< 1 & -1 <p`2 & p`2< 1)}&
 Cb={p2 where p2 is Point of TOP-REAL 2: |.p2.|>=1};
   thus Sq_Circ.:Kb c= Cb
    proof let y be set;assume y in Sq_Circ.:Kb;
      then consider x being set such that
      A2: x in dom Sq_Circ & x in Kb & y=Sq_Circ.x
                   by FUNCT_1:def 12;
      consider q being Point of TOP-REAL 2 such that
      A3:q=x &( not(-1 <q`1 & q`1< 1 & -1 <q`2 & q`2< 1) ) by A1,A2;
        now per cases;
      case q=0.REAL 2;
        then q`1=0 & q`2=0 by EUCLID:56,58;
       hence contradiction by A3;
      case A4:q<>0.REAL 2 &
          (q`2<=q`1 & -q`1<=q`2 or q`2>=q`1 & q`2<=-q`1);
        then A5:Sq_Circ.q=|[q`1/sqrt(1+(q`2/q`1)^2),q`2/sqrt(1+(q`2/q`1)^2)]|
                                 by JGRAPH_3:def 1;
        A6: not(-1 <q`2 & q`2< 1) implies -1>=q`1 or q`1>=1
        proof assume A7: not(-1 <q`2 & q`2< 1);
            now per cases by A7;
          case A8: -1>=q`2;
            then -q`1<= -1 or q`2>=q`1 & q`2<= -q`1 by A4,AXIOMS:22;
           hence -1>=q`1 or q`1>=1 by A8,AXIOMS:22,REAL_1:50;
          case q`2>=1;
            then 1<=q`1 or 1<= -q`1 by A4,AXIOMS:22;
            then 1<=q`1 or --q`1<= -1 by REAL_1:50;
           hence -1>=q`1 or q`1>=1;
          end;
         hence -1>=q`1 or q`1>=1;
        end;
        A9: (|[q`1/sqrt(1+(q`2/q`1)^2),q`2/sqrt(1+(q`2/q`1)^2)]|)`1
         = q`1/sqrt(1+(q`2/q`1)^2) &
         (|[q`1/sqrt(1+(q`2/q`1)^2),q`2/sqrt(1+(q`2/q`1)^2)]|)`2
         = q`2/sqrt(1+(q`2/q`1)^2) by EUCLID:56;
          0<=(q`2/q`1)^2 by SQUARE_1:72;
        then A10:1+(q`2/q`1)^2>0 by Th3;
        A11: now assume A12: q`1=0;
          then q`2=0 by A4;
         hence contradiction by A4,A12,EUCLID:57,58;
        end;
        then A13: (q`1)^2>0 by SQUARE_1:74;
        A14: (q`1)^2+(q`2)^2 <>0 by A11,COMPLEX1:2;
           (q`1)^2>=1^2 by A3,A6,JGRAPH_2:6;
          then A15: sqrt((q`1)^2)>=1 by SQUARE_1:59,83,94;
           A16: |.(|[q`1/sqrt(1+(q`2/q`1)^2),q`2/sqrt(1+(q`2/q`1)^2)]|).|^2
          =((q`1)/sqrt(1+(q`2/(q`1))^2))^2+(q`2/sqrt(1+(q`2/(q`1))^2))^2
                        by A9,JGRAPH_3:10
          .=(q`1)^2/(sqrt(1+(q`2/(q`1))^2))^2+(q`2/sqrt(1+(q`2/(q`1))^2))^2
                        by SQUARE_1:69
          .=(q`1)^2/(sqrt(1+(q`2/(q`1))^2))^2
                  +(q`2)^2/(sqrt(1+(q`2/(q`1))^2))^2 by SQUARE_1:69
          .=(q`1)^2/(1+(q`2/q`1)^2)+(q`2)^2/(sqrt(1+(q`2/q`1)^2))^2
                               by A10,SQUARE_1:def 4
          .=(q`1)^2/(1+(q`2/q`1)^2)+(q`2)^2/(1+(q`2/q`1)^2)
                           by A10,SQUARE_1:def 4
          .=((q`1)^2+(q`2)^2)/(1+(q`2/q`1)^2) by XCMPLX_1:63
          .=((q`1)^2+(q`2)^2)/(1+(q`2)^2/(q`1)^2) by SQUARE_1:69
          .=((q`1)^2+(q`2)^2)/((q`1)^2/(q`1)^2+(q`2)^2/(q`1)^2)
                                by A13,XCMPLX_1:60
          .=((q`1)^2+(q`2)^2)/(((q`1)^2+(q`2)^2)/(q`1)^2)
                                by XCMPLX_1:63
          .=(q`1)^2*(((q`1)^2+(q`2)^2)/((q`1)^2+(q`2)^2))
                                by XCMPLX_1:82
          .=(q`1)^2*1 by A14,XCMPLX_1:60
          .=(q`1)^2;
            0<=|.(|[q`1/sqrt(1+(q`2/q`1)^2),q`2/sqrt(1+(q`2/q`1)^2)]|).|
                               by TOPRNS_1:26;
          then |.(|[q`1/sqrt(1+(q`2/q`1)^2),q`2/sqrt(1+(q`2/q`1)^2)]|).|>=1
                               by A15,A16,SQUARE_1:89;
         hence ex p2 being Point of TOP-REAL 2 st p2=y & |.p2.|>=1
                                     by A2,A3,A5;
      case A17:q<>0.REAL 2 &
        not(q`2<=q`1 & -q`1<=q`2 or q`2>=q`1 & q`2<=-q`1);
        then A18:Sq_Circ.q=|[q`1/sqrt(1+(q`1/q`2)^2),q`2/sqrt(1+(q`1/q`2)^2)]|
                                 by JGRAPH_3:def 1;
        A19: (|[q`1/sqrt(1+(q`1/q`2)^2),q`2/sqrt(1+(q`1/q`2)^2)]|)`1
         = q`1/sqrt(1+(q`1/q`2)^2) &
         (|[q`1/sqrt(1+(q`1/q`2)^2),q`2/sqrt(1+(q`1/q`2)^2)]|)`2
         = q`2/sqrt(1+(q`1/q`2)^2) by EUCLID:56;
          0<=(q`1/q`2)^2 by SQUARE_1:72;
        then A20:1+(q`1/q`2)^2>0 by Th3;
        A21: q`2 <> 0 by A17,REAL_1:66;
        then A22: (q`2)^2>0 by SQUARE_1:74;
        A23: (q`1)^2+(q`2)^2 <>0 by A21,COMPLEX1:2;
          not(-1 <q`1 & q`1< 1) implies -1>=q`2 or q`2>=1
        proof assume A24: not(-1 <q`1 & q`1< 1);
            now per cases by A24;
          case A25: -1>=q`1;
           then q`2<= -1 or q`1<q`2 & -q`2<= --q`1 by A17,AXIOMS:22,REAL_1:50;
            then -q`2<= -1 or -1>=q`2 by A25,AXIOMS:22;
           hence -1>=q`2 or q`2>=1 by REAL_1:50;
          case A26: q`1>=1;
             --q`1<= -q`2 & q`2<=q`1 or q`2>=q`1 & q`2>= -q`1
                        by A17,REAL_1:50;
            then 1<= -q`2 or q`2>=q`1 & q`2>= -q`1 by A26,AXIOMS:22;
            then -1>= --q`2 or 1<=q`2 by A26,AXIOMS:22,REAL_1:50;
           hence -1>=q`2 or q`2>=1;
          end;
         hence -1>=q`2 or q`2>=1;
        end;
          then (q`2)^2>=1^2 by A3,JGRAPH_2:6;
          then A27: sqrt((q`2)^2)>=1 by SQUARE_1:59,83,94;
           A28: |.(|[q`1/sqrt(1+(q`1/q`2)^2),q`2/sqrt(1+(q`1/q`2)^2)]|).|^2
          =((q`1)/sqrt(1+(q`1/(q`2))^2))^2+(q`2/sqrt(1+(q`1/(q`2))^2))^2
                        by A19,JGRAPH_3:10
          .=(q`1)^2/(sqrt(1+(q`1/(q`2))^2))^2+(q`2/sqrt(1+(q`1/(q`2))^2))^2
                        by SQUARE_1:69
          .=(q`1)^2/(sqrt(1+(q`1/(q`2))^2))^2
                  +(q`2)^2/(sqrt(1+(q`1/(q`2))^2))^2 by SQUARE_1:69
          .=(q`1)^2/(1+(q`1/q`2)^2)+(q`2)^2/(sqrt(1+(q`1/q`2)^2))^2
                               by A20,SQUARE_1:def 4
          .=(q`1)^2/(1+(q`1/q`2)^2)+(q`2)^2/(1+(q`1/q`2)^2)
                           by A20,SQUARE_1:def 4
          .=((q`1)^2+(q`2)^2)/(1+(q`1/q`2)^2) by XCMPLX_1:63
          .=((q`1)^2+(q`2)^2)/(1+(q`1)^2/(q`2)^2) by SQUARE_1:69
          .=((q`1)^2+(q`2)^2)/((q`1)^2/(q`2)^2+(q`2)^2/(q`2)^2)
                                by A22,XCMPLX_1:60
          .=((q`1)^2+(q`2)^2)/(((q`1)^2+(q`2)^2)/(q`2)^2) by XCMPLX_1:63
          .=(q`2)^2*(((q`1)^2+(q`2)^2)/((q`1)^2+(q`2)^2)) by XCMPLX_1:82
          .=(q`2)^2*1 by A23,XCMPLX_1:60
          .=(q`2)^2;
            0<=|.(|[q`1/sqrt(1+(q`1/q`2)^2),q`2/sqrt(1+(q`1/q`2)^2)]|).|
                               by TOPRNS_1:26;
          then |.(|[q`1/sqrt(1+(q`1/q`2)^2),q`2/sqrt(1+(q`1/q`2)^2)]|).|>=1
                               by A27,A28,SQUARE_1:89;
         hence ex p2 being Point of TOP-REAL 2 st p2=y & |.p2.|>=1
                                     by A2,A3,A18;
      end;
     hence y in Cb by A1;
    end;
   let y be set;assume y in Cb;
     then consider p2 being Point of TOP-REAL 2 such that
     A29: p2=y & |.p2.|>=1 by A1;
     set q=p2;
       now per cases;
     case q=0.REAL 2;
       then |.q.|=0 by TOPRNS_1:24;
      hence contradiction by A29;
     case A30:q<>0.REAL 2 &
          (q`2<=q`1 & -q`1<=q`2 or q`2>=q`1 & q`2<=-q`1);
          set px=|[q`1*sqrt(1+(q`2/q`1)^2),q`2*sqrt(1+(q`2/q`1)^2)]|;
        A31: px`1 = q`1*sqrt(1+(q`2/q`1)^2) &
         px`2 = q`2*sqrt(1+(q`2/q`1)^2) by EUCLID:56;
          0<=(q`2/q`1)^2 by SQUARE_1:72;
        then 1+(q`2/q`1)^2>0 by Th3;
        then A32:sqrt(1+(q`2/q`1)^2)>0 by SQUARE_1:93;
          0<=(px`2/px`1)^2 by SQUARE_1:72;
        then A33:1+(px`2/px`1)^2>0 by Th3;
        A34:px`2/px`1=q`2/q`1 by A31,A32,XCMPLX_1:92;
        A35: q`1=q`1*sqrt(1+(q`2/q`1)^2)/(sqrt(1+(q`2/q`1)^2))by A32,XCMPLX_1:
90
          .=px`1/(sqrt(1+(q`2/q`1)^2)) by EUCLID:56;
        A36: q`2=q`2*sqrt(1+(q`2/q`1)^2)/(sqrt(1+(q`2/q`1)^2))by A32,XCMPLX_1:
90
          .=px`2/(sqrt(1+(q`2/q`1)^2)) by EUCLID:56;
        A37: |.q.|^2=q`1^2+q`2^2 by JGRAPH_3:10;
        A38: |.q.|^2>=1 by A29,SQUARE_1:59,77;
        A39:now assume px`1=0 & px`2=0;
          then A40:q`1*sqrt(1+(q`2/q`1)^2)=0 & q`2*sqrt(1+(q`2/q`1)^2)=0
                                    by EUCLID:56;
          then A41:q`1=0 by A32,XCMPLX_1:6;
            q`2=0 by A32,A40,XCMPLX_1:6;
         hence contradiction by A30,A41,EUCLID:57,58;
        end;
          q`2<=q`1 & -q`1<=q`2 or q`2>=q`1 &
        q`2*sqrt(1+(q`2/q`1)^2) <= (-q`1)*sqrt(1+(q`2/q`1)^2)
                             by A30,A32,AXIOMS:25;
        then q`2<=q`1 & (-q`1)*sqrt(1+(q`2/q`1)^2) <= q`2*sqrt(1+(q`2/q`1)^2)
           or px`2>=px`1 & px`2<=-px`1 by A31,A32,AXIOMS:25,XCMPLX_1:175;
        then A42:px`2<=px`1 & -px`1<=px`2 or px`2>=px`1 & px`2<=-px`1
                                by A31,A32,AXIOMS:25,XCMPLX_1:175;
    then A43:Sq_Circ.px=|[px`1/sqrt(1+(px`2/px`1)^2),px`2/sqrt(1+(px`2/px`1)^2)
]|
                                 by A39,JGRAPH_2:11,JGRAPH_3:def 1;
        A44:px`1/sqrt(1+(px`2/px`1)^2)=q`1 by A31,A32,A34,XCMPLX_1:90;
         px`2/sqrt(1+(px`2/px`1)^2)=q`2 by A31,A32,A34,XCMPLX_1:90;
        then A45:q=Sq_Circ.px by A43,A44,EUCLID:57;
        A46: dom Sq_Circ=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
          not px`1=0 by A39,A42;
        then A47: (px`1)^2>0 by SQUARE_1:74;
          (px`2)^2>=0 by SQUARE_1:72;
        then A48: (px`1)^2+(px`2)^2>0 by A47,Th3;
           (px`1)^2/(sqrt(1+(px`2/px`1)^2))^2+(px`2/sqrt(1+(px`2/px`1)^2))^2>=1
                            by A34,A35,A36,A37,A38,SQUARE_1:69;
      then (px`1)^2/(sqrt(1+(px`2/px`1)^2))^2+(px`2)^2/(sqrt(1+(px`2/px`1)^2))
^2>=1
                            by SQUARE_1:69;
         then (px`1)^2/(1+(px`2/px`1)^2)+(px`2)^2/(sqrt(1+(px`2/px`1)^2))^2>=1
                            by A33,SQUARE_1:def 4;
         then (px`1)^2/(1+(px`2/px`1)^2)+(px`2)^2/(1+(px`2/px`1)^2)>=1
                            by A33,SQUARE_1:def 4;
         then ((px`1)^2/(1+(px`2/px`1)^2)
          +(px`2)^2/(1+(px`2/px`1)^2))*(1+(px`2/px`1)^2)>=1 *(1+(px`2/px`1)^2)
                            by A33,AXIOMS:25;
         then (px`1)^2/(1+(px`2/px`1)^2)*(1+(px`2/px`1)^2)
           +(px`2)^2/(1+(px`2/px`1)^2)*(1+(px`2/px`1)^2)>=1 *(1+(px`2/px`1)^2)
                            by XCMPLX_1:8;
    then (px`1)^2+(px`2)^2/(1+(px`2/px`1)^2)*(1+(px`2/px`1)^2)
                       >=1 *(1+(px`2/px`1)^2) by A33,XCMPLX_1:88;
         then A49: (px`1)^2+(px`2)^2>=1 *(1+(px`2/px`1)^2) by A33,XCMPLX_1:88;
           1 *(1+(px`2/px`1)^2)
         =1+(px`2)^2/(px`1)^2 by SQUARE_1:69;
         then (px`1)^2+(px`2)^2-1>=1+(px`2)^2/(px`1)^2-1 by A49,REAL_1:49;
         then (px`1)^2+(px`2)^2-1>=(px`2)^2/(px`1)^2 by XCMPLX_1:26;
         then ((px`1)^2+(px`2)^2-1)*(px`1)^2>=((px`2)^2/(px`1)^2)*(px`1)^2
                            by A47,AXIOMS:25;
         then ((px`1)^2+(px`2)^2-1)*(px`1)^2>=(px`2)^2 by A47,XCMPLX_1:88;
         then ((px`1)^2+((px`2)^2-1))*(px`1)^2>=(px`2)^2 by XCMPLX_1:29;
         then (px`1)^2*(px`1)^2+((px`2)^2-1)*(px`1)^2>=(px`2)^2 by XCMPLX_1:8;
         then (px`1)^2*(px`1)^2+(px`1)^2*((px`2)^2-1)-(px`2)^2>=0
                             by SQUARE_1:12;
         then A50: 0<=(px`1)^2*(px`1)^2+((px`1)^2*(px`2)^2-(px`1)^2*1)-(px`2)^2
                                   by XCMPLX_1:40;
            (px`1)^2*(px`1)^2+((px`1)^2*(px`2)^2-(px`1)^2*1)-(px`2)^2
         = (px`1)^2*(px`1)^2+(px`1)^2*(px`2)^2-(px`1)^2*1-(px`2)^2
                                   by XCMPLX_1:29
         .= (px`1)^2*(px`1)^2-(px`1)^2*1+(px`1)^2*(px`2)^2-1 *(px`2)^2
                                   by XCMPLX_1:29
         .= (px`1)^2*((px`1)^2-1)+(px`1)^2*(px`2)^2-1 *(px`2)^2 by XCMPLX_1:40
         .= (px`1)^2*((px`1)^2-1)+((px`1)^2*(px`2)^2-1 *(px`2)^2) by XCMPLX_1:
29
         .= (px`1)^2*((px`1)^2-1)+((px`1)^2-1)*(px`2)^2 by XCMPLX_1:40
         .= ((px`1)^2-1)*((px`1)^2+(px`2)^2) by XCMPLX_1:8;
         then (px`1)^2-1>=0 & (px`1)^2+(px`2)^2>=0 by A48,A50,SQUARE_1:25;
         then (px`1)^2-1+1>=0+1 & ((px`1)^2+(px`2)^2)>=0 by REAL_1:55;
         then (px`1)^2>=1 by XCMPLX_1:27;
          then px`1>=1 or px`1<= -1 by JGRAPH_2:8,SQUARE_1:59;
        then px in Kb by A1;
      hence ex x being set st x in dom Sq_Circ & x in Kb & y=Sq_Circ.x
                               by A29,A45,A46;
     case A51:q<>0.REAL 2 & not(q`2<=q`1 & -q`1<=q`2 or q`2>=q`1 & q`2<=-q`1);
          set px=|[q`1*sqrt(1+(q`1/q`2)^2),q`2*sqrt(1+(q`1/q`2)^2)]|;
        A52:q<>0.REAL 2 & (q`1<=q`2 & -q`2<=q`1 or q`1>=q`2 & q`1<=-q`2)
                          by A51,JGRAPH_2:23;
        A53: px`2 = q`2*sqrt(1+(q`1/q`2)^2) &
         px`1 = q`1*sqrt(1+(q`1/q`2)^2) by EUCLID:56;
          0<=(q`1/q`2)^2 by SQUARE_1:72;
        then 1+(q`1/q`2)^2>0 by Th3;
        then A54:sqrt(1+(q`1/q`2)^2)>0 by SQUARE_1:93;
          0<=(px`1/px`2)^2 by SQUARE_1:72;
        then A55:1+(px`1/px`2)^2>0 by Th3;
        A56:px`1/px`2=q`1/q`2 by A53,A54,XCMPLX_1:92;
        A57: q`2=q`2*sqrt(1+(q`1/q`2)^2)/(sqrt(1+(q`1/q`2)^2))by A54,XCMPLX_1:
90
          .=px`2/(sqrt(1+(q`1/q`2)^2)) by EUCLID:56;
        A58: q`1=q`1*sqrt(1+(q`1/q`2)^2)/(sqrt(1+(q`1/q`2)^2))by A54,XCMPLX_1:
90
          .=px`1/(sqrt(1+(q`1/q`2)^2)) by EUCLID:56;
        A59: |.q.|^2=q`2^2+q`1^2 by JGRAPH_3:10;
        A60: |.q.|^2>=1 by A29,SQUARE_1:59,77;
        A61:now assume px`2=0 & px`1=0;
          then A62:q`2*sqrt(1+(q`1/q`2)^2)=0 & q`1*sqrt(1+(q`1/q`2)^2)=0
                                    by EUCLID:56;
          then q`1=0 by A54,XCMPLX_1:6;
         hence contradiction by A51,A54,A62,XCMPLX_1:6;
        end;
          q`1<=q`2 & -q`2<=q`1 or q`1>=q`2 &
        q`1*sqrt(1+(q`1/q`2)^2) <= (-q`2)*sqrt(1+(q`1/q`2)^2)
                             by A52,A54,AXIOMS:25;
        then q`1<=q`2 & (-q`2)*sqrt(1+(q`1/q`2)^2) <= q`1*sqrt(1+(q`1/q`2)^2)
           or px`1>=px`2 & px`1<=-px`2 by A53,A54,AXIOMS:25,XCMPLX_1:175;
        then A63:px`1<=px`2 & -px`2<=px`1 or px`1>=px`2 & px`1<=-px`2
                                by A53,A54,AXIOMS:25,XCMPLX_1:175;
    then A64:Sq_Circ.px=|[px`1/sqrt(1+(px`1/px`2)^2),px`2/sqrt(1+(px`1/px`2)^2)
]|
                                 by A61,JGRAPH_2:11,JGRAPH_3:14;
        A65:px`2/sqrt(1+(px`1/px`2)^2)=q`2 by A53,A54,A56,XCMPLX_1:90;
         px`1/sqrt(1+(px`1/px`2)^2)=q`1 by A53,A54,A56,XCMPLX_1:90;
        then A66:q=Sq_Circ.px by A64,A65,EUCLID:57;
        A67: dom Sq_Circ=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
          not px`2=0 by A61,A63;
        then A68: (px`2)^2>0 by SQUARE_1:74;
        A69: (px`1)^2>=0 by SQUARE_1:72;
           (px`2)^2/(sqrt(1+(px`1/px`2)^2))^2+(px`1/sqrt(1+(px`1/px`2)^2))^2>=1
                            by A56,A57,A58,A59,A60,SQUARE_1:69;
      then (px`2)^2/(sqrt(1+(px`1/px`2)^2))^2+(px`1)^2/(sqrt(1+(px`1/px`2)^2))
^2>=1
                            by SQUARE_1:69;
         then (px`2)^2/(1+(px`1/px`2)^2)+(px`1)^2/(sqrt(1+(px`1/px`2)^2))^2>=1
                            by A55,SQUARE_1:def 4;
         then (px`2)^2/(1+(px`1/px`2)^2)+(px`1)^2/(1+(px`1/px`2)^2)>=1
                            by A55,SQUARE_1:def 4;
         then ((px`2)^2/(1+(px`1/px`2)^2)
          +(px`1)^2/(1+(px`1/px`2)^2))*(1+(px`1/px`2)^2)>=1 *(1+(px`1/px`2)^2)
                            by A55,AXIOMS:25;
         then (px`2)^2/(1+(px`1/px`2)^2)*(1+(px`1/px`2)^2)
           +(px`1)^2/(1+(px`1/px`2)^2)*(1+(px`1/px`2)^2)>=1 *(1+(px`1/px`2)^2)
                            by XCMPLX_1:8;
    then (px`2)^2+(px`1)^2/(1+(px`1/px`2)^2)*(1+(px`1/px`2)^2)
                      >=1 *(1+(px`1/px`2)^2) by A55,XCMPLX_1:88;
         then A70: (px`2)^2+(px`1)^2>=1 *(1+(px`1/px`2)^2) by A55,XCMPLX_1:88;
           1 *(1+(px`1/px`2)^2)
         =1+(px`1)^2/(px`2)^2 by SQUARE_1:69;
         then (px`2)^2+(px`1)^2-1>=1+(px`1)^2/(px`2)^2-1 by A70,REAL_1:49;
         then (px`2)^2+(px`1)^2-1>=(px`1)^2/(px`2)^2 by XCMPLX_1:26;

         then ((px`2)^2+(px`1)^2-1)*(px`2)^2>=((px`1)^2/(px`2)^2)*(px`2)^2
                            by A68,AXIOMS:25;
         then ((px`2)^2+(px`1)^2-1)*(px`2)^2>=(px`1)^2 by A68,XCMPLX_1:88;
         then ((px`2)^2+((px`1)^2-1))*(px`2)^2>=(px`1)^2 by XCMPLX_1:29;
         then (px`2)^2*(px`2)^2+((px`1)^2-1)*(px`2)^2>=(px`1)^2 by XCMPLX_1:8;
         then (px`2)^2*(px`2)^2+(px`2)^2*((px`1)^2-1)-(px`1)^2>=0
                             by SQUARE_1:12;
         then A71: 0<=(px`2)^2*(px`2)^2+((px`2)^2*(px`1)^2-(px`2)^2*1)-(px`1)^2
                                   by XCMPLX_1:40;
            (px`2)^2*(px`2)^2+((px`2)^2*(px`1)^2-(px`2)^2*1)-(px`1)^2
         = (px`2)^2*(px`2)^2+(px`2)^2*(px`1)^2-(px`2)^2*1-(px`1)^2
                                   by XCMPLX_1:29
         .= (px`2)^2*(px`2)^2-(px`2)^2*1+(px`2)^2*(px`1)^2-1 *(px`1)^2
                                   by XCMPLX_1:29
         .= (px`2)^2*((px`2)^2-1)+(px`2)^2*(px`1)^2-1 *(px`1)^2 by XCMPLX_1:40
         .= (px`2)^2*((px`2)^2-1)+((px`2)^2*(px`1)^2-1 *(px`1)^2) by XCMPLX_1:
29
         .= (px`2)^2*((px`2)^2-1)+((px`2)^2-1)*(px`1)^2 by XCMPLX_1:40
         .= ((px`2)^2-1)*((px`2)^2+(px`1)^2) by XCMPLX_1:8;
         then ( (px`2)^2-1>=0 & (px`1)^2+(px`2)^2>=0 or
           (px`2)^2-1<=0 & (px`1)^2+(px`2)^2<=0) by A71,SQUARE_1:25;
         then (px`2)^2-1+1>=0+1 & ((px`1)^2+(px`2)^2)>=0 by A68,A69,Th3,REAL_1:
55;
         then (px`2)^2>=1^2 by SQUARE_1:59,XCMPLX_1:27;
          then px`2>=1 or px`2<= -1 by JGRAPH_2:8;
        then px in Kb by A1;
      hence ex x being set st x in dom Sq_Circ & x in Kb & y=Sq_Circ.x
                               by A29,A66,A67;
     end;
    hence thesis by FUNCT_1:def 12;
 end;

theorem for P0,P1,P01,P11,K0,K1,K01,K11 being Subset of TOP-REAL 2,
P,K being non empty compact Subset of TOP-REAL 2,
f being map of TOP-REAL 2,TOP-REAL 2 st
P= circle(0,0,1) &
P0= inside_of_circle(0,0,1) &
P1= outside_of_circle(0,0,1) &
P01= closed_inside_of_circle(0,0,1) &
P11= closed_outside_of_circle(0,0,1) &
K=rectangle(-1,1,-1,1) &
K0=inside_of_rectangle(-1,1,-1,1) &
K1=outside_of_rectangle(-1,1,-1,1) &
K01=closed_inside_of_rectangle(-1,1,-1,1) &
K11=closed_outside_of_rectangle(-1,1,-1,1) &
f=Sq_Circ
holds f.:K=P & f".:P=K & f.:K0=P0 & f".:P0=K0 & f.:K1=P1 & f".:P1=K1 &
f.:K01=P01 & f.:K11=P11 & f".:P01=K01 & f".:P11=K11
proof let P0,P1,P01,P11,K0,K1,K01,K11 be Subset of TOP-REAL 2,
P,K being non empty compact Subset of TOP-REAL 2,
f be map of TOP-REAL 2,TOP-REAL 2;
 assume A1: P= circle(0,0,1) & P0= inside_of_circle(0,0,1) &
  P1= outside_of_circle(0,0,1) &
  P01= closed_inside_of_circle(0,0,1) &
  P11= closed_outside_of_circle(0,0,1) &
  K=rectangle(-1,1,-1,1) &
  K0=inside_of_rectangle(-1,1,-1,1) & K1=outside_of_rectangle(-1,1,-1,1)&
  K01=closed_inside_of_rectangle(-1,1,-1,1) &
  K11=closed_outside_of_rectangle(-1,1,-1,1) &
  f=Sq_Circ;
  then A2: K0={p: -1 <p`1 & p`1< 1 & -1 <p`2 & p`2< 1} by Def2;
  A3: P0={p: |.p.| <1} by A1,Th33;
  A4: K01={p: -1 <=p`1 & p`1<= 1 & -1 <=p`2 & p`2<= 1} by A1,Def3;
  A5: P01={p: |.p.| <=1} by A1,Th33;
  A6: K1={p: not(-1 <=p`1 & p`1<= 1 & -1 <=p`2 & p`2<= 1)} by A1,Def4;
  A7: P1={p: |.p.| >1} by A1,Th33;
  A8: K11={p: not(-1 <p`1 & p`1< 1 & -1 <p`2 & p`2< 1)} by A1,Def5;
  A9: P11={p: |.p.| >=1} by A1,Th33;
  defpred P[Point of TOP-REAL 2] means
  $1`1=-1 & -1 <=$1`2 & $1`2<=1 or $1`2=1 & -1<=$1`1 & $1`1<=1 or
      $1`1=1 & -1 <=$1`2 & $1`2<=1 or $1`2=-1 & -1<=$1`1 & $1`1<=1;
  defpred Q[Point of TOP-REAL 2] means
  -1=$1`1 & -1<=$1`2 & $1`2<=1 or $1`1=1 & -1<=$1`2 & $1`2<=1
        or -1=$1`2 & -1<=$1`1 & $1`1<=1 or 1=$1`2 & -1<=$1`1 & $1`1<=1;
deffunc F(set)=$1;
  A10: for p being Element of TOP-REAL 2
  holds P[p] iff Q[p];
  A11: K= {F(p): P[p]} by A1,Def1
  .= {F(q):Q[q]}
                        from Fraenkel6'(A10);
  defpred Q[Point of TOP-REAL 2] means
  |.$1.|=1;
  defpred P[Point of TOP-REAL 2] means
  |.$1- |[0,0]| .|=1;
  A12: for p holds P[p] iff Q[p] by EUCLID:58,JORDAN2C:13;
  A13: P= {F(p): P[p]} by A1,Def6
  .= {F(p2) where p2 is Point of TOP-REAL 2: Q[p2]} from Fraenkel6'(A12);
  then A14: f.:K=P by A1,A11,JGRAPH_3:33;
  A16: dom f=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
  then A17: (f qua Function)".:P=K by A1,A14,Th22;
    f.:K0 = P0 by A1,A2,A3,Th34;
  then A18: (f qua Function)".:P0=K0 by A1,A16,Th22;
    f.:K1 = P1 by A1,A6,A7,Th35;
  then (f qua Function)".:P1=K1 by A1,A16,Th22;
 hence f.:K=P & f".:P=K & f.:K0=P0 & f".:P0=K0 & f.:K1=P1 & f".:P1=K1
                         by A1,A2,A3,A6,A7,A11,A13,A17,A18,Th34,Th35,
JGRAPH_3:33;
    f.:K01 = P01 by A1,A4,A5,Th36;
  then A19: (f qua Function)".:P01=K01 by A1,A16,Th22;
    f.:K11 = P11 by A1,A8,A9,Th37;
  then (f qua Function)".:P11=K11 by A1,A16,Th22;
 hence thesis by A1,A4,A5,A8,A9,A19,Th36,Th37;
end;

begin :: Order of Points on Rectangle

theorem Th39:
  for a,b,c,d being real number st a <= b & c <= d holds
   LSeg(|[ a,c ]|, |[ a,d ]|) = { p1 : p1`1 = a & p1`2 <= d & p1`2 >= c}
 & LSeg(|[ a,d ]|, |[ b,d ]|) = { p2 : p2`1 <= b & p2`1 >= a & p2`2 = d}
 & LSeg(|[ a,c ]|, |[ b,c ]|) = { q1 : q1`1 <= b & q1`1 >= a & q1`2 = c}
 & LSeg(|[ b,c ]|, |[ b,d ]|) = { q2 : q2`1 = b & q2`2 <= d & q2`2 >= c}
proof let a,b,c,d be real number;assume A1: a <= b & c <= d;
  set L1 = { p : p`1 = a & p`2 <= d & p`2 >= c},
      L2 = { p : p`1 <= b & p`1 >= a & p`2 = d},
      L3 = { p : p`1 <= b & p`1 >= a & p`2 = c},
      L4 = { p : p`1 = b & p`2 <= d & p`2 >= c};
  set p0 = |[ a,c ]|, p01 = |[ a,d ]|, p10 = |[ b,c ]|, p1 = |[ b,d ]|;
 A2: p01`1 = a & p01`2 = d by EUCLID:56;
 A3: p10`1 = b & p10`2 = c by EUCLID:56;
  thus L1 = LSeg(p0,p01)
   proof
    A4: L1 c= LSeg(p0,p01)
     proof
      let a2 be set; assume a2 in L1;
      then consider p such that A5: a2 = p and
      A6: p`1 = a & p`2 <= d & p`2 >= c;
        now per cases;
      case A7: d <>c;
      reconsider lambda = (p`2-c)/(d-c) as Real by XREAL_0:def 1;
        d>=c by A6,AXIOMS:22;
      then d>c by A7,REAL_1:def 5;
      then A8: d-c>0 by SQUARE_1:11;
        p`2-c>=0 by A6,SQUARE_1:12;
      then A9: lambda>=0 by A8,REAL_2:125;
        d-c>=p`2-c by A6,REAL_1:49;
      then (d-c)/(d-c)>=(p`2-c)/(d-c) by A8,REAL_1:73;
      then A10: 1>=lambda by A8,XCMPLX_1:60;
      A11: (1-lambda)*c+lambda*d
        =((d-c)/(d-c)- (p`2-c)/(d-c))*c+(p`2-c)/(d-c)*d by A8,XCMPLX_1:60
        .=(((d-c)- (p`2-c))/(d-c))*c+(p`2-c)/(d-c)*d by XCMPLX_1:121
        .=((d- p`2)/(d-c))*c+(p`2-c)/(d-c)*d by XCMPLX_1:22
        .=c*((d- p`2)/(d-c))+d*(p`2-c)/(d-c) by XCMPLX_1:75
        .=c*(d- p`2)/(d-c)+d*(p`2-c)/(d-c) by XCMPLX_1:75
        .=(c*d- c*p`2)/(d-c)+d*(p`2-c)/(d-c) by XCMPLX_1:40
        .=(c*d- c*p`2)/(d-c)+(d*p`2-d*c)/(d-c) by XCMPLX_1:40
        .=((c*d- c*p`2)+(d*p`2-d*c))/(d-c) by XCMPLX_1:63
        .=((d*p`2-c*d)+c*d- c*p`2)/(d-c) by XCMPLX_1:29
        .=(d*p`2- c*p`2)/(d-c) by XCMPLX_1:27
        .=(d- c)*p`2/(d-c) by XCMPLX_1:40
        .=p`2*((d- c)/(d-c)) by XCMPLX_1:75
        .=p`2*1 by A8,XCMPLX_1:60
        .=p`2;
        (1-lambda)*p0 + lambda*p01
       =|[(1-lambda)*a,(1-lambda)*c]| + lambda*(|[a,d]|)
                       by EUCLID:62
       .=|[(1-lambda)*a,(1-lambda)*c]| +(|[ lambda*a, lambda*d]|)
                       by EUCLID:62
       .=|[(1-lambda)*a+lambda*a,(1-lambda)*c+lambda*d]| by EUCLID:60
       .=|[((1-lambda)+lambda)*a,(1-lambda)*c+lambda*d]| by XCMPLX_1:8
       .=|[1*a,(1-lambda)*c+lambda*d]| by XCMPLX_1:27
       .= p by A6,A11,EUCLID:57;
       then a2 in {(1-r)*p0 + r*p01: 0 <= r & r <= 1 } by A5,A9,A10;
      hence a2 in LSeg(p0,p01) by TOPREAL1:def 4;
      case d =c;
       then A12: p`2=c by A6,AXIOMS:21;
      reconsider lambda = 0 as Real;
        (1-lambda)*p0 + lambda*p01 =
       |[(1-lambda)*a,(1-lambda)*c]| + lambda*(|[a,d]|)
                       by EUCLID:62
       .=|[(1-lambda)*a,(1-lambda)*c]| +(|[ lambda*a, lambda*d]|)
                       by EUCLID:62
       .=|[(1-lambda)*a+lambda*a,(1-lambda)*c+lambda*d]|
                       by EUCLID:60
       .= p by A6,A12,EUCLID:57;
       then a2 in {(1-r)*p0 + r*p01: 0 <= r & r <= 1 } by A5;
      hence a2 in LSeg(p0,p01) by TOPREAL1:def 4;
      end;
      hence a2 in LSeg(p0,p01);
     end;
      LSeg(p0,p01) c= L1
     proof
      let a2 be set; assume a2 in LSeg(p0,p01);
      then a2 in {(1-lambda)*p0 + lambda*p01: 0 <= lambda & lambda <= 1 }
       by TOPREAL1:def 4;
      then consider lambda such that
       A13: a2 = (1-lambda)*p0 + lambda*p01 & 0 <= lambda & lambda <= 1;
      set q = (1-lambda)*p0 + lambda*p01;
      A14: q`1= ((1-lambda)*p0)`1 + (lambda*p01)`1 by TOPREAL3:7
         .= (1-lambda)*(p0)`1 + (lambda*p01)`1 by TOPREAL3:9
         .= (1-lambda)*(p0)`1 + lambda*(p01)`1 by TOPREAL3:9
         .=(1-lambda)*a +lambda*a by A2,EUCLID:56
         .=((1-lambda) +lambda)*a by XCMPLX_1:8
         .=1*a by XCMPLX_1:27 .=a;
         q`2= ((1-lambda)*p0)`2 + (lambda*p01)`2 by TOPREAL3:7
         .= (1-lambda)*(p0)`2 + (lambda*p01)`2 by TOPREAL3:9
         .= (1-lambda)*(p0)`2 + lambda*(p01)`2 by TOPREAL3:9
         .= (1-lambda)*c + lambda*d by A2,EUCLID:56;
      then q`1 = a & q`2 <= d & q`2 >= c by A1,A13,A14,Th2;
      hence a2 in L1 by A13;
     end;
    hence thesis by A4,XBOOLE_0:def 10;
   end;
  thus L2 = LSeg(p01,p1)
   proof
    A15: L2 c= LSeg(p01,p1)
     proof
      let a2 be set; assume a2 in L2;
      then consider p such that A16: a2 = p and
      A17: p`1 <= b & p`1 >= a & p`2=d;
        now per cases;
      case A18: b <>a;
      reconsider lambda = (p`1-a)/(b-a) as Real by XREAL_0:def 1;
        b>=a by A17,AXIOMS:22;
      then b>a by A18,REAL_1:def 5;
      then A19: b-a>0 by SQUARE_1:11;
        p`1-a>=0 by A17,SQUARE_1:12;
      then A20: lambda>=0 by A19,REAL_2:125;
        b-a>=p`1-a by A17,REAL_1:49;
      then (b-a)/(b-a)>=(p`1-a)/(b-a) by A19,REAL_1:73;
      then A21: 1>=lambda by A19,XCMPLX_1:60;
      A22: (1-lambda)*a+lambda*b
        =((b-a)/(b-a)- (p`1-a)/(b-a))*a+(p`1-a)/(b-a)*b by A19,XCMPLX_1:60
        .=(((b-a)- (p`1-a))/(b-a))*a+(p`1-a)/(b-a)*b by XCMPLX_1:121
        .=((b- p`1)/(b-a))*a+(p`1-a)/(b-a)*b by XCMPLX_1:22
        .=a*((b- p`1)/(b-a))+b*(p`1-a)/(b-a) by XCMPLX_1:75
        .=a*(b- p`1)/(b-a)+b*(p`1-a)/(b-a) by XCMPLX_1:75
        .=(a*b- a*p`1)/(b-a)+b*(p`1-a)/(b-a) by XCMPLX_1:40
        .=(a*b- a*p`1)/(b-a)+(b*p`1-b*a)/(b-a) by XCMPLX_1:40
        .=((a*b- a*p`1)+(b*p`1-b*a))/(b-a) by XCMPLX_1:63
        .=((b*p`1-a*b)+a*b- a*p`1)/(b-a) by XCMPLX_1:29
        .=(b*p`1- a*p`1)/(b-a) by XCMPLX_1:27
        .=(b- a)*p`1/(b-a) by XCMPLX_1:40
        .=p`1*((b- a)/(b-a)) by XCMPLX_1:75
        .=p`1*1 by A19,XCMPLX_1:60
        .=p`1;
        (1-lambda)*p01 + lambda*p1
       =|[(1-lambda)*a,(1-lambda)*d]| + lambda*(|[b,d]|)
                       by EUCLID:62
       .=|[(1-lambda)*a,(1-lambda)*d]| +(|[ lambda*b, lambda*d]|)
                       by EUCLID:62
       .=|[(1-lambda)*a+lambda*b,(1-lambda)*d+lambda*d]|
                       by EUCLID:60
       .=|[(1-lambda)*a+lambda*b,((1-lambda)+lambda)*d]|
                       by XCMPLX_1:8
       .=|[(1-lambda)*a+lambda*b,1*d]|
                       by XCMPLX_1:27
       .= p by A17,A22,EUCLID:57;
       then a2 in {(1-r)*p01 + r*p1: 0 <= r & r <= 1 } by A16,A20,A21;
      hence a2 in LSeg(p01,p1) by TOPREAL1:def 4;
      case b =a;
       then A23: p`1=a by A17,AXIOMS:21;
      reconsider lambda = 0 as Real;
        (1-lambda)*p01 + lambda*p1
       =|[(1-lambda)*a,(1-lambda)*d]| + lambda*(|[b,d]|)
                       by EUCLID:62
       .=|[(1-lambda)*a,(1-lambda)*d]| +(|[ lambda*b, lambda*d]|)
                       by EUCLID:62
       .=|[(1-lambda)*a+lambda*b,(1-lambda)*d+lambda*d]|
                       by EUCLID:60
       .= p by A17,A23,EUCLID:57;
       then a2 in {(1-r)*p01 + r*p1: 0 <= r & r <= 1 } by A16;
      hence a2 in LSeg(p01,p1) by TOPREAL1:def 4;
      end;
      hence a2 in LSeg(p01,p1);
     end;
      LSeg(p01,p1) c= L2
     proof
      let a2 be set; assume a2 in LSeg(p01,p1);
      then a2 in {(1-lambda)*p01 + lambda*p1: 0 <= lambda & lambda <= 1 }
       by TOPREAL1:def 4;
      then consider lambda such that
       A24: a2 = (1-lambda)*p01 + lambda*p1 & 0 <= lambda & lambda <= 1;
      set q = (1-lambda)*p01 + lambda*p1;
      A25: q`2= ((1-lambda)*p01)`2 + (lambda*p1)`2 by TOPREAL3:7
         .= (1-lambda)*(p01)`2 + (lambda*p1)`2 by TOPREAL3:9
         .= (1-lambda)*(p01)`2 + lambda*(p1)`2 by TOPREAL3:9
         .=(1-lambda)*d +lambda*d by A2,EUCLID:56
         .=((1-lambda) +lambda)*d by XCMPLX_1:8
         .=1*d by XCMPLX_1:27 .=d;
         q`1= ((1-lambda)*p01)`1 + (lambda*p1)`1 by TOPREAL3:7
         .= (1-lambda)*(p01)`1 + (lambda*p1)`1 by TOPREAL3:9
         .= (1-lambda)*(p01)`1 + lambda*(p1)`1 by TOPREAL3:9
         .= (1-lambda)*a + lambda*b by A2,EUCLID:56;
      then q`2 = d & q`1 <= b & q`1 >= a by A1,A24,A25,Th2;
      hence a2 in L2 by A24;
     end;
    hence thesis by A15,XBOOLE_0:def 10;
   end;
 thus L3 = LSeg(p0,p10)
   proof
    A26: L3 c= LSeg(p0,p10)
     proof
      let a2 be set; assume a2 in L3;
      then consider p such that A27: a2 = p and
      A28: p`1 <= b & p`1 >= a & p`2=c;
        now per cases;
      case A29: b <>a;
      reconsider lambda = (p`1-a)/(b-a) as Real by XREAL_0:def 1;
        b>=a by A28,AXIOMS:22;
      then b>a by A29,REAL_1:def 5;
      then A30: b-a>0 by SQUARE_1:11;
        p`1-a>=0 by A28,SQUARE_1:12;
      then A31: lambda>=0 by A30,REAL_2:125;
        b-a>=p`1-a by A28,REAL_1:49;
      then (b-a)/(b-a)>=(p`1-a)/(b-a) by A30,REAL_1:73;
      then A32: 1>=lambda by A30,XCMPLX_1:60;
      A33: (1-lambda)*a+lambda*b
        =((b-a)/(b-a)- (p`1-a)/(b-a))*a+(p`1-a)/(b-a)*b by A30,XCMPLX_1:60
        .=(((b-a)- (p`1-a))/(b-a))*a+(p`1-a)/(b-a)*b by XCMPLX_1:121
        .=((b- p`1)/(b-a))*a+(p`1-a)/(b-a)*b by XCMPLX_1:22
        .=a*((b- p`1)/(b-a))+b*(p`1-a)/(b-a) by XCMPLX_1:75
        .=a*(b- p`1)/(b-a)+b*(p`1-a)/(b-a) by XCMPLX_1:75
        .=(a*b- a*p`1)/(b-a)+b*(p`1-a)/(b-a) by XCMPLX_1:40
        .=(a*b- a*p`1)/(b-a)+(b*p`1-b*a)/(b-a) by XCMPLX_1:40
        .=((a*b- a*p`1)+(b*p`1-b*a))/(b-a) by XCMPLX_1:63
        .=((b*p`1-a*b)+a*b- a*p`1)/(b-a) by XCMPLX_1:29
        .=(b*p`1- a*p`1)/(b-a) by XCMPLX_1:27
        .=(b- a)*p`1/(b-a) by XCMPLX_1:40
        .=p`1*((b- a)/(b-a)) by XCMPLX_1:75
        .=p`1*1 by A30,XCMPLX_1:60
        .=p`1;
        (1-lambda)*p0 + lambda*p10
       =|[(1-lambda)*a,(1-lambda)*c]| + lambda*(|[b,c]|) by EUCLID:62
       .=|[(1-lambda)*a,(1-lambda)*c]| +(|[ lambda*b, lambda*c]|)
                       by EUCLID:62
       .=|[(1-lambda)*a+lambda*b,(1-lambda)*c+lambda*c]| by EUCLID:60
       .=|[(1-lambda)*a+lambda*b,((1-lambda)+lambda)*c]| by XCMPLX_1:8
       .=|[(1-lambda)*a+lambda*b,1*c]| by XCMPLX_1:27
       .= p by A28,A33,EUCLID:57;
       then a2 in {(1-r)*p0 + r*p10: 0 <= r & r <= 1 } by A27,A31,A32;
      hence a2 in LSeg(p0,p10) by TOPREAL1:def 4;
      case b =a;
       then A34: p`1=a by A28,AXIOMS:21;
      reconsider lambda = 0 as Real;
        (1-lambda)*p0 + lambda*p10 =
|[(1-lambda)*a,(1-lambda)*c]| + lambda*(|[b,c]|)
                       by EUCLID:62
       .=|[(1-lambda)*a,(1-lambda)*c]| +(|[ lambda*b, lambda*c]|)
                       by EUCLID:62
       .=|[(1-lambda)*a+lambda*b,(1-lambda)*c+lambda*c]|
                       by EUCLID:60
       .= p by A28,A34,EUCLID:57;
       then a2 in {(1-r)*p0 + r*p10: 0 <= r & r <= 1 } by A27;
      hence a2 in LSeg(p0,p10) by TOPREAL1:def 4;
      end;
      hence a2 in LSeg(p0,p10);
     end;
      LSeg(p0,p10) c= L3
     proof
      let a2 be set; assume a2 in LSeg(p0,p10);
      then a2 in {(1-lambda)*p0 + lambda*p10: 0 <= lambda & lambda <= 1 }
       by TOPREAL1:def 4;
      then consider lambda such that
       A35: a2 = (1-lambda)*p0 + lambda*p10 & 0 <= lambda & lambda <= 1;
      set q = (1-lambda)*p0 + lambda*p10;
      A36: q`2= ((1-lambda)*p0)`2 + (lambda*p10)`2 by TOPREAL3:7
         .= (1-lambda)*(p0)`2 + (lambda*p10)`2 by TOPREAL3:9
         .= (1-lambda)*(p0)`2 + lambda*(p10)`2 by TOPREAL3:9
         .=(1-lambda)*c +lambda*c by A3,EUCLID:56
         .=((1-lambda) +lambda)*c by XCMPLX_1:8
         .=1*c by XCMPLX_1:27 .=c;
         q`1= ((1-lambda)*p0)`1 + (lambda*p10)`1 by TOPREAL3:7
         .= (1-lambda)*(p0)`1 + (lambda*p10)`1 by TOPREAL3:9
         .= (1-lambda)*(p0)`1 + lambda*(p10)`1 by TOPREAL3:9
         .= (1-lambda)*a + lambda*b by A3,EUCLID:56;
      then q`2 = c & q`1 <= b & q`1 >= a by A1,A35,A36,Th2;
      hence a2 in L3 by A35;
     end;
    hence thesis by A26,XBOOLE_0:def 10;
   end;
 thus L4 = LSeg(p10,p1)
   proof
    A37: L4 c= LSeg(p10,p1)
     proof
      let a2 be set; assume a2 in L4;
      then consider p such that A38: a2 = p and
      A39: p`1 = b & p`2 <= d & p`2 >= c;
        now per cases;
      case A40: d <>c;
      reconsider lambda = (p`2-c)/(d-c) as Real by XREAL_0:def 1;
        d>=c by A39,AXIOMS:22;
      then d>c by A40,REAL_1:def 5;
      then A41: d-c>0 by SQUARE_1:11;
        p`2-c>=0 by A39,SQUARE_1:12;
      then A42: lambda>=0 by A41,REAL_2:125;
        d-c>=p`2-c by A39,REAL_1:49;
      then (d-c)/(d-c)>=(p`2-c)/(d-c) by A41,REAL_1:73;
      then A43: 1>=lambda by A41,XCMPLX_1:60;
      A44: (1-lambda)*c+lambda*d
        =((d-c)/(d-c)- (p`2-c)/(d-c))*c+(p`2-c)/(d-c)*d by A41,XCMPLX_1:60
        .=(((d-c)- (p`2-c))/(d-c))*c+(p`2-c)/(d-c)*d by XCMPLX_1:121
        .=((d- p`2)/(d-c))*c+(p`2-c)/(d-c)*d by XCMPLX_1:22
        .=c*((d- p`2)/(d-c))+d*(p`2-c)/(d-c) by XCMPLX_1:75
        .=c*(d- p`2)/(d-c)+d*(p`2-c)/(d-c) by XCMPLX_1:75
        .=(c*d- c*p`2)/(d-c)+d*(p`2-c)/(d-c) by XCMPLX_1:40
        .=(c*d- c*p`2)/(d-c)+(d*p`2-d*c)/(d-c) by XCMPLX_1:40
        .=((c*d- c*p`2)+(d*p`2-d*c))/(d-c) by XCMPLX_1:63
        .=((d*p`2-c*d)+c*d- c*p`2)/(d-c) by XCMPLX_1:29
        .=(d*p`2- c*p`2)/(d-c) by XCMPLX_1:27
        .=(d- c)*p`2/(d-c) by XCMPLX_1:40
        .=p`2*((d- c)/(d-c)) by XCMPLX_1:75
        .=p`2*1 by A41,XCMPLX_1:60
        .=p`2;
        (1-lambda)*p10 + lambda*p1
       =|[(1-lambda)*b,(1-lambda)*c]| + lambda*(|[b,d]|)
                       by EUCLID:62
       .=|[(1-lambda)*b,(1-lambda)*c]| +(|[ lambda*b, lambda*d]|)
                       by EUCLID:62
       .=|[(1-lambda)*b+lambda*b,(1-lambda)*c+lambda*d]|
                       by EUCLID:60
       .=|[((1-lambda)+lambda)*b,(1-lambda)*c+lambda*d]|
                       by XCMPLX_1:8
       .=|[1*b,(1-lambda)*c+lambda*d]|
                       by XCMPLX_1:27
       .= p by A39,A44,EUCLID:57;
       then a2 in {(1-r)*p10 + r*p1: 0 <= r & r <= 1 } by A38,A42,A43;
      hence a2 in LSeg(p10,p1) by TOPREAL1:def 4;
      case d =c;
       then A45: p`2=c by A39,AXIOMS:21;
      reconsider lambda = 0 as Real;
        (1-lambda)*p10 + lambda*p1
       =|[(1-lambda)*b,(1-lambda)*c]| + lambda*(|[b,d]|)
                       by EUCLID:62
       .=|[(1-lambda)*b,(1-lambda)*c]| +(|[ lambda*b, lambda*d]|)
                       by EUCLID:62
       .=|[(1-lambda)*b+lambda*b,(1-lambda)*c+lambda*d]|
                       by EUCLID:60
       .= p by A39,A45,EUCLID:57;
       then a2 in {(1-r)*p10 + r*p1: 0 <= r & r <= 1 } by A38;
      hence a2 in LSeg(p10,p1) by TOPREAL1:def 4;
      end;
      hence a2 in LSeg(p10,p1);
     end;
      LSeg(p10,p1) c= L4
     proof
      let a2 be set; assume a2 in LSeg(p10,p1);
      then a2 in {(1-lambda)*p10 + lambda*p1: 0 <= lambda & lambda <= 1 }
       by TOPREAL1:def 4;
      then consider lambda such that
       A46: a2 = (1-lambda)*p10 + lambda*p1 & 0 <= lambda & lambda <= 1;
      set q = (1-lambda)*p10 + lambda*p1;
      A47: q`1= ((1-lambda)*p10)`1 + (lambda*p1)`1 by TOPREAL3:7
         .= (1-lambda)*(p10)`1 + (lambda*p1)`1 by TOPREAL3:9
         .= (1-lambda)*(p10)`1 + lambda*(p1)`1 by TOPREAL3:9
         .=(1-lambda)*b +lambda*b by A3,EUCLID:56
         .=((1-lambda) +lambda)*b by XCMPLX_1:8
         .=1*b by XCMPLX_1:27 .=b;
         q`2= ((1-lambda)*p10)`2 + (lambda*p1)`2 by TOPREAL3:7
         .= (1-lambda)*(p10)`2 + (lambda*p1)`2 by TOPREAL3:9
         .= (1-lambda)*(p10)`2 + lambda*(p1)`2 by TOPREAL3:9
         .= (1-lambda)*c + lambda*d by A3,EUCLID:56;
      then q`1 = b & q`2 <= d & q`2 >= c by A1,A46,A47,Th2;
      hence a2 in L4 by A46;
     end;
    hence thesis by A37,XBOOLE_0:def 10;
   end;
 end;

theorem Th40: for a,b,c,d being real number st a<=b & c <=d holds
 {p: p`1=a & c <=p`2 & p`2<=d or p`2=d & a<=p`1 & p`1<=b or
      p`1=b & c <=p`2 & p`2<=d or p`2=c & a<=p`1 & p`1<=b}
 = (LSeg(|[ a,c ]|, |[ a,d ]|) \/ LSeg(|[ a,d ]|, |[ b,d ]|))
 \/ (LSeg(|[ a,c ]|, |[ b,c ]|) \/ LSeg(|[ b,c ]|, |[ b,d ]|))
proof let a,b,c,d be real number;
 assume A1: a<=b & c <=d;
  set L1 = { p : p`1 = a & p`2 <= d & p`2 >= c},
      L2 = { p : p`1 <= b & p`1 >= a & p`2 = d},
      L3 = { p : p`1 <= b & p`1 >= a & p`2 = c},
      L4 = { p : p`1 = b & p`2 <= d & p`2 >= c};
defpred P1[Point of TOP-REAL 2] means $1`1=a & c <=$1`2 & $1`2<=d;
defpred P2[Point of TOP-REAL 2] means $1`2=d & a<=$1`1 & $1`1 <= b;
defpred Q1[Point of TOP-REAL 2] means $1`1=b & c <=$1`2 & $1`2<=d;
defpred Q2[Point of TOP-REAL 2] means $1`2=c & a<=$1`1 & $1`1<=b;
set M1 = { p : P1[p]},
    M2 = { p : P2[p]},
    M3 = { p : Q2[p]},
    M4 = { p : Q1[p]};
  A2: L1 = M1
   proof
    thus L1 c= M1
     proof let x be set;assume x in L1;
       then consider p such that
       A3: x=p & p`1 = a & p`2 <= d & p`2 >= c;
      thus x in M1 by A3;
     end;
    thus M1 c= L1
     proof let x be set;assume x in M1;
       then consider p such that
       A4: x=p & p`1 = a & c <=p`2 & p`2 <= d;
      thus x in L1 by A4;
     end;
   end;
  A5: L2 = M2
   proof
    thus L2 c= M2
     proof let x be set;assume x in L2;
       then consider p such that
       A6: x=p & p`1 <= b & p`1 >= a & p`2 = d;
      thus x in M2 by A6;
     end;
    thus M2 c= L2
     proof let x be set;assume x in M2;
       then consider p such that
       A7: x=p & p`2=d & a<=p`1 & p`1 <= b;
      thus x in L2 by A7;
     end;
   end;
  A8: L3 = M3
   proof
    thus L3 c= M3
     proof let x be set;assume x in L3;
       then consider p such that
       A9: x=p & p`1 <= b & p`1 >= a & p`2 = c;
      thus x in M3 by A9;
     end;
    thus M3 c= L3
     proof let x be set;assume x in M3;
       then consider p such that
       A10: x=p & p`2=c & a<=p`1 & p`1 <= b;
      thus x in L3 by A10;
     end;
   end;
  A11: L4 = M4
   proof
    thus L4 c= M4
     proof let x be set;assume x in L4;
       then consider p such that
       A12: x=p & p`1 = b & p`2 <= d & p`2 >= c;
      thus x in M4 by A12;
     end;
    thus M4 c= L4
     proof let x be set;assume x in M4;
       then consider p such that
       A13: x=p & p`1 = b & c <=p`2 & p`2 <= d;
      thus x in L4 by A13;
     end;
   end;
   defpred P[Point of TOP-REAL 2] means
   $1`1=a & c <=$1`2 & $1`2<=d or $1`2=d & a<=$1`1 & $1`1<=b;
   defpred Q[Point of TOP-REAL 2] means
   $1`1=b & c <=$1`2 & $1`2<=d or $1`2=c & a<=$1`1 & $1`1<=b;

 {p2: P[p2] or Q[p2]}
   =
 {p: P[p]} \/ {q1: Q[q1]} from Fraenkel_Alt; then
A14:
 {p2: p2`1=a & c <=p2`2 & p2`2<=d or p2`2=d & a<=p2`1 & p2`1<=b or
      p2`1=b & c <=p2`2 & p2`2<=d or p2`2=c & a<=p2`1 & p2`1<=b}
   =
 {p: p`1=a & c <=p`2 & p`2<=d or p`2=d & a<=p`1 & p`1<=b} \/
   {q1: q1`1=b & c <=q1`2 & q1`2<=d or q1`2=c & a<=q1`1 & q1`1<=b};
A15:
 {p: P1[p] or P2[p]}
      = M1 \/ M2 from Fraenkel_Alt
     .= LSeg(|[a,c]|,|[a,d]|) \/ L2 by A1,A2,A5,Th39
     .= LSeg(|[a,c]|,|[a,d]|) \/ LSeg(|[ a,d ]|, |[ b,d ]|) by A1,Th39;
     {q1: Q1[q1] or Q2[q1]}
      = M4 \/ M3 from Fraenkel_Alt
     .= LSeg(|[ a,c ]|, |[ b,c ]|) \/ L4 by A1,A8,A11,Th39
     .= LSeg(|[ a,c ]|, |[ b,c ]|) \/ LSeg(|[ b,c ]|, |[ b,d ]|) by A1,Th39;
 hence thesis by A14,A15;
end;

theorem Th41: for a,b,c,d being real number st a <= b & c <= d holds
 LSeg(|[a,c]|,|[a,d]|) /\ LSeg(|[a,c]|,|[b,c]|) = {|[a,c]|}
proof let a,b,c,d be real number; assume A1: a <= b & c <= d;
    for ax being set
   holds ax in LSeg(|[a,c]|,|[a,d]|) /\ LSeg(|[a,c]|,|[b,c]|) iff ax = |[a,c]|
  proof let ax be set;
   thus ax in LSeg(|[a,c]|,|[a,d]|) /\ LSeg(|[a,c]|,|[b,c]|)
             implies ax = |[a,c]|
    proof
     assume ax in LSeg(|[a,c]|,|[a,d]|) /\ LSeg(|[a,c]|,|[b,c]|);
     then A2: ax in LSeg(|[a,c]|,|[a,d]|)& ax in LSeg(|[a,c]|,|[b,c]|)
      by XBOOLE_0:def 3;
     then ax in { p2 : p2`1 <= b & p2`1 >= a & p2`2 = c }
                  by A1,Th39;
     then A3: ex p2 st p2 = ax & p2`1 <= b & p2`1 >= a & p2`2 = c;
       ax in { p2 : p2`1 = a & p2`2 <= d & p2`2 >= c }
                  by A1,A2,Th39;
     then ex p st p = ax & p`1 = a & p`2 <= d & p`2 >= c;
     hence ax = |[a,c]| by A3,EUCLID:57;
    end;
   assume ax = |[a,c]|;
   then ax in LSeg(|[a,c]|,|[a,d]|) & ax in LSeg(|[a,c]|,|[b,c]|)
                     by TOPREAL1:6;
   hence ax in LSeg(|[a,c]|,|[a,d]|) /\ LSeg(|[a,c]|,|[b,c]|)
                   by XBOOLE_0:def 3;
  end;
 hence thesis by TARSKI:def 1;
end;

theorem Th42: for a,b,c,d being real number st a <= b & c <= d holds
 LSeg(|[a,c]|,|[b,c]|) /\ LSeg(|[b,c]|,|[b,d]|) = {|[b,c]|}
proof let a,b,c,d be real number; assume A1: a <= b & c <= d;
    for ax being set
   holds ax in LSeg(|[a,c]|,|[b,c]|) /\ LSeg(|[b,c]|,|[b,d]|) iff ax = |[b,c]|
  proof let ax be set;
   thus ax in LSeg(|[a,c]|,|[b,c]|) /\ LSeg(|[b,c]|,|[b,d]|)
             implies ax = |[b,c]|
    proof
     assume ax in LSeg(|[a,c]|,|[b,c]|) /\ LSeg(|[b,c]|,|[b,d]|);
     then A2: ax in LSeg(|[a,c]|,|[b,c]|)& ax in LSeg(|[b,c]|,|[b,d]|)
      by XBOOLE_0:def 3;
     then ax in { q1 : q1`1 <= b & q1`1 >= a & q1`2 = c}
                  by A1,Th39;
     then A3: ex p2 st p2 = ax & p2`1 <= b & p2`1 >= a & p2`2 = c;
       ax in { q2 : q2`1 = b & q2`2 <= d & q2`2 >= c}
                  by A1,A2,Th39;
     then ex p st p = ax & p`1 = b & p`2 <= d & p`2 >= c;
     hence ax = |[b,c]| by A3,EUCLID:57;
    end;
   assume ax = |[b,c]|;
   then ax in LSeg(|[a,c]|,|[b,c]|) & ax in LSeg(|[b,c]|,|[b,d]|)
                     by TOPREAL1:6;
   hence ax in LSeg(|[a,c]|,|[b,c]|) /\ LSeg(|[b,c]|,|[b,d]|)
                   by XBOOLE_0:def 3;
  end;
 hence thesis by TARSKI:def 1;
end;

theorem Th43: for a,b,c,d being real number st a <= b & c <= d holds
 LSeg(|[a,d]|,|[b,d]|) /\ LSeg(|[b,c]|,|[b,d]|) = {|[b,d]|}
 proof let a,b,c,d be real number;assume A1: a <= b & c <= d;
    for ax being set
   holds ax in LSeg(|[a,d]|,|[b,d]|) /\ LSeg(|[b,c]|,|[b,d]|) iff ax = |[b,d]|
  proof let ax be set;
   thus ax in LSeg(|[a,d]|,|[b,d]|) /\ LSeg(|[b,c]|,|[b,d]|)
      implies ax = |[b,d]|
    proof
     assume ax in LSeg(|[a,d]|,|[b,d]|) /\ LSeg(|[b,c]|,|[b,d]|);
     then A2: ax in LSeg(|[b,c]|,|[b,d]|) & ax in LSeg(|[a,d]|,|[b,d]|)
      by XBOOLE_0:def 3;
     then ax in { p : p`1 <= b & p`1 >= a & p`2 = d} by A1,Th39;
     then A3: ex p st p = ax & p`1 <= b & p`1 >= a & p`2 = d;
       ax in { p : p`1 = b & p`2 <= d & p`2 >= c} by A1,A2,Th39;
     then ex p2 st p2 = ax & p2`1 = b & p2`2 <= d & p2`2 >= c;
     hence ax = |[b,d]| by A3,EUCLID:57;
    end;
   assume ax = |[b,d]|;
   then ax in LSeg(|[a,d]|,|[b,d]|) & ax in LSeg(|[b,c]|,|[b,d]|)
                        by TOPREAL1:6;
   hence ax in LSeg(|[a,d]|,|[b,d]|) /\ LSeg(|[b,c]|,|[b,d]|)
                        by XBOOLE_0:def 3;
  end;
  hence thesis by TARSKI:def 1;
 end;

theorem Th44: for a,b,c,d being real number st a <= b & c <= d holds
 LSeg(|[a,c]|,|[a,d]|) /\ LSeg(|[a,d]|,|[b,d]|) = {|[a,d]|}
proof let a,b,c,d be real number; assume A1: a <= b & c <= d;
    for ax being set
   holds ax in LSeg(|[a,c]|,|[a,d]|) /\ LSeg(|[a,d]|,|[b,d]|) iff ax = |[a,d]|
  proof let ax be set;
   thus ax in LSeg(|[a,c]|,|[a,d]|) /\ LSeg(|[a,d]|,|[b,d]|)
             implies ax = |[a,d]|
    proof
     assume ax in LSeg(|[a,c]|,|[a,d]|) /\ LSeg(|[a,d]|,|[b,d]|);
     then A2: ax in LSeg(|[a,c]|,|[a,d]|)& ax in LSeg(|[a,d]|,|[b,d]|)
      by XBOOLE_0:def 3;
     then ax in { p2 : p2`1 <= b & p2`1 >= a & p2`2 = d }
                  by A1,Th39;
     then A3: ex p2 st p2 = ax & p2`1 <= b & p2`1 >= a & p2`2 = d;
       ax in { p2 : p2`1 = a & p2`2 <= d & p2`2 >= c }
                  by A1,A2,Th39;
     then ex p st p = ax & p`1 = a & p`2 <= d & p`2 >= c;
     hence ax = |[a,d]| by A3,EUCLID:57;
    end;
   assume ax = |[a,d]|;
   then ax in LSeg(|[a,c]|,|[a,d]|) & ax in LSeg(|[a,d]|,|[b,d]|)
                     by TOPREAL1:6;
   hence ax in LSeg(|[a,c]|,|[a,d]|) /\ LSeg(|[a,d]|,|[b,d]|)
                   by XBOOLE_0:def 3;
  end;
 hence thesis by TARSKI:def 1;
end;

theorem Th45:
 {q: -1=q`1 & -1<=q`2 & q`2<=1 or q`1=1 & -1<=q`2 & q`2<=1
     or -1=q`2 & -1<=q`1 & q`1<=1 or 1=q`2 & -1<=q`1 & q`1<=1}
 = {p: p`1=-1 & -1 <=p`2 & p`2<=1 or p`2=1 & -1<=p`1 & p`1<=1 or
      p`1=1 & -1 <=p`2 & p`2<=1 or p`2=-1 & -1<=p`1 & p`1<=1}
 proof
  thus {q: -1=q`1 & -1<=q`2 & q`2<=1 or q`1=1 & -1<=q`2 & q`2<=1
     or -1=q`2 & -1<=q`1 & q`1<=1 or 1=q`2 & -1<=q`1 & q`1<=1}
   c= {p: p`1=-1 & -1 <=p`2 & p`2<=1 or p`2=1 & -1<=p`1 & p`1<=1 or
      p`1=1 & -1 <=p`2 & p`2<=1 or p`2=-1 & -1<=p`1 & p`1<=1}
   proof let x be set;assume x in
    {q: -1=q`1 & -1<=q`2 & q`2<=1 or q`1=1 & -1<=q`2 & q`2<=1
      or -1=q`2 & -1<=q`1 & q`1<=1 or 1=q`2 & -1<=q`1 & q`1<=1};
    then consider q such that
    A1: x=q & ( -1=q`1 & -1<=q`2 & q`2<=1 or q`1=1 & -1<=q`2 & q`2<=1
      or -1=q`2 & -1<=q`1 & q`1<=1 or 1=q`2 & -1<=q`1 & q`1<=1);
    thus x in
     {p: p`1=-1 & -1 <=p`2 & p`2<=1 or p`2=1 & -1<=p`1 & p`1<=1 or
      p`1=1 & -1 <=p`2 & p`2<=1 or p`2=-1 & -1<=p`1 & p`1<=1} by A1;
   end;
  thus {p: p`1=-1 & -1 <=p`2 & p`2<=1 or p`2=1 & -1<=p`1 & p`1<=1 or
      p`1=1 & -1 <=p`2 & p`2<=1 or p`2=-1 & -1<=p`1 & p`1<=1}
  c= {q: -1=q`1 & -1<=q`2 & q`2<=1 or q`1=1 & -1<=q`2 & q`2<=1
     or -1=q`2 & -1<=q`1 & q`1<=1 or 1=q`2 & -1<=q`1 & q`1<=1}
   proof let x be set;assume x in
     {p: p`1=-1 & -1 <=p`2 & p`2<=1 or p`2=1 & -1<=p`1 & p`1<=1 or
      p`1=1 & -1 <=p`2 & p`2<=1 or p`2=-1 & -1<=p`1 & p`1<=1};
     then consider p such that
     A2: p=x & (
      p`1=-1 & -1 <=p`2 & p`2<=1 or p`2=1 & -1<=p`1 & p`1<=1 or
      p`1=1 & -1 <=p`2 & p`2<=1 or p`2=-1 & -1<=p`1 & p`1<=1);
    thus x in
    {q: -1=q`1 & -1<=q`2 & q`2<=1 or q`1=1 & -1<=q`2 & q`2<=1
      or -1=q`2 & -1<=q`1 & q`1<=1 or 1=q`2 & -1<=q`1 & q`1<=1} by A2;
   end;
 end;

theorem Th46:
for K being non empty compact Subset of TOP-REAL 2,
a,b,c,d being real number st K=rectangle(a,b,c,d) & a<=b & c <=d
holds W-bound K = a
proof let K be non empty compact Subset of TOP-REAL 2,
a,b,c,d be real number;
assume A1: K=rectangle(a,b,c,d) & a<=b & c <=d;
  set X = K;
  reconsider Z = (proj1||X).:the carrier of ((TOP-REAL 2)|X)
                            as Subset of REAL;
A2: X = the carrier of ((TOP-REAL 2)|X) by JORDAN1:1;
A3: for p be real number st p in Z holds p >= a
    proof let p be real number;
     assume p in Z;
      then consider p0 being set such that
A4:    p0 in the carrier of (TOP-REAL 2)|X and
         p0 in the carrier of (TOP-REAL 2)|X and
A5:    p = (proj1||X).p0 by FUNCT_2:115;
      reconsider p0 as Point of TOP-REAL 2 by A2,A4;
        X= {q :
         q`1=a & c <=q`2 & q`2<=d or q`2=d & a<=q`1 & q`1<=b or
          q`1=b & c <=q`2 & q`2<=d or q`2=c & a<=q`1 & q`1<=b} by A1,Def1;
        then ex q being Point of TOP-REAL 2 st p0 = q &
        ( q`1=a & c <=q`2 & q`2<=d or q`2=d & a<=q`1 & q`1<=b or
          q`1=b & c <=q`2 & q`2<=d or q`2=c & a<=q`1 & q`1<=b)
                          by A2,A4;
     hence p >= a by A1,A2,A4,A5,PSCOMP_1:69;
    end;
A6: for q being real number st
   for p being real number st p in Z holds p >= q holds a >= q
   proof let q be real number such that
A7:    for p being real number st p in Z holds p >= q;
       |[a,c]| in LSeg(|[ a,c ]|, |[ b,c ]|) by TOPREAL1:6;
     then A8: |[a,c]| in LSeg(|[ a,c ]|, |[ b,c ]|) \/ LSeg(|[ b,c ]|, |[ b,d
]|)
                 by XBOOLE_0:def 2;
        X= {q2 :
         q2`1=a & c <=q2`2 & q2`2<=d or q2`2=d & a<=q2`1 & q2`1<=b or
          q2`1=b & c <=q2`2 & q2`2<=d or q2`2=c & a<=q2`1 & q2`1<=b}
                                 by A1,Def1;
       then X= (LSeg(|[ a,c ]|, |[ a,d ]|) \/ LSeg(|[ a,d ]|, |[ b,d ]|))
          \/ (LSeg(|[ a,c ]|, |[ b,c ]|) \/ LSeg(|[ b,c ]|, |[ b,d ]|))
                       by A1,Th40;
then A9:   |[a,c]| in X by A8,XBOOLE_0:def 2;
     then (proj1||X). |[a,c]| = |[a,c]|`1 by PSCOMP_1:69
             .= a by EUCLID:56;
     then a in Z by A2,A9,FUNCT_2:43;
    hence thesis by A7;
   end;
 thus W-bound X = inf (proj1||X) by PSCOMP_1:def 30
     .= inf Z by PSCOMP_1:def 20
     .= a by A3,A6,PSCOMP_1:9;
end;

theorem Th47:
 for K being non empty compact Subset of TOP-REAL 2,
 a,b,c,d being real number st K=rectangle(a,b,c,d) & a<=b & c <=d
holds N-bound K = d
proof let K be non empty compact Subset of TOP-REAL 2,
a,b,c,d be real number;
assume A1: K=rectangle(a,b,c,d) & a<=b & c <=d;
  set X = K;
  reconsider Z = (proj2||X).:the carrier of ((TOP-REAL 2)|X)
                                as Subset of REAL;
A2: X = the carrier of ((TOP-REAL 2)|X) by JORDAN1:1;
A3: for p be real number st p in Z holds p <= d
    proof let p be real number;
     assume p in Z;
      then consider p0 being set such that
A4:    p0 in the carrier of (TOP-REAL 2)|X and
         p0 in the carrier of (TOP-REAL 2)|X and
A5:    p = (proj2||X).p0 by FUNCT_2:115;
      reconsider p0 as Point of TOP-REAL 2 by A2,A4;
        X= {q :
        q`1=a & c <=q`2 & q`2<=d or q`2=d & a<=q`1 & q`1<=b or
        q`1=b & c <=q`2 & q`2<=d or q`2=c & a<=q`1 & q`1<=b} by A1,Def1;
        then ex q being Point of TOP-REAL 2 st p0 = q &
        ( q`1=a & c <=q`2 & q`2<=d or q`2=d & a<=q`1 & q`1<=b or
          q`1=b & c <=q`2 & q`2<=d or q`2=c & a<=q`1 & q`1<=b) by A2,A4;
     hence p <= d by A1,A2,A4,A5,PSCOMP_1:70;
    end;
A6: for q being real number st
    for p being real number
                st p in Z holds p <= q holds d <= q
   proof let q be real number such that
A7:    for p be real number st p in Z holds p <= q;
         |[b,d]| in LSeg(|[ b,c ]|, |[ b,d ]|) by TOPREAL1:6;
then A8:   |[b,d]| in LSeg(|[ a,c ]|, |[ b,c ]|) \/ LSeg(|[ b,c ]|, |[ b,d ]|)
                 by XBOOLE_0:def 2;
        X= {q2 :
         q2`1=a & c <=q2`2 & q2`2<=d or q2`2=d & a<=q2`1 & q2`1<=b or
          q2`1=b & c <=q2`2 & q2`2<=d or q2`2=c & a<=q2`1 & q2`1<=b}
                                 by A1,Def1;
       then X= (LSeg(|[ a,c ]|, |[ a,d ]|) \/ LSeg(|[ a,d ]|, |[ b,d ]|))
          \/ (LSeg(|[ a,c ]|, |[ b,c ]|) \/ LSeg(|[ b,c ]|, |[ b,d ]|))
                       by A1,Th40;
then A9:   |[b,d]| in X by A8,XBOOLE_0:def 2;
     then (proj2||X). |[b,d]| = |[b,d]|`2 by PSCOMP_1:70
             .= d by EUCLID:56;
     then d in Z by A2,A9,FUNCT_2:43;
    hence thesis by A7;
   end;
 thus N-bound X = sup (proj2||X) by PSCOMP_1:def 31
     .= sup Z by PSCOMP_1:def 21
     .= d by A3,A6,PSCOMP_1:11;
end;

theorem Th48:
 for K being non empty compact Subset of TOP-REAL 2,
 a,b,c,d being real number st K=rectangle(a,b,c,d) & a<=b & c <=d
holds E-bound K = b
proof let K be non empty compact Subset of TOP-REAL 2,
a,b,c,d be real number;
assume A1: K=rectangle(a,b,c,d) & a<=b & c <=d;
  set X = K;
  reconsider Z = (proj1||X).:the carrier of ((TOP-REAL 2)|X)
                       as Subset of REAL;
A2: X = the carrier of ((TOP-REAL 2)|X) by JORDAN1:1;
A3: for p be real number st p in Z holds p <= b
    proof let p be real number;
     assume p in Z;
      then consider p0 being set such that
A4:    p0 in the carrier of (TOP-REAL 2)|X and
         p0 in the carrier of (TOP-REAL 2)|X and
A5:    p = (proj1||X).p0 by FUNCT_2:115;
      reconsider p0 as Point of TOP-REAL 2 by A2,A4;
        X= {q :
         q`1=a & c <=q`2 & q`2<=d or q`2=d & a<=q`1 & q`1<=b or
          q`1=b & c <=q`2 & q`2<=d or q`2=c & a<=q`1 & q`1<=b} by A1,Def1;
        then ex q being Point of TOP-REAL 2 st p0 = q &
        ( q`1=a & c <=q`2 & q`2<=d or q`2=d & a<=q`1 & q`1<=b or
          q`1=b & c <=q`2 & q`2<=d or q`2=c & a<=q`1 & q`1<=b)
                         by A2,A4;
     hence p <= b by A1,A2,A4,A5,PSCOMP_1:69;
    end;
A6: for q being real number st for p being real number
                st p in Z holds p <= q holds b <= q
   proof let q be real number such that
A7:    for p be real number st p in Z holds p <= q;
         |[b,d]| in LSeg(|[ b,c ]|, |[ b,d ]|) by TOPREAL1:6;
then A8:   |[b,d]| in LSeg(|[ a,c ]|, |[ b,c ]|) \/ LSeg(|[ b,c ]|, |[ b,d ]|)
                 by XBOOLE_0:def 2;
        X= {q2 :
         q2`1=a & c <=q2`2 & q2`2<=d or q2`2=d & a<=q2`1 & q2`1<=b or
          q2`1=b & c <=q2`2 & q2`2<=d or q2`2=c & a<=q2`1 & q2`1<=b}
                                 by A1,Def1;
       then X= (LSeg(|[ a,c ]|, |[ a,d ]|) \/ LSeg(|[ a,d ]|, |[ b,d ]|))
          \/ (LSeg(|[ a,c ]|, |[ b,c ]|) \/ LSeg(|[ b,c ]|, |[ b,d ]|))
                       by A1,Th40;
then A9:   |[b,d]| in X by A8,XBOOLE_0:def 2;
     then (proj1||X). |[b,d]| = |[b,d]|`1 by PSCOMP_1:69
             .= b by EUCLID:56;
     then b in Z by A2,A9,FUNCT_2:43;
    hence thesis by A7;
   end;
 thus E-bound X = sup (proj1||X) by PSCOMP_1:def 32
     .= sup Z by PSCOMP_1:def 21
     .= b by A3,A6,PSCOMP_1:11;
end;

theorem Th49:
for K being non empty compact Subset of TOP-REAL 2,
a,b,c,d being real number st K=rectangle(a,b,c,d) & a<=b & c <=d
holds S-bound K = c
proof let K be non empty compact Subset of TOP-REAL 2,
a,b,c,d be real number;
assume A1: K=rectangle(a,b,c,d) & a<=b & c <=d;
  set X = K;
  reconsider Z = (proj2||X).:the carrier of ((TOP-REAL 2)|X)
                                    as Subset of REAL;
A2: X = the carrier of ((TOP-REAL 2)|X) by JORDAN1:1;
A3: for p be real number st p in Z holds p >= c
    proof let p be real number;
     assume p in Z;
      then consider p0 being set such that
A4:    p0 in the carrier of (TOP-REAL 2)|X and
         p0 in the carrier of (TOP-REAL 2)|X and
A5:    p = (proj2||X).p0 by FUNCT_2:115;
      reconsider p0 as Point of TOP-REAL 2 by A2,A4;
        X= {q :
         q`1=a & c <=q`2 & q`2<=d or q`2=d & a<=q`1 & q`1<=b or
          q`1=b & c <=q`2 & q`2<=d or q`2=c & a<=q`1 & q`1<=b} by A1,Def1;
        then ex q being Point of TOP-REAL 2 st p0 = q &
        ( q`1=a & c <=q`2 & q`2<=d or q`2=d & a<=q`1 & q`1<=b or
          q`1=b & c <=q`2 & q`2<=d or q`2=c & a<=q`1 & q`1<=b)
                       by A2,A4;
     hence p >= c by A1,A2,A4,A5,PSCOMP_1:70;
    end;
A6: for q being real number st
   for p being real number st p in Z holds p >= q holds c >= q
   proof let q be real number such that
A7:    for p being real number st p in Z holds p >= q;
       |[b,c]| in LSeg(|[ b,c ]|, |[ b,d ]|) by TOPREAL1:6;
     then A8: |[b,c]| in LSeg(|[ a,c ]|, |[ b,c ]|) \/ LSeg(|[ b,c ]|, |[ b,d
]|)
                 by XBOOLE_0:def 2;
        X= {q2 :
         q2`1=a & c <=q2`2 & q2`2<=d or q2`2=d & a<=q2`1 & q2`1<=b or
          q2`1=b & c <=q2`2 & q2`2<=d or q2`2=c & a<=q2`1 & q2`1<=b}
                                 by A1,Def1;
       then X= (LSeg(|[ a,c ]|, |[ a,d ]|) \/ LSeg(|[ a,d ]|, |[ b,d ]|))
          \/ (LSeg(|[ a,c ]|, |[ b,c ]|) \/ LSeg(|[ b,c ]|, |[ b,d ]|))
                       by A1,Th40;
then A9:   |[b,c]| in X by A8,XBOOLE_0:def 2;
     then (proj2||X). |[b,c]| = |[b,c]|`2 by PSCOMP_1:70
             .= c by EUCLID:56;
     then c in Z by A2,A9,FUNCT_2:43;
    hence thesis by A7;
   end;
 thus S-bound X = inf (proj2||X) by PSCOMP_1:def 33
     .= inf Z by PSCOMP_1:def 20
     .= c by A3,A6,PSCOMP_1:9;
end;

theorem Th50: for K being non empty compact Subset of TOP-REAL 2,
a,b,c,d being real number st K=rectangle(a,b,c,d) & a<=b & c <=d
holds
 NW-corner K = |[a,d]|
proof let K be non empty compact Subset of TOP-REAL 2,
a,b,c,d be real number;
 assume A1: K=rectangle(a,b,c,d) & a<=b & c <=d;
   A2: NW-corner K= |[W-bound K, N-bound K]| by PSCOMP_1:def 35;
     W-bound K=a by A1,Th46;
 hence NW-corner K = |[a,d]| by A1,A2,Th47;
end;

theorem Th51: for K being non empty compact Subset of TOP-REAL 2,
a,b,c,d being real number st K=rectangle(a,b,c,d) & a<=b & c <=d
holds
 NE-corner K = |[b,d]|
proof let K be non empty compact Subset of TOP-REAL 2,
a,b,c,d be real number;
 assume A1: K=rectangle(a,b,c,d) & a<=b & c <=d;
   A2: NE-corner K= |[E-bound K, N-bound K]| by PSCOMP_1:def 36;
     E-bound K=b by A1,Th48;
 hence NE-corner K = |[b,d]| by A1,A2,Th47;
end;

theorem for K being non empty compact Subset of TOP-REAL 2,
a,b,c,d being real number st K=rectangle(a,b,c,d) & a<=b & c <=d
holds
 SW-corner K = |[a,c]|
proof let K be non empty compact Subset of TOP-REAL 2,
a,b,c,d be real number;
 assume A1: K=rectangle(a,b,c,d) & a<=b & c <=d;
   A2: SW-corner K= |[W-bound K, S-bound K]| by PSCOMP_1:def 34;
     W-bound K=a by A1,Th46;
 hence SW-corner K = |[a,c]| by A1,A2,Th49;
end;

theorem for K being non empty compact Subset of TOP-REAL 2,
a,b,c,d being real number st K=rectangle(a,b,c,d) & a<=b & c <=d
holds
 SE-corner K = |[b,c]|
proof let K be non empty compact Subset of TOP-REAL 2,
a,b,c,d be real number;
 assume A1: K=rectangle(a,b,c,d) & a<=b & c <=d;
   A2: SE-corner K= |[E-bound K, S-bound K]| by PSCOMP_1:def 37;
     E-bound K=b by A1,Th48;
 hence SE-corner K = |[b,c]| by A1,A2,Th49;
end;

theorem Th54: for K being non empty compact Subset of TOP-REAL 2,
a,b,c,d being real number st K=rectangle(a,b,c,d) & a<=b & c <=d
 holds W-most K = LSeg(|[a,c]|,|[a,d]|)
proof let K be non empty compact Subset of TOP-REAL 2,
a,b,c,d be real number;
 assume A1: K=rectangle(a,b,c,d) & a<=b & c <=d;
  then K= {p: p`1=a & c <=p`2 & p`2<=d or p`2=d & a<=p`1 & p`1<=b or
      p`1=b & c <=p`2 & p`2<=d or p`2=c & a<=p`1 & p`1<=b} by Def1;
  then K= LSeg(|[ a,c ]|, |[ a,d ]|) \/ LSeg(|[ a,d ]|, |[ b,d ]|)
    \/ (LSeg(|[ a,c ]|, |[ b,c ]|) \/ LSeg(|[ b,c ]|, |[ b,d ]|))
                  by A1,Th40;
then A2: LSeg(|[ a,c ]|, |[ a,d ]|) \/ LSeg(|[ a,d ]|, |[ b,d ]|) c= K
                                  by XBOOLE_1:7;
     LSeg(|[a,c]|,|[a,d]|) c=
   LSeg(|[ a,c ]|, |[ a,d ]|) \/ LSeg(|[ a,d ]|, |[ b,d ]|)
                                       by XBOOLE_1:7;
then A3: LSeg(|[a,c]|,|[a,d]|) c= K by A2,XBOOLE_1:1;
   A4: SW-corner K= |[W-bound K, S-bound K]| by PSCOMP_1:def 34;
   A5: NW-corner K = |[a,d]| by A1,Th50;
   A6: W-bound K=a by A1,Th46;
   A7: S-bound K= c by A1,Th49;
 thus W-most K = LSeg(SW-corner K, NW-corner K)/\K by PSCOMP_1:def 38
       .= LSeg(|[a,c]|,|[a,d]|) by A3,A4,A5,A6,A7,XBOOLE_1:28;
end;

theorem Th55: for K being non empty compact Subset of TOP-REAL 2,
a,b,c,d being real number st K=rectangle(a,b,c,d) & a<=b & c <=d
 holds E-most K = LSeg(|[b,c]|,|[b,d]|)
proof let K be non empty compact Subset of TOP-REAL 2,
a,b,c,d be real number;
 assume A1: K=rectangle(a,b,c,d) & a<=b & c <=d;
  then K= {p: p`1=a & c <=p`2 & p`2<=d or p`2=d & a<=p`1 & p`1<=b or
      p`1=b & c <=p`2 & p`2<=d or p`2=c & a<=p`1 & p`1<=b} by Def1;
  then K= LSeg(|[ a,c ]|, |[ a,d ]|) \/ LSeg(|[ a,d ]|, |[ b,d ]|)
    \/ (LSeg(|[ a,c ]|, |[ b,c ]|) \/ LSeg(|[ b,c ]|, |[ b,d ]|))
                  by A1,Th40;
then A2: LSeg(|[ a,c ]|, |[ b,c ]|) \/ LSeg(|[ b,c ]|, |[ b,d ]|) c= K
                                  by XBOOLE_1:7;
     LSeg(|[b,c]|,|[b,d]|) c=
   LSeg(|[ a,c ]|, |[ b,c ]|) \/ LSeg(|[ b,c ]|, |[ b,d ]|)
                                  by XBOOLE_1:7;
then A3: LSeg(|[b,c]|,|[b,d]|) c= K by A2,XBOOLE_1:1;
   A4: SE-corner K= |[E-bound K, S-bound K]| by PSCOMP_1:def 37;
   A5: NE-corner K = |[b,d]| by A1,Th51;
   A6: E-bound K=b by A1,Th48;
   A7: S-bound K= c by A1,Th49;
 thus E-most K = LSeg(SE-corner K, NE-corner K)/\K by PSCOMP_1:def 40
       .= LSeg(|[b,c]|,|[b,d]|) by A3,A4,A5,A6,A7,XBOOLE_1:28;
end;

theorem Th56: for K being non empty compact Subset of TOP-REAL 2,
a,b,c,d being real number st K=rectangle(a,b,c,d) & a<=b & c <=d
holds W-min K=|[a,c]| & E-max K=|[b,d]|
proof let K be non empty compact Subset of TOP-REAL 2,
a,b,c,d be real number;
 assume A1: K=rectangle(a,b,c,d) & a<=b & c <=d;
  A2: W-min K=|[a,c]|
proof
A3: inf (proj2||LSeg(|[a,c]|,|[a,d]|)) = c
proof
  set X = LSeg(|[a,c]|,|[a,d]|);
  reconsider Z = (proj2||X).:the carrier of ((TOP-REAL 2)|X)
                     as Subset of REAL;
A4: X = the carrier of ((TOP-REAL 2)|X) by JORDAN1:1;
A5: for p be real number st p in Z holds p >= c
    proof let p be real number;
     assume p in Z;
      then consider p0 being set such that
A6:    p0 in the carrier of (TOP-REAL 2)|X and
         p0 in the carrier of (TOP-REAL 2)|X and
A7:    p = (proj2||X).p0 by FUNCT_2:115;
      reconsider p0 as Point of TOP-REAL 2 by A4,A6;
        |[a,c]|`2 = c & |[a,d]|`2 = d by EUCLID:56;
      then p0`2 >= c by A1,A4,A6,TOPREAL1:10;
     hence p >= c by A4,A6,A7,PSCOMP_1:70;
    end;
A8: for q being real number st
    for p being real number st p in Z holds p >= q holds c >= q
   proof let q be real number such that
A9:    for p being real number st p in Z holds p >= q;
A10:   |[a,c]| in X by TOPREAL1:6;
     then (proj2||X). |[a,c]| = |[a,c]|`2 by PSCOMP_1:70
             .= c by EUCLID:56;
     then c in Z by A4,A10,FUNCT_2:43;
    hence thesis by A9;
   end;
 thus inf (proj2||X) = inf Z by PSCOMP_1:def 20
     .= c by A5,A8,PSCOMP_1:9;
end;
 A11: W-most K = LSeg(|[a,c]|,|[a,d]|) by A1,Th54;
   W-bound K = a by A1,Th46;
 hence W-min K= |[a,c]| by A3,A11,PSCOMP_1:def 42;
end;
   E-max K=|[b,d]|
proof
A12: sup (proj2||LSeg(|[b,c]|,|[b,d]|)) = d
proof
  set X = LSeg(|[b,c]|,|[b,d]|);
  reconsider Z = (proj2||X).:the carrier of ((TOP-REAL 2)|X)
                     as Subset of REAL;
A13: X = the carrier of ((TOP-REAL 2)|X) by JORDAN1:1;
A14: for p be real number st p in Z holds p <= d
    proof let p be real number;
     assume p in Z;
      then consider p0 being set such that
A15:    p0 in the carrier of (TOP-REAL 2)|X and
         p0 in the carrier of (TOP-REAL 2)|X and
A16:    p = (proj2||X).p0 by FUNCT_2:115;
      reconsider p0 as Point of TOP-REAL 2 by A13,A15;
        |[b,c]|`2 = c & |[b,d]|`2 = d by EUCLID:56;
      then p0`2 <= d by A1,A13,A15,TOPREAL1:10;
     hence p <= d by A13,A15,A16,PSCOMP_1:70;
    end;
A17: for q being real number st
    for p being real number st p in Z holds p <= q holds d <= q
   proof let q be real number such that
A18:    for p being real number st p in Z holds p <= q;
A19:   |[b,d]| in X by TOPREAL1:6;
     then (proj2||X). |[b,d]| = |[b,d]|`2 by PSCOMP_1:70
             .= d by EUCLID:56;
     then d in Z by A13,A19,FUNCT_2:43;
    hence thesis by A18;
   end;
 thus sup (proj2||X) = sup Z by PSCOMP_1:def 21
     .= d by A14,A17,PSCOMP_1:11;
end;
 A20: E-most K = LSeg(|[b,c]|,|[b,d]|) by A1,Th55;
   E-bound K = b by A1,Th48;
 hence E-max K= |[b,d]| by A12,A20,PSCOMP_1:def 46;
end;
 hence W-min K=|[a,c]| & E-max K=|[b,d]| by A2;
end;

theorem Th57: for K being non empty compact Subset of TOP-REAL 2,
a,b,c,d being real number st K=rectangle(a,b,c,d) & a<b & c <d
holds LSeg(|[a,c]|,|[a,d]|) \/ LSeg(|[a,d]|,|[b,d]|)
 is_an_arc_of W-min(K),E-max(K)
 &
 LSeg(|[a,c]|,|[b,c]|) \/ LSeg(|[b,c]|,|[b,d]|)
 is_an_arc_of E-max(K),W-min(K)
proof let K be non empty compact Subset of TOP-REAL 2,
a,b,c,d be real number;
 assume A1: K=rectangle(a,b,c,d) & a<b & c <d;
  then A2: W-min(K)= |[a,c]| by Th56;
  A3: E-max(K)= |[b,d]| by A1,Th56;
    (|[a,c]|)`2=c by EUCLID:56;
  then A4: |[a,c]| <> |[a,d]| by A1,EUCLID:56;
  set p1= |[a,c]|,p2= |[a,d]|,q1=|[b,d]|;
  A5: LSeg(p1,p2) /\ LSeg(p2,q1) ={p2} by A1,Th44;
    (|[a,c]|)`1=a by EUCLID:56;
  then A6: |[a,c]| <> |[b,c]| by A1,EUCLID:56;
  set q2=|[b,c]|;
    LSeg(q1,q2) /\ LSeg(q2,p1) ={q2} by A1,Th42;
 hence LSeg(|[a,c]|,|[a,d]|) \/ LSeg(|[a,d]|,|[b,d]|)
 is_an_arc_of W-min(K),E-max(K)
 &
 LSeg(|[a,c]|,|[b,c]|) \/ LSeg(|[b,c]|,|[b,d]|)
 is_an_arc_of E-max(K),W-min(K) by A2,A3,A4,A5,A6,TOPREAL1:18;
end;

theorem Th58: for P,P1,P2 being Subset of TOP-REAL 2,
a,b,c,d being real number,
f1,f2 being FinSequence of TOP-REAL 2,
p0,p1,p01,p10 being Point of TOP-REAL 2 st
 a < b & c < d &
P={p: p`1=a & c <=p`2 & p`2<=d or p`2=d & a<=p`1 & p`1<=b or
      p`1=b & c <=p`2 & p`2<=d or p`2=c & a<=p`1 & p`1<=b} &
p0=|[a,c]| & p1=|[b,d]| & p01=|[a,d]| & p10=|[b,c]|
& f1=<*p0,p01,p1*>
& f2=<*p0,p10,p1*> holds
  f1 is_S-Seq &
  L~f1 = LSeg(p0,p01) \/ LSeg(p01,p1) &
  f2 is_S-Seq &
  L~f2 = LSeg(p0,p10) \/ LSeg(p10,p1) &
  P = L~f1 \/ L~f2 &
  L~f1 /\ L~f2 = {p0,p1} & f1/.1 = p0 &
  f1/.len f1=p1 & f2/.1 = p0 & f2/.len f2 = p1
 proof let P,P1,P2 be Subset of TOP-REAL 2,
   a,b,c,d be real number,
   f1,f2 be FinSequence of TOP-REAL 2,
   p0,p1,p01,p10 be Point of TOP-REAL 2;
   assume A1: a < b & c < d &
    P={p: p`1=a & c <=p`2 & p`2<=d or p`2=d & a<=p`1 & p`1<=b or
      p`1=b & c <=p`2 & p`2<=d or p`2=c & a<=p`1 & p`1<=b} &
    p0=|[a,c]| & p1=|[b,d]| & p01=|[a,d]| & p10=|[b,c]|
    & f1=<*p0,p01,p1*>
    & f2=<*p0,p10,p1*>;
  set L1 = { p : p`1 = a & p`2 <= d & p`2 >= c};
  set L2 = { p : p`1 <= b & p`1 >= a & p`2 = d};
  set L3 = { p : p`1 <= b & p`1 >= a & p`2 = c};
  set L4 = { p : p`1 = b & p`2 <= d & p`2 >= c};
 A2: p1`1 = b & p1`2 = d by A1,EUCLID:56;
 A3: p10`1 = b & p10`2 = c by A1,EUCLID:56;
 A4: p0`1 = a & p0`2 = c by A1,EUCLID:56;
A5: len f1 = 1 + 2 by A1,FINSEQ_1:62;
A6: f1/.1 = p0 & f1/.2 = p01 & f1/.3 = p1 by A1,FINSEQ_4:27;
thus f1 is_S-Seq
    proof
       p0 <> p01 & p01 <> p1 & p0 <> p1 by A1,A2,A4,EUCLID:56;
     hence f1 is one-to-one by A1,FINSEQ_3:104;
     thus len f1 >= 2 by A5;
     thus f1 is unfolded
      proof
       let i be Nat; assume A7: 1 <= i & i + 2 <= len f1;
       then i <= 1 by A5,REAL_1:53;
       then A8: i = 1 by A7,AXIOMS:21;
       reconsider n2=1+1 as Nat;
         n2 in Seg len f1 by A5,FINSEQ_1:3;
       then A9: LSeg(f1,1) = LSeg(p0,p01) & LSeg(f1,n2) = LSeg(p01,p1)
        by A5,A6,TOPREAL1:def 5;
         for x being set holds x in LSeg(p0,p01) /\ LSeg(p01,p1) iff x = p01
        proof
         let x be set;
         thus x in LSeg(p0,p01) /\ LSeg(p01,p1) implies x = p01
          proof
           assume x in LSeg(p0,p01) /\ LSeg(p01,p1);
           then x in LSeg(p0,p01) & x in LSeg(p01,p1) by XBOOLE_0:def 3;
           then A10: x in {p : p`1 = a & p`2 <= d & p`2 >= c} &
             x in {p2 : p2`1 <= b & p2`1 >= a & p2`2 = d} by A1,Th39;
           then A11: ex p st p = x & p`1 = a & p`2 <= d & p`2 >= c;
             ex p2 st p2=x & p2`1<=b & p2`1>=a & p2`2=d by A10;
           hence x = p01 by A1,A11,EUCLID:57;
          end;
          assume x = p01;
          then x in LSeg(p0,p01) & x in LSeg(p01,p1) by TOPREAL1:6;
         hence x in LSeg(p0,p01) /\ LSeg(p01,p1) by XBOOLE_0:def 3;
        end;
       hence thesis by A6,A8,A9,TARSKI:def 1;
      end;
     thus f1 is s.n.c.
      proof
       let i,j be Nat such that A12: i+1 < j;
         now per cases;
        suppose 1 <= i;
         then A13:      1+1 <= i+1 by AXIOMS:24;
           now per cases;
          case 1 <= j & j+1 <= len f1;
          then j <= 2 by A5,REAL_1:53;
          hence contradiction by A12,A13,AXIOMS:22;
          case not (1 <= j & j+1 <= len f1);
          then LSeg(f1,j) = {} by TOPREAL1:def 5;
          hence LSeg(f1,i) /\ LSeg(f1,j) = {};
         end;
        hence LSeg(f1,i) /\ LSeg(f1,j) = {};
        suppose not (1 <= i & i+1 <= len f1);
        then LSeg(f1,i) = {} by TOPREAL1:def 5;
        hence LSeg(f1,i) /\ LSeg(f1,j) = {};
       end;
       hence LSeg(f1,i) /\ LSeg(f1,j) = {};
      end;
     let i be Nat; assume A14: 1 <= i & i + 1 <= len f1;
      then A15: i <= 1 + 1 by A5,REAL_1:53;
        now per cases by A14,A15,NAT_1:27;
       suppose A16: i = 1;
        then (f1/.i)`1 = p0`1 by A1,FINSEQ_4:27 .= a
                                  by A1,EUCLID:56
                 .= (f1/.(i+1))`1 by A1,A6,A16,EUCLID:56;
        hence (f1/.i)`1 = (f1/.(i+1))`1 or (f1/.i)`2 = (f1/.(i+1))`2;
       suppose A17: i = 2;
        then (f1/.i)`2 = p01`2 by A1,FINSEQ_4:27 .= d by A1,EUCLID:56
                 .= (f1/.(i+1))`2 by A1,A6,A17,EUCLID:56;
        hence (f1/.i)`1 = (f1/.(i+1))`1 or (f1/.i)`2 = (f1/.(i+1))`2;
       end;
      hence thesis;
    end;
     L~f1 = union {LSeg(p0,p01),LSeg(p01,p1)}
    proof
       len f1 = 2+1 & 1+1 in Seg len f1
      by A5,FINSEQ_1:3;
     then 1+1 <= len f1 & LSeg(p0,p01) = LSeg(f1,1)
      by A6,TOPREAL1:def 5;
     then A18: LSeg(p0,p01) in {LSeg(f1,i): 1 <= i & i+1 <= len f1};
       2+1 <= len f1 & LSeg(p01,p1) = LSeg(f1,2)
      by A5,A6,TOPREAL1:def 5;
     then LSeg(p01,p1) in {LSeg(f1,i): 1 <= i & i+1 <= len f1};
     then A19: {LSeg(p0,p01),LSeg(p01,p1)} c= {LSeg(f1,i):1<=i & i+1<=len f1}
      by A18,ZFMISC_1:38;
       {LSeg(f1,i): 1 <= i & i+1 <= len f1} c= {LSeg(p0,p01),LSeg(p01,p1)}
      proof
       let a be set; assume a in {LSeg(f1,i): 1 <= i & i+1 <= len f1};
       then consider i such that A20: a = LSeg(f1,i) & 1<=i & i+1<=len f1;
         i+1 <= 2 + 1 by A1,A20,FINSEQ_1:62; then i <= 1 + 1 by REAL_1:53;
       then i = 1 or i = 2 by A20,NAT_1:27;
       then a = LSeg(p0,p01) or a = LSeg(p01,p1) by A6,A20,TOPREAL1:def 5;
       hence a in {LSeg(p0,p01),LSeg(p01,p1)} by TARSKI:def 2;
      end;
    hence union {LSeg(p0,p01),LSeg(p01,p1)} =
     union {LSeg(f1,i): 1 <= i & i+1 <= len f1} by A19,XBOOLE_0:def 10
 .= L~f1 by TOPREAL1:def 6;
    end;
hence A21: L~f1 = (LSeg(p0,p01) \/ LSeg(p01,p1)) by ZFMISC_1:93;
then A22: L~f1=L1 \/ LSeg(p01,p1) by A1,Th39
          .=L1 \/ L2 by A1,Th39;
A23: len f2 = 1 + 2 by A1,FINSEQ_1:62;
A24: f2/.1 = p0 & f2/.2 = p10 & f2/.3 = p1 by A1,FINSEQ_4:27;
thus f2 is_S-Seq
    proof
     thus f2 is one-to-one by A1,A2,A3,A4,FINSEQ_3:104;
     thus len f2 >= 2 by A23;
     thus f2 is unfolded
      proof
       let i; assume A25: 1 <= i & i + 2 <= len f2;
       then i <= 1 by A23,REAL_1:53;
       then A26: i = 1 by A25,AXIOMS:21;
         1+1 in Seg len f2
        by A23,FINSEQ_1:3;
       then A27: LSeg(f2,1) = LSeg(p0,p10) & LSeg(f2,1+1) = LSeg(p10,p1)
        by A23,A24,TOPREAL1:def 5;
         for x being set holds x in LSeg(p0,p10) /\ LSeg(p10,p1) iff x = p10
        proof
         let x be set;
         thus x in LSeg(p0,p10) /\ LSeg(p10,p1) implies x = p10
          proof
           assume x in LSeg(p0,p10) /\ LSeg(p10,p1);
           then x in LSeg(p0,p10) &
             x in LSeg(p10,p1) by XBOOLE_0:def 3;
           then A28: x in { p : p`1 <= b & p`1 >= a & p`2 = c} &
             x in { p2 : p2`1 = b & p2`2 <= d & p2`2 >= c} by A1,Th39;
           then A29: ex p st p = x & p`1 <= b & p`1 >= a & p`2 = c;
             ex p2 st p2=x & p2`1=b & p2`2<=d & p2`2>=c by A28;
           hence x = p10 by A1,A29,EUCLID:57;
          end;
          assume x = p10;
          then x in LSeg(p0,p10) & x in LSeg(p10,p1) by TOPREAL1:6;
         hence x in LSeg(p0,p10) /\ LSeg(p10,p1) by XBOOLE_0:def 3;
        end;
       hence thesis by A24,A26,A27,TARSKI:def 1;
      end;
     thus f2 is s.n.c.
      proof
       let i,j such that A30: i+1 < j;
         now per cases;
        suppose 1 <= i;
         then A31:      1+1 <= i+1 by AXIOMS:24;
           now per cases;
          case 1 <= j & j+1 <= len f2;
          then j <= 2 by A23,REAL_1:53;
          hence contradiction by A30,A31,AXIOMS:22;
          case not (1 <= j & j+1 <= len f2);
          then LSeg(f2,j) = {} by TOPREAL1:def 5;
          hence LSeg(f2,i) /\ LSeg(f2,j) = {};
         end;
        hence LSeg(f2,i) /\ LSeg(f2,j) = {};
        suppose not (1 <= i & i+1 <= len f2);
        then LSeg(f2,i) = {} by TOPREAL1:def 5;
        hence LSeg(f2,i) /\ LSeg(f2,j) = {};
       end;
       hence LSeg(f2,i) /\ LSeg(f2,j) = {};
      end;
     let i; assume A32: 1 <= i & i + 1 <= len f2;
      then A33: i <= 1 + 1 by A23,REAL_1:53;
       per cases by A32,A33,NAT_1:27;
       suppose A34: i = 1;
        then (f2/.i)`2 = p0`2 by A1,FINSEQ_4:27 .= c
                           by A1,EUCLID:56
                 .= (f2/.(i+1))`2 by A1,A24,A34,EUCLID:56;
        hence (f2/.i)`1 = (f2/.(i+1))`1 or (f2/.i)`2 = (f2/.(i+1))`2;
       suppose A35: i = 2;
        then (f2/.i)`1 = p10`1 by A1,FINSEQ_4:27 .= b
                             by A1,EUCLID:56
                 .= (f2/.(i+1))`1 by A1,A24,A35,EUCLID:56;
        hence (f2/.i)`1 = (f2/.(i+1))`1 or (f2/.i)`2 = (f2/.(i+1))`2;
    end;
A36:   L~f2 = union {LSeg(p0,p10),LSeg(p10,p1)}
    proof
       len f2 = 2+1 & 1+1 in Seg len f2
      by A23,FINSEQ_1:3;
     then 1+1 <= len f2 & LSeg(p0,p10) = LSeg(f2,1)
      by A24,TOPREAL1:def 5;
     then A37: LSeg(p0,p10) in {LSeg(f2,i): 1 <= i & i+1 <= len f2};
       2+1 <= len f2 & LSeg(p10,p1) = LSeg(f2,2)
      by A23,A24,TOPREAL1:def 5;
     then LSeg(p10,p1) in {LSeg(f2,i): 1 <= i & i+1 <= len f2};
     then A38: {LSeg(p0,p10),LSeg(p10,p1)} c= {LSeg(f2,i):1<=i & i+1<=len f2}
      by A37,ZFMISC_1:38;
       {LSeg(f2,i): 1 <= i & i+1 <= len f2} c= {LSeg(p0,p10),LSeg(p10,p1)}
      proof
       let ax be set; assume ax in {LSeg(f2,i): 1 <= i & i+1 <= len f2};
       then consider i such that A39: ax = LSeg(f2,i) & 1<=i & i+1<=len f2;
         i+1 <= 2 + 1 by A1,A39,FINSEQ_1:62; then i <= 1 + 1 by REAL_1:53;
       then i = 1 or i = 2 by A39,NAT_1:27;
       then ax = LSeg(p0,p10) or ax = LSeg(p10,p1) by A24,A39,TOPREAL1:def 5;
       hence ax in {LSeg(p0,p10),LSeg(p10,p1)} by TARSKI:def 2;
      end;
    hence union {LSeg(p0,p10),LSeg(p10,p1)} =
     union {LSeg(f2,i): 1 <= i & i+1 <= len f2} by A38,XBOOLE_0:def 10
 .= L~f2 by TOPREAL1:def 6;
    end;
hence L~f2 = LSeg(p0,p10) \/ LSeg(p10,p1) by ZFMISC_1:93;
     L~f2 = (LSeg(p0,p10) \/ LSeg(p10,p1)) by A36,ZFMISC_1:93;
   then A40: L~f2=L3 \/ LSeg(p10,p1) by A1,Th39
             .=L3 \/ L4 by A1,Th39;
     P = LSeg(p0,p01) \/ LSeg(p01,p1) \/ (LSeg(p0,p10) \/ LSeg(p10,p1))
                 by A1,Th40;
   hence P = L~f1 \/ L~f2 by A21,A36,ZFMISC_1:93;
   now assume L2 meets L3;
     then consider x being set such that
     A41: x in L2 & x in L3 by XBOOLE_0:3;
     A42: ex p st p = x & p`1 <= b & p`1 >= a & p`2 = d by A41;
       ex p2 st p2 = x & p2`1 <= b & p2`1 >= a & p2`2 = c by A41;
     hence contradiction by A1,A42;
    end;
then A43: L2 /\ L3 = {} by XBOOLE_0:def 7;
  A44: LSeg(|[ a,c ]|, |[ a,d ]|) = { p3 : p3`1 = a & p3`2 <= d & p3`2 >= c}
 & LSeg(|[ a,d ]|, |[ b,d ]|) = { p2 : p2`1 <= b & p2`1 >= a & p2`2 = d}
 & LSeg(|[ a,c ]|, |[ b,c ]|) = { q1 : q1`1 <= b & q1`1 >= a & q1`2 = c}
 & LSeg(|[ b,c ]|, |[ b,d ]|) = { q2 : q2`1 = b & q2`2 <= d & q2`2 >= c}
 & LSeg(|[a,c]|,|[a,d]|) /\ LSeg(|[a,c]|,|[b,c]|) = {|[a,c]|}
 & LSeg(|[a,d]|,|[b,d]|) /\ LSeg(|[b,c]|,|[b,d]|) = {|[b,d]|}
             by A1,Th39,Th41,Th43;
   now assume L1 meets L4;
     then consider x being set such that
     A45: x in L1 & x in L4 by XBOOLE_0:3;
     A46: ex p st p = x & p`1 = a & p`2 <= d & p`2 >= c by A45;
       ex p2 st p2 = x & p2`1 = b & p2`2 <= d & p2`2 >= c by A45;
     hence contradiction by A1,A46;
    end;
then A47: L1 /\ L4 = {} by XBOOLE_0:def 7;
   thus L~f1 /\ L~f2
       = (L1 \/ L2) /\ L3 \/ (L1 \/ L2) /\ L4 by A22,A40,XBOOLE_1:23
      .= L1 /\ L3 \/ L2 /\ L3 \/ (L1 \/ L2) /\ L4 by XBOOLE_1:23
      .= L1 /\ L3 \/ (L1 /\ L4 \/ L2 /\ L4) by A43,XBOOLE_1:23
      .= {p0, p1} by A1,A44,A47,ENUMSET1:41;
  thus f1/.1 = p0 & f1/.len f1=p1 & f2/.1 = p0 &
       f2/.len f2 = p1 by A1,A5,A23,FINSEQ_4:27;
 end;

theorem Th59: for P,P1,P2 being Subset of TOP-REAL 2,
a,b,c,d being real number,
f1,f2 being FinSequence of TOP-REAL 2,p1,p2 being Point of TOP-REAL 2 st
 a < b & c < d &
P={p: p`1=a & c <=p`2 & p`2<=d or p`2=d & a<=p`1 & p`1<=b or
      p`1=b & c <=p`2 & p`2<=d or p`2=c & a<=p`1 & p`1<=b} &
p1=|[a,c]| & p2=|[b,d]|
& f1=<*|[a,c]|,|[a,d]|,|[b,d]|*>
& f2=<*|[a,c]|,|[b,c]|,|[b,d]|*>
& P1=L~f1
& P2=L~f2
holds P1 is_an_arc_of p1,p2 & P2 is_an_arc_of p1,p2
& P1 is non empty & P2 is non empty & P = P1 \/ P2 & P1 /\ P2 = {p1,p2}
proof let P,P1,P2 be Subset of TOP-REAL 2,
  a,b,c,d be real number,
  f1,f2 be FinSequence of TOP-REAL 2,p1,p2 be Point of TOP-REAL 2;
 assume A1: a < b & c < d &
   P={p: p`1=a & c <=p`2 & p`2<=d or p`2=d & a<=p`1 & p`1<=b or
      p`1=b & c <=p`2 & p`2<=d or p`2=c & a<=p`1 & p`1<=b} &
   p1=|[a,c]| & p2=|[b,d]|
   & f1=<*|[a,c]|,|[a,d]|,|[b,d]|*>
   & f2=<*|[a,c]|,|[b,c]|,|[b,d]|*>
   & P1=L~f1 & P2=L~f2;
     (|[a,c]|)`2=c by EUCLID:56;
   then A2: |[a,c]|<>|[a,d]| or |[a,d]|<>|[b,d]| by A1,EUCLID:56;
   A3: P1=LSeg(|[a,c]|,|[a,d]|) \/ LSeg(|[a,d]|,|[b,d]|) by A1,Th58;
   A4: LSeg(|[a,c]|,|[a,d]|) /\ LSeg(|[a,d]|,|[b,d]|)={|[a,d]|} by A1,Th44;
     (|[b,c]|)`2=c by EUCLID:56;
   then A5: |[a,c]|<>|[b,c]| or |[b,c]|<>|[b,d]| by A1,EUCLID:56;
   A6: P2=LSeg(|[a,c]|,|[b,c]|) \/ LSeg(|[b,c]|,|[b,d]|) by A1,Th58;
   LSeg(|[a,c]|,|[b,c]|) /\ LSeg(|[b,c]|,|[b,d]|)={|[b,c]|} by A1,Th42;
 hence thesis by A1,A2,A3,A4,A5,A6,Th58,TOPREAL1:18;
end;

theorem Th60: for a,b,c,d being real number st
a < b & c < d
holds rectangle(a,b,c,d) is_simple_closed_curve
proof let a,b,c,d be real number;
  assume A1: a < b & c < d;
   set P=rectangle(a,b,c,d);

   set p1=|[a,c]|,p2=|[b,d]|;
   reconsider f1=<*|[a,c]|,|[a,d]|,|[b,d]|*> as FinSequence of TOP-REAL 2;
   reconsider f2=<*|[a,c]|,|[b,c]|,|[b,d]|*> as FinSequence of TOP-REAL 2;
   set P1=L~f1,P2=L~f2;
   A2: a < b & c < d &
    P={p: p`1=a & c <=p`2 & p`2<=d or p`2=d & a<=p`1 & p`1<=b or
      p`1=b & c <=p`2 & p`2<=d or p`2=c & a<=p`1 & p`1<=b} &
    p1=|[a,c]| & p2=|[b,d]|
    & f1=<*|[a,c]|,|[a,d]|,|[b,d]|*>
    & f2=<*|[a,c]|,|[b,c]|,|[b,d]|*>
    & P1=L~f1 & P2=L~f2
    implies P1 is_an_arc_of p1,p2 & P2 is_an_arc_of p1,p2
    & P1 is non empty & P2 is non empty &
    P = P1 \/ P2 & P1 /\ P2 = {p1,p2} by Th59;
     (|[a,c]|)`1=a by EUCLID:56;
   then A3: p1<>p2 by A1,EUCLID:56;
     p1 in P1 /\ P2 by A1,A2,Def1,TARSKI:def 2;
   then p1 in P1 by XBOOLE_0:def 3;
   then A4: p1 in P by A1,A2,Def1,XBOOLE_0:def 2;
     p2 in P1 /\ P2 by A1,A2,Def1,TARSKI:def 2;
   then p2 in P1 by XBOOLE_0:def 3;
   then p2 in P by A1,A2,Def1,XBOOLE_0:def 2;
 hence thesis by A1,A2,A3,A4,Def1,TOPREAL2:6;
end;

theorem Th61: for K being non empty compact Subset of TOP-REAL 2,
a,b,c,d being real number st K=rectangle(a,b,c,d) & a<b & c <d
holds Upper_Arc(K)= LSeg(|[a,c]|,|[a,d]|) \/ LSeg(|[a,d]|,|[b,d]|)
proof let K be non empty compact Subset of TOP-REAL 2,
a,b,c,d be real number;
 assume A1: K=rectangle(a,b,c,d) & a<b & c <d;
  then A2: K is_simple_closed_curve by Th60;
  set P=K;
  A3: W-min(K)= |[a,c]| by A1,Th56;
  A4: E-max(K)= |[b,d]| by A1,Th56;
  reconsider U= LSeg(|[a,c]|,|[a,d]|) \/ LSeg(|[a,d]|,|[b,d]|)
            as non empty Subset of TOP-REAL 2;
  A5: U is_an_arc_of W-min(P),E-max(P) by A1,Th57;
  reconsider P3= LSeg(|[a,c]|,|[b,c]|) \/ LSeg(|[b,c]|,|[b,d]|)
            as non empty Subset of TOP-REAL 2;
  A6: P3 is_an_arc_of E-max(P),W-min(P) by A1,Th57;
  reconsider f1=<* |[a,c]|,|[a,d]|,|[b,d]| *>,
     f2=<* |[a,c]|,|[b,c]|,|[b,d]| *> as FinSequence of TOP-REAL 2;
  set p0=|[a,c]|,p01=|[a,d]|,p10=|[b,c]|,p1=|[b,d]|;
 A7: a < b & c < d &
K={p: p`1=a & c <=p`2 & p`2<=d or p`2=d & a<=p`1 & p`1<=b or
      p`1=b & c <=p`2 & p`2<=d or p`2=c & a<=p`1 & p`1<=b} &
p0=|[a,c]| & p1=|[b,d]| & p01=|[a,d]| & p10=|[b,c]|
& f1=<*p0,p01,p1*>
& f2=<*p0,p10,p1*> implies
  f1 is_S-Seq &
  L~f1 = LSeg(p0,p01) \/ LSeg(p01,p1) &
  f2 is_S-Seq &
  L~f2 = LSeg(p0,p10) \/ LSeg(p10,p1) &
  K = L~f1 \/ L~f2 &
  L~f1 /\ L~f2 = {p0,p1} & f1/.1 = p0 &
  f1/.len f1=p1 & f2/.1 = p0 & f2/.len f2 = p1 by Th58;
  A8: Vertical_Line((W-bound(P)+E-bound(P))/2)
   = Vertical_Line((a+E-bound(P))/2) by A1,Th46
  .= Vertical_Line((a+b)/2) by A1,Th48;
   set Q=Vertical_Line((W-bound(P)+E-bound(P))/2);
   reconsider a2=a,b2=b,c2=c,d2=d as Real by XREAL_0:def 1;
  A9: U /\ Vertical_Line((W-bound(P)+E-bound(P))/2)
   ={|[(a+b)/2,d]|}
   proof
     thus U /\ Vertical_Line((W-bound(P)+E-bound(P))/2)
      c= {|[(a+b)/2,d]|}
      proof let x be set;assume
          x in U /\ Vertical_Line((W-bound(P)+E-bound(P))/2);
        then A10: x in U & x in Vertical_Line((W-bound(P)+E-bound(P))/2)
                                  by XBOOLE_0:def 3;
        then x in {p where p is Point of TOP-REAL 2: p`1=(a+b)/2}
                                  by A8,JORDAN6:def 6;
        then consider p such that
        A11: x=p & p`1=(a+b)/2;
          now assume p in LSeg(|[a,c]|,|[a,d]|);
          then p in LSeg(|[a2,c2]|,|[a2,d2]|);
          then p`1=a by TOPREAL3:17;
          then a+a=(a+b)/2*2 by A11,XCMPLX_1:11 .=a+b by XCMPLX_1:88;
          then a+b-a=a by XCMPLX_1:26;
         hence contradiction by A1,XCMPLX_1:26;
        end;
        then p in LSeg(|[a2,d2]|,|[b2,d2]|) by A10,A11,XBOOLE_0:def 2;
        then p`2=d by TOPREAL3:18;
        then x=|[(a+b)/2,d]| by A11,EUCLID:57;
       hence x in {|[(a+b)/2,d]|} by TARSKI:def 1;
      end;
     let x be set;assume x in {|[(a+b)/2,d]|};
      then A12: x= |[(a+b)/2,d]| by TARSKI:def 1;
        (|[(a+b)/2,d]|)`1= (a+b)/2 by EUCLID:56;
      then x in {p where p is Point of TOP-REAL 2: p`1=(a+b)/2} by A12;
      then A13: x in Vertical_Line((W-bound(P)+E-bound(P))/2) by A8,JORDAN6:def
6;
      A14: (|[b,d]|)`1=b & (|[b,d]|)`2=d by EUCLID:56;
        (|[a,d]|)`1=a & (|[a,d]|)`2=d by EUCLID:56;
      then x in LSeg(|[b,d]|,|[a,d]|) by A1,A12,A14,TOPREAL3:19;
      then x in U by XBOOLE_0:def 2;
     hence
        x in U /\ Vertical_Line((W-bound(P)+E-bound(P))/2) by A13,XBOOLE_0:def
3
;
   end;
  then |[(a+b)/2,d]| in U /\ Q by TARSKI:def 1;
   then A15: U meets Q by XBOOLE_0:4;
   A16: Q is closed by JORDAN6:33;
     U is closed by A5,JORDAN6:12;
   then U /\ Q is closed by A16,TOPS_1:35;
   then First_Point(U,W-min(P),E-max(P),
        Vertical_Line((W-bound(P)+E-bound(P))/2)) in {|[(a+b)/2,d]|}
                            by A5,A9,A15,JORDAN5C:def 1;
      then First_Point(U,W-min(P),E-max(P),
        Vertical_Line((W-bound(P)+E-bound(P))/2)) = |[(a+b)/2,d]|
                            by TARSKI:def 1;
      then A17: First_Point(U,W-min(P),E-max(P),
        Vertical_Line((W-bound(P)+E-bound(P))/2))`2=d by EUCLID:56;
  A18: P3 /\ Vertical_Line((W-bound(P)+E-bound(P))/2)
   ={|[(a+b)/2,c]|}
   proof
     thus P3 /\ Vertical_Line((W-bound(P)+E-bound(P))/2)
      c= {|[(a+b)/2,c]|}
      proof let x be set;assume
          x in P3 /\ Vertical_Line((W-bound(P)+E-bound(P))/2);
        then A19: x in P3 & x in Vertical_Line((W-bound(P)+E-bound(P))/2)
                                  by XBOOLE_0:def 3;
        then x in {p where p is Point of TOP-REAL 2: p`1=(a+b)/2}
                                  by A8,JORDAN6:def 6;
        then consider p such that
        A20: x=p & p`1=(a+b)/2;
          now assume p in LSeg(|[b,c]|,|[b,d]|);
          then p in LSeg(|[b2,c2]|,|[b2,d2]|);
          then p`1= b by TOPREAL3:17;
          then b+b=(a+b)/2*2 by A20,XCMPLX_1:11 .=a+b by XCMPLX_1:88;
          then a+b-b=b by XCMPLX_1:26;
         hence contradiction by A1,XCMPLX_1:26;
        end;
        then p in LSeg(|[a2,c2]|,|[b2,c2]|) by A19,A20,XBOOLE_0:def 2;
        then p`2= c by TOPREAL3:18;
        then x=|[(a+b)/2,c]| by A20,EUCLID:57;
       hence x in {|[(a+b)/2,c]|} by TARSKI:def 1;
      end;
     let x be set;assume x in {|[(a+b)/2,c]|};
      then A21: x= |[(a+b)/2,c]| by TARSKI:def 1;
        (|[(a+b)/2,c]|)`1= (a+b)/2 by EUCLID:56;
      then x in {p where p is Point of TOP-REAL 2: p`1=(a+b)/2} by A21;
      then A22: x in Vertical_Line((W-bound(P)+E-bound(P))/2) by A8,JORDAN6:def
6;
      A23: (|[b,c]|)`1=b & (|[b,c]|)`2=c by EUCLID:56;
        (|[a,c]|)`1=a & (|[a,c]|)`2=c by EUCLID:56;
      then |[(b+a)/2,c]| in LSeg(|[a,c]|,|[b,c]|) by A1,A23,TOPREAL3:19;
      then x in P3 by A21,XBOOLE_0:def 2;
     hence
        x in P3 /\ Vertical_Line((W-bound(P)+E-bound(P))/2)
                         by A22,XBOOLE_0:def 3;
   end;
   then |[(a+b)/2,c]| in P3 /\ Q by TARSKI:def 1;
   then A24: P3 meets Q by XBOOLE_0:4;
     P3 is closed by A6,JORDAN6:12;
   then P3 /\ Q is closed by A16,TOPS_1:35;
   then Last_Point(P3,E-max(P),W-min(P),
        Vertical_Line((W-bound(P)+E-bound(P))/2)) in {|[(a+b)/2,c]|}
                            by A6,A18,A24,JORDAN5C:def 2;
      then Last_Point(P3,E-max(P),W-min(P),
        Vertical_Line((W-bound(P)+E-bound(P))/2)) = |[(a+b)/2,c]|
                            by TARSKI:def 1;
      then Last_Point(P3,E-max(P),W-min(P),
        Vertical_Line((W-bound(P)+E-bound(P))/2))`2 = c by EUCLID:56;
 hence Upper_Arc(K)= LSeg(|[a,c]|,|[a,d]|) \/ LSeg(|[a,d]|,|[b,d]|)
                   by A1,A2,A3,A4,A5,A6,A7,A17,Def1,JORDAN6:def 8;
end;

theorem Th62: for K being non empty compact Subset of TOP-REAL 2,
a,b,c,d being real number st K=rectangle(a,b,c,d) & a<b & c <d
holds Lower_Arc(K)= LSeg(|[a,c]|,|[b,c]|) \/ LSeg(|[b,c]|,|[b,d]|)
proof let K be non empty compact Subset of TOP-REAL 2,
a,b,c,d be real number;
 assume A1: K=rectangle(a,b,c,d) & a<b & c <d;
  then A2: K is_simple_closed_curve by Th60;
  set P=K;
  A3: W-min(K)= |[a,c]| by A1,Th56;
  A4: E-max(K)= |[b,d]| by A1,Th56;
  reconsider U= LSeg(|[a,c]|,|[a,d]|) \/ LSeg(|[a,d]|,|[b,d]|)
            as non empty Subset of TOP-REAL 2;
  A5: U is_an_arc_of W-min(P),E-max(P) by A1,Th57;
  reconsider P3= LSeg(|[a,c]|,|[b,c]|) \/ LSeg(|[b,c]|,|[b,d]|)
            as non empty Subset of TOP-REAL 2;
  A6: P3 is_an_arc_of E-max(P),W-min(P) by A1,Th57;
  reconsider f1=<* |[a,c]|,|[a,d]|,|[b,d]| *>,
     f2=<* |[a,c]|,|[b,c]|,|[b,d]| *> as FinSequence of TOP-REAL 2;
  set p0=|[a,c]|,p01=|[a,d]|,p10=|[b,c]|,p1=|[b,d]|;
 A7: a < b & c < d &
K={p: p`1=a & c <=p`2 & p`2<=d or p`2=d & a<=p`1 & p`1<=b or
      p`1=b & c <=p`2 & p`2<=d or p`2=c & a<=p`1 & p`1<=b} &
p0=|[a,c]| & p1=|[b,d]| & p01=|[a,d]| & p10=|[b,c]|
& f1=<*p0,p01,p1*>
& f2=<*p0,p10,p1*> implies
  f1 is_S-Seq &
  L~f1 = LSeg(p0,p01) \/ LSeg(p01,p1) &
  f2 is_S-Seq &
  L~f2 = LSeg(p0,p10) \/ LSeg(p10,p1) &
  K = L~f1 \/ L~f2 &
  L~f1 /\ L~f2 = {p0,p1} & f1/.1 = p0 &
  f1/.len f1=p1 & f2/.1 = p0 & f2/.len f2 = p1 by Th58;
  A8: Vertical_Line((W-bound(P)+E-bound(P))/2)
   = Vertical_Line((a+E-bound(P))/2) by A1,Th46
  .= Vertical_Line((a+b)/2) by A1,Th48;
   set Q=Vertical_Line((W-bound(P)+E-bound(P))/2);
   reconsider a2=a,b2=b,c2=c,d2=d as Real by XREAL_0:def 1;
  A9: U /\ Vertical_Line((W-bound(P)+E-bound(P))/2)
   ={|[(a+b)/2,d]|}
   proof
     thus U /\ Vertical_Line((W-bound(P)+E-bound(P))/2)
      c= {|[(a+b)/2,d]|}
      proof let x be set;assume
          x in U /\ Vertical_Line((W-bound(P)+E-bound(P))/2);
        then A10: x in U & x in Vertical_Line((W-bound(P)+E-bound(P))/2)
                                  by XBOOLE_0:def 3;
        then x in {p where p is Point of TOP-REAL 2: p`1=(a+b)/2}
                                  by A8,JORDAN6:def 6;
        then consider p such that
        A11: x=p & p`1=(a+b)/2;
          now assume p in LSeg(|[a,c]|,|[a,d]|);
          then p in LSeg(|[a2,c2]|,|[a2,d2]|);
          then p`1=a by TOPREAL3:17;
          then a+a=(a+b)/2*2 by A11,XCMPLX_1:11 .=a+b by XCMPLX_1:88;
          then a+b-a=a by XCMPLX_1:26;
         hence contradiction by A1,XCMPLX_1:26;
        end;
        then p in LSeg(|[a2,d2]|,|[b2,d2]|) by A10,A11,XBOOLE_0:def 2;
        then p`2=d by TOPREAL3:18;
        then x=|[(a+b)/2,d]| by A11,EUCLID:57;
       hence x in {|[(a+b)/2,d]|} by TARSKI:def 1;
      end;
     let x be set;assume x in {|[(a+b)/2,d]|};
      then A12: x= |[(a+b)/2,d]| by TARSKI:def 1;
        (|[(a+b)/2,d]|)`1= (a+b)/2 by EUCLID:56;
      then x in {p where p is Point of TOP-REAL 2: p`1=(a+b)/2} by A12;
      then A13: x in Vertical_Line((W-bound(P)+E-bound(P))/2) by A8,JORDAN6:def
6;
      A14: (|[b,d]|)`1=b & (|[b,d]|)`2=d by EUCLID:56;
        (|[a,d]|)`1=a & (|[a,d]|)`2=d by EUCLID:56;
      then x in LSeg(|[b,d]|,|[a,d]|) by A1,A12,A14,TOPREAL3:19;
      then x in U by XBOOLE_0:def 2;
     hence
        x in U /\ Vertical_Line((W-bound(P)+E-bound(P))/2) by A13,XBOOLE_0:def
3
;
   end;
  then |[(a+b)/2,d]| in U /\ Q by TARSKI:def 1;
   then A15: U meets Q by XBOOLE_0:4;
   A16: Q is closed by JORDAN6:33;
     U is closed by A5,JORDAN6:12;
   then U /\ Q is closed by A16,TOPS_1:35;
   then First_Point(U,W-min(P),E-max(P),
        Vertical_Line((W-bound(P)+E-bound(P))/2)) in {|[(a+b)/2,d]|}
                            by A5,A9,A15,JORDAN5C:def 1;
      then First_Point(U,W-min(P),E-max(P),
        Vertical_Line((W-bound(P)+E-bound(P))/2)) = |[(a+b)/2,d]|
                            by TARSKI:def 1;
      then A17: First_Point(U,W-min(P),E-max(P),
        Vertical_Line((W-bound(P)+E-bound(P))/2))`2=d by EUCLID:56;

  A18: P3 /\ Vertical_Line((W-bound(P)+E-bound(P))/2)
   ={|[(a+b)/2,c]|}
   proof
     thus P3 /\ Vertical_Line((W-bound(P)+E-bound(P))/2)
      c= {|[(a+b)/2,c]|}
      proof let x be set;assume
          x in P3 /\ Vertical_Line((W-bound(P)+E-bound(P))/2);
        then A19: x in P3 & x in Vertical_Line((W-bound(P)+E-bound(P))/2)
                                  by XBOOLE_0:def 3;
        then x in {p where p is Point of TOP-REAL 2: p`1=(a+b)/2}
                                  by A8,JORDAN6:def 6;
        then consider p such that
        A20: x=p & p`1=(a+b)/2;
          now assume p in LSeg(|[b,c]|,|[b,d]|);
          then p in LSeg(|[b2,c2]|,|[b2,d2]|);
          then p`1= b by TOPREAL3:17;
          then b+b=(a+b)/2*2 by A20,XCMPLX_1:11 .=a+b by XCMPLX_1:88;
          then a+b-b=b by XCMPLX_1:26;
         hence contradiction by A1,XCMPLX_1:26;
        end;
        then p in LSeg(|[a2,c2]|,|[b2,c2]|) by A19,A20,XBOOLE_0:def 2;
        then p`2= c by TOPREAL3:18;
        then x=|[(a+b)/2,c]| by A20,EUCLID:57;
       hence x in {|[(a+b)/2,c]|} by TARSKI:def 1;
      end;
     let x be set;assume x in {|[(a+b)/2,c]|};
      then A21: x= |[(a+b)/2,c]| by TARSKI:def 1;
        (|[(a+b)/2,c]|)`1= (a+b)/2 by EUCLID:56;
      then x in {p where p is Point of TOP-REAL 2: p`1=(a+b)/2} by A21;
      then A22: x in Vertical_Line((W-bound(P)+E-bound(P))/2) by A8,JORDAN6:def
6;
      A23: (|[b,c]|)`1=b & (|[b,c]|)`2=c by EUCLID:56;
        (|[a,c]|)`1=a & (|[a,c]|)`2=c by EUCLID:56;
      then |[(a+b)/2,c]| in LSeg(|[a,c]|,|[b,c]|) by A1,A23,TOPREAL3:19;
      then x in P3 by A21,XBOOLE_0:def 2;
     hence
        x in P3 /\ Vertical_Line((W-bound(P)+E-bound(P))/2)
                         by A22,XBOOLE_0:def 3;
   end;
   then |[(a+b)/2,c]| in P3 /\ Q by TARSKI:def 1;
   then A24: P3 meets Q by XBOOLE_0:4;
     P3 is closed by A6,JORDAN6:12;
   then P3 /\ Q is closed by A16,TOPS_1:35;
   then Last_Point(P3,E-max(P),W-min(P),
        Vertical_Line((W-bound(P)+E-bound(P))/2)) in {|[(a+b)/2,c]|}
                            by A6,A18,A24,JORDAN5C:def 2;
      then Last_Point(P3,E-max(P),W-min(P),
        Vertical_Line((W-bound(P)+E-bound(P))/2)) = |[(a+b)/2,c]|
                            by TARSKI:def 1;
      then Last_Point(P3,E-max(P),W-min(P),
        Vertical_Line((W-bound(P)+E-bound(P))/2))`2 = c by EUCLID:56;
       then P3 is_an_arc_of E-max(P),W-min(P)
         & Upper_Arc(P) /\ P3={W-min(P),E-max(P)}
         & Upper_Arc(P) \/ P3=P
         & First_Point(Upper_Arc(P),W-min(P),E-max(P),
         Vertical_Line((W-bound(P)+E-bound(P))/2))`2>
        Last_Point(P3,E-max(P),W-min(P),
         Vertical_Line((W-bound(P)+E-bound(P))/2))`2 by A1,A3,A4,A7,A17,Def1,
Th57,Th61;
 hence Lower_Arc(K)= LSeg(|[a,c]|,|[b,c]|) \/ LSeg(|[b,c]|,|[b,d]|)
                   by A2,JORDAN6:def 9;
end;

theorem Th63: for K being non empty compact Subset of TOP-REAL 2,
a,b,c,d being real number,p1,p2 being Point of TOP-REAL 2
 st K=rectangle(a,b,c,d) & a<b & c <d
ex f being map of I[01],(TOP-REAL 2)|(Upper_Arc(K)) st f is_homeomorphism
& f.0=W-min(K) & f.1=E-max(K) & rng f=Upper_Arc(K) &
(for r being Real st r in [.0,1/2 .] holds
              f.r=(1-2*r)*|[a,c]|+(2*r)*|[a,d]|)&
(for r being Real st r in [.1/2,1 .] holds
    f.r=(1-(2*r-1))*|[a,d]|+(2*r-1)*|[b,d]|)&
(for p being Point of TOP-REAL 2 st p in LSeg(|[a,c]|,|[a,d]|)
  holds 0<=((p`2)-c)/(d-c)/2 & ((p`2)-c)/(d-c)/2<=1
        & f.(((p`2)-c)/(d-c)/2)=p)&
(for p being Point of TOP-REAL 2 st p in LSeg(|[a,d]|,|[b,d]|)
  holds 0<=((p`1)-a)/(b-a)/2 + 1/2 & ((p`1)-a)/(b-a)/2 + 1/2<=1
        & f.(((p`1)-a)/(b-a)/2 + 1/2)=p)
proof let K be non empty compact Subset of TOP-REAL 2,
a,b,c,d be real number,p1,p2 be Point of TOP-REAL 2;
 assume A1: K=rectangle(a,b,c,d) & a<b & c <d;
  reconsider a2=a,b2=b,c2=c,d2=d as Real by XREAL_0:def 1;
  defpred P[set,set] means for r being Real st $1=r holds
         (r in [.0,1/2 .] implies $2=(1-2*r)*|[a,c]|+(2*r)*|[a,d]|) &
         (r in [.1/2,1 .] implies $2=(1-(2*r-1))*|[a,d]|+(2*r-1)*|[b,d]|);
  A2: [.0,1.]=[.0,1/2 .] \/ [.1/2,1 .] by HEINE:2;
  A3: for x,y1,y2 being set st x in [.0,1.] & P[x,y1] & P[x,y2]
                    holds y1 = y2
   proof let x,y1,y2 be set;
    assume A4: x in [.0,1.] & P[x,y1] & P[x,y2];
      now per cases by A2,A4,XBOOLE_0:def 2;
    case A5: x in [.0,1/2.];
      then reconsider r=x as Real;
        y1= (1-2*r)*|[a,c]|+(2*r)*|[a,d]| by A4,A5;
     hence y1=y2 by A4,A5;
    case A6: x in [.1/2,1.];
      then reconsider r=x as Real;
        y1= (1-(2*r-1))*|[a,d]|+(2*r-1)*|[b,d]| by A4,A6;
     hence y1=y2 by A4,A6;
     end;
    hence y1 = y2;
   end;
  A7: for x being set st x in [.0,1.] ex y being set st P[x,y]
   proof let x be set;assume
    A8: x in [.0,1.];
      now per cases by A2,A8,XBOOLE_0:def 2;
    case A9: x in [.0,1/2.];
      then reconsider r=x as Real;
      A10: 0<=r & r<=1/2 by A9,TOPREAL5:1;
      set y0= (1-2*r)*|[a,c]|+(2*r)*|[a,d]|;
        r in [.1/2,1.] implies y0=(1-(2*r-1))*|[a,d]|+(2*r-1)*|[b,d]|
       proof assume r in [.1/2,1.]; then 1/2 <=r & r<=1 by TOPREAL5:1;
         then A11: r=1/2 by A10,AXIOMS:21;
         then A12: y0= (0)*|[a,c]|+|[a,d]| by EUCLID:33
         .= (0.REAL 2) + |[a,d]| by EUCLID:33
         .= |[a,d]| by EUCLID:31;
           (1-(2*r-1))*|[a,d]|+(2*r-1)*|[b,d]|
          = (1)*|[a,d]|+0.REAL 2 by A11,EUCLID:33
         .= |[a,d]|+0.REAL 2 by EUCLID:33
         .= |[a,d]| by EUCLID:31;
        hence y0=(1-(2*r-1))*|[a,d]|+(2*r-1)*|[b,d]| by A12;
       end;
      then for r2 being Real st x=r2 holds
         (r2 in [.0,1/2.] implies y0=(1-2*r2)*|[a,c]|+(2*r2)*|[a,d]|) &
         (r2 in [.1/2,1.] implies y0=(1-(2*r2-1))*|[a,d]|+(2*r2-1)*|[b,d]|);
     hence ex y being set st P[x,y];
    case A13: x in [.1/2,1.];
      then reconsider r=x as Real;
      A14: 1/2<=r & r<=1 by A13,TOPREAL5:1;
      set y0= (1-(2*r-1))*|[a,d]|+(2*r-1)*|[b,d]|;
        r in [.0,1/2.] implies y0=(1-2*r)*|[a,c]|+(2*r)*|[a,d]|
       proof assume r in [.0,1/2.]; then 0 <=r & r<=1/2 by TOPREAL5:1;
         then A15: r=1/2 by A14,AXIOMS:21;
         then A16: y0= |[a,d]|+(0)*|[b,d]| by EUCLID:33
         .= |[a,d]|+(0.REAL 2) by EUCLID:33
         .= |[a,d]| by EUCLID:31;
           (1-2*r)*|[a,c]|+(2*r)*|[a,d]|
          = 0.REAL 2+(1)*|[a,d]| by A15,EUCLID:33
         .= 0.REAL 2+|[a,d]| by EUCLID:33
         .= |[a,d]| by EUCLID:31;
        hence y0=(1-2*r)*|[a,c]|+(2*r)*|[a,d]| by A16;
       end;
      then for r2 being Real st x=r2 holds
         (r2 in [.0,1/2.] implies y0=(1-2*r2)*|[a,c]|+(2*r2)*|[a,d]|) &
         (r2 in [.1/2,1.] implies y0=(1-(2*r2-1))*|[a,d]|+(2*r2-1)*|[b,d]|);
     hence ex y being set st P[x,y];
     end;
    hence ex y being set st P[x,y];
   end;
    ex f2 being Function st dom f2 = [.0,1.] &
     for x being set st x in [.0,1.] holds P[x,f2.x] from FuncEx(A3,A7);
  then consider f2 being Function such that
  A17: dom f2 = [.0,1.] &
   for x being set st x in [.0,1.] holds P[x,f2.x];
    rng f2 c= the carrier of (TOP-REAL 2)|(Upper_Arc(K))
  proof let y be set;assume y in rng f2;
    then consider x being set such that
    A18: x in dom f2 & y=f2.x by FUNCT_1:def 5;
      now per cases by A2,A17,A18,XBOOLE_0:def 2;
    case A19: x in [.0,1/2.];
      then reconsider r=x as Real;
      A20: 0<=r & r<=1/2 by A19,TOPREAL5:1;
      then A21: r*2<=1/2*2 by AXIOMS:25;
      A22: 2*0<=2*r by A20,AXIOMS:25;
     f2.x= (1-2*r)*|[a,c]|+(2*r)*|[a,d]| by A17,A18,A19;
     then y in { (1-lambda)*|[a,c]| + lambda*|[a,d]| where lambda is Real:
            0 <= lambda & lambda <= 1 } by A18,A21,A22;
     then A23: y in LSeg(|[a,c]|,|[a,d]|) by TOPREAL1:def 4;
       Upper_Arc(K)= LSeg(|[a,c]|,|[a,d]|) \/ LSeg(|[a,d]|,|[b,d]|)
                               by A1,Th61;
     then y in Upper_Arc(K) by A23,XBOOLE_0:def 2;
     hence y in the carrier of (TOP-REAL 2)|(Upper_Arc(K)) by JORDAN1:1;
    case A24: x in [.1/2,1.];
      then reconsider r=x as Real;
      A25: 1/2<=r & r<=1 by A24,TOPREAL5:1;
      then r*2>=1/2*2 by AXIOMS:25;
      then A26: 2*r-1>=0 by SQUARE_1:12;
        2*1>=2*r by A25,AXIOMS:25;
      then A27: 1+1-1>=2*r-1 by REAL_1:49;
     f2.x= (1-(2*r-1))*|[a,d]|+(2*r-1)*|[b,d]| by A17,A18,A24;
     then y in { (1-lambda)*|[a,d]| + lambda*|[b,d]| where lambda is Real:
            0 <= lambda & lambda <= 1 } by A18,A26,A27;
     then A28: y in LSeg(|[a,d]|,|[b,d]|) by TOPREAL1:def 4;
       Upper_Arc(K)= LSeg(|[a,c]|,|[a,d]|) \/ LSeg(|[a,d]|,|[b,d]|)
                               by A1,Th61;
     then y in Upper_Arc(K) by A28,XBOOLE_0:def 2;
     hence y in the carrier of (TOP-REAL 2)|(Upper_Arc(K)) by JORDAN1:1;
    end;
   hence y in the carrier of (TOP-REAL 2)|(Upper_Arc(K));
  end;
  then f2 is Function of the carrier of I[01],
          the carrier of (TOP-REAL 2)|(Upper_Arc(K)) by A17,BORSUK_1:83,FUNCT_2
:4;
  then reconsider f3=f2 as map of I[01],(TOP-REAL 2)|(Upper_Arc(K))
                          ;
  A29: 0 in [.0,1.] by TOPREAL5:1;
    0 in [.0,1/2.] by TOPREAL5:1;
  then A30: f3.0= (1-2*0)*|[a,c]|+(2*0)*|[a,d]| by A17,A29
         .= (1)*|[a,c]|+0.REAL 2 by EUCLID:33
         .= |[a,c]|+0.REAL 2 by EUCLID:33
         .= |[a,c]| by EUCLID:31
         .= W-min(K) by A1,Th56;
  A31: 1 in [.0,1.] by TOPREAL5:1;
    1 in [.1/2,1.] by TOPREAL5:1;
  then A32: f3.1= (1-(2*1-1))*|[a,d]|+(2*1-1)*|[b,d]| by A17,A31
         .= (0)*|[a,d]|+|[b,d]| by EUCLID:33
         .= (0.REAL 2) + |[b,d]| by EUCLID:33
         .= |[b,d]| by EUCLID:31
         .= E-max(K) by A1,Th56;
  A33: for r being Real st r in [.0,1/2.] holds
    f3.r=(1-2*r)*|[a,c]|+(2*r)*|[a,d]|
   proof let r be Real;
    assume A34: r in [.0,1/2.];
     then A35: 0<=r & r<=1/2 by TOPREAL5:1;
     then r<=1 by AXIOMS:22;
     then r in [.0,1.] by A35,TOPREAL5:1;
    hence f3.r=(1-2*r)*|[a,c]|+(2*r)*|[a,d]| by A17,A34;
   end;
  A36: for r being Real st r in [.1/2,1.] holds
    f3.r=(1-(2*r-1))*|[a,d]|+(2*r-1)*|[b,d]|
   proof let r be Real;
    assume A37: r in [.1/2,1.];
     then A38: 1/2<=r & r<=1 by TOPREAL5:1;
     then 0<=r by AXIOMS:22;
     then r in [.0,1.] by A38,TOPREAL5:1;
    hence f3.r=(1-(2*r-1))*|[a,d]|+(2*r-1)*|[b,d]| by A17,A37;
   end;
  A39: (for p being Point of TOP-REAL 2 st p in LSeg(|[a,c]|,|[a,d]|)
  holds 0<=((p`2)-c)/(d-c)/2 & ((p`2)-c)/(d-c)/2<=1
        & f3.(((p`2)-c)/(d-c)/2)=p)
  proof let p be Point of TOP-REAL 2;
    assume A40: p in LSeg(|[a,c]|,|[a,d]|);
     then p in LSeg(|[a2,c2]|,|[a2,d2]|);
     then A41: p`1=a by TOPREAL3:17;
     A42: (|[a,c]|)`2= c by EUCLID:56;
       (|[a,d]|)`2= d by EUCLID:56;
     then A43: c <=p`2 & p`2<=d by A1,A40,A42,TOPREAL1:10;
     A44: d-c>0 by A1,SQUARE_1:11;
       (p`2) -c >=0 by A43,SQUARE_1:12;
     then ((p`2) -c)/(d-c) >=0/(d-c) by A44,REAL_1:73;
     then A45: ((p`2) -c)/(d-c)/2 >=0/2 by REAL_1:73;
     A46: d-c>0 by A1,SQUARE_1:11;
       (p`2) -c <=d-c by A43,REAL_1:49;
     then ((p`2) -c)/(d-c) <=(d-c)/(d-c) by A46,REAL_1:73;
     then ((p`2) -c)/(d-c) <=1 by A46,XCMPLX_1:60;
     then A47: ((p`2) -c)/(d-c)/2 <=1/2 by REAL_1:73;
     set r=((p`2)-c)/(d-c)/2;
     A48: ((2*r)*d)-(2*r)*c = (2*r)*(d-c) by XCMPLX_1:40;
       r in [.0,1/2.] by A45,A47,TOPREAL5:1;
     then f3.(((p`2)-c)/(d-c)/2)=(1-2*r)*|[a,c]|+(2*r)*|[a,d]| by A33
       .=|[(1-2*r)*a,(1-2*r)*c]|+(2*r)*|[a,d]| by EUCLID:62
       .=|[(1-2*r)*a,(1-2*r)*c]|+|[(2*r)*a,(2*r)*d]| by EUCLID:62
       .=|[(1-2*r)*a+(2*r)*a,(1-2*r)*c+(2*r)*d]| by EUCLID:60
       .=|[1*a-(2*r)*a+(2*r)*a,(1-2*r)*c+(2*r)*d]| by XCMPLX_1:40
       .=|[1*a,(1-2*r)*c+(2*r)*d]| by XCMPLX_1:27
       .=|[a,1*c-(2*r)*c+(2*r)*d]| by XCMPLX_1:40
       .=|[a,1*c+-((2*r)*c)+(2*r)*d]| by XCMPLX_0:def 8
       .=|[a,1*c+((-((2*r)*c))+(2*r)*d)]| by XCMPLX_1:1
       .=|[a,1*c+((2*r)*(d-c))]| by A48,XCMPLX_0:def 8
       .=|[a,1*c+(((p`2)-c)/(d-c))*(d-c)]| by XCMPLX_1:88
       .=|[a,1*c+((p`2)-c)]| by A46,XCMPLX_1:88
       .=|[p`1,p`2]| by A41,XCMPLX_1:27
       .= p by EUCLID:57;
    hence 0<=((p`2)-c)/(d-c)/2 & ((p`2)-c)/(d-c)/2<=1
        & f3.(((p`2)-c)/(d-c)/2)=p by A45,A47,AXIOMS:22;
  end;
 A49: for p being Point of TOP-REAL 2 st p in LSeg(|[a,d]|,|[b,d]|)
  holds 0<=((p`1)-a)/(b-a)/2+1/2 & ((p`1)-a)/(b-a)/2+1/2<=1
        & f3.(((p`1)-a)/(b-a)/2+1/2)=p
 proof let p be Point of TOP-REAL 2;
  assume A50: p in LSeg(|[a,d]|,|[b,d]|);
     then p in LSeg(|[a2,d2]|,|[b2,d2]|);
     then A51: p`2=d by TOPREAL3:18;
     A52: (|[a,d]|)`1= a by EUCLID:56;
       (|[b,d]|)`1= b by EUCLID:56;
     then A53: a <=p`1 & p`1<=b by A1,A50,A52,TOPREAL1:9;
     A54: b-a>0 by A1,SQUARE_1:11;
       (p`1) -a >=0 by A53,SQUARE_1:12;
     then ((p`1) -a)/(b-a) >=0/(b-a) by A54,REAL_1:73;
     then ((p`1) -a)/(b-a)/2 >=0/2 by REAL_1:73;
     then A55: ((p`1) -a)/(b-a)/2+1/2 >=0+1/2 by REAL_1:55;
     A56: b-a>0 by A1,SQUARE_1:11;
       (p`1) -a <=b-a by A53,REAL_1:49;
     then ((p`1) -a)/(b-a) <=(b-a)/(b-a) by A56,REAL_1:73;
     then ((p`1) -a)/(b-a) <=1 by A56,XCMPLX_1:60;
     then ((p`1) -a)/(b-a)/2 <=1/2 by REAL_1:73;
     then A57: ((p`1) -a)/(b-a)/2+1/2 <=1/2+1/2 by REAL_1:55;
     set r=((p`1)-a)/(b-a)/2+1/2;
     A58: r= (((p`1)-a)/(b-a)+1)/2 by XCMPLX_1:63;
       r in [.1/2,1.] by A55,A57,TOPREAL5:1;
     then f3.(((p`1)-a)/(b-a)/2+1/2)=(1-(2*r-1))*|[a,d]|+(2*r-1)*|[b,d]| by A36
       .=|[(1-(2*r-1))*a,(1-(2*r-1))*d]|+((2*r-1))*|[b,d]| by EUCLID:62
       .=|[(1-(2*r-1))*a,(1-(2*r-1))*d]|+|[((2*r-1))*b,((2*r-1))*d]|
                                     by EUCLID:62
       .=|[(1-(2*r-1))*a+((2*r-1))*b,(1-(2*r-1))*d+((2*r-1))*d]| by EUCLID:60
       .=|[(1-(2*r-1))*a+((2*r-1))*b,1*d-(2*r-1)*d+((2*r-1))*d]| by XCMPLX_1:40
       .=|[(1-(2*r-1))*a+((2*r-1))*b,1*d]| by XCMPLX_1:27
       .=|[1*a-(2*r-1)*a+((2*r-1))*b,d]| by XCMPLX_1:40
       .=|[1*a+-(((2*r-1))*a)+((2*r-1))*b,d]| by XCMPLX_0:def 8
       .=|[1*a+((-(((2*r-1))*a))+((2*r-1))*b),d]| by XCMPLX_1:1
       .=|[1*a+((((2*r-1))*b)-((2*r-1))*a),d]| by XCMPLX_0:def 8
       .=|[1*a+(((2*r-1))*(b-a)),d]| by XCMPLX_1:40
       .=|[1*a+(((p`1)-a)/(b-a)+1-1)*(b-a),d]| by A58,XCMPLX_1:88
       .=|[1*a+(((p`1)-a)/(b-a))*(b-a),d]| by XCMPLX_1:26
       .=|[1*a+((p`1)-a),d]| by A56,XCMPLX_1:88
       .=|[p`1,p`2]| by A51,XCMPLX_1:27
       .= p by EUCLID:57;
  hence 0<=((p`1)-a)/(b-a)/2+1/2 & ((p`1)-a)/(b-a)/2+1/2<=1
        & f3.(((p`1)-a)/(b-a)/2+1/2)=p by A55,A57,AXIOMS:22;
 end;
  reconsider B00=[.0,1.] as Subset of R^1
                          by TOPMETR:24;
  reconsider B01=B00 as non empty Subset of R^1
              by TOPREAL5:1;
    I[01]=(R^1)|B01 by TOPMETR:26,27;
  then consider h1 being map of I[01],R^1 such that
  A59: (for p being Point of I[01] holds
  h1.p=p) & h1 is continuous by Th14;
  consider h2 being map of I[01],R^1 such that
  A60: (for p being Point of I[01],r1 being real number st
  h1.p=r1 holds h2.p=2*r1) & h2 is continuous by A59,JGRAPH_2:33;
  consider h5 being map of I[01],R^1 such that
  A61: (for p being Point of I[01],r1 being real number st
  h2.p=r1 holds h5.p=1-r1) & h5 is continuous by A60,Th16;
  consider h3 being map of I[01],R^1 such that
  A62: (for p being Point of I[01],r1 being real number st
  h2.p=r1 holds h3.p=r1-1) & h3 is continuous by A60,Th15;
  consider h4 being map of I[01],R^1 such that
  A63: (for p being Point of I[01],r1 being real number st
  h3.p=r1 holds h4.p=1-r1) & h4 is continuous by A62,Th16;
  consider g1 being map of I[01],TOP-REAL 2 such that
  A64: (for r being Point of I[01] holds
       g1.r=(h5.r)*|[a,c]|+(h2.r)*|[a,d]|) &
   g1 is continuous by A60,A61,Th21;
  A65: for r being Point of I[01],s being real number st r=s holds
       g1.r=(1-2*s)*|[a,c]|+(2*s)*|[a,d]|
    proof let r be Point of I[01],s be real number;
      assume A66: r=s;
        g1.r=(h5.r)*|[a,c]|+(h2.r)*|[a,d]| by A64
         .=(1-(h2.r))*|[a,c]|+(h2.r)*|[a,d]| by A61
         .=(1-2*(h1.r))*|[a,c]|+(h2.r)*|[a,d]| by A60
         .=(1-2*(h1.r))*|[a,c]|+(2*(h1.r))*|[a,d]| by A60
         .=(1-2*s)*|[a,c]|+(2*(h1.r))*|[a,d]| by A59,A66
         .=(1-2*s)*|[a,c]|+(2*s)*|[a,d]| by A59,A66;
     hence g1.r=(1-2*s)*|[a,c]|+(2*s)*|[a,d]|;
    end;

  consider g2 being map of I[01],TOP-REAL 2 such that
  A67: (for r being Point of I[01] holds
       g2.r=(h4.r)*|[a,d]|+(h3.r)*|[b,d]|) &
   g2 is continuous by A62,A63,Th21;

  A68: for r being Point of I[01],s being real number st r=s holds
       g2.r=(1-(2*s-1))*|[a,d]|+(2*s-1)*|[b,d]|
    proof let r be Point of I[01],s be real number;
      assume A69: r=s;
        g2.r=(h4.r)*|[a,d]|+(h3.r)*|[b,d]| by A67
         .=(1-h3.r)*|[a,d]|+(h3.r)*|[b,d]| by A63
         .=(1-((h2.r)-1))*|[a,d]|+(h3.r)*|[b,d]| by A62
         .=(1-((h2.r)-1))*|[a,d]|+((h2.r)-1)*|[b,d]| by A62
         .=(1-(2*(h1.r)-1))*|[a,d]|+((h2.r)-1)*|[b,d]| by A60
         .=(1-(2*(h1.r)-1))*|[a,d]|+(2*(h1.r)-1)*|[b,d]| by A60
         .=(1-(2*s-1))*|[a,d]|+(2*(h1.r)-1)*|[b,d]| by A59,A69
         .=(1-(2*s-1))*|[a,d]|+(2*s-1)*|[b,d]| by A59,A69;
     hence g2.r=(1-(2*s-1))*|[a,d]|+(2*s-1)*|[b,d]|;
    end;
     [.0,1/2.] c= [.0,1.] by A2,XBOOLE_1:7;
  then reconsider B11=[.0,1/2.] as non empty Subset of I[01]
                       by BORSUK_1:83,TOPREAL5:1;
  A70: dom (g1|B11)=dom g1 /\ B11 by FUNCT_1:68
  .= (the carrier of I[01]) /\ B11 by FUNCT_2:def 1
  .=B11 by XBOOLE_1:28
  .=the carrier of (I[01]|B11) by JORDAN1:1;
    rng (g1|B11) c= rng g1 by FUNCT_1:76;
  then rng (g1|B11) c= the carrier of TOP-REAL 2 by XBOOLE_1:1;
  then g1|B11 is Function of the carrier of (I[01]|B11),the carrier of
TOP-REAL 2
                              by A70,FUNCT_2:4;
  then reconsider g11=g1|B11 as map of I[01]|B11,TOP-REAL 2 ;
  A71: TOP-REAL 2=(TOP-REAL 2)|([#](TOP-REAL 2)) by TSEP_1:3;
  then A72: g11 is continuous by A64,BORSUK_4:69;
    [.1/2,1.] c= the carrier of I[01] by A2,BORSUK_1:83,XBOOLE_1:7;
  then reconsider B22=[.1/2,1.] as non empty Subset of I[01] by TOPREAL5:1;
  A73: dom (g2|B22)=dom g2 /\ B22 by FUNCT_1:68
  .= (the carrier of I[01]) /\ B22 by FUNCT_2:def 1
  .=B22 by XBOOLE_1:28
  .=the carrier of (I[01]|B22) by JORDAN1:1;
    rng (g2|B22) c= rng g2 by FUNCT_1:76;
  then rng (g2|B22) c= the carrier of TOP-REAL 2 by XBOOLE_1:1;
  then g2|B22 is Function of the carrier of (I[01]|B22),the carrier of
TOP-REAL 2
                              by A73,FUNCT_2:4;
  then reconsider g22=g2|B22 as map of I[01]|B22,TOP-REAL 2 ;
  A74: g22 is continuous by A67,A71,BORSUK_4:69;
  A75: B11=[#](I[01]|B11) by PRE_TOPC:def 10;
  A76: B22=[#](I[01]|B22) by PRE_TOPC:def 10;
  A77: B11 is closed by Th12;
  A78: B22 is closed by Th12;
    B11 \/ B22=[.0,1.] by HEINE:2;
  then A79: [#](I[01]|B11) \/ [#](I[01]|B22)=[#]I[01]
                   by A75,A76,BORSUK_1:83,PRE_TOPC:12;
    for p being set st p in ([#](I[01]|B11)) /\ ([#](I[01]|B22))
               holds g11.p = g22.p
  proof let p be set;assume
      p in ([#](I[01]|B11)) /\ ([#](I[01]|B22));
    then p in [#](I[01]|B11) & p in [#](I[01]|B22) by XBOOLE_0:def 3;
    then A80: p in B11 & p in B22 by PRE_TOPC:def 10;
    then reconsider rp=p as Real;
    A81: rp<=1/2 by A80,TOPREAL5:1;
      rp>=1/2 by A80,TOPREAL5:1;
    then rp=1/2 by A81,AXIOMS:21;
    then A82: 2*rp=1;
   thus g11.p=g1.p by A80,FUNCT_1:72
   .= (1-1)*|[a,c]|+(1)*|[a,d]| by A65,A80,A82
   .=0.REAL 2 +1*|[a,d]| by EUCLID:33
   .=(1-0)*|[a,d]| +(1-1)*|[b,d]| by EUCLID:33
   .=g2.p by A68,A80,A82 .=g22.p by A80,FUNCT_1:72;
  end;
  then consider h being map of I[01],TOP-REAL 2 such that
  A83: h=g11+*g22 & h is continuous
                       by A72,A74,A75,A76,A77,A78,A79,JGRAPH_2:9;
  A84: dom f3=dom h & dom f3=the carrier of I[01] by Th13;
    for x being set st x in dom f2 holds f3.x=h.x
   proof let x be set;
    assume A85: x in dom f2;
     then reconsider rx=x as Real by A84,BORSUK_1:83;
     A86: 0<=rx & rx<=1 by A84,A85,BORSUK_1:83,TOPREAL5:1;
       now per cases;
     case A87: rx<1/2;
      then A88: rx in [.0,1/2.] by A86,TOPREAL5:1;
        now assume rx in dom g22; then rx in B22 by A76,Th13;
       hence contradiction by A87,TOPREAL5:1;
      end;
      then h.rx=g11.rx by A83,FUNCT_4:12 .=g1.rx by A88,FUNCT_1:72
      .=(1-(2*rx))*|[a,c]|+(2*rx)*|[a,d]| by A65,A84,A85
      .=f3.rx by A33,A88;
      hence f3.x=h.x;
     case rx >=1/2;
      then A89: rx in [.1/2,1.] by A86,TOPREAL5:1;
      then rx in [#](I[01]|B22) by PRE_TOPC:def 10;
      then h.rx=g22.rx by A73,A83,FUNCT_4:14 .=g2.rx by A89,FUNCT_1:72
      .=(1-(2*rx-1))*|[a,d]|+(2*rx-1)*|[b,d]| by A68,A84,A85
      .=f3.rx by A36,A89;
      hence f3.x=h.x;
     end;
    hence f3.x=h.x;
   end;
  then f2=h by A84,FUNCT_1:9;
  then A90: f3 is continuous by A83,JGRAPH_1:63;
  A91: dom f3=[#](I[01]) by A17,BORSUK_1:83,PRE_TOPC:12;
    for x1,x2 being set st x1 in dom f3 & x2 in dom f3 & f3.x1=f3.x2
  holds x1=x2
  proof let x1,x2 be set;
   assume A92: x1 in dom f3 & x2 in dom f3 & f3.x1=f3.x2;
    then reconsider r1=x1 as Real by A17;
    reconsider r2=x2 as Real by A17,A92;
    A93: LSeg(|[a,c]|,|[a,d]|) /\ LSeg(|[a,d]|,|[b,d]|) = {|[a,d]|}
                  by A1,Th44;
      now per cases by A2,A17,A92,XBOOLE_0:def 2;
    case A94: x1 in [.0,1/2.] & x2 in [.0,1/2.];
      then f3.r1=(1-2*r1)*|[a,c]|+(2*r1)*|[a,d]| by A33;
      then (1-2*r2)*|[a,c]|+(2*r2)*|[a,d]|= (1-2*r1)*|[a,c]|+(2*r1)*|[a,d]|
                               by A33,A92,A94;
      then (1-2*r2)*|[a,c]|+(2*r2)*|[a,d]| -(2*r1)*|[a,d]|
      = (1-2*r1)*|[a,c]| by EUCLID:52;
      then (1-2*r2)*|[a,c]|+((2*r2)*|[a,d]| -(2*r1)*|[a,d]|)
      = (1-2*r1)*|[a,c]| by EUCLID:49;
      then (1-2*r2)*|[a,c]|+(2*r2-2*r1)*(|[a,d]|)
      = (1-2*r1)*|[a,c]| by EUCLID:54;
      then (2*r2-2*r1)*(|[a,d]|)+((1-2*r2)*|[a,c]|-(1-2*r1)*|[a,c]|)
      = (1-2*r1)*|[a,c]|-(1-2*r1)*|[a,c]| by EUCLID:49;
      then (2*r2-2*r1)*(|[a,d]|)+((1-2*r2)*|[a,c]|-(1-2*r1)*|[a,c]|)
      = 0.REAL 2 by EUCLID:46;
      then (2*r2-2*r1)*(|[a,d]|)+((1-2*r2)-(1-2*r1))*|[a,c]|
      = 0.REAL 2 by EUCLID:54;
      then (2*r2-2*r1)*(|[a,d]|)+(-(2*r2-2*r1))*|[a,c]|
      = 0.REAL 2 by XCMPLX_1:149;
      then (2*r2-2*r1)*(|[a,d]|)+-((2*r2-2*r1)*|[a,c]|)
      = 0.REAL 2 by EUCLID:44;
      then (2*r2-2*r1)*(|[a,d]|)-((2*r2-2*r1)*|[a,c]|)
      = 0.REAL 2 by EUCLID:45;
      then (2*r2-2*r1)*((|[a,d]|)-(|[a,c]|))
      = 0.REAL 2 by EUCLID:53;
      then (2*r2-2*r1)=0 or (|[a,d]|)-(|[a,c]|)=0.REAL 2 by EUCLID:35;
      then (2*r2-2*r1)=0 or |[a,d]|=|[a,c]| by EUCLID:47;
      then (2*r2-2*r1)=0 or d =|[a,c]|`2 by EUCLID:56;
      then 2*r2=2*r1 by A1,EUCLID:56,XCMPLX_1:15;
      then r2=r1*2/2 by XCMPLX_1:90;
     hence x1=x2 by XCMPLX_1:90;
    case A95: x1 in [.0,1/2.] & x2 in [.1/2,1.];
      then A96: f3.r1=(1-2*r1)*|[a,c]|+(2*r1)*|[a,d]| by A33;
      A97: 0<=r1 & r1<=1/2 by A95,TOPREAL5:1;
      then A98: r1*2<=1/2*2 by AXIOMS:25;
        2*0<=2*r1 by A97,AXIOMS:25;
      then f3.r1 in { (1-lambda)*|[a,c]| + lambda*|[a,d]| where lambda is Real
:
            0 <= lambda & lambda <= 1 } by A96,A98;
      then A99: f3.r1 in LSeg(|[a,c]|,|[a,d]|) by TOPREAL1:def 4;
      A100: f3.r2=(1-(2*r2-1))*|[a,d]|+(2*r2-1)*|[b,d]| by A36,A95;
      A101: 1/2<=r2 & r2<=1 by A95,TOPREAL5:1;
      then r2*2>=1/2*2 by AXIOMS:25;
      then A102: 2*r2-1>=0 by SQUARE_1:12;
        2*1>=2*r2 by A101,AXIOMS:25;
      then 1+1-1>=2*r2-1 by REAL_1:49;
      then f3.r2 in { (1-lambda)*|[a,d]| + lambda*|[b,d]| where lambda is Real
:
            0 <= lambda & lambda <= 1 } by A100,A102;
      then f3.r2 in LSeg(|[a,d]|,|[b,d]|) by TOPREAL1:def 4;
      then f3.r1 in LSeg(|[a,c]|,|[a,d]|) /\ LSeg(|[a,d]|,|[b,d]|)
                                 by A92,A99,XBOOLE_0:def 3;
      then A103: f3.r1= |[a,d]| by A93,TARSKI:def 1;
      then (1-2*r1)*|[a,c]|+(2*r1)*|[a,d]|+-(|[a,d]|)=0.REAL 2
                                   by A96,EUCLID:40;
      then (1-2*r1)*|[a,c]|+(2*r1)*|[a,d]|+(-1)*|[a,d]|=0.REAL 2
                                   by EUCLID:43;
      then (1-2*r1)*|[a,c]|+((2*r1)*|[a,d]|+(-1)*|[a,d]|)=0.REAL 2
                                   by EUCLID:30;
      then (1-2*r1)*|[a,c]|+((2*r1)+(-1))*|[a,d]|=0.REAL 2
                                   by EUCLID:37;
      then (1-2*r1)*|[a,c]|+(-(1-(2*r1)))*|[a,d]|=0.REAL 2 by XCMPLX_1:162;
      then (1-2*r1)*|[a,c]|+-((1-(2*r1))*|[a,d]|)=0.REAL 2 by EUCLID:44;
      then (1-2*r1)*|[a,c]|-((1-(2*r1))*|[a,d]|)=0.REAL 2 by EUCLID:45;
      then (1-2*r1)*(|[a,c]|-(|[a,d]|))=0.REAL 2 by EUCLID:53;
      then 1-2*r1=0 or (|[a,c]|-(|[a,d]|))=0.REAL 2 by EUCLID:35;
      then 1-2*r1=0 or |[a,c]|=|[a,d]| by EUCLID:47;
      then 1-2*r1=0 or c =|[a,d]|`2 by EUCLID:56;
      then 1=2*r1 by A1,EUCLID:56,XCMPLX_1:15;
      then A104: r1=1/2 by XCMPLX_1:90;
        (1-(2*r2-1))*|[a,d]|+(2*r2-1)*|[b,d]|+-(|[a,d]|)=0.REAL 2
                                   by A92,A100,A103,EUCLID:40;
      then (1-(2*r2-1))*|[a,d]|+((2*r2-1))*|[b,d]|+(-1)*|[a,d]|=0.REAL 2
                                   by EUCLID:43;
      then ((2*r2-1))*|[b,d]|+((1-(2*r2-1))*|[a,d]|+(-1)*|[a,d]|)=0.REAL 2
                                   by EUCLID:30;
      then ((2*r2-1))*|[b,d]|+((1-(2*r2-1))+(-1))*|[a,d]|=0.REAL 2
                                   by EUCLID:37;
      then ((2*r2-1))*|[b,d]|+(-1+1-(2*r2-1))*|[a,d]|=0.REAL 2 by XCMPLX_1:29;
      then ((2*r2-1))*|[b,d]|+(-(2*r2-1))*|[a,d]|=0.REAL 2 by XCMPLX_1:150;
      then ((2*r2-1))*|[b,d]|+-((2*r2-1)*|[a,d]|)=0.REAL 2 by EUCLID:44;
      then ((2*r2-1))*|[b,d]|-((2*r2-1)*|[a,d]|)=0.REAL 2 by EUCLID:45;
      then ((2*r2-1))*(|[b,d]|-(|[a,d]|))=0.REAL 2 by EUCLID:53;
      then (2*r2-1)=0 or (|[b,d]|-(|[a,d]|))=0.REAL 2 by EUCLID:35;
      then (2*r2-1)=0 or |[b,d]|=|[a,d]| by EUCLID:47;
      then (2*r2-1)=0 or b =|[a,d]|`1 by EUCLID:56;
      then 1=2*r2 by A1,EUCLID:56,XCMPLX_1:15;
     hence x1=x2 by A104,XCMPLX_1:90;
    case A105: x1 in [.1/2,1.] & x2 in [.0,1/2.];
      then A106: f3.r2=(1-2*r2)*|[a,c]|+(2*r2)*|[a,d]| by A33;
      A107: 0<=r2 & r2<=1/2 by A105,TOPREAL5:1;
      then A108: r2*2<=1/2*2 by AXIOMS:25;
        2*0<=2*r2 by A107,AXIOMS:25;
      then f3.r2 in { (1-lambda)*|[a,c]| + lambda*|[a,d]| where lambda is Real
:
            0 <= lambda & lambda <= 1 } by A106,A108;
      then A109: f3.r2 in LSeg(|[a,c]|,|[a,d]|) by TOPREAL1:def 4;
      A110: f3.r1=(1-(2*r1-1))*|[a,d]|+(2*r1-1)*|[b,d]| by A36,A105;
      A111: 1/2<=r1 & r1<=1 by A105,TOPREAL5:1;
      then r1*2>=1/2*2 by AXIOMS:25;
      then A112: 2*r1-1>=0 by SQUARE_1:12;
        2*1>=2*r1 by A111,AXIOMS:25;
      then 1+1-1>=2*r1-1 by REAL_1:49;
      then f3.r1 in { (1-lambda)*|[a,d]| + lambda*|[b,d]| where lambda is Real
:
            0 <= lambda & lambda <= 1 } by A110,A112;
      then f3.r1 in LSeg(|[a,d]|,|[b,d]|) by TOPREAL1:def 4;
      then f3.r2 in LSeg(|[a,c]|,|[a,d]|) /\ LSeg(|[a,d]|,|[b,d]|)
                                 by A92,A109,XBOOLE_0:def 3;
      then A113: f3.r2= |[a,d]| by A93,TARSKI:def 1;
      then (1-2*r2)*|[a,c]|+(2*r2)*|[a,d]|+-(|[a,d]|)=0.REAL 2
                                   by A106,EUCLID:40;
      then (1-2*r2)*|[a,c]|+(2*r2)*|[a,d]|+(-1)*|[a,d]|=0.REAL 2
                                   by EUCLID:43;
      then (1-2*r2)*|[a,c]|+((2*r2)*|[a,d]|+(-1)*|[a,d]|)=0.REAL 2
                                   by EUCLID:30;
      then (1-2*r2)*|[a,c]|+((2*r2)+(-1))*|[a,d]|=0.REAL 2
                                   by EUCLID:37;
      then (1-2*r2)*|[a,c]|+(-(1-(2*r2)))*|[a,d]|=0.REAL 2 by XCMPLX_1:162;
      then (1-2*r2)*|[a,c]|+-((1-(2*r2))*|[a,d]|)=0.REAL 2 by EUCLID:44;
      then (1-2*r2)*|[a,c]|-((1-(2*r2))*|[a,d]|)=0.REAL 2 by EUCLID:45;
      then (1-2*r2)*(|[a,c]|-(|[a,d]|))=0.REAL 2 by EUCLID:53;
      then 1-2*r2=0 or (|[a,c]|-(|[a,d]|))=0.REAL 2 by EUCLID:35;
      then 1-2*r2=0 or |[a,c]|=|[a,d]| by EUCLID:47;
      then 1-2*r2=0 or c =|[a,d]|`2 by EUCLID:56;
      then 1=2*r2 by A1,EUCLID:56,XCMPLX_1:15;
      then A114: r2=1/2 by XCMPLX_1:90;
        (1-(2*r1-1))*|[a,d]|+(2*r1-1)*|[b,d]|+-(|[a,d]|)=0.REAL 2
                                   by A92,A110,A113,EUCLID:40;
      then (1-(2*r1-1))*|[a,d]|+((2*r1-1))*|[b,d]|+(-1)*|[a,d]|=0.REAL 2
                                   by EUCLID:43;
      then ((2*r1-1))*|[b,d]|+((1-(2*r1-1))*|[a,d]|+(-1)*|[a,d]|)=0.REAL 2
                                   by EUCLID:30;
      then ((2*r1-1))*|[b,d]|+(-1+(1-(2*r1-1)))*|[a,d]|=0.REAL 2 by EUCLID:37;
      then ((2*r1-1))*|[b,d]|+(1+-1-(2*r1-1))*|[a,d]|=0.REAL 2 by XCMPLX_1:29;
      then ((2*r1-1))*|[b,d]|+(-(2*r1-1))*|[a,d]|=0.REAL 2 by XCMPLX_1:150;
      then ((2*r1-1))*|[b,d]|+-((2*r1-1)*|[a,d]|)=0.REAL 2 by EUCLID:44;
      then ((2*r1-1))*|[b,d]|-((2*r1-1)*|[a,d]|)=0.REAL 2 by EUCLID:45;
      then ((2*r1-1))*(|[b,d]|-(|[a,d]|))=0.REAL 2 by EUCLID:53;
      then (2*r1-1)=0 or (|[b,d]|-(|[a,d]|))=0.REAL 2 by EUCLID:35;
      then (2*r1-1)=0 or |[b,d]|=|[a,d]| by EUCLID:47;
      then (2*r1-1)=0 or b =|[a,d]|`1 by EUCLID:56;
      then 1=2*r1 by A1,EUCLID:56,XCMPLX_1:15;
     hence x1=x2 by A114,XCMPLX_1:90;
    case A115: x1 in [.1/2,1.] & x2 in [.1/2,1.];
      then f3.r1=(1-(2*r1-1))*|[a,d]|+((2*r1-1))*|[b,d]| by A36;
      then (1-(2*r2-1))*|[a,d]|+((2*r2-1))*|[b,d]|
      = (1-(2*r1-1))*|[a,d]|+((2*r1-1))*|[b,d]| by A36,A92,A115;
      then (1-(2*r2-1))*|[a,d]|+((2*r2-1))*|[b,d]| -((2*r1-1))*|[b,d]|
      = (1-(2*r1-1))*|[a,d]| by EUCLID:52;
      then (1-(2*r2-1))*|[a,d]|+(((2*r2-1))*|[b,d]| -((2*r1-1))*|[b,d]|)
      = (1-(2*r1-1))*|[a,d]| by EUCLID:49;
      then (1-(2*r2-1))*|[a,d]|+((2*r2-1)-(2*r1-1))*(|[b,d]|)
      = (1-(2*r1-1))*|[a,d]| by EUCLID:54;
      then ((2*r2-1)-(2*r1-1))*(|[b,d]|)
          +((1-(2*r2-1))*|[a,d]|-(1-(2*r1-1))*|[a,d]|)
      = (1-(2*r1-1))*|[a,d]|-(1-(2*r1-1))*|[a,d]| by EUCLID:49;
      then ((2*r2-1)-(2*r1-1))*(|[b,d]|)
          +((1-(2*r2-1))*|[a,d]|-(1-(2*r1-1))*|[a,d]|)
      = 0.REAL 2 by EUCLID:46;
      then ((2*r2-1)-(2*r1-1))*(|[b,d]|)+((1-(2*r2-1))-(1-(2*r1-1)))*|[a,d]|
      = 0.REAL 2 by EUCLID:54;
      then ((2*r2-1)-(2*r1-1))*(|[b,d]|)+(-((2*r2-1)-(2*r1-1)))*|[a,d]|
      = 0.REAL 2 by XCMPLX_1:149;
      then ((2*r2-1)-(2*r1-1))*(|[b,d]|)+-(((2*r2-1)-(2*r1-1))*|[a,d]|)
      = 0.REAL 2 by EUCLID:44;
      then ((2*r2-1)-(2*r1-1))*(|[b,d]|)-(((2*r2-1)-(2*r1-1))*|[a,d]|)
      = 0.REAL 2 by EUCLID:45;
      then ((2*r2-1)-(2*r1-1))*((|[b,d]|)-(|[a,d]|))
      = 0.REAL 2 by EUCLID:53;
      then ((2*r2-1)-(2*r1-1))=0 or (|[b,d]|)-(|[a,d]|)=0.REAL 2
                 by EUCLID:35;
      then ((2*r2-1)-(2*r1-1))=0 or |[b,d]|=|[a,d]| by EUCLID:47;
      then ((2*r2-1)-(2*r1-1))=0 or b =|[a,d]|`1 by EUCLID:56;
      then (2*r2-1)=(2*r1-1) by A1,EUCLID:56,XCMPLX_1:15;
      then 2*r2=2*r1-1+1 by XCMPLX_1:27;
      then 2*r2=2*r1 by XCMPLX_1:27;
      then r2=r1*2/2 by XCMPLX_1:90;
     hence x1=x2 by XCMPLX_1:90;
    end;
   hence x1=x2;
  end;
  then A116: f3 is one-to-one by FUNCT_1:def 8;
  A117: the carrier of ((TOP-REAL 2)|(Upper_Arc(K)))
    =[#]((TOP-REAL 2)|(Upper_Arc(K))) by PRE_TOPC:12;
    [#]((TOP-REAL 2)|(Upper_Arc(K))) c= rng f3
   proof let y be set;assume y in [#]((TOP-REAL 2)|(Upper_Arc(K)));
     then A118: y in Upper_Arc(K) by PRE_TOPC:def 10;
     then reconsider q=y as Point of TOP-REAL 2;
     A119: Upper_Arc(K)=LSeg(|[a,c]|,|[a,d]|) \/ LSeg(|[a,d]|,|[b,d]|)
                         by A1,Th61;
       now per cases by A118,A119,XBOOLE_0:def 2;
     case q in LSeg(|[a,c]|,|[a,d]|);
       then A120: 0<=((q`2)-c)/(d-c)/2 & ((q`2)-c)/(d-c)/2<=1
        & f3.(((q`2)-c)/(d-c)/2)=q by A39;
       then ((q`2)-c)/(d-c)/2 in [.0,1.] by TOPREAL5:1;
      hence y in rng f3 by A17,A120,FUNCT_1:def 5;
     case q in LSeg(|[a,d]|,|[b,d]|);
      then A121: 0<=((q`1)-a)/(b-a)/2+1/2 & ((q`1)-a)/(b-a)/2+1/2<=1
        & f3.(((q`1)-a)/(b-a)/2+1/2)=q by A49;
       then ((q`1)-a)/(b-a)/2+1/2 in [.0,1.] by TOPREAL5:1;
      hence y in rng f3 by A17,A121,FUNCT_1:def 5;
     end;
    hence y in rng f3;
   end;
  then A122: rng f3=[#]((TOP-REAL 2)|(Upper_Arc(K)))
                             by A117,XBOOLE_0:def 10;
  A123: I[01] is compact by HEINE:11,TOPMETR:27;
    ((TOP-REAL 2)|(Upper_Arc(K))) is_T2 by TOPMETR:3;
  then A124: f3 is_homeomorphism by A90,A91,A116,A122,A123,COMPTS_1:26;
    rng f3=Upper_Arc(K) by A122,PRE_TOPC:def 10;
 hence
  ex f being map of I[01],(TOP-REAL 2)|(Upper_Arc(K)) st f is_homeomorphism
& f.0=W-min(K) & f.1=E-max(K) & rng f=Upper_Arc(K) &
(for r being Real st r in [.0,1/2.] holds f.r=(1-2*r)*|[a,c]|+(2*r)*|[a,d]|)&
(for r being Real st r in [.1/2,1.] holds
    f.r=(1-(2*r-1))*|[a,d]|+(2*r-1)*|[b,d]|)&
(for p being Point of TOP-REAL 2 st p in LSeg(|[a,c]|,|[a,d]|)
  holds 0<=((p`2)-c)/(d-c)/2 & ((p`2)-c)/(d-c)/2<=1
        & f.(((p`2)-c)/(d-c)/2)=p)&
(for p being Point of TOP-REAL 2 st p in LSeg(|[a,d]|,|[b,d]|)
  holds 0<=((p`1)-a)/(b-a)/2+1/2 & ((p`1)-a)/(b-a)/2+1/2<=1
        & f.(((p`1)-a)/(b-a)/2+1/2)=p) by A30,A32,A33,A36,A39,A49,A124;
end;

theorem Th64: for K being non empty compact Subset of TOP-REAL 2,
a,b,c,d being real number,p1,p2 being Point of TOP-REAL 2
 st K=rectangle(a,b,c,d) & a<b & c <d
ex f being map of I[01],(TOP-REAL 2)|(Lower_Arc(K)) st f is_homeomorphism
& f.0=E-max(K) & f.1=W-min(K) & rng f=Lower_Arc(K) &
(for r being Real st r in [.0,1/2.] holds f.r=(1-2*r)*|[b,d]|+(2*r)*|[b,c]|)&
(for r being Real st r in [.1/2,1.] holds
    f.r=(1-(2*r-1))*|[b,c]|+(2*r-1)*|[a,c]|)&
(for p being Point of TOP-REAL 2 st p in LSeg(|[b,d]|,|[b,c]|)
  holds 0<=((p`2)-d)/(c-d)/2 & ((p`2)-d)/(c-d)/2<=1
        & f.(((p`2)-d)/(c-d)/2)=p)&
(for p being Point of TOP-REAL 2 st p in LSeg(|[b,c]|,|[a,c]|)
  holds 0<=((p`1)-b)/(a-b)/2+1/2 & ((p`1)-b)/(a-b)/2+1/2<=1
        & f.(((p`1)-b)/(a-b)/2+1/2)=p)
proof let K be non empty compact Subset of TOP-REAL 2,
a,b,c,d be real number,p1,p2 be Point of TOP-REAL 2;
 assume A1: K=rectangle(a,b,c,d) & a<b & c <d;
  reconsider a2=a,b2=b,c2=c,d2=d as Real by XREAL_0:def 1;
  defpred P[set,set] means for r being Real st $1=r holds
         (r in [.0,1/2.] implies $2=(1-2*r)*|[b,d]|+(2*r)*|[b,c]|) &
         (r in [.1/2,1.] implies $2=(1-(2*r-1))*|[b,c]|+(2*r-1)*|[a,c]|);
  A2: [.0,1.]=[.0,1/2.] \/ [.1/2,1.] by HEINE:2;
  A3: for x,y1,y2 being set st x in [.0,1.] & P[x,y1] & P[x,y2]
                    holds y1 = y2
   proof let x,y1,y2 be set;
    assume A4: x in [.0,1.] & P[x,y1] & P[x,y2];
      now per cases by A2,A4,XBOOLE_0:def 2;
    case A5: x in [.0,1/2.];
      then reconsider r=x as Real;
        y1= (1-2*r)*|[b,d]|+(2*r)*|[b,c]| by A4,A5;
     hence y1=y2 by A4,A5;
    case A6: x in [.1/2,1.];
      then reconsider r=x as Real;
        y1= (1-(2*r-1))*|[b,c]|+(2*r-1)*|[a,c]| by A4,A6;
     hence y1=y2 by A4,A6;
     end;
    hence y1 = y2;
   end;
  A7: for x being set st x in [.0,1.] ex y being set st P[x,y]
   proof let x be set;assume
    A8: x in [.0,1.];
      now per cases by A2,A8,XBOOLE_0:def 2;
    case A9: x in [.0,1/2.];
      then reconsider r=x as Real;
      A10: 0<=r & r<=1/2 by A9,TOPREAL5:1;
      set y0= (1-2*r)*|[b,d]|+(2*r)*|[b,c]|;
        r in [.1/2,1.] implies y0=(1-(2*r-1))*|[b,c]|+(2*r-1)*|[a,c]|
       proof assume r in [.1/2,1.]; then 1/2 <=r & r<=1 by TOPREAL5:1;
         then A11: r=1/2 by A10,AXIOMS:21;
         then A12: y0= (0)*|[b,d]|+|[b,c]| by EUCLID:33
         .= (0.REAL 2) + |[b,c]| by EUCLID:33
         .= |[b,c]| by EUCLID:31;
           (1-(2*r-1))*|[b,c]|+(2*r-1)*|[a,c]|
          = (1)*|[b,c]|+0.REAL 2 by A11,EUCLID:33
         .= |[b,c]|+0.REAL 2 by EUCLID:33
         .= |[b,c]| by EUCLID:31;
        hence y0=(1-(2*r-1))*|[b,c]|+(2*r-1)*|[a,c]| by A12;
       end;
      then for r2 being Real st x=r2 holds
         (r2 in [.0,1/2.] implies y0=(1-2*r2)*|[b,d]|+(2*r2)*|[b,c]|) &
         (r2 in [.1/2,1.] implies y0=(1-(2*r2-1))*|[b,c]|+(2*r2-1)*|[a,c]|);
     hence ex y being set st P[x,y];
    case A13: x in [.1/2,1.];
      then reconsider r=x as Real;
      A14: 1/2<=r & r<=1 by A13,TOPREAL5:1;
      set y0= (1-(2*r-1))*|[b,c]|+(2*r-1)*|[a,c]|;
        r in [.0,1/2.] implies y0=(1-2*r)*|[b,d]|+(2*r)*|[b,c]|
       proof assume r in [.0,1/2.]; then 0 <=r & r<=1/2 by TOPREAL5:1;
         then A15: r=1/2 by A14,AXIOMS:21;
         then A16: y0= |[b,c]|+(0)*|[a,c]| by EUCLID:33
         .= |[b,c]|+(0.REAL 2) by EUCLID:33
         .= |[b,c]| by EUCLID:31;
           (1-2*r)*|[b,d]|+(2*r)*|[b,c]|
          = 0.REAL 2+(1)*|[b,c]| by A15,EUCLID:33
         .= 0.REAL 2+|[b,c]| by EUCLID:33
         .= |[b,c]| by EUCLID:31;
        hence y0=(1-2*r)*|[b,d]|+(2*r)*|[b,c]| by A16;
       end;
      then for r2 being Real st x=r2 holds
         (r2 in [.0,1/2.] implies y0=(1-2*r2)*|[b,d]|+(2*r2)*|[b,c]|) &
         (r2 in [.1/2,1.] implies y0=(1-(2*r2-1))*|[b,c]|+(2*r2-1)*|[a,c]|);
     hence ex y being set st P[x,y];
     end;
    hence ex y being set st P[x,y];
   end;
    ex f2 being Function st dom f2 = [.0,1.] &
     for x being set st x in [.0,1.] holds P[x,f2.x] from FuncEx(A3,A7);
  then consider f2 being Function such that
  A17: dom f2 = [.0,1.] &
   for x being set st x in [.0,1.] holds P[x,f2.x];
    rng f2 c= the carrier of (TOP-REAL 2)|(Lower_Arc(K))
  proof let y be set;assume y in rng f2;
    then consider x being set such that
    A18: x in dom f2 & y=f2.x by FUNCT_1:def 5;
      now per cases by A2,A17,A18,XBOOLE_0:def 2;
    case A19: x in [.0,1/2.];
      then reconsider r=x as Real;
      A20: 0<=r & r<=1/2 by A19,TOPREAL5:1;
      then A21: r*2<=1/2*2 by AXIOMS:25;
      A22: 2*0<=2*r by A20,AXIOMS:25;
     f2.x= (1-2*r)*|[b,d]|+(2*r)*|[b,c]| by A17,A18,A19;
     then y in { (1-lambda)*|[b,d]| + lambda*|[b,c]| where lambda is Real:
            0 <= lambda & lambda <= 1 } by A18,A21,A22;
     then A23: y in LSeg(|[b,d]|,|[b,c]|) by TOPREAL1:def 4;
       Lower_Arc(K)= LSeg(|[a,c]|,|[b,c]|) \/ LSeg(|[b,c]|,|[b,d]|)
                               by A1,Th62;
     then y in Lower_Arc(K) by A23,XBOOLE_0:def 2;
     hence y in the carrier of (TOP-REAL 2)|(Lower_Arc(K)) by JORDAN1:1;
    case A24: x in [.1/2,1.];
      then reconsider r=x as Real;
      A25: 1/2<=r & r<=1 by A24,TOPREAL5:1;
      then r*2>=1/2*2 by AXIOMS:25;
      then A26: 2*r-1>=0 by SQUARE_1:12;
        2*1>=2*r by A25,AXIOMS:25;
      then A27: 1+1-1>=2*r-1 by REAL_1:49;
     f2.x= (1-(2*r-1))*|[b,c]|+(2*r-1)*|[a,c]| by A17,A18,A24;
     then y in { (1-lambda)*|[b,c]| + lambda*|[a,c]| where lambda is Real:
            0 <= lambda & lambda <= 1 } by A18,A26,A27;
     then A28: y in LSeg(|[b,c]|,|[a,c]|) by TOPREAL1:def 4;
       Lower_Arc(K)= LSeg(|[a,c]|,|[b,c]|) \/ LSeg(|[b,c]|,|[b,d]|)
                               by A1,Th62;
     then y in Lower_Arc(K) by A28,XBOOLE_0:def 2;
     hence y in the carrier of (TOP-REAL 2)|(Lower_Arc(K)) by JORDAN1:1;
    end;
   hence y in the carrier of (TOP-REAL 2)|(Lower_Arc(K));
  end;
  then f2 is Function of the carrier of I[01],
          the carrier of (TOP-REAL 2)|(Lower_Arc(K)) by A17,BORSUK_1:83,FUNCT_2
:4;
  then reconsider f3=f2 as map of I[01],(TOP-REAL 2)|(Lower_Arc(K))
                          ;
  A29: 0 in [.0,1.] by TOPREAL5:1;
    0 in [.0,1/2.] by TOPREAL5:1;
  then A30: f3.0= (1-2*0)*|[b,d]|+(2*0)*|[b,c]| by A17,A29
         .= (1)*|[b,d]|+0.REAL 2 by EUCLID:33
         .= |[b,d]|+0.REAL 2 by EUCLID:33
         .= |[b,d]| by EUCLID:31
         .= E-max(K) by A1,Th56;
  A31: 1 in [.0,1.] by TOPREAL5:1;
    1 in [.1/2,1.] by TOPREAL5:1;
  then A32: f3.1= (1-(2*1-1))*|[b,c]|+(2*1-1)*|[a,c]| by A17,A31
         .= (0)*|[b,c]|+|[a,c]| by EUCLID:33
         .= (0.REAL 2) + |[a,c]| by EUCLID:33
         .= |[a,c]| by EUCLID:31
         .= W-min(K) by A1,Th56;
  A33: for r being Real st r in [.0,1/2.] holds
    f3.r=(1-2*r)*|[b,d]|+(2*r)*|[b,c]|
   proof let r be Real;
    assume A34: r in [.0,1/2.];
     then A35: 0<=r & r<=1/2 by TOPREAL5:1;
     then r<=1 by AXIOMS:22;
     then r in [.0,1.] by A35,TOPREAL5:1;
    hence f3.r=(1-2*r)*|[b,d]|+(2*r)*|[b,c]| by A17,A34;
   end;
  A36: for r being Real st r in [.1/2,1.] holds
    f3.r=(1-(2*r-1))*|[b,c]|+(2*r-1)*|[a,c]|
   proof let r be Real;
    assume A37: r in [.1/2,1.];
     then A38: 1/2<=r & r<=1 by TOPREAL5:1;
     then 0<=r by AXIOMS:22;
     then r in [.0,1.] by A38,TOPREAL5:1;
    hence f3.r=(1-(2*r-1))*|[b,c]|+(2*r-1)*|[a,c]| by A17,A37;
   end;
  A39: (for p being Point of TOP-REAL 2 st p in LSeg(|[b,d]|,|[b,c]|)
  holds 0<=((p`2)-d)/(c-d)/2 & ((p`2)-d)/(c-d)/2<=1
        & f3.(((p`2)-d)/(c-d)/2)=p)
  proof let p be Point of TOP-REAL 2;
    assume A40: p in LSeg(|[b,d]|,|[b,c]|);
     then p in LSeg(|[b2,d2]|,|[b2,c2]|);
     then A41: p`1=b by TOPREAL3:17;
     A42: (|[b,d]|)`2= d by EUCLID:56;
       (|[b,c]|)`2= c by EUCLID:56;
     then A43: c <=p`2 & p`2<=d by A1,A40,A42,TOPREAL1:10;
       d-c>0 by A1,SQUARE_1:11;
     then -(d-c)< -0 by REAL_1:50;
     then 0-(d-c) < 0 by XCMPLX_1:150;
     then 0-d+c <0 by XCMPLX_1:37;
     then -d+c <0 by XCMPLX_1:150;
     then A44: c -d <0 by XCMPLX_0:def 8;
       d-(p`2) >=0 by A43,SQUARE_1:12;
     then -(d-(p`2)) <= -0 by REAL_1:50;
     then 0-(d-(p`2)) <= 0 by XCMPLX_1:150;
     then 0-d+p`2 <=0 by XCMPLX_1:37;
     then -d+p`2 <=0 by XCMPLX_1:150;
     then (p`2) -d <=0 by XCMPLX_0:def 8;
     then ((p`2) -d)/(c-d) >=0/(c-d) by A44,REAL_1:74;
     then A45: ((p`2) -d)/(c-d)/2 >=0/2 by REAL_1:73;
       (p`2) -d >=c-d by A43,REAL_1:49;
     then ((p`2) -d)/(c-d) <=(c-d)/(c-d) by A44,REAL_1:74;
     then ((p`2) -d)/(c-d) <=1 by A44,XCMPLX_1:60;
     then A46: ((p`2) -d)/(c-d)/2 <=1/2 by REAL_1:73;
     set r=((p`2)-d)/(c-d)/2;
       r in [.0,1/2.] by A45,A46,TOPREAL5:1;
     then f3.(((p`2)-d)/(c-d)/2)=(1-2*r)*|[b,d]|+(2*r)*|[b,c]| by A33
       .=|[(1-2*r)*b,(1-2*r)*d]|+(2*r)*|[b,c]| by EUCLID:62
       .=|[(1-2*r)*b,(1-2*r)*d]|+|[(2*r)*b,(2*r)*c]| by EUCLID:62
       .=|[(1-2*r)*b+(2*r)*b,(1-2*r)*d+(2*r)*c]| by EUCLID:60
       .=|[1*b-(2*r)*b+(2*r)*b,(1-2*r)*d+(2*r)*c]| by XCMPLX_1:40
       .=|[1*b,(1-2*r)*d+(2*r)*c]| by XCMPLX_1:27
       .=|[b,1*d-(2*r)*d+(2*r)*c]| by XCMPLX_1:40
       .=|[b,1*d+-((2*r)*d)+(2*r)*c]| by XCMPLX_0:def 8
       .=|[b,1*d+((-((2*r)*d))+(2*r)*c)]| by XCMPLX_1:1
       .=|[b,1*d+(((2*r)*c)-(2*r)*d)]| by XCMPLX_0:def 8
       .=|[b,1*d+((2*r)*(c-d))]| by XCMPLX_1:40
       .=|[b,1*d+(((p`2)-d)/(c-d))*(c-d)]| by XCMPLX_1:88
       .=|[b,1*d+((p`2)-d)]| by A44,XCMPLX_1:88
       .=|[p`1,p`2]| by A41,XCMPLX_1:27
       .= p by EUCLID:57;
    hence 0<=((p`2)-d)/(c-d)/2 & ((p`2)-d)/(c-d)/2<=1
        & f3.(((p`2)-d)/(c-d)/2)=p by A45,A46,AXIOMS:22;
  end;
 A47: for p being Point of TOP-REAL 2 st p in LSeg(|[b,c]|,|[a,c]|)
  holds 0<=((p`1)-b)/(a-b)/2+1/2 & ((p`1)-b)/(a-b)/2+1/2<=1
        & f3.(((p`1)-b)/(a-b)/2+1/2)=p
 proof let p be Point of TOP-REAL 2;
  assume A48: p in LSeg(|[b,c]|,|[a,c]|);
     then p in LSeg(|[b2,c2]|,|[a2,c2]|);
     then A49: p`2=c by TOPREAL3:18;
     A50: (|[b,c]|)`1= b by EUCLID:56;
       (|[a,c]|)`1= a by EUCLID:56;
     then A51: a <=p`1 & p`1<=b by A1,A48,A50,TOPREAL1:9;
       b-a>0 by A1,SQUARE_1:11;
     then -(b-a)< -0 by REAL_1:50;
     then 0-(b-a) < 0 by XCMPLX_1:150;
     then 0-b+a <0 by XCMPLX_1:37;
     then -b+a <0 by XCMPLX_1:150;
     then A52: a -b <0 by XCMPLX_0:def 8;
       b-(p`1) >=0 by A51,SQUARE_1:12;
     then -(b-(p`1)) <= -0 by REAL_1:50;
     then 0-(b-(p`1)) <= 0 by XCMPLX_1:150;
     then 0-b+p`1 <=0 by XCMPLX_1:37;
     then -b+p`1 <=0 by XCMPLX_1:150;
     then (p`1) -b <=0 by XCMPLX_0:def 8;
     then ((p`1) -b)/(a-b) >=0/(a-b) by A52,REAL_1:74;
     then ((p`1) -b)/(a-b)/2 >=0/2 by REAL_1:73;
     then A53: ((p`1) -b)/(a-b)/2+1/2 >=0+1/2 by REAL_1:55;
       (p`1) -b >=a-b by A51,REAL_1:49;
     then ((p`1) -b)/(a-b) <=(a-b)/(a-b) by A52,REAL_1:74;
     then ((p`1) -b)/(a-b) <=1 by A52,XCMPLX_1:60;
     then ((p`1) -b)/(a-b)/2 <=1/2 by REAL_1:73;
     then A54: ((p`1) -b)/(a-b)/2+1/2 <=1/2+1/2 by REAL_1:55;
     set r=((p`1)-b)/(a-b)/2+1/2;
     A55: r= (((p`1)-b)/(a-b)+1)/2 by XCMPLX_1:63;
       r in [.1/2,1.] by A53,A54,TOPREAL5:1;
     then f3.(((p`1)-b)/(a-b)/2+1/2)=(1-(2*r-1))*|[b,c]|+(2*r-1)*|[a,c]| by A36
       .=|[(1-(2*r-1))*b,(1-(2*r-1))*c]|+((2*r-1))*|[a,c]| by EUCLID:62
       .=|[(1-(2*r-1))*b,(1-(2*r-1))*c]|+|[((2*r-1))*a,((2*r-1))*c]|
                                     by EUCLID:62
       .=|[(1-(2*r-1))*b+((2*r-1))*a,(1-(2*r-1))*c+((2*r-1))*c]| by EUCLID:60
       .=|[(1-(2*r-1))*b+((2*r-1))*a,1*c-(2*r-1)*c+((2*r-1))*c]| by XCMPLX_1:40
       .=|[(1-(2*r-1))*b+((2*r-1))*a,1*c]| by XCMPLX_1:27
       .=|[1*b-(2*r-1)*b+((2*r-1))*a,c]| by XCMPLX_1:40
       .=|[1*b+-(((2*r-1))*b)+((2*r-1))*a,c]| by XCMPLX_0:def 8
       .=|[1*b+((-(((2*r-1))*b))+((2*r-1))*a),c]| by XCMPLX_1:1
       .=|[1*b+((((2*r-1))*a)-((2*r-1))*b),c]| by XCMPLX_0:def 8
       .=|[1*b+(((2*r-1))*(a-b)),c]| by XCMPLX_1:40
       .=|[1*b+(((p`1)-b)/(a-b)+1-1)*(a-b),c]| by A55,XCMPLX_1:88
       .=|[1*b+(((p`1)-b)/(a-b))*(a-b),c]| by XCMPLX_1:26
       .=|[1*b+((p`1)-b),c]| by A52,XCMPLX_1:88
       .=|[p`1,p`2]| by A49,XCMPLX_1:27
       .= p by EUCLID:57;
  hence 0<=((p`1)-b)/(a-b)/2+1/2 & ((p`1)-b)/(a-b)/2+1/2<=1
        & f3.(((p`1)-b)/(a-b)/2+1/2)=p by A53,A54,AXIOMS:22;
 end;
  reconsider B00=[.0,1.] as Subset of R^1
                          by TOPMETR:24;
  reconsider B01=B00 as non empty Subset of R^1
              by TOPREAL5:1;
    I[01]=(R^1)|B01 by TOPMETR:26,27;
  then consider h1 being map of I[01],R^1 such that
  A56: (for p being Point of I[01] holds
  h1.p=p) & h1 is continuous by Th14;
  consider h2 being map of I[01],R^1 such that
  A57: (for p being Point of I[01],r1 being real number st
  h1.p=r1 holds h2.p=2*r1) & h2 is continuous by A56,JGRAPH_2:33;
  consider h5 being map of I[01],R^1 such that
  A58: (for p being Point of I[01],r1 being real number st
  h2.p=r1 holds h5.p=1-r1) & h5 is continuous by A57,Th16;
  consider h3 being map of I[01],R^1 such that
  A59: (for p being Point of I[01],r1 being real number st
  h2.p=r1 holds h3.p=r1-1) & h3 is continuous by A57,Th15;
  consider h4 being map of I[01],R^1 such that
  A60: (for p being Point of I[01],r1 being real number st
  h3.p=r1 holds h4.p=1-r1) & h4 is continuous by A59,Th16;
  consider g1 being map of I[01],TOP-REAL 2 such that
  A61: (for r being Point of I[01] holds
       g1.r=(h5.r)*|[b,d]|+(h2.r)*|[b,c]|) &
   g1 is continuous by A57,A58,Th21;
  A62: for r being Point of I[01],s being real number st r=s holds
       g1.r=(1-2*s)*|[b,d]|+(2*s)*|[b,c]|
    proof let r be Point of I[01],s be real number;
      assume A63: r=s;
        g1.r=(h5.r)*|[b,d]|+(h2.r)*|[b,c]| by A61
         .=(1-(h2.r))*|[b,d]|+(h2.r)*|[b,c]| by A58
         .=(1-2*(h1.r))*|[b,d]|+(h2.r)*|[b,c]| by A57
         .=(1-2*(h1.r))*|[b,d]|+(2*(h1.r))*|[b,c]| by A57
         .=(1-2*s)*|[b,d]|+(2*(h1.r))*|[b,c]| by A56,A63
         .=(1-2*s)*|[b,d]|+(2*s)*|[b,c]| by A56,A63;
     hence g1.r=(1-2*s)*|[b,d]|+(2*s)*|[b,c]|;
    end;

  consider g2 being map of I[01],TOP-REAL 2 such that
  A64: (for r being Point of I[01] holds
       g2.r=(h4.r)*|[b,c]|+(h3.r)*|[a,c]|) &
   g2 is continuous by A59,A60,Th21;

  A65: for r being Point of I[01],s being real number st r=s holds
       g2.r=(1-(2*s-1))*|[b,c]|+(2*s-1)*|[a,c]|
    proof let r be Point of I[01],s be real number;
      assume A66: r=s;
        g2.r=(h4.r)*|[b,c]|+(h3.r)*|[a,c]| by A64
         .=(1-h3.r)*|[b,c]|+(h3.r)*|[a,c]| by A60
         .=(1-((h2.r)-1))*|[b,c]|+(h3.r)*|[a,c]| by A59
         .=(1-((h2.r)-1))*|[b,c]|+((h2.r)-1)*|[a,c]| by A59
         .=(1-(2*(h1.r)-1))*|[b,c]|+((h2.r)-1)*|[a,c]| by A57
         .=(1-(2*(h1.r)-1))*|[b,c]|+(2*(h1.r)-1)*|[a,c]| by A57
         .=(1-(2*s-1))*|[b,c]|+(2*(h1.r)-1)*|[a,c]| by A56,A66
         .=(1-(2*s-1))*|[b,c]|+(2*s-1)*|[a,c]| by A56,A66;
     hence g2.r=(1-(2*s-1))*|[b,c]|+(2*s-1)*|[a,c]|;
    end;
     [.0,1/2.] c= [.0,1.] by A2,XBOOLE_1:7;
  then reconsider B11=[.0,1/2.] as non empty Subset of I[01]
                    by BORSUK_1:83,TOPREAL5:1;
  A67: dom (g1|B11)=dom g1 /\ B11 by FUNCT_1:68
  .= (the carrier of I[01]) /\ B11 by FUNCT_2:def 1
  .=B11 by XBOOLE_1:28
  .=the carrier of (I[01]|B11) by JORDAN1:1;
    rng (g1|B11) c= rng g1 by FUNCT_1:76;
  then rng (g1|B11) c= the carrier of TOP-REAL 2 by XBOOLE_1:1;
  then g1|B11 is Function of the carrier of (I[01]|B11),the carrier of
TOP-REAL 2
                              by A67,FUNCT_2:4;
  then reconsider g11=g1|B11 as map of I[01]|B11,TOP-REAL 2 ;
  A68: TOP-REAL 2=(TOP-REAL 2)|([#](TOP-REAL 2)) by TSEP_1:3;
  then A69: g11 is continuous by A61,BORSUK_4:69;
    [.1/2,1.] c= [.0,1.] by A2,XBOOLE_1:7;
  then reconsider B22=[.1/2,1.] as non empty Subset of I[01]
                    by BORSUK_1:83,TOPREAL5:1;
  A70: dom (g2|B22)=dom g2 /\ B22 by FUNCT_1:68
  .= (the carrier of I[01]) /\ B22 by FUNCT_2:def 1
  .=B22 by XBOOLE_1:28
  .=the carrier of (I[01]|B22) by JORDAN1:1;
    rng (g2|B22) c= rng g2 by FUNCT_1:76;
  then rng (g2|B22) c= the carrier of TOP-REAL 2 by XBOOLE_1:1;
  then g2|B22 is Function of the carrier of (I[01]|B22),the carrier of
TOP-REAL 2
                              by A70,FUNCT_2:4;
  then reconsider g22=g2|B22 as map of I[01]|B22,TOP-REAL 2 ;
  A71: g22 is continuous by A64,A68,BORSUK_4:69;
  A72: B11=[#](I[01]|B11) by PRE_TOPC:def 10;
  A73: B22=[#](I[01]|B22) by PRE_TOPC:def 10;
  A74: B11 is closed by Th12;
  A75: B22 is closed by Th12;
    B11 \/ B22=[.0,1.] by HEINE:2;
  then A76: [#](I[01]|B11) \/ [#](I[01]|B22)=[#]I[01]
                by A72,A73,BORSUK_1:83,PRE_TOPC:12;
    for p being set st p in ([#](I[01]|B11)) /\ ([#](I[01]|B22))
               holds g11.p = g22.p
  proof let p be set;assume
      p in ([#](I[01]|B11)) /\ ([#](I[01]|B22));
    then p in [#](I[01]|B11) & p in [#](I[01]|B22) by XBOOLE_0:def 3;
    then A77: p in B11 & p in B22 by PRE_TOPC:def 10;
    then reconsider rp=p as Real;
    A78: rp<=1/2 by A77,TOPREAL5:1;
      rp>=1/2 by A77,TOPREAL5:1;
    then rp=1/2 by A78,AXIOMS:21;
    then A79: 2*rp=1;
   thus g11.p=g1.p by A77,FUNCT_1:72
   .= (1-1)*|[b,d]|+(1)*|[b,c]| by A62,A77,A79
   .=0.REAL 2 +1*|[b,c]| by EUCLID:33
   .=(1-0)*|[b,c]| +(1-1)*|[a,c]| by EUCLID:33
   .=g2.p by A65,A77,A79 .=g22.p by A77,FUNCT_1:72;
  end;
  then consider h being map of I[01],TOP-REAL 2 such that
  A80: h=g11+*g22 & h is continuous
                       by A69,A71,A72,A73,A74,A75,A76,JGRAPH_2:9;
  A81: dom f3=dom h & dom f3=the carrier of I[01] by Th13;
    for x being set st x in dom f2 holds f3.x=h.x
   proof let x be set;
    assume A82: x in dom f2;
     then reconsider rx=x as Real by A81,BORSUK_1:83;
     A83: 0<=rx & rx<=1 by A81,A82,BORSUK_1:83,TOPREAL5:1;
       now per cases;
     case A84: rx<1/2;
      then A85: rx in [.0,1/2.] by A83,TOPREAL5:1;
        now assume rx in dom g22; then rx in B22 by A73,Th13;
       hence contradiction by A84,TOPREAL5:1;
      end;
      then h.rx=g11.rx by A80,FUNCT_4:12 .=g1.rx by A85,FUNCT_1:72
      .=(1-(2*rx))*|[b,d]|+(2*rx)*|[b,c]| by A62,A81,A82
      .=f3.rx by A33,A85;
      hence f3.x=h.x;
     case rx >=1/2;
      then A86: rx in [.1/2,1.] by A83,TOPREAL5:1;
      then rx in [#](I[01]|B22) by PRE_TOPC:def 10;
      then h.rx=g22.rx by A70,A80,FUNCT_4:14 .=g2.rx by A86,FUNCT_1:72
      .=(1-(2*rx-1))*|[b,c]|+(2*rx-1)*|[a,c]| by A65,A81,A82
      .=f3.rx by A36,A86;
      hence f3.x=h.x;
     end;
    hence f3.x=h.x;
   end;
  then f2=h by A81,FUNCT_1:9;
  then A87: f3 is continuous by A80,JGRAPH_1:63;
  A88: dom f3=[#](I[01]) by A17,BORSUK_1:83,PRE_TOPC:12;
    for x1,x2 being set st x1 in dom f3 & x2 in dom f3 & f3.x1=f3.x2
  holds x1=x2
  proof let x1,x2 be set;
   assume A89: x1 in dom f3 & x2 in dom f3 & f3.x1=f3.x2;
    then reconsider r1=x1 as Real by A17;
    reconsider r2=x2 as Real by A17,A89;
    A90: LSeg(|[b,d]|,|[b,c]|) /\ LSeg(|[b,c]|,|[a,c]|) = {|[b,c]|}
                  by A1,Th42;
      now per cases by A2,A17,A89,XBOOLE_0:def 2;
    case A91: x1 in [.0,1/2.] & x2 in [.0,1/2.];
      then f3.r1=(1-2*r1)*|[b,d]|+(2*r1)*|[b,c]| by A33;
      then (1-2*r2)*|[b,d]|+(2*r2)*|[b,c]|= (1-2*r1)*|[b,d]|+(2*r1)*|[b,c]|
                               by A33,A89,A91;
      then (1-2*r2)*|[b,d]|+(2*r2)*|[b,c]| -(2*r1)*|[b,c]|
      = (1-2*r1)*|[b,d]| by EUCLID:52;
      then (1-2*r2)*|[b,d]|+((2*r2)*|[b,c]| -(2*r1)*|[b,c]|)
      = (1-2*r1)*|[b,d]| by EUCLID:49;
      then (1-2*r2)*|[b,d]|+(2*r2-2*r1)*(|[b,c]|)
      = (1-2*r1)*|[b,d]| by EUCLID:54;
      then (2*r2-2*r1)*(|[b,c]|)+((1-2*r2)*|[b,d]|-(1-2*r1)*|[b,d]|)
      = (1-2*r1)*|[b,d]|-(1-2*r1)*|[b,d]| by EUCLID:49;
      then (2*r2-2*r1)*(|[b,c]|)+((1-2*r2)*|[b,d]|-(1-2*r1)*|[b,d]|)
      = 0.REAL 2 by EUCLID:46;
      then (2*r2-2*r1)*(|[b,c]|)+((1-2*r2)-(1-2*r1))*|[b,d]|
      = 0.REAL 2 by EUCLID:54;
      then (2*r2-2*r1)*(|[b,c]|)+(-(2*r2-2*r1))*|[b,d]|
      = 0.REAL 2 by XCMPLX_1:149;
      then (2*r2-2*r1)*(|[b,c]|)+-((2*r2-2*r1)*|[b,d]|)
      = 0.REAL 2 by EUCLID:44;
      then (2*r2-2*r1)*(|[b,c]|)-((2*r2-2*r1)*|[b,d]|)
      = 0.REAL 2 by EUCLID:45;
      then (2*r2-2*r1)*((|[b,c]|)-(|[b,d]|))
      = 0.REAL 2 by EUCLID:53;
      then (2*r2-2*r1)=0 or (|[b,c]|)-(|[b,d]|)=0.REAL 2 by EUCLID:35;
      then (2*r2-2*r1)=0 or |[b,c]|=|[b,d]| by EUCLID:47;
      then (2*r2-2*r1)=0 or d =|[b,c]|`2 by EUCLID:56;
      then 2*r2=2*r1 by A1,EUCLID:56,XCMPLX_1:15;
      then r2=r1*2/2 by XCMPLX_1:90;
     hence x1=x2 by XCMPLX_1:90;
    case A92: x1 in [.0,1/2.] & x2 in [.1/2,1.];
      then A93: f3.r1=(1-2*r1)*|[b,d]|+(2*r1)*|[b,c]| by A33;
      A94: 0<=r1 & r1<=1/2 by A92,TOPREAL5:1;
      then A95: r1*2<=1/2*2 by AXIOMS:25;
        2*0<=2*r1 by A94,AXIOMS:25;
      then f3.r1 in { (1-lambda)*|[b,d]| + lambda*|[b,c]| where lambda is Real
:
            0 <= lambda & lambda <= 1 } by A93,A95;
      then A96: f3.r1 in LSeg(|[b,d]|,|[b,c]|) by TOPREAL1:def 4;
      A97: f3.r2=(1-(2*r2-1))*|[b,c]|+(2*r2-1)*|[a,c]| by A36,A92;
      A98: 1/2<=r2 & r2<=1 by A92,TOPREAL5:1;
      then r2*2>=1/2*2 by AXIOMS:25;
      then A99: 2*r2-1>=0 by SQUARE_1:12;
        2*1>=2*r2 by A98,AXIOMS:25;
      then 1+1-1>=2*r2-1 by REAL_1:49;
      then f3.r2 in { (1-lambda)*|[b,c]| + lambda*|[a,c]| where lambda is Real
:
            0 <= lambda & lambda <= 1 } by A97,A99;
      then f3.r2 in LSeg(|[b,c]|,|[a,c]|) by TOPREAL1:def 4;
      then f3.r1 in LSeg(|[b,d]|,|[b,c]|) /\ LSeg(|[b,c]|,|[a,c]|)
                                 by A89,A96,XBOOLE_0:def 3;
      then A100: f3.r1= |[b,c]| by A90,TARSKI:def 1;
      then (1-2*r1)*|[b,d]|+(2*r1)*|[b,c]|+-(|[b,c]|)=0.REAL 2
                                   by A93,EUCLID:40;
      then (1-2*r1)*|[b,d]|+(2*r1)*|[b,c]|+(-1)*|[b,c]|=0.REAL 2
                                   by EUCLID:43;
      then (1-2*r1)*|[b,d]|+((2*r1)*|[b,c]|+(-1)*|[b,c]|)=0.REAL 2
                                   by EUCLID:30;
      then (1-2*r1)*|[b,d]|+((2*r1)+(-1))*|[b,c]|=0.REAL 2
                                   by EUCLID:37;
      then (1-2*r1)*|[b,d]|+(-(1-(2*r1)))*|[b,c]|=0.REAL 2 by XCMPLX_1:162;
      then (1-2*r1)*|[b,d]|+-((1-(2*r1))*|[b,c]|)=0.REAL 2 by EUCLID:44;
      then (1-2*r1)*|[b,d]|-((1-(2*r1))*|[b,c]|)=0.REAL 2 by EUCLID:45;
      then (1-2*r1)*(|[b,d]|-(|[b,c]|))=0.REAL 2 by EUCLID:53;
      then 1-2*r1=0 or (|[b,d]|-(|[b,c]|))=0.REAL 2 by EUCLID:35;
      then 1-2*r1=0 or |[b,d]|=|[b,c]| by EUCLID:47;
      then 1-2*r1=0 or d =|[b,c]|`2 by EUCLID:56;
      then 1=2*r1 by A1,EUCLID:56,XCMPLX_1:15;
      then A101: r1=1/2 by XCMPLX_1:90;
        (1-(2*r2-1))*|[b,c]|+(2*r2-1)*|[a,c]|+-(|[b,c]|)=0.REAL 2
                                   by A89,A97,A100,EUCLID:40;
      then (1-(2*r2-1))*|[b,c]|+((2*r2-1))*|[a,c]|+(-1)*|[b,c]|=0.REAL 2
                                   by EUCLID:43;
      then ((2*r2-1))*|[a,c]|+((1-(2*r2-1))*|[b,c]|+(-1)*|[b,c]|)=0.REAL 2
                                   by EUCLID:30;
      then ((2*r2-1))*|[a,c]|+((1-(2*r2-1))+(-1))*|[b,c]|=0.REAL 2
                                   by EUCLID:37;
      then ((2*r2-1))*|[a,c]|+(-1+1-(2*r2-1))*|[b,c]|=0.REAL 2 by XCMPLX_1:29;
      then ((2*r2-1))*|[a,c]|+(-(2*r2-1))*|[b,c]|=0.REAL 2 by XCMPLX_1:150;
      then ((2*r2-1))*|[a,c]|+-((2*r2-1)*|[b,c]|)=0.REAL 2 by EUCLID:44;
      then ((2*r2-1))*|[a,c]|-((2*r2-1)*|[b,c]|)=0.REAL 2 by EUCLID:45;
      then ((2*r2-1))*(|[a,c]|-(|[b,c]|))=0.REAL 2 by EUCLID:53;
      then (2*r2-1)=0 or (|[a,c]|-(|[b,c]|))=0.REAL 2 by EUCLID:35;
      then (2*r2-1)=0 or |[a,c]|=|[b,c]| by EUCLID:47;
      then (2*r2-1)=0 or a =|[b,c]|`1 by EUCLID:56;
      then 1=2*r2 by A1,EUCLID:56,XCMPLX_1:15;
     hence x1=x2 by A101,XCMPLX_1:90;
    case A102: x1 in [.1/2,1.] & x2 in [.0,1/2.];
      then A103: f3.r2=(1-2*r2)*|[b,d]|+(2*r2)*|[b,c]| by A33;
      A104: 0<=r2 & r2<=1/2 by A102,TOPREAL5:1;
      then A105: r2*2<=1/2*2 by AXIOMS:25;
        2*0<=2*r2 by A104,AXIOMS:25;
      then f3.r2 in { (1-lambda)*|[b,d]| + lambda*|[b,c]| where lambda is Real
:
            0 <= lambda & lambda <= 1 } by A103,A105;
      then A106: f3.r2 in LSeg(|[b,d]|,|[b,c]|) by TOPREAL1:def 4;
      A107: f3.r1=(1-(2*r1-1))*|[b,c]|+(2*r1-1)*|[a,c]| by A36,A102;
      A108: 1/2<=r1 & r1<=1 by A102,TOPREAL5:1;
      then r1*2>=1/2*2 by AXIOMS:25;
      then A109: 2*r1-1>=0 by SQUARE_1:12;
        2*1>=2*r1 by A108,AXIOMS:25;
      then 1+1-1>=2*r1-1 by REAL_1:49;
      then f3.r1 in { (1-lambda)*|[b,c]| + lambda*|[a,c]| where lambda is Real
:
            0 <= lambda & lambda <= 1 } by A107,A109;
      then f3.r1 in LSeg(|[b,c]|,|[a,c]|) by TOPREAL1:def 4;
      then f3.r2 in LSeg(|[b,d]|,|[b,c]|) /\ LSeg(|[b,c]|,|[a,c]|)
                                 by A89,A106,XBOOLE_0:def 3;
      then A110: f3.r2= |[b,c]| by A90,TARSKI:def 1;
      then (1-2*r2)*|[b,d]|+(2*r2)*|[b,c]|+-(|[b,c]|)=0.REAL 2
                                   by A103,EUCLID:40;
      then (1-2*r2)*|[b,d]|+(2*r2)*|[b,c]|+(-1)*|[b,c]|=0.REAL 2
                                   by EUCLID:43;
      then (1-2*r2)*|[b,d]|+((2*r2)*|[b,c]|+(-1)*|[b,c]|)=0.REAL 2
                                   by EUCLID:30;
      then (1-2*r2)*|[b,d]|+((2*r2)+(-1))*|[b,c]|=0.REAL 2
                                   by EUCLID:37;
      then (1-2*r2)*|[b,d]|+(-(1-(2*r2)))*|[b,c]|=0.REAL 2 by XCMPLX_1:162;
      then (1-2*r2)*|[b,d]|+-((1-(2*r2))*|[b,c]|)=0.REAL 2 by EUCLID:44;
      then (1-2*r2)*|[b,d]|-((1-(2*r2))*|[b,c]|)=0.REAL 2 by EUCLID:45;
      then (1-2*r2)*(|[b,d]|-(|[b,c]|))=0.REAL 2 by EUCLID:53;
      then 1-2*r2=0 or (|[b,d]|-(|[b,c]|))=0.REAL 2 by EUCLID:35;
      then 1-2*r2=0 or |[b,d]|=|[b,c]| by EUCLID:47;
      then 1-2*r2=0 or d =|[b,c]|`2 by EUCLID:56;
      then 1=2*r2 by A1,EUCLID:56,XCMPLX_1:15;
      then A111: r2=1/2 by XCMPLX_1:90;
        (1-(2*r1-1))*|[b,c]|+(2*r1-1)*|[a,c]|+-(|[b,c]|)=0.REAL 2
                                   by A89,A107,A110,EUCLID:40;
      then (1-(2*r1-1))*|[b,c]|+((2*r1-1))*|[a,c]|+(-1)*|[b,c]|=0.REAL 2
                                   by EUCLID:43;
      then ((2*r1-1))*|[a,c]|+((1-(2*r1-1))*|[b,c]|+(-1)*|[b,c]|)=0.REAL 2
                                   by EUCLID:30;
      then ((2*r1-1))*|[a,c]|+((1-(2*r1-1))+(-1))*|[b,c]|=0.REAL 2
                                   by EUCLID:37;
      then ((2*r1-1))*|[a,c]|+(-1+1-(2*r1-1))*|[b,c]|=0.REAL 2 by XCMPLX_1:29;
      then ((2*r1-1))*|[a,c]|+(-(2*r1-1))*|[b,c]|=0.REAL 2 by XCMPLX_1:150;
      then ((2*r1-1))*|[a,c]|+-((2*r1-1)*|[b,c]|)=0.REAL 2 by EUCLID:44;
      then ((2*r1-1))*|[a,c]|-((2*r1-1)*|[b,c]|)=0.REAL 2 by EUCLID:45;
      then ((2*r1-1))*(|[a,c]|-(|[b,c]|))=0.REAL 2 by EUCLID:53;
      then (2*r1-1)=0 or (|[a,c]|-(|[b,c]|))=0.REAL 2 by EUCLID:35;
      then (2*r1-1)=0 or |[a,c]|=|[b,c]| by EUCLID:47;
      then (2*r1-1)=0 or a =|[b,c]|`1 by EUCLID:56;
      then 1=2*r1 by A1,EUCLID:56,XCMPLX_1:15;
     hence x1=x2 by A111,XCMPLX_1:90;
    case A112: x1 in [.1/2,1.] & x2 in [.1/2,1.];
      then f3.r1=(1-(2*r1-1))*|[b,c]|+((2*r1-1))*|[a,c]| by A36;
      then (1-(2*r2-1))*|[b,c]|+((2*r2-1))*|[a,c]|
         = (1-(2*r1-1))*|[b,c]|+((2*r1-1))*|[a,c]| by A36,A89,A112;
      then (1-(2*r2-1))*|[b,c]|+((2*r2-1))*|[a,c]| -((2*r1-1))*|[a,c]|
      = (1-(2*r1-1))*|[b,c]| by EUCLID:52;
      then (1-(2*r2-1))*|[b,c]|+(((2*r2-1))*|[a,c]| -((2*r1-1))*|[a,c]|)
      = (1-(2*r1-1))*|[b,c]| by EUCLID:49;
      then (1-(2*r2-1))*|[b,c]|+((2*r2-1)-(2*r1-1))*(|[a,c]|)
      = (1-(2*r1-1))*|[b,c]| by EUCLID:54;
      then ((2*r2-1)-(2*r1-1))*(|[a,c]|)
          +((1-(2*r2-1))*|[b,c]|-(1-(2*r1-1))*|[b,c]|)
      = (1-(2*r1-1))*|[b,c]|-(1-(2*r1-1))*|[b,c]| by EUCLID:49;
      then ((2*r2-1)-(2*r1-1))*(|[a,c]|)
          +((1-(2*r2-1))*|[b,c]|-(1-(2*r1-1))*|[b,c]|)
      = 0.REAL 2 by EUCLID:46;
      then ((2*r2-1)-(2*r1-1))*(|[a,c]|)+((1-(2*r2-1))-(1-(2*r1-1)))*|[b,c]|
      = 0.REAL 2 by EUCLID:54;
      then ((2*r2-1)-(2*r1-1))*(|[a,c]|)+(-((2*r2-1)-(2*r1-1)))*|[b,c]|
      = 0.REAL 2 by XCMPLX_1:149;
      then ((2*r2-1)-(2*r1-1))*(|[a,c]|)+-(((2*r2-1)-(2*r1-1))*|[b,c]|)
      = 0.REAL 2 by EUCLID:44;
      then ((2*r2-1)-(2*r1-1))*(|[a,c]|)-(((2*r2-1)-(2*r1-1))*|[b,c]|)
      = 0.REAL 2 by EUCLID:45;
      then ((2*r2-1)-(2*r1-1))*((|[a,c]|)-(|[b,c]|))
      = 0.REAL 2 by EUCLID:53;
      then ((2*r2-1)-(2*r1-1))=0 or (|[a,c]|)-(|[b,c]|)=0.REAL 2
                 by EUCLID:35;
      then ((2*r2-1)-(2*r1-1))=0 or |[a,c]|=|[b,c]| by EUCLID:47;
      then ((2*r2-1)-(2*r1-1))=0 or a =|[b,c]|`1 by EUCLID:56;
      then (2*r2-1)=(2*r1-1) by A1,EUCLID:56,XCMPLX_1:15;
      then 2*r2=2*r1-1+1 by XCMPLX_1:27;
      then 2*r2=2*r1 by XCMPLX_1:27;
      then r2=r1*2/2 by XCMPLX_1:90;
     hence x1=x2 by XCMPLX_1:90;
    end;
   hence x1=x2;
  end;
  then A113: f3 is one-to-one by FUNCT_1:def 8;
  A114: the carrier of ((TOP-REAL 2)|(Lower_Arc(K)))
    =[#]((TOP-REAL 2)|(Lower_Arc(K))) by PRE_TOPC:12;
    [#]((TOP-REAL 2)|(Lower_Arc(K))) c= rng f3
   proof let y be set;assume y in [#]((TOP-REAL 2)|(Lower_Arc(K)));
     then A115: y in Lower_Arc(K) by PRE_TOPC:def 10;
     then reconsider q=y as Point of TOP-REAL 2;
     A116: Lower_Arc(K)=LSeg(|[b,d]|,|[b,c]|) \/ LSeg(|[b,c]|,|[a,c]|) by A1,
Th62
;
       now per cases by A115,A116,XBOOLE_0:def 2;
     case q in LSeg(|[b,d]|,|[b,c]|);
       then A117: 0<=((q`2)-d)/(c-d)/2 & ((q`2)-d)/(c-d)/2<=1
        & f3.(((q`2)-d)/(c-d)/2)=q by A39;
       then ((q`2)-d)/(c-d)/2 in [.0,1.] by TOPREAL5:1;
      hence y in rng f3 by A17,A117,FUNCT_1:def 5;
     case q in LSeg(|[b,c]|,|[a,c]|);
      then A118: 0<=((q`1)-b)/(a-b)/2+1/2 & ((q`1)-b)/(a-b)/2+1/2<=1
        & f3.(((q`1)-b)/(a-b)/2+1/2)=q by A47;
       then ((q`1)-b)/(a-b)/2+1/2 in [.0,1.] by TOPREAL5:1;
      hence y in rng f3 by A17,A118,FUNCT_1:def 5;
     end;
    hence y in rng f3;
   end;
  then A119: rng f3=[#]((TOP-REAL 2)|(Lower_Arc(K)))
                             by A114,XBOOLE_0:def 10;
  A120: I[01] is compact by HEINE:11,TOPMETR:27;
    ((TOP-REAL 2)|(Lower_Arc(K))) is_T2 by TOPMETR:3;
  then A121: f3 is_homeomorphism by A87,A88,A113,A119,A120,COMPTS_1:26;
    rng f3=Lower_Arc(K) by A119,PRE_TOPC:def 10;
 hence
  ex f being map of I[01],(TOP-REAL 2)|(Lower_Arc(K)) st f is_homeomorphism
& f.0=E-max(K) & f.1=W-min(K) & rng f=Lower_Arc(K) &
(for r being Real st r in [.0,1/2.] holds f.r=(1-2*r)*|[b,d]|+(2*r)*|[b,c]|)&
(for r being Real st r in [.1/2,1.] holds
    f.r=(1-(2*r-1))*|[b,c]|+(2*r-1)*|[a,c]|)&
(for p being Point of TOP-REAL 2 st p in LSeg(|[b,d]|,|[b,c]|)
  holds 0<=((p`2)-d)/(c-d)/2 & ((p`2)-d)/(c-d)/2<=1
        & f.(((p`2)-d)/(c-d)/2)=p)&
(for p being Point of TOP-REAL 2 st p in LSeg(|[b,c]|,|[a,c]|)
  holds 0<=((p`1)-b)/(a-b)/2+1/2 & ((p`1)-b)/(a-b)/2+1/2<=1
        & f.(((p`1)-b)/(a-b)/2+1/2)=p) by A30,A32,A33,A36,A39,A47,A121;
end;

theorem Th65: for K being non empty compact Subset of TOP-REAL 2,
a,b,c,d being real number,p1,p2 being Point of TOP-REAL 2
 st K=rectangle(a,b,c,d) & a<b & c <d
 & p1 in LSeg(|[a,c]|,|[a,d]|)
 & p2 in LSeg(|[a,c]|,|[a,d]|)
holds LE p1,p2,K iff p1`2<=p2`2
proof let K be non empty compact Subset of TOP-REAL 2,
a,b,c,d be real number,p1,p2 be Point of TOP-REAL 2;
 assume A1: K=rectangle(a,b,c,d) & a<b & c <d
 & p1 in LSeg(|[a,c]|,|[a,d]|)
 & p2 in LSeg(|[a,c]|,|[a,d]|);
 then A2: K is_simple_closed_curve by Th60;
 A3: p1`1=a & c <=p1`2 & p1`2 <= d by A1,Th9;
 A4: p2`1=a & c <=p2`2 & p2`2 <= d by A1,Th9;
  A5: E-max(K)= |[b,d]| by A1,Th56;
    A6: Upper_Arc(K)=LSeg(|[a,c]|,|[a,d]|) \/ LSeg(|[a,d]|,|[b,d]|)
                        by A1,Th61;
    then A7: LSeg(|[a,c]|,|[a,d]|) c= Upper_Arc(K) by XBOOLE_1:7;
    A8: Upper_Arc(K) /\ Lower_Arc(K)={W-min(K),E-max(K)}
                              by A2,JORDAN6:def 9;
   reconsider a2=a,c2=c,d2=d as Real by XREAL_0:def 1;
    A9: now assume p2 in Lower_Arc(K);
      then A10: p2 in Upper_Arc(K) /\ Lower_Arc(K) by A1,A7,XBOOLE_0:def 3;
        now assume p2=E-max(K);
        then A11: p2`1=b by A5,EUCLID:56;
          p2 in LSeg(|[a2,c2]|,|[a2,d2]|) by A1;
       hence contradiction by A1,A11,TOPREAL3:17;
      end;
     hence p2=W-min(K) by A8,A10,TARSKI:def 2;
    end;
 thus LE p1,p2,K implies p1`2<=p2`2
  proof assume LE p1,p2,K;
    then A12: p1 in Upper_Arc(K) & p2 in Lower_Arc(K)& not p2=W-min(K) or
     p1 in Upper_Arc(K) & p2 in Upper_Arc(K) &
     LE p1,p2,Upper_Arc(K),W-min(K),E-max(K) or
     p1 in Lower_Arc(K) & p2 in Lower_Arc(K)& not p2=W-min(K) &
     LE p1,p2,Lower_Arc(K),E-max(K),W-min(K) by JORDAN6:def 10;
    consider f being map of I[01],(TOP-REAL 2)|(Upper_Arc(K)) such that
    A13: f is_homeomorphism
     & f.0=W-min(K) & f.1=E-max(K) & rng f=Upper_Arc(K) &
(for r being Real st r in [.0,1/2.] holds f.r=(1-2*r)*|[a,c]|+(2*r)*|[a,d]|)&
(for r being Real st r in [.1/2,1.] holds
    f.r=(1-(2*r-1))*|[a,d]|+(2*r-1)*|[b,d]|)&
(for p being Point of TOP-REAL 2 st p in LSeg(|[a,c]|,|[a,d]|)
  holds 0<=((p`2)-c)/(d-c)/2 & ((p`2)-c)/(d-c)/2<=1
        & f.(((p`2)-c)/(d-c)/2)=p)&
(for p being Point of TOP-REAL 2 st p in LSeg(|[a,d]|,|[b,d]|)
  holds 0<=((p`1)-a)/(b-a)/2+1/2 & ((p`1)-a)/(b-a)/2+1/2<=1
        & f.(((p`1)-a)/(b-a)/2+1/2)=p)
                         by A1,Th63;
     reconsider s1=((p1`2)-c)/(d-c)/2,s2=((p2`2)-c)/(d-c)/2 as
         Real by XREAL_0:def 1;
     A14: f.s1=p1 by A1,A13;
     A15: f.s2=p2 by A1,A13;
     A16: d-c >0 by A1,SQUARE_1:11;
     A17: 0<=s1 & s1<=1 by A1,A13;
       0<=s2 & s2<=1 by A1,A13;
     then s1<=s2 by A9,A12,A13,A14,A15,A17,JORDAN5C:def 3;
     then ((p1`2)-c)/(d-c)/2*2<=((p2`2)-c)/(d-c)/2*2 by AXIOMS:25;
     then ((p1`2)-c)/(d-c)<= ((p2`2)-c)/(d-c)/2*2 by XCMPLX_1:88;
     then ((p1`2)-c)/(d-c)<= ((p2`2)-c)/(d-c) by XCMPLX_1:88;
     then ((p1`2)-c)/(d-c)*(d-c)<= ((p2`2)-c)/(d-c)*(d-c) by A16,AXIOMS:25;
     then ((p1`2)-c)<= ((p2`2)-c)/(d-c)*(d-c) by A16,XCMPLX_1:88;
     then ((p1`2)-c)<= ((p2`2)-c) by A16,XCMPLX_1:88;
     then ((p1`2)-c)+c <= ((p2`2)-c)+c by REAL_1:55;
     then (p1`2) <= ((p2`2)-c)+c by XCMPLX_1:27;
   hence p1`2<=p2`2 by XCMPLX_1:27;
  end;
 thus p1`2<=p2`2 implies LE p1,p2,K
  proof assume A18: p1`2<=p2`2;
      for g being map of I[01], (TOP-REAL 2)|Upper_Arc(K),
        s1, s2 being Real st
        g is_homeomorphism
        & g.0 = W-min(K) & g.1 = E-max(K)
        & g.s1 = p1 & 0 <= s1 & s1 <= 1
        & g.s2 = p2 & 0 <= s2 & s2 <= 1
        holds s1 <= s2
     proof let g be map of I[01], (TOP-REAL 2)|Upper_Arc(K),
        s1, s2 be Real;
        assume A19: g is_homeomorphism
        & g.0 = W-min(K) & g.1 = E-max(K)
        & g.s1 = p1 & 0 <= s1 & s1 <= 1
        & g.s2 = p2 & 0 <= s2 & s2 <= 1;
        A20: dom g=the carrier of I[01] by FUNCT_2:def 1;
        A21: g is one-to-one by A19,TOPS_2:def 5;
        A22: the carrier of ((TOP-REAL 2)|Upper_Arc(K))
        =Upper_Arc(K) by JORDAN1:1;
        then g is Function of the carrier of I[01],
              the carrier of TOP-REAL 2 by FUNCT_2:9;
        then reconsider g1=g as map of I[01],TOP-REAL 2 ;
          g is continuous by A19,TOPS_2:def 5;
        then A23: g1 is continuous by TOPMETR:7;
        reconsider h1=proj1 as map of TOP-REAL 2,R^1 by JGRAPH_2:16;
        reconsider h2=proj2 as map of TOP-REAL 2,R^1 by JGRAPH_2:17;
        A24: (TOP-REAL 2)|([#](TOP-REAL 2))=TOP-REAL 2 by TSEP_1:3;
        then A25: (for p being Point of (TOP-REAL 2)|([#](TOP-REAL 2)) holds
        h1.p=proj1.p) implies h1 is continuous by JGRAPH_2:39;
          (for p being Point of (TOP-REAL 2)|([#](TOP-REAL 2)) holds
        h2.p=proj2.p) implies h2 is continuous by A24,JGRAPH_2:40;
        then consider h being map of TOP-REAL 2,R^1
        such that
        A26: (for p being Point of TOP-REAL 2,
          r1,r2 being real number st
        h1.p=r1 & h2.p=r2 holds h.p=r1+r2) & h is continuous
                          by A25,JGRAPH_2:29;
        reconsider k=h*g1 as map of I[01],R^1;
        A27: k is continuous map of I[01],R^1 by A23,A26,TOPS_2:58;
        A28: W-min K=|[a,c]| by A1,Th56;
          now assume A29: s1>s2;
          A30: dom g=[.0,1.] by BORSUK_1:83,FUNCT_2:def 1;
            0 in [.0,1.] by TOPREAL5:1;
          then A31: k.0=h.(W-min(K)) by A19,A30,FUNCT_1:23
          .=h1.(W-min(K))+h2.(W-min(K)) by A26
          .=(W-min(K))`1+proj2.(W-min(K)) by PSCOMP_1:def 28
          .=(W-min(K))`1+(W-min(K))`2 by PSCOMP_1:def 29
          .=a+(W-min(K))`2 by A28,EUCLID:56
          .=a+c by A28,EUCLID:56;
            s1 in [.0,1.] by A19,TOPREAL5:1;
          then A32: k.s1=h.p1 by A19,A30,FUNCT_1:23
          .=h1.p1+h2.p1 by A26
          .=p1`1+proj2.p1 by PSCOMP_1:def 28
          .=a+p1`2 by A3,PSCOMP_1:def 29;
          A33: s2 in [.0,1.] by A19,TOPREAL5:1;
          then k.s2=h.p2 by A19,A30,FUNCT_1:23
          .=h1.p2+h2.p2 by A26
          .=p2`1+proj2.p2 by PSCOMP_1:def 28
          .=a+p2`2 by A4,PSCOMP_1:def 29;
          then A34: k.0<=k.s1 & k.s1<=k.s2 by A3,A18,A31,A32,REAL_1:55;
          A35: 0 in [.0,1.] by TOPREAL5:1;
          then A36: [.0,s2.] c= [.0,1.] by A33,RCOMP_1:16;
          then reconsider B=[.0,s2.] as Subset of I[01] by BORSUK_1:83
;
          A37: B is connected by A19,A33,A35,BORSUK_1:83,BORSUK_4:49;
          A38: 0 in B by A19,TOPREAL5:1;
          A39: s2 in B by A19,TOPREAL5:1;
          A40: k.0 is Real by XREAL_0:def 1;
          A41: k.s2 is Real by XREAL_0:def 1;
            k.s1 is Real by XREAL_0:def 1;
          then consider xc being Point of I[01] such that
          A42: xc in B & k.xc =k.s1
                   by A27,A34,A37,A38,A39,A40,A41,TOPREAL5:11;
            xc in [.0,1.] by BORSUK_1:83;
          then reconsider rxc=xc as Real;
          A43: k is one-to-one
          proof
              for x1,x2 being set st x1 in dom k & x2 in dom k &
            k.x1=k.x2 holds x1=x2
            proof let x1,x2 be set;
             assume A44: x1 in dom k & x2 in dom k &
               k.x1=k.x2;
              then reconsider r1=x1 as Point of I[01] by FUNCT_2:def 1;
              reconsider r2=x2 as Point of I[01] by A44,FUNCT_2:def 1;
              A45: k.x1=h.(g1.x1) by A44,FUNCT_1:22
                 .=h1.(g1.r1)+h2.(g1.r1) by A26
                 .=(g1.r1)`1+proj2.(g1.r1) by PSCOMP_1:def 28
                 .=(g1.r1)`1+(g1.r1)`2 by PSCOMP_1:def 29;
              A46: k.x2=h.(g1.x2) by A44,FUNCT_1:22
                 .=h1.(g1.r2)+h2.(g1.r2) by A26
                 .=(g1.r2)`1+proj2.(g1.r2) by PSCOMP_1:def 28
                 .=(g1.r2)`1+(g1.r2)`2 by PSCOMP_1:def 29;
              A47: g.r1 in Upper_Arc(K) by A22;
              A48: g.r2 in Upper_Arc(K) by A22;
              reconsider gr1=g.r1 as Point of TOP-REAL 2 by A47;
              reconsider gr2=g.r2 as Point of TOP-REAL 2 by A48;
                now per cases by A6,A22,XBOOLE_0:def 2;
              case A49: g.r1 in LSeg(|[a,c]|,|[a,d]|) &
                  g.r2 in LSeg(|[a,c]|,|[a,d]|);
                then A50: (gr1)`1=a & c <=(gr1)`2 & (gr1)`2 <=d by A1,Th9;
                A51: (gr2)`1=a & c <=(gr2)`2 & (gr2)`2 <=d
                                  by A1,A49,Th9;
                 then (gr1)`2=(gr2)`2 by A44,A45,A46,A50,XCMPLX_1:2;
                then |[(gr1)`1,(gr1)`2]|=g.r2 by A50,A51,EUCLID:57;
                then g.r1=g.r2 by EUCLID:57;
               hence x1=x2 by A20,A21,FUNCT_1:def 8;
              case A52: g.r1 in LSeg(|[a,c]|,|[a,d]|) &
                  g.r2 in LSeg(|[a,d]|,|[b,d]|);
                then A53: (gr1)`1=a & c <=(gr1)`2 & (gr1)`2 <=d by A1,Th9;
                A54: (gr2)`2=d & a <=(gr2)`1 & (gr2)`1 <=b
                                  by A1,A52,Th11;
                A55: a+(gr1)`2=(gr2)`1 +d by A1,A44,A45,A46,A52,A53,Th11;
                A56: now assume a<>gr2`1; then a<gr2`1 by A54,REAL_1:def 5;
                 hence contradiction by A53,A55,REAL_1:67;
                end;
                  now assume gr1`2<>d; then d>gr1`2 by A53,REAL_1:def 5;
                 hence contradiction by A44,A45,A46,A53,A54,REAL_1:67;
                end;
                then |[(gr1)`1,(gr1)`2]|=g.r2 by A53,A54,A56,EUCLID:57;
                then g.r1=g.r2 by EUCLID:57;
               hence x1=x2 by A20,A21,FUNCT_1:def 8;
              case A57: g.r1 in LSeg(|[a,d]|,|[b,d]|) &
                  g.r2 in LSeg(|[a,c]|,|[a,d]|);
                then A58: (gr2)`1=a & c <=(gr2)`2 & (gr2)`2 <=d by A1,Th9;
                A59: (gr1)`2=d & a <=(gr1)`1 & (gr1)`1 <=b
                                  by A1,A57,Th11;
                A60: a+(gr2)`2=(gr1)`1 +d by A1,A44,A45,A46,A57,A58,Th11;
                A61: now assume a<>gr1`1; then a<gr1`1 by A59,REAL_1:def 5;
                 hence contradiction by A58,A60,REAL_1:67;
                end;
                  now assume gr2`2<>d; then d>gr2`2 by A58,REAL_1:def 5;
                 hence contradiction by A44,A45,A46,A58,A59,REAL_1:67;
                end;
                then |[(gr2)`1,(gr2)`2]|=g.r1 by A58,A59,A61,EUCLID:57;
                then g.r1=g.r2 by EUCLID:57;
               hence x1=x2 by A20,A21,FUNCT_1:def 8;
              case A62: g.r1 in LSeg(|[a,d]|,|[b,d]|) &
                  g.r2 in LSeg(|[a,d]|,|[b,d]|);
                then A63: (gr1)`2=d & a <=(gr1)`1 & (gr1)`1 <=b by A1,Th11;
                A64: (gr2)`2=d & a <=(gr2)`1 & (gr2)`1 <=b
                                  by A1,A62,Th11;
                 then (gr1)`1=(gr2)`1 by A44,A45,A46,A63,XCMPLX_1:2;
                then |[(gr1)`1,(gr1)`2]|=g.r2 by A63,A64,EUCLID:57;
                then g.r1=g.r2 by EUCLID:57;
               hence x1=x2 by A20,A21,FUNCT_1:def 8;
              end;
             hence x1=x2;
            end;
           hence thesis by FUNCT_1:def 8;
          end;
          A65: dom k=[.0,1.] by BORSUK_1:83,FUNCT_2:def 1;
          then s1 in dom k by A19,TOPREAL5:1;
          then rxc=s1 by A36,A42,A43,A65,FUNCT_1:def 8;
         hence contradiction by A29,A42,TOPREAL5:1;
        end;
      hence s1 <= s2;
     end;
    then p1 in Upper_Arc(K) & p2 in Upper_Arc(K) &
     LE p1,p2,Upper_Arc(K),W-min(K),E-max(K)
                        by A1,A7,JORDAN5C:def 3;
   hence LE p1,p2,K by JORDAN6:def 10;
  end;
end;

theorem Th66: for K being non empty compact Subset of TOP-REAL 2,
a,b,c,d being real number,p1,p2 being Point of TOP-REAL 2
 st K=rectangle(a,b,c,d) & a<b & c <d
 & p1 in LSeg(|[a,d]|,|[b,d]|)
 & p2 in LSeg(|[a,d]|,|[b,d]|)
holds LE p1,p2,K iff p1`1<=p2`1
proof let K be non empty compact Subset of TOP-REAL 2,
a,b,c,d be real number,p1,p2 be Point of TOP-REAL 2;
 assume A1: K=rectangle(a,b,c,d) & a<b & c <d
 & p1 in LSeg(|[a,d]|,|[b,d]|)
 & p2 in LSeg(|[a,d]|,|[b,d]|);
 then A2: K is_simple_closed_curve by Th60;
 A3: p1`2=d & a <=p1`1 & p1`1 <= b by A1,Th11;
 A4: p2`2=d & a <=p2`1 & p2`1 <= b by A1,Th11;
  A5: W-min(K)= |[a,c]| by A1,Th56;
  A6: E-max(K)= |[b,d]| by A1,Th56;
    A7: Upper_Arc(K)=LSeg(|[a,c]|,|[a,d]|) \/ LSeg(|[a,d]|,|[b,d]|)
                        by A1,Th61;
    then A8: LSeg(|[a,d]|,|[b,d]|) c= Upper_Arc(K) by XBOOLE_1:7;
    A9: Upper_Arc(K) /\ Lower_Arc(K)={W-min(K),E-max(K)}
                              by A2,JORDAN6:def 9;
   reconsider a2=a,b2=b,d2=d as Real by XREAL_0:def 1;
    A10: now assume p2 in Lower_Arc(K);
      then A11: p2 in Upper_Arc(K) /\ Lower_Arc(K) by A1,A8,XBOOLE_0:def 3;
        now assume p2=W-min(K);
        then A12: p2`2=c by A5,EUCLID:56;
          p2 in LSeg(|[a2,d2]|,|[b2,d2]|) by A1;
       hence contradiction by A1,A12,TOPREAL3:18;
      end;
     hence p2=E-max(K) by A9,A11,TARSKI:def 2;
    end;
 thus LE p1,p2,K implies p1`1<=p2`1
  proof assume LE p1,p2,K;
    then A13: p1 in Upper_Arc(K) & p2 in Lower_Arc(K)& not p2=W-min(K) or
     p1 in Upper_Arc(K) & p2 in Upper_Arc(K) &
     LE p1,p2,Upper_Arc(K),W-min(K),E-max(K) or
     p1 in Lower_Arc(K) & p2 in Lower_Arc(K)& not p2=W-min(K) &
     LE p1,p2,Lower_Arc(K),E-max(K),W-min(K) by JORDAN6:def 10;
      now per cases;
    case p2=E-max(K);
     hence p1`1<=p2`1 by A3,A6,EUCLID:56;
    case A14: p2<>E-max(K);
    consider f being map of I[01],(TOP-REAL 2)|(Upper_Arc(K)) such that
    A15: f is_homeomorphism
     & f.0=W-min(K) & f.1=E-max(K) & rng f=Upper_Arc(K) &
(for r being Real st r in [.0,1/2.] holds f.r=(1-2*r)*|[a,c]|+(2*r)*|[a,d]|)&
(for r being Real st r in [.1/2,1.] holds
    f.r=(1-(2*r-1))*|[a,d]|+(2*r-1)*|[b,d]|)&
(for p being Point of TOP-REAL 2 st p in LSeg(|[a,c]|,|[a,d]|)
  holds 0<=((p`2)-c)/(d-c)/2 & ((p`2)-c)/(d-c)/2<=1
        & f.(((p`2)-c)/(d-c)/2)=p)&
(for p being Point of TOP-REAL 2 st p in LSeg(|[a,d]|,|[b,d]|)
  holds 0<=((p`1)-a)/(b-a)/2+1/2 & ((p`1)-a)/(b-a)/2+1/2<=1
        & f.(((p`1)-a)/(b-a)/2+1/2)=p)
                         by A1,Th63;
     reconsider s1=((p1`1)-a)/(b-a)/2+1/2,s2=((p2`1)-a)/(b-a)/2+1/2 as
         Real by XREAL_0:def 1;
     A16: f.s1=p1 by A1,A15;
     A17: f.s2=p2 by A1,A15;
     A18: b-a >0 by A1,SQUARE_1:11;
     A19: 0<=s1 & s1<=1 by A1,A15;
       0<=s2 & s2<=1 by A1,A15;
     then s1<=s2 by A10,A13,A14,A15,A16,A17,A19,JORDAN5C:def 3;
     then ((p1`1)-a)/(b-a)/2<= ((p2`1)-a)/(b-a)/2 by REAL_1:53;
     then ((p1`1)-a)/(b-a)/2*2<=((p2`1)-a)/(b-a)/2*2 by AXIOMS:25;
     then ((p1`1)-a)/(b-a)<= ((p2`1)-a)/(b-a)/2*2 by XCMPLX_1:88;
     then ((p1`1)-a)/(b-a)<= ((p2`1)-a)/(b-a) by XCMPLX_1:88;
     then ((p1`1)-a)/(b-a)*(b-a)<= ((p2`1)-a)/(b-a)*(b-a) by A18,AXIOMS:25;
     then ((p1`1)-a)<= ((p2`1)-a)/(b-a)*(b-a) by A18,XCMPLX_1:88;
     then ((p1`1)-a)<= ((p2`1)-a) by A18,XCMPLX_1:88;
     then ((p1`1)-a)+a <= ((p2`1)-a)+a by REAL_1:55;
     then (p1`1) <= ((p2`1)-a)+a by XCMPLX_1:27;
   hence p1`1<=p2`1 by XCMPLX_1:27;
   end;
   hence thesis;
  end;
 thus p1`1<=p2`1 implies LE p1,p2,K
  proof assume A20: p1`1<=p2`1;
      for g being map of I[01], (TOP-REAL 2)|Upper_Arc(K),
        s1, s2 being Real st
        g is_homeomorphism
        & g.0 = W-min(K) & g.1 = E-max(K)
        & g.s1 = p1 & 0 <= s1 & s1 <= 1
        & g.s2 = p2 & 0 <= s2 & s2 <= 1
        holds s1 <= s2
     proof let g be map of I[01], (TOP-REAL 2)|Upper_Arc(K),
        s1, s2 be Real;
        assume A21: g is_homeomorphism
        & g.0 = W-min(K) & g.1 = E-max(K)
        & g.s1 = p1 & 0 <= s1 & s1 <= 1
        & g.s2 = p2 & 0 <= s2 & s2 <= 1;
        A22: dom g=the carrier of I[01] by FUNCT_2:def 1;
        A23: g is one-to-one by A21,TOPS_2:def 5;
        A24: the carrier of ((TOP-REAL 2)|Upper_Arc(K))
        =Upper_Arc(K) by JORDAN1:1;
        then g is Function of the carrier of I[01],
              the carrier of TOP-REAL 2 by FUNCT_2:9;
        then reconsider g1=g as map of I[01],TOP-REAL 2 ;
          g is continuous by A21,TOPS_2:def 5;
        then A25: g1 is continuous by TOPMETR:7;
        reconsider h1=proj1 as map of TOP-REAL 2,R^1 by JGRAPH_2:16;
        reconsider h2=proj2 as map of TOP-REAL 2,R^1 by JGRAPH_2:17;
        A26: (TOP-REAL 2)|([#](TOP-REAL 2))=TOP-REAL 2 by TSEP_1:3;
        then A27: (for p being Point of (TOP-REAL 2)|([#](TOP-REAL 2)) holds
        h1.p=proj1.p) implies h1 is continuous by JGRAPH_2:39;
          (for p being Point of (TOP-REAL 2)|([#](TOP-REAL 2)) holds
        h2.p=proj2.p) implies h2 is continuous by A26,JGRAPH_2:40;
        then consider h being map of TOP-REAL 2,R^1
        such that
        A28: (for p being Point of TOP-REAL 2,
          r1,r2 being real number st
        h1.p=r1 & h2.p=r2 holds h.p=r1+r2) & h is continuous
                          by A27,JGRAPH_2:29;
        reconsider k=h*g1 as map of I[01],R^1;
        A29: k is continuous map of I[01],R^1 by A25,A28,TOPS_2:58;
        A30: W-min K=|[a,c]| by A1,Th56;
          now assume A31: s1>s2;
          A32: dom g=[.0,1.] by BORSUK_1:83,FUNCT_2:def 1;
            0 in [.0,1.] by TOPREAL5:1;
          then A33: k.0=h.(W-min(K)) by A21,A32,FUNCT_1:23
          .=h1.(W-min(K))+h2.(W-min(K)) by A28
          .=(W-min(K))`1+proj2.(W-min(K)) by PSCOMP_1:def 28
          .=(W-min(K))`1+(W-min(K))`2 by PSCOMP_1:def 29
          .=(W-min(K))`1+c by A30,EUCLID:56
          .=a+c by A30,EUCLID:56;
            s1 in [.0,1.] by A21,TOPREAL5:1;
          then A34: k.s1=h.p1 by A21,A32,FUNCT_1:23
          .=h1.p1+h2.p1 by A28
          .=p1`1+proj2.p1 by PSCOMP_1:def 28
          .=p1`1 +d by A3,PSCOMP_1:def 29;
          A35: s2 in [.0,1.] by A21,TOPREAL5:1;
          then k.s2=h.p2 by A21,A32,FUNCT_1:23
          .=h1.p2+h2.p2 by A28
          .=p2`1+proj2.p2 by PSCOMP_1:def 28
          .=p2`1 +d by A4,PSCOMP_1:def 29;
          then A36: k.0<=k.s1 & k.s1<=k.s2 by A1,A3,A20,A33,A34,REAL_1:55;
          A37: 0 in [.0,1.] by TOPREAL5:1;
          then A38: [.0,s2.] c= [.0,1.] by A35,RCOMP_1:16;
          then reconsider B=[.0,s2.] as Subset of I[01] by BORSUK_1:83
;
          A39: B is connected by A21,A35,A37,BORSUK_1:83,BORSUK_4:49;
          A40: 0 in B by A21,TOPREAL5:1;
          A41: s2 in B by A21,TOPREAL5:1;
          A42: k.0 is Real by XREAL_0:def 1;
          A43: k.s2 is Real by XREAL_0:def 1;
            k.s1 is Real by XREAL_0:def 1;
          then consider xc being Point of I[01] such that
          A44: xc in B & k.xc =k.s1
                   by A29,A36,A39,A40,A41,A42,A43,TOPREAL5:11;
            xc in [.0,1.] by BORSUK_1:83;
          then reconsider rxc=xc as Real;
          A45: k is one-to-one
          proof
              for x1,x2 being set st x1 in dom k & x2 in dom k &
            k.x1=k.x2 holds x1=x2
            proof let x1,x2 be set;
             assume A46: x1 in dom k & x2 in dom k &
               k.x1=k.x2;
              then reconsider r1=x1 as Point of I[01] by FUNCT_2:def 1;
              reconsider r2=x2 as Point of I[01] by A46,FUNCT_2:def 1;
              A47: k.x1=h.(g1.x1) by A46,FUNCT_1:22
                 .=h1.(g1.r1)+h2.(g1.r1) by A28
                 .=(g1.r1)`1+proj2.(g1.r1) by PSCOMP_1:def 28
                 .=(g1.r1)`1+(g1.r1)`2 by PSCOMP_1:def 29;
              A48: k.x2=h.(g1.x2) by A46,FUNCT_1:22
                 .=h1.(g1.r2)+h2.(g1.r2) by A28
                 .=(g1.r2)`1+proj2.(g1.r2) by PSCOMP_1:def 28
                 .=(g1.r2)`1+(g1.r2)`2 by PSCOMP_1:def 29;
              A49: g.r1 in Upper_Arc(K) by A24;
              A50: g.r2 in Upper_Arc(K) by A24;
              reconsider gr1=g.r1 as Point of TOP-REAL 2 by A49;
              reconsider gr2=g.r2 as Point of TOP-REAL 2 by A50;
                now per cases by A7,A24,XBOOLE_0:def 2;
              case A51: g.r1 in LSeg(|[a,c]|,|[a,d]|) &
                  g.r2 in LSeg(|[a,c]|,|[a,d]|);
                then A52: (gr1)`1=a & c <=(gr1)`2 & (gr1)`2 <=d by A1,Th9;
                A53: (gr2)`1=a & c <=(gr2)`2 & (gr2)`2 <=d
                                  by A1,A51,Th9;
                 then (gr1)`2=(gr2)`2 by A46,A47,A48,A52,XCMPLX_1:2;
                then |[(gr1)`1,(gr1)`2]|=g.r2 by A52,A53,EUCLID:57;
                then g.r1=g.r2 by EUCLID:57;
               hence x1=x2 by A22,A23,FUNCT_1:def 8;
              case A54: g.r1 in LSeg(|[a,c]|,|[a,d]|) &
                  g.r2 in LSeg(|[a,d]|,|[b,d]|);
                then A55: (gr1)`1=a & c <=(gr1)`2 & (gr1)`2 <=d by A1,Th9;
                A56: (gr2)`2=d & a <=(gr2)`1 & (gr2)`1 <=b
                                  by A1,A54,Th11;
                A57: a+(gr1)`2=(gr2)`1 +d by A1,A46,A47,A48,A54,A55,Th11;
                A58: now assume a<>gr2`1; then a<gr2`1 by A56,REAL_1:def 5;
                 hence contradiction by A55,A57,REAL_1:67;
                end;
                  now assume gr1`2<>d; then d>gr1`2 by A55,REAL_1:def 5;
                 hence contradiction by A46,A47,A48,A55,A56,REAL_1:67;
                end;
                then |[(gr1)`1,(gr1)`2]|=g.r2 by A55,A56,A58,EUCLID:57;
                then g.r1=g.r2 by EUCLID:57;
               hence x1=x2 by A22,A23,FUNCT_1:def 8;
              case A59: g.r1 in LSeg(|[a,d]|,|[b,d]|) &
                  g.r2 in LSeg(|[a,c]|,|[a,d]|);
                then A60: (gr2)`1=a & c <=(gr2)`2 & (gr2)`2 <=d by A1,Th9;
                A61: (gr1)`2=d & a <=(gr1)`1 & (gr1)`1 <=b
                                  by A1,A59,Th11;
                A62: a+(gr2)`2=(gr1)`1 +d by A1,A46,A47,A48,A59,A60,Th11;
                A63: now assume a<>gr1`1; then a<gr1`1 by A61,REAL_1:def 5;
                 hence contradiction by A60,A62,REAL_1:67;
                end;
                  now assume gr2`2<>d; then d>gr2`2 by A60,REAL_1:def 5;
                 hence contradiction by A46,A47,A48,A60,A61,REAL_1:67;
                end;
                then |[(gr2)`1,(gr2)`2]|=g.r1 by A60,A61,A63,EUCLID:57;
                then g.r1=g.r2 by EUCLID:57;
               hence x1=x2 by A22,A23,FUNCT_1:def 8;
              case A64: g.r1 in LSeg(|[a,d]|,|[b,d]|) &
                  g.r2 in LSeg(|[a,d]|,|[b,d]|);
                then A65: (gr1)`2=d & a <=(gr1)`1 & (gr1)`1 <=b by A1,Th11;
                A66: (gr2)`2=d & a <=(gr2)`1 & (gr2)`1 <=b
                                  by A1,A64,Th11;
                 then (gr1)`1=(gr2)`1 by A46,A47,A48,A65,XCMPLX_1:2;
                then |[(gr1)`1,(gr1)`2]|=g.r2 by A65,A66,EUCLID:57;
                then g.r1=g.r2 by EUCLID:57;
               hence x1=x2 by A22,A23,FUNCT_1:def 8;
              end;
             hence x1=x2;
            end;
           hence thesis by FUNCT_1:def 8;
          end;
          A67: dom k=[.0,1.] by BORSUK_1:83,FUNCT_2:def 1;
          then s1 in dom k by A21,TOPREAL5:1;
          then rxc=s1 by A38,A44,A45,A67,FUNCT_1:def 8;
         hence contradiction by A31,A44,TOPREAL5:1;
        end;
      hence s1 <= s2;
     end;
    then p1 in Upper_Arc(K) & p2 in Upper_Arc(K) &
     LE p1,p2,Upper_Arc(K),W-min(K),E-max(K)
                        by A1,A8,JORDAN5C:def 3;
   hence LE p1,p2,K by JORDAN6:def 10;
  end;
end;

theorem Th67: for K being non empty compact Subset of TOP-REAL 2,
a,b,c,d being real number,p1,p2 being Point of TOP-REAL 2
 st K=rectangle(a,b,c,d) & a<b & c <d
 & p1 in LSeg(|[b,c]|,|[b,d]|)
 & p2 in LSeg(|[b,c]|,|[b,d]|)
holds LE p1,p2,K iff p1`2>=p2`2
proof let K be non empty compact Subset of TOP-REAL 2,
a,b,c,d be real number,p1,p2 be Point of TOP-REAL 2;
 assume A1: K=rectangle(a,b,c,d) & a<b & c <d
 & p1 in LSeg(|[b,c]|,|[b,d]|)
 & p2 in LSeg(|[b,c]|,|[b,d]|);
 then A2: K is_simple_closed_curve by Th60;
 A3: p1`1=b & c <=p1`2 & p1`2 <= d by A1,Th9;
 A4: p2`1=b & c <=p2`2 & p2`2 <= d by A1,Th9;
  A5: W-min(K)= |[a,c]| by A1,Th56;
  A6: E-max(K)= |[b,d]| by A1,Th56;
  A7: Lower_Arc(K)=LSeg(|[b,d]|,|[b,c]|) \/ LSeg(|[b,c]|,|[a,c]|) by A1,Th62
;
    then A8: LSeg(|[b,d]|,|[b,c]|) c= Lower_Arc(K) by XBOOLE_1:7;
    A9: Upper_Arc(K) /\ Lower_Arc(K)={W-min(K),E-max(K)}
                              by A2,JORDAN6:def 9;
   reconsider b2=b,c2=c,d2=d as Real by XREAL_0:def 1;
    A10: now assume p1 in Upper_Arc(K);
      then A11: p1 in Upper_Arc(K) /\ Lower_Arc(K) by A1,A8,XBOOLE_0:def 3;
        now assume p1=W-min(K);
        then A12: p1`1=a by A5,EUCLID:56;
          p1 in LSeg(|[b2,c2]|,|[b2,d2]|) by A1;
       hence contradiction by A1,A12,TOPREAL3:17;
      end;
     hence p1=E-max(K) by A9,A11,TARSKI:def 2;
    end;
 thus LE p1,p2,K implies p1`2>=p2`2
  proof assume LE p1,p2,K;
    then A13: p1 in Upper_Arc(K) & p2 in Lower_Arc(K)& not p2=W-min(K) or
     p1 in Upper_Arc(K) & p2 in Upper_Arc(K) &
     LE p1,p2,Upper_Arc(K),W-min(K),E-max(K) or
     p1 in Lower_Arc(K) & p2 in Lower_Arc(K)& not p2=W-min(K) &
     LE p1,p2,Lower_Arc(K),E-max(K),W-min(K) by JORDAN6:def 10;
      now per cases;
    case p1=E-max(K);
      hence p1`2>=p2`2 by A4,A6,EUCLID:56;
    case A14: p1<>E-max(K);
    consider f being map of I[01],(TOP-REAL 2)|(Lower_Arc(K)) such that
    A15: f is_homeomorphism
     & f.0=E-max(K) & f.1=W-min(K) & rng f=Lower_Arc(K) &
(for r being Real st r in [.0,1/2.] holds f.r=(1-2*r)*|[b,d]|+(2*r)*|[b,c]|)&
(for r being Real st r in [.1/2,1.] holds
    f.r=(1-(2*r-1))*|[b,c]|+(2*r-1)*|[a,c]|)&
(for p being Point of TOP-REAL 2 st p in LSeg(|[b,d]|,|[b,c]|)
  holds 0<=((p`2)-d)/(c-d)/2 & ((p`2)-d)/(c-d)/2<=1
        & f.(((p`2)-d)/(c-d)/2)=p)&
(for p being Point of TOP-REAL 2 st p in LSeg(|[b,c]|,|[a,c]|)
  holds 0<=((p`1)-b)/(a-b)/2+1/2 & ((p`1)-b)/(a-b)/2+1/2<=1
        & f.(((p`1)-b)/(a-b)/2+1/2)=p)
                         by A1,Th64;
     reconsider s1=((p1`2)-d)/(c-d)/2,s2=((p2`2)-d)/(c-d)/2 as
         Real by XREAL_0:def 1;
     A16: f.s1=p1 by A1,A15;
     A17: f.s2=p2 by A1,A15;
       d-c>0 by A1,SQUARE_1:11;
     then -(d-c)< -0 by REAL_1:50;
     then 0-(d-c) < 0 by XCMPLX_1:150;
     then 0-d+c <0 by XCMPLX_1:37;
     then -d+c <0 by XCMPLX_1:150;
     then A18: c -d <0 by XCMPLX_0:def 8;
     A19: 0<=s1 & s1<=1 by A1,A15;
       0<=s2 & s2<=1 by A1,A15;
     then s1<=s2 by A10,A13,A14,A15,A16,A17,A19,JORDAN5C:def 3;
     then ((p1`2)-d)/(c-d)/2*2<=((p2`2)-d)/(c-d)/2*2 by AXIOMS:25;
     then ((p1`2)-d)/(c-d)<= ((p2`2)-d)/(c-d)/2*2 by XCMPLX_1:88;
     then ((p1`2)-d)/(c-d)<= ((p2`2)-d)/(c-d) by XCMPLX_1:88;
     then ((p1`2)-d)/(c-d)*(c-d)>= ((p2`2)-d)/(c-d)*(c-d) by A18,REAL_1:52;
     then ((p1`2)-d)>= ((p2`2)-d)/(c-d)*(c-d) by A18,XCMPLX_1:88;
     then ((p1`2)-d)>= ((p2`2)-d) by A18,XCMPLX_1:88;
     then ((p1`2)-d)+d >= ((p2`2)-d)+d by REAL_1:55;
     then (p1`2) >= ((p2`2)-d)+d by XCMPLX_1:27;
   hence p1`2>=p2`2 by XCMPLX_1:27;
   end;
   hence thesis;
  end;
 thus p1`2>=p2`2 implies LE p1,p2,K
  proof assume A20: p1`2>=p2`2;
     now per cases;
   case p2=W-min (K); then p2=|[a,c]| by A1,Th56;
    hence contradiction by A1,A4,EUCLID:56;
   case A21: p2<>W-min(K);
      for g being map of I[01], (TOP-REAL 2)|Lower_Arc(K),
        s1, s2 being Real st
        g is_homeomorphism
        & g.0 = E-max(K) & g.1 = W-min(K)
        & g.s1 = p1 & 0 <= s1 & s1 <= 1
        & g.s2 = p2 & 0 <= s2 & s2 <= 1
        holds s1 <= s2
     proof let g be map of I[01], (TOP-REAL 2)|Lower_Arc(K),
        s1, s2 be Real;
        assume A22: g is_homeomorphism
        & g.0 = E-max(K) & g.1 = W-min(K)
        & g.s1 = p1 & 0 <= s1 & s1 <= 1
        & g.s2 = p2 & 0 <= s2 & s2 <= 1;
        A23: dom g=the carrier of I[01] by FUNCT_2:def 1;
        A24: g is one-to-one by A22,TOPS_2:def 5;
        A25: the carrier of ((TOP-REAL 2)|Lower_Arc(K))
        =Lower_Arc(K) by JORDAN1:1;
        then g is Function of the carrier of I[01],
              the carrier of TOP-REAL 2 by FUNCT_2:9;
        then reconsider g1=g as map of I[01],TOP-REAL 2 ;
          g is continuous by A22,TOPS_2:def 5;
        then A26: g1 is continuous by TOPMETR:7;
        reconsider h1=proj1 as map of TOP-REAL 2,R^1 by JGRAPH_2:16;
        reconsider h2=proj2 as map of TOP-REAL 2,R^1 by JGRAPH_2:17;
        A27: (TOP-REAL 2)|([#](TOP-REAL 2))=TOP-REAL 2 by TSEP_1:3;
        then A28: (for p being Point of (TOP-REAL 2)|([#](TOP-REAL 2)) holds
        h1.p=proj1.p) implies h1 is continuous by JGRAPH_2:39;
          (for p being Point of (TOP-REAL 2)|([#](TOP-REAL 2)) holds
        h2.p=proj2.p) implies h2 is continuous by A27,JGRAPH_2:40;
        then consider h being map of TOP-REAL 2,R^1
        such that
        A29: (for p being Point of TOP-REAL 2,
          r1,r2 being real number st
        h1.p=r1 & h2.p=r2 holds h.p=r1+r2) & h is continuous
                          by A28,JGRAPH_2:29;
        reconsider k=h*g1 as map of I[01],R^1;
        A30: k is continuous map of I[01],R^1 by A26,A29,TOPS_2:58;
        A31: E-max K=|[b,d]| by A1,Th56;
          now assume A32: s1>s2;
          A33: dom g=[.0,1.] by BORSUK_1:83,FUNCT_2:def 1;
            0 in [.0,1.] by TOPREAL5:1;
          then A34: k.0=h.(E-max(K)) by A22,A33,FUNCT_1:23
          .=h1.(E-max(K))+h2.(E-max(K)) by A29
          .=(E-max(K))`1+proj2.(E-max(K)) by PSCOMP_1:def 28
          .=(E-max(K))`1+(E-max(K))`2 by PSCOMP_1:def 29
          .=b+(E-max(K))`2 by A31,EUCLID:56
          .=b+d by A31,EUCLID:56;
            s1 in [.0,1.] by A22,TOPREAL5:1;
          then A35: k.s1=h.p1 by A22,A33,FUNCT_1:23
          .=h1.p1+h2.p1 by A29
          .=p1`1+proj2.p1 by PSCOMP_1:def 28
          .=b+p1`2 by A3,PSCOMP_1:def 29;
          A36: s2 in [.0,1.] by A22,TOPREAL5:1;
          then k.s2=h.p2 by A22,A33,FUNCT_1:23
          .=proj1.p2 +proj2.p2 by A29 .=p2`1+proj2.p2 by PSCOMP_1:def 28
          .=b+p2`2 by A4,PSCOMP_1:def 29;
          then A37: k.0>=k.s1 & k.s1>=k.s2 by A3,A20,A34,A35,REAL_1:55;
          A38: 0 in [.0,1.] by TOPREAL5:1;
          then A39: [.0,s2.] c= [.0,1.] by A36,RCOMP_1:16;
          then reconsider B=[.0,s2.] as Subset of I[01] by BORSUK_1:83
;
          A40: B is connected by A22,A36,A38,BORSUK_1:83,BORSUK_4:49;
          A41: 0 in B by A22,TOPREAL5:1;
          A42: s2 in B by A22,TOPREAL5:1;
          A43: k.0 is Real by XREAL_0:def 1;
          A44: k.s2 is Real by XREAL_0:def 1;
            k.s1 is Real by XREAL_0:def 1;
          then consider xc being Point of I[01] such that
          A45: xc in B & k.xc =k.s1
                   by A30,A37,A40,A41,A42,A43,A44,TOPREAL5:11;
            xc in [.0,1.] by BORSUK_1:83;
          then reconsider rxc=xc as Real;
          A46: k is one-to-one
          proof
              for x1,x2 being set st x1 in dom k & x2 in dom k &
            k.x1=k.x2 holds x1=x2
            proof let x1,x2 be set;
             assume A47: x1 in dom k & x2 in dom k &
               k.x1=k.x2;
              then reconsider r1=x1 as Point of I[01] by FUNCT_2:def 1;
              reconsider r2=x2 as Point of I[01] by A47,FUNCT_2:def 1;
              A48: k.x1=h.(g1.x1) by A47,FUNCT_1:22
                 .=h1.(g1.r1)+h2.(g1.r1) by A29
                 .=(g1.r1)`1+proj2.(g1.r1) by PSCOMP_1:def 28
                 .=(g1.r1)`1+(g1.r1)`2 by PSCOMP_1:def 29;
              A49: k.x2=h.(g1.x2) by A47,FUNCT_1:22
                 .=h1.(g1.r2)+h2.(g1.r2) by A29
                 .=(g1.r2)`1+proj2.(g1.r2) by PSCOMP_1:def 28
                 .=(g1.r2)`1+(g1.r2)`2 by PSCOMP_1:def 29;
              A50: g.r1 in Lower_Arc(K) by A25;
              A51: g.r2 in Lower_Arc(K) by A25;
              reconsider gr1=g.r1 as Point of TOP-REAL 2 by A50;
              reconsider gr2=g.r2 as Point of TOP-REAL 2 by A51;
                now per cases by A7,A25,XBOOLE_0:def 2;
              case A52: g.r1 in LSeg(|[b,c]|,|[b,d]|) &
                  g.r2 in LSeg(|[b,c]|,|[b,d]|);
                then A53: (gr1)`1=b & c <=(gr1)`2 & (gr1)`2 <=d by A1,Th9;
                A54: (gr2)`1=b & c <=(gr2)`2 & (gr2)`2 <=d
                                  by A1,A52,Th9;
                 then (gr1)`2=(gr2)`2 by A47,A48,A49,A53,XCMPLX_1:2;
                then |[(gr1)`1,(gr1)`2]|=g.r2 by A53,A54,EUCLID:57;
                then g.r1=g.r2 by EUCLID:57;
               hence x1=x2 by A23,A24,FUNCT_1:def 8;
              case A55: g.r1 in LSeg(|[b,c]|,|[b,d]|) &
                  g.r2 in LSeg(|[b,c]|,|[a,c]|);
                then A56: (gr1)`1=b & c <=(gr1)`2 & (gr1)`2 <=d by A1,Th9;
                A57: (gr2)`2=c & a <=(gr2)`1 & (gr2)`1 <=b
                                  by A1,A55,Th11;
                A58: b+(gr1)`2=(gr2)`1 +c by A1,A47,A48,A49,A55,A56,Th11;
                A59: now assume b<>gr2`1; then b>gr2`1 by A57,REAL_1:def 5;
                 hence contradiction by A56,A58,REAL_1:67;
                end;
                  now assume gr1`2<>c; then c <gr1`2 by A56,REAL_1:def 5;
                 hence contradiction by A47,A48,A49,A56,A57,REAL_1:67;
                end;
                then |[(gr1)`1,(gr1)`2]|=g.r2 by A56,A57,A59,EUCLID:57;
                then g.r1=g.r2 by EUCLID:57;
               hence x1=x2 by A23,A24,FUNCT_1:def 8;
              case A60: g.r1 in LSeg(|[b,c]|,|[a,c]|) &
                  g.r2 in LSeg(|[b,c]|,|[b,d]|);
                then A61: (gr2)`1=b & c <=(gr2)`2 & (gr2)`2 <=d by A1,Th9;
                A62: (gr1)`2=c & a <=(gr1)`1 & (gr1)`1 <=b
                                  by A1,A60,Th11;
                A63: b+(gr2)`2=(gr1)`1 +c by A1,A47,A48,A49,A60,A61,Th11;
                A64: now assume b<>gr1`1; then b> gr1`1 by A62,REAL_1:def 5;
                 hence contradiction by A61,A63,REAL_1:67;
                end;
                  now assume gr2`2<> c; then c <gr2`2 by A61,REAL_1:def 5;
                 hence contradiction by A47,A48,A49,A61,A62,REAL_1:67;
                end;
                then |[(gr2)`1,(gr2)`2]|=g.r1 by A61,A62,A64,EUCLID:57;
                then g.r1=g.r2 by EUCLID:57;
               hence x1=x2 by A23,A24,FUNCT_1:def 8;
              case A65: g.r1 in LSeg(|[b,c]|,|[a,c]|) &
                  g.r2 in LSeg(|[b,c]|,|[a,c]|);
                then A66: (gr1)`2=c & a <=(gr1)`1 & (gr1)`1 <=b by A1,Th11;
                A67: (gr2)`2=c & a <=(gr2)`1 & (gr2)`1 <=b
                                  by A1,A65,Th11;
                 then (gr1)`1=(gr2)`1 by A47,A48,A49,A66,XCMPLX_1:2;
                then |[(gr1)`1,(gr1)`2]|=g.r2 by A66,A67,EUCLID:57;
                then g.r1=g.r2 by EUCLID:57;
               hence x1=x2 by A23,A24,FUNCT_1:def 8;
              end;
             hence x1=x2;
            end;
           hence thesis by FUNCT_1:def 8;
          end;
          A68: dom k=[.0,1.] by BORSUK_1:83,FUNCT_2:def 1;
          then s1 in dom k by A22,TOPREAL5:1;
          then rxc=s1 by A39,A45,A46,A68,FUNCT_1:def 8;
         hence contradiction by A32,A45,TOPREAL5:1;
        end;
      hence s1 <= s2;
     end;
    then p1 in Lower_Arc(K) & p2 in Lower_Arc(K) & not p2=W-min(K) &
     LE p1,p2,Lower_Arc(K),E-max(K),W-min(K)
                        by A1,A8,A21,JORDAN5C:def 3;
   hence LE p1,p2,K by JORDAN6:def 10;
   end;
   hence thesis;
  end;
end;

theorem Th68: for K being non empty compact Subset of TOP-REAL 2,
a,b,c,d being real number,p1,p2 being Point of TOP-REAL 2
 st K=rectangle(a,b,c,d) & a<b & c <d
 & p1 in LSeg(|[a,c]|,|[b,c]|)
 & p2 in LSeg(|[a,c]|,|[b,c]|)
holds LE p1,p2,K & p1<>W-min(K) iff p1`1>=p2`1 & p2<>W-min(K)
proof let K be non empty compact Subset of TOP-REAL 2,
a,b,c,d be real number,p1,p2 be Point of TOP-REAL 2;
 assume A1: K=rectangle(a,b,c,d) & a<b & c <d
 & p1 in LSeg(|[a,c]|,|[b,c]|)
 & p2 in LSeg(|[a,c]|,|[b,c]|);
 then A2: K is_simple_closed_curve by Th60;
 A3: p1`2=c & a <=p1`1 & p1`1 <= b by A1,Th11;
 A4: p2`2=c & a <=p2`1 & p2`1 <= b by A1,Th11;
  A5: W-min(K)= |[a,c]| by A1,Th56;
  A6: E-max(K)= |[b,d]| by A1,Th56;
  A7: Lower_Arc(K)=LSeg(|[b,d]|,|[b,c]|) \/ LSeg(|[b,c]|,|[a,c]|) by A1,Th62
;
    then A8: LSeg(|[b,c]|,|[a,c]|) c= Lower_Arc(K) by XBOOLE_1:7;
    then A9: p1 in Lower_Arc(K) by A1;
    A10: Lower_Arc(K) c= K by A2,JORDAN1A:16;
    A11: Upper_Arc(K) /\ Lower_Arc(K)={W-min(K),E-max(K)}
                              by A2,JORDAN6:def 9;
    A12: now assume p1 in Upper_Arc(K);
      then p1 in Upper_Arc(K) /\ Lower_Arc(K) by A1,A8,XBOOLE_0:def 3;
      then p1=W-min(K) or p1=E-max(K) by A11,TARSKI:def 2;
     hence p1=W-min(K) by A1,A3,A6,EUCLID:56;
    end;
 thus LE p1,p2,K & p1<>W-min(K) implies p1`1>=p2`1 & p2<>W-min(K)
  proof assume A13: LE p1,p2,K & p1<>W-min(K);
     then A14: p1 in Upper_Arc(K) & p2 in Lower_Arc(K)& not p2=W-min(K) or
     p1 in Upper_Arc(K) & p2 in Upper_Arc(K) &
     LE p1,p2,Upper_Arc(K),W-min(K),E-max(K) or
     p1 in Lower_Arc(K) & p2 in Lower_Arc(K)& not p2=W-min(K) &
     LE p1,p2,Lower_Arc(K),E-max(K),W-min(K) by JORDAN6:def 10;
    consider f being map of I[01],(TOP-REAL 2)|(Lower_Arc(K)) such that
    A15: f is_homeomorphism
     & f.0=E-max(K) & f.1=W-min(K) & rng f=Lower_Arc(K) &
(for r being Real st r in [.0,1/2.] holds f.r=(1-2*r)*|[b,d]|+(2*r)*|[b,c]|)&
(for r being Real st r in [.1/2,1.] holds
    f.r=(1-(2*r-1))*|[b,c]|+(2*r-1)*|[a,c]|)&
(for p being Point of TOP-REAL 2 st p in LSeg(|[b,d]|,|[b,c]|)
  holds 0<=((p`2)-d)/(c-d)/2 & ((p`2)-d)/(c-d)/2<=1
        & f.(((p`2)-d)/(c-d)/2)=p)&
(for p being Point of TOP-REAL 2 st p in LSeg(|[b,c]|,|[a,c]|)
  holds 0<=((p`1)-b)/(a-b)/2+1/2 & ((p`1)-b)/(a-b)/2+1/2<=1
        & f.(((p`1)-b)/(a-b)/2+1/2)=p)
                         by A1,Th64;
     reconsider s1=((p1`1)-b)/(a-b)/2+1/2,s2=((p2`1)-b)/(a-b)/2+1/2 as
         Real by XREAL_0:def 1;
     A16: f.s1=p1 by A1,A15;
     A17: f.s2=p2 by A1,A15;
       b-a>0 by A1,SQUARE_1:11;
     then -(b-a)< -0 by REAL_1:50;
     then 0-(b-a) < 0 by XCMPLX_1:150;
     then 0-b+a <0 by XCMPLX_1:37;
     then -b+a <0 by XCMPLX_1:150;
     then A18: a -b <0 by XCMPLX_0:def 8;
     A19: 0<=s1 & s1<=1 by A1,A15;
       0<=s2 & s2<=1 by A1,A15;
     then s1<=s2 by A12,A13,A14,A15,A16,A17,A19,JORDAN5C:def 3;
     then ((p1`1)-b)/(a-b)/2<= ((p2`1)-b)/(a-b)/2 by REAL_1:53;
     then ((p1`1)-b)/(a-b)/2*2<=((p2`1)-b)/(a-b)/2*2 by AXIOMS:25;
     then ((p1`1)-b)/(a-b)<= ((p2`1)-b)/(a-b)/2*2 by XCMPLX_1:88;
     then ((p1`1)-b)/(a-b)<= ((p2`1)-b)/(a-b) by XCMPLX_1:88;
     then ((p1`1)-b)/(a-b)*(a-b)>= ((p2`1)-b)/(a-b)*(a-b)
                   by A18,REAL_1:52;
     then ((p1`1)-b)>= ((p2`1)-b)/(a-b)*(a-b) by A18,XCMPLX_1:88;
     then ((p1`1)-b)>= ((p2`1)-b) by A18,XCMPLX_1:88;
     then ((p1`1)-b)+b >= ((p2`1)-b)+b by REAL_1:55;
     then (p1`1) >= ((p2`1)-b)+b by XCMPLX_1:27;
   hence p1`1>=p2`1 by XCMPLX_1:27;
     now assume A20: p2=W-min(K);
     then LE p2,p1,K by A2,A9,A10,JORDAN7:3;
    hence contradiction by A2,A13,A20,JORDAN6:72;
   end;
   hence thesis;
  end;

 thus p1`1>=p2`1 & p2<>W-min(K) implies LE p1,p2,K & p1<>W-min(K)
  proof assume A21: p1`1>=p2`1 & p2<>W-min(K);
    A22: for g being map of I[01], (TOP-REAL 2)|Lower_Arc(K),
        s1, s2 being Real st
        g is_homeomorphism
        & g.0 = E-max(K) & g.1 = W-min(K)
        & g.s1 = p1 & 0 <= s1 & s1 <= 1
        & g.s2 = p2 & 0 <= s2 & s2 <= 1
        holds s1 <= s2
     proof let g be map of I[01], (TOP-REAL 2)|Lower_Arc(K),
        s1, s2 be Real;
        assume A23: g is_homeomorphism
        & g.0 = E-max(K) & g.1 = W-min(K)
        & g.s1 = p1 & 0 <= s1 & s1 <= 1
        & g.s2 = p2 & 0 <= s2 & s2 <= 1;
        A24: dom g=the carrier of I[01] by FUNCT_2:def 1;
        A25: g is one-to-one by A23,TOPS_2:def 5;
        A26: the carrier of ((TOP-REAL 2)|Lower_Arc(K))
        =Lower_Arc(K) by JORDAN1:1;
        then g is Function of the carrier of I[01],
              the carrier of TOP-REAL 2 by FUNCT_2:9;
        then reconsider g1=g as map of I[01],TOP-REAL 2 ;
          g is continuous by A23,TOPS_2:def 5;
        then A27: g1 is continuous by TOPMETR:7;
        reconsider h1=proj1 as map of TOP-REAL 2,R^1 by JGRAPH_2:16;
        reconsider h2=proj2 as map of TOP-REAL 2,R^1 by JGRAPH_2:17;
        A28: (TOP-REAL 2)|([#](TOP-REAL 2))=TOP-REAL 2 by TSEP_1:3;
        then A29: (for p being Point of (TOP-REAL 2)|([#](TOP-REAL 2)) holds
        h1.p=proj1.p) implies h1 is continuous by JGRAPH_2:39;
          (for p being Point of (TOP-REAL 2)|([#](TOP-REAL 2)) holds
        h2.p=proj2.p) implies h2 is continuous by A28,JGRAPH_2:40;
        then consider h being map of TOP-REAL 2,R^1
        such that
        A30: (for p being Point of TOP-REAL 2,
          r1,r2 being real number st
        h1.p=r1 & h2.p=r2 holds h.p=r1+r2) & h is continuous
                          by A29,JGRAPH_2:29;
        reconsider k=h*g1 as map of I[01],R^1;
        A31: k is continuous map of I[01],R^1 by A27,A30,TOPS_2:58;
        A32: E-max K=|[b,d]| by A1,Th56;
          now assume A33: s1>s2;
          A34: dom g=[.0,1.] by BORSUK_1:83,FUNCT_2:def 1;
            0 in [.0,1.] by TOPREAL5:1;
          then A35: k.0=h.(E-max(K)) by A23,A34,FUNCT_1:23
          .=h1.(E-max(K))+h2.(E-max(K)) by A30
          .=(E-max(K))`1+proj2.(E-max(K)) by PSCOMP_1:def 28
          .=(E-max(K))`1+(E-max(K))`2 by PSCOMP_1:def 29
          .=(E-max(K))`1+d by A32,EUCLID:56
          .=b+d by A32,EUCLID:56;
            s1 in [.0,1.] by A23,TOPREAL5:1;
          then A36: k.s1=h.p1 by A23,A34,FUNCT_1:23
          .=proj1.p1 +proj2.p1 by A30 .=p1`1+proj2.p1 by PSCOMP_1:def 28
          .=p1`1 +c by A3,PSCOMP_1:def 29;
          A37: s2 in [.0,1.] by A23,TOPREAL5:1;
          then k.s2=h.p2 by A23,A34,FUNCT_1:23
          .=h1.p2+h2.p2 by A30
          .=p2`1+proj2.p2 by PSCOMP_1:def 28
          .=p2`1 +c by A4,PSCOMP_1:def 29;
          then A38: k.0>=k.s1 & k.s1>=k.s2 by A1,A3,A21,A35,A36,REAL_1:55;
          A39: 0 in [.0,1.] by TOPREAL5:1;
          then A40: [.0,s2.] c= [.0,1.] by A37,RCOMP_1:16;
          then reconsider B=[.0,s2.] as Subset of I[01] by BORSUK_1:83
;
          A41: B is connected by A23,A37,A39,BORSUK_1:83,BORSUK_4:49;
          A42: 0 in B by A23,TOPREAL5:1;
          A43: s2 in B by A23,TOPREAL5:1;
          A44: k.0 is Real by XREAL_0:def 1;
          A45: k.s2 is Real by XREAL_0:def 1;
            k.s1 is Real by XREAL_0:def 1;
          then consider xc being Point of I[01] such that
          A46: xc in B & k.xc =k.s1
                   by A31,A38,A41,A42,A43,A44,A45,TOPREAL5:11;
            xc in [.0,1.] by BORSUK_1:83;
          then reconsider rxc=xc as Real;
          A47: k is one-to-one
          proof
              for x1,x2 being set st x1 in dom k & x2 in dom k &
            k.x1=k.x2 holds x1=x2
            proof let x1,x2 be set;
             assume A48: x1 in dom k & x2 in dom k &
               k.x1=k.x2;
              then reconsider r1=x1 as Point of I[01] by FUNCT_2:def 1;
              reconsider r2=x2 as Point of I[01] by A48,FUNCT_2:def 1;
              A49: k.x1=h.(g1.x1) by A48,FUNCT_1:22
                 .=h1.(g1.r1)+h2.(g1.r1) by A30
                 .=(g1.r1)`1+proj2.(g1.r1) by PSCOMP_1:def 28
                 .=(g1.r1)`1+(g1.r1)`2 by PSCOMP_1:def 29;
              A50: k.x2=h.(g1.x2) by A48,FUNCT_1:22
                 .=h1.(g1.r2)+h2.(g1.r2) by A30
                 .=(g1.r2)`1+proj2.(g1.r2) by PSCOMP_1:def 28
                 .=(g1.r2)`1+(g1.r2)`2 by PSCOMP_1:def 29;
              A51: g.r1 in Lower_Arc(K) by A26;
              A52: g.r2 in Lower_Arc(K) by A26;
              reconsider gr1=g.r1 as Point of TOP-REAL 2 by A51;
              reconsider gr2=g.r2 as Point of TOP-REAL 2 by A52;
                now per cases by A7,A26,XBOOLE_0:def 2;
              case A53: g.r1 in LSeg(|[b,d]|,|[b,c]|) &
                  g.r2 in LSeg(|[b,d]|,|[b,c]|);
                then A54: (gr1)`1=b & c <=(gr1)`2 & (gr1)`2 <=d by A1,Th9;
                A55: (gr2)`1=b & c <=(gr2)`2 & (gr2)`2 <=d
                                  by A1,A53,Th9;
                 then (gr1)`2=(gr2)`2 by A48,A49,A50,A54,XCMPLX_1:2;
                then |[(gr1)`1,(gr1)`2]|=g.r2 by A54,A55,EUCLID:57;
                then g.r1=g.r2 by EUCLID:57;
               hence x1=x2 by A24,A25,FUNCT_1:def 8;
              case A56: g.r1 in LSeg(|[b,d]|,|[b,c]|) &
                  g.r2 in LSeg(|[b,c]|,|[a,c]|);
                then A57: (gr1)`1=b & c <=(gr1)`2 & (gr1)`2 <=d by A1,Th9;
                A58: (gr2)`2=c & a <=(gr2)`1 & (gr2)`1 <=b
                                  by A1,A56,Th11;
                A59: b+(gr1)`2=(gr2)`1 +c by A1,A48,A49,A50,A56,A57,Th11;
                A60: now assume b<>gr2`1; then b>gr2`1 by A58,REAL_1:def 5;
                 hence contradiction by A57,A59,REAL_1:67;
                end;
                  now assume gr1`2<>c; then c <gr1`2 by A57,REAL_1:def 5;
                 hence contradiction by A48,A49,A50,A57,A58,REAL_1:67;
                end;
                then |[(gr1)`1,(gr1)`2]|=g.r2 by A57,A58,A60,EUCLID:57;
                then g.r1=g.r2 by EUCLID:57;
               hence x1=x2 by A24,A25,FUNCT_1:def 8;
              case A61: g.r1 in LSeg(|[b,c]|,|[a,c]|) &
                  g.r2 in LSeg(|[b,d]|,|[b,c]|);
                then A62: (gr2)`1=b & c <=(gr2)`2 & (gr2)`2 <=d by A1,Th9;
                A63: (gr1)`2=c & a <=(gr1)`1 & (gr1)`1 <=b
                                  by A1,A61,Th11;
                A64: b+(gr2)`2=(gr1)`1 +c by A1,A48,A49,A50,A61,A62,Th11;
                A65: now assume b<>gr1`1; then b>gr1`1 by A63,REAL_1:def 5;
                 hence contradiction by A62,A64,REAL_1:67;
                end;
                  now assume gr2`2<>c; then c <gr2`2 by A62,REAL_1:def 5;
                 hence contradiction by A48,A49,A50,A62,A63,REAL_1:67;
                end;
                then |[(gr2)`1,(gr2)`2]|=g.r1 by A62,A63,A65,EUCLID:57;
                then g.r1=g.r2 by EUCLID:57;
               hence x1=x2 by A24,A25,FUNCT_1:def 8;
              case A66: g.r1 in LSeg(|[b,c]|,|[a,c]|) &
                  g.r2 in LSeg(|[b,c]|,|[a,c]|);
                then A67: (gr1)`2=c & a <=(gr1)`1 & (gr1)`1 <=b by A1,Th11;
                A68: (gr2)`2=c & a <=(gr2)`1 & (gr2)`1 <=b
                                  by A1,A66,Th11;
                 then (gr1)`1=(gr2)`1 by A48,A49,A50,A67,XCMPLX_1:2;
                then |[(gr1)`1,(gr1)`2]|=g.r2 by A67,A68,EUCLID:57;
                then g.r1=g.r2 by EUCLID:57;
               hence x1=x2 by A24,A25,FUNCT_1:def 8;
              end;
             hence x1=x2;
            end;
           hence thesis by FUNCT_1:def 8;
          end;
          A69: dom k=[.0,1.] by BORSUK_1:83,FUNCT_2:def 1;
          then s1 in dom k by A23,TOPREAL5:1;
          then rxc=s1 by A40,A46,A47,A69,FUNCT_1:def 8;
         hence contradiction by A33,A46,TOPREAL5:1;
        end;
      hence s1 <= s2;
     end;
    A70: now assume A71: p1=W-min(K);
      then p1`1=a & p1`2=c by A5,EUCLID:56;
      then p1`1=p2`1 by A4,A21,AXIOMS:21;
      then |[(p1)`1,(p1)`2]|=p2 by A3,A4,EUCLID:57;
     hence contradiction by A21,A71,EUCLID:57;
    end;
      p1 in Lower_Arc(K) & p2 in Lower_Arc(K) & not p2=W-min(K) &
     LE p1,p2,Lower_Arc(K),E-max(K),W-min(K)
                        by A1,A8,A21,A22,JORDAN5C:def 3;
   hence LE p1,p2,K by JORDAN6:def 10;
   thus thesis by A70;
  end;
end;

theorem Th69: for K being non empty compact Subset of TOP-REAL 2,
a,b,c,d being real number,p1,p2 being Point of TOP-REAL 2
 st K=rectangle(a,b,c,d) & a<b & c <d
 & p1 in LSeg(|[a,c]|,|[a,d]|)
holds LE p1,p2,K iff p2 in LSeg(|[a,c]|,|[a,d]|) & p1`2<=p2`2
or p2 in LSeg(|[a,d]|,|[b,d]|) or p2 in LSeg(|[b,d]|,|[b,c]|)
or p2 in LSeg(|[b,c]|,|[a,c]|) & p2<>W-min(K)
proof let K be non empty compact Subset of TOP-REAL 2,
a,b,c,d be real number,p1,p2 be Point of TOP-REAL 2;
 assume A1: K=rectangle(a,b,c,d) & a<b & c <d
    & p1 in LSeg(|[a,c]|,|[a,d]|);
  then A2: K is_simple_closed_curve by Th60;
    Upper_Arc(K)=LSeg(|[a,c]|,|[a,d]|) \/ LSeg(|[a,d]|,|[b,d]|)
                        by A1,Th61;
  then A3: LSeg(|[a,c]|,|[a,d]|) c= Upper_Arc(K) by XBOOLE_1:7;

  A4: p1`1=a & c <=p1`2 & p1`2 <= d by A1,Th9;
  thus LE p1,p2,K implies p2 in LSeg(|[a,c]|,|[a,d]|) & p1`2<=p2`2
  or p2 in LSeg(|[a,d]|,|[b,d]|) or p2 in LSeg(|[b,d]|,|[b,c]|)
  or p2 in LSeg(|[b,c]|,|[a,c]|) & p2<>W-min(K)
  proof assume A5: LE p1,p2,K;
    then A6: p1 in K & p2 in K by A2,JORDAN7:5;
      K= {p: p`1=a & c <=p`2 & p`2<=d or p`2=d & a<=p`1 & p`1<=b or
      p`1=b & c <=p`2 & p`2<=d or p`2=c & a<=p`1 & p`1<=b} by A1,Def1;
    then K= LSeg(|[ a,c ]|, |[ a,d ]|) \/ LSeg(|[ a,d ]|, |[ b,d ]|)
    \/ (LSeg(|[ a,c ]|, |[ b,c ]|) \/ LSeg(|[ b,c ]|, |[ b,d ]|)) by A1,Th40
    .=LSeg(|[a,c]|,|[a,d]|) \/ LSeg(|[a,d]|,|[b,d]|)
    \/ LSeg(|[b,d]|,|[b,c]|) \/ LSeg(|[b,c]|,|[a,c]|) by XBOOLE_1:4;
    then p2 in LSeg(|[a,c]|,|[a,d]|) \/ LSeg(|[a,d]|,|[b,d]|)
    \/ LSeg(|[b,d]|,|[b,c]|) or
    p2 in LSeg(|[b,c]|,|[a,c]|) by A6,XBOOLE_0:def 2;
    then A7: p2 in LSeg(|[a,c]|,|[a,d]|) \/ LSeg(|[a,d]|,|[b,d]|) or
    p2 in LSeg(|[b,d]|,|[b,c]|) or
    p2 in LSeg(|[b,c]|,|[a,c]|) by XBOOLE_0:def 2;
      now per cases by A7,XBOOLE_0:def 2;
    case p2 in LSeg(|[a,c]|,|[a,d]|);
     hence p2 in LSeg(|[a,c]|,|[a,d]|) & p1`2<=p2`2
      or p2 in LSeg(|[a,d]|,|[b,d]|) or p2 in LSeg(|[b,d]|,|[b,c]|)
      or p2 in LSeg(|[b,c]|,|[a,c]|) & p2<>W-min(K) by A1,A5,Th65;
    case p2 in LSeg(|[a,d]|,|[b,d]|);
     hence p2 in LSeg(|[a,c]|,|[a,d]|) & p1`2<=p2`2
      or p2 in LSeg(|[a,d]|,|[b,d]|) or p2 in LSeg(|[b,d]|,|[b,c]|)
      or p2 in LSeg(|[b,c]|,|[a,c]|) & p2<>W-min(K);

    case p2 in LSeg(|[b,d]|,|[b,c]|);
     hence p2 in LSeg(|[a,c]|,|[a,d]|) & p1`2<=p2`2
      or p2 in LSeg(|[a,d]|,|[b,d]|) or p2 in LSeg(|[b,d]|,|[b,c]|)
      or p2 in LSeg(|[b,c]|,|[a,c]|) & p2<>W-min(K);
    case A8: p2 in LSeg(|[b,c]|,|[a,c]|);
        now per cases;
      case p2=W-min(K); then LE p2,p1,K by A2,A6,JORDAN7:3;
       hence p2 in LSeg(|[a,c]|,|[a,d]|) & p1`2<=p2`2
        or p2 in LSeg(|[a,d]|,|[b,d]|) or p2 in LSeg(|[b,d]|,|[b,c]|)
        or p2 in LSeg(|[b,c]|,|[a,c]|) & p2<>W-min(K)
                         by A1,A2,A5,JORDAN6:72;
      case p2<>W-min(K);
       hence p2 in LSeg(|[a,c]|,|[a,d]|) & p1`2<=p2`2
        or p2 in LSeg(|[a,d]|,|[b,d]|) or p2 in LSeg(|[b,d]|,|[b,c]|)
        or p2 in LSeg(|[b,c]|,|[a,c]|) & p2<>W-min(K) by A8;
      end;
     hence p2 in LSeg(|[a,c]|,|[a,d]|) & p1`2<=p2`2
      or p2 in LSeg(|[a,d]|,|[b,d]|) or p2 in LSeg(|[b,d]|,|[b,c]|)
      or p2 in LSeg(|[b,c]|,|[a,c]|) & p2<>W-min(K);
    end;
   hence p2 in LSeg(|[a,c]|,|[a,d]|) & p1`2<=p2`2
  or p2 in LSeg(|[a,d]|,|[b,d]|) or p2 in LSeg(|[b,d]|,|[b,c]|)
  or p2 in LSeg(|[b,c]|,|[a,c]|) & p2<>W-min(K);
  end;
  A9: W-min(K)= |[a,c]| by A1,Th56;
  thus p2 in LSeg(|[a,c]|,|[a,d]|) & p1`2<=p2`2
  or p2 in LSeg(|[a,d]|,|[b,d]|) or p2 in LSeg(|[b,d]|,|[b,c]|)
  or p2 in LSeg(|[b,c]|,|[a,c]|) & p2<>W-min(K) implies LE p1,p2,K
  proof assume A10: p2 in LSeg(|[a,c]|,|[a,d]|) & p1`2<=p2`2
    or p2 in LSeg(|[a,d]|,|[b,d]|) or p2 in LSeg(|[b,d]|,|[b,c]|)
    or p2 in LSeg(|[b,c]|,|[a,c]|) & p2<>W-min(K);
      now per cases by A10;
    case p2 in LSeg(|[a,c]|,|[a,d]|) & p1`2<=p2`2;
     hence LE p1,p2,K by A1,Th65;
    case A11: p2 in LSeg(|[a,d]|,|[b,d]|);
      then A12: p2`2=d & a <=p2`1 & p2`1 <= b by A1,Th11;

      A13: Upper_Arc(K)=LSeg(|[a,c]|,|[a,d]|) \/ LSeg(|[a,d]|,|[b,d]|)
                        by A1,Th61;
      then A14: p2 in Upper_Arc(K) by A11,XBOOLE_0:def 2;
      A15: p1 in Upper_Arc(K) by A1,A13,XBOOLE_0:def 2;
        LE p1,p2,Upper_Arc(K),W-min(K),E-max(K)
       proof
      for g being map of I[01], (TOP-REAL 2)|Upper_Arc(K),
        s1, s2 being Real st
        g is_homeomorphism
        & g.0 = W-min(K) & g.1 = E-max(K)
        & g.s1 = p1 & 0 <= s1 & s1 <= 1
        & g.s2 = p2 & 0 <= s2 & s2 <= 1
        holds s1 <= s2
     proof let g be map of I[01], (TOP-REAL 2)|Upper_Arc(K),
        s1, s2 be Real;
        assume A16: g is_homeomorphism
        & g.0 = W-min(K) & g.1 = E-max(K)
        & g.s1 = p1 & 0 <= s1 & s1 <= 1
        & g.s2 = p2 & 0 <= s2 & s2 <= 1;
        A17: dom g=the carrier of I[01] by FUNCT_2:def 1;
        A18: g is one-to-one by A16,TOPS_2:def 5;
        A19: the carrier of ((TOP-REAL 2)|Upper_Arc(K))
        =Upper_Arc(K) by JORDAN1:1;
        then g is Function of the carrier of I[01],
              the carrier of TOP-REAL 2 by FUNCT_2:9;
        then reconsider g1=g as map of I[01],TOP-REAL 2 ;
          g is continuous by A16,TOPS_2:def 5;
        then A20: g1 is continuous by TOPMETR:7;
        reconsider h1=proj1 as map of TOP-REAL 2,R^1 by JGRAPH_2:16;
        reconsider h2=proj2 as map of TOP-REAL 2,R^1 by JGRAPH_2:17;
        A21: (TOP-REAL 2)|([#](TOP-REAL 2))=TOP-REAL 2 by TSEP_1:3;
        then A22: (for p being Point of (TOP-REAL 2)|([#](TOP-REAL 2)) holds
        h1.p=proj1.p) implies h1 is continuous by JGRAPH_2:39;
          (for p being Point of (TOP-REAL 2)|([#](TOP-REAL 2)) holds
        h2.p=proj2.p) implies h2 is continuous by A21,JGRAPH_2:40;
        then consider h being map of TOP-REAL 2,R^1
        such that
        A23: (for p being Point of TOP-REAL 2,
          r1,r2 being real number st
        h1.p=r1 & h2.p=r2 holds h.p=r1+r2) & h is continuous
                          by A22,JGRAPH_2:29;
        reconsider k=h*g1 as map of I[01],R^1;
        A24: k is continuous map of I[01],R^1 by A20,A23,TOPS_2:58;
        A25: W-min K=|[a,c]| by A1,Th56;
          now assume A26: s1>s2;
          A27: dom g=[.0,1.] by BORSUK_1:83,FUNCT_2:def 1;
            0 in [.0,1.] by TOPREAL5:1;
          then A28: k.0=h.(W-min(K)) by A16,A27,FUNCT_1:23
          .=h1.(W-min(K))+h2.(W-min(K)) by A23
          .=(W-min(K))`1+proj2.(W-min(K)) by PSCOMP_1:def 28
          .=(W-min(K))`1+(W-min(K))`2 by PSCOMP_1:def 29
          .=(W-min(K))`1+c by A25,EUCLID:56
          .=a+c by A25,EUCLID:56;
            s1 in [.0,1.] by A16,TOPREAL5:1;
          then A29: k.s1=h.p1 by A16,A27,FUNCT_1:23
          .=proj1.p1 +proj2.p1 by A23 .=p1`1+proj2.p1 by PSCOMP_1:def 28
          .=a +p1`2 by A4,PSCOMP_1:def 29;
          A30: s2 in [.0,1.] by A16,TOPREAL5:1;
          then k.s2=h.p2 by A16,A27,FUNCT_1:23
          .=proj1.p2 +proj2.p2 by A23 .=p2`1+proj2.p2 by PSCOMP_1:def 28
          .=p2`1 +d by A12,PSCOMP_1:def 29;
          then A31: k.0<=k.s1 & k.s1<=k.s2 by A4,A12,A28,A29,REAL_1:55;
          A32: 0 in [.0,1.] by TOPREAL5:1;
          then A33: [.0,s2.] c= [.0,1.] by A30,RCOMP_1:16;
          then reconsider B=[.0,s2.] as Subset of I[01] by BORSUK_1:83
;
          A34: B is connected by A16,A30,A32,BORSUK_1:83,BORSUK_4:49;
          A35: 0 in B by A16,TOPREAL5:1;
          A36: s2 in B by A16,TOPREAL5:1;
          A37: k.0 is Real by XREAL_0:def 1;
          A38: k.s2 is Real by XREAL_0:def 1;
            k.s1 is Real by XREAL_0:def 1;
          then consider xc being Point of I[01] such that
          A39: xc in B & k.xc =k.s1
                   by A24,A31,A34,A35,A36,A37,A38,TOPREAL5:11;
            xc in [.0,1.] by BORSUK_1:83;
          then reconsider rxc=xc as Real;
          A40: k is one-to-one
          proof
              for x1,x2 being set st x1 in dom k & x2 in dom k &
            k.x1=k.x2 holds x1=x2
            proof let x1,x2 be set;
             assume A41: x1 in dom k & x2 in dom k &
               k.x1=k.x2;
              then reconsider r1=x1 as Point of I[01] by FUNCT_2:def 1;
              reconsider r2=x2 as Point of I[01] by A41,FUNCT_2:def 1;
              A42: k.x1=h.(g1.x1) by A41,FUNCT_1:22
                 .=h1.(g1.r1)+h2.(g1.r1) by A23
                 .=(g1.r1)`1+proj2.(g1.r1) by PSCOMP_1:def 28
                 .=(g1.r1)`1+(g1.r1)`2 by PSCOMP_1:def 29;
              A43: k.x2=h.(g1.x2) by A41,FUNCT_1:22
                 .=h1.(g1.r2)+h2.(g1.r2) by A23
                 .=(g1.r2)`1+proj2.(g1.r2) by PSCOMP_1:def 28
                 .=(g1.r2)`1+(g1.r2)`2 by PSCOMP_1:def 29;
              A44: g.r1 in Upper_Arc(K) by A19;
              A45: g.r2 in Upper_Arc(K) by A19;
              reconsider gr1=g.r1 as Point of TOP-REAL 2 by A44;
              reconsider gr2=g.r2 as Point of TOP-REAL 2 by A45;
                now per cases by A13,A19,XBOOLE_0:def 2;
              case A46: g.r1 in LSeg(|[a,c]|,|[a,d]|) &
                  g.r2 in LSeg(|[a,c]|,|[a,d]|);
                then A47: (gr1)`1=a & c <=(gr1)`2 & (gr1)`2 <=d by A1,Th9;
                A48: (gr2)`1=a & c <=(gr2)`2 & (gr2)`2 <=d
                                  by A1,A46,Th9;
                 then (gr1)`2=(gr2)`2 by A41,A42,A43,A47,XCMPLX_1:2;
                then |[(gr1)`1,(gr1)`2]|=g.r2 by A47,A48,EUCLID:57;
                then g.r1=g.r2 by EUCLID:57;
               hence x1=x2 by A17,A18,FUNCT_1:def 8;
              case A49: g.r1 in LSeg(|[a,c]|,|[a,d]|) &
                  g.r2 in LSeg(|[a,d]|,|[b,d]|);
                then A50: (gr1)`1=a & c <=(gr1)`2 & (gr1)`2 <=d by A1,Th9;
                A51: (gr2)`2=d & a <=(gr2)`1 & (gr2)`1 <=b
                                  by A1,A49,Th11;
                A52: a+(gr1)`2=(gr2)`1 +d by A1,A41,A42,A43,A49,A50,Th11;
                A53: now assume a<>gr2`1; then a<gr2`1 by A51,REAL_1:def 5;
                 hence contradiction by A50,A52,REAL_1:67;
                end;
                  now assume gr1`2<>d; then d>gr1`2 by A50,REAL_1:def 5;
                 hence contradiction by A41,A42,A43,A50,A51,REAL_1:67;
                end;
                then |[(gr1)`1,(gr1)`2]|=g.r2 by A50,A51,A53,EUCLID:57;
                then g.r1=g.r2 by EUCLID:57;
               hence x1=x2 by A17,A18,FUNCT_1:def 8;
              case A54: g.r1 in LSeg(|[a,d]|,|[b,d]|) &
                  g.r2 in LSeg(|[a,c]|,|[a,d]|);
                then A55: (gr2)`1=a & c <=(gr2)`2 & (gr2)`2 <=d by A1,Th9;
                A56: (gr1)`2=d & a <=(gr1)`1 & (gr1)`1 <=b
                                  by A1,A54,Th11;
                A57: a+(gr2)`2=(gr1)`1 +d by A1,A41,A42,A43,A54,A55,Th11;
                A58: now assume a<>gr1`1; then a<gr1`1 by A56,REAL_1:def 5;
                 hence contradiction by A55,A57,REAL_1:67;
                end;
                  now assume gr2`2<>d; then d>gr2`2 by A55,REAL_1:def 5;
                 hence contradiction by A41,A42,A43,A55,A56,REAL_1:67;
                end;
                then |[(gr2)`1,(gr2)`2]|=g.r1 by A55,A56,A58,EUCLID:57;
                then g.r1=g.r2 by EUCLID:57;
               hence x1=x2 by A17,A18,FUNCT_1:def 8;
              case A59: g.r1 in LSeg(|[a,d]|,|[b,d]|) &
                  g.r2 in LSeg(|[a,d]|,|[b,d]|);
                then A60: (gr1)`2=d & a <=(gr1)`1 & (gr1)`1 <=b by A1,Th11;
                A61: (gr2)`2=d & a <=(gr2)`1 & (gr2)`1 <=b
                                  by A1,A59,Th11;
                 then (gr1)`1=(gr2)`1 by A41,A42,A43,A60,XCMPLX_1:2;
                then |[(gr1)`1,(gr1)`2]|=g.r2 by A60,A61,EUCLID:57;
                then g.r1=g.r2 by EUCLID:57;
               hence x1=x2 by A17,A18,FUNCT_1:def 8;
              end;
             hence x1=x2;
            end;
           hence thesis by FUNCT_1:def 8;
          end;
          A62: dom k=[.0,1.] by BORSUK_1:83,FUNCT_2:def 1;
          then s1 in dom k by A16,TOPREAL5:1;
          then rxc=s1 by A33,A39,A40,A62,FUNCT_1:def 8;
         hence contradiction by A26,A39,TOPREAL5:1;
        end;
      hence s1 <= s2;
     end;
     hence thesis by A14,A15,JORDAN5C:def 3;
       end;
     hence LE p1,p2,K by A14,A15,JORDAN6:def 10;
    case A63: p2 in LSeg(|[b,d]|,|[b,c]|);
   reconsider b2=b,c2=c,d2=d as Real by XREAL_0:def 1;
         p2 in LSeg(|[b2,d2]|,|[b2,c2]|) by A63;
      then A64: p2`1 =b by TOPREAL3:17;
     Lower_Arc(K)=LSeg(|[b,d]|,|[b,c]|) \/ LSeg(|[b,c]|,|[a,c]|) by A1,Th62;
    then A65: LSeg(|[b,d]|,|[b,c]|) c= Lower_Arc(K) by XBOOLE_1:7;
      p2 <> W-min(K) by A1,A9,A64,EUCLID:56;
     hence LE p1,p2,K by A1,A3,A63,A65,JORDAN6:def 10;
    case A66: p2 in LSeg(|[b,c]|,|[a,c]|) & p2<>W-min(K);
     Lower_Arc(K)=LSeg(|[b,d]|,|[b,c]|) \/ LSeg(|[b,c]|,|[a,c]|) by A1,Th62;
    then LSeg(|[b,c]|,|[a,c]|) c= Lower_Arc(K) by XBOOLE_1:7;
     hence LE p1,p2,K by A1,A3,A66,JORDAN6:def 10;
    end;
   hence LE p1,p2,K;
  end;
end;

theorem Th70: for K being non empty compact Subset of TOP-REAL 2,
a,b,c,d being real number,p1,p2 being Point of TOP-REAL 2
 st K=rectangle(a,b,c,d) & a<b & c <d
 & p1 in LSeg(|[a,d]|,|[b,d]|)
holds LE p1,p2,K iff p2 in LSeg(|[a,d]|,|[b,d]|) & p1`1<=p2`1
or p2 in LSeg(|[b,d]|,|[b,c]|)
or p2 in LSeg(|[b,c]|,|[a,c]|) & p2<>W-min(K)
proof let K be non empty compact Subset of TOP-REAL 2,
a,b,c,d be real number,p1,p2 be Point of TOP-REAL 2;
 assume A1: K=rectangle(a,b,c,d) & a<b & c <d
    & p1 in LSeg(|[a,d]|,|[b,d]|);
  then A2: K is_simple_closed_curve by Th60;
    Upper_Arc(K)=LSeg(|[a,c]|,|[a,d]|) \/ LSeg(|[a,d]|,|[b,d]|)
                        by A1,Th61;
  then A3: LSeg(|[a,d]|,|[b,d]|) c= Upper_Arc(K) by XBOOLE_1:7;
  A4: p1`2=d & a <=p1`1 & p1`1 <= b by A1,Th11;
  thus LE p1,p2,K implies
  p2 in LSeg(|[a,d]|,|[b,d]|) & p1`1<=p2`1
  or p2 in LSeg(|[b,d]|,|[b,c]|)
  or p2 in LSeg(|[b,c]|,|[a,c]|) & p2<>W-min(K)
  proof assume A5: LE p1,p2,K;
    then A6: p1 in K & p2 in K by A2,JORDAN7:5;
      K= {p: p`1=a & c <=p`2 & p`2<=d or p`2=d & a<=p`1 & p`1<=b or
      p`1=b & c <=p`2 & p`2<=d or p`2=c & a<=p`1 & p`1<=b} by A1,Def1;
    then K= LSeg(|[ a,c ]|, |[ a,d ]|) \/ LSeg(|[ a,d ]|, |[ b,d ]|)
    \/ (LSeg(|[ a,c ]|, |[ b,c ]|) \/ LSeg(|[ b,c ]|, |[ b,d ]|)) by A1,Th40
    .=LSeg(|[a,c]|,|[a,d]|) \/ LSeg(|[a,d]|,|[b,d]|)
    \/ LSeg(|[b,d]|,|[b,c]|) \/ LSeg(|[b,c]|,|[a,c]|) by XBOOLE_1:4;
    then p2 in LSeg(|[a,c]|,|[a,d]|) \/ LSeg(|[a,d]|,|[b,d]|)
    \/ LSeg(|[b,d]|,|[b,c]|) or
    p2 in LSeg(|[b,c]|,|[a,c]|) by A6,XBOOLE_0:def 2;
    then A7: p2 in LSeg(|[a,c]|,|[a,d]|) \/ LSeg(|[a,d]|,|[b,d]|) or
    p2 in LSeg(|[b,d]|,|[b,c]|) or
    p2 in LSeg(|[b,c]|,|[a,c]|) by XBOOLE_0:def 2;
      now per cases by A7,XBOOLE_0:def 2;
    case p2 in LSeg(|[a,c]|,|[a,d]|);
      then LE p2,p1,K by A1,Th69;
     hence p2 in LSeg(|[a,d]|,|[b,d]|)& p1`1<=p2`1
      or p2 in LSeg(|[b,d]|,|[b,c]|)
      or p2 in LSeg(|[b,c]|,|[a,c]|) & p2<>W-min(K) by A1,A2,A5,JORDAN6:72;
    case p2 in LSeg(|[a,d]|,|[b,d]|);
     hence p2 in LSeg(|[a,d]|,|[b,d]|)& p1`1<=p2`1
      or p2 in LSeg(|[b,d]|,|[b,c]|)
      or p2 in LSeg(|[b,c]|,|[a,c]|) & p2<>W-min(K) by A1,A5,Th66;
    case p2 in LSeg(|[b,d]|,|[b,c]|);
     hence p2 in LSeg(|[a,d]|,|[b,d]|)& p1`1<=p2`1
      or p2 in LSeg(|[b,d]|,|[b,c]|)
      or p2 in LSeg(|[b,c]|,|[a,c]|) & p2<>W-min(K);
    case A8: p2 in LSeg(|[b,c]|,|[a,c]|);
        now per cases;
      case p2=W-min(K); then LE p2,p1,K by A2,A6,JORDAN7:3;
       hence p2 in LSeg(|[a,d]|,|[b,d]|)& p1`1<=p2`1
        or p2 in LSeg(|[b,d]|,|[b,c]|)
        or p2 in LSeg(|[b,c]|,|[a,c]|) & p2<>W-min(K) by A1,A2,A5,JORDAN6:72;
      case p2<>W-min(K);
       hence p2 in LSeg(|[a,d]|,|[b,d]|)& p1`1<=p2`1
        or p2 in LSeg(|[b,d]|,|[b,c]|)
        or p2 in LSeg(|[b,c]|,|[a,c]|) & p2<>W-min(K) by A8;
      end;
     hence p2 in LSeg(|[a,d]|,|[b,d]|)& p1`1<=p2`1
      or p2 in LSeg(|[b,d]|,|[b,c]|)
      or p2 in LSeg(|[b,c]|,|[a,c]|) & p2<>W-min(K);
    end;
    hence p2 in LSeg(|[a,d]|,|[b,d]|)& p1`1<=p2`1
      or p2 in LSeg(|[b,d]|,|[b,c]|)
      or p2 in LSeg(|[b,c]|,|[a,c]|) & p2<>W-min(K);
  end;
  A9: W-min(K)= |[a,c]| by A1,Th56;
  thus
    p2 in LSeg(|[a,d]|,|[b,d]|)& p1`1<=p2`1 or p2 in LSeg(|[b,d]|,|[b,c]|)
  or p2 in LSeg(|[b,c]|,|[a,c]|) & p2<>W-min(K) implies LE p1,p2,K
  proof assume A10: p2 in LSeg(|[a,d]|,|[b,d]|)& p1`1<=p2`1
    or p2 in LSeg(|[b,d]|,|[b,c]|)
    or p2 in LSeg(|[b,c]|,|[a,c]|) & p2<>W-min(K);
      now per cases by A10;
    case A11: p2 in LSeg(|[a,d]|,|[b,d]|)& p1`1<=p2`1;
      then A12: p2`2=d & a <=p2`1 & p2`1 <= b by A1,Th11;

      A13: Upper_Arc(K)=LSeg(|[a,c]|,|[a,d]|) \/ LSeg(|[a,d]|,|[b,d]|)
                        by A1,Th61;
      then A14: p2 in Upper_Arc(K) by A11,XBOOLE_0:def 2;
      A15: p1 in Upper_Arc(K) by A1,A13,XBOOLE_0:def 2;
        LE p1,p2,Upper_Arc(K),W-min(K),E-max(K)
       proof
      for g being map of I[01], (TOP-REAL 2)|Upper_Arc(K),
        s1, s2 being Real st
        g is_homeomorphism
        & g.0 = W-min(K) & g.1 = E-max(K)
        & g.s1 = p1 & 0 <= s1 & s1 <= 1
        & g.s2 = p2 & 0 <= s2 & s2 <= 1
        holds s1 <= s2
     proof let g be map of I[01], (TOP-REAL 2)|Upper_Arc(K),
        s1, s2 be Real;
        assume A16: g is_homeomorphism
        & g.0 = W-min(K) & g.1 = E-max(K)
        & g.s1 = p1 & 0 <= s1 & s1 <= 1
        & g.s2 = p2 & 0 <= s2 & s2 <= 1;
        A17: dom g=the carrier of I[01] by FUNCT_2:def 1;
        A18: g is one-to-one by A16,TOPS_2:def 5;
        A19: the carrier of ((TOP-REAL 2)|Upper_Arc(K))
        =Upper_Arc(K) by JORDAN1:1;
        then g is Function of the carrier of I[01],
              the carrier of TOP-REAL 2 by FUNCT_2:9;
        then reconsider g1=g as map of I[01],TOP-REAL 2 ;
          g is continuous by A16,TOPS_2:def 5;
        then A20: g1 is continuous by TOPMETR:7;
        reconsider h1=proj1 as map of TOP-REAL 2,R^1 by JGRAPH_2:16;
        reconsider h2=proj2 as map of TOP-REAL 2,R^1 by JGRAPH_2:17;
        A21: (TOP-REAL 2)|([#](TOP-REAL 2))=TOP-REAL 2 by TSEP_1:3;
        then A22: (for p being Point of (TOP-REAL 2)|([#](TOP-REAL 2)) holds
        h1.p=proj1.p) implies h1 is continuous by JGRAPH_2:39;
          (for p being Point of (TOP-REAL 2)|([#](TOP-REAL 2)) holds
        h2.p=proj2.p) implies h2 is continuous by A21,JGRAPH_2:40;
        then consider h being map of TOP-REAL 2,R^1
        such that
        A23: (for p being Point of TOP-REAL 2,
          r1,r2 being real number st
        h1.p=r1 & h2.p=r2 holds h.p=r1+r2) & h is continuous
                          by A22,JGRAPH_2:29;
        reconsider k=h*g1 as map of I[01],R^1;
        A24: k is continuous map of I[01],R^1 by A20,A23,TOPS_2:58;
        A25: W-min K=|[a,c]| by A1,Th56;
          now assume A26: s1>s2;
          A27: dom g=[.0,1.] by BORSUK_1:83,FUNCT_2:def 1;
            0 in [.0,1.] by TOPREAL5:1;
          then A28: k.0=h.(W-min(K)) by A16,A27,FUNCT_1:23
          .=h1.(W-min(K))+h2.(W-min(K)) by A23
          .=(W-min(K))`1+proj2.(W-min(K)) by PSCOMP_1:def 28
          .=(W-min(K))`1+(W-min(K))`2 by PSCOMP_1:def 29
          .=(W-min(K))`1+c by A25,EUCLID:56
          .=a+c by A25,EUCLID:56;
            s1 in [.0,1.] by A16,TOPREAL5:1;
          then A29: k.s1=h.p1 by A16,A27,FUNCT_1:23
          .=proj1.p1 +proj2.p1 by A23 .=p1`1+proj2.p1 by PSCOMP_1:def 28
          .=p1`1+d by A4,PSCOMP_1:def 29;
          A30: s2 in [.0,1.] by A16,TOPREAL5:1;
          then k.s2=h.p2 by A16,A27,FUNCT_1:23
          .=proj1.p2 +proj2.p2 by A23 .=p2`1+proj2.p2 by PSCOMP_1:def 28
          .=p2`1 +d by A12,PSCOMP_1:def 29;
          then A31: k.0<=k.s1 & k.s1<=k.s2 by A1,A4,A11,A28,A29,REAL_1:55;
          A32: 0 in [.0,1.] by TOPREAL5:1;
          then A33: [.0,s2.] c= [.0,1.] by A30,RCOMP_1:16;
          then reconsider B=[.0,s2.] as Subset of I[01] by BORSUK_1:83
;
          A34: B is connected by A16,A30,A32,BORSUK_1:83,BORSUK_4:49;
          A35: 0 in B by A16,TOPREAL5:1;
          A36: s2 in B by A16,TOPREAL5:1;
          A37: k.0 is Real by XREAL_0:def 1;
          A38: k.s2 is Real by XREAL_0:def 1;
            k.s1 is Real by XREAL_0:def 1;
          then consider xc being Point of I[01] such that
          A39: xc in B & k.xc =k.s1
                   by A24,A31,A34,A35,A36,A37,A38,TOPREAL5:11;
            xc in [.0,1.] by BORSUK_1:83;
          then reconsider rxc=xc as Real;
          A40: k is one-to-one
          proof
              for x1,x2 being set st x1 in dom k & x2 in dom k &
            k.x1=k.x2 holds x1=x2
            proof let x1,x2 be set;
             assume A41: x1 in dom k & x2 in dom k &
               k.x1=k.x2;
              then reconsider r1=x1 as Point of I[01] by FUNCT_2:def 1;
              reconsider r2=x2 as Point of I[01] by A41,FUNCT_2:def 1;
              A42: k.x1=h.(g1.x1) by A41,FUNCT_1:22
                 .=h1.(g1.r1)+h2.(g1.r1) by A23
                 .=(g1.r1)`1+proj2.(g1.r1) by PSCOMP_1:def 28
                 .=(g1.r1)`1+(g1.r1)`2 by PSCOMP_1:def 29;
              A43: k.x2=h.(g1.x2) by A41,FUNCT_1:22
                 .=h1.(g1.r2)+h2.(g1.r2) by A23
                 .=(g1.r2)`1+proj2.(g1.r2) by PSCOMP_1:def 28
                 .=(g1.r2)`1+(g1.r2)`2 by PSCOMP_1:def 29;
              A44: g.r1 in Upper_Arc(K) by A19;
              A45: g.r2 in Upper_Arc(K) by A19;
              reconsider gr1=g.r1 as Point of TOP-REAL 2 by A44;
              reconsider gr2=g.r2 as Point of TOP-REAL 2 by A45;
                now per cases by A13,A19,XBOOLE_0:def 2;
              case A46: g.r1 in LSeg(|[a,c]|,|[a,d]|) &
                  g.r2 in LSeg(|[a,c]|,|[a,d]|);
                then A47: (gr1)`1=a & c <=(gr1)`2 & (gr1)`2 <=d by A1,Th9;
                A48: (gr2)`1=a & c <=(gr2)`2 & (gr2)`2 <=d
                                  by A1,A46,Th9;
                 then (gr1)`2=(gr2)`2 by A41,A42,A43,A47,XCMPLX_1:2;
                then |[(gr1)`1,(gr1)`2]|=g.r2 by A47,A48,EUCLID:57;
                then g.r1=g.r2 by EUCLID:57;
               hence x1=x2 by A17,A18,FUNCT_1:def 8;
              case A49: g.r1 in LSeg(|[a,c]|,|[a,d]|) &
                  g.r2 in LSeg(|[a,d]|,|[b,d]|);
                then A50: (gr1)`1=a & c <=(gr1)`2 & (gr1)`2 <=d by A1,Th9;
                A51: (gr2)`2=d & a <=(gr2)`1 & (gr2)`1 <=b
                                  by A1,A49,Th11;
                A52: a+(gr1)`2=(gr2)`1 +d by A1,A41,A42,A43,A49,A50,Th11;
                A53: now assume a<>gr2`1; then a<gr2`1 by A51,REAL_1:def 5;
                 hence contradiction by A50,A52,REAL_1:67;
                end;
                  now assume gr1`2<>d; then d>gr1`2 by A50,REAL_1:def 5;
                 hence contradiction by A41,A42,A43,A50,A51,REAL_1:67;
                end;
                then |[(gr1)`1,(gr1)`2]|=g.r2 by A50,A51,A53,EUCLID:57;
                then g.r1=g.r2 by EUCLID:57;
               hence x1=x2 by A17,A18,FUNCT_1:def 8;
              case A54: g.r1 in LSeg(|[a,d]|,|[b,d]|) &
                  g.r2 in LSeg(|[a,c]|,|[a,d]|);
                then A55: (gr2)`1=a & c <=(gr2)`2 & (gr2)`2 <=d by A1,Th9;
                A56: (gr1)`2=d & a <=(gr1)`1 & (gr1)`1 <=b
                                  by A1,A54,Th11;
                A57: a+(gr2)`2=(gr1)`1 +d by A1,A41,A42,A43,A54,A55,Th11;
                A58: now assume a<>gr1`1; then a<gr1`1 by A56,REAL_1:def 5;
                 hence contradiction by A55,A57,REAL_1:67;
                end;
                  now assume gr2`2<>d; then d>gr2`2 by A55,REAL_1:def 5;
                 hence contradiction by A41,A42,A43,A55,A56,REAL_1:67;
                end;
                then |[(gr2)`1,(gr2)`2]|=g.r1 by A55,A56,A58,EUCLID:57;
                then g.r1=g.r2 by EUCLID:57;
               hence x1=x2 by A17,A18,FUNCT_1:def 8;
              case A59: g.r1 in LSeg(|[a,d]|,|[b,d]|) &
                  g.r2 in LSeg(|[a,d]|,|[b,d]|);
                then A60: (gr1)`2=d & a <=(gr1)`1 & (gr1)`1 <=b by A1,Th11;
                A61: (gr2)`2=d & a <=(gr2)`1 & (gr2)`1 <=b
                                  by A1,A59,Th11;
                 then (gr1)`1=(gr2)`1 by A41,A42,A43,A60,XCMPLX_1:2;
                then |[(gr1)`1,(gr1)`2]|=g.r2 by A60,A61,EUCLID:57;
                then g.r1=g.r2 by EUCLID:57;
               hence x1=x2 by A17,A18,FUNCT_1:def 8;
              end;
             hence x1=x2;
            end;
           hence thesis by FUNCT_1:def 8;
          end;
          A62: dom k=[.0,1.] by BORSUK_1:83,FUNCT_2:def 1;
          then s1 in dom k by A16,TOPREAL5:1;
          then rxc=s1 by A33,A39,A40,A62,FUNCT_1:def 8;
         hence contradiction by A26,A39,TOPREAL5:1;
        end;
      hence s1 <= s2;
     end;
     hence thesis by A14,A15,JORDAN5C:def 3;
       end;
     hence LE p1,p2,K by A14,A15,JORDAN6:def 10;
    case A63: p2 in LSeg(|[b,d]|,|[b,c]|);
   reconsider b2=b,c2=c,d2=d as Real by XREAL_0:def 1;
         p2 in LSeg(|[b2,d2]|,|[b2,c2]|) by A63;
      then A64: p2`1 =b by TOPREAL3:17;
     Lower_Arc(K)=LSeg(|[b,d]|,|[b,c]|) \/ LSeg(|[b,c]|,|[a,c]|) by A1,Th62;
    then A65: LSeg(|[b,d]|,|[b,c]|) c= Lower_Arc(K) by XBOOLE_1:7;
      p2 <> W-min(K) by A1,A9,A64,EUCLID:56;
     hence LE p1,p2,K by A1,A3,A63,A65,JORDAN6:def 10;
    case A66: p2 in LSeg(|[b,c]|,|[a,c]|) & p2<>W-min(K);
     Lower_Arc(K)=LSeg(|[b,d]|,|[b,c]|) \/ LSeg(|[b,c]|,|[a,c]|) by A1,Th62;
    then LSeg(|[b,c]|,|[a,c]|) c= Lower_Arc(K) by XBOOLE_1:7;
     hence LE p1,p2,K by A1,A3,A66,JORDAN6:def 10;
    end;
   hence LE p1,p2,K;
  end;
end;

theorem Th71: for K being non empty compact Subset of TOP-REAL 2,
a,b,c,d being real number,p1,p2 being Point of TOP-REAL 2
 st K=rectangle(a,b,c,d) & a<b & c <d
 & p1 in LSeg(|[b,d]|,|[b,c]|)
holds LE p1,p2,K iff p2 in LSeg(|[b,d]|,|[b,c]|) & p1`2>=p2`2
or p2 in LSeg(|[b,c]|,|[a,c]|) & p2<>W-min(K)
proof let K be non empty compact Subset of TOP-REAL 2,
a,b,c,d be real number,p1,p2 be Point of TOP-REAL 2;
 assume A1: K=rectangle(a,b,c,d) & a<b & c <d
    & p1 in LSeg(|[b,d]|,|[b,c]|);
  then A2: K is_simple_closed_curve by Th60;
  A3: p1`1=b & c <=p1`2 & p1`2 <= d by A1,Th9;
  thus LE p1,p2,K implies
  p2 in LSeg(|[b,d]|,|[b,c]|) & p1`2>=p2`2
  or p2 in LSeg(|[b,c]|,|[a,c]|) & p2<>W-min(K)
  proof assume A4: LE p1,p2,K;
    then A5: p1 in K & p2 in K by A2,JORDAN7:5;
      K= {p: p`1=a & c <=p`2 & p`2<=d or p`2=d & a<=p`1 & p`1<=b or
      p`1=b & c <=p`2 & p`2<=d or p`2=c & a<=p`1 & p`1<=b} by A1,Def1;
    then K= LSeg(|[ a,c ]|, |[ a,d ]|) \/ LSeg(|[ a,d ]|, |[ b,d ]|)
    \/ (LSeg(|[ a,c ]|, |[ b,c ]|) \/ LSeg(|[ b,c ]|, |[ b,d ]|)) by A1,Th40
    .=LSeg(|[a,c]|,|[a,d]|) \/ LSeg(|[a,d]|,|[b,d]|)
    \/ LSeg(|[b,d]|,|[b,c]|) \/ LSeg(|[b,c]|,|[a,c]|) by XBOOLE_1:4;
    then p2 in LSeg(|[a,c]|,|[a,d]|) \/ LSeg(|[a,d]|,|[b,d]|)
    \/ LSeg(|[b,d]|,|[b,c]|) or
    p2 in LSeg(|[b,c]|,|[a,c]|) by A5,XBOOLE_0:def 2;
    then A6: p2 in LSeg(|[a,c]|,|[a,d]|) \/ LSeg(|[a,d]|,|[b,d]|) or
    p2 in LSeg(|[b,d]|,|[b,c]|) or
    p2 in LSeg(|[b,c]|,|[a,c]|) by XBOOLE_0:def 2;
      now per cases by A6,XBOOLE_0:def 2;
    case p2 in LSeg(|[a,c]|,|[a,d]|);
      then LE p2,p1,K by A1,Th69;
     hence p2 in LSeg(|[b,d]|,|[b,c]|)& p1`2>=p2`2
      or p2 in LSeg(|[b,c]|,|[a,c]|) & p2<>W-min(K) by A1,A2,A4,JORDAN6:72;
    case p2 in LSeg(|[a,d]|,|[b,d]|);
      then LE p2,p1,K by A1,Th70;
     hence p2 in LSeg(|[b,d]|,|[b,c]|)& p1`2>=p2`2
      or p2 in LSeg(|[b,c]|,|[a,c]|) & p2<>W-min(K) by A1,A2,A4,JORDAN6:72;
    case p2 in LSeg(|[b,d]|,|[b,c]|);
     hence p2 in LSeg(|[b,d]|,|[b,c]|)& p1`2>=p2`2
      or p2 in LSeg(|[b,c]|,|[a,c]|) & p2<>W-min(K) by A1,A4,Th67;
    case A7: p2 in LSeg(|[b,c]|,|[a,c]|);
        now per cases;
      case p2=W-min(K); then LE p2,p1,K by A2,A5,JORDAN7:3;
       hence p2 in LSeg(|[b,d]|,|[b,c]|)& p1`2>=p2`2
        or p2 in LSeg(|[b,c]|,|[a,c]|) & p2<>W-min(K) by A1,A2,A4,JORDAN6:72;
      case p2<>W-min(K);
       hence p2 in LSeg(|[b,d]|,|[b,c]|)& p1`2>=p2`2
        or p2 in LSeg(|[b,c]|,|[a,c]|) & p2<>W-min(K) by A7;
      end;
     hence p2 in LSeg(|[b,d]|,|[b,c]|)& p1`2>=p2`2
      or p2 in LSeg(|[b,c]|,|[a,c]|) & p2<>W-min(K);
    end;
    hence p2 in LSeg(|[b,d]|,|[b,c]|)& p1`2>=p2`2
      or p2 in LSeg(|[b,c]|,|[a,c]|) & p2<>W-min(K);
  end;
  thus
    p2 in LSeg(|[b,d]|,|[b,c]|)& p1`2>=p2`2
  or p2 in LSeg(|[b,c]|,|[a,c]|) & p2<>W-min(K) implies LE p1,p2,K
  proof assume A8: p2 in LSeg(|[b,d]|,|[b,c]|)& p1`2>=p2`2
    or p2 in LSeg(|[b,c]|,|[a,c]|) & p2<>W-min(K);
      now per cases by A8;
    case A9: p2 in LSeg(|[b,d]|,|[b,c]|)& p1`2>=p2`2;
      then A10: p2`1=b & c <=p2`2 & p2`2 <= d by A1,Th9;
        W-min K=|[a,c]| by A1,Th56;
      then A11: p2 <> W-min(K) by A1,A10,EUCLID:56;
      A12: Lower_Arc(K)=LSeg(|[b,d]|,|[b,c]|) \/ LSeg(|[b,c]|,|[a,c]|)
                        by A1,Th62;
      then A13: p2 in Lower_Arc(K) by A9,XBOOLE_0:def 2;
      A14: p1 in Lower_Arc(K) by A1,A12,XBOOLE_0:def 2;
        LE p1,p2,Lower_Arc(K),E-max(K),W-min(K)
       proof
      for g being map of I[01], (TOP-REAL 2)|Lower_Arc(K),
        s1, s2 being Real st
        g is_homeomorphism
        & g.0 = E-max(K) & g.1 = W-min(K)
        & g.s1 = p1 & 0 <= s1 & s1 <= 1
        & g.s2 = p2 & 0 <= s2 & s2 <= 1
        holds s1 <= s2
     proof let g be map of I[01], (TOP-REAL 2)|Lower_Arc(K),
        s1, s2 be Real;
        assume A15: g is_homeomorphism
        & g.0 = E-max(K) & g.1 = W-min(K)
        & g.s1 = p1 & 0 <= s1 & s1 <= 1
        & g.s2 = p2 & 0 <= s2 & s2 <= 1;
        A16: dom g=the carrier of I[01] by FUNCT_2:def 1;
        A17: g is one-to-one by A15,TOPS_2:def 5;
        A18: the carrier of ((TOP-REAL 2)|Lower_Arc(K))
        =Lower_Arc(K) by JORDAN1:1;
        then g is Function of the carrier of I[01],
              the carrier of TOP-REAL 2 by FUNCT_2:9;
        then reconsider g1=g as map of I[01],TOP-REAL 2 ;
          g is continuous by A15,TOPS_2:def 5;
        then A19: g1 is continuous by TOPMETR:7;
        reconsider h1=proj1 as map of TOP-REAL 2,R^1 by JGRAPH_2:16;
        reconsider h2=proj2 as map of TOP-REAL 2,R^1 by JGRAPH_2:17;
        A20: (TOP-REAL 2)|([#](TOP-REAL 2))=TOP-REAL 2 by TSEP_1:3;
        then A21: (for p being Point of (TOP-REAL 2)|([#](TOP-REAL 2)) holds
        h1.p=proj1.p) implies h1 is continuous by JGRAPH_2:39;
          (for p being Point of (TOP-REAL 2)|([#](TOP-REAL 2)) holds
        h2.p=proj2.p) implies h2 is continuous by A20,JGRAPH_2:40;
        then consider h being map of TOP-REAL 2,R^1
        such that
        A22: (for p being Point of TOP-REAL 2,
          r1,r2 being real number st
        h1.p=r1 & h2.p=r2 holds h.p=r1+r2) & h is continuous
                          by A21,JGRAPH_2:29;
        reconsider k=h*g1 as map of I[01],R^1;
        A23: k is continuous map of I[01],R^1 by A19,A22,TOPS_2:58;
        A24: E-max K=|[b,d]| by A1,Th56;
          now assume A25: s1>s2;
          A26: dom g=[.0,1.] by BORSUK_1:83,FUNCT_2:def 1;
            0 in [.0,1.] by TOPREAL5:1;
          then A27: k.0=h.(E-max(K)) by A15,A26,FUNCT_1:23
          .=h1.(E-max(K))+h2.(E-max(K)) by A22
          .=(E-max(K))`1+proj2.(E-max(K)) by PSCOMP_1:def 28
          .=(E-max(K))`1+(E-max(K))`2 by PSCOMP_1:def 29
          .=(E-max(K))`1+d by A24,EUCLID:56
          .=b+d by A24,EUCLID:56;
            s1 in [.0,1.] by A15,TOPREAL5:1;
          then A28: k.s1=h.p1 by A15,A26,FUNCT_1:23
          .=proj1.p1 +proj2.p1 by A22 .=p1`1+proj2.p1 by PSCOMP_1:def 28
          .=b+p1`2 by A3,PSCOMP_1:def 29;
          A29: s2 in [.0,1.] by A15,TOPREAL5:1;
          then k.s2=h.p2 by A15,A26,FUNCT_1:23
          .=proj1.p2 +proj2.p2 by A22 .=p2`1+proj2.p2 by PSCOMP_1:def 28
          .=b+p2`2 by A10,PSCOMP_1:def 29;
          then A30: k.0>=k.s1 & k.s1>=k.s2 by A3,A9,A27,A28,REAL_1:55;
          A31: 0 in [.0,1.] by TOPREAL5:1;
          then A32: [.0,s2.] c= [.0,1.] by A29,RCOMP_1:16;
          then reconsider B=[.0,s2.] as Subset of I[01] by BORSUK_1:83
;
          A33: B is connected by A15,A29,A31,BORSUK_1:83,BORSUK_4:49;
          A34: 0 in B by A15,TOPREAL5:1;
          A35: s2 in B by A15,TOPREAL5:1;
          A36: k.0 is Real by XREAL_0:def 1;
          A37: k.s2 is Real by XREAL_0:def 1;
            k.s1 is Real by XREAL_0:def 1;
          then consider xc being Point of I[01] such that
          A38: xc in B & k.xc =k.s1
                   by A23,A30,A33,A34,A35,A36,A37,TOPREAL5:11;
            xc in [.0,1.] by BORSUK_1:83;
          then reconsider rxc=xc as Real;
          A39: k is one-to-one
          proof
              for x1,x2 being set st x1 in dom k & x2 in dom k &
            k.x1=k.x2 holds x1=x2
            proof let x1,x2 be set;
             assume A40: x1 in dom k & x2 in dom k &
               k.x1=k.x2;
              then reconsider r1=x1 as Point of I[01] by FUNCT_2:def 1;
              reconsider r2=x2 as Point of I[01] by A40,FUNCT_2:def 1;
              A41: k.x1=h.(g1.x1) by A40,FUNCT_1:22
                 .=h1.(g1.r1)+h2.(g1.r1) by A22
                 .=(g1.r1)`1+proj2.(g1.r1) by PSCOMP_1:def 28
                 .=(g1.r1)`1+(g1.r1)`2 by PSCOMP_1:def 29;
              A42: k.x2=h.(g1.x2) by A40,FUNCT_1:22
                 .=h1.(g1.r2)+h2.(g1.r2) by A22
                 .=(g1.r2)`1+proj2.(g1.r2) by PSCOMP_1:def 28
                 .=(g1.r2)`1+(g1.r2)`2 by PSCOMP_1:def 29;
              A43: g.r1 in Lower_Arc(K) by A18;
              A44: g.r2 in Lower_Arc(K) by A18;
              reconsider gr1=g.r1 as Point of TOP-REAL 2 by A43;
              reconsider gr2=g.r2 as Point of TOP-REAL 2 by A44;
                now per cases by A12,A18,XBOOLE_0:def 2;
              case A45: g.r1 in LSeg(|[b,d]|,|[b,c]|) &
                  g.r2 in LSeg(|[b,d]|,|[b,c]|);
                then A46: (gr1)`1=b & c <=(gr1)`2 & (gr1)`2 <=d by A1,Th9;
                A47: (gr2)`1=b & c <=(gr2)`2 & (gr2)`2 <=d
                                  by A1,A45,Th9;
                 then (gr1)`2=(gr2)`2 by A40,A41,A42,A46,XCMPLX_1:2;
                then |[(gr1)`1,(gr1)`2]|=g.r2 by A46,A47,EUCLID:57;
                then g.r1=g.r2 by EUCLID:57;
               hence x1=x2 by A16,A17,FUNCT_1:def 8;
              case A48: g.r1 in LSeg(|[b,d]|,|[b,c]|) &
                  g.r2 in LSeg(|[b,c]|,|[a,c]|);
                then A49: (gr1)`1=b & c <=(gr1)`2 & (gr1)`2 <=d by A1,Th9;
                A50: (gr2)`2=c & a <=(gr2)`1 & (gr2)`1 <=b
                                  by A1,A48,Th11;
                then A51: b+(gr1)`2=(gr2)`1 +c by A1,A40,A41,A42,A48,Th9;
                A52: now assume b<>gr2`1; then b>gr2`1 by A50,REAL_1:def 5;
                 hence contradiction by A40,A41,A42,A49,A50,REAL_1:67;
                end;
                  now assume gr1`2<> c; then c <gr1`2 by A49,REAL_1:def 5;
                 hence contradiction by A50,A51,REAL_1:67;
                end;
                then |[(gr1)`1,(gr1)`2]|=g.r2 by A49,A50,A52,EUCLID:57;
                then g.r1=g.r2 by EUCLID:57;
               hence x1=x2 by A16,A17,FUNCT_1:def 8;
              case A53: g.r1 in LSeg(|[b,c]|,|[a,c]|) &
                  g.r2 in LSeg(|[b,d]|,|[b,c]|);
                then A54: (gr2)`1=b & c <=(gr2)`2 & (gr2)`2 <=d by A1,Th9;
                A55: (gr1)`2=c & a <=(gr1)`1 & (gr1)`1 <=b
                                  by A1,A53,Th11;
                A56: b+(gr2)`2=(gr1)`1 +c by A1,A40,A41,A42,A53,A54,Th11;
                A57: now assume b<>gr1`1; then b>gr1`1 by A55,REAL_1:def 5;
                 hence contradiction by A54,A56,REAL_1:67;
                end;
                  now assume gr2`2<> c; then c < gr2`2 by A54,REAL_1:def 5;
                 hence contradiction by A40,A41,A42,A54,A55,REAL_1:67;
                end;
                then |[(gr2)`1,(gr2)`2]|=g.r1 by A54,A55,A57,EUCLID:57;
                then g.r1=g.r2 by EUCLID:57;
               hence x1=x2 by A16,A17,FUNCT_1:def 8;
              case A58: g.r1 in LSeg(|[b,c]|,|[a,c]|) &
                  g.r2 in LSeg(|[b,c]|,|[a,c]|);
                then A59: (gr1)`2=c & a <=(gr1)`1 & (gr1)`1 <=b by A1,Th11;
                A60: (gr2)`2=c & a <=(gr2)`1 & (gr2)`1 <=b
                                  by A1,A58,Th11;
                 then (gr1)`1=(gr2)`1 by A40,A41,A42,A59,XCMPLX_1:2;
                then |[(gr1)`1,(gr1)`2]|=g.r2 by A59,A60,EUCLID:57;
                then g.r1=g.r2 by EUCLID:57;
               hence x1=x2 by A16,A17,FUNCT_1:def 8;
              end;
             hence x1=x2;
            end;
           hence thesis by FUNCT_1:def 8;
          end;
          A61: dom k=[.0,1.] by BORSUK_1:83,FUNCT_2:def 1;
          then s1 in dom k by A15,TOPREAL5:1;
          then rxc=s1 by A32,A38,A39,A61,FUNCT_1:def 8;
         hence contradiction by A25,A38,TOPREAL5:1;
        end;
      hence s1 <= s2;
     end;
     hence thesis by A13,A14,JORDAN5C:def 3;
       end;
     hence LE p1,p2,K by A11,A13,A14,JORDAN6:def 10;
    case A62: p2 in LSeg(|[b,c]|,|[a,c]|) & p2<>W-min(K);
      then A63: p2`2=c & a <=p2`1 & p2`1 <= b by A1,Th11;
      A64: Lower_Arc(K)=LSeg(|[b,d]|,|[b,c]|) \/ LSeg(|[b,c]|,|[a,c]|)
                        by A1,Th62;
      then A65: p2 in Lower_Arc(K) by A62,XBOOLE_0:def 2;
      A66: p1 in Lower_Arc(K) by A1,A64,XBOOLE_0:def 2;
        LE p1,p2,Lower_Arc(K),E-max(K),W-min(K)
       proof
      for g being map of I[01], (TOP-REAL 2)|Lower_Arc(K),
        s1, s2 being Real st
        g is_homeomorphism
        & g.0 = E-max(K) & g.1 = W-min(K)
        & g.s1 = p1 & 0 <= s1 & s1 <= 1
        & g.s2 = p2 & 0 <= s2 & s2 <= 1
        holds s1 <= s2
     proof let g be map of I[01], (TOP-REAL 2)|Lower_Arc(K),
        s1, s2 be Real;
        assume A67: g is_homeomorphism
        & g.0 = E-max(K) & g.1 = W-min(K)
        & g.s1 = p1 & 0 <= s1 & s1 <= 1
        & g.s2 = p2 & 0 <= s2 & s2 <= 1;
        A68: dom g=the carrier of I[01] by FUNCT_2:def 1;
        A69: g is one-to-one by A67,TOPS_2:def 5;
        A70: the carrier of ((TOP-REAL 2)|Lower_Arc(K))
        =Lower_Arc(K) by JORDAN1:1;
        then g is Function of the carrier of I[01],
              the carrier of TOP-REAL 2 by FUNCT_2:9;
        then reconsider g1=g as map of I[01],TOP-REAL 2 ;
          g is continuous by A67,TOPS_2:def 5;
        then A71: g1 is continuous by TOPMETR:7;
        reconsider h1=proj1 as map of TOP-REAL 2,R^1 by JGRAPH_2:16;
        reconsider h2=proj2 as map of TOP-REAL 2,R^1 by JGRAPH_2:17;
        A72: (TOP-REAL 2)|([#](TOP-REAL 2))=TOP-REAL 2 by TSEP_1:3;
        then A73: (for p being Point of (TOP-REAL 2)|([#](TOP-REAL 2)) holds
        h1.p=proj1.p) implies h1 is continuous by JGRAPH_2:39;
          (for p being Point of (TOP-REAL 2)|([#](TOP-REAL 2)) holds
        h2.p=proj2.p) implies h2 is continuous by A72,JGRAPH_2:40;
        then consider h being map of TOP-REAL 2,R^1
        such that
        A74: (for p being Point of TOP-REAL 2,
          r1,r2 being real number st
        h1.p=r1 & h2.p=r2 holds h.p=r1+r2) & h is continuous
                          by A73,JGRAPH_2:29;
        reconsider k=h*g1 as map of I[01],R^1;
        A75: k is continuous map of I[01],R^1 by A71,A74,TOPS_2:58;
        A76: E-max K=|[b,d]| by A1,Th56;
          now assume A77: s1>s2;
          A78: dom g=[.0,1.] by BORSUK_1:83,FUNCT_2:def 1;
            0 in [.0,1.] by TOPREAL5:1;
          then A79: k.0=h.(E-max(K)) by A67,A78,FUNCT_1:23
          .=h1.(E-max(K))+h2.(E-max(K)) by A74
          .=(E-max(K))`1+proj2.(E-max(K)) by PSCOMP_1:def 28
          .=(E-max(K))`1+(E-max(K))`2 by PSCOMP_1:def 29
          .=(E-max(K))`1+d by A76,EUCLID:56
          .=b+d by A76,EUCLID:56;
            s1 in [.0,1.] by A67,TOPREAL5:1;
          then A80: k.s1=h.p1 by A67,A78,FUNCT_1:23
          .=proj1.p1 +proj2.p1 by A74 .=p1`1+proj2.p1 by PSCOMP_1:def 28
          .=b+p1`2 by A3,PSCOMP_1:def 29;
          A81: s2 in [.0,1.] by A67,TOPREAL5:1;
          then k.s2=h.p2 by A67,A78,FUNCT_1:23
          .=proj1.p2 +proj2.p2 by A74 .=p2`1+proj2.p2 by PSCOMP_1:def 28
          .=p2`1+c by A63,PSCOMP_1:def 29;
          then A82: k.0>=k.s1 & k.s1>=k.s2 by A3,A63,A79,A80,REAL_1:55;
          A83: 0 in [.0,1.] by TOPREAL5:1;
          then A84: [.0,s2.] c= [.0,1.] by A81,RCOMP_1:16;
          then reconsider B=[.0,s2.] as Subset of I[01] by BORSUK_1:83
;
          A85: B is connected by A67,A81,A83,BORSUK_1:83,BORSUK_4:49;
          A86: 0 in B by A67,TOPREAL5:1;
          A87: s2 in B by A67,TOPREAL5:1;
          A88: k.0 is Real by XREAL_0:def 1;
          A89: k.s2 is Real by XREAL_0:def 1;
            k.s1 is Real by XREAL_0:def 1;
          then consider xc being Point of I[01] such that
          A90: xc in B & k.xc =k.s1
                   by A75,A82,A85,A86,A87,A88,A89,TOPREAL5:11;
            xc in [.0,1.] by BORSUK_1:83;
          then reconsider rxc=xc as Real;
          A91: k is one-to-one
          proof
              for x1,x2 being set st x1 in dom k & x2 in dom k &
            k.x1=k.x2 holds x1=x2
            proof let x1,x2 be set;
             assume A92: x1 in dom k & x2 in dom k &
               k.x1=k.x2;
              then reconsider r1=x1 as Point of I[01] by FUNCT_2:def 1;
              reconsider r2=x2 as Point of I[01] by A92,FUNCT_2:def 1;
              A93: k.x1=h.(g1.x1) by A92,FUNCT_1:22
                 .=h1.(g1.r1)+h2.(g1.r1) by A74
                 .=(g1.r1)`1+proj2.(g1.r1) by PSCOMP_1:def 28
                 .=(g1.r1)`1+(g1.r1)`2 by PSCOMP_1:def 29;
              A94: k.x2=h.(g1.x2) by A92,FUNCT_1:22
                 .=h1.(g1.r2)+h2.(g1.r2) by A74
                 .=(g1.r2)`1+proj2.(g1.r2) by PSCOMP_1:def 28
                 .=(g1.r2)`1+(g1.r2)`2 by PSCOMP_1:def 29;
              A95: g.r1 in Lower_Arc(K) by A70;
              A96: g.r2 in Lower_Arc(K) by A70;
              reconsider gr1=g.r1 as Point of TOP-REAL 2 by A95;
              reconsider gr2=g.r2 as Point of TOP-REAL 2 by A96;
                now per cases by A64,A70,XBOOLE_0:def 2;
              case A97: g.r1 in LSeg(|[b,d]|,|[b,c]|) &
                  g.r2 in LSeg(|[b,d]|,|[b,c]|);
                then A98: (gr1)`1=b & c <=(gr1)`2 & (gr1)`2 <=d by A1,Th9;
                A99: (gr2)`1=b & c <=(gr2)`2 & (gr2)`2 <=d
                                  by A1,A97,Th9;
                 then (gr1)`2=(gr2)`2 by A92,A93,A94,A98,XCMPLX_1:2;
                then |[(gr1)`1,(gr1)`2]|=g.r2 by A98,A99,EUCLID:57;
                then g.r1=g.r2 by EUCLID:57;
               hence x1=x2 by A68,A69,FUNCT_1:def 8;
              case A100: g.r1 in LSeg(|[b,d]|,|[b,c]|) &
                  g.r2 in LSeg(|[b,c]|,|[a,c]|);
                then A101: (gr1)`1=b & c <=(gr1)`2 & (gr1)`2 <=d by A1,Th9;
                A102: (gr2)`2=c & a <=(gr2)`1 & (gr2)`1 <=b
                                  by A1,A100,Th11;
                then A103: b+(gr1)`2=(gr2)`1 +c by A1,A92,A93,A94,A100,Th9;
                A104: now assume b<>gr2`1; then b>gr2`1 by A102,REAL_1:def 5;
                 hence contradiction by A92,A93,A94,A101,A102,REAL_1:67;
                end;
                  now assume gr1`2<> c; then c <gr1`2 by A101,REAL_1:def 5;
                 hence contradiction by A102,A103,REAL_1:67;
                end;
                then |[(gr1)`1,(gr1)`2]|=g.r2 by A101,A102,A104,EUCLID:57;
                then g.r1=g.r2 by EUCLID:57;
               hence x1=x2 by A68,A69,FUNCT_1:def 8;
              case A105: g.r1 in LSeg(|[b,c]|,|[a,c]|) &
                  g.r2 in LSeg(|[b,d]|,|[b,c]|);
                then A106: (gr2)`1=b & c <=(gr2)`2 & (gr2)`2 <=d by A1,Th9;
                A107: (gr1)`2=c & a <=(gr1)`1 & (gr1)`1 <=b
                                  by A1,A105,Th11;
                A108: b+(gr2)`2=(gr1)`1 +c by A1,A92,A93,A94,A105,A106,Th11;
                A109: now assume b<>gr1`1; then b>gr1`1 by A107,REAL_1:def 5;
                 hence contradiction by A106,A108,REAL_1:67;
                end;
                  now assume gr2`2<> c; then c < gr2`2 by A106,REAL_1:def 5;
                 hence contradiction by A92,A93,A94,A106,A107,REAL_1:67;
                end;
                then |[(gr2)`1,(gr2)`2]|=g.r1 by A106,A107,A109,EUCLID:57;
                then g.r1=g.r2 by EUCLID:57;
               hence x1=x2 by A68,A69,FUNCT_1:def 8;
              case A110: g.r1 in LSeg(|[b,c]|,|[a,c]|) &
                  g.r2 in LSeg(|[b,c]|,|[a,c]|);
                then A111: (gr1)`2=c & a <=(gr1)`1 & (gr1)`1 <=b by A1,Th11;
                A112: (gr2)`2=c & a <=(gr2)`1 & (gr2)`1 <=b
                                  by A1,A110,Th11;
                 then (gr1)`1=(gr2)`1 by A92,A93,A94,A111,XCMPLX_1:2;
                then |[(gr1)`1,(gr1)`2]|=g.r2 by A111,A112,EUCLID:57;
                then g.r1=g.r2 by EUCLID:57;
               hence x1=x2 by A68,A69,FUNCT_1:def 8;
              end;
             hence x1=x2;
            end;
           hence thesis by FUNCT_1:def 8;
          end;
          A113: dom k=[.0,1.] by BORSUK_1:83,FUNCT_2:def 1;
          then s1 in dom k by A67,TOPREAL5:1;
          then rxc=s1 by A84,A90,A91,A113,FUNCT_1:def 8;
         hence contradiction by A77,A90,TOPREAL5:1;
        end;
      hence s1 <= s2;
     end;
     hence thesis by A65,A66,JORDAN5C:def 3;
       end;
     hence LE p1,p2,K by A62,A65,A66,JORDAN6:def 10;
    end;
   hence LE p1,p2,K;
  end;
end;

theorem Th72: for K being non empty compact Subset of TOP-REAL 2,
a,b,c,d being real number,p1,p2 being Point of TOP-REAL 2
 st K=rectangle(a,b,c,d) & a<b & c < d
 & p1 in LSeg(|[b,c]|,|[a,c]|)& p1<>W-min(K)
holds LE p1,p2,K iff p2 in LSeg(|[b,c]|,|[a,c]|) & p1`1>=p2`1 & p2<>W-min(K)
proof let K be non empty compact Subset of TOP-REAL 2,
a,b,c,d be real number,p1,p2 be Point of TOP-REAL 2;
 assume A1: K=rectangle(a,b,c,d) & a<b & c <d
    & p1 in LSeg(|[b,c]|,|[a,c]|) & p1<>W-min(K);
  then A2: K is_simple_closed_curve by Th60;
  A3: p1`2=c & a <=p1`1 & p1`1 <= b by A1,Th11;
  thus LE p1,p2,K implies
  p2 in LSeg(|[b,c]|,|[a,c]|) & p1`1>=p2`1 & p2<>W-min(K)
  proof assume A4: LE p1,p2,K;
    then A5: p1 in K & p2 in K by A2,JORDAN7:5;
      K= {p: p`1=a & c <=p`2 & p`2<=d or p`2=d & a<=p`1 & p`1<=b or
      p`1=b & c <=p`2 & p`2<=d or p`2=c & a<=p`1 & p`1<=b} by A1,Def1;
    then K= LSeg(|[ a,c ]|, |[ a,d ]|) \/ LSeg(|[ a,d ]|, |[ b,d ]|)
    \/ (LSeg(|[ a,c ]|, |[ b,c ]|) \/ LSeg(|[ b,c ]|, |[ b,d ]|)) by A1,Th40
    .=LSeg(|[a,c]|,|[a,d]|) \/ LSeg(|[a,d]|,|[b,d]|)
    \/ LSeg(|[b,d]|,|[b,c]|) \/ LSeg(|[b,c]|,|[a,c]|) by XBOOLE_1:4;
    then p2 in LSeg(|[a,c]|,|[a,d]|) \/ LSeg(|[a,d]|,|[b,d]|)
    \/ LSeg(|[b,d]|,|[b,c]|) or
    p2 in LSeg(|[b,c]|,|[a,c]|) by A5,XBOOLE_0:def 2;
    then A6: p2 in LSeg(|[a,c]|,|[a,d]|) \/ LSeg(|[a,d]|,|[b,d]|) or
    p2 in LSeg(|[b,d]|,|[b,c]|) or
    p2 in LSeg(|[b,c]|,|[a,c]|) by XBOOLE_0:def 2;
      now per cases by A6,XBOOLE_0:def 2;
    case p2 in LSeg(|[a,c]|,|[a,d]|);
      then LE p2,p1,K by A1,Th69;
     hence p2 in LSeg(|[b,c]|,|[a,c]|)& p1`1>=p2`1
      & p2<>W-min(K) by A1,A2,A4,JORDAN6:72;
    case p2 in LSeg(|[a,d]|,|[b,d]|);
      then LE p2,p1,K by A1,Th70;
     hence p2 in LSeg(|[b,c]|,|[a,c]|)& p1`1>=p2`1
      & p2<>W-min(K) by A1,A2,A4,JORDAN6:72;
    case p2 in LSeg(|[b,d]|,|[b,c]|);
      then LE p2,p1,K by A1,Th71;
     hence p2 in LSeg(|[b,c]|,|[a,c]|)& p1`1>=p2`1
     & p2<>W-min(K) by A1,A2,A4,JORDAN6:72;
    case p2 in LSeg(|[b,c]|,|[a,c]|);
     hence p2 in LSeg(|[b,c]|,|[a,c]|)& p1`1>=p2`1
      & p2<>W-min(K) by A1,A4,Th68;
    end;
    hence p2 in LSeg(|[b,c]|,|[a,c]|)& p1`1>=p2`1
    & p2<>W-min(K);
  end;
  thus
    p2 in LSeg(|[b,c]|,|[a,c]|)& p1`1>=p2`1
  & p2<>W-min(K) implies LE p1,p2,K
  proof assume A7: p2 in LSeg(|[b,c]|,|[a,c]|)& p1`1>=p2`1
    & p2<>W-min(K);
      now per cases by A7;
    case A8: p2 in LSeg(|[b,c]|,|[a,c]|)& p1`1>=p2`1;
      then A9: p2`2=c & a <=p2`1 & p2`1 <= b by A1,Th11;
      A10: Lower_Arc(K)=LSeg(|[b,d]|,|[b,c]|) \/ LSeg(|[b,c]|,|[a,c]|)
                        by A1,Th62;
      then A11: p2 in Lower_Arc(K) by A8,XBOOLE_0:def 2;
      A12: p1 in Lower_Arc(K) by A1,A10,XBOOLE_0:def 2;
        LE p1,p2,Lower_Arc(K),E-max(K),W-min(K)
       proof
      for g being map of I[01], (TOP-REAL 2)|Lower_Arc(K),
        s1, s2 being Real st
        g is_homeomorphism
        & g.0 = E-max(K) & g.1 = W-min(K)
        & g.s1 = p1 & 0 <= s1 & s1 <= 1
        & g.s2 = p2 & 0 <= s2 & s2 <= 1
        holds s1 <= s2
     proof let g be map of I[01], (TOP-REAL 2)|Lower_Arc(K),
        s1, s2 be Real;
        assume A13: g is_homeomorphism
        & g.0 = E-max(K) & g.1 = W-min(K)
        & g.s1 = p1 & 0 <= s1 & s1 <= 1
        & g.s2 = p2 & 0 <= s2 & s2 <= 1;
        A14: dom g=the carrier of I[01] by FUNCT_2:def 1;
        A15: g is one-to-one by A13,TOPS_2:def 5;
        A16: the carrier of ((TOP-REAL 2)|Lower_Arc(K))
        =Lower_Arc(K) by JORDAN1:1;
        then g is Function of the carrier of I[01],
              the carrier of TOP-REAL 2 by FUNCT_2:9;
        then reconsider g1=g as map of I[01],TOP-REAL 2 ;
          g is continuous by A13,TOPS_2:def 5;
        then A17: g1 is continuous by TOPMETR:7;
        reconsider h1=proj1 as map of TOP-REAL 2,R^1 by JGRAPH_2:16;
        reconsider h2=proj2 as map of TOP-REAL 2,R^1 by JGRAPH_2:17;
        A18: (TOP-REAL 2)|([#](TOP-REAL 2))=TOP-REAL 2 by TSEP_1:3;
        then A19: (for p being Point of (TOP-REAL 2)|([#](TOP-REAL 2)) holds
        h1.p=proj1.p) implies h1 is continuous by JGRAPH_2:39;
          (for p being Point of (TOP-REAL 2)|([#](TOP-REAL 2)) holds
        h2.p=proj2.p) implies h2 is continuous by A18,JGRAPH_2:40;
        then consider h being map of TOP-REAL 2,R^1
        such that
        A20: (for p being Point of TOP-REAL 2,
          r1,r2 being real number st
        h1.p=r1 & h2.p=r2 holds h.p=r1+r2) & h is continuous
                          by A19,JGRAPH_2:29;
        reconsider k=h*g1 as map of I[01],R^1;
        A21: k is continuous map of I[01],R^1 by A17,A20,TOPS_2:58;
        A22: E-max K=|[b,d]| by A1,Th56;
          now assume A23: s1>s2;
          A24: dom g=[.0,1.] by BORSUK_1:83,FUNCT_2:def 1;
            0 in [.0,1.] by TOPREAL5:1;
          then A25: k.0=h.(E-max(K)) by A13,A24,FUNCT_1:23
          .=h1.(E-max(K))+h2.(E-max(K)) by A20
          .=(E-max(K))`1+proj2.(E-max(K)) by PSCOMP_1:def 28
          .=(E-max(K))`1+(E-max(K))`2 by PSCOMP_1:def 29
          .=(E-max(K))`1+d by A22,EUCLID:56
          .=b+d by A22,EUCLID:56;
            s1 in [.0,1.] by A13,TOPREAL5:1;
          then A26: k.s1=h.p1 by A13,A24,FUNCT_1:23
          .=proj1.p1 +proj2.p1 by A20 .=p1`1+proj2.p1 by PSCOMP_1:def 28
          .=p1`1+c by A3,PSCOMP_1:def 29;
          A27: s2 in [.0,1.] by A13,TOPREAL5:1;
          then k.s2=h.p2 by A13,A24,FUNCT_1:23
          .=proj1.p2 +proj2.p2 by A20 .=p2`1+proj2.p2 by PSCOMP_1:def 28
          .=p2`1+c by A9,PSCOMP_1:def 29;
          then A28: k.0>=k.s1 & k.s1>=k.s2 by A1,A3,A8,A25,A26,REAL_1:55;
          A29: 0 in [.0,1.] by TOPREAL5:1;
          then A30: [.0,s2.] c= [.0,1.] by A27,RCOMP_1:16;
          then reconsider B=[.0,s2.] as Subset of I[01] by BORSUK_1:83
;
          A31: B is connected by A13,A27,A29,BORSUK_1:83,BORSUK_4:49;
          A32: 0 in B by A13,TOPREAL5:1;
          A33: s2 in B by A13,TOPREAL5:1;
          A34: k.0 is Real by XREAL_0:def 1;
          A35: k.s2 is Real by XREAL_0:def 1;
            k.s1 is Real by XREAL_0:def 1;
          then consider xc being Point of I[01] such that
          A36: xc in B & k.xc =k.s1
                   by A21,A28,A31,A32,A33,A34,A35,TOPREAL5:11;
            xc in [.0,1.] by BORSUK_1:83;
          then reconsider rxc=xc as Real;
          A37: k is one-to-one
          proof
              for x1,x2 being set st x1 in dom k & x2 in dom k &
            k.x1=k.x2 holds x1=x2
            proof let x1,x2 be set;
             assume A38: x1 in dom k & x2 in dom k &
               k.x1=k.x2;
              then reconsider r1=x1 as Point of I[01] by FUNCT_2:def 1;
              reconsider r2=x2 as Point of I[01] by A38,FUNCT_2:def 1;
              A39: k.x1=h.(g1.x1) by A38,FUNCT_1:22
                 .=h1.(g1.r1)+h2.(g1.r1) by A20
                 .=(g1.r1)`1+proj2.(g1.r1) by PSCOMP_1:def 28
                 .=(g1.r1)`1+(g1.r1)`2 by PSCOMP_1:def 29;
              A40: k.x2=h.(g1.x2) by A38,FUNCT_1:22
                 .=h1.(g1.r2)+h2.(g1.r2) by A20
                 .=(g1.r2)`1+proj2.(g1.r2) by PSCOMP_1:def 28
                 .=(g1.r2)`1+(g1.r2)`2 by PSCOMP_1:def 29;
              A41: g.r1 in Lower_Arc(K) by A16;
              A42: g.r2 in Lower_Arc(K) by A16;
              reconsider gr1=g.r1 as Point of TOP-REAL 2 by A41;
              reconsider gr2=g.r2 as Point of TOP-REAL 2 by A42;
                now per cases by A10,A16,XBOOLE_0:def 2;
              case A43: g.r1 in LSeg(|[b,d]|,|[b,c]|) &
                  g.r2 in LSeg(|[b,d]|,|[b,c]|);
                then A44: (gr1)`1=b & c <=(gr1)`2 & (gr1)`2 <=d by A1,Th9;
                A45: (gr2)`1=b & c <=(gr2)`2 & (gr2)`2 <=d
                                  by A1,A43,Th9;
                 then (gr1)`2=(gr2)`2 by A38,A39,A40,A44,XCMPLX_1:2;
                then |[(gr1)`1,(gr1)`2]|=g.r2 by A44,A45,EUCLID:57;
                then g.r1=g.r2 by EUCLID:57;
               hence x1=x2 by A14,A15,FUNCT_1:def 8;
              case A46: g.r1 in LSeg(|[b,d]|,|[b,c]|) &
                  g.r2 in LSeg(|[b,c]|,|[a,c]|);
                then A47: (gr1)`1=b & c <=(gr1)`2 & (gr1)`2 <=d by A1,Th9;
                A48: (gr2)`2=c & a <=(gr2)`1 & (gr2)`1 <=b
                                  by A1,A46,Th11;
                then A49: b+(gr1)`2=(gr2)`1 +c by A1,A38,A39,A40,A46,Th9;
                A50: now assume b<>gr2`1; then b>gr2`1 by A48,REAL_1:def 5;
                 hence contradiction by A38,A39,A40,A47,A48,REAL_1:67;
                end;
                  now assume gr1`2<> c; then c <gr1`2 by A47,REAL_1:def 5;
                 hence contradiction by A48,A49,REAL_1:67;
                end;
                then |[(gr1)`1,(gr1)`2]|=g.r2 by A47,A48,A50,EUCLID:57;
                then g.r1=g.r2 by EUCLID:57;
               hence x1=x2 by A14,A15,FUNCT_1:def 8;
              case A51: g.r1 in LSeg(|[b,c]|,|[a,c]|) &
                  g.r2 in LSeg(|[b,d]|,|[b,c]|);
                then A52: (gr2)`1=b & c <=(gr2)`2 & (gr2)`2 <=d by A1,Th9;
                A53: (gr1)`2=c & a <=(gr1)`1 & (gr1)`1 <=b
                                  by A1,A51,Th11;
                A54: b+(gr2)`2=(gr1)`1 +c by A1,A38,A39,A40,A51,A52,Th11;
                A55: now assume b<>gr1`1; then b>gr1`1 by A53,REAL_1:def 5;
                 hence contradiction by A52,A54,REAL_1:67;
                end;
                  now assume gr2`2<> c; then c < gr2`2 by A52,REAL_1:def 5;
                 hence contradiction by A38,A39,A40,A52,A53,REAL_1:67;
                end;
                then |[(gr2)`1,(gr2)`2]|=g.r1 by A52,A53,A55,EUCLID:57;
                then g.r1=g.r2 by EUCLID:57;
               hence x1=x2 by A14,A15,FUNCT_1:def 8;
              case A56: g.r1 in LSeg(|[b,c]|,|[a,c]|) &
                  g.r2 in LSeg(|[b,c]|,|[a,c]|);
                then A57: (gr1)`2=c & a <=(gr1)`1 & (gr1)`1 <=b by A1,Th11;
                A58: (gr2)`2=c & a <=(gr2)`1 & (gr2)`1 <=b
                                  by A1,A56,Th11;
                 then (gr1)`1=(gr2)`1 by A38,A39,A40,A57,XCMPLX_1:2;
                then |[(gr1)`1,(gr1)`2]|=g.r2 by A57,A58,EUCLID:57;
                then g.r1=g.r2 by EUCLID:57;
               hence x1=x2 by A14,A15,FUNCT_1:def 8;
              end;
             hence x1=x2;
            end;
           hence thesis by FUNCT_1:def 8;
          end;
          A59: dom k=[.0,1.] by BORSUK_1:83,FUNCT_2:def 1;
          then s1 in dom k by A13,TOPREAL5:1;
          then rxc=s1 by A30,A36,A37,A59,FUNCT_1:def 8;
         hence contradiction by A23,A36,TOPREAL5:1;
        end;
      hence s1 <= s2;
     end;
     hence thesis by A11,A12,JORDAN5C:def 3;
       end;
     hence LE p1,p2,K by A7,A11,A12,JORDAN6:def 10;
    case A60: p2 in LSeg(|[b,c]|,|[a,c]|) & p2<>W-min(K);
      then A61: p2`2=c & a <=p2`1 & p2`1 <= b by A1,Th11;
      A62: Lower_Arc(K)=LSeg(|[b,d]|,|[b,c]|) \/ LSeg(|[b,c]|,|[a,c]|)
                        by A1,Th62;
      then A63: p2 in Lower_Arc(K) by A60,XBOOLE_0:def 2;
      A64: p1 in Lower_Arc(K) by A1,A62,XBOOLE_0:def 2;
        LE p1,p2,Lower_Arc(K),E-max(K),W-min(K)
       proof
      for g being map of I[01], (TOP-REAL 2)|Lower_Arc(K),
        s1, s2 being Real st
        g is_homeomorphism
        & g.0 = E-max(K) & g.1 = W-min(K)
        & g.s1 = p1 & 0 <= s1 & s1 <= 1
        & g.s2 = p2 & 0 <= s2 & s2 <= 1
        holds s1 <= s2
     proof let g be map of I[01], (TOP-REAL 2)|Lower_Arc(K),
        s1, s2 be Real;
        assume A65: g is_homeomorphism
        & g.0 = E-max(K) & g.1 = W-min(K)
        & g.s1 = p1 & 0 <= s1 & s1 <= 1
        & g.s2 = p2 & 0 <= s2 & s2 <= 1;
        A66: dom g=the carrier of I[01] by FUNCT_2:def 1;
        A67: g is one-to-one by A65,TOPS_2:def 5;
        A68: the carrier of ((TOP-REAL 2)|Lower_Arc(K))
        =Lower_Arc(K) by JORDAN1:1;
        then g is Function of the carrier of I[01],
              the carrier of TOP-REAL 2 by FUNCT_2:9;
        then reconsider g1=g as map of I[01],TOP-REAL 2 ;
          g is continuous by A65,TOPS_2:def 5;
        then A69: g1 is continuous by TOPMETR:7;
        reconsider h1=proj1 as map of TOP-REAL 2,R^1 by JGRAPH_2:16;
        reconsider h2=proj2 as map of TOP-REAL 2,R^1 by JGRAPH_2:17;
        A70: (TOP-REAL 2)|([#](TOP-REAL 2))=TOP-REAL 2 by TSEP_1:3;
        then A71: (for p being Point of (TOP-REAL 2)|([#](TOP-REAL 2)) holds
        h1.p=proj1.p) implies h1 is continuous by JGRAPH_2:39;
          (for p being Point of (TOP-REAL 2)|([#](TOP-REAL 2)) holds
        h2.p=proj2.p) implies h2 is continuous by A70,JGRAPH_2:40;
        then consider h being map of TOP-REAL 2,R^1
        such that
        A72: (for p being Point of TOP-REAL 2,
          r1,r2 being real number st
        h1.p=r1 & h2.p=r2 holds h.p=r1+r2) & h is continuous
                          by A71,JGRAPH_2:29;
        reconsider k=h*g1 as map of I[01],R^1;
        A73: k is continuous map of I[01],R^1 by A69,A72,TOPS_2:58;
        A74: E-max K=|[b,d]| by A1,Th56;
          now assume A75: s1>s2;
          A76: dom g=[.0,1.] by BORSUK_1:83,FUNCT_2:def 1;
            0 in [.0,1.] by TOPREAL5:1;
          then A77: k.0=h.(E-max(K)) by A65,A76,FUNCT_1:23
          .=h1.(E-max(K))+h2.(E-max(K)) by A72
          .=(E-max(K))`1+proj2.(E-max(K)) by PSCOMP_1:def 28
          .=(E-max(K))`1+(E-max(K))`2 by PSCOMP_1:def 29
          .=(E-max(K))`1+d by A74,EUCLID:56
          .=b+d by A74,EUCLID:56;
            s1 in [.0,1.] by A65,TOPREAL5:1;
          then A78: k.s1=h.p1 by A65,A76,FUNCT_1:23
          .=proj1.p1 +proj2.p1 by A72 .=p1`1+proj2.p1 by PSCOMP_1:def 28
          .=p1`1+c by A3,PSCOMP_1:def 29;
          A79: s2 in [.0,1.] by A65,TOPREAL5:1;
          then k.s2=h.p2 by A65,A76,FUNCT_1:23
          .=proj1.p2 +proj2.p2 by A72 .=p2`1+proj2.p2 by PSCOMP_1:def 28
          .=p2`1+c by A61,PSCOMP_1:def 29;
          then A80: k.0>=k.s1 & k.s1>=k.s2 by A1,A3,A7,A77,A78,REAL_1:55;
          A81: 0 in [.0,1.] by TOPREAL5:1;
          then A82: [.0,s2.] c= [.0,1.] by A79,RCOMP_1:16;
          then reconsider B=[.0,s2.] as Subset of I[01] by BORSUK_1:83
;
          A83: B is connected by A65,A79,A81,BORSUK_1:83,BORSUK_4:49;
          A84: 0 in B by A65,TOPREAL5:1;
          A85: s2 in B by A65,TOPREAL5:1;
          A86: k.0 is Real by XREAL_0:def 1;
          A87: k.s2 is Real by XREAL_0:def 1;
            k.s1 is Real by XREAL_0:def 1;
          then consider xc being Point of I[01] such that
          A88: xc in B & k.xc =k.s1
                   by A73,A80,A83,A84,A85,A86,A87,TOPREAL5:11;
            xc in [.0,1.] by BORSUK_1:83;
          then reconsider rxc=xc as Real;
          A89: k is one-to-one
          proof
              for x1,x2 being set st x1 in dom k & x2 in dom k &
            k.x1=k.x2 holds x1=x2
            proof let x1,x2 be set;
             assume A90: x1 in dom k & x2 in dom k &
               k.x1=k.x2;
              then reconsider r1=x1 as Point of I[01] by FUNCT_2:def 1;
              reconsider r2=x2 as Point of I[01] by A90,FUNCT_2:def 1;
              A91: k.x1=h.(g1.x1) by A90,FUNCT_1:22
                 .=h1.(g1.r1)+h2.(g1.r1) by A72
                 .=(g1.r1)`1+proj2.(g1.r1) by PSCOMP_1:def 28
                 .=(g1.r1)`1+(g1.r1)`2 by PSCOMP_1:def 29;
              A92: k.x2=h.(g1.x2) by A90,FUNCT_1:22
                 .=h1.(g1.r2)+h2.(g1.r2) by A72
                 .=(g1.r2)`1+proj2.(g1.r2) by PSCOMP_1:def 28
                 .=(g1.r2)`1+(g1.r2)`2 by PSCOMP_1:def 29;
              A93: g.r1 in Lower_Arc(K) by A68;
              A94: g.r2 in Lower_Arc(K) by A68;
              reconsider gr1=g.r1 as Point of TOP-REAL 2 by A93;
              reconsider gr2=g.r2 as Point of TOP-REAL 2 by A94;
                now per cases by A62,A68,XBOOLE_0:def 2;
              case A95: g.r1 in LSeg(|[b,d]|,|[b,c]|) &
                  g.r2 in LSeg(|[b,d]|,|[b,c]|);
                then A96: (gr1)`1=b & c <=(gr1)`2 & (gr1)`2 <=d by A1,Th9;
                A97: (gr2)`1=b & c <=(gr2)`2 & (gr2)`2 <=d
                                  by A1,A95,Th9;
                 then (gr1)`2=(gr2)`2 by A90,A91,A92,A96,XCMPLX_1:2;
                then |[(gr1)`1,(gr1)`2]|=g.r2 by A96,A97,EUCLID:57;
                then g.r1=g.r2 by EUCLID:57;
               hence x1=x2 by A66,A67,FUNCT_1:def 8;
              case A98: g.r1 in LSeg(|[b,d]|,|[b,c]|) &
                  g.r2 in LSeg(|[b,c]|,|[a,c]|);
                then A99: (gr1)`1=b & c <=(gr1)`2 & (gr1)`2 <=d by A1,Th9;
                A100: (gr2)`2=c & a <=(gr2)`1 & (gr2)`1 <=b
                                  by A1,A98,Th11;
                then A101: b+(gr1)`2=(gr2)`1 +c by A1,A90,A91,A92,A98,Th9;
                A102: now assume b<>gr2`1; then b>gr2`1 by A100,REAL_1:def 5;
                 hence contradiction by A90,A91,A92,A99,A100,REAL_1:67;
                end;
                  now assume gr1`2<> c; then c <gr1`2 by A99,REAL_1:def 5;
                 hence contradiction by A100,A101,REAL_1:67;
                end;
                then |[(gr1)`1,(gr1)`2]|=g.r2 by A99,A100,A102,EUCLID:57;
                then g.r1=g.r2 by EUCLID:57;
               hence x1=x2 by A66,A67,FUNCT_1:def 8;
              case A103: g.r1 in LSeg(|[b,c]|,|[a,c]|) &
                  g.r2 in LSeg(|[b,d]|,|[b,c]|);
                then A104: (gr2)`1=b & c <=(gr2)`2 & (gr2)`2 <=d by A1,Th9;
                A105: (gr1)`2=c & a <=(gr1)`1 & (gr1)`1 <=b
                                  by A1,A103,Th11;
                A106: b+(gr2)`2=(gr1)`1 +c by A1,A90,A91,A92,A103,A104,Th11;
                A107: now assume b<>gr1`1; then b>gr1`1 by A105,REAL_1:def 5;
                 hence contradiction by A104,A106,REAL_1:67;
                end;
                  now assume gr2`2<> c; then c < gr2`2 by A104,REAL_1:def 5;
                 hence contradiction by A90,A91,A92,A104,A105,REAL_1:67;
                end;
                then |[(gr2)`1,(gr2)`2]|=g.r1 by A104,A105,A107,EUCLID:57;
                then g.r1=g.r2 by EUCLID:57;
               hence x1=x2 by A66,A67,FUNCT_1:def 8;
              case A108: g.r1 in LSeg(|[b,c]|,|[a,c]|) &
                  g.r2 in LSeg(|[b,c]|,|[a,c]|);
                then A109: (gr1)`2=c & a <=(gr1)`1 & (gr1)`1 <=b by A1,Th11;
                A110: (gr2)`2=c & a <=(gr2)`1 & (gr2)`1 <=b
                                  by A1,A108,Th11;
                 then (gr1)`1=(gr2)`1 by A90,A91,A92,A109,XCMPLX_1:2;
                then |[(gr1)`1,(gr1)`2]|=g.r2 by A109,A110,EUCLID:57;
                then g.r1=g.r2 by EUCLID:57;
               hence x1=x2 by A66,A67,FUNCT_1:def 8;
              end;
             hence x1=x2;
            end;
           hence thesis by FUNCT_1:def 8;
          end;
          A111: dom k=[.0,1.] by BORSUK_1:83,FUNCT_2:def 1;
          then s1 in dom k by A65,TOPREAL5:1;
          then rxc=s1 by A82,A88,A89,A111,FUNCT_1:def 8;
         hence contradiction by A75,A88,TOPREAL5:1;
        end;
      hence s1 <= s2;
     end;
     hence thesis by A63,A64,JORDAN5C:def 3;
       end;
     hence LE p1,p2,K by A60,A63,A64,JORDAN6:def 10;
    end;
   hence LE p1,p2,K;
  end;
end;

theorem Th73: for x being set,a,b,c,d being real number
 st x in rectangle(a,b,c,d) & a<b & c <d
 holds x in LSeg(|[a,c]|,|[a,d]|) or x in LSeg(|[a,d]|,|[b,d]|)
 or x in LSeg(|[b,d]|,|[b,c]|) or x in LSeg(|[b,c]|,|[a,c]|)
proof let x be set,a,b,c,d be real number;assume
  A1: x in rectangle(a,b,c,d) & a<b & c <d;
  then x in {p: p`1=a & c <=p`2 & p`2<=d or p`2=d & a<=p`1 & p`1<=b or
      p`1=b & c <=p`2 & p`2<=d or p`2=c & a<=p`1 & p`1<=b} by Def1;
  then consider p such that
  A2: p=x &( p`1=a & c <=p`2 & p`2<=d or p`2=d & a<=p`1 & p`1<=b or
      p`1=b & c <=p`2 & p`2<=d or p`2=c & a<=p`1 & p`1<=b);
    now per cases by A2;
  case A3: p`1=a & c <=p`2 & p`2<=d;
    A4: d-c >0 by A1,SQUARE_1:11;
    then A5: (d-c)">0 by REAL_1:72;
    A6: p`2 -c >=0 by A3,SQUARE_1:12;
    A7: d-p`2 >=0 by A3,SQUARE_1:12;
      (p`2-c)/(d-c)=(p`2 -c)*(d-c)" by XCMPLX_0:def 9;
    then A8: (p`2-c)/(d-c)>=0 by A5,A6,REAL_2:121;
    reconsider r=(p`2-c)/(d-c) as Real by XREAL_0:def 1;
    A9: 1-r=(d-c)/(d-c)-(p`2-c)/(d-c) by A4,XCMPLX_1:60
          .=((d-c)-(p`2 -c))/(d-c) by XCMPLX_1:121
          .=(d-p`2)/(d-c) by XCMPLX_1:22;
    then 1-r=(d-p`2)*(d-c)" by XCMPLX_0:def 9;
    then 1-r >=0 by A5,A7,REAL_2:121;
    then 1-r+r>=0+r by REAL_1:55;
    then A10: 1>=r by XCMPLX_1:27;
    A11: ((1-r)*(|[a,c]|)+r*(|[a,d]|))`1
      =((1-r)*(|[a,c]|))`1+(r*(|[a,d]|))`1 by TOPREAL3:7
      .=(1-r)*((|[a,c]|)`1)+(r*(|[a,d]|))`1 by TOPREAL3:9
      .=(1-r)*a+(r*(|[a,d]|))`1 by EUCLID:56
      .=(1-r)*a+r*((|[a,d]|)`1) by TOPREAL3:9
      .=(1-r)*a+r*a by EUCLID:56
      .=(1-r+r)*a by XCMPLX_1:8
      .=1*a by XCMPLX_1:27
      .=p`1 by A3;
      ((1-r)*(|[a,c]|)+r*(|[a,d]|))`2
      =((1-r)*(|[a,c]|))`2+(r*(|[a,d]|))`2 by TOPREAL3:7
      .=(1-r)*((|[a,c]|)`2)+(r*(|[a,d]|))`2 by TOPREAL3:9
      .=(1-r)*c+(r*(|[a,d]|))`2 by EUCLID:56
      .=(1-r)*c+r*((|[a,d]|)`2) by TOPREAL3:9
      .=(d-p`2)/(d-c)*c+(p`2-c)/(d-c)*d by A9,EUCLID:56
      .=(d-p`2)*(d-c)"*c + (p`2-c)/(d-c)*d by XCMPLX_0:def 9
      .=(d-p`2)*(d-c)"*c + (p`2-c)*(d-c)"*d by XCMPLX_0:def 9
      .=(d-c)"*((d-p`2)*c)+ (d-c)"*(p`2-c)*d by XCMPLX_1:4
      .=(d-c)"*((d-p`2)*c)+ (d-c)"*((p`2-c)*d) by XCMPLX_1:4
      .=(d-c)"*(((d-p`2)*c)+ ((p`2-c)*d)) by XCMPLX_1:8
      .=(d-c)"*((d*c -p`2*c)+ ((p`2-c)*d)) by XCMPLX_1:40
      .=(d-c)"*((d*c -p`2*c)+ (p`2*d-c*d)) by XCMPLX_1:40
      .=(d-c)"*(d*c+ (p`2*d-c*d) -p`2*c) by XCMPLX_1:29
      .=(d-c)"*(d*c+ p`2*d-c*d -p`2*c) by XCMPLX_1:29
      .=(d-c)"*( p`2*d -p`2*c) by XCMPLX_1:26
      .=(d-c)"*( p`2*(d -c)) by XCMPLX_1:40
      .=(d-c)"*(d -c)*p`2 by XCMPLX_1:4
      .=1*p`2 by A4,XCMPLX_0:def 7
      .=p`2;
      then p=|[((1-r)*(|[a,c]|)+r*(|[a,d]|))`1,
               ((1-r)*(|[a,c]|)+r*(|[a,d]|))`2]| by A11,EUCLID:57
      .=(1-r)*(|[a,c]|)+r*(|[a,d]|) by EUCLID:57;
    then x in {(1-l)*(|[a,c]|)+l*(|[a,d]|) where l is Real: 0<=l & l<=1}
                               by A2,A8,A10;
   hence x in LSeg(|[a,c]|,|[a,d]|) or x in LSeg(|[a,d]|,|[b,d]|)
    or x in LSeg(|[b,d]|,|[b,c]|) or x in LSeg(|[b,c]|,|[a,c]|)
                                  by TOPREAL1:def 4;
  case A12: p`2=d & a<=p`1 & p`1<=b;
    A13: b-a >0 by A1,SQUARE_1:11;
    then A14: (b-a)">0 by REAL_1:72;
    A15: p`1 -a >=0 by A12,SQUARE_1:12;
    A16: b-p`1 >=0 by A12,SQUARE_1:12;
      (p`1-a)/(b-a)=(p`1 -a)*(b-a)" by XCMPLX_0:def 9;
    then A17: (p`1-a)/(b-a)>=0 by A14,A15,REAL_2:121;
    reconsider r=(p`1-a)/(b-a) as Real by XREAL_0:def 1;
    A18: 1-r=(b-a)/(b-a)-(p`1-a)/(b-a) by A13,XCMPLX_1:60
          .=((b-a)-(p`1 -a))/(b-a) by XCMPLX_1:121
          .=(b-p`1)/(b-a) by XCMPLX_1:22;
    then 1-r=(b-p`1)*(b-a)" by XCMPLX_0:def 9;
    then 1-r >=0 by A14,A16,REAL_2:121;
    then 1-r+r>=0+r by REAL_1:55;
    then A19: 1>=r by XCMPLX_1:27;
    A20: ((1-r)*(|[a,d]|)+r*(|[b,d]|))`1
      =((1-r)*(|[a,d]|))`1+(r*(|[b,d]|))`1 by TOPREAL3:7
      .=(1-r)*((|[a,d]|)`1)+(r*(|[b,d]|))`1 by TOPREAL3:9
      .=(1-r)*a+(r*(|[b,d]|))`1 by EUCLID:56
      .=(1-r)*a+r*((|[b,d]|)`1) by TOPREAL3:9
      .=(b-p`1)/(b-a)*a+(p`1-a)/(b-a)*b by A18,EUCLID:56
      .=(b-p`1)*(b-a)"*a + (p`1-a)/(b-a)*b by XCMPLX_0:def 9
      .=(b-p`1)*(b-a)"*a + (p`1-a)*(b-a)"*b by XCMPLX_0:def 9
      .=(b-a)"*((b-p`1)*a)+ (b-a)"*(p`1-a)*b by XCMPLX_1:4
      .=(b-a)"*((b-p`1)*a)+ (b-a)"*((p`1-a)*b) by XCMPLX_1:4
      .=(b-a)"*(((b-p`1)*a)+ ((p`1-a)*b)) by XCMPLX_1:8
      .=(b-a)"*((b*a -p`1*a)+ ((p`1-a)*b)) by XCMPLX_1:40
      .=(b-a)"*((b*a -p`1*a)+ (p`1*b-a*b)) by XCMPLX_1:40
      .=(b-a)"*(b*a+ (p`1*b-a*b) -p`1*a) by XCMPLX_1:29
      .=(b-a)"*(b*a+ p`1*b-a*b -p`1*a) by XCMPLX_1:29
      .=(b-a)"*( p`1*b -p`1*a) by XCMPLX_1:26
      .=(b-a)"*( p`1*(b -a)) by XCMPLX_1:40
      .=(b-a)"*(b -a)*p`1 by XCMPLX_1:4
      .=1*p`1 by A13,XCMPLX_0:def 7
      .=p`1;
      ((1-r)*(|[a,d]|)+r*(|[b,d]|))`2
      =((1-r)*(|[a,d]|))`2+(r*(|[b,d]|))`2 by TOPREAL3:7
      .=(1-r)*((|[a,d]|)`2)+(r*(|[b,d]|))`2 by TOPREAL3:9
      .=(1-r)*d+(r*(|[b,d]|))`2 by EUCLID:56
      .=(1-r)*d+r*((|[b,d]|)`2) by TOPREAL3:9
      .=(1-r)*d+r*d by EUCLID:56
      .=(1-r+r)*d by XCMPLX_1:8
      .=1*d by XCMPLX_1:27
      .=p`2 by A12;
     then p=|[((1-r)*(|[a,d]|)+r*(|[b,d]|))`1,
               ((1-r)*(|[a,d]|)+r*(|[b,d]|))`2]| by A20,EUCLID:57
      .=(1-r)*(|[a,d]|)+r*(|[b,d]|) by EUCLID:57;
    then x in {(1-l)*(|[a,d]|)+l*(|[b,d]|) where l is Real: 0<=l & l<=1}
                               by A2,A17,A19;
   hence x in LSeg(|[a,c]|,|[a,d]|) or x in LSeg(|[a,d]|,|[b,d]|)
    or x in LSeg(|[b,d]|,|[b,c]|) or x in LSeg(|[b,c]|,|[a,c]|)
                               by TOPREAL1:def 4;
  case A21: p`1=b & c <=p`2 & p`2<=d;
    A22: d-c >0 by A1,SQUARE_1:11;
    then A23: (d-c)">0 by REAL_1:72;
    A24: p`2 -c >=0 by A21,SQUARE_1:12;
    A25: d-p`2 >=0 by A21,SQUARE_1:12;
      (d-p`2)/(d-c)=(d-p`2)*(d-c)" by XCMPLX_0:def 9;
    then A26: (d-p`2)/(d-c)>=0 by A23,A25,REAL_2:121;
    reconsider r=(d-p`2)/(d-c) as Real by XREAL_0:def 1;
    A27: 1-r=(d-c)/(d-c)-(d-p`2)/(d-c) by A22,XCMPLX_1:60
          .=((d-c)-(d-p`2))/(d-c) by XCMPLX_1:121
          .=((d-c)-d+p`2)/(d-c) by XCMPLX_1:37
          .= ((d+-c)-d+p`2)/(d-c) by XCMPLX_0:def 8
          .= (-c+p`2)/(d-c) by XCMPLX_1:26
          .=(p`2-c)/(d-c) by XCMPLX_0:def 8;
    then 1-r=(p`2-c)*(d-c)" by XCMPLX_0:def 9;
    then 1-r >=0 by A23,A24,REAL_2:121;
    then 1-r+r>=0+r by REAL_1:55;
    then A28: 1>=r by XCMPLX_1:27;
    A29: ((1-r)*(|[b,d]|)+r*(|[b,c]|))`1
      =((1-r)*(|[b,d]|))`1+(r*(|[b,c]|))`1 by TOPREAL3:7
      .=(1-r)*((|[b,d]|)`1)+(r*(|[b,c]|))`1 by TOPREAL3:9
      .=(1-r)*b+(r*(|[b,c]|))`1 by EUCLID:56
      .=(1-r)*b+r*((|[b,c]|)`1) by TOPREAL3:9
      .=(1-r)*b+r*b by EUCLID:56
      .=(1-r+r)*b by XCMPLX_1:8
      .=1*b by XCMPLX_1:27
      .=p`1 by A21;
      ((1-r)*(|[b,d]|)+r*(|[b,c]|))`2
      =((1-r)*(|[b,d]|))`2+(r*(|[b,c]|))`2 by TOPREAL3:7
      .=(1-r)*((|[b,d]|)`2)+(r*(|[b,c]|))`2 by TOPREAL3:9
      .=(1-r)*d+(r*(|[b,c]|))`2 by EUCLID:56
      .=(1-r)*d+r*((|[b,c]|)`2) by TOPREAL3:9
      .=(p`2-c)/(d-c)*d+(d-p`2)/(d-c)*c by A27,EUCLID:56
      .=(p`2-c)*(d-c)"*d + (d-p`2)/(d-c)*c by XCMPLX_0:def 9
      .=(p`2-c)*(d-c)"*d + (d-p`2)*(d-c)"*c by XCMPLX_0:def 9
      .=(d-c)"*((p`2-c)*d)+ (d-c)"*(d-p`2)*c by XCMPLX_1:4
      .=(d-c)"*((p`2-c)*d)+ (d-c)"*((d-p`2)*c) by XCMPLX_1:4
      .=(d-c)"*(((p`2-c)*d)+ ((d-p`2)*c)) by XCMPLX_1:8
      .=(d-c)"*((p`2*d -c*d)+ ((d-p`2)*c)) by XCMPLX_1:40
      .=(d-c)"*((p`2*d -d*c )+ (d*c-p`2*c)) by XCMPLX_1:40
      .=(d-c)"*(p`2*d -d*c+ d*c-p`2*c) by XCMPLX_1:29
      .=(d-c)"*(p`2*d -p`2*c) by XCMPLX_1:27
      .=(d-c)"*( p`2*(d -c)) by XCMPLX_1:40
      .=(d-c)"*(d -c)*p`2 by XCMPLX_1:4
      .=1*p`2 by A22,XCMPLX_0:def 7
      .=p`2;
     then p=|[((1-r)*(|[b,d]|)+r*(|[b,c]|))`1,
               ((1-r)*(|[b,d]|)+r*(|[b,c]|))`2]| by A29,EUCLID:57
      .=(1-r)*(|[b,d]|)+r*(|[b,c]|) by EUCLID:57;
    then x in {(1-l)*(|[b,d]|)+l*(|[b,c]|) where l is Real: 0<=l & l<=1}
                               by A2,A26,A28;
   hence x in LSeg(|[a,c]|,|[a,d]|) or x in LSeg(|[a,d]|,|[b,d]|)
    or x in LSeg(|[b,d]|,|[b,c]|) or x in LSeg(|[b,c]|,|[a,c]|)
                               by TOPREAL1:def 4;
  case A30: p`2=c & a<=p`1 & p`1<=b;
    A31: b-a >0 by A1,SQUARE_1:11;
    then A32: (b-a)">0 by REAL_1:72;
    A33: p`1 -a >=0 by A30,SQUARE_1:12;
    A34: b-p`1 >=0 by A30,SQUARE_1:12;
      (b-p`1)/(b-a)=(b-p`1)*(b-a)" by XCMPLX_0:def 9;
    then A35: (b-p`1)/(b-a)>=0 by A32,A34,REAL_2:121;
    reconsider r=(b-p`1)/(b-a) as Real by XREAL_0:def 1;
    A36: 1-r=(b-a)/(b-a)-(b-p`1)/(b-a) by A31,XCMPLX_1:60
          .=((b-a)-(b-p`1))/(b-a) by XCMPLX_1:121
          .=((b-a)-b+p`1)/(b-a) by XCMPLX_1:37
          .= ((b+-a)-b+p`1)/(b-a) by XCMPLX_0:def 8
          .= (-a+p`1)/(b-a) by XCMPLX_1:26
          .=(p`1-a)/(b-a) by XCMPLX_0:def 8;
    then 1-r=(p`1-a)*(b-a)" by XCMPLX_0:def 9;
    then 1-r >=0 by A32,A33,REAL_2:121;
    then 1-r+r>=0+r by REAL_1:55;
    then A37: 1>=r by XCMPLX_1:27;
    A38: ((1-r)*(|[b,c]|)+r*(|[a,c]|))`1
      =((1-r)*(|[b,c]|))`1+(r*(|[a,c]|))`1 by TOPREAL3:7
      .=(1-r)*((|[b,c]|)`1)+(r*(|[a,c]|))`1 by TOPREAL3:9
      .=(1-r)*b+(r*(|[a,c]|))`1 by EUCLID:56
      .=(1-r)*b+r*((|[a,c]|)`1) by TOPREAL3:9
      .=(p`1-a)/(b-a)*b+(b-p`1)/(b-a)*a by A36,EUCLID:56
      .=(p`1-a)*(b-a)"*b + (b-p`1)/(b-a)*a by XCMPLX_0:def 9
      .=(p`1-a)*(b-a)"*b + (b-p`1)*(b-a)"*a by XCMPLX_0:def 9
      .=(b-a)"*((p`1-a)*b)+ (b-a)"*(b-p`1)*a by XCMPLX_1:4
      .=(b-a)"*((p`1-a)*b)+ (b-a)"*((b-p`1)*a) by XCMPLX_1:4
      .=(b-a)"*(((p`1-a)*b)+ ((b-p`1)*a)) by XCMPLX_1:8
      .=(b-a)"*((p`1*b -a*b)+ ((b-p`1)*a)) by XCMPLX_1:40
      .=(b-a)"*((p`1*b -a*b)+ (b*a-p`1*a)) by XCMPLX_1:40
      .=(b-a)"*((p`1*b -a*b)+ b*a-p`1*a) by XCMPLX_1:29
      .=(b-a)"*(p`1*b -p`1*a) by XCMPLX_1:27
      .=(b-a)"*( p`1*(b -a)) by XCMPLX_1:40
      .=(b-a)"*(b -a)*p`1 by XCMPLX_1:4
      .=1*p`1 by A31,XCMPLX_0:def 7
      .=p`1;
      ((1-r)*(|[b,c]|)+r*(|[a,c]|))`2
      =((1-r)*(|[b,c]|))`2+(r*(|[a,c]|))`2 by TOPREAL3:7
      .=(1-r)*((|[b,c]|)`2)+(r*(|[a,c]|))`2 by TOPREAL3:9
      .=(1-r)*c+(r*(|[a,c]|))`2 by EUCLID:56
      .=(1-r)*c+r*((|[a,c]|)`2) by TOPREAL3:9
      .=(1-r)*c+r*c by EUCLID:56
      .=(1-r+r)*c by XCMPLX_1:8
      .=1*c by XCMPLX_1:27
      .=p`2 by A30;
     then p=|[((1-r)*(|[b,c]|)+r*(|[a,c]|))`1,
               ((1-r)*(|[b,c]|)+r*(|[a,c]|))`2]| by A38,EUCLID:57
      .=(1-r)*(|[b,c]|)+r*(|[a,c]|) by EUCLID:57;
    then x in {(1-l)*(|[b,c]|)+l*(|[a,c]|) where l is Real: 0<=l & l<=1}
                               by A2,A35,A37;
   hence x in LSeg(|[a,c]|,|[a,d]|) or x in LSeg(|[a,d]|,|[b,d]|)
    or x in LSeg(|[b,d]|,|[b,c]|) or x in LSeg(|[b,c]|,|[a,c]|)
                                by TOPREAL1:def 4;
  end;
 hence x in LSeg(|[a,c]|,|[a,d]|) or x in LSeg(|[a,d]|,|[b,d]|)
 or x in LSeg(|[b,d]|,|[b,c]|) or x in LSeg(|[b,c]|,|[a,c]|);
end;

begin :: General Fashoda Theorem for Unit Square

theorem Th74: for p1,p2 being Point of TOP-REAL 2,
K being non empty compact Subset of TOP-REAL 2
st K=rectangle(-1,1,-1,1) & LE p1,p2,K &
p1 in LSeg(|[-1,-1]|,|[-1,1]|)
holds p2 in LSeg(|[-1,-1]|,|[-1,1]|)& p2`2>=p1`2 or
p2 in LSeg(|[-1,1]|,|[1,1]|) or p2 in LSeg(|[1,1]|,|[1,-1]|)
or p2 in LSeg(|[1,-1]|,|[-1,-1]|)& p2<>|[-1,-1]|
proof let p1,p2 be Point of TOP-REAL 2,
  K be non empty compact Subset of TOP-REAL 2;
  assume A1: K=rectangle(-1,1,-1,1) & LE p1,p2,K &
     p1 in LSeg(|[-1,-1]|,|[-1,1]|);
     then p2 in LSeg(|[-1,-1]|,|[-1,1]|) & p1`2<=p2`2
    or p2 in LSeg(|[-1,1]|,|[1,1]|) or p2 in LSeg(|[1,1]|,|[1,-1]|)
     or p2 in LSeg(|[1,-1]|,|[-1,-1]|) & p2<>W-min(K) by Th69;
 hence p2 in LSeg(|[-1,-1]|,|[-1,1]|)& p2`2>=p1`2 or
p2 in LSeg(|[-1,1]|,|[1,1]|) or p2 in LSeg(|[1,1]|,|[1,-1]|)
or p2 in LSeg(|[1,-1]|,|[-1,-1]|)& p2<>|[-1,-1]| by A1,Th56;
end;

theorem Th75: for p1,p2 being Point of TOP-REAL 2,
P,K being non empty compact Subset of TOP-REAL 2,
f being map of TOP-REAL 2,TOP-REAL 2 st
P= circle(0,0,1) & K=rectangle(-1,1,-1,1) & f=Sq_Circ
& p1 in LSeg(|[-1,-1]|,|[-1,1]|)& p1`2>=0 & LE p1,p2,K
holds LE f.p1,f.p2,P
proof let p1,p2 be Point of TOP-REAL 2,
  P,K be non empty compact Subset of TOP-REAL 2,
  f be map of TOP-REAL 2,TOP-REAL 2;
  assume A1: P= circle(0,0,1) & K=rectangle(-1,1,-1,1) & f=Sq_Circ
    & p1 in LSeg(|[-1,-1]|,|[-1,1]|)& p1`2>=0 & LE p1,p2,K;
    then A2: K is_simple_closed_curve by Th60;
    A3: P={p: |.p.|=1} by A1,Th33;
    A4: p1`1=-1 & -1<=p1`2 & p1`2<=1 by A1,Th9;
    A5: p1 in K by A1,A2,JORDAN7:5;
    A6: p2 in K by A1,A2,JORDAN7:5;
      K= {p: p`1=-1 & -1 <=p`2 & p`2<=1 or p`2=1 & -1<=p`1 & p`1<=1 or
               p`1=1 & -1 <=p`2 & p`2<=1 or p`2=-1 & -1<=p`1 & p`1<=1}
                                          by A1,Def1;
    then A7: f.:K=P by A1,A3,Th45,JGRAPH_3:33;
    A8: dom f = the carrier of TOP-REAL 2 by FUNCT_2:def 1;
    then A9: f.p1 in P by A5,A7,FUNCT_1:def 12;
    A10: f.p2 in P by A6,A7,A8,FUNCT_1:def 12;
    A11: p1`1=-1 & -1<=p1`2 & p1`2 <=1 by A1,Th9;
    A12: (p1`2)^2 >=0 by SQUARE_1:72;
      1+(p1`2)^2 >(p1`2)^2 by REAL_1:69;
    then A13: sqrt(1+(p1`2)^2)>0 by A12,SQUARE_1:93;
    A14: p1`2<=-p1`1 by A4;
      p1<>0.REAL 2 by A4,EUCLID:56,58;
    then A15: f.p1= |[p1`1/sqrt(1+(p1`2/p1`1)^2),p1`2/sqrt(1+(p1`2/p1`1)^2)]|
                          by A1,A11,A14,JGRAPH_3:def 1;
    then A16: (f.p1)`1= p1`1/sqrt(1+(p1`2/(-1))^2) by A11,EUCLID:56
               .=(p1`1)/sqrt(1+(-(p1`2/1))^2) by XCMPLX_1:189
               .=(p1`1)/sqrt(1+(p1`2)^2) by SQUARE_1:61;
    A17: (f.p1)`2= p1`2/sqrt(1+(p1`2/(-1))^2) by A11,A15,EUCLID:56
               .=(p1`2)/sqrt(1+(-(p1`2/1))^2) by XCMPLX_1:189
               .=(p1`2)/sqrt(1+(p1`2)^2) by SQUARE_1:61;
    A18: (f.p1)`1<0 by A11,A13,A16,REAL_2:128;
    A19: (f.p1)`2>=0 by A1,A13,A17,REAL_2:125;
    then f.p1 in {p9 where p9 is Point of TOP-REAL 2:p9 in P & p9`2>=0} by A9;
    then A20: f.p1 in Upper_Arc(P) by A3,JGRAPH_5:37;
      now per cases by A1,Th74;
    case A21: p2 in LSeg(|[-1,-1]|,|[-1,1]|)& p2`2>=p1`2;
      A22: (p2`2)^2 >=0 by SQUARE_1:72;
        1+(p2`2)^2 >(p2`2)^2 by REAL_1:69;
      then A23: sqrt(1+(p2`2)^2)>0 by A22,SQUARE_1:93;
      A24: p2`1=-1 & -1<=p2`2 & p2`2 <=1 by A21,Th9;
      then A25: p2`2<=-p2`1;
        p2<>0.REAL 2 by A24,EUCLID:56,58;
      then A26: f.p2= |[p2`1/sqrt(1+(p2`2/p2`1)^2),p2`2/sqrt(1+(p2`2/p2`1)^2)]|
                          by A1,A24,A25,JGRAPH_3:def 1;
      then A27: (f.p2)`1= p2`1/sqrt(1+(p2`2/(-1))^2) by A24,EUCLID:56
               .=(p2`1)/sqrt(1+(-(p2`2/1))^2) by XCMPLX_1:189
               .=(p2`1)/sqrt(1+(p2`2)^2) by SQUARE_1:61;
      A28: (f.p2)`2= p2`2/sqrt(1+(p2`2/(-1))^2) by A24,A26,EUCLID:56
               .=(p2`2)/sqrt(1+(-(p2`2/1))^2) by XCMPLX_1:189
               .=(p2`2)/sqrt(1+(p2`2)^2) by SQUARE_1:61;
      A29: (f.p2)`1<0 by A23,A24,A27,REAL_2:128;
        (p1`2)*sqrt(1+(p2`2)^2)<= (p2`2)*sqrt(1+(p1`2)^2)
                                by A1,A21,Lm4;
      then (p1`2)*sqrt(1+(p2`2)^2)/sqrt(1+(p2`2)^2)
        <= (p2`2)*sqrt(1+(p1`2)^2)/sqrt(1+(p2`2)^2) by A23,REAL_1:73;
      then (p1`2)
        <= (p2`2)*sqrt(1+(p1`2)^2)/sqrt(1+(p2`2)^2) by A23,XCMPLX_1:90
;
      then (p1`2)/sqrt(1+(p1`2)^2)
        <= (p2`2)*sqrt(1+(p1`2)^2)/sqrt(1+(p2`2)^2)/sqrt(1+(p1`2)^2)
                 by A13,REAL_1:73;
      then (p1`2)/sqrt(1+(p1`2)^2)
        <= (p2`2)*sqrt(1+(p1`2)^2)/sqrt(1+(p1`2)^2)/sqrt(1+(p2`2)^2)
                 by XCMPLX_1:48;
      then (f.p1)`2<=(f.p2)`2 by A13,A17,A28,XCMPLX_1:90;
     hence LE f.p1,f.p2,P by A3,A9,A10,A18,A19,A29,JGRAPH_5:56;
    case p2 in LSeg(|[-1,1]|,|[1,1]|);
      then A30: p2`2=1 & -1<=p2`1 & p2`1<=1 by Th11;
      A31: (p2`1)^2 >=0 by SQUARE_1:72;
        1+(p2`1)^2 >(p2`1)^2 by REAL_1:69;
      then A32: sqrt(1+(p2`1)^2)>0 by A31,SQUARE_1:93;
        p2<>0.REAL 2 by A30,EUCLID:56,58;
      then A33: f.p2= |[p2`1/sqrt(1+(p2`1/p2`2)^2),p2`2/sqrt(1+(p2`1/p2`2)^2)]|
                          by A1,A30,JGRAPH_3:14;
      then A34: (f.p2)`1=(p2`1)/sqrt(1+(p2`1)^2) by A30,EUCLID:56;
        (f.p2)`2=(p2`2)/sqrt(1+(p2`1)^2) by A30,A33,EUCLID:56;
      then A35: (f.p2)`2>=0 by A30,A32,REAL_2:125;
        -sqrt(1+(p2`1)^2)<= (p2`1)*sqrt(1+(p1`2)^2)
                                by A4,A30,Th6;
      then (p1`1)*sqrt(1+(p2`1)^2)<= (p2`1)*sqrt(1+(p1`2)^2) by A4,XCMPLX_1:180
;
      then (p1`1)*sqrt(1+(p2`1)^2)/sqrt(1+(p2`1)^2)
        <= (p2`1)*sqrt(1+(p1`2)^2)/sqrt(1+(p2`1)^2) by A32,REAL_1:73;
      then (p1`1)
        <= (p2`1)*sqrt(1+(p1`2)^2)/sqrt(1+(p2`1)^2) by A32,XCMPLX_1:90
;
      then (p1`1)/sqrt(1+(p1`2)^2)
        <= (p2`1)*sqrt(1+(p1`2)^2)/sqrt(1+(p2`1)^2)/sqrt(1+(p1`2)^2)
                 by A13,REAL_1:73;
      then (p1`1)/sqrt(1+(p1`2)^2)
        <= (p2`1)*sqrt(1+(p1`2)^2)/sqrt(1+(p1`2)^2)/sqrt(1+(p2`1)^2)
                 by XCMPLX_1:48;
      then (f.p1)`1<=(f.p2)`1 by A13,A16,A34,XCMPLX_1:90;
     hence LE f.p1,f.p2,P by A3,A9,A10,A19,A35,JGRAPH_5:57;
    case p2 in LSeg(|[1,1]|,|[1,-1]|);
      then A36: p2`1=1 & -1<=p2`2 & p2`2<=1 by Th9;
      A37: (p2`2)^2 >=0 by SQUARE_1:72;
        1+(p2`2)^2 >(p2`2)^2 by REAL_1:69;
      then A38: sqrt(1+(p2`2)^2)>0 by A37,SQUARE_1:93;
        p2<>0.REAL 2 by A36,EUCLID:56,58;
      then f.p2= |[p2`1/sqrt(1+(p2`2/p2`1)^2),p2`2/sqrt(1+(p2`2/p2`1)^2)]|
                          by A1,A36,JGRAPH_3:def 1;
      then A39: (f.p2)`1=(p2`1)/sqrt(1+(p2`2)^2) by A36,EUCLID:56;
        (p1`1)*sqrt(1+(p2`2)^2)<=0 by A4,A38,SQUARE_1:23;
      then (p1`1)*sqrt(1+(p2`2)^2)/sqrt(1+(p2`2)^2)
        <= (p2`1)*sqrt(1+(p1`2)^2)/sqrt(1+(p2`2)^2) by A13,A36,A38,REAL_1:73;
      then (p1`1)
        <= (p2`1)*sqrt(1+(p1`2)^2)/sqrt(1+(p2`2)^2) by A38,XCMPLX_1:90
;
      then (p1`1)/sqrt(1+(p1`2)^2)
        <= (p2`1)*sqrt(1+(p1`2)^2)/sqrt(1+(p2`2)^2)/sqrt(1+(p1`2)^2)
                 by A13,REAL_1:73;
      then (p1`1)/sqrt(1+(p1`2)^2)
        <= (p2`1)*sqrt(1+(p1`2)^2)/sqrt(1+(p1`2)^2)/sqrt(1+(p2`2)^2)
                 by XCMPLX_1:48;
      then A40: (f.p1)`1<=(f.p2)`1 by A13,A16,A39,XCMPLX_1:90;
        now per cases;
      case (f.p2)`2>=0;
       hence LE f.p1,f.p2,P by A3,A9,A10,A19,A40,JGRAPH_5:57;
      case A41: (f.p2)`2<0;
        then f.p2 in {p9 where p9 is Point of TOP-REAL 2:p9 in P & p9`2<=0}
                                 by A10;
        then A42: f.p2 in Lower_Arc(P) by A3,JGRAPH_5:38;
          W-min(P)=|[-1,0]| by A3,JGRAPH_5:32;
        then f.p2 <> W-min(P) by A41,EUCLID:56;
       hence LE f.p1,f.p2,P by A20,A42,JORDAN6:def 10;
      end;
     hence LE f.p1,f.p2,P;
    case p2 in LSeg(|[1,-1]|,|[-1,-1]|)& p2<>|[-1,-1]|;
      then A43: p2`2=-1 & -1<=p2`1 & p2`1<=1 by Th11;
      A44: (p2`1)^2 >=0 by SQUARE_1:72;
        1+(p2`1)^2 >(p2`1)^2 by REAL_1:69;
      then A45: sqrt(1+(p2`1)^2)>0 by A44,SQUARE_1:93;
      A46: p2`1<=-p2`2 by A43;
        p2<>0.REAL 2 by A43,EUCLID:56,58;
      then f.p2= |[p2`1/sqrt(1+(p2`1/p2`2)^2),p2`2/sqrt(1+(p2`1/p2`2)^2)]|
                          by A1,A43,A46,JGRAPH_3:14;
      then (f.p2)`2= p2`2/sqrt(1+(p2`1/(-1))^2) by A43,EUCLID:56
               .=(p2`2)/sqrt(1+(-(p2`1/1))^2) by XCMPLX_1:189
               .=(p2`2)/sqrt(1+(p2`1)^2) by SQUARE_1:61;
      then A47: (f.p2)`2<0 by A43,A45,REAL_2:128;
      then f.p2 in {p9 where p9 is Point of TOP-REAL 2:p9 in P & p9`2<=0}
                                 by A10;
      then A48: f.p2 in Lower_Arc(P) by A3,JGRAPH_5:38;
        W-min(P)=|[-1,0]| by A3,JGRAPH_5:32;
      then f.p2 <> W-min(P) by A47,EUCLID:56;
     hence LE f.p1,f.p2,P by A20,A48,JORDAN6:def 10;
    end;
 hence LE f.p1,f.p2,P;
end;

theorem Th76: for p1,p2,p3 being Point of TOP-REAL 2,
P,K being non empty compact Subset of TOP-REAL 2,
f being map of TOP-REAL 2,TOP-REAL 2 st
P= circle(0,0,1) & K=rectangle(-1,1,-1,1) & f=Sq_Circ
& p1 in LSeg(|[-1,-1]|,|[-1,1]|)& p1`2>=0 & LE p1,p2,K & LE p2,p3,K
holds LE f.p2,f.p3,P
proof let p1,p2,p3 be Point of TOP-REAL 2,
  P,K be non empty compact Subset of TOP-REAL 2,
  f be map of TOP-REAL 2,TOP-REAL 2;
  assume A1: P= circle(0,0,1) & K=rectangle(-1,1,-1,1) & f=Sq_Circ
& p1 in LSeg(|[-1,-1]|,|[-1,1]|)& p1`2>=0 & LE p1,p2,K & LE p2,p3,K;
    then A2: K is_simple_closed_curve by Th60;
    A3: P={p: |.p.|=1} by A1,Th33;
    A4: p3 in K by A1,A2,JORDAN7:5;
    A5: p2 in K by A1,A2,JORDAN7:5;
      K= {p: p`1=-1 & -1 <=p`2 & p`2<=1 or p`2=1 & -1<=p`1 & p`1<=1 or
               p`1=1 & -1 <=p`2 & p`2<=1 or p`2=-1 & -1<=p`1 & p`1<=1}
                                          by A1,Def1;
    then A6: f.:K=P by A1,A3,Th45,JGRAPH_3:33;
    A7: dom f = the carrier of TOP-REAL 2 by FUNCT_2:def 1;
    then A8: f.p2 in P by A5,A6,FUNCT_1:def 12;
    A9: f.p3 in P by A4,A6,A7,FUNCT_1:def 12;
      now per cases by A1,Th74;
    case p2 in LSeg(|[-1,-1]|,|[-1,1]|)& p2`2>=p1`2;
     hence LE f.p2,f.p3,P by A1,Th75;
    case A10: p2 in LSeg(|[-1,1]|,|[1,1]|);
      then A11: p2`2=1 & -1<=p2`1 & p2`1<=1 by Th11;
      A12: (p2`1)^2 >=0 by SQUARE_1:72;
        1+(p2`1)^2 >(p2`1)^2 by REAL_1:69;
      then A13: sqrt(1+(p2`1)^2)>0 by A12,SQUARE_1:93;
        p2<>0.REAL 2 by A11,EUCLID:56,58;
      then A14: f.p2= |[p2`1/sqrt(1+(p2`1/p2`2)^2),p2`2/sqrt(1+(p2`1/p2`2)^2)]|
                          by A1,A11,JGRAPH_3:14;
      then A15: (f.p2)`1=(p2`1)/sqrt(1+(p2`1)^2) by A11,EUCLID:56;
        (f.p2)`2=(p2`2)/sqrt(1+(p2`1)^2) by A11,A14,EUCLID:56;
      then A16: (f.p2)`2>=0 by A11,A13,REAL_2:125;
      then f.p2 in {p9 where p9 is Point of TOP-REAL 2:p9 in P & p9`2>=0}
                                by A8;
      then A17: f.p2 in Upper_Arc(P) by A3,JGRAPH_5:37;
        now per cases by A1,A10,Th70;
      case A18: p3 in LSeg(|[-1,1]|,|[1,1]|) & p2`1<=p3`1;
        then A19: p3`2=1 & -1<=p3`1 & p3`1<=1 by Th11;
        A20: (p3`1)^2 >=0 by SQUARE_1:72;
          1+(p3`1)^2 >(p3`1)^2 by REAL_1:69;
        then A21: sqrt(1+(p3`1)^2)>0 by A20,SQUARE_1:93;
          p3<>0.REAL 2 by A19,EUCLID:56,58;
        then A22: f.p3= |[p3`1/sqrt(1+(p3`1/p3`2)^2),p3`2/sqrt(1+(p3`1/p3`2)^2)
]|
                          by A1,A19,JGRAPH_3:14;
        then A23: (f.p3)`1=(p3`1)/sqrt(1+(p3`1)^2) by A19,EUCLID:56;
          (f.p3)`2=(p3`2)/sqrt(1+(p3`1)^2) by A19,A22,EUCLID:56;
        then A24: (f.p3)`2>=0 by A19,A21,REAL_2:125;
          (p2`1)*sqrt(1+(p3`1)^2)<= (p3`1)*sqrt(1+(p2`1)^2)
                                by A18,Th8;
        then (p2`1)*sqrt(1+(p3`1)^2)/sqrt(1+(p3`1)^2)
          <= (p3`1)*sqrt(1+(p2`1)^2)/sqrt(1+(p3`1)^2) by A21,REAL_1:73;
        then (p2`1)
          <= (p3`1)*sqrt(1+(p2`1)^2)/sqrt(1+(p3`1)^2) by A21,XCMPLX_1:90
;
        then (p2`1)/sqrt(1+(p2`1)^2)
          <= (p3`1)*sqrt(1+(p2`1)^2)/sqrt(1+(p3`1)^2)/sqrt(1+(p2`1)^2)
                 by A13,REAL_1:73;
        then (p2`1)/sqrt(1+(p2`1)^2)
          <= (p3`1)*sqrt(1+(p2`1)^2)/sqrt(1+(p2`1)^2)/sqrt(1+(p3`1)^2)
                 by XCMPLX_1:48;
        then (f.p2)`1<=(f.p3)`1 by A13,A15,A23,XCMPLX_1:90;
       hence LE f.p2,f.p3,P by A3,A8,A9,A16,A24,JGRAPH_5:57;
      case p3 in LSeg(|[1,1]|,|[1,-1]|);
        then A25: p3`1=1 & -1<=p3`2 & p3`2<=1 by Th9;
        A26: (p3`2)^2 >=0 by SQUARE_1:72;
          1+(p3`2)^2 >(p3`2)^2 by REAL_1:69;
        then A27: sqrt(1+(p3`2)^2)>0 by A26,SQUARE_1:93;
          p3<>0.REAL 2 by A25,EUCLID:56,58;
        then f.p3= |[p3`1/sqrt(1+(p3`2/p3`1)^2),p3`2/sqrt(1+(p3`2/p3`1)^2)]|
                          by A1,A25,JGRAPH_3:def 1;
        then A28: (f.p3)`1=(p3`1)/sqrt(1+(p3`2)^2) by A25,EUCLID:56;
        A29: -1<=-p2`1 by A11,REAL_1:50;
        A30: --1>=-p2`1 by A11,REAL_1:50;
          (p2`1)^2 = (-p2`1)^2 by SQUARE_1:61;
        then (--p2`1)*sqrt(1+(p3`2)^2)<= sqrt(1+(p2`1)^2)
                           by A25,A29,A30,Th6;
        then (p2`1)*sqrt(1+(p3`2)^2)/sqrt(1+(p3`2)^2)
          <= (p3`1)*sqrt(1+(p2`1)^2)/sqrt(1+(p3`2)^2) by A25,A27,REAL_1:73;
        then (p2`1)
          <= (p3`1)*sqrt(1+(p2`1)^2)/sqrt(1+(p3`2)^2) by A27,XCMPLX_1:90
;
        then (p2`1)/sqrt(1+(p2`1)^2)
          <= (p3`1)*sqrt(1+(p2`1)^2)/sqrt(1+(p3`2)^2)/sqrt(1+(p2`1)^2)
                 by A13,REAL_1:73;
        then (p2`1)/sqrt(1+(p2`1)^2)
          <= (p3`1)*sqrt(1+(p2`1)^2)/sqrt(1+(p2`1)^2)/sqrt(1+(p3`2)^2)
                 by XCMPLX_1:48;
        then A31: (f.p2)`1<=(f.p3)`1 by A13,A15,A28,XCMPLX_1:90;
          now per cases;
        case (f.p3)`2>=0;
         hence LE f.p2,f.p3,P by A3,A8,A9,A16,A31,JGRAPH_5:57;
        case A32: (f.p3)`2<0;
          then f.p3 in {p9 where p9 is Point of TOP-REAL 2:p9 in P & p9`2<=0}
                                 by A9;
          then A33: f.p3 in Lower_Arc(P) by A3,JGRAPH_5:38;
            W-min(P)=|[-1,0]| by A3,JGRAPH_5:32;
          then f.p3 <> W-min(P) by A32,EUCLID:56;
         hence LE f.p2,f.p3,P by A17,A33,JORDAN6:def 10;
        end;
       hence LE f.p2,f.p3,P;
      case p3 in LSeg(|[1,-1]|,|[-1,-1]|) & p3<>W-min(K);
        then A34: p3`2=-1 & -1<=p3`1 & p3`1<=1 by Th11;
        A35: (p3`1)^2 >=0 by SQUARE_1:72;
          1+(p3`1)^2 >(p3`1)^2 by REAL_1:69;
        then A36: sqrt(1+(p3`1)^2)>0 by A35,SQUARE_1:93;
        A37: -p3`2>=p3`1 by A34;
          p3<>0.REAL 2 by A34,EUCLID:56,58;
        then f.p3= |[p3`1/sqrt(1+(p3`1/p3`2)^2),p3`2/sqrt(1+(p3`1/p3`2)^2)]|
                          by A1,A34,A37,JGRAPH_3:14;
        then (f.p3)`2= p3`2/sqrt(1+(p3`1/(-1))^2) by A34,EUCLID:56
               .=(p3`2)/sqrt(1+(-(p3`1/1))^2) by XCMPLX_1:189
               .=(p3`2)/sqrt(1+(p3`1)^2) by SQUARE_1:61;
        then A38: (f.p3)`2<0 by A34,A36,REAL_2:128;
        then f.p3 in {p9 where p9 is Point of TOP-REAL 2:p9 in P & p9`2<=0}
                                 by A9;
        then A39: f.p3 in Lower_Arc(P) by A3,JGRAPH_5:38;
          W-min(P)=|[-1,0]| by A3,JGRAPH_5:32;
        then f.p3 <> W-min(P) by A38,EUCLID:56;
       hence LE f.p2,f.p3,P by A17,A39,JORDAN6:def 10;
      end;
     hence LE f.p2,f.p3,P;
    case A40: p2 in LSeg(|[1,1]|,|[1,-1]|);
      then A41: p2`1=1 & -1<=p2`2 & p2`2<=1 by Th9;
      A42: (p2`2)^2 >=0 by SQUARE_1:72;
        1+(p2`2)^2 >(p2`2)^2 by REAL_1:69;
      then A43: sqrt(1+(p2`2)^2)>0 by A42,SQUARE_1:93;
        p2<>0.REAL 2 by A41,EUCLID:56,58;
      then A44: f.p2= |[p2`1/sqrt(1+(p2`2/p2`1)^2),p2`2/sqrt(1+(p2`2/p2`1)^2)]|
                          by A1,A41,JGRAPH_3:def 1;
      then A45: (f.p2)`1=(p2`1)/sqrt(1+(p2`2)^2) by A41,EUCLID:56;
      A46: (f.p2)`2=(p2`2)/sqrt(1+(p2`2)^2) by A41,A44,EUCLID:56;
      A47: (f.p2)`1>=0 by A41,A43,A45,REAL_2:125;
        now per cases by A1,A40,Th71;
      case A48: p3 in LSeg(|[1,1]|,|[1,-1]|) & p2`2>=p3`2;
        then A49: p3`1=1 & -1<=p3`2 & p3`2<=1 by Th9;
        A50: (p3`2)^2 >=0 by SQUARE_1:72;
          1+(p3`2)^2 >(p3`2)^2 by REAL_1:69;
        then A51: sqrt(1+(p3`2)^2)>0 by A50,SQUARE_1:93;
          p3<>0.REAL 2 by A49,EUCLID:56,58;
        then A52: f.p3= |[p3`1/sqrt(1+(p3`2/p3`1)^2),p3`2/sqrt(1+(p3`2/p3`1)^2)
]|
                          by A1,A49,JGRAPH_3:def 1;
        then A53: (f.p3)`1=(p3`1)/sqrt(1+(p3`2)^2) by A49,EUCLID:56;
        A54: (f.p3)`2=(p3`2)/sqrt(1+(p3`2)^2) by A49,A52,EUCLID:56;
        A55: (f.p3)`1>=0 by A49,A51,A53,REAL_2:125;
          (p2`2)*sqrt(1+(p3`2)^2)>= (p3`2)*sqrt(1+(p2`2)^2)
                                by A48,Th8;
        then (p2`2)*sqrt(1+(p3`2)^2)/sqrt(1+(p3`2)^2)
          >= (p3`2)*sqrt(1+(p2`2)^2)/sqrt(1+(p3`2)^2) by A51,REAL_1:73;
        then (p2`2)
          >= (p3`2)*sqrt(1+(p2`2)^2)/sqrt(1+(p3`2)^2) by A51,XCMPLX_1:90
;
        then (p2`2)/sqrt(1+(p2`2)^2)
          >= (p3`2)*sqrt(1+(p2`2)^2)/sqrt(1+(p3`2)^2)/sqrt(1+(p2`2)^2)
                 by A43,REAL_1:73;
        then (p2`2)/sqrt(1+(p2`2)^2)
          >= (p3`2)*sqrt(1+(p2`2)^2)/sqrt(1+(p2`2)^2)/sqrt(1+(p3`2)^2)
                 by XCMPLX_1:48;
        then (p2`2)/sqrt(1+(p2`2)^2)
          >= (p3`2)/sqrt(1+(p3`2)^2)
                 by A43,XCMPLX_1:90;
       hence LE f.p2,f.p3,P by A3,A8,A9,A46,A47,A54,A55,JGRAPH_5:58;
      case p3 in LSeg(|[1,-1]|,|[-1,-1]|) & p3<>W-min(K);
        then A56: p3`2=-1 & -1<=p3`1 & p3`1<=1 by Th11;
        A57: (p3`1)^2 >=0 by SQUARE_1:72;
          1+(p3`1)^2 >(p3`1)^2 by REAL_1:69;
        then A58: sqrt(1+(p3`1)^2)>0 by A57,SQUARE_1:93;
        A59: -p3`2>=p3`1 by A56;
          p3<>0.REAL 2 by A56,EUCLID:56,58;
        then A60: f.p3= |[p3`1/sqrt(1+(p3`1/p3`2)^2),p3`2/sqrt(1+(p3`1/p3`2)^2)
]|
                          by A1,A56,A59,JGRAPH_3:14;
        then A61: (f.p3)`1= p3`1/sqrt(1+(p3`1/(-1))^2) by A56,EUCLID:56
               .=(p3`1)/sqrt(1+(-(p3`1/1))^2) by XCMPLX_1:189
               .=(p3`1)/sqrt(1+(p3`1)^2) by SQUARE_1:61;
          (f.p3)`2= p3`2/sqrt(1+(p3`1/(-1))^2) by A56,A60,EUCLID:56
               .=(p3`2)/sqrt(1+(-(p3`1/1))^2) by XCMPLX_1:189
               .=(p3`2)/sqrt(1+(p3`1)^2) by SQUARE_1:61;
        then A62: (f.p3)`2<0 by A56,A58,REAL_2:128;
          then f.p3 in {p9 where p9 is Point of TOP-REAL 2:p9 in P & p9`2<=0}
                                 by A9;
          then A63: f.p3 in Lower_Arc(P) by A3,JGRAPH_5:38;
            W-min(P)=|[-1,0]| by A3,JGRAPH_5:32;
         then A64: f.p3 <> W-min(P) by A62,EUCLID:56;
          now per cases;
        case (f.p2)`2>=0;
          then f.p2 in {p9 where p9 is Point of TOP-REAL 2:p9 in P & p9`2>=0}
                                 by A8;
          then f.p2 in Upper_Arc(P) by A3,JGRAPH_5:37;
         hence LE f.p2,f.p3,P by A63,A64,JORDAN6:def 10;
        case A65: (f.p2)`2<0;
            sqrt(1+(p3`1)^2)>= (p3`1)*sqrt(1+(p2`2)^2) by A41,A56,Th7;
          then (p2`1)*sqrt(1+(p3`1)^2)/sqrt(1+(p3`1)^2)
            >= (p3`1)*sqrt(1+(p2`2)^2)/sqrt(1+(p3`1)^2)
                                by A41,A58,REAL_1:73;
          then (p2`1)
           >= (p3`1)*sqrt(1+(p2`2)^2)/sqrt(1+(p3`1)^2) by A58,XCMPLX_1:90
;
          then (p2`1)/sqrt(1+(p2`2)^2)
           >= (p3`1)*sqrt(1+(p2`2)^2)/sqrt(1+(p3`1)^2)/sqrt(1+(p2`2)^2)
                 by A43,REAL_1:73;
          then (p2`1)/sqrt(1+(p2`2)^2)
           >= (p3`1)*sqrt(1+(p2`2)^2)/sqrt(1+(p2`2)^2)/sqrt(1+(p3`1)^2)
                 by XCMPLX_1:48;
          then (p2`1)/sqrt(1+(p2`2)^2)
           >= (p3`1)/sqrt(1+(p3`1)^2)
                 by A43,XCMPLX_1:90;
         hence LE f.p2,f.p3,P by A3,A8,A9,A45,A61,A62,A64,A65,JGRAPH_5:59;
        end;
       hence LE f.p2,f.p3,P;
      end;
     hence LE f.p2,f.p3,P;
    case A66: p2 in LSeg(|[1,-1]|,|[-1,-1]|)& p2<>|[-1,-1]|;
      then A67: p2`2=-1 & -1<=p2`1 & p2`1<=1 by Th11;
      A68: (p2`1)^2 >=0 by SQUARE_1:72;
        1+(p2`1)^2 >(p2`1)^2 by REAL_1:69;
      then A69: sqrt(1+(p2`1)^2)>0 by A68,SQUARE_1:93;
      A70: -p2`2>=p2`1 by A67;
        p2<>0.REAL 2 by A67,EUCLID:56,58;
      then A71: f.p2= |[p2`1/sqrt(1+(p2`1/p2`2)^2),p2`2/sqrt(1+(p2`1/p2`2)^2)]|
                          by A1,A67,A70,JGRAPH_3:14;
      then A72: (f.p2)`1= p2`1/sqrt(1+(p2`1/(-1))^2) by A67,EUCLID:56
               .=(p2`1)/sqrt(1+(-(p2`1/1))^2) by XCMPLX_1:189
               .=(p2`1)/sqrt(1+(p2`1)^2) by SQUARE_1:61;
        (f.p2)`2= p2`2/sqrt(1+(p2`1/(-1))^2) by A67,A71,EUCLID:56
               .=(p2`2)/sqrt(1+(-(p2`1/1))^2) by XCMPLX_1:189
               .=(p2`2)/sqrt(1+(p2`1)^2) by SQUARE_1:61;
      then A73: (f.p2)`2<0 by A67,A69,REAL_2:128;
        W-min(K)=|[-1,-1]| by A1,Th56;
      then A74: p3 in LSeg(|[1,-1]|,|[-1,-1]|) & p2`1>=p3`1 & p3<>W-min(K)
                                 by A1,A66,Th72;
        then A75: p3`2=-1 & -1<=p3`1 & p3`1<=1 by Th11;
        A76: (p3`1)^2 >=0 by SQUARE_1:72;
          1+(p3`1)^2 >(p3`1)^2 by REAL_1:69;
        then A77: sqrt(1+(p3`1)^2)>0 by A76,SQUARE_1:93;
        A78: -p3`2>=p3`1 by A75;
          p3<>0.REAL 2 by A75,EUCLID:56,58;
        then A79: f.p3= |[p3`1/sqrt(1+(p3`1/p3`2)^2),p3`2/sqrt(1+(p3`1/p3`2)^2)
]|
                          by A1,A75,A78,JGRAPH_3:14;
        then A80: (f.p3)`1= p3`1/sqrt(1+(p3`1/(-1))^2) by A75,EUCLID:56
               .=(p3`1)/sqrt(1+(-(p3`1/1))^2) by XCMPLX_1:189
               .=(p3`1)/sqrt(1+(p3`1)^2) by SQUARE_1:61;
          (f.p3)`2= p3`2/sqrt(1+(p3`1/(-1))^2) by A75,A79,EUCLID:56
               .=(p3`2)/sqrt(1+(-(p3`1/1))^2) by XCMPLX_1:189
               .=(p3`2)/sqrt(1+(p3`1)^2) by SQUARE_1:61;
        then A81: (f.p3)`2<0 by A75,A77,REAL_2:128;
            W-min(P)=|[-1,0]| by A3,JGRAPH_5:32;
         then A82: f.p3 <> W-min(P) by A81,EUCLID:56;
            (p2`1)*sqrt(1+(p3`1)^2)>= (p3`1)*sqrt(1+(p2`1)^2)
                                by A74,Th8;
          then (p2`1)*sqrt(1+(p3`1)^2)/sqrt(1+(p3`1)^2)
            >= (p3`1)*sqrt(1+(p2`1)^2)/sqrt(1+(p3`1)^2) by A77,REAL_1:73;
          then (p2`1)
           >= (p3`1)*sqrt(1+(p2`1)^2)/sqrt(1+(p3`1)^2) by A77,XCMPLX_1:90
;
          then (p2`1)/sqrt(1+(p2`1)^2)
           >= (p3`1)*sqrt(1+(p2`1)^2)/sqrt(1+(p3`1)^2)/sqrt(1+(p2`1)^2)
                 by A69,REAL_1:73;
          then (p2`1)/sqrt(1+(p2`1)^2)
           >= (p3`1)*sqrt(1+(p2`1)^2)/sqrt(1+(p2`1)^2)/sqrt(1+(p3`1)^2)
                 by XCMPLX_1:48;
          then (p2`1)/sqrt(1+(p2`1)^2)
           >= (p3`1)/sqrt(1+(p3`1)^2)
                 by A69,XCMPLX_1:90;
         hence LE f.p2,f.p3,P by A3,A8,A9,A72,A73,A80,A81,A82,JGRAPH_5:59;
    end;
 hence LE f.p2,f.p3,P;
end;

theorem Th77: for p being Point of TOP-REAL 2,
f being map of TOP-REAL 2,TOP-REAL 2 st
f=Sq_Circ & p`1=-1 & p`2<0 holds (f.p)`1<0 & (f.p)`2<0
proof let p be Point of TOP-REAL 2,
  f be map of TOP-REAL 2,TOP-REAL 2;
  assume A1: f=Sq_Circ & p`1=-1 & p`2<0;
      now per cases;
    case p=0.REAL 2;
     hence contradiction by A1,EUCLID:56,58;
    case A2: p<> 0.REAL 2;
        now per cases;
      case (p`2<=p`1 & -p`1<=p`2 or p`2>=p`1 & p`2<=-p`1);
        then f.p=|[p`1/sqrt(1+(p`2/p`1)^2),p`2/sqrt(1+(p`2/p`1)^2)]|
                          by A1,A2,JGRAPH_3:def 1;
        then A3: (f.p)`1= p`1/sqrt(1+(p`2/p`1)^2) &
            (f.p)`2= p`2/sqrt(1+(p`2/p`1)^2) by EUCLID:56;
        A4: (p`2/p`1)^2 >=0 by SQUARE_1:72;
          1+(p`2/p`1)^2 >(p`2/p`1)^2 by REAL_1:69;
        then sqrt(1+(p`2/p`1)^2)>0 by A4,SQUARE_1:93;
       hence (f.p)`1<0 & (f.p)`2<0 by A1,A3,REAL_2:128;
      case not (p`2<=p`1 & -p`1<=p`2 or p`2>=p`1 & p`2<=-p`1);
        then f.p=|[p`1/sqrt(1+(p`1/p`2)^2),p`2/sqrt(1+(p`1/p`2)^2)]|
                          by A1,A2,JGRAPH_3:def 1;
        then A5: (f.p)`1= p`1/sqrt(1+(p`1/p`2)^2) &
           (f.p)`2= p`2/sqrt(1+(p`1/p`2)^2) by EUCLID:56;
        A6: (p`1/p`2)^2 >=0 by SQUARE_1:72;
          1+(p`1/p`2)^2 >(p`1/p`2)^2 by REAL_1:69;
        then sqrt(1+(p`1/p`2)^2)>0 by A6,SQUARE_1:93;
       hence (f.p)`1<0 & (f.p)`2<0 by A1,A5,REAL_2:128;
      end;
     hence (f.p)`1<0 & (f.p)`2<0;
    end;
  hence (f.p)`1<0 & (f.p)`2<0;
end;

theorem Th78: for p being Point of TOP-REAL 2,
P,K being non empty compact Subset of TOP-REAL 2,
f being map of TOP-REAL 2,TOP-REAL 2 st
P= circle(0,0,1) & K=rectangle(-1,1,-1,1) & f=Sq_Circ
holds (f.p)`1>=0 iff p`1>=0
proof let p be Point of TOP-REAL 2,
  P,K be non empty compact Subset of TOP-REAL 2,
  f be map of TOP-REAL 2,TOP-REAL 2;
  assume A1: P= circle(0,0,1) & K=rectangle(-1,1,-1,1) & f=Sq_Circ;
  thus (f.p)`1>=0 implies p`1>=0
   proof assume A2: (f.p)`1>=0;
     reconsider g=(Sq_Circ") as map of TOP-REAL 2,TOP-REAL 2 by JGRAPH_3:39;
     A3: dom f = the carrier of TOP-REAL 2 by FUNCT_2:def 1;
     set q=(f.p);
      now per cases;
    case q=0.REAL 2;
     hence (g.q)`1>=0 by A2,JGRAPH_3:38;
    case A4: q<> 0.REAL 2;
        now per cases;
      case (q`2<=q`1 & -q`1<=q`2 or q`2>=q`1 & q`2<=-q`1);
        then g.q=|[q`1*sqrt(1+(q`2/q`1)^2),q`2*sqrt(1+(q`2/q`1)^2)]|
                          by A4,JGRAPH_3:38;
        then A5: (g.q)`1= q`1*sqrt(1+(q`2/q`1)^2) by EUCLID:56;
        A6: (q`2/q`1)^2 >=0 by SQUARE_1:72;
          1+(q`2/q`1)^2 >(q`2/q`1)^2 by REAL_1:69;
        then sqrt(1+(q`2/q`1)^2)>0 by A6,SQUARE_1:93;
       hence (g.q)`1>=0 by A2,A5,SQUARE_1:19;
      case not (q`2<=q`1 & -q`1<=q`2 or q`2>=q`1 & q`2<=-q`1);
        then g.q=|[q`1*sqrt(1+(q`1/q`2)^2),q`2*sqrt(1+(q`1/q`2)^2)]|
                          by A4,JGRAPH_3:38;
        then A7: (g.q)`1= q`1*sqrt(1+(q`1/q`2)^2) by EUCLID:56;
        A8: (q`1/q`2)^2 >=0 by SQUARE_1:72;
          1+(q`1/q`2)^2 >(q`1/q`2)^2 by REAL_1:69;
        then sqrt(1+(q`1/q`2)^2)>0 by A8,SQUARE_1:93;
       hence (g.q)`1>=0 by A2,A7,SQUARE_1:19;
      end;
     hence (g.q)`1>=0;
    end;
    hence p`1>=0 by A1,A3,FUNCT_1:56;
   end;
  thus p`1>=0 implies (f.p)`1>=0
   proof assume A9: p`1>=0;
      now per cases;
    case p=0.REAL 2;
     hence (f.p)`1>=0 by A1,A9,JGRAPH_3:def 1;
    case A10: p<> 0.REAL 2;
        now per cases;
      case (p`2<=p`1 & -p`1<=p`2 or p`2>=p`1 & p`2<=-p`1);
        then f.p=|[p`1/sqrt(1+(p`2/p`1)^2),p`2/sqrt(1+(p`2/p`1)^2)]|
                          by A1,A10,JGRAPH_3:def 1;
        then A11: (f.p)`1= p`1/sqrt(1+(p`2/p`1)^2) by EUCLID:56;
        A12: (p`2/p`1)^2 >=0 by SQUARE_1:72;
          1+(p`2/p`1)^2 >(p`2/p`1)^2 by REAL_1:69;
        then sqrt(1+(p`2/p`1)^2)>0 by A12,SQUARE_1:93;
       hence (f.p)`1>=0 by A9,A11,REAL_2:125;
      case not (p`2<=p`1 & -p`1<=p`2 or p`2>=p`1 & p`2<=-p`1);
        then f.p=|[p`1/sqrt(1+(p`1/p`2)^2),p`2/sqrt(1+(p`1/p`2)^2)]|
                          by A1,A10,JGRAPH_3:def 1;
        then A13: (f.p)`1= p`1/sqrt(1+(p`1/p`2)^2) by EUCLID:56;
        A14: (p`1/p`2)^2 >=0 by SQUARE_1:72;
          1+(p`1/p`2)^2 >(p`1/p`2)^2 by REAL_1:69;
        then sqrt(1+(p`1/p`2)^2)>0 by A14,SQUARE_1:93;
       hence (f.p)`1>=0 by A9,A13,REAL_2:125;
      end;
     hence (f.p)`1>=0;
    end;
    hence (f.p)`1>=0;
   end;
end;

theorem Th79: for p being Point of TOP-REAL 2,
P,K being non empty compact Subset of TOP-REAL 2,
f being map of TOP-REAL 2,TOP-REAL 2 st
P= circle(0,0,1) & K=rectangle(-1,1,-1,1) & f=Sq_Circ
holds (f.p)`2>=0 iff p`2>=0
proof let p be Point of TOP-REAL 2,
  P,K be non empty compact Subset of TOP-REAL 2,
  f be map of TOP-REAL 2,TOP-REAL 2;
  assume A1: P= circle(0,0,1) & K=rectangle(-1,1,-1,1) & f=Sq_Circ;
  thus (f.p)`2>=0 implies p`2>=0
   proof assume A2: (f.p)`2>=0;
     reconsider g=(Sq_Circ") as map of TOP-REAL 2,TOP-REAL 2 by JGRAPH_3:39;
     A3: dom f = the carrier of TOP-REAL 2 by FUNCT_2:def 1;
     set q=(f.p);
      now per cases;
    case q=0.REAL 2;
     hence (g.q)`2>=0 by A2,JGRAPH_3:38;
    case A4: q<> 0.REAL 2;
        now per cases;
      case (q`2<=q`1 & -q`1<=q`2 or q`2>=q`1 & q`2<=-q`1);
        then g.q=|[q`1*sqrt(1+(q`2/q`1)^2),q`2*sqrt(1+(q`2/q`1)^2)]|
                          by A4,JGRAPH_3:38;
        then A5: (g.q)`2= q`2*sqrt(1+(q`2/q`1)^2) by EUCLID:56;
        A6: (q`2/q`1)^2 >=0 by SQUARE_1:72;
          1+(q`2/q`1)^2 >(q`2/q`1)^2 by REAL_1:69;
        then sqrt(1+(q`2/q`1)^2)>0 by A6,SQUARE_1:93;
       hence (g.q)`2>=0 by A2,A5,SQUARE_1:19;
      case not (q`2<=q`1 & -q`1<=q`2 or q`2>=q`1 & q`2<=-q`1);
        then g.q=|[q`1*sqrt(1+(q`1/q`2)^2),q`2*sqrt(1+(q`1/q`2)^2)]|
                          by A4,JGRAPH_3:38;
        then A7: (g.q)`2= q`2*sqrt(1+(q`1/q`2)^2) by EUCLID:56;
        A8: (q`1/q`2)^2 >=0 by SQUARE_1:72;
          1+(q`1/q`2)^2 >(q`1/q`2)^2 by REAL_1:69;
        then sqrt(1+(q`1/q`2)^2)>0 by A8,SQUARE_1:93;
       hence (g.q)`2>=0 by A2,A7,SQUARE_1:19;
      end;
     hence (g.q)`2>=0;
    end;
    hence p`2>=0 by A1,A3,FUNCT_1:56;
   end;
  thus p`2>=0 implies (f.p)`2>=0
   proof assume A9: p`2>=0;
      now per cases;
    case p=0.REAL 2;
     hence (f.p)`2>=0 by A1,A9,JGRAPH_3:def 1;
    case A10: p<> 0.REAL 2;
        now per cases;
      case (p`2<=p`1 & -p`1<=p`2 or p`2>=p`1 & p`2<=-p`1);
        then f.p=|[p`1/sqrt(1+(p`2/p`1)^2),p`2/sqrt(1+(p`2/p`1)^2)]|
                          by A1,A10,JGRAPH_3:def 1;
        then A11: (f.p)`2= p`2/sqrt(1+(p`2/p`1)^2) by EUCLID:56;
        A12: (p`2/p`1)^2 >=0 by SQUARE_1:72;
          1+(p`2/p`1)^2 >(p`2/p`1)^2 by REAL_1:69;
        then sqrt(1+(p`2/p`1)^2)>0 by A12,SQUARE_1:93;
       hence (f.p)`2>=0 by A9,A11,REAL_2:125;
      case not (p`2<=p`1 & -p`1<=p`2 or p`2>=p`1 & p`2<=-p`1);
        then f.p=|[p`1/sqrt(1+(p`1/p`2)^2),p`2/sqrt(1+(p`1/p`2)^2)]|
                          by A1,A10,JGRAPH_3:def 1;
        then A13: (f.p)`2= p`2/sqrt(1+(p`1/p`2)^2) by EUCLID:56;
        A14: (p`1/p`2)^2 >=0 by SQUARE_1:72;
          1+(p`1/p`2)^2 >(p`1/p`2)^2 by REAL_1:69;
        then sqrt(1+(p`1/p`2)^2)>0 by A14,SQUARE_1:93;
       hence (f.p)`2>=0 by A9,A13,REAL_2:125;
      end;
     hence (f.p)`2>=0;
    end;
    hence (f.p)`2>=0;
   end;
end;

theorem Th80: for p,q being Point of TOP-REAL 2,
 f being map of TOP-REAL 2,TOP-REAL 2 st
 f=Sq_Circ
 & p in LSeg(|[-1,-1]|,|[-1,1]|)
 & q in LSeg(|[1,-1]|,|[-1,-1]|) holds (f.p)`1<=(f.q)`1
proof let p,q be Point of TOP-REAL 2,
  f be map of TOP-REAL 2,TOP-REAL 2;
  assume A1: f=Sq_Circ &
    p in LSeg(|[-1,-1]|,|[-1,1]|)
    & q in LSeg(|[1,-1]|,|[-1,-1]|);
    then A2: p`1=-1 & -1<=p`2 & p`2<=1 by Th9;
    A3: q`2=-1 & -1<=q`1 & q`1<=1 by A1,Th11;
    A4: p<>0.REAL 2 by A2,EUCLID:56,58;
    A5: q<>0.REAL 2 by A3,EUCLID:56,58;
      p`2<=-p`1 by A2;
    then f.p= |[p`1/sqrt(1+(p`2/p`1)^2),p`2/sqrt(1+(p`2/p`1)^2)]|
                          by A1,A2,A4,JGRAPH_3:def 1;
    then A6: (f.p)`1=(-1)/sqrt(1+(p`2/(-1))^2) by A2,EUCLID:56
               .=(-1)/sqrt(1+(-(p`2/1))^2) by XCMPLX_1:189
               .=(-1)/sqrt(1+(p`2)^2) by SQUARE_1:61;
    A7: (p`2)^2 >=0 by SQUARE_1:72;
      1+(p`2)^2 >(p`2)^2 by REAL_1:69;
    then A8: sqrt(1+(p`2)^2)>0 by A7,SQUARE_1:93;
    A9: (q`1)^2 >=0 by SQUARE_1:72;
      1+(q`1)^2 >(q`1)^2 by REAL_1:69;
    then A10: sqrt(1+(q`1)^2)>0 by A9,SQUARE_1:93;
      q`1<=-q`2 by A3;
    then f.q= |[q`1/sqrt(1+(q`1/q`2)^2),q`2/sqrt(1+(q`1/q`2)^2)]|
                          by A1,A3,A5,JGRAPH_3:14;
    then A11: (f.q)`1=q`1/sqrt(1+(q`1/(-1))^2) by A3,EUCLID:56
               .=q`1/sqrt(1+(-(q`1/1))^2) by XCMPLX_1:189
               .=q`1/sqrt(1+(q`1)^2) by SQUARE_1:61;
      -sqrt(1+(q`1)^2)<= q`1*sqrt(1+(p`2)^2) by A2,A3,Th6;
    then (-1)*sqrt(1+(q`1)^2)<= q`1*sqrt(1+(p`2)^2) by XCMPLX_1:180;
    then (-1)*sqrt(1+(q`1)^2)/sqrt(1+(q`1)^2)
      <= q`1*sqrt(1+(p`2)^2)/sqrt(1+(q`1)^2) by A10,REAL_1:73;
    then (-1)
      <= q`1*sqrt(1+(p`2)^2)/sqrt(1+(q`1)^2) by A10,XCMPLX_1:90;
    then -1<= q`1/sqrt(1+(q`1)^2)*sqrt(1+(p`2)^2) by XCMPLX_1:75;
   then (-1)/sqrt(1+(p`2)^2)
     <= q`1/sqrt(1+(q`1)^2)*sqrt(1+(p`2)^2)/sqrt(1+(p`2)^2)
                      by A8,REAL_1:73;
  hence (f.p)`1<=(f.q)`1 by A6,A8,A11,XCMPLX_1:90;
end;

theorem Th81: for p,q being Point of TOP-REAL 2,
f being map of TOP-REAL 2,TOP-REAL 2 st
f=Sq_Circ & p in LSeg(|[-1,-1]|,|[-1,1]|) & q in LSeg(|[-1,-1]|,|[-1,1]|)
& p`2>=q`2 & p`2<0
holds (f.p)`2>=(f.q)`2
proof let p,q be Point of TOP-REAL 2,
  f be map of TOP-REAL 2,TOP-REAL 2;
  assume A1: f=Sq_Circ & p in LSeg(|[-1,-1]|,|[-1,1]|) &
   q in LSeg(|[-1,-1]|,|[-1,1]|) & p`2>=q`2 & p`2<0;
   then A2: p`1=-1 & -1<=p`2 & p`2 <=1 by Th9;
   A3: q`2<0 by A1;
    A4: (p`2)^2 >=0 by SQUARE_1:72;
      1+(p`2)^2 >(p`2)^2 by REAL_1:69;
    then A5: sqrt(1+(p`2)^2)>0 by A4,SQUARE_1:93;
    A6: (q`2)^2 >=0 by SQUARE_1:72;
      1+(q`2)^2 >(q`2)^2 by REAL_1:69;
    then A7: sqrt(1+(q`2)^2)>0 by A6,SQUARE_1:93;
   A8: p`2<=-p`1 by A2;
     p<>0.REAL 2 by A1,EUCLID:56,58;
    then f.p= |[p`1/sqrt(1+(p`2/p`1)^2),p`2/sqrt(1+(p`2/p`1)^2)]|
                          by A1,A2,A8,JGRAPH_3:def 1;
    then A9: (f.p)`2= p`2/sqrt(1+(p`2/(-1))^2) by A2,EUCLID:56
               .=(p`2)/sqrt(1+(-(p`2/1))^2) by XCMPLX_1:189
               .=(p`2)/sqrt(1+(p`2)^2) by SQUARE_1:61;
   A10: q`1=-1 & -1<=q`2 & q`2 <=1 by A1,Th9;
   then A11: q`2<=-q`1;
     q<>0.REAL 2 by A1,EUCLID:56,58;
    then f.q= |[q`1/sqrt(1+(q`2/q`1)^2),q`2/sqrt(1+(q`2/q`1)^2)]|
                          by A1,A10,A11,JGRAPH_3:def 1;
    then A12: (f.q)`2= q`2/sqrt(1+(q`2/(-1))^2) by A10,EUCLID:56
               .=(q`2)/sqrt(1+(-(q`2/1))^2) by XCMPLX_1:189
               .=(q`2)/sqrt(1+(q`2)^2) by SQUARE_1:61;
     (p`2)*sqrt(1+(q`2)^2)>= (q`2)*sqrt(1+(p`2)^2) by A1,A3,Lm3;
   then (p`2)*sqrt(1+(q`2)^2)/sqrt(1+(q`2)^2)
     >= (q`2)*sqrt(1+(p`2)^2)/sqrt(1+(q`2)^2) by A7,REAL_1:73;
   then (p`2)
     >= (q`2)*sqrt(1+(p`2)^2)/sqrt(1+(q`2)^2) by A7,XCMPLX_1:90;
   then (p`2)/sqrt(1+(p`2)^2)
     >= (q`2)*sqrt(1+(p`2)^2)/sqrt(1+(q`2)^2)/sqrt(1+(p`2)^2)
                 by A5,REAL_1:73;
   then (p`2)/sqrt(1+(p`2)^2)
     >= (q`2)*sqrt(1+(p`2)^2)/sqrt(1+(p`2)^2)/sqrt(1+(q`2)^2)
                 by XCMPLX_1:48;
 hence (f.p)`2>=(f.q)`2 by A5,A9,A12,XCMPLX_1:90;
end;

theorem Th82: for p1,p2,p3,p4 being Point of TOP-REAL 2,
P,K being non empty compact Subset of TOP-REAL 2,
f being map of TOP-REAL 2,TOP-REAL 2 st
P= circle(0,0,1) & K=rectangle(-1,1,-1,1) & f=Sq_Circ
holds LE p1,p2,K & LE p2,p3,K & LE p3,p4,K implies
 f.p1,f.p2,f.p3,f.p4 are_in_this_order_on P
proof let p1,p2,p3,p4 be Point of TOP-REAL 2,
  P,K be non empty compact Subset of TOP-REAL 2,
  f be map of TOP-REAL 2,TOP-REAL 2;
  assume A1: P= circle(0,0,1) & K=rectangle(-1,1,-1,1) & f=Sq_Circ;
  then A2: K is_simple_closed_curve by Th60;
  A3: K= {p: p`1=-1 & -1 <=p`2 & p`2<=1 or p`2=1 & -1<=p`1 & p`1<=1 or
      p`1=1 & -1 <=p`2 & p`2<=1 or p`2=-1 & -1<=p`1 & p`1<=1} by A1,Def1;

  A4: P={p: |.p.|=1} by A1,Th33;
  thus LE p1,p2,K & LE p2,p3,K & LE p3,p4,K implies
    f.p1,f.p2,f.p3,f.p4 are_in_this_order_on P
   proof assume A5: LE p1,p2,K & LE p2,p3,K & LE p3,p4,K;
     then A6: p1 in K & p2 in K by A2,JORDAN7:5;
     A7: p3 in K & p4 in K by A2,A5,JORDAN7:5;
     then consider q8 being Point of TOP-REAL 2 such that
     A8: q8=p4 &
       (q8`1=-1 & -1 <=q8`2 & q8`2<=1 or q8`2=1 & -1<=q8`1 & q8`1<=1 or
       q8`1=1 & -1 <=q8`2 & q8`2<=1 or q8`2=-1 & -1<=q8`1 & q8`1<=1)
                             by A3;
     A9: LE p1,p3,K by A2,A5,JORDAN6:73;
     A10: LE p2,p4,K by A2,A5,JORDAN6:73;
     A11: W-min(K)=|[-1,-1]| by A1,Th56;
     A12: (|[-1,0]|)`1=-1 & (|[-1,0]|)`2=0 by EUCLID:56;
       1/2*(|[-1,-1]|+|[-1,1]|)=1/2*(|[-1,-1]|)+1/2*(|[-1,1]|)
                        by EUCLID:36
        .= (|[1/2*(-1),1/2*(-1)]|)+1/2*(|[-1,1]|) by EUCLID:62
        .= (|[1/2*(-1),1/2*(-1)]|)+(|[1/2*(-1),1/2*1]|) by EUCLID:62
        .= (|[1/2*(-1)+1/2*(-1),1/2*(-1)+1/2*1]|) by EUCLID:60
        .= (|[(-1),0]|);
     then A13: |[-1,0]| in LSeg(|[-1,-1]|,|[-1,1]|) by GOBOARD7:7;
       now per cases by A1,A6,A11,Th73,TOPREAL1:6;
     case A14: p1 in LSeg(|[-1,-1]|,|[-1,1]|);
      then A15: p1`1=-1 & -1<=p1`2 & p1`2<=1 by Th9;
      then A16: (f.p1)`1<0 by A1,Th78;
               K= {p: p`1=-1 & -1 <=p`2 & p`2<=1 or p`2=1 & -1<=p`1 & p`1<=1 or
               p`1=1 & -1 <=p`2 & p`2<=1 or p`2=-1 & -1<=p`1 & p`1<=1}
                                          by A1,Def1;
             then A17: f.:K=P by A1,A4,Th45,JGRAPH_3:33;
             A18: dom f = the carrier of TOP-REAL 2 by FUNCT_2:def 1;
             then A19: f.p1 in P by A6,A17,FUNCT_1:def 12;
        now per cases;
      case A20: p1`2>=0;
        then A21: LE f.p1,f.p2,P by A1,A5,A14,Th75;
        A22: LE f.p2,f.p3,P by A1,A5,A14,A20,Th76;
          LE f.p3,f.p4,P by A1,A5,A9,A14,A20,Th76;
       hence f.p1,f.p2,f.p3,f.p4 are_in_this_order_on P
                            by A21,A22,JORDAN17:def 1;
      case A23: p1`2<0;
        then A24: (f.p1)`2<0 by A1,Th79;
          now per cases;
        case A25: p2`2<0 & p2 in LSeg(|[-1,-1]|,|[-1,1]|);
         then A26: p2`1=-1 & -1<=p2`2 & p2`2<=1 by Th9;
         A27: f.p2 in P by A6,A17,A18,FUNCT_1:def 12;
         A28: p1`2<=p2`2 by A1,A5,A14,A25,Th65;
           now per cases;
         case A29: p3`2<0 & p3 in LSeg(|[-1,-1]|,|[-1,1]|);
          then A30: p3`1=-1 & -1<=p3`2 & p3`2<=1 by Th9;
          A31: f.p3 in P by A7,A17,A18,FUNCT_1:def 12;
          A32: p2`2<=p3`2 by A1,A5,A25,A29,Th65;
             now per cases;
           case A33: p4`2<0 & p4 in LSeg(|[-1,-1]|,|[-1,1]|);
             then A34: p4`1=-1 & -1<=p4`2 & p4`2<=1 by Th9;
             A35: (f.p2)`1<0 & (f.p2)`2<0 by A1,A25,A26,Th77;
             A36: (f.p3)`1<0 & (f.p3)`2<0 by A1,A29,A30,Th77;
             A37: (f.p4)`1<0 & (f.p4)`2<0 by A1,A33,A34,Th77;
               (f.p1)`2<=(f.p2)`2 by A1,A14,A25,A28,Th81;
             then A38: LE f.p1,f.p2,P by A4,A16,A19,A24,A27,A35,JGRAPH_5:54;
               (f.p2)`2<=(f.p3)`2 by A1,A25,A29,A32,Th81;
             then A39: LE f.p2,f.p3,P by A4,A27,A31,A35,A36,JGRAPH_5:54;
             A40: f.p4 in P by A7,A17,A18,FUNCT_1:def 12;
               p3`2<=p4`2 by A1,A5,A29,A33,Th65;
             then (f.p3)`2<=(f.p4)`2 by A1,A29,A33,Th81;
             then LE f.p3,f.p4,P by A4,A31,A36,A37,A40,JGRAPH_5:54;
            hence f.p1,f.p2,f.p3,f.p4 are_in_this_order_on P
                         by A38,A39,JORDAN17:def 1;
           case A41: not(p4`2<0 & p4 in LSeg(|[-1,-1]|,|[-1,1]|));
       A42: now per cases by A1,A7,Th73;
       case p4 in LSeg(|[-1,-1]|,|[-1,1]|);
        hence p4 in LSeg(|[-1,-1]|,|[-1,1]|) & (|[-1,0]|)`2<=p4`2
       or p4 in LSeg(|[-1,1]|,|[1,1]|) or p4 in LSeg(|[1,1]|,|[1,-1]|)
       or p4 in LSeg(|[1,-1]|,|[-1,-1]|) & p4<>W-min(K) by A41,EUCLID:56;
       case p4 in LSeg(|[-1,1]|,|[1,1]|);
        hence p4 in LSeg(|[-1,-1]|,|[-1,1]|) & (|[-1,0]|)`2<=p4`2
       or p4 in LSeg(|[-1,1]|,|[1,1]|) or p4 in LSeg(|[1,1]|,|[1,-1]|)
       or p4 in LSeg(|[1,-1]|,|[-1,-1]|) & p4<>W-min(K);
       case p4 in LSeg(|[1,1]|,|[1,-1]|);
        hence p4 in LSeg(|[-1,-1]|,|[-1,1]|) & (|[-1,0]|)`2<=p4`2
       or p4 in LSeg(|[-1,1]|,|[1,1]|) or p4 in LSeg(|[1,1]|,|[1,-1]|)
       or p4 in LSeg(|[1,-1]|,|[-1,-1]|) & p4<>W-min(K);
       case A43: p4 in LSeg(|[1,-1]|,|[-1,-1]|);
         A44: W-min(K)=|[-1,-1]| by A1,Th56;
           now assume A45: p4= W-min(K);
           then p4`2=-1 by A44,EUCLID:56;
          hence contradiction by A41,A44,A45,TOPREAL1:6;
         end;
        hence p4 in LSeg(|[-1,-1]|,|[-1,1]|) & (|[-1,0]|)`2<=p4`2
       or p4 in LSeg(|[-1,1]|,|[1,1]|) or p4 in LSeg(|[1,1]|,|[1,-1]|)
       or p4 in LSeg(|[1,-1]|,|[-1,-1]|) & p4<>W-min(K) by A43;
       end;
       A46: (f.p2)`1<0 & (f.p2)`2<0 by A1,A25,A26,Th77;
       A47: (f.p3)`1<0 & (f.p3)`2<0 by A1,A29,A30,Th77;
         (f.p1)`2<=(f.p2)`2 by A1,A14,A25,A28,Th81;
       then A48: LE f.p1,f.p2,P by A4,A16,A19,A24,A27,A46,JGRAPH_5:54;
         (f.p2)`2<=(f.p3)`2 by A1,A25,A29,A32,Th81;
       then A49: LE f.p2,f.p3,P by A4,A27,A31,A46,A47,JGRAPH_5:54;
       A50: now per cases;
       case A51: p4`1=-1 & p4`2<0 & p1`2<=p4`2;
           now per cases by A42;
         case p4 in LSeg(|[-1,-1]|,|[-1,1]|) & (|[-1,0]|)`2<=p4`2;
          hence contradiction by A51,EUCLID:56;
         case p4 in LSeg(|[-1,1]|,|[1,1]|);
           then p4`2=1 by Th11;
          hence contradiction by A51;
         case p4 in LSeg(|[1,1]|,|[1,-1]|);
          hence contradiction by A51,Th9;
         case A52: p4 in LSeg(|[1,-1]|,|[-1,-1]|) & p4<>W-min(K);
           then A53: p4`2= -1 by Th11;
             W-min(K)= |[-1,-1]| by A1,Th56;
           then (W-min(K))`1=-1 & (W-min(K))`2=-1 by EUCLID:56;
          hence contradiction by A51,A52,A53,TOPREAL3:11;
         end;
        hence contradiction;
       case A54: not (p4`1=-1 & p4`2<0 & p1`2<=p4`2);
         A55: p4 in LSeg(|[-1,-1]|,|[-1,1]|) or p4 in LSeg(|[-1,1]|,|[1,1]|)
          or p4 in LSeg(|[1,1]|,|[1,-1]|) or p4 in LSeg(|[1,-1]|,|[-1,-1]|)
                             by A1,A7,Th73;
           now per cases by A54;
         case A56: p4`1<>-1;
               K= {p: p`1=-1 & -1 <=p`2 & p`2<=1 or p`2=1 & -1<=p`1 & p`1<=1 or
               p`1=1 & -1 <=p`2 & p`2<=1 or p`2=-1 & -1<=p`1 & p`1<=1}
                                          by A1,Def1;
             then A57: f.:K=P by A1,A4,Th45,JGRAPH_3:33;
             A58: dom f = the carrier of TOP-REAL 2 by FUNCT_2:def 1;
             A59: (f.p1)`1<0 & (f.p1)`2<=0 by A1,A15,A23,Th77;
             A60: Upper_Arc(P)
                ={p where p is Point of TOP-REAL 2:p in P & p`2>=0}
                                     by A4,JGRAPH_5:37;
             A61: f.p1 in P by A6,A57,A58,FUNCT_1:def 12;
               Lower_Arc(P)
                ={p where p is Point of TOP-REAL 2:p in P & p`2<=0}
                                     by A4,JGRAPH_5:38;
             then A62: f.p1 in Lower_Arc(P) by A59,A61;
             A63: f.(|[-1,0]|)=W-min(P) by A1,A4,Th18,JGRAPH_5:32;
             A64: now assume f.p1=W-min(P);
               then p1=|[-1,0]| by A1,A58,A63,FUNCT_1:def 8;
              hence contradiction by A23,EUCLID:56;
             end;
             now per cases by A55,A56,Th9;
           case p4 in LSeg(|[-1,1]|,|[1,1]|);
             then A65: p4`2=1 by Th11;
             A66: f.p4 in P by A7,A57,A58,FUNCT_1:def 12;
               (f.p4)`2>=0 by A1,A65,Th79;
             then f.p4 in Upper_Arc(P) by A60,A66;
            hence LE f.p4,f.p1,P by A62,A64,JORDAN6:def 10;
           case p4 in LSeg(|[1,1]|,|[1,-1]|);
             then A67: p4`1=1 by Th9;
               K= {p: p`1=-1 & -1 <=p`2 & p`2<=1 or p`2=1 & -1<=p`1 & p`1<=1 or
               p`1=1 & -1 <=p`2 & p`2<=1 or p`2=-1 & -1<=p`1 & p`1<=1}
                                          by A1,Def1;
             then A68: f.:K=P by A1,A4,Th45,JGRAPH_3:33;
             A69: dom f = the carrier of TOP-REAL 2 by FUNCT_2:def 1;
             then A70: f.p4 in P by A7,A68,FUNCT_1:def 12;
             A71: f.p1 in P by A6,A68,A69,FUNCT_1:def 12;
             A72: (f.p1)`1<0 & (f.p1)`2<=0 by A1,A15,A23,Th77;
             A73: f.(|[-1,0]|)=W-min(P) by A1,A4,Th18,JGRAPH_5:32;
             A74: now assume f.p1=W-min(P);
               then p1=|[-1,0]| by A1,A69,A73,FUNCT_1:def 8;
              hence contradiction by A23,EUCLID:56;
             end;
             A75: (f.p4)`1>=0 by A1,A67,Th78;
               now per cases;
             case A76: (f.p4)`2>=0;
             A77: (f.p1)`1<0 & (f.p1)`2<=0 by A1,A15,A23,Th77;
             A78: Upper_Arc(P)
                ={p where p is Point of TOP-REAL 2:p in P & p`2>=0}
                                     by A4,JGRAPH_5:37;
             A79: f.(|[-1,0]|)=W-min(P) by A1,A4,Th18,JGRAPH_5:32;
             A80: now assume f.p1=W-min(P);
               then p1=|[-1,0]| by A1,A69,A79,FUNCT_1:def 8;
              hence contradiction by A23,EUCLID:56;
             end;
             A81: f.p4 in Upper_Arc(P) by A70,A76,A78;
               Lower_Arc(P)
                ={p where p is Point of TOP-REAL 2:p in P & p`2<=0}
                                     by A4,JGRAPH_5:38;
             then f.p1 in Lower_Arc(P) by A71,A77;
            hence LE f.p4,f.p1,P by A80,A81,JORDAN6:def 10;
            case A82: (f.p4)`2<0;
               (f.p1)`1<=0 &
             (f.p4)`1>=(f.p1)`1 by A72,A75,AXIOMS:22;
            hence LE f.p4,f.p1,P
                         by A4,A70,A71,A72,A74,A82,JGRAPH_5:59;
            end;
            hence LE f.p4,f.p1,P;
           case A83: p4 in LSeg(|[1,-1]|,|[-1,-1]|);
             then p4`2=-1 & -1 <=p4`1 & p4`1<=1 by Th11;
             then A84: (f.p4)`2<0 by A1,Th79;
             A85: f.p4 in P by A7,A57,A58,FUNCT_1:def 12;
             A86: (f.p1)`1<0 & (f.p1)`2<=0 by A1,A15,A23,Th77;
               (f.p4)`1>=(f.p1)`1 by A1,A14,A83,Th80;
            hence LE f.p4,f.p1,P by A4,A61,A64,A84,A85,A86,JGRAPH_5:59;
           end;
          hence LE f.p4,f.p1,P;
         case A87: p4`1=-1 & p4`2>=0;
               K= {p: p`1=-1 & -1 <=p`2 & p`2<=1 or p`2=1 & -1<=p`1 & p`1<=1 or
               p`1=1 & -1 <=p`2 & p`2<=1 or p`2=-1 & -1<=p`1 & p`1<=1}
                                          by A1,Def1;
             then A88: f.:K=P by A1,A4,Th45,JGRAPH_3:33;
             A89: dom f = the carrier of TOP-REAL 2 by FUNCT_2:def 1;
             then A90: f.p4 in P by A7,A88,FUNCT_1:def 12;
             A91: f.p1 in P by A6,A88,A89,FUNCT_1:def 12;
            A92: (f.p4)`2>=0 by A1,A87,Th79;
             A93: (f.p1)`1<0 & (f.p1)`2<=0 by A1,A15,A23,Th77;
             A94: Upper_Arc(P)
                ={p where p is Point of TOP-REAL 2:p in P & p`2>=0}
                                     by A4,JGRAPH_5:37;
             A95: f.(|[-1,0]|)=W-min(P) by A1,A4,Th18,JGRAPH_5:32;
             A96: now assume f.p1=W-min(P);
               then p1=|[-1,0]| by A1,A89,A95,FUNCT_1:def 8;
              hence contradiction by A23,EUCLID:56;
             end;
             A97: f.p4 in Upper_Arc(P) by A90,A92,A94;
               Lower_Arc(P)
                ={p where p is Point of TOP-REAL 2:p in P & p`2<=0}
                                     by A4,JGRAPH_5:38;
             then f.p1 in Lower_Arc(P) by A91,A93;
            hence LE f.p4,f.p1,P by A96,A97,JORDAN6:def 10;
         case A98: p4`1=-1 & p4`2<0 & p1`2>p4`2;
             then A99: p4 in LSeg(|[-1,-1]|,|[-1,1]|) by A8,Th10;
               K= {p: p`1=-1 & -1 <=p`2 & p`2<=1 or p`2=1 & -1<=p`1 & p`1<=1 or
               p`1=1 & -1 <=p`2 & p`2<=1 or p`2=-1 & -1<=p`1 & p`1<=1}
                                          by A1,Def1;
             then A100: f.:K=P by A1,A4,Th45,JGRAPH_3:33;
             A101: dom f = the carrier of TOP-REAL 2 by FUNCT_2:def 1;
             then A102: f.p4 in P by A7,A100,FUNCT_1:def 12;
             A103: f.p1 in P by A6,A100,A101,FUNCT_1:def 12;
           A104: (f.p1)`1<0 & (f.p1)`2<0 by A1,A15,A23,Th77;
           A105: (f.p4)`2<=(f.p1)`2 by A1,A14,A23,A98,A99,Th81;
           A106: (f.p4)`1<0 by A1,A98,Th78;
             (f.p4)`2<0 by A1,A98,Th79;
          hence LE f.p4,f.p1,P by A4,A102,A103,A104,A105,A106,JGRAPH_5:54;
         end;
        hence LE f.p4,f.p1,P;
       end;
       A107: K={p: p`1=-1 & -1 <=p`2 & p`2<=1 or p`2=1 & -1<=p`1 & p`1<=1 or
          p`1=1 & -1 <=p`2 & p`2<=1 or p`2=-1 & -1<=p`1 & p`1<=1}
                         by A1,Def1;
       thus K={p: p`1=-1 & -1 <=p`2 & p`2<=1 or
          p`1=1 & -1 <=p`2 & p`2<=1 or -1=p`2 & -1<=p`1 & p`1<=1 or
          1=p`2 & -1<=p`1 & p`1<=1}
       proof thus K c=
         {p: p`1=-1 & -1 <=p`2 & p`2<=1 or
          p`1=1 & -1 <=p`2 & p`2<=1 or -1=p`2 & -1<=p`1 & p`1<=1 or
          1=p`2 & -1<=p`1 & p`1<=1}
          proof let x be set;assume x in K;
            then consider p such that
            A108: p=x & (p`1=-1 & -1 <=p`2 & p`2<=1 or
            p`2=1 & -1<=p`1 & p`1<=1 or
            p`1=1 & -1 <=p`2 & p`2<=1 or p`2=-1 & -1<=p`1 & p`1<=1)
                        by A107;
           thus x in
             {q: q`1=-1 & -1 <=q`2 & q`2<=1 or
             q`1=1 & -1 <=q`2 & q`2<=1 or -1=q`2 & -1<=q`1 & q`1<=1 or
             1=q`2 & -1<=q`1 & q`1<=1} by A108;
          end;
        thus
           {p: p`1=-1 & -1 <=p`2 & p`2<=1 or
          p`1=1 & -1 <=p`2 & p`2<=1 or -1=p`2 & -1<=p`1 & p`1<=1 or
          1=p`2 & -1<=p`1 & p`1<=1} c= K
          proof let x be set;assume x in
           {p: p`1=-1 & -1 <=p`2 & p`2<=1 or
           p`1=1 & -1 <=p`2 & p`2<=1 or -1=p`2 & -1<=p`1 & p`1<=1 or
           1=p`2 & -1<=p`1 & p`1<=1};
           then consider p such that
           A109: p=x &
             (p`1=-1 & -1 <=p`2 & p`2<=1 or
            p`1=1 & -1 <=p`2 & p`2<=1 or -1=p`2 & -1<=p`1 & p`1<=1 or
            1=p`2 & -1<=p`1 & p`1<=1);
            thus x in K by A107,A109;
          end;
       end;
         thus f.p1,f.p2,f.p3,f.p4 are_in_this_order_on P
                         by A48,A49,A50,JORDAN17:def 1;
           end;
          hence f.p1,f.p2,f.p3,f.p4 are_in_this_order_on P;
         case A110: not(p3`2<0 & p3 in LSeg(|[-1,-1]|,|[-1,1]|));
       A111: now per cases by A1,A7,Th73;
       case p3 in LSeg(|[-1,-1]|,|[-1,1]|);
        hence p3 in LSeg(|[-1,-1]|,|[-1,1]|) & (|[-1,0]|)`2<=p3`2
       or p3 in LSeg(|[-1,1]|,|[1,1]|) or p3 in LSeg(|[1,1]|,|[1,-1]|)
       or p3 in LSeg(|[1,-1]|,|[-1,-1]|) & p3<>W-min(K) by A110,EUCLID:56;
       case p3 in LSeg(|[-1,1]|,|[1,1]|);
        hence p3 in LSeg(|[-1,-1]|,|[-1,1]|) & (|[-1,0]|)`2<=p3`2
       or p3 in LSeg(|[-1,1]|,|[1,1]|) or p3 in LSeg(|[1,1]|,|[1,-1]|)
       or p3 in LSeg(|[1,-1]|,|[-1,-1]|) & p3<>W-min(K);
       case p3 in LSeg(|[1,1]|,|[1,-1]|);
        hence p3 in LSeg(|[-1,-1]|,|[-1,1]|) & (|[-1,0]|)`2<=p3`2
       or p3 in LSeg(|[-1,1]|,|[1,1]|) or p3 in LSeg(|[1,1]|,|[1,-1]|)
       or p3 in LSeg(|[1,-1]|,|[-1,-1]|) & p3<>W-min(K);
       case A112: p3 in LSeg(|[1,-1]|,|[-1,-1]|);
         A113: W-min(K)=|[-1,-1]| by A1,Th56;
           now assume A114: p3= W-min(K);
           then p3`2=-1 by A113,EUCLID:56;
          hence contradiction by A110,A113,A114,TOPREAL1:6;
         end;
        hence p3 in LSeg(|[-1,-1]|,|[-1,1]|) & (|[-1,0]|)`2<=p3`2
       or p3 in LSeg(|[-1,1]|,|[1,1]|) or p3 in LSeg(|[1,1]|,|[1,-1]|)
       or p3 in LSeg(|[1,-1]|,|[-1,-1]|) & p3<>W-min(K) by A112;
       end;
       then A115: LE |[-1,0]|,p3,K by A1,A13,Th69;
       A116: (f.p2)`1<0 & (f.p2)`2<0 by A1,A25,A26,Th77;
         (f.p1)`2<=(f.p2)`2 by A1,A14,A25,A28,Th81;
       then A117: LE f.p1,f.p2,P by A4,A16,A19,A24,A27,A116,JGRAPH_5:54;
       A118: LE f.p3,f.p4,P by A1,A5,A12,A13,A115,Th76;
       A119: now per cases;
       case A120: p4`1=-1 & p4`2<0 & p1`2<=p4`2;
         A121: (|[-1,-1]|)`1=-1 by EUCLID:56;
         A122: (|[-1,-1]|)`2=-1 by EUCLID:56;
         A123: (|[-1,1]|)`1=-1 by EUCLID:56;
         A124: (|[-1,1]|)`2=1 by EUCLID:56;

           -1<=p4`2 & p4`2<=1 by A1,A7,Th28;
         then A125: p4 in LSeg(|[-1,-1]|,|[-1,1]|) by A120,A121,A122,A123,A124,
GOBOARD7:8;
           now per cases by A111;
         case A126: p3 in LSeg(|[-1,-1]|,|[-1,1]|) & (|[-1,0]|)`2<=p3`2;
           then 0<=p3`2 by EUCLID:56;
          hence contradiction by A1,A5,A120,A125,A126,Th65;
         case A127: p3 in LSeg(|[-1,1]|,|[1,1]|);
           then LE p4,p3,K by A1,A125,Th69;
           then A128: p3=p4 by A2,A5,JORDAN6:72;
             p3`2=1 by A127,Th11;
          hence contradiction by A120,A128;
         case A129: p3 in LSeg(|[1,1]|,|[1,-1]|);
           then LE p4,p3,K by A1,A125,Th69;
           then p3=p4 by A2,A5,JORDAN6:72;
          hence contradiction by A120,A129,Th9;
         case A130: p3 in LSeg(|[1,-1]|,|[-1,-1]|) & p3<>W-min(K);
           then LE p4,p3,K by A1,A125,Th69;
           then A131: p3=p4 by A2,A5,JORDAN6:72;
           A132: p3`2= -1 by A130,Th11;
             W-min(K)= |[-1,-1]| by A1,Th56;
           then (W-min(K))`1=-1 & (W-min(K))`2=-1 by EUCLID:56;
          hence contradiction by A120,A130,A131,A132,TOPREAL3:11;
         end;
        hence contradiction;
       case A133: not (p4`1=-1 & p4`2<0 & p1`2<=p4`2);
         A134: p4 in LSeg(|[-1,-1]|,|[-1,1]|) or p4 in LSeg(|[-1,1]|,|[1,1]|)
          or p4 in LSeg(|[1,1]|,|[1,-1]|) or p4 in LSeg(|[1,-1]|,|[-1,-1]|)
                             by A1,A7,Th73;
           now per cases by A133;
         case A135: p4`1<>-1;
               K= {p: p`1=-1 & -1 <=p`2 & p`2<=1 or p`2=1 & -1<=p`1 & p`1<=1 or
               p`1=1 & -1 <=p`2 & p`2<=1 or p`2=-1 & -1<=p`1 & p`1<=1}
                                          by A1,Def1;
             then A136: f.:K=P by A1,A4,Th45,JGRAPH_3:33;
             A137: dom f = the carrier of TOP-REAL 2 by FUNCT_2:def 1;
             A138: (f.p1)`1<0 & (f.p1)`2<=0 by A1,A15,A23,Th77;
             A139: Upper_Arc(P)
                ={p where p is Point of TOP-REAL 2:p in P & p`2>=0}
                                     by A4,JGRAPH_5:37;
             A140: f.p1 in P by A6,A136,A137,FUNCT_1:def 12;
               Lower_Arc(P)
                ={p where p is Point of TOP-REAL 2:p in P & p`2<=0}
                                     by A4,JGRAPH_5:38;
             then A141: f.p1 in Lower_Arc(P) by A138,A140;
             A142: f.(|[-1,0]|)=W-min(P) by A1,A4,Th18,JGRAPH_5:32;
             A143: now assume f.p1=W-min(P);
               then p1=|[-1,0]| by A1,A137,A142,FUNCT_1:def 8;
              hence contradiction by A23,EUCLID:56;
             end;
             now per cases by A134,A135,Th9;
           case p4 in LSeg(|[-1,1]|,|[1,1]|);
             then A144: p4`2=1 by Th11;
             A145: f.p4 in P by A7,A136,A137,FUNCT_1:def 12;
               (f.p4)`2>=0 by A1,A144,Th79;
             then f.p4 in Upper_Arc(P) by A139,A145;
            hence LE f.p4,f.p1,P by A141,A143,JORDAN6:def 10;
           case p4 in LSeg(|[1,1]|,|[1,-1]|);
             then A146: p4`1=1 by Th9;
               K= {p: p`1=-1 & -1 <=p`2 & p`2<=1 or p`2=1 & -1<=p`1 & p`1<=1 or
               p`1=1 & -1 <=p`2 & p`2<=1 or p`2=-1 & -1<=p`1 & p`1<=1}
                                          by A1,Def1;
             then A147: f.:K=P by A1,A4,Th45,JGRAPH_3:33;
             A148: dom f = the carrier of TOP-REAL 2 by FUNCT_2:def 1;
             then A149: f.p4 in P by A7,A147,FUNCT_1:def 12;
             A150: f.p1 in P by A6,A147,A148,FUNCT_1:def 12;
             A151: (f.p1)`1<0 & (f.p1)`2<=0 by A1,A15,A23,Th77;
             A152: f.(|[-1,0]|)=W-min(P) by A1,A4,Th18,JGRAPH_5:32;
             A153: now assume f.p1=W-min(P);
               then p1=|[-1,0]| by A1,A148,A152,FUNCT_1:def 8;
              hence contradiction by A23,EUCLID:56;
             end;
             A154: (f.p4)`1>=0 by A1,A146,Th78;
               now per cases;
             case A155: (f.p4)`2>=0;
             A156: (f.p1)`1<0 & (f.p1)`2<=0 by A1,A15,A23,Th77;
             A157: Upper_Arc(P)
                ={p where p is Point of TOP-REAL 2:p in P & p`2>=0}
                                     by A4,JGRAPH_5:37;
             A158: f.(|[-1,0]|)=W-min(P) by A1,A4,Th18,JGRAPH_5:32;
             A159: now assume f.p1=W-min(P);
               then p1=|[-1,0]| by A1,A148,A158,FUNCT_1:def 8;
              hence contradiction by A23,EUCLID:56;
             end;
             A160: f.p4 in Upper_Arc(P) by A149,A155,A157;
               Lower_Arc(P)
                ={p where p is Point of TOP-REAL 2:p in P & p`2<=0}
                                     by A4,JGRAPH_5:38;
             then f.p1 in Lower_Arc(P) by A150,A156;
            hence LE f.p4,f.p1,P by A159,A160,JORDAN6:def 10;
            case A161: (f.p4)`2<0;
               (f.p1)`1<=0 &
             (f.p4)`1>=(f.p1)`1 by A151,A154,AXIOMS:22;
            hence LE f.p4,f.p1,P
                         by A4,A149,A150,A151,A153,A161,JGRAPH_5:59;
            end;
            hence LE f.p4,f.p1,P;
           case A162: p4 in LSeg(|[1,-1]|,|[-1,-1]|);
             then p4`2=-1 & -1 <=p4`1 & p4`1<=1 by Th11;
             then A163: (f.p4)`2<0 by A1,Th79;
             A164: f.p4 in P by A7,A136,A137,FUNCT_1:def 12;
             A165: (f.p1)`1<0 & (f.p1)`2<=0 by A1,A15,A23,Th77;
               (f.p4)`1>=(f.p1)`1 by A1,A14,A162,Th80;
            hence LE f.p4,f.p1,P by A4,A140,A143,A163,A164,A165,JGRAPH_5:59;
           end;
          hence LE f.p4,f.p1,P;
         case A166: p4`1=-1 & p4`2>=0;
               K= {p: p`1=-1 & -1 <=p`2 & p`2<=1 or p`2=1 & -1<=p`1 & p`1<=1 or
               p`1=1 & -1 <=p`2 & p`2<=1 or p`2=-1 & -1<=p`1 & p`1<=1}
                                          by A1,Def1;
             then A167: f.:K=P by A1,A4,Th45,JGRAPH_3:33;
             A168: dom f = the carrier of TOP-REAL 2 by FUNCT_2:def 1;
             then A169: f.p4 in P by A7,A167,FUNCT_1:def 12;
             A170: f.p1 in P by A6,A167,A168,FUNCT_1:def 12;
            A171: (f.p4)`2>=0 by A1,A166,Th79;
             A172: (f.p1)`1<0 & (f.p1)`2<=0 by A1,A15,A23,Th77;
             A173: Upper_Arc(P)
                ={p where p is Point of TOP-REAL 2:p in P & p`2>=0}
                                     by A4,JGRAPH_5:37;
             A174: f.(|[-1,0]|)=W-min(P) by A1,A4,Th18,JGRAPH_5:32;
             A175: now assume f.p1=W-min(P);
               then p1=|[-1,0]| by A1,A168,A174,FUNCT_1:def 8;
              hence contradiction by A23,EUCLID:56;
             end;
             A176: f.p4 in Upper_Arc(P) by A169,A171,A173;
               Lower_Arc(P)
                ={p where p is Point of TOP-REAL 2:p in P & p`2<=0}
                                     by A4,JGRAPH_5:38;
             then f.p1 in Lower_Arc(P) by A170,A172;
            hence LE f.p4,f.p1,P by A175,A176,JORDAN6:def 10;
         case A177: p4`1=-1 & p4`2<0 & p1`2>p4`2;
             then A178: p4 in LSeg(|[-1,-1]|,|[-1,1]|) by A8,Th10;
               K= {p: p`1=-1 & -1 <=p`2 & p`2<=1 or p`2=1 & -1<=p`1 & p`1<=1 or
               p`1=1 & -1 <=p`2 & p`2<=1 or p`2=-1 & -1<=p`1 & p`1<=1}
                                          by A1,Def1;
             then A179: f.:K=P by A1,A4,Th45,JGRAPH_3:33;
             A180: dom f = the carrier of TOP-REAL 2 by FUNCT_2:def 1;
             then A181: f.p4 in P by A7,A179,FUNCT_1:def 12;
             A182: f.p1 in P by A6,A179,A180,FUNCT_1:def 12;

           A183: (f.p1)`1<0 & (f.p1)`2<0 by A1,A15,A23,Th77;

           A184: (f.p4)`2<=(f.p1)`2 by A1,A14,A23,A177,A178,Th81;
           A185: (f.p4)`1<0 by A1,A177,Th78;
             (f.p4)`2<0 by A1,A177,Th79;
          hence LE f.p4,f.p1,P by A4,A181,A182,A183,A184,A185,JGRAPH_5:54;
         end;
        hence LE f.p4,f.p1,P;
       end;
       A186: K={p: p`1=-1 & -1 <=p`2 & p`2<=1 or p`2=1 & -1<=p`1 & p`1<=1 or
          p`1=1 & -1 <=p`2 & p`2<=1 or p`2=-1 & -1<=p`1 & p`1<=1}
                         by A1,Def1;
       thus K={p: p`1=-1 & -1 <=p`2 & p`2<=1 or
          p`1=1 & -1 <=p`2 & p`2<=1 or -1=p`2 & -1<=p`1 & p`1<=1 or
          1=p`2 & -1<=p`1 & p`1<=1}
       proof thus K c=
         {p: p`1=-1 & -1 <=p`2 & p`2<=1 or
          p`1=1 & -1 <=p`2 & p`2<=1 or -1=p`2 & -1<=p`1 & p`1<=1 or
          1=p`2 & -1<=p`1 & p`1<=1}
          proof let x be set;assume x in K;
            then consider p such that
            A187: p=x & (p`1=-1 & -1 <=p`2 & p`2<=1 or
            p`2=1 & -1<=p`1 & p`1<=1 or
            p`1=1 & -1 <=p`2 & p`2<=1 or p`2=-1 & -1<=p`1 & p`1<=1)
                        by A186;
           thus x in
             {q: q`1=-1 & -1 <=q`2 & q`2<=1 or
             q`1=1 & -1 <=q`2 & q`2<=1 or -1=q`2 & -1<=q`1 & q`1<=1 or
             1=q`2 & -1<=q`1 & q`1<=1} by A187;
          end;
        thus
           {p: p`1=-1 & -1 <=p`2 & p`2<=1 or
          p`1=1 & -1 <=p`2 & p`2<=1 or -1=p`2 & -1<=p`1 & p`1<=1 or
          1=p`2 & -1<=p`1 & p`1<=1} c= K
          proof let x be set;assume x in
           {p: p`1=-1 & -1 <=p`2 & p`2<=1 or
           p`1=1 & -1 <=p`2 & p`2<=1 or -1=p`2 & -1<=p`1 & p`1<=1 or
           1=p`2 & -1<=p`1 & p`1<=1};
           then consider p such that
           A188: p=x &
             (p`1=-1 & -1 <=p`2 & p`2<=1 or
            p`1=1 & -1 <=p`2 & p`2<=1 or -1=p`2 & -1<=p`1 & p`1<=1 or
            1=p`2 & -1<=p`1 & p`1<=1);
            thus x in K by A186,A188;
          end;
       end;
         thus f.p1,f.p2,f.p3,f.p4 are_in_this_order_on P
                         by A117,A118,A119,JORDAN17:def 1;
         end;
         hence f.p1,f.p2,f.p3,f.p4 are_in_this_order_on P;
        case A189: not(p2`2<0 & p2 in LSeg(|[-1,-1]|,|[-1,1]|));
       A190: now per cases by A1,A6,Th73;
       case p2 in LSeg(|[-1,-1]|,|[-1,1]|);
        hence p2 in LSeg(|[-1,-1]|,|[-1,1]|) & (|[-1,0]|)`2<=p2`2
       or p2 in LSeg(|[-1,1]|,|[1,1]|) or p2 in LSeg(|[1,1]|,|[1,-1]|)
       or p2 in LSeg(|[1,-1]|,|[-1,-1]|) & p2<>W-min(K) by A189,EUCLID:56;
       case p2 in LSeg(|[-1,1]|,|[1,1]|);
        hence p2 in LSeg(|[-1,-1]|,|[-1,1]|) & (|[-1,0]|)`2<=p2`2
       or p2 in LSeg(|[-1,1]|,|[1,1]|) or p2 in LSeg(|[1,1]|,|[1,-1]|)
       or p2 in LSeg(|[1,-1]|,|[-1,-1]|) & p2<>W-min(K);
       case p2 in LSeg(|[1,1]|,|[1,-1]|);
        hence p2 in LSeg(|[-1,-1]|,|[-1,1]|) & (|[-1,0]|)`2<=p2`2
       or p2 in LSeg(|[-1,1]|,|[1,1]|) or p2 in LSeg(|[1,1]|,|[1,-1]|)
       or p2 in LSeg(|[1,-1]|,|[-1,-1]|) & p2<>W-min(K);
       case A191: p2 in LSeg(|[1,-1]|,|[-1,-1]|);
         A192: W-min(K)=|[-1,-1]| by A1,Th56;
           now assume A193: p2= W-min(K);
           then p2`2=-1 by A192,EUCLID:56;
          hence contradiction by A189,A192,A193,TOPREAL1:6;
         end;
        hence p2 in LSeg(|[-1,-1]|,|[-1,1]|) & (|[-1,0]|)`2<=p2`2
       or p2 in LSeg(|[-1,1]|,|[1,1]|) or p2 in LSeg(|[1,1]|,|[1,-1]|)
       or p2 in LSeg(|[1,-1]|,|[-1,-1]|) & p2<>W-min(K) by A191;
       end;

       then A194: LE |[-1,0]|,p2,K by A1,A13,Th69;
       then A195: LE f.p2,f.p3,P by A1,A5,A12,A13,Th76;
         LE |[-1,0]|,p3,K by A2,A5,A194,JORDAN6:73;
       then A196: LE f.p3,f.p4,P by A1,A5,A12,A13,Th76;
       A197: now per cases;
       case A198: p4`1=-1 & p4`2<0 & p1`2<=p4`2;
         A199: (|[-1,-1]|)`1=-1 by EUCLID:56;
         A200: (|[-1,-1]|)`2=-1 by EUCLID:56;
         A201: (|[-1,1]|)`1=-1 by EUCLID:56;
         A202: (|[-1,1]|)`2=1 by EUCLID:56;

           -1<=p4`2 & p4`2<=1 by A1,A7,Th28;
         then A203: p4 in LSeg(|[-1,-1]|,|[-1,1]|) by A198,A199,A200,A201,A202,
GOBOARD7:8;
           now per cases by A190;
         case A204: p2 in LSeg(|[-1,-1]|,|[-1,1]|) & (|[-1,0]|)`2<=p2`2;
           then 0<=p2`2 by EUCLID:56;
          hence contradiction by A1,A10,A198,A203,A204,Th65;
         case A205: p2 in LSeg(|[-1,1]|,|[1,1]|);
           then LE p4,p2,K by A1,A203,Th69;
           then A206: p2=p4 by A2,A10,JORDAN6:72;
             p2`2=1 by A205,Th11;
          hence contradiction by A198,A206;
         case A207: p2 in LSeg(|[1,1]|,|[1,-1]|);
           then LE p4,p2,K by A1,A203,Th69;
           then p2=p4 by A2,A10,JORDAN6:72;
          hence contradiction by A198,A207,Th9;
         case A208: p2 in LSeg(|[1,-1]|,|[-1,-1]|) & p2<>W-min(K);
           then LE p4,p2,K by A1,A203,Th69;
           then A209: p2=p4 by A2,A10,JORDAN6:72;
           A210: p2`2= -1 by A208,Th11;
             W-min(K)= |[-1,-1]| by A1,Th56;
           then (W-min(K))`1=-1 & (W-min(K))`2=-1 by EUCLID:56;
          hence contradiction by A198,A208,A209,A210,TOPREAL3:11;
         end;
        hence contradiction;
       case A211: not (p4`1=-1 & p4`2<0 & p1`2<=p4`2);
         A212: p4 in LSeg(|[-1,-1]|,|[-1,1]|) or p4 in LSeg(|[-1,1]|,|[1,1]|)
          or p4 in LSeg(|[1,1]|,|[1,-1]|) or p4 in LSeg(|[1,-1]|,|[-1,-1]|)
                             by A1,A7,Th73;
           now per cases by A211;
         case A213: p4`1<>-1;
               K= {p: p`1=-1 & -1 <=p`2 & p`2<=1 or p`2=1 & -1<=p`1 & p`1<=1 or
               p`1=1 & -1 <=p`2 & p`2<=1 or p`2=-1 & -1<=p`1 & p`1<=1}
                                          by A1,Def1;
             then A214: f.:K=P by A1,A4,Th45,JGRAPH_3:33;
             A215: dom f = the carrier of TOP-REAL 2 by FUNCT_2:def 1;
             A216: (f.p1)`1<0 & (f.p1)`2<=0 by A1,A15,A23,Th77;
             A217: Upper_Arc(P)
                ={p where p is Point of TOP-REAL 2:p in P & p`2>=0}
                                     by A4,JGRAPH_5:37;
             A218: f.p1 in P by A6,A214,A215,FUNCT_1:def 12;
               Lower_Arc(P)
                ={p where p is Point of TOP-REAL 2:p in P & p`2<=0}
                                     by A4,JGRAPH_5:38;
             then A219: f.p1 in Lower_Arc(P) by A216,A218;
             A220: f.(|[-1,0]|)=W-min(P) by A1,A4,Th18,JGRAPH_5:32;
             A221: now assume f.p1=W-min(P);
               then p1=|[-1,0]| by A1,A215,A220,FUNCT_1:def 8;
              hence contradiction by A23,EUCLID:56;
             end;
             now per cases by A212,A213,Th9;
           case p4 in LSeg(|[-1,1]|,|[1,1]|);
             then A222: p4`2=1 by Th11;
             A223: f.p4 in P by A7,A214,A215,FUNCT_1:def 12;
               (f.p4)`2>=0 by A1,A222,Th79;
              then f.p4 in Upper_Arc(P) by A217,A223;
            hence LE f.p4,f.p1,P by A219,A221,JORDAN6:def 10;
           case p4 in LSeg(|[1,1]|,|[1,-1]|);
             then A224: p4`1=1 by Th9;
               K= {p: p`1=-1 & -1 <=p`2 & p`2<=1 or p`2=1 & -1<=p`1 & p`1<=1 or
               p`1=1 & -1 <=p`2 & p`2<=1 or p`2=-1 & -1<=p`1 & p`1<=1}
                                          by A1,Def1;
             then A225: f.:K=P by A1,A4,Th45,JGRAPH_3:33;
             A226: dom f = the carrier of TOP-REAL 2 by FUNCT_2:def 1;
             then A227: f.p4 in P by A7,A225,FUNCT_1:def 12;
             A228: f.p1 in P by A6,A225,A226,FUNCT_1:def 12;
             A229: (f.p1)`1<0 & (f.p1)`2<=0 by A1,A15,A23,Th77;
             A230: f.(|[-1,0]|)=W-min(P) by A1,A4,Th18,JGRAPH_5:32;
             A231: now assume f.p1=W-min(P);
               then p1=|[-1,0]| by A1,A226,A230,FUNCT_1:def 8;
              hence contradiction by A23,EUCLID:56;
             end;
             A232: (f.p4)`1>=0 by A1,A224,Th78;
               now per cases;
             case A233: (f.p4)`2>=0;
             A234: (f.p1)`1<0 & (f.p1)`2<=0 by A1,A15,A23,Th77;
             A235: Upper_Arc(P)
                ={p where p is Point of TOP-REAL 2:p in P & p`2>=0}
                                     by A4,JGRAPH_5:37;
             A236: f.(|[-1,0]|)=W-min(P) by A1,A4,Th18,JGRAPH_5:32;
             A237: now assume f.p1=W-min(P);
               then p1=|[-1,0]| by A1,A226,A236,FUNCT_1:def 8;
              hence contradiction by A23,EUCLID:56;
             end;
             A238: f.p4 in Upper_Arc(P) by A227,A233,A235;
               Lower_Arc(P)
                ={p where p is Point of TOP-REAL 2:p in P & p`2<=0}
                                     by A4,JGRAPH_5:38;
             then f.p1 in Lower_Arc(P) by A228,A234;
            hence LE f.p4,f.p1,P by A237,A238,JORDAN6:def 10;
            case A239: (f.p4)`2<0;
               (f.p1)`1<=0 &
             (f.p4)`1>=(f.p1)`1 by A229,A232,AXIOMS:22;
            hence LE f.p4,f.p1,P
                         by A4,A227,A228,A229,A231,A239,JGRAPH_5:59;
            end;
            hence LE f.p4,f.p1,P;
           case A240: p4 in LSeg(|[1,-1]|,|[-1,-1]|);
             then p4`2=-1 & -1 <=p4`1 & p4`1<=1 by Th11;
             then A241: (f.p4)`2<0 by A1,Th79;
             A242: f.p4 in P by A7,A214,A215,FUNCT_1:def 12;

             A243: (f.p1)`1<0 & (f.p1)`2<=0 by A1,A15,A23,Th77;
               (f.p4)`1>=(f.p1)`1 by A1,A14,A240,Th80;
            hence LE f.p4,f.p1,P by A4,A218,A221,A241,A242,A243,JGRAPH_5:59;
           end;
          hence LE f.p4,f.p1,P;
         case A244: p4`1=-1 & p4`2>=0;
               K= {p: p`1=-1 & -1 <=p`2 & p`2<=1 or p`2=1 & -1<=p`1 & p`1<=1 or
               p`1=1 & -1 <=p`2 & p`2<=1 or p`2=-1 & -1<=p`1 & p`1<=1}
                                          by A1,Def1;
             then A245: f.:K=P by A1,A4,Th45,JGRAPH_3:33;
             A246: dom f = the carrier of TOP-REAL 2 by FUNCT_2:def 1;
             then A247: f.p4 in P by A7,A245,FUNCT_1:def 12;
             A248: f.p1 in P by A6,A245,A246,FUNCT_1:def 12;
             A249: (f.p4)`2>=0 by A1,A244,Th79;
             A250: (f.p1)`1<0 & (f.p1)`2<=0 by A1,A15,A23,Th77;
             A251: Upper_Arc(P)
                ={p where p is Point of TOP-REAL 2:p in P & p`2>=0}
                                     by A4,JGRAPH_5:37;
             A252: f.(|[-1,0]|)=W-min(P) by A1,A4,Th18,JGRAPH_5:32;
             A253: now assume f.p1=W-min(P);
               then p1=|[-1,0]| by A1,A246,A252,FUNCT_1:def 8;
              hence contradiction by A23,EUCLID:56;
             end;
             A254: f.p4 in Upper_Arc(P) by A247,A249,A251;
               Lower_Arc(P)
                ={p where p is Point of TOP-REAL 2:p in P & p`2<=0}
                                     by A4,JGRAPH_5:38;
             then f.p1 in Lower_Arc(P) by A248,A250;
            hence LE f.p4,f.p1,P by A253,A254,JORDAN6:def 10;
         case A255: p4`1=-1 & p4`2<0 & p1`2>p4`2;
             then A256: p4 in LSeg(|[-1,-1]|,|[-1,1]|) by A8,Th10;
               K= {p: p`1=-1 & -1 <=p`2 & p`2<=1 or p`2=1 & -1<=p`1 & p`1<=1 or
               p`1=1 & -1 <=p`2 & p`2<=1 or p`2=-1 & -1<=p`1 & p`1<=1}
                                          by A1,Def1;
             then A257: f.:K=P by A1,A4,Th45,JGRAPH_3:33;
             A258: dom f = the carrier of TOP-REAL 2 by FUNCT_2:def 1;
             then A259: f.p4 in P by A7,A257,FUNCT_1:def 12;
             A260: f.p1 in P by A6,A257,A258,FUNCT_1:def 12;

           A261: (f.p1)`1<0 & (f.p1)`2<0 by A1,A15,A23,Th77;
           A262: (f.p4)`2<=(f.p1)`2 by A1,A14,A23,A255,A256,Th81;
           A263: (f.p4)`1<0 by A1,A255,Th78;
             (f.p4)`2<0 by A1,A255,Th79;
          hence LE f.p4,f.p1,P by A4,A259,A260,A261,A262,A263,JGRAPH_5:54;
         end;
        hence LE f.p4,f.p1,P;
       end;
       A264: K={p: p`1=-1 & -1 <=p`2 & p`2<=1 or p`2=1 & -1<=p`1 & p`1<=1 or
          p`1=1 & -1 <=p`2 & p`2<=1 or p`2=-1 & -1<=p`1 & p`1<=1}
                         by A1,Def1;
       thus K={p: p`1=-1 & -1 <=p`2 & p`2<=1 or
          p`1=1 & -1 <=p`2 & p`2<=1 or -1=p`2 & -1<=p`1 & p`1<=1 or
          1=p`2 & -1<=p`1 & p`1<=1}
       proof thus K c=
         {p: p`1=-1 & -1 <=p`2 & p`2<=1 or
          p`1=1 & -1 <=p`2 & p`2<=1 or -1=p`2 & -1<=p`1 & p`1<=1 or
          1=p`2 & -1<=p`1 & p`1<=1}
          proof let x be set;assume x in K;
            then consider p such that
            A265: p=x & (p`1=-1 & -1 <=p`2 & p`2<=1 or
            p`2=1 & -1<=p`1 & p`1<=1 or
            p`1=1 & -1 <=p`2 & p`2<=1 or p`2=-1 & -1<=p`1 & p`1<=1)
                        by A264;
           thus x in
             {q: q`1=-1 & -1 <=q`2 & q`2<=1 or
             q`1=1 & -1 <=q`2 & q`2<=1 or -1=q`2 & -1<=q`1 & q`1<=1 or
             1=q`2 & -1<=q`1 & q`1<=1} by A265;
          end;
        thus
           {p: p`1=-1 & -1 <=p`2 & p`2<=1 or
          p`1=1 & -1 <=p`2 & p`2<=1 or -1=p`2 & -1<=p`1 & p`1<=1 or
          1=p`2 & -1<=p`1 & p`1<=1} c= K
          proof let x be set;assume x in
           {p: p`1=-1 & -1 <=p`2 & p`2<=1 or
           p`1=1 & -1 <=p`2 & p`2<=1 or -1=p`2 & -1<=p`1 & p`1<=1 or
           1=p`2 & -1<=p`1 & p`1<=1};
           then consider p such that
           A266: p=x &
             (p`1=-1 & -1 <=p`2 & p`2<=1 or
            p`1=1 & -1 <=p`2 & p`2<=1 or -1=p`2 & -1<=p`1 & p`1<=1 or
            1=p`2 & -1<=p`1 & p`1<=1);
            thus x in K by A264,A266;
          end;
       end;
         thus f.p1,f.p2,f.p3,f.p4 are_in_this_order_on P
                         by A195,A196,A197,JORDAN17:def 1;
        end;
       hence f.p1,f.p2,f.p3,f.p4 are_in_this_order_on P;
      end;
      hence f.p1,f.p2,f.p3,f.p4 are_in_this_order_on P;
     case A267: p1 in LSeg(|[-1,1]|,|[1,1]|);
       A268: |[-1,1]| in LSeg(|[-1,1]|,|[1,1]|) by TOPREAL1:6;
       A269: |[-1,1]| in LSeg(|[-1,-1]|,|[-1,1]|) by TOPREAL1:6;
       A270: (|[-1,1]|)`1=-1 & (|[-1,1]|)`2=1 by EUCLID:56;
         p1`2=1 & -1 <=p1`1 & p1`1<=1 by A267,Th11;
       then A271: LE |[-1,1]|,p1,K by A1,A267,A268,A270,Th70;
       then A272: LE f.p1,f.p2,P by A1,A5,A269,A270,Th76;
       A273: LE |[-1,1]|,p2,K by A2,A5,A271,JORDAN6:73;
       then A274: LE f.p2,f.p3,P by A1,A5,A269,A270,Th76;
         LE |[-1,1]|,p3,K by A2,A5,A273,JORDAN6:73;
       then LE f.p3,f.p4,P by A1,A5,A269,A270,Th76;
      hence f.p1,f.p2,f.p3,f.p4 are_in_this_order_on P
                               by A272,A274,JORDAN17:def 1;
     case A275: p1 in LSeg(|[1,1]|,|[1,-1]|);
       A276: |[-1,1]| in LSeg(|[-1,1]|,|[1,1]|) by TOPREAL1:6;
       A277: |[-1,1]| in LSeg(|[-1,-1]|,|[-1,1]|) by TOPREAL1:6;
       A278: (|[-1,1]|)`1=-1 & (|[-1,1]|)`2=1 by EUCLID:56;
       A279: |[1,1]| in LSeg(|[1,1]|,|[1,-1]|) by TOPREAL1:6;
       A280: |[1,1]| in LSeg(|[-1,1]|,|[1,1]|) by TOPREAL1:6;
       A281: (|[1,1]|)`1=1 & (|[1,1]|)`2=1 by EUCLID:56;
       then A282: LE |[-1,1]|,|[1,1]|,K by A1,A276,A278,A280,Th70;
         p1`1=1 & -1 <=p1`2 & p1`2<=1 by A275,Th9;
       then LE |[1,1]|,p1,K by A1,A275,A279,A281,Th71;
       then A283: LE |[-1,1]|,p1,K by A2,A282,JORDAN6:73;
       then A284: LE f.p1,f.p2,P by A1,A5,A277,A278,Th76;
       A285: LE |[-1,1]|,p2,K by A2,A5,A283,JORDAN6:73;
       then A286: LE f.p2,f.p3,P by A1,A5,A277,A278,Th76;
         LE |[-1,1]|,p3,K by A2,A5,A285,JORDAN6:73;
       then LE f.p3,f.p4,P by A1,A5,A277,A278,Th76;
      hence f.p1,f.p2,f.p3,f.p4 are_in_this_order_on P
                       by A284,A286,JORDAN17:def 1;
     case A287: p1 in LSeg(|[1,-1]|,|[-1,-1]|) & p1 <> W-min(K);
       A288: |[-1,1]| in LSeg(|[-1,1]|,|[1,1]|) by TOPREAL1:6;
       A289: |[-1,1]| in LSeg(|[-1,-1]|,|[-1,1]|) by TOPREAL1:6;
       A290: (|[-1,1]|)`1=-1 & (|[-1,1]|)`2=1 by EUCLID:56;
       A291: |[1,1]| in LSeg(|[-1,1]|,|[1,1]|) by TOPREAL1:6;
         (|[1,1]|)`1=1 & (|[1,1]|)`2=1 by EUCLID:56;
       then A292: LE |[-1,1]|,|[1,1]|,K by A1,A288,A290,A291,Th70;
       A293: |[1,-1]| in LSeg(|[1,1]|,|[1,-1]|) by TOPREAL1:6;
       A294: |[1,-1]| in LSeg(|[1,-1]|,|[-1,-1]|) by TOPREAL1:6;
       A295: (|[1,-1]|)`1=1 & (|[1,-1]|)`2= -1 by EUCLID:56;
         LE |[1,1]|,|[1,-1]|,K by A1,A291,A293,Th70;
       then A296: LE |[-1,1]|,|[1,-1]|,K by A2,A292,JORDAN6:73;
         W-min(K)=|[-1,-1]| by A1,Th56;
       then (W-min(K))`1=-1 & (W-min(K))`2=-1 by EUCLID:56;
       then A297: |[1,-1]| <> W-min(K) by EUCLID:56;
         p1`2=-1 & -1 <=p1`1 & p1`1<=1 by A287,Th11;
       then LE |[1,-1]|,p1,K by A1,A287,A294,A295,A297,Th72;
       then A298: LE |[-1,1]|,p1,K by A2,A296,JORDAN6:73;
       then A299: LE f.p1,f.p2,P by A1,A5,A289,A290,Th76;
       A300: LE |[-1,1]|,p2,K by A2,A5,A298,JORDAN6:73;
       then A301: LE f.p2,f.p3,P by A1,A5,A289,A290,Th76;
         LE |[-1,1]|,p3,K by A2,A5,A300,JORDAN6:73;
       then LE f.p3,f.p4,P by A1,A5,A289,A290,Th76;
      hence f.p1,f.p2,f.p3,f.p4 are_in_this_order_on P
                       by A299,A301,JORDAN17:def 1;
     end;
    hence f.p1,f.p2,f.p3,f.p4 are_in_this_order_on P;
   end;
end;

theorem Th83: for p1,p2 being Point of TOP-REAL 2,
P being non empty compact Subset of TOP-REAL 2 st
P is_simple_closed_curve & p1 in P & p2 in P & not LE p1,p2,P
holds LE p2,p1,P
proof let p1,p2 be Point of TOP-REAL 2,
  P be non empty compact Subset of TOP-REAL 2;
  assume A1: P is_simple_closed_curve & p1 in P & p2 in P
       & not LE p1,p2,P;
  then A2: P=Upper_Arc(P) \/ Lower_Arc(P) by JORDAN6:def 9;
  A3: not p1=W-min(P) by A1,JORDAN7:3;
  per cases by A1,A2,XBOOLE_0:def 2;
  suppose A4: p1 in Upper_Arc(P) & p2 in Upper_Arc(P);
    A5: Upper_Arc(P) is_an_arc_of W-min(P),E-max(P) by A1,JORDAN6:def 8;
    set q1=W-min(P),q2=E-max(P);
    set Q= Upper_Arc(P);
      now per cases;
    case A6: p1<>p2;
        now per cases by A4,A5,A6,JORDAN5C:14;
      case LE p1,p2,Q,q1,q2 & not LE p2,p1,Q,q1,q2;
       hence contradiction by A1,A4,JORDAN6:def 10;
      case LE p2,p1,Q,q1,q2 & not LE p1,p2,Q,q1,q2;
       hence LE p2,p1,P by A4,JORDAN6:def 10;
      end;
     hence LE p2,p1,P;
    case p1=p2;
     hence LE p2,p1,P by A1,JORDAN6:71;
    end;
   hence LE p2,p1,P;
  suppose A7: p1 in Upper_Arc(P) & p2 in Lower_Arc(P);
     now per cases;
   case p2=W-min(P);
    hence LE p2,p1,P by A1,JORDAN7:3;
   case p2<>W-min(P);
    hence contradiction by A1,A7,JORDAN6:def 10;
   end;
   hence LE p2,p1,P;
  suppose p1 in Lower_Arc(P) & p2 in Upper_Arc(P);
   hence LE p2,p1,P by A3,JORDAN6:def 10;
  suppose A8: p1 in Lower_Arc(P) & p2 in Lower_Arc(P);
    A9: Lower_Arc(P) is_an_arc_of E-max(P),W-min(P) by A1,JORDAN6:65;
    set q2=W-min(P),q1=E-max(P);
    set Q= Lower_Arc(P);
      now per cases;
    case A10: p1<>p2;
        now per cases by A8,A9,A10,JORDAN5C:14;
      case A11: LE p1,p2,Q,q1,q2 & not LE p2,p1,Q,q1,q2;
          now per cases;
        case p2=W-min(P);
         hence LE p2,p1,P by A1,JORDAN7:3;
        case p2<>W-min(P);
         hence contradiction by A1,A8,A11,JORDAN6:def 10;
        end;
       hence LE p2,p1,P;
      case LE p2,p1,Q,q1,q2 & not LE p1,p2,Q,q1,q2;
       hence LE p2,p1,P by A3,A8,JORDAN6:def 10;
      end;
     hence LE p2,p1,P;
    case p1=p2;
     hence LE p2,p1,P by A1,JORDAN6:71;
    end;
   hence LE p2,p1,P;
end;

theorem for p1,p2,p3 being Point of TOP-REAL 2,
P being non empty compact Subset of TOP-REAL 2 st
P is_simple_closed_curve & p1 in P & p2 in P & p3 in P holds
  LE p1,p2,P & LE p2,p3,P or LE p1,p3,P & LE p3,p2,P or
  LE p2,p1,P & LE p1,p3,P or LE p2,p3,P & LE p3,p1,P or
  LE p3,p1,P & LE p1,p2,P or LE p3,p2,P & LE p2,p1,P by Th83;

theorem for p1,p2,p3 being Point of TOP-REAL 2,
P being non empty compact Subset of TOP-REAL 2 st
P is_simple_closed_curve & p1 in P & p2 in P & p3 in P &
LE p2,p3,P holds
LE p1,p2,P or LE p2,p1,P & LE p1,p3,P or LE p3,p1,P by Th83;

theorem for p1,p2,p3,p4 being Point of TOP-REAL 2,
P being non empty compact Subset of TOP-REAL 2 st
P is_simple_closed_curve & p1 in P & p2 in P & p3 in P & p4 in P &
LE p2,p3,P & LE p3,p4,P holds
LE p1,p2,P or LE p2,p1,P & LE p1,p3,P or LE p3,p1,P & LE p1,p4,P or
LE p4,p1,P by Th83;

theorem Th87: for p1,p2,p3,p4 being Point of TOP-REAL 2,
 P,K being non empty compact Subset of TOP-REAL 2,
 f being map of TOP-REAL 2,TOP-REAL 2 st
 P= circle(0,0,1) & K=rectangle(-1,1,-1,1) & f=Sq_Circ &
 LE f.p1,f.p2,P & LE f.p2,f.p3,P & LE f.p3,f.p4,P
 holds p1,p2,p3,p4 are_in_this_order_on K
proof let p1,p2,p3,p4 be Point of TOP-REAL 2,
  P,K be non empty compact Subset of TOP-REAL 2,
  f be map of TOP-REAL 2,TOP-REAL 2;
  assume A1: P= circle(0,0,1) & K=rectangle(-1,1,-1,1) & f=Sq_Circ &
   LE f.p1,f.p2,P & LE f.p2,f.p3,P & LE f.p3,f.p4,P;
   then A2: K is_simple_closed_curve by Th60;
    A3: P={p: |.p.|=1} by A1,Th33;
    then A4: P is_simple_closed_curve by JGRAPH_3:36;
   then A5: LE f.p1,f.p3,P by A1,JORDAN6:73;
   A6: LE f.p2,f.p4,P by A1,A4,JORDAN6:73;
     K= {p: p`1=-1 & -1 <=p`2 & p`2<=1 or p`2=1 & -1<=p`1 & p`1<=1 or
      p`1=1 & -1 <=p`2 & p`2<=1 or p`2=-1 & -1<=p`1 & p`1<=1} by A1,Def1;
    then A7: f.:K=P by A1,A3,Th45,JGRAPH_3:33;
    A8: dom f = the carrier of TOP-REAL 2 by FUNCT_2:def 1;
     P={p: |.p.|=1} by A1,Th33;
   then A9: P is_simple_closed_curve by JGRAPH_3:36;
   then f.p1 in P by A1,JORDAN7:5;
   then consider x1 being set such that
   A10: x1 in dom f & x1 in K & f.p1=f.x1 by A7,FUNCT_1:def 12;
   A11: p1 in K by A1,A8,A10,FUNCT_1:def 8;
     f.p2 in P by A1,A9,JORDAN7:5;
   then consider x2 being set such that
   A12: x2 in dom f & x2 in K & f.p2=f.x2 by A7,FUNCT_1:def 12;
   A13: p2 in K by A1,A8,A12,FUNCT_1:def 8;
     f.p3 in P by A1,A9,JORDAN7:5;
   then consider x3 being set such that
   A14: x3 in dom f & x3 in K & f.p3=f.x3 by A7,FUNCT_1:def 12;
   A15: p3 in K by A1,A8,A14,FUNCT_1:def 8;
     f.p4 in P by A1,A9,JORDAN7:5;
   then consider x4 being set such that
   A16: x4 in dom f & x4 in K & f.p4=f.x4 by A7,FUNCT_1:def 12;
   A17: p4 in K by A1,A8,A16,FUNCT_1:def 8;
     now assume A18: not p1,p2,p3,p4 are_in_this_order_on K;
     A19: now assume A20: p1,p2,p4,p3 are_in_this_order_on K;
         now per cases by A20,JORDAN17:def 1;
       case A21: LE p1,p2,K & LE p2,p4,K & LE p4,p3,K;
         then f.p1,f.p2,f.p4,f.p3 are_in_this_order_on P by A1,Th82;
         then LE f.p1,f.p2,P & LE f.p2,f.p4,P & LE f.p4,f.p3,P or
         LE f.p2,f.p4,P & LE f.p4,f.p3,P & LE f.p3,f.p1,P or
         LE f.p4,f.p3,P & LE f.p3,f.p1,P & LE f.p1,f.p2,P or
         LE f.p3,f.p1,P & LE f.p1,f.p2,P & LE f.p2,f.p4,P
                       by JORDAN17:def 1;
         then f.p3=f.p4 or f.p3=f.p1 by A1,A5,A9,JORDAN6:72;
         then A22: p3=p4 or p3=p1 by A1,A8,FUNCT_1:def 8;
           LE p1,p4,K by A2,A21,JORDAN6:73;
         then p3=p1 & p1=p4 by A2,A18,A20,A21,A22,JORDAN6:72;
        hence contradiction by A18,A20;
       case A23: LE p2,p4,K & LE p4,p3,K & LE p3,p1,K;
         then f.p2,f.p4,f.p3,f.p1 are_in_this_order_on P by A1,Th82;
         then LE f.p2,f.p4,P & LE f.p4,f.p3,P & LE f.p3,f.p1,P or
         LE f.p4,f.p3,P & LE f.p3,f.p1,P & LE f.p1,f.p2,P or
         LE f.p3,f.p1,P & LE f.p1,f.p2,P & LE f.p2,f.p4,P or
         LE f.p1,f.p2,P & LE f.p2,f.p4,P & LE f.p4,f.p3,P
                       by JORDAN17:def 1;
         then f.p3=f.p4 or LE f.p3,f.p2,P by A1,A9,JORDAN6:72,73;
         then f.p3=f.p4 or f.p3=f.p2 by A1,A9,JORDAN6:72;
         then p3=p4 or p3=p2 by A1,A8,FUNCT_1:def 8;
         then p3=p2 & p4=p2 by A2,A18,A20,A23,JORDAN6:72;
        hence contradiction by A18,A20;
       case A24: LE p4,p3,K & LE p3,p1,K & LE p1,p2,K;
         then f.p4,f.p3,f.p1,f.p2 are_in_this_order_on P by A1,Th82;
         then LE f.p4,f.p3,P & LE f.p3,f.p1,P & LE f.p1,f.p2,P or
         LE f.p3,f.p1,P & LE f.p1,f.p2,P & LE f.p2,f.p4,P or
         LE f.p1,f.p2,P & LE f.p2,f.p4,P & LE f.p4,f.p3,P or
         LE f.p2,f.p4,P & LE f.p4,f.p3,P & LE f.p3,f.p1,P
                       by JORDAN17:def 1;
         then f.p3=f.p4 or LE f.p3,f.p2,P by A1,A9,JORDAN6:72,73;
         then f.p3=f.p4 or f.p3=f.p2 by A1,A9,JORDAN6:72;
         then p3=p4 or p3=p2 by A1,A8,FUNCT_1:def 8;
         then p3=p2 & p3=p1 by A2,A18,A20,A24,JORDAN6:72;
        hence contradiction by A2,A18,A20,JORDAN17:12;
       case A25: LE p3,p1,K & LE p1,p2,K & LE p2,p4,K;
         then f.p3,f.p1,f.p2,f.p4 are_in_this_order_on P by A1,Th82;
         then LE f.p4,f.p3,P & LE f.p3,f.p1,P & LE f.p1,f.p2,P or
         LE f.p3,f.p1,P & LE f.p1,f.p2,P & LE f.p2,f.p4,P or
         LE f.p1,f.p2,P & LE f.p2,f.p4,P & LE f.p4,f.p3,P or
         LE f.p2,f.p4,P & LE f.p4,f.p3,P & LE f.p3,f.p1,P
                       by JORDAN17:def 1;
         then f.p3=f.p4 or LE f.p3,f.p2,P by A1,A9,JORDAN6:72,73;
         then f.p3=f.p4 or f.p3=f.p2 by A1,A9,JORDAN6:72;
         then p3=p4 or p3=p2 by A1,A8,FUNCT_1:def 8;
         then p3=p2 & p3=p1 by A2,A18,A20,A25,JORDAN6:72;
        hence contradiction by A2,A18,A20,JORDAN17:12;
       end;
      hence contradiction;
     end;
     A26: now assume A27: p1,p3,p4,p2 are_in_this_order_on K;
         now per cases by A27,JORDAN17:def 1;
       case LE p1,p3,K & LE p3,p4,K & LE p4,p2,K;
         then f.p1,f.p3,f.p4,f.p2 are_in_this_order_on P by A1,Th82;
         then LE f.p1,f.p3,P & LE f.p3,f.p4,P & LE f.p4,f.p2,P or
         LE f.p3,f.p4,P & LE f.p4,f.p2,P & LE f.p2,f.p1,P or
         LE f.p4,f.p2,P & LE f.p2,f.p1,P & LE f.p1,f.p3,P or
         LE f.p2,f.p1,P & LE f.p1,f.p3,P & LE f.p3,f.p4,P
                       by JORDAN17:def 1;
         then f.p4=f.p2 or f.p2=f.p1 by A1,A6,A9,JORDAN6:72;
         then A28: p4=p2 or p2=p1 by A1,A8,FUNCT_1:def 8;
         then f.p3=f.p2 or f.p4=f.p1 by A1,A2,A9,A18,A27,JORDAN17:12,JORDAN6:72
;
         then p3=p2 or p4=p1 by A1,A8,FUNCT_1:def 8;
        hence contradiction by A2,A18,A27,A28,JORDAN17:12;
       case LE p3,p4,K & LE p4,p2,K & LE p2,p1,K;
         then f.p3,f.p4,f.p2,f.p1 are_in_this_order_on P by A1,Th82;
         then A29: LE f.p1,f.p3,P & LE f.p3,f.p4,P & LE f.p4,f.p2,P or
         LE f.p3,f.p4,P & LE f.p4,f.p2,P & LE f.p2,f.p1,P or
         LE f.p4,f.p2,P & LE f.p2,f.p1,P & LE f.p1,f.p3,P or
         LE f.p2,f.p1,P & LE f.p1,f.p3,P & LE f.p3,f.p4,P
                       by JORDAN17:def 1;
         then f.p4=f.p2 or f.p2=f.p1 by A1,A6,A9,JORDAN6:72;
         then A30: p4=p2 or p2=p1 by A1,A8,FUNCT_1:def 8;
           f.p2=f.p1 or LE f.p3,f.p2,P by A1,A9,A29,JORDAN6:72,73;
         then f.p2=f.p1 or f.p3=f.p2 by A1,A9,JORDAN6:72;
         then p2=p1 or p3=p2 by A1,A8,FUNCT_1:def 8;
        hence contradiction by A2,A18,A27,A30,JORDAN17:12;
       case LE p4,p2,K & LE p2,p1,K & LE p1,p3,K;
         then f.p4,f.p2,f.p1,f.p3 are_in_this_order_on P by A1,Th82;
         then A31: LE f.p1,f.p3,P & LE f.p3,f.p4,P & LE f.p4,f.p2,P or
         LE f.p3,f.p4,P & LE f.p4,f.p2,P & LE f.p2,f.p1,P or
         LE f.p4,f.p2,P & LE f.p2,f.p1,P & LE f.p1,f.p3,P or
         LE f.p2,f.p1,P & LE f.p1,f.p3,P & LE f.p3,f.p4,P
                       by JORDAN17:def 1;
         then f.p4=f.p2 or f.p2=f.p1 by A1,A6,A9,JORDAN6:72;
         then A32: p4=p2 or p2=p1 by A1,A8,FUNCT_1:def 8;
           f.p2=f.p1 or LE f.p3,f.p2,P by A1,A9,A31,JORDAN6:72,73;
         then f.p2=f.p1 or f.p3=f.p2 by A1,A9,JORDAN6:72;
         then p2=p1 or p3=p2 by A1,A8,FUNCT_1:def 8;
        hence contradiction by A2,A18,A27,A32,JORDAN17:12;
       case A33: LE p2,p1,K & LE p1,p3,K & LE p3,p4,K;
         then f.p2,f.p1,f.p3,f.p4 are_in_this_order_on P by A1,Th82;
         then LE f.p1,f.p3,P & LE f.p3,f.p4,P & LE f.p4,f.p2,P or
         LE f.p3,f.p4,P & LE f.p4,f.p2,P & LE f.p2,f.p1,P or
         LE f.p4,f.p2,P & LE f.p2,f.p1,P & LE f.p1,f.p3,P or
         LE f.p2,f.p1,P & LE f.p1,f.p3,P & LE f.p3,f.p4,P
                       by JORDAN17:def 1;
         then f.p4=f.p2 or f.p2=f.p1 by A1,A6,A9,JORDAN6:72;
         then A34: p4=p2 or p2=p1 by A1,A8,FUNCT_1:def 8;
         then LE p2,p3,K & LE p3,p2,K by A2,A18,A27,A33,JORDAN17:12,JORDAN6:73
;
         then p2=p3 by A2,JORDAN6:72;
        hence contradiction by A2,A18,A27,A34,JORDAN17:12;
       end;
      hence contradiction;
     end;
       now per cases by A2,A11,A13,A15,A17,A18,JORDAN17:27;
     case p1,p2,p4,p3 are_in_this_order_on K;
      hence contradiction by A19;
     case A35: p1,p3,p2,p4 are_in_this_order_on K;
         now per cases by A35,JORDAN17:def 1;
       case A36: LE p1,p3,K & LE p3,p2,K & LE p2,p4,K;
         then f.p1,f.p3,f.p2,f.p4 are_in_this_order_on P by A1,Th82;
         then LE f.p1,f.p3,P & LE f.p3,f.p2,P & LE f.p2,f.p4,P or
         LE f.p3,f.p2,P & LE f.p2,f.p4,P & LE f.p4,f.p1,P or
         LE f.p2,f.p4,P & LE f.p4,f.p1,P & LE f.p1,f.p3,P or
         LE f.p4,f.p1,P & LE f.p1,f.p3,P & LE f.p3,f.p2,P
                       by JORDAN17:def 1;
         then f.p3=f.p2 or LE f.p2,f.p1,P by A1,A9,JORDAN6:72,73;
         then f.p3=f.p2 or f.p2=f.p1 by A1,A9,JORDAN6:72;
         then p3=p2 or p2=p1 by A1,A8,FUNCT_1:def 8;
         then p2=p1 & p3=p1 by A2,A18,A35,A36,JORDAN6:72;
        hence contradiction by A18,A35;
       case A37: LE p3,p2,K & LE p2,p4,K & LE p4,p1,K;
         then f.p3,f.p2,f.p4,f.p1 are_in_this_order_on P by A1,Th82;
         then LE f.p1,f.p3,P & LE f.p3,f.p2,P & LE f.p2,f.p4,P or
         LE f.p3,f.p2,P & LE f.p2,f.p4,P & LE f.p4,f.p1,P or
         LE f.p2,f.p4,P & LE f.p4,f.p1,P & LE f.p1,f.p3,P or
         LE f.p4,f.p1,P & LE f.p1,f.p3,P & LE f.p3,f.p2,P
                       by JORDAN17:def 1;
         then f.p3=f.p2 or LE f.p2,f.p1,P by A1,A9,JORDAN6:72,73;
         then f.p3=f.p2 or f.p2=f.p1 by A1,A9,JORDAN6:72;
         then p3=p2 or p2=p1 by A1,A8,FUNCT_1:def 8;
         then p2=p1 & p4=p1 by A2,A18,A35,A37,JORDAN6:72;
        hence contradiction by A2,A18,A35,JORDAN17:12;
       case A38: LE p2,p4,K & LE p4,p1,K & LE p1,p3,K;
         then f.p2,f.p4,f.p1,f.p3 are_in_this_order_on P by A1,Th82;
         then LE f.p1,f.p3,P & LE f.p3,f.p2,P & LE f.p2,f.p4,P or
         LE f.p3,f.p2,P & LE f.p2,f.p4,P & LE f.p4,f.p1,P or
         LE f.p2,f.p4,P & LE f.p4,f.p1,P & LE f.p1,f.p3,P or
         LE f.p4,f.p1,P & LE f.p1,f.p3,P & LE f.p3,f.p2,P
                       by JORDAN17:def 1;
         then f.p3=f.p2 or LE f.p2,f.p1,P by A1,A9,JORDAN6:72,73;
         then f.p3=f.p2 or f.p2=f.p1 by A1,A9,JORDAN6:72;
         then p3=p2 or p2=p1 by A1,A8,FUNCT_1:def 8;
         then p2=p1 & p4=p1 by A2,A18,A35,A38,JORDAN6:72;
        hence contradiction by A2,A18,A35,JORDAN17:12;
       case A39: LE p4,p1,K & LE p1,p3,K & LE p3,p2,K;
         then f.p4,f.p1,f.p3,f.p2 are_in_this_order_on P by A1,Th82;
         then LE f.p1,f.p3,P & LE f.p3,f.p2,P & LE f.p2,f.p4,P or
         LE f.p3,f.p2,P & LE f.p2,f.p4,P & LE f.p4,f.p1,P or
         LE f.p2,f.p4,P & LE f.p4,f.p1,P & LE f.p1,f.p3,P or
         LE f.p4,f.p1,P & LE f.p1,f.p3,P & LE f.p3,f.p2,P
                       by JORDAN17:def 1;
         then f.p3=f.p2 or LE f.p2,f.p1,P by A1,A9,JORDAN6:72,73;
         then f.p3=f.p2 or f.p2=f.p1 by A1,A9,JORDAN6:72;
         then p3=p2 or p2=p1 by A1,A8,FUNCT_1:def 8;
         then p2=p1 & p3=p1 by A2,A18,A35,A39,JORDAN6:72;
        hence contradiction by A18,A35;
       end;
      hence contradiction;
     case p1,p3,p4,p2 are_in_this_order_on K;
       hence contradiction by A26;
     case A40: p1,p4,p2,p3 are_in_this_order_on K;
         now per cases by A40,JORDAN17:def 1;
       case A41: LE p1,p4,K & LE p4,p2,K & LE p2,p3,K;
         then f.p1,f.p4,f.p2,f.p3 are_in_this_order_on P by A1,Th82;
         then A42: LE f.p1,f.p4,P & LE f.p4,f.p2,P & LE f.p2,f.p3,P or
         LE f.p4,f.p2,P & LE f.p2,f.p3,P & LE f.p3,f.p1,P or
         LE f.p2,f.p3,P & LE f.p3,f.p1,P & LE f.p1,f.p4,P or
         LE f.p3,f.p1,P & LE f.p1,f.p4,P & LE f.p4,f.p2,P
                       by JORDAN17:def 1;
         then f.p4=f.p2 or LE f.p2,f.p1,P by A6,A9,JORDAN6:72,73;
         then f.p4=f.p2 or f.p2=f.p1 by A1,A9,JORDAN6:72;
         then A43: p4=p2 or p2=p1 by A1,A8,FUNCT_1:def 8;
         then A44: p4=p2 by A2,A41,JORDAN6:72;
           f.p3=f.p1 or LE f.p4,f.p3,P by A5,A9,A42,JORDAN6:72,73;
         then f.p3=f.p1 or f.p4=f.p3 by A1,A9,JORDAN6:72;
         then A45: p3=p1 or p4=p3 by A1,A8,FUNCT_1:def 8;
         then p1=p2 by A2,A18,A40,A41,A43,JORDAN6:72;
        hence contradiction by A18,A40,A44,A45;
       case A46: LE p4,p2,K & LE p2,p3,K & LE p3,p1,K;
         then f.p4,f.p2,f.p3,f.p1 are_in_this_order_on P by A1,Th82;
         then A47: LE f.p1,f.p4,P & LE f.p4,f.p2,P & LE f.p2,f.p3,P or
         LE f.p4,f.p2,P & LE f.p2,f.p3,P & LE f.p3,f.p1,P or
         LE f.p2,f.p3,P & LE f.p3,f.p1,P & LE f.p1,f.p4,P or
         LE f.p3,f.p1,P & LE f.p1,f.p4,P & LE f.p4,f.p2,P
                   by JORDAN17:def 1;
         then f.p4=f.p2 or LE f.p2,f.p1,P by A6,A9,JORDAN6:72,73;
         then f.p4=f.p2 or f.p2=f.p1 by A1,A9,JORDAN6:72;
         then p4=p2 or p2=p1 by A1,A8,FUNCT_1:def 8;
         then A48: p4=p2 or p2=p1 & p3=p1 by A2,A46,JORDAN6:72;
           f.p3=f.p1 or LE f.p4,f.p3,P by A5,A9,A47,JORDAN6:72,73;
         then f.p3=f.p1 or f.p4=f.p3 by A1,A9,JORDAN6:72;
         then p3=p1 or p4=p3 by A1,A8,FUNCT_1:def 8;
        hence contradiction by A2,A26,A40,A48,JORDAN17:12;
       case A49: LE p2,p3,K & LE p3,p1,K & LE p1,p4,K;
         then f.p2,f.p3,f.p1,f.p4 are_in_this_order_on P by A1,Th82;
         then A50: LE f.p1,f.p4,P & LE f.p4,f.p2,P & LE f.p2,f.p3,P or
         LE f.p4,f.p2,P & LE f.p2,f.p3,P & LE f.p3,f.p1,P or
         LE f.p2,f.p3,P & LE f.p3,f.p1,P & LE f.p1,f.p4,P or
         LE f.p3,f.p1,P & LE f.p1,f.p4,P & LE f.p4,f.p2,P
                   by JORDAN17:def 1;
         then f.p4=f.p2 or LE f.p2,f.p1,P by A6,A9,JORDAN6:72,73;
         then f.p4=f.p2 or f.p2=f.p1 by A1,A9,JORDAN6:72;
         then p4=p2 or p2=p1 by A1,A8,FUNCT_1:def 8;
         then A51: p4=p2 or p2=p1 & p3=p1 by A2,A49,JORDAN6:72;
           f.p3=f.p1 or LE f.p4,f.p3,P by A5,A9,A50,JORDAN6:72,73;
         then f.p3=f.p1 or f.p4=f.p3 by A1,A9,JORDAN6:72;
         then p3=p1 or p4=p3 by A1,A8,FUNCT_1:def 8;
        hence contradiction by A2,A26,A40,A51,JORDAN17:12;
       case A52: LE p3,p1,K & LE p1,p4,K & LE p4,p2,K;
         then f.p3,f.p1,f.p4,f.p2 are_in_this_order_on P by A1,Th82;
         then A53: LE f.p1,f.p4,P & LE f.p4,f.p2,P & LE f.p2,f.p3,P or
         LE f.p4,f.p2,P & LE f.p2,f.p3,P & LE f.p3,f.p1,P or
         LE f.p2,f.p3,P & LE f.p3,f.p1,P & LE f.p1,f.p4,P or
         LE f.p3,f.p1,P & LE f.p1,f.p4,P & LE f.p4,f.p2,P
                   by JORDAN17:def 1;
         then f.p4=f.p2 or LE f.p2,f.p1,P by A6,A9,JORDAN6:72,73;
         then f.p4=f.p2 or f.p2=f.p1 by A1,A9,JORDAN6:72;
         then p4=p2 or p2=p1 by A1,A8,FUNCT_1:def 8;
         then A54: p4=p2 by A2,A52,JORDAN6:72;
           f.p3=f.p1 or LE f.p4,f.p3,P by A5,A9,A53,JORDAN6:72,73;
         then f.p3=f.p1 or f.p4=f.p3 by A1,A9,JORDAN6:72;
         then p3=p1 or p4=p3 by A1,A8,FUNCT_1:def 8;
        hence contradiction by A2,A26,A40,A54,JORDAN17:12;
       end;
      hence contradiction;
     case A55: p1,p4,p3,p2 are_in_this_order_on K;
        now per cases by A55,JORDAN17:def 1;
      case A56: LE p1,p4,K & LE p4,p3,K & LE p3,p2,K;
         then f.p1,f.p4,f.p3,f.p2 are_in_this_order_on P by A1,Th82;
         then A57: LE f.p1,f.p4,P & LE f.p4,f.p3,P & LE f.p3,f.p2,P or
         LE f.p4,f.p3,P & LE f.p3,f.p2,P & LE f.p2,f.p1,P or
         LE f.p3,f.p2,P & LE f.p2,f.p1,P & LE f.p1,f.p4,P or
         LE f.p2,f.p1,P & LE f.p1,f.p4,P & LE f.p4,f.p3,P
                   by JORDAN17:def 1;
         then f.p3=f.p2 or f.p2=f.p1 by A1,A9,JORDAN6:72;
         then A58: p3=p2 or p2=p1 by A1,A8,FUNCT_1:def 8;
           LE p1,p3,K by A2,A56,JORDAN6:73;
         then A59: p3=p2 by A2,A56,A58,JORDAN6:72;
           f.p4=f.p3 or f.p2=f.p1 by A1,A9,A57,JORDAN6:72;
         then p4=p3 or p2=p1 by A1,A8,FUNCT_1:def 8;
         then p4=p3 or p2,p3,p4,p1 are_in_this_order_on K
                     by A2,A55,A59,JORDAN17:12;
        hence contradiction by A2,A18,A55,A58,JORDAN17:12;
      case LE p4,p3,K & LE p3,p2,K & LE p2,p1,K;
         then f.p4,f.p3,f.p2,f.p1 are_in_this_order_on P by A1,Th82;
         then A60: LE f.p1,f.p4,P & LE f.p4,f.p3,P & LE f.p3,f.p2,P or
         LE f.p4,f.p3,P & LE f.p3,f.p2,P & LE f.p2,f.p1,P or
         LE f.p3,f.p2,P & LE f.p2,f.p1,P & LE f.p1,f.p4,P or
         LE f.p2,f.p1,P & LE f.p1,f.p4,P & LE f.p4,f.p3,P
                   by JORDAN17:def 1;
         then f.p3=f.p2 or f.p2=f.p1 by A1,A9,JORDAN6:72;
         then A61: p3=p2 or p2=p1 by A1,A8,FUNCT_1:def 8;
           f.p4=f.p3 or f.p2=f.p1 by A1,A9,A60,JORDAN6:72;
         then p4=p3 or p2=p1 by A1,A8,FUNCT_1:def 8;
        hence contradiction by A2,A19,A55,A61,JORDAN17:12;
      case LE p3,p2,K & LE p2,p1,K & LE p1,p4,K;
         then f.p3,f.p2,f.p1,f.p4 are_in_this_order_on P by A1,Th82;
         then LE f.p1,f.p4,P & LE f.p4,f.p3,P & LE f.p3,f.p2,P or
         LE f.p4,f.p3,P & LE f.p3,f.p2,P & LE f.p2,f.p1,P or
         LE f.p3,f.p2,P & LE f.p2,f.p1,P & LE f.p1,f.p4,P or
         LE f.p2,f.p1,P & LE f.p1,f.p4,P & LE f.p4,f.p3,P
                   by JORDAN17:def 1;
         then f.p4=f.p3 or f.p2=f.p1 by A1,A9,JORDAN6:72;
         then p4=p3 or p2=p1 by A1,A8,FUNCT_1:def 8;
        hence contradiction by A2,A19,A26,A55,JORDAN17:12;
      case A62: LE p2,p1,K & LE p1,p4,K & LE p4,p3,K;
         then f.p2,f.p1,f.p4,f.p3 are_in_this_order_on P by A1,Th82;
         then A63: LE f.p1,f.p4,P & LE f.p4,f.p3,P & LE f.p3,f.p2,P or
         LE f.p4,f.p3,P & LE f.p3,f.p2,P & LE f.p2,f.p1,P or
         LE f.p3,f.p2,P & LE f.p2,f.p1,P & LE f.p1,f.p4,P or
         LE f.p2,f.p1,P & LE f.p1,f.p4,P & LE f.p4,f.p3,P
                   by JORDAN17:def 1;
         then f.p3=f.p2 or f.p2=f.p1 by A1,A9,JORDAN6:72;
         then A64: p3=p2 or p2=p1 by A1,A8,FUNCT_1:def 8;
           LE p1,p3,K by A2,A62,JORDAN6:73;
         then A65: p1=p2 by A2,A62,A64,JORDAN6:72;
           f.p4=f.p3 or f.p2=f.p3 by A1,A9,A63,JORDAN6:72;
         then p4=p3 or p2=p3 by A1,A8,FUNCT_1:def 8;
        hence contradiction by A2,A26,A55,A65,JORDAN17:12;
      end;
      hence contradiction;
     end;
    hence contradiction;
   end;
 hence p1,p2,p3,p4 are_in_this_order_on K;
end;

theorem Th88: for p1,p2,p3,p4 being Point of TOP-REAL 2,
P,K being non empty compact Subset of TOP-REAL 2,
f being map of TOP-REAL 2,TOP-REAL 2 st
P= circle(0,0,1) & K=rectangle(-1,1,-1,1) & f=Sq_Circ
holds p1,p2,p3,p4 are_in_this_order_on K
  iff f.p1,f.p2,f.p3,f.p4 are_in_this_order_on P
proof let p1,p2,p3,p4 be Point of TOP-REAL 2,
P,K be non empty compact Subset of TOP-REAL 2,
f be map of TOP-REAL 2,TOP-REAL 2;
 assume A1: P= circle(0,0,1) & K=rectangle(-1,1,-1,1) & f=Sq_Circ;
  then A2: K is_simple_closed_curve by Th60;
    circle(0,0,1)={p5 where p5 is Point of TOP-REAL 2: |.p5.|=1}
                      by Th33;
  then A3: P is_simple_closed_curve by A1,JGRAPH_3:36;
 thus p1,p2,p3,p4 are_in_this_order_on K implies
  f.p1,f.p2,f.p3,f.p4 are_in_this_order_on P
  proof assume A4: p1,p2,p3,p4 are_in_this_order_on K;
      now per cases by A4,JORDAN17:def 1;
    case LE p1,p2,K & LE p2,p3,K & LE p3,p4,K;
     hence f.p1,f.p2,f.p3,f.p4 are_in_this_order_on P by A1,Th82;
    case LE p2,p3,K & LE p3,p4,K & LE p4,p1,K;
      then f.p2,f.p3,f.p4,f.p1 are_in_this_order_on P by A1,Th82;
     hence f.p1,f.p2,f.p3,f.p4 are_in_this_order_on P by A3,JORDAN17:12;
    case LE p3,p4,K & LE p4,p1,K & LE p1,p2,K;
      then f.p3,f.p4,f.p1,f.p2 are_in_this_order_on P by A1,Th82;
     hence f.p1,f.p2,f.p3,f.p4 are_in_this_order_on P by A3,JORDAN17:11;
    case LE p4,p1,K & LE p1,p2,K & LE p2,p3,K;
      then f.p4,f.p1,f.p2,f.p3 are_in_this_order_on P by A1,Th82;
     hence f.p1,f.p2,f.p3,f.p4 are_in_this_order_on P by A3,JORDAN17:10;
    end;
   hence thesis;
  end;
 thus f.p1,f.p2,f.p3,f.p4 are_in_this_order_on P implies
  p1,p2,p3,p4 are_in_this_order_on K
  proof assume A5: f.p1,f.p2,f.p3,f.p4 are_in_this_order_on P;
      now per cases by A5,JORDAN17:def 1;
    case LE f.p1,f.p2,P & LE f.p2,f.p3,P & LE f.p3,f.p4,P;
     hence p1,p2,p3,p4 are_in_this_order_on K by A1,Th87;
    case LE f.p2,f.p3,P & LE f.p3,f.p4,P & LE f.p4,f.p1,P;
      then p2,p3,p4,p1 are_in_this_order_on K by A1,Th87;
     hence p1,p2,p3,p4 are_in_this_order_on K by A2,JORDAN17:12;
    case LE f.p3,f.p4,P & LE f.p4,f.p1,P & LE f.p1,f.p2,P;
      then p3,p4,p1,p2 are_in_this_order_on K by A1,Th87;
     hence p1,p2,p3,p4 are_in_this_order_on K by A2,JORDAN17:11;
    case LE f.p4,f.p1,P & LE f.p1,f.p2,P & LE f.p2,f.p3,P;
      then p4,p1,p2,p3 are_in_this_order_on K by A1,Th87;
     hence p1,p2,p3,p4 are_in_this_order_on K by A2,JORDAN17:10;
    end;
   hence p1,p2,p3,p4 are_in_this_order_on K;
  end;
end;

theorem
  for p1,p2,p3,p4 being Point of TOP-REAL 2,
K being compact non empty Subset of TOP-REAL 2,K0 being Subset of TOP-REAL 2
 st K=rectangle(-1,1,-1,1)
 & p1,p2,p3,p4 are_in_this_order_on K
 holds
 (for f,g being map of I[01],TOP-REAL 2 st
  f is continuous one-to-one & g is continuous one-to-one &
  K0= closed_inside_of_rectangle(-1,1,-1,1) &
  f.0=p1 & f.1=p3 & g.0=p2 & g.1=p4 &
  rng f c= K0 & rng g c= K0 holds rng f meets rng g)
proof let p1,p2,p3,p4 be Point of TOP-REAL 2,
K be compact non empty Subset of TOP-REAL 2,K0 be Subset of TOP-REAL 2;
 assume A1: K=rectangle(-1,1,-1,1)
   & p1,p2,p3,p4 are_in_this_order_on K;
   reconsider P= circle(0,0,1) as compact non empty Subset of TOP-REAL 2;
   A2: P={p6 where p6 is Point of TOP-REAL 2: |.p6.|=1} by Th33;
  thus
    (for f,g being map of I[01],TOP-REAL 2 st
  f is continuous one-to-one & g is continuous one-to-one &
  K0= closed_inside_of_rectangle(-1,1,-1,1) &
  f.0=p1 & f.1=p3 & g.0=p2 & g.1=p4 &
  rng f c= K0 & rng g c= K0 holds rng f meets rng g)
  proof let f,g be map of I[01],TOP-REAL 2;
   assume A3: f is continuous one-to-one & g is continuous one-to-one &
    K0= closed_inside_of_rectangle(-1,1,-1,1) &
    f.0=p1 & f.1=p3 & g.0=p2 & g.1=p4 &
    rng f c= K0 & rng g c= K0;
    then A4: K0={p: -1 <=p`1 & p`1<= 1 & -1 <=p`2 & p`2<= 1} by Def3;
    reconsider s=Sq_Circ as map of TOP-REAL 2,TOP-REAL 2;
    A5: dom s=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
    reconsider f1=s*f as map of I[01],TOP-REAL 2;
    reconsider g1=s*g as map of I[01],TOP-REAL 2;
      s is_homeomorphism by JGRAPH_3:54;
    then A6: s is continuous by TOPS_2:def 5;
    then A7: f1 is continuous by A3,TOPS_2:58;
    A8: g1 is continuous by A3,A6,TOPS_2:58;
    A9: f1 is one-to-one by A3,FUNCT_1:46;
    A10: g1 is one-to-one by A3,FUNCT_1:46;
    A11: dom f=[.0,1.] by BORSUK_1:83,FUNCT_2:def 1;
    then 0 in dom f by TOPREAL5:1;
    then A12: f1.0=Sq_Circ.p1 by A3,FUNCT_1:23;
      1 in dom f by A11,TOPREAL5:1;
    then A13: f1.1=Sq_Circ.p3 by A3,FUNCT_1:23;
    A14: dom g=[.0,1.] by BORSUK_1:83,FUNCT_2:def 1;
    then 0 in dom g by TOPREAL5:1;
    then A15: g1.0=Sq_Circ.p2 by A3,FUNCT_1:23;
      1 in dom g by A14,TOPREAL5:1;
    then A16: g1.1=Sq_Circ.p4 by A3,FUNCT_1:23;
    defpred P[Point of TOP-REAL 2] means |.$1.|<=1;
      {p8 where p8 is Point of TOP-REAL 2: P[p8]}
               is Subset of TOP-REAL 2
                    from SubsetD;
    then reconsider C0={p8 where p8 is Point of TOP-REAL 2: |.p8.|<=1} as
          Subset of TOP-REAL 2;
    A17: s.:K0 = C0 by A4,Th36;
    A18: rng f1 c= C0
    proof let y be set;assume y in rng f1;
      then consider x being set such that
      A19: x in dom f1 & y=f1.x by FUNCT_1:def 5;
      A20: x in dom f & f.x in dom s by A19,FUNCT_1:21;
      then f.x in rng f by FUNCT_1:12;
      then s.(f.x) in s.:K0 by A3,A20,FUNCT_1:def 12;
     hence y in C0 by A17,A19,FUNCT_1:22;
    end;
    A21: rng g1 c= C0
    proof let y be set;assume y in rng g1;
      then consider x being set such that
      A22: x in dom g1 & y=g1.x by FUNCT_1:def 5;
      A23: x in dom g & g.x in dom s by A22,FUNCT_1:21;
      then g.x in rng g by FUNCT_1:12;
      then s.(g.x) in s.:K0 by A3,A23,FUNCT_1:def 12;
     hence y in C0 by A17,A22,FUNCT_1:22;
    end;
    reconsider q1=s.p1,q2=s.p2,q3=s.p3,q4=s.p4
          as Point of TOP-REAL 2;
      q1,q2,q3,q4 are_in_this_order_on P
           by A1,Th88;
    then rng f1 meets rng g1
         by A2,A7,A8,A9,A10,A12,A13,A15,A16,A18,A21,Th27;
    then consider y being set such that
    A24: y in rng f1 & y in rng g1 by XBOOLE_0:3;
    consider x1 being set such that
    A25: x1 in dom f1 & y=f1.x1 by A24,FUNCT_1:def 5;
    consider x2 being set such that
    A26: x2 in dom g1 & y=g1.x2 by A24,FUNCT_1:def 5;
      dom f1 c= dom f by RELAT_1:44;
    then A27: f.x1 in rng f by A25,FUNCT_1:12;
      dom g1 c= dom g by RELAT_1:44;
    then A28: g.x2 in rng g by A26,FUNCT_1:12;
      y=(Sq_Circ).(f.x1) by A25,FUNCT_1:22;
    then A29: Sq_Circ".y=f.x1 by A5,A27,FUNCT_1:54;
      x1 in dom f by A25,FUNCT_1:21;
    then A30: f.x1 in rng f by FUNCT_1:def 5;
      y=(Sq_Circ).(g.x2) by A26,FUNCT_1:22;
    then A31: Sq_Circ".y=g.x2 by A5,A28,FUNCT_1:54;
      x2 in dom g by A26,FUNCT_1:21;
    then g.x2 in rng g by FUNCT_1:def 5;
   hence rng f meets rng g by A29,A30,A31,XBOOLE_0:3;
  end;
end;

Back to top