The Mizar article:

General Fashoda Meet Theorem for Unit Circle

by
Yatsuka Nakamura

Received June 24, 2002

Copyright (c) 2002 Association of Mizar Users

MML identifier: JGRAPH_5
[ 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,
      JGRAPH_4, ORDINAL2, TOPS_2, ARYTM_1, COMPLEX1, MCART_1, PCOMPS_1,
      JGRAPH_3, BORSUK_1, TOPREAL1, TOPREAL2, JORDAN3, PSCOMP_1, REALSET1,
      JORDAN5C, JORDAN6, ARYTM, SEQ_1;
 notation ORDINAL1, NUMBERS, XCMPLX_0, XREAL_0, REAL_1, XBOOLE_0, ABSVALUE,
      EUCLID, TARSKI, RELAT_1, TOPS_2, FUNCT_1, FUNCT_2, NAT_1, STRUCT_0,
      TOPMETR, PCOMPS_1, COMPTS_1, METRIC_1, SQUARE_1, RCOMP_1, PSCOMP_1,
      BINOP_1, PRE_TOPC, JGRAPH_1, JGRAPH_3, TOPREAL1, JORDAN5C, JORDAN6,
      TOPREAL2, JGRAPH_4, GRCAT_1;
 constructors REAL_1, ABSVALUE, TOPREAL1, TOPS_2, RCOMP_1, PSCOMP_1, TOPREAL2,
      WELLFND1, JGRAPH_3, JORDAN5C, JORDAN6, JGRAPH_4, GRCAT_1, BORSUK_3,
      TOPRNS_1;
 clusters XREAL_0, STRUCT_0, RELSET_1, FUNCT_1, EUCLID, PRE_TOPC, TOPMETR,
      SQUARE_1, PSCOMP_1, BORSUK_1, METRIC_1, BORSUK_2, BORSUK_3, MEMBERED;
 requirements REAL, NUMERALS, BOOLE, SUBSET, ARITHM;
 definitions TARSKI, JORDAN6;
 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, SEQ_4, SQUARE_1, PSCOMP_1, METRIC_1, JGRAPH_2, RCOMP_1,
      COMPTS_1, RFUNCT_2, SETWISEO, BORSUK_1, TOPREAL1, TOPREAL3, TOPREAL5,
      JGRAPH_3, ABSVALUE, COMPLEX1, JORDAN5A, JORDAN5B, JORDAN7, HEINE,
      JGRAPH_4, PCOMPS_1, JORDAN5C, JORDAN1B, XREAL_0, TREAL_1, GRCAT_1,
      TSEP_1, JORDAN1A, JORDAN1, TOPRNS_1, XCMPLX_0, XCMPLX_1;
 schemes FUNCT_2, JGRAPH_2;

begin :: Preliminaries

reserve x,a for real number;

theorem Th1:
  a>=0 & (x-a)*(x+a)>=0 implies -a>=x or x>=a
proof assume a>=0 & (x-a)*(x+a)>=0;
   then x-a>=0 & x+a>=0 or x-a<=0 & x+a<=0 by SQUARE_1:25;
   then x-a+a>=0+a or x+a<=0 by REAL_1:55;
   then x-(a-a)>=0+a or x+a<=0 by XCMPLX_1:37;
   then x>=0+a or x+a-a<=0-a by REAL_1:49,XCMPLX_1:17;
   then x>=0+a or x+(a-a)<=0-a by XCMPLX_1:29;
   then x>=a or x<=0-a by XCMPLX_1:25;
  hence -a>=x or x>=a by XCMPLX_1:150;
end;

theorem Th2: a<=0 & x<a implies x^2>a^2
proof assume A1: a<=0 & x<a;
  then --a<=0;
  then A2: -a>=0 by REAL_1:66;
    -x>-a by A1,REAL_1:50;
  then (-x)^2>(-a)^2 by A2,SQUARE_1:78;
  then x^2>(-a)^2 by SQUARE_1:61;
 hence thesis by SQUARE_1:61;
end;

theorem Th3: for p being Point of TOP-REAL 2 st |.p.|<=1
 holds -1<=p`1 & p`1<=1 & -1<=p`2 & p`2<=1
proof let p be Point of TOP-REAL 2;
 assume A1: |.p.|<=1;
 set a=|.p.|;
 A2: a>=0 by TOPRNS_1:26;
 A3: (|.p.|)^2 =(p`1)^2+(p`2)^2 by JGRAPH_3:10;
 then a^2-(p`1)^2=(p`2)^2 by XCMPLX_1:26;
 then a^2-(p`1)^2>=0 by SQUARE_1:72;
 then a^2-(p`1)^2+(p`1)^2>=0+(p`1)^2 by REAL_1:55;
 then a^2>=(p`1)^2 by XCMPLX_1:27;
 then A4: -a<=p`1 & p`1<=a by A2,JGRAPH_2:5;
   a^2-(p`2)^2=(p`1)^2 by A3,XCMPLX_1:26;
 then a^2-(p`2)^2>=0 by SQUARE_1:72;
 then a^2-(p`2)^2+(p`2)^2>=0+(p`2)^2 by REAL_1:55;
 then a^2>=(p`2)^2 by XCMPLX_1:27;
 then A5: -a<=p`2 & p`2<=a by A2,JGRAPH_2:5;
   -a>=-1 by A1,REAL_1:50;
 hence thesis by A1,A4,A5,AXIOMS:22;
end;

theorem Th4: for p being Point of TOP-REAL 2 st |.p.|<=1 & p`1<>0 & p`2<>0
 holds -1<p`1 & p`1<1 & -1<p`2 & p`2<1
proof let p be Point of TOP-REAL 2;
 assume A1: |.p.|<=1 & p`1<>0 & p`2<>0;
 set a=|.p.|;
 A2: a>=0 by TOPRNS_1:26;
 A3: (|.p.|)^2 =(p`1)^2+(p`2)^2 by JGRAPH_3:10;
 then a^2-(p`1)^2=(p`2)^2 by XCMPLX_1:26;
 then a^2-(p`1)^2>0 by A1,SQUARE_1:74;
 then a^2-(p`1)^2+(p`1)^2>0+(p`1)^2 by REAL_1:67;
 then a^2>(p`1)^2 by XCMPLX_1:27;
 then A4: -a<p`1 & p`1<a by A2,JGRAPH_2:6;
   a^2-(p`2)^2=(p`1)^2 by A3,XCMPLX_1:26;
 then a^2-(p`2)^2>0 by A1,SQUARE_1:74;
 then a^2-(p`2)^2+(p`2)^2>0+(p`2)^2 by REAL_1:67;
 then a^2>(p`2)^2 by XCMPLX_1:27;
 then A5: -a<p`2 & p`2<a by A2,JGRAPH_2:6;
   -a>=-1 by A1,REAL_1:50;
 hence thesis by A1,A4,A5,AXIOMS:22;
end;

theorem for a,b,d,e,r3 being Real,PM,PM2 being non empty MetrStruct,
 x being Element of PM,
 x2 being Element of PM2
 st d<=a & a<=b & b<=e
 & PM=Closed-Interval-MSpace(a,b)
 & PM2=Closed-Interval-MSpace(d,e)
 & x=x2 & x in the carrier of PM & x2 in the carrier of PM2
 holds Ball(x,r3) c= Ball(x2,r3)
proof let a,b,d,e,r3 be Real,PM,PM2 be non empty MetrStruct,
 x be Element of PM,
 x2 be Element of PM2;
 assume A1: d<=a & a<=b & b<=e
 & PM=Closed-Interval-MSpace(a,b)
 & PM2=Closed-Interval-MSpace(d,e) &
 x=x2 & x in the carrier of PM & x2 in the carrier of PM2;
  then A2: d<=b by AXIOMS:22;
  then A3: d<=e by A1,AXIOMS:22;
  A4: a<=e by A1,AXIOMS:22;
 let z be set;assume z in Ball(x,r3);
   then z in {y where y is Element of PM: dist(x,y) < r3 }
                                 by METRIC_1:18;
   then consider y being Element of PM such that
   A5: y=z & dist(x,y)<r3;
   A6: the carrier of PM=[.a,b.] by A1,TOPMETR:14;
   A7: a in [.d,e.] by A1,A4,TOPREAL5:1;
     b in [.d,e.] by A1,A2,TOPREAL5:1;
   then A8: [.a,b.] c= [.d,e.] by A7,RCOMP_1:16;
   A9: (the distance of PM).(x,y)
                    = real_dist.(x,y) by A1,METRIC_1:def 14,TOPMETR:def 1;
   A10: dist(x,y)= (the distance of PM).(x,y) by METRIC_1:def 1;
     y in [.a,b.] by A6;
   then reconsider y3=y as Element of PM2 by A1,A3,A8,TOPMETR:14
;
     real_dist.(x,y)=
(the distance of PM2).(x2,y3) by A1,METRIC_1:def 14,TOPMETR:def 1
                .=dist(x2,y3) by METRIC_1:def 1;
   then z in {y2 where y2 is Element of PM2:
  dist(x2,y2)<r3} by A5,A9,A10;
  hence thesis by METRIC_1:18;
end;

theorem Th6: for a,b,d,e being real number,
 B being Subset of Closed-Interval-TSpace(d,e)
 st d<=a & a<=b & b<=e & B=[.a,b.] holds
 Closed-Interval-TSpace(a,b)=Closed-Interval-TSpace(d,e)|B
proof let a,b,d,e be real number,
 B be Subset of Closed-Interval-TSpace(d,e);
 assume A1: d<=a & a<=b & b<=e & B=[.a,b.];
  then A2: d<=b by AXIOMS:22;
  then A3: d<=e by A1,AXIOMS:22;
  A4: a<=e by A1,AXIOMS:22;
  reconsider A=[.d,e.] as non empty Subset of R^1
                           by A1,A2,TOPMETR:24,TOPREAL5:1;
  reconsider B2=[.a,b.] as non empty Subset of R^1
                    by A1,TOPMETR:24,TOPREAL5:1;
  A5: a in [.d,e.] by A1,A4,TOPREAL5:1;
    b in [.d,e.] by A1,A2,TOPREAL5:1;
  then A6: [.a,b.] c= [.d,e.] by A5,RCOMP_1:16;
  A7: Closed-Interval-TSpace(a,b)=R^1|B2 by A1,TOPMETR:26;
    Closed-Interval-TSpace(d,e)=R^1|A by A3,TOPMETR:26;
 hence thesis by A1,A6,A7,JORDAN6:47;
end;

theorem for a,b being real number, B being Subset of I[01]
 st 0<=a & a<=b & b<=1 & B=[.a,b.] holds
 Closed-Interval-TSpace(a,b)=I[01]|B by Th6,TOPMETR:27;

theorem Th8: for X being TopStruct,
 Y,Z being non empty TopStruct,f being map of X,Y,
 h being map of Y,Z st h is_homeomorphism & f is continuous
 holds h*f is continuous
 proof let X be TopStruct,Y,Z be non empty TopStruct,f be map of X,Y,
 h be map of Y,Z;
 assume A1: h is_homeomorphism &
   f is continuous;
   then h is continuous by TOPS_2:def 5;
  hence h*f is continuous by A1,TOPS_2:58;
 end;

theorem Th9: for X,Y,Z being TopStruct, f being map of X,Y,
 h being map of Y,Z st h is_homeomorphism & f is one-to-one
 holds h*f is one-to-one
 proof let X,Y,Z be TopStruct, f be map of X,Y, h be map of Y,Z;
  assume A1: h is_homeomorphism & f is one-to-one;
   then h is one-to-one by TOPS_2:def 5;
  hence h*f is one-to-one by A1,FUNCT_1:46;
 end;

theorem Th10: for X being TopStruct,S,V being non empty TopStruct,
 B being non empty Subset of S,f being map of X,S|B, g being map of S,V,
 h being map of X,V
 st h=g*f & f is continuous & g is continuous holds h is continuous
    proof let X be TopStruct,S,V be non empty TopStruct,
     B be non empty Subset of S,
     f be map of X,S|B, g be map of S,V,
     h being map of X,V;
     assume that A1: h=g*f & f is continuous and A2: g is continuous;
        now let P be Subset of V;
      A3: (g*f)"P = f"(g"P) by RELAT_1:181;
         now assume P is closed;
        then A4: g"P is closed by A2,PRE_TOPC:def 12;
        A5: [#](S|B)=B by PRE_TOPC:def 10;
        A6: the carrier of S|B =B by JORDAN1:1;
        then B /\ g"P c= the carrier of S|B by XBOOLE_1:17;
        then reconsider F=B /\ g"P as Subset of S|B;
        A7: F is closed by A4,A5,PRE_TOPC:43;
        A8: rng f /\ (the carrier of S|B)= rng f by XBOOLE_1:28;
          h"P=f"(rng f /\ g"P) by A1,A3,RELAT_1:168
        .=f"(rng f /\ ((the carrier of S|B) /\ g"P)) by A8,XBOOLE_1:16
        .=f"F by A6,RELAT_1:168;
        hence h"P is closed by A1,A7,PRE_TOPC:def 12;
       end;
       hence P is closed implies h"P is closed;
      end;
     hence thesis by PRE_TOPC:def 12;
    end;

theorem Th11:for a,b,d,e,s1,s2,t1,t2 being Real,h being map of
 Closed-Interval-TSpace(a,b),Closed-Interval-TSpace(d,e) st
 h is_homeomorphism & h.s1=t1 & h.s2=t2 & h.a=d & h.b=e & d<=e &
 t1<=t2 & s1 in [.a,b.] & s2 in [.a,b.] holds s1<=s2
 proof let a,b,d,e,s1,s2,t1,t2 be Real,h be map of
   Closed-Interval-TSpace(a,b),Closed-Interval-TSpace(d,e);
   assume A1: h is_homeomorphism & h.s1=t1 & h.s2=t2 & h.a=d & h.b=e & d<=e
    & t1<=t2 & s1 in [.a,b.] & s2 in [.a,b.];
    then A2: h is one-to-one by TOPS_2:def 5;
    A3: a<=s2 & s2<=b by A1,TOPREAL5:1;
    A4: a<=s1 & s1<=b by A1,TOPREAL5:1;
    then A5: a<=b by AXIOMS:22;
    A6: dom h=[#](Closed-Interval-TSpace(a,b)) by A1,TOPS_2:def 5
    .=the carrier of Closed-Interval-TSpace(a,b) by PRE_TOPC:12
    .=[.a,b.] by A5,TOPMETR:25;
    A7: h is continuous by A1,TOPS_2:def 5;
    A8: the carrier of Closed-Interval-TSpace(a,b)
              =[.a,b.] by A5,TOPMETR:25;
    A9: the carrier of Closed-Interval-TSpace(d,e)
              =[.d,e.] by A1,TOPMETR:25;
    A10: h is one-to-one by A1,TOPS_2:def 5;
      [.s2,s1.] c= the carrier of Closed-Interval-TSpace(a,b)
                        by A3,A4,A8,TREAL_1:1;
    then reconsider B=[.s2,s1.] as Subset of Closed-Interval-TSpace(a,b)
                         ;
    reconsider Bb=[.s2,s1.] as Subset of
             Closed-Interval-TSpace(a,b) by A3,A4,A8,TREAL_1:1;
    assume A11: s1>s2;
      reconsider f3=h|Bb as map of Closed-Interval-TSpace(a,b)|B,
         Closed-Interval-TSpace(d,e) by JGRAPH_3:12;
      A12: f3 is continuous by A7,TOPMETR:10;
      reconsider C=[.d,e.] as non empty Subset of R^1
                                 by A1,TOPMETR:24,TOPREAL5:1;
      A13: R^1|C=Closed-Interval-TSpace(d,e) by A1,TOPMETR:26;
      A14: Closed-Interval-TSpace(s2,s1)
      =Closed-Interval-TSpace(a,b)|B by A3,A4,A11,Th6;
      then f3 is map of Closed-Interval-TSpace(s2,s1),R^1
                     by A13,JORDAN6:4;
      then reconsider f=h|B as map of Closed-Interval-TSpace(s2,s1),R^1;
        dom f=the carrier of Closed-Interval-TSpace(s2,s1)
                       by FUNCT_2:def 1;
      then A15: dom f=[.s2,s1.] by A11,TOPMETR:25;
      A16: f is continuous by A12,A13,A14,JORDAN6:4;
      set t=(t1+t2)/2;
        s2 in B by A11,TOPREAL5:1;
      then A17: f.s2=t2 by A1,FUNCT_1:72;
        s1 in B by A11,TOPREAL5:1;
      then A18: f.s1=t1 by A1,FUNCT_1:72;
        t1<>t2 by A1,A2,A6,A11,FUNCT_1:def 8;
      then A19: t1<t2 by A1,REAL_1:def 5;
      then t1+t1<t1+t2 by REAL_1:67;
      then (t1+t1)/2<(t1+t2)/2 by REAL_1:73;
      then A20: (2*t1)/2<t by XCMPLX_1:11;
        t1+t2<t2+t2 by A19,REAL_1:67;
      then (t1+t2)/2<(t2+t2)/2 by REAL_1:73;
      then (2*t2)/2>t by XCMPLX_1:11;
      then A21: t2>t & t>t1 by A20,XCMPLX_1:90;
      then consider r being Real such that
      A22: f.r =t & s2<r & r <s1 by A11,A16,A17,A18,TOPREAL5:13;
      A23: r<b by A4,A22,AXIOMS:22;
        a<r by A3,A22,AXIOMS:22;
      then A24: r in [.a,b.] by A23,TOPREAL5:1;
          [.s1,b.] c= the carrier of Closed-Interval-TSpace(a,b)
                                   by A4,A8,TREAL_1:1;
        then reconsider B1=[.s1,b.] as Subset of Closed-Interval-TSpace(a,b)
                              ;
        reconsider B1b=[.s1,b.] as Subset of
             Closed-Interval-TSpace(a,b) by A4,A8,TREAL_1:1;
        reconsider f4=h|B1b as map of Closed-Interval-TSpace(a,b)|B1,
         Closed-Interval-TSpace(d,e) by JGRAPH_3:12;
        A25: Closed-Interval-TSpace(s1,b)
          =Closed-Interval-TSpace(a,b)|B1 by A4,Th6;
        A26: f4 is continuous by A7,TOPMETR:10;
          f4 is map of Closed-Interval-TSpace(s1,b),R^1
                     by A13,A25,JORDAN6:4;
        then reconsider f1=h|B1 as map of Closed-Interval-TSpace(s1,b),R^1;
        A27: f1 is continuous by A13,A25,A26,JORDAN6:4;
          s2 in dom f by A11,A15,TOPREAL5:1;
        then t2 in rng f3 by A17,FUNCT_1:def 5;
        then A28: d<=t2 & t2<=e by A9,TOPREAL5:1;
        then A29: s1<b by A1,A4,A19,REAL_1:def 5;
        A30: s1 in B1 by A4,TOPREAL5:1;
        A31: b in B1 by A4,TOPREAL5:1;
        A32: f1.s1= t1 by A1,A30,FUNCT_1:72;
        A33: f1.b= e by A1,A31,FUNCT_1:72;
          e>t & t>t1 by A21,A28,AXIOMS:22;
        then consider r1 being Real such that
        A34: f1.r1 =t & s1<r1 & r1 <b by A27,A29,A32,A33,TOPREAL5:12;
          a<r1 by A4,A34,AXIOMS:22;
        then A35: r1 in [.a,b.] by A34,TOPREAL5:1;
        A36: r1 in B1 by A34,TOPREAL5:1;
          r in [.s2,s1.] by A22,TOPREAL5:1;
        then h.r = t by A22,FUNCT_1:72 .=h.r1 by A34,A36,FUNCT_1:72;
       hence contradiction by A6,A10,A22,A24,A34,A35,FUNCT_1:def 8;
 end;

theorem Th12:for a,b,d,e,s1,s2,t1,t2 being Real,h being map of
 Closed-Interval-TSpace(a,b),Closed-Interval-TSpace(d,e) st
 h is_homeomorphism & h.s1=t1 & h.s2=t2 & h.a=e & h.b=d & e>=d &
 t1>=t2 & s1 in [.a,b.] & s2 in [.a,b.] holds s1<=s2
 proof let a,b,d,e,s1,s2,t1,t2 be Real,h be map of
   Closed-Interval-TSpace(a,b),Closed-Interval-TSpace(d,e);
   assume A1: h is_homeomorphism & h.s1=t1 & h.s2=t2 & h.a=e & h.b=d & e>=d
    & t1>=t2 & s1 in [.a,b.] & s2 in [.a,b.];
    then A2: h is one-to-one by TOPS_2:def 5;
    A3: a<=s2 & s2<=b by A1,TOPREAL5:1;
    A4: a<=s1 & s1<=b by A1,TOPREAL5:1;
    then A5: a<=b by AXIOMS:22;
    A6: dom h=[#](Closed-Interval-TSpace(a,b)) by A1,TOPS_2:def 5
    .=the carrier of Closed-Interval-TSpace(a,b) by PRE_TOPC:12
    .=[.a,b.] by A5,TOPMETR:25;
    A7: h is continuous by A1,TOPS_2:def 5;
    A8: the carrier of Closed-Interval-TSpace(a,b)
              =[.a,b.] by A5,TOPMETR:25;
    A9: the carrier of Closed-Interval-TSpace(d,e)
              =[.d,e.] by A1,TOPMETR:25;
    A10: h is one-to-one by A1,TOPS_2:def 5;
      [.s2,s1.] c= the carrier of Closed-Interval-TSpace(a,b)
                          by A3,A4,A8,TREAL_1:1;
    then reconsider B=[.s2,s1.] as Subset of Closed-Interval-TSpace(a,b)
                         ;
    reconsider Bb=[.s2,s1.] as Subset of
             Closed-Interval-TSpace(a,b) by A3,A4,A8,TREAL_1:1;
    assume A11: s1>s2;
      reconsider f3=h|Bb as map of Closed-Interval-TSpace(a,b)|B,
         Closed-Interval-TSpace(d,e) by JGRAPH_3:12;
      A12: f3 is continuous by A7,TOPMETR:10;
      reconsider C=[.d,e.] as non empty Subset of R^1
                                 by A1,TOPMETR:24,TOPREAL5:1;
      A13: R^1|C=Closed-Interval-TSpace(d,e) by A1,TOPMETR:26;
      A14: Closed-Interval-TSpace(s2,s1)
      =Closed-Interval-TSpace(a,b)|B by A3,A4,A11,Th6;
      then f3 is map of Closed-Interval-TSpace(s2,s1),R^1
                     by A13,JORDAN6:4;
      then reconsider f=h|B as map of Closed-Interval-TSpace(s2,s1),R^1;
        dom f=the carrier of Closed-Interval-TSpace(s2,s1)
                       by FUNCT_2:def 1;
      then A15: dom f=[.s2,s1.] by A11,TOPMETR:25;
      A16: f is continuous by A12,A13,A14,JORDAN6:4;
      set t=(t1+t2)/2;
        s2 in B by A11,TOPREAL5:1;
      then A17: f.s2=t2 by A1,FUNCT_1:72;
        s1 in B by A11,TOPREAL5:1;
      then A18: f.s1=t1 by A1,FUNCT_1:72;
        t1<>t2 by A1,A2,A6,A11,FUNCT_1:def 8;
      then A19: t1>t2 by A1,REAL_1:def 5;
      then t1+t1>t1+t2 by REAL_1:67;
      then (t1+t1)/2>(t1+t2)/2 by REAL_1:73;
      then A20: (2*t1)/2>t by XCMPLX_1:11;
        t1+t2>t2+t2 by A19,REAL_1:67;
      then (t1+t2)/2>(t2+t2)/2 by REAL_1:73;
      then (2*t2)/2<t by XCMPLX_1:11;
      then A21: t2<t & t<t1 by A20,XCMPLX_1:90;
      then consider r being Real such that
      A22: f.r =t & s2<r & r <s1 by A11,A16,A17,A18,TOPREAL5:12;
      A23: r<b by A4,A22,AXIOMS:22;
        a<r by A3,A22,AXIOMS:22;
      then A24: r in [.a,b.] by A23,TOPREAL5:1;
        [.s1,b.] c= the carrier of Closed-Interval-TSpace(a,b)
                                   by A4,A8,TREAL_1:1;
      then reconsider B1=[.s1,b.] as Subset of Closed-Interval-TSpace(a,b)
                              ;
      reconsider B1b=[.s1,b.] as Subset of
             Closed-Interval-TSpace(a,b) by A4,A8,TREAL_1:1;
      reconsider f4=h|B1b as map of Closed-Interval-TSpace(a,b)|B1,
         Closed-Interval-TSpace(d,e) by JGRAPH_3:12;
      A25: Closed-Interval-TSpace(s1,b)
          =Closed-Interval-TSpace(a,b)|B1 by A4,Th6;
      A26: f4 is continuous by A7,TOPMETR:10;
        f4 is map of Closed-Interval-TSpace(s1,b),R^1
                     by A13,A25,JORDAN6:4;
      then reconsider f1=h|B1 as map of Closed-Interval-TSpace(s1,b),R^1;
      A27: f1 is continuous by A13,A25,A26,JORDAN6:4;
        s2 in dom f by A11,A15,TOPREAL5:1;
      then t2 in rng f3 by A17,FUNCT_1:def 5;
      then A28: d<=t2 & t2<=e by A9,TOPREAL5:1;
      then A29: s1<b by A1,A4,A19,REAL_1:def 5;
      A30: s1 in B1 by A4,TOPREAL5:1;
      A31: b in B1 by A4,TOPREAL5:1;
      A32: f1.s1= t1 by A1,A30,FUNCT_1:72;
      A33: f1.b= d by A1,A31,FUNCT_1:72;
        d<t & t<t1 by A21,A28,AXIOMS:22;
      then consider r1 being Real such that
         A34: f1.r1 =t & s1<r1 & r1 <b by A27,A29,A32,A33,TOPREAL5:13;
        a<r1 by A4,A34,AXIOMS:22;
      then A35: r1 in [.a,b.] by A34,TOPREAL5:1;
      A36: r1 in B1 by A34,TOPREAL5:1;
        r in [.s2,s1.] by A22,TOPREAL5:1;
      then h.r= t by A22,FUNCT_1:72 .=h.r1 by A34,A36,FUNCT_1:72;
     hence contradiction by A6,A10,A22,A24,A34,A35,FUNCT_1:def 8;
end;

theorem for n being Nat holds
 -(0.REAL n)=0.REAL n
 proof let n be Nat;
     0.REAL n+0.REAL n=0.REAL n by EUCLID:31;
  hence thesis by EUCLID:41;
 end;

begin :: Fashoda Meet Theorems for Circle in Special Case

theorem Th14:
 for f,g being map of I[01],TOP-REAL 2,a,b,c,d being Real,
   O,I being Point of I[01] st O=0 & I=1 &
   f is continuous one-to-one &
   g is continuous one-to-one & a <> b & c <> d &
   (f.O)`1=a & (c <=(f.O)`2 & (f.O)`2 <=d) &
   (f.I)`1=b & (c <=(f.I)`2 & (f.I)`2 <=d) &
   (g.O)`2=c & (a <=(g.O)`1 & (g.O)`1 <=b) &
   (g.I)`2=d & (a <=(g.I)`1 & (g.I)`1 <=b) &
   (for r being Point of I[01] holds
                (a >=(f.r)`1 or (f.r)`1>=b or c >=(f.r)`2 or (f.r)`2>=d) &
                (a >=(g.r)`1 or (g.r)`1>=b or c >=(g.r)`2 or (g.r)`2>=d))
   holds rng f meets rng g
proof let f,g be map of I[01],TOP-REAL 2,a,b,c,d be Real,
   O,I be Point of I[01];
 assume A1: O=0 & I=1 &
  f is continuous one-to-one & g is continuous one-to-one & a <> b & c <> d &
   (f.O)`1=a & (c <=(f.O)`2 & (f.O)`2 <=d) &
   (f.I)`1=b & (c <=(f.I)`2 & (f.I)`2 <=d) &
   (g.O)`2=c & (a <=(g.O)`1 & (g.O)`1 <=b) &
   (g.I)`2=d & (a <=(g.I)`1 & (g.I)`1 <=b) &
   (for r being Point of I[01] holds
                (a >=(f.r)`1 or (f.r)`1>=b or c >=(f.r)`2 or (f.r)`2>=d) &
                (a >=(g.r)`1 or (g.r)`1>=b or c >=(g.r)`2 or (g.r)`2>=d));
  then A2: a <= b by AXIOMS:22;
    c <= d by A1,AXIOMS:22;
  then a < b & c < d by A1,A2,REAL_1:def 5;
 hence thesis by A1,JGRAPH_2:55;
end;

Lm1: 0 in [.0,1.] & 1 in [.0,1.] by RCOMP_1:15;

theorem Th15: for f being map of I[01],TOP-REAL 2 st
  f is continuous one-to-one
  ex f2 being map of I[01],TOP-REAL 2 st f2.0=f.1 & f2.1=f.0 &
  rng f2=rng f & f2 is continuous & f2 is one-to-one
proof let f be map of I[01],TOP-REAL 2;
 assume A1: f is continuous one-to-one;
   A2: I[01] is compact by HEINE:11,TOPMETR:27;
   A3: dom f=the carrier of I[01] by FUNCT_2:def 1;
   then A4: f.1 in rng f by Lm1,BORSUK_1:83,FUNCT_1:12;
   reconsider P=rng f as non empty Subset of TOP-REAL 2 by A3,Lm1,BORSUK_1:83,
FUNCT_1:12;
   consider f1 being map of I[01],(TOP-REAL 2)|P such that
   A5: f1=f & f1 is_homeomorphism by A1,A2,JGRAPH_1:64;
     f.0 in rng f by A3,Lm1,BORSUK_1:83,FUNCT_1:12;
   then reconsider p1=f.0,p2=f.1 as Point of TOP-REAL 2 by A4;
     P is_an_arc_of p1,p2 by A5,TOPREAL1:def 2;
   then P is_an_arc_of p2,p1 by JORDAN5B:14;
   then consider f3 being map of I[01], (TOP-REAL 2)|P such that
   A6: f3 is_homeomorphism & f3.0 = p2 & f3.1 = p1 by TOPREAL1:def 2;
   A7: rng f3=[#]((TOP-REAL 2)|P) by A6,TOPS_2:def 5
         .=P by PRE_TOPC:def 10;
   consider f4 being map of I[01], (TOP-REAL 2) such that
   A8: f3=f4 & f4 is continuous & f4 is one-to-one by A6,JORDAN7:15;
  thus thesis by A6,A7,A8;
end;

reserve p,q for Point of TOP-REAL 2;

theorem Th16:
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 KXN & f.I in KXP & 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 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 KXN & f.I in KXP & g.O in KYP & g.I in KYN &
  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 one-to-one by Th15;
 thus thesis by A1,A2,JGRAPH_3:55;
end;

theorem Th17: 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 KXN & f.I in KXP & 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 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 KXN & f.I in KXP & g.O in KYN & g.I in KYP &
  rng f c= C0 & rng g c= C0;
      Sq_Circ"*f is Function of the carrier of I[01],
           the carrier of TOP-REAL 2 by FUNCT_2:19,JGRAPH_3:39;
    then reconsider ff=Sq_Circ"*f as map of I[01],TOP-REAL 2;
      Sq_Circ"*g is Function of the carrier of I[01],
           the carrier of TOP-REAL 2 by FUNCT_2:19,JGRAPH_3:39;
    then reconsider gg=Sq_Circ"*g as map of I[01],TOP-REAL 2;
    consider h1 being map of (TOP-REAL 2),(TOP-REAL 2)
      such that A2:h1=(Sq_Circ") & h1 is continuous by JGRAPH_3:52;
    A3:dom ff=the carrier of I[01] by FUNCT_2:def 1;
    A4:dom gg=the carrier of I[01] by FUNCT_2:def 1;
    A5:dom f=the carrier of I[01] by FUNCT_2:def 1;
    A6:dom g=the carrier of I[01] by FUNCT_2:def 1;
    A7:ff is continuous by A1,A2,TOPS_2:58;
    A8: Sq_Circ" is one-to-one by FUNCT_1:62,JGRAPH_3:32;
    then A9:ff is one-to-one by A1,FUNCT_1:46;
    A10:gg is continuous by A1,A2,TOPS_2:58;
    A11:gg is one-to-one by A1,A8,FUNCT_1:46;
    A12: (ff.O)`1=-1 & (ff.I)`1=1 & (gg.O)`2=-1 & (gg.I)`2=1
    proof
     A13: (ff.O)=(Sq_Circ").(f.O) by A3,FUNCT_1:22;
     consider p1 being Point of TOP-REAL 2 such that
     A14:  f.O=p1 &( |.p1.|=1 & p1`2>=p1`1 & p1`2<=-p1`1) by A1;
     A15:p1<>0.REAL 2 &
          (p1`2<=p1`1 & -p1`1<=p1`2 or p1`2>=p1`1 & p1`2<=-p1`1)
                            by A14,TOPRNS_1:24;
    then A16:Sq_Circ".p1=|[p1`1*sqrt(1+(p1`2/p1`1)^2),p1`2*sqrt(1+(p1`2/p1`1)^2
)]|
                                 by JGRAPH_3:38;
        reconsider px=ff.O as Point of TOP-REAL 2;
        set q=px;
        A17: px`1 = p1`1*sqrt(1+(p1`2/p1`1)^2) &
         px`2 = p1`2*sqrt(1+(p1`2/p1`1)^2) by A13,A14,A16,EUCLID:56;
          (p1`2/p1`1)^2 >=0 by SQUARE_1:72;
        then 1+(p1`2/p1`1)^2>=1+0 by REAL_1:55;
        then 1+(p1`2/p1`1)^2>0 by AXIOMS:22;
        then A18:sqrt(1+(p1`2/p1`1)^2)>0 by SQUARE_1:93;
        A19:now assume
A20:        px`1=0 & px`2=0;
          then A21:p1`1=0 by A17,A18,XCMPLX_1:6;
            p1`2=0 by A17,A18,A20,XCMPLX_1:6;
         hence contradiction by A15,A21,EUCLID:57,58;
        end;
          p1`2<=p1`1 & -p1`1<=p1`2 or p1`2>=p1`1 &
        p1`2*sqrt(1+(p1`2/p1`1)^2) <= (-p1`1)*sqrt(1+(p1`2/p1`1)^2)
                             by A14,A18,AXIOMS:25;
     then p1`2<=p1`1 & (-p1`1)*sqrt(1+(p1`2/p1`1)^2) <= p1`2*sqrt(1+(p1`2/p1`1
)^2)
           or px`2>=px`1 & px`2<=-px`1 by A17,A18,AXIOMS:25,XCMPLX_1:175;
        then A22:px`2<=px`1 & -px`1<=px`2 or px`2>=px`1 & px`2<=-px`1
                                by A17,A18,AXIOMS:25,XCMPLX_1:175;
        A23:p1=Sq_Circ.px by A13,A14,FUNCT_1:54,JGRAPH_3:32,54;
        A24:Sq_Circ.q=|[q`1/sqrt(1+(q`2/q`1)^2),q`2/sqrt(1+(q`2/q`1)^2)]|
                                 by A19,A22,JGRAPH_2:11,JGRAPH_3:def 1;
        A25: (|[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;
          (q`2/q`1)^2 >=0 by SQUARE_1:72;
        then 1+(q`2/q`1)^2>=1+0 by REAL_1:55;
        then A26:1+(q`2/q`1)^2>0 by AXIOMS:22;
        then A27:sqrt(1+(q`2/q`1)^2)>0 by SQUARE_1:93;
         px`1<>0 by A19,A22;
        then A28: (px`1)^2<>0 by SQUARE_1:73;
          now
          (|.p1.|)^2= (q`1/sqrt(1+(q`2/q`1)^2))^2+(q`2/sqrt(1+(q`2/q`1)^2))^2
                              by A23,A24,A25,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 A26,SQUARE_1:def 4
               .= (q`1)^2/(1+(q`2/q`1)^2)+(q`2)^2/(1+(q`2/q`1)^2)
                              by A26,SQUARE_1:def 4
               .= ((q`1)^2+(q`2)^2)/(1+(q`2/q`1)^2)
                              by XCMPLX_1:63;
        then ((q`1)^2+(q`2)^2)/(1+(q`2/q`1)^2)*(1+(q`2/q`1)^2)=(1)*(1+(q`2/q`1
)^2)
                              by A14,SQUARE_1:59;
        then ((q`1)^2+(q`2)^2)=(1+(q`2/q`1)^2) by A26,XCMPLX_1:88;
         then (px`1)^2+(px`2)^2=1+(px`2)^2/(px`1)^2 by SQUARE_1:69;
         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 by A28,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 XCMPLX_1:14
;
         then (px`1)^2*(px`1)^2+((px`1)^2*(px`2)^2-(px`1)^2*1)-(px`2)^2=0
                                   by XCMPLX_1:40;
         then (px`1)^2*(px`1)^2+(px`1)^2*(px`2)^2-(px`1)^2*1-(px`2)^2=0
                                   by XCMPLX_1:29;
         then (px`1)^2*(px`1)^2-(px`1)^2*1+(px`1)^2*(px`2)^2-(1)*(px`2)^2=0
                                   by XCMPLX_1:29;
         then (px`1)^2*((px`1)^2-1)+(px`1)^2*(px`2)^2-(1)*(px`2)^2=0
                                   by XCMPLX_1:40;
         then (px`1)^2*((px`1)^2-1)+((px`1)^2*(px`2)^2-(1)*(px`2)^2)=0
                                   by XCMPLX_1:29;
         hence ((px`1)^2-1)*(px`1)^2+((px`1)^2-1)*(px`2)^2=0 by XCMPLX_1:40;
         end;
         then A29:((px`1)^2-1)*((px`1)^2+(px`2)^2)=0 by XCMPLX_1:8;
           ((px`1)^2+(px`2)^2)<>0 by A19,COMPLEX1:2;
         then (px`1)^2-1=0 by A29,XCMPLX_1:6;
         then (px`1-1)*(px`1+1)=0 by SQUARE_1:59,67;
         then A30: px`1-1=0 or px`1+1=0 by XCMPLX_1:6;
         A31: now assume A32: px`1=1;
           then A33:p1`1>0 by A23,A24,A25,A27,REAL_2:127;
             -p1`2>=--p1`1 by A14,REAL_1:50;
           then -p1`2>0 by A23,A24,A25,A27,A32,REAL_2:127;
           then --p1`2<-0 by REAL_1:50;
          hence contradiction by A14,A33,AXIOMS:22;
         end;
    A34: (ff.I)=(Sq_Circ").(f.I) by A3,FUNCT_1:22;
    consider p2 being Point of TOP-REAL 2 such that
    A35:  f.I=p2 & (|.p2.|=1 & p2`2<=p2`1 & p2`2>=-p2`1) by A1;
    A36:p2<>0.REAL 2 &
          (p2`2<=p2`1 & -p2`1<=p2`2 or p2`2>=p2`1 & p2`2<=-p2`1)
                            by A35,TOPRNS_1:24;
    then A37:Sq_Circ".p2=|[p2`1*sqrt(1+(p2`2/p2`1)^2),p2`2*sqrt(1+(p2`2/p2`1)^2
)]|
                                 by JGRAPH_3:38;
        reconsider py=ff.I as Point of TOP-REAL 2;
        A38: py`1 = p2`1*sqrt(1+(p2`2/p2`1)^2) &
         py`2 = p2`2*sqrt(1+(p2`2/p2`1)^2) by A34,A35,A37,EUCLID:56;
          (p2`2/p2`1)^2 >=0 by SQUARE_1:72;
        then 1+(p2`2/p2`1)^2>=1+0 by REAL_1:55;
        then 1+(p2`2/p2`1)^2>0 by AXIOMS:22;
        then A39:sqrt(1+(p2`2/p2`1)^2)>0 by SQUARE_1:93;
        A40:now assume
A41:        py`1=0 & py`2=0;
          then A42:p2`1=0 by A38,A39,XCMPLX_1:6;
            p2`2=0 by A38,A39,A41,XCMPLX_1:6;
         hence contradiction by A36,A42,EUCLID:57,58;
        end;
        A43: now
       p2`2<=p2`1 & (-p2`1)*sqrt(1+(p2`2/p2`1)^2) <= p2`2*sqrt(1+(p2`2/p2`1)^2)
           or py`2>=py`1 & py`2<=-py`1 by A35,A39,AXIOMS:25;
           hence
          p2`2*sqrt(1+(p2`2/p2`1)^2) <= p2`1*sqrt(1+(p2`2/p2`1)^2) & -py`1<=py
`2
           or py`2>=py`1 & py`2<=-py`1 by A38,A39,AXIOMS:25,XCMPLX_1:175;
        end;
        A44:p2=Sq_Circ.py by A34,A35,FUNCT_1:54,JGRAPH_3:32,54;
     A45:Sq_Circ.py=|[py`1/sqrt(1+(py`2/py`1)^2),py`2/sqrt(1+(py`2/py`1)^2)]|
                                 by A38,A40,A43,JGRAPH_2:11,JGRAPH_3:def 1;
        A46: (|[py`1/sqrt(1+(py`2/py`1)^2),py`2/sqrt(1+(py`2/py`1)^2)]|)`1
         = py`1/sqrt(1+(py`2/py`1)^2) &
         (|[py`1/sqrt(1+(py`2/py`1)^2),py`2/sqrt(1+(py`2/py`1)^2)]|)`2
         = py`2/sqrt(1+(py`2/py`1)^2) by EUCLID:56;
          (py`2/py`1)^2 >=0 by SQUARE_1:72;
        then 1+(py`2/py`1)^2>=1+0 by REAL_1:55;
        then A47:1+(py`2/py`1)^2>0 by AXIOMS:22;
        then A48:sqrt(1+(py`2/py`1)^2)>0 by SQUARE_1:93;
         py`1<>0 by A38,A40,A43;
        then A49: (py`1)^2<>0 by SQUARE_1:73;
          now
      (|.p2.|)^2= (py`1/sqrt(1+(py`2/py`1)^2))^2+(py`2/sqrt(1+(py`2/py`1)^2))^2
                              by A44,A45,A46,JGRAPH_3:10
         .= (py`1)^2/(sqrt(1+(py`2/py`1)^2))^2+(py`2/sqrt(1+(py`2/py`1)^2))^2
                              by SQUARE_1:69
               .= (py`1)^2/(sqrt(1+(py`2/py`1)^2))^2
                             +(py`2)^2/(sqrt(1+(py`2/py`1)^2))^2
                              by SQUARE_1:69
               .= (py`1)^2/(1+(py`2/py`1)^2)+(py`2)^2/(sqrt(1+(py`2/py`1)^2))^2
                              by A47,SQUARE_1:def 4
               .= (py`1)^2/(1+(py`2/py`1)^2)+(py`2)^2/(1+(py`2/py`1)^2)
                              by A47,SQUARE_1:def 4
               .= ((py`1)^2+(py`2)^2)/(1+(py`2/py`1)^2) by XCMPLX_1:63;
       then ((py`1)^2+(py`2)^2)/(1+(py`2/py`1)^2)*(1+(py`2/py`1)^2)
                =(1)*(1+(py`2/py`1)^2) by A35,SQUARE_1:59;
        then ((py`1)^2+(py`2)^2)=(1+(py`2/py`1)^2)
                              by A47,XCMPLX_1:88;
         then (py`1)^2+(py`2)^2=1+(py`2)^2/(py`1)^2 by SQUARE_1:69;
         then (py`1)^2+(py`2)^2-1=(py`2)^2/(py`1)^2 by XCMPLX_1:26;
         then ((py`1)^2+(py`2)^2-1)*(py`1)^2=(py`2)^2 by A49,XCMPLX_1:88;
         then ((py`1)^2+((py`2)^2-1))*(py`1)^2=(py`2)^2 by XCMPLX_1:29;
         then (py`1)^2*(py`1)^2+((py`2)^2-1)*(py`1)^2=(py`2)^2 by XCMPLX_1:8;
         then (py`1)^2*(py`1)^2+(py`1)^2*((py`2)^2-1)-(py`2)^2=0 by XCMPLX_1:14
;
         then (py`1)^2*(py`1)^2+((py`1)^2*(py`2)^2-(py`1)^2*1)-(py`2)^2=0
                                   by XCMPLX_1:40;
         then (py`1)^2*(py`1)^2+(py`1)^2*(py`2)^2-(py`1)^2*1-(py`2)^2=0
                                   by XCMPLX_1:29;
         then (py`1)^2*(py`1)^2-(py`1)^2*1+(py`1)^2*(py`2)^2-(1)*(py`2)^2=0
                                   by XCMPLX_1:29;
         then (py`1)^2*((py`1)^2-1)+(py`1)^2*(py`2)^2-(1)*(py`2)^2=0
                                   by XCMPLX_1:40;
         then (py`1)^2*((py`1)^2-1)+((py`1)^2*(py`2)^2-(1)*(py`2)^2)=0
                                   by XCMPLX_1:29;
          hence ((py`1)^2-1)*(py`1)^2+((py`1)^2-1)*(py`2)^2=0 by XCMPLX_1:40;
         end;
         then A50:((py`1)^2-1)*((py`1)^2+(py`2)^2)=0 by XCMPLX_1:8;
           ((py`1)^2+(py`2)^2)<>0 by A40,COMPLEX1:2;
         then ((py`1)^2-1)=0 by A50,XCMPLX_1:6;
         then (py`1-1)*(py`1+1)=0 by SQUARE_1:59,67;
         then A51: py`1-1=0 or py`1+1=0 by XCMPLX_1:6;
         A52: now assume A53: py`1=-1;
           then A54:p2`1<0 by A44,A45,A46,A48,REAL_2:128;
             -p2`2<=--p2`1 by A35,REAL_1:50;
           then -p2`2<0 by A44,A45,A46,A48,A53,REAL_2:128;
           then --p2`2>-0 by REAL_1:50;
          hence contradiction by A35,A54,AXIOMS:22;
         end;
    A55: (gg.O)=(Sq_Circ").(g.O) by A4,FUNCT_1:22;
    consider p3 being Point of TOP-REAL 2 such that
    A56:  g.O=p3 &( |.p3.|=1 & p3`2<=p3`1 & p3`2<=-p3`1) by A1;
    A57: -p3`2>=--p3`1 by A56,REAL_1:50;
    then A58:p3<>0.REAL 2 &
          (p3`1<=p3`2 & -p3`2<=p3`1 or p3`1>=p3`2 & p3`1<=-p3`2)
                                 by A56,TOPRNS_1:24;
   then A59:Sq_Circ".p3=|[p3`1*sqrt(1+(p3`1/p3`2)^2),p3`2*sqrt(1+(p3`1/p3`2)^2)
]|
                                 by JGRAPH_3:40;
        reconsider pz=gg.O as Point of TOP-REAL 2;
        A60: pz`2 = p3`2*sqrt(1+(p3`1/p3`2)^2) &
         pz`1 = p3`1*sqrt(1+(p3`1/p3`2)^2) by A55,A56,A59,EUCLID:56;
          (p3`1/p3`2)^2 >=0 by SQUARE_1:72;
        then 1+(p3`1/p3`2)^2>=1+0 by REAL_1:55;
        then 1+(p3`1/p3`2)^2>0 by AXIOMS:22;
        then A61:sqrt(1+(p3`1/p3`2)^2)>0 by SQUARE_1:93;
        A62:now assume
A63:        pz`2=0 & pz`1=0;
          then A64:p3`2=0 by A60,A61,XCMPLX_1:6;
            p3`1=0 by A60,A61,A63,XCMPLX_1:6;
         hence contradiction by A58,A64,EUCLID:57,58;
        end;
        A65: now
          p3`1<=p3`2 & -p3`2<=p3`1 or p3`1>=p3`2 &
        p3`1*sqrt(1+(p3`1/p3`2)^2) <= (-p3`2)*sqrt(1+(p3`1/p3`2)^2)
                             by A56,A57,A61,AXIOMS:25;
     then p3`1<=p3`2 & (-p3`2)*sqrt(1+(p3`1/p3`2)^2) <= p3`1*sqrt(1+(p3`1/p3`2
)^2)
           or pz`1>=pz`2 & pz`1<=-pz`2 by A60,A61,AXIOMS:25,XCMPLX_1:175;
           hence
         p3`1*sqrt(1+(p3`1/p3`2)^2) <= p3`2*sqrt(1+(p3`1/p3`2)^2) & -pz`2<=pz`1
           or pz`1>=pz`2 & pz`1<=-pz`2 by A60,A61,AXIOMS:25,XCMPLX_1:175;
        end;
        A66:p3=Sq_Circ.pz by A55,A56,FUNCT_1:54,JGRAPH_3:32,54;
     A67:Sq_Circ.pz=|[pz`1/sqrt(1+(pz`1/pz`2)^2),pz`2/sqrt(1+(pz`1/pz`2)^2)]|
                                 by A60,A62,A65,JGRAPH_2:11,JGRAPH_3:14;
        A68: (|[pz`1/sqrt(1+(pz`1/pz`2)^2),pz`2/sqrt(1+(pz`1/pz`2)^2)]|)`2
         = pz`2/sqrt(1+(pz`1/pz`2)^2) &
         (|[pz`1/sqrt(1+(pz`1/pz`2)^2),pz`2/sqrt(1+(pz`1/pz`2)^2)]|)`1
         = pz`1/sqrt(1+(pz`1/pz`2)^2) by EUCLID:56;
          (pz`1/pz`2)^2 >=0 by SQUARE_1:72;
        then 1+(pz`1/pz`2)^2>=1+0 by REAL_1:55;
        then A69:1+(pz`1/pz`2)^2>0 by AXIOMS:22;
        then A70:sqrt(1+(pz`1/pz`2)^2)>0 by SQUARE_1:93;
         pz`2<>0 by A60,A62,A65;
        then A71: (pz`2)^2<>0 by SQUARE_1:73;
 A72:(|.p3.|)^2= (pz`2/sqrt(1+(pz`1/pz`2)^2))^2+(pz`1/sqrt(1+(pz`1/pz`2)^2))^2
                              by A66,A67,A68,JGRAPH_3:10
        .= (pz`2)^2/(sqrt(1+(pz`1/pz`2)^2))^2+(pz`1/sqrt(1+(pz`1/pz`2)^2))^2
                              by SQUARE_1:69
               .= (pz`2)^2/(sqrt(1+(pz`1/pz`2)^2))^2
                      +(pz`1)^2/(sqrt(1+(pz`1/pz`2)^2))^2
                              by SQUARE_1:69
               .= (pz`2)^2/(1+(pz`1/pz`2)^2)+(pz`1)^2/(sqrt(1+(pz`1/pz`2)^2))^2
                              by A69,SQUARE_1:def 4
               .= (pz`2)^2/(1+(pz`1/pz`2)^2)+(pz`1)^2/(1+(pz`1/pz`2)^2)
                              by A69,SQUARE_1:def 4
               .= ((pz`2)^2+(pz`1)^2)/(1+(pz`1/pz`2)^2) by XCMPLX_1:63;
          now
   ((pz`2)^2+(pz`1)^2)/(1+(pz`1/pz`2)^2)*(1+(pz`1/pz`2)^2)=(1)*(1+(pz`1/pz`2)^2
)
                           by A56,A72,SQUARE_1:59;
        then ((pz`2)^2+(pz`1)^2)=(1+(pz`1/pz`2)^2) by A69,XCMPLX_1:88;
         then (pz`2)^2+(pz`1)^2=1+(pz`1)^2/(pz`2)^2 by SQUARE_1:69;
         then (pz`2)^2+(pz`1)^2-1=(pz`1)^2/(pz`2)^2 by XCMPLX_1:26;
         then ((pz`2)^2+(pz`1)^2-1)*(pz`2)^2=(pz`1)^2 by A71,XCMPLX_1:88;
         then ((pz`2)^2+((pz`1)^2-1))*(pz`2)^2=(pz`1)^2 by XCMPLX_1:29;
         then (pz`2)^2*(pz`2)^2+((pz`1)^2-1)*(pz`2)^2=(pz`1)^2 by XCMPLX_1:8;
         then (pz`2)^2*(pz`2)^2+(pz`2)^2*((pz`1)^2-1)-(pz`1)^2=0 by XCMPLX_1:14
;
         then (pz`2)^2*(pz`2)^2+((pz`2)^2*(pz`1)^2-(pz`2)^2*1)-(pz`1)^2=0
                                   by XCMPLX_1:40;
         then (pz`2)^2*(pz`2)^2+(pz`2)^2*(pz`1)^2-(pz`2)^2*1-(pz`1)^2=0
                                   by XCMPLX_1:29;
         then (pz`2)^2*(pz`2)^2-(pz`2)^2*1+(pz`2)^2*(pz`1)^2-(1)*(pz`1)^2=0
                                   by XCMPLX_1:29;hence
           (pz`2)^2*((pz`2)^2-1)+(pz`2)^2*(pz`1)^2-(1)*(pz`1)^2=0
                                   by XCMPLX_1:40;
         end;
         then (pz`2)^2*((pz`2)^2-1)+((pz`2)^2*(pz`1)^2-(1)*(pz`1)^2)=0
                                   by XCMPLX_1:29;
         then (pz`2)^2*((pz`2)^2-1)+((pz`2)^2-1)*(pz`1)^2=0 by XCMPLX_1:40;
         then A73:((pz`2)^2-1)*((pz`2)^2+(pz`1)^2)=0 by XCMPLX_1:8;
           ((pz`2)^2+(pz`1)^2)<>0 by A62,COMPLEX1:2;
         then ((pz`2)^2-1)=0 by A73,XCMPLX_1:6;
         then (pz`2-1)*(pz`2+1)=0 by SQUARE_1:59,67;
         then pz`2-1=0 or pz`2+1=0 by XCMPLX_1:6;
         then pz`2=0+1 or pz`2+1=0 by XCMPLX_1:27;
         then A74: pz`2=1 or pz`2=0-1 by XCMPLX_1:26;
         A75: now assume A76: pz`2=1;
           then A77:p3`2>0 by A66,A67,A68,A70,REAL_2:127;
            -p3`1>0 by A56,A66,A67,A68,A70,A76,REAL_2:127;
           then --p3`1<-0 by REAL_1:50;
          hence contradiction by A56,A77,AXIOMS:22;
         end;
    A78: (gg.I)=(Sq_Circ").(g.I) by A4,FUNCT_1:22;
    consider p4 being Point of TOP-REAL 2 such that
    A79:  g.I=p4 &(
    |.p4.|=1 & p4`2>=p4`1 & p4`2>=-p4`1) by A1;
    A80: -p4`2<=--p4`1 by A79,REAL_1:50;
    then A81:p4<>0.REAL 2 & (p4`1<=p4`2 & -p4`2<=p4`1 or p4`1>=p4`2 & p4`1<=-p4
`2)
                      by A79,TOPRNS_1:24;
    then A82:Sq_Circ".p4=|[p4`1*sqrt(1+(p4`1/p4`2)^2),p4`2*sqrt(1+(p4`1/p4`2)^2
)]|
                                 by JGRAPH_3:40;
        reconsider pu=gg.I as Point of TOP-REAL 2;
        A83: pu`2 = p4`2*sqrt(1+(p4`1/p4`2)^2) &
         pu`1 = p4`1*sqrt(1+(p4`1/p4`2)^2) by A78,A79,A82,EUCLID:56;
          (p4`1/p4`2)^2 >=0 by SQUARE_1:72;
        then 1+(p4`1/p4`2)^2>=1+0 by REAL_1:55;
        then 1+(p4`1/p4`2)^2>0 by AXIOMS:22;
        then A84:sqrt(1+(p4`1/p4`2)^2)>0 by SQUARE_1:93;
        A85:now assume
A86:        pu`2=0 & pu`1=0;
          then A87:p4`2=0 by A83,A84,XCMPLX_1:6;
            p4`1=0 by A83,A84,A86,XCMPLX_1:6;
         hence contradiction by A81,A87,EUCLID:57,58;
        end;
        A88: now
       p4`1<=p4`2 & (-p4`2)*sqrt(1+(p4`1/p4`2)^2) <= p4`1*sqrt(1+(p4`1/p4`2)^2)
           or pu`1>=pu`2 & pu`1<=-pu`2 by A79,A80,A84,AXIOMS:25;hence
         p4`1*sqrt(1+(p4`1/p4`2)^2) <= p4`2*sqrt(1+(p4`1/p4`2)^2) & -pu`2<=pu`1
           or pu`1>=pu`2 & pu`1<=-pu`2 by A83,A84,AXIOMS:25,XCMPLX_1:175;
        end;
        A89:p4=Sq_Circ.pu by A78,A79,FUNCT_1:54,JGRAPH_3:32,54;
    A90:Sq_Circ.pu=|[pu`1/sqrt(1+(pu`1/pu`2)^2),pu`2/sqrt(1+(pu`1/pu`2)^2)]|
                                 by A83,A85,A88,JGRAPH_2:11,JGRAPH_3:14;
        A91: (|[pu`1/sqrt(1+(pu`1/pu`2)^2),pu`2/sqrt(1+(pu`1/pu`2)^2)]|)`2
         = pu`2/sqrt(1+(pu`1/pu`2)^2) &
         (|[pu`1/sqrt(1+(pu`1/pu`2)^2),pu`2/sqrt(1+(pu`1/pu`2)^2)]|)`1
         = pu`1/sqrt(1+(pu`1/pu`2)^2) by EUCLID:56;
          (pu`1/pu`2)^2 >=0 by SQUARE_1:72;
        then 1+(pu`1/pu`2)^2>=1+0 by REAL_1:55;
        then A92:1+(pu`1/pu`2)^2>0 by AXIOMS:22;
        then A93:sqrt(1+(pu`1/pu`2)^2)>0 by SQUARE_1:93;
         pu`2<>0 by A83,A85,A88;
        then A94: (pu`2)^2<>0 by SQUARE_1:73;
          now
      (|.p4.|)^2= (pu`2/sqrt(1+(pu`1/pu`2)^2))^2+(pu`1/sqrt(1+(pu`1/pu`2)^2))^2
                              by A89,A90,A91,JGRAPH_3:10
        .= (pu`2)^2/(sqrt(1+(pu`1/pu`2)^2))^2+(pu`1/sqrt(1+(pu`1/pu`2)^2))^2
                              by SQUARE_1:69
               .= (pu`2)^2/(sqrt(1+(pu`1/pu`2)^2))^2
                             +(pu`1)^2/(sqrt(1+(pu`1/pu`2)^2))^2
                              by SQUARE_1:69
            .= (pu`2)^2/(1+(pu`1/pu`2)^2)+(pu`1)^2/(sqrt(1+(pu`1/pu`2)^2))^2
                              by A92,SQUARE_1:def 4
               .= (pu`2)^2/(1+(pu`1/pu`2)^2)+(pu`1)^2/(1+(pu`1/pu`2)^2)
                              by A92,SQUARE_1:def 4
               .= ((pu`2)^2+(pu`1)^2)/(1+(pu`1/pu`2)^2) by XCMPLX_1:63;
  then ((pu`2)^2+(pu`1)^2)/(1+(pu`1/pu`2)^2)*(1+(pu`1/pu`2)^2)=(1)*(1+(pu`1/pu
`2)^2)
                               by A79,SQUARE_1:59;
        then ((pu`2)^2+(pu`1)^2)=(1+(pu`1/pu`2)^2)
                              by A92,XCMPLX_1:88;
         then (pu`2)^2+(pu`1)^2=1+(pu`1)^2/(pu`2)^2 by SQUARE_1:69;
         then (pu`2)^2+(pu`1)^2-1=(pu`1)^2/(pu`2)^2 by XCMPLX_1:26;
         then ((pu`2)^2+(pu`1)^2-1)*(pu`2)^2=(pu`1)^2
                            by A94,XCMPLX_1:88;
         then ((pu`2)^2+((pu`1)^2-1))*(pu`2)^2=(pu`1)^2 by XCMPLX_1:29;
         then (pu`2)^2*(pu`2)^2+((pu`1)^2-1)*(pu`2)^2=(pu`1)^2 by XCMPLX_1:8;
         then (pu`2)^2*(pu`2)^2+(pu`2)^2*((pu`1)^2-1)-(pu`1)^2=0 by XCMPLX_1:14
;
         then (pu`2)^2*(pu`2)^2+((pu`2)^2*(pu`1)^2-(pu`2)^2*1)-(pu`1)^2=0
                                   by XCMPLX_1:40;
         then (pu`2)^2*(pu`2)^2+(pu`2)^2*(pu`1)^2-(pu`2)^2*1-(pu`1)^2=0
                                   by XCMPLX_1:29;
         then (pu`2)^2*(pu`2)^2-(pu`2)^2*1+(pu`2)^2*(pu`1)^2-(1)*(pu`1)^2=0
                                   by XCMPLX_1:29;
         then (pu`2)^2*((pu`2)^2-1)+(pu`2)^2*(pu`1)^2-(1)*(pu`1)^2=0
                                   by XCMPLX_1:40;
         then (pu`2)^2*((pu`2)^2-1)+((pu`2)^2*(pu`1)^2-(1)*(pu`1)^2)=0
                                   by XCMPLX_1:29;
          hence ((pu`2)^2-1)*(pu`2)^2+((pu`2)^2-1)*(pu`1)^2=0 by XCMPLX_1:40;
         end;
         then A95:((pu`2)^2-1)*((pu`2)^2+(pu`1)^2)=0 by XCMPLX_1:8;
           ((pu`2)^2+(pu`1)^2)<>0 by A85,COMPLEX1:2;
         then ((pu`2)^2-1)=0 by A95,XCMPLX_1:6;
         then (pu`2-1)*(pu`2+1)=0 by SQUARE_1:59,67;
         then pu`2-1=0 or pu`2+1=0 by XCMPLX_1:6;
         then A96: pu`2=0+1 or pu`2+1=0 by XCMPLX_1:27;
           now assume A97: pu`2=-1;
           then A98:p4`2<0 by A89,A90,A91,A93,REAL_2:128;
            -p4`1<0 by A79,A89,A90,A91,A93,A97,REAL_2:128;
           then --p4`1>-0 by REAL_1:50;
          hence contradiction by A79,A98,AXIOMS:22;
         end;
     hence thesis by A30,A31,A51,A52,A74,A75,A96,XCMPLX_1:26,27;
     end;
     A99: for r being Point of I[01] holds
        (-1>=(ff.r)`1 or (ff.r)`1>=1 or -1 >=(ff.r)`2 or (ff.r)`2>=1)
      & (-1>=(gg.r)`1 or (gg.r)`1>=1 or -1 >=(gg.r)`2 or (gg.r)`2>=1)
      proof let r be Point of I[01];
        A100: (ff.r)=(Sq_Circ").(f.r) by A3,FUNCT_1:22;
          f.r in rng f by A5,FUNCT_1:12;
        then f.r in C0 by A1;
        then consider p1 being Point of TOP-REAL 2 such that
        A101:  f.r=p1 & |.p1.|>=1 by A1;
       A102:now per cases;
       case A103: p1=0.REAL 2;
           |.0.REAL 2.|=0 by TOPRNS_1:24;
        hence contradiction by A101,A103;
       case
        A104:p1<>0.REAL 2 &
          (p1`2<=p1`1 & -p1`1<=p1`2 or p1`2>=p1`1 & p1`2<=-p1`1);
   then A105:Sq_Circ".p1=|[p1`1*sqrt(1+(p1`2/p1`1)^2),p1`2*sqrt(1+(p1`2/p1`1)^2
)]|
                                 by JGRAPH_3:38;
        reconsider px=ff.r as Point of TOP-REAL 2;
        set q=px;
        A106: px`1 = p1`1*sqrt(1+(p1`2/p1`1)^2) &
         px`2 = p1`2*sqrt(1+(p1`2/p1`1)^2) by A100,A101,A105,EUCLID:56;
          (p1`2/p1`1)^2 >=0 by SQUARE_1:72;
        then 1+(p1`2/p1`1)^2>=1+0 by REAL_1:55;
        then 1+(p1`2/p1`1)^2>0 by AXIOMS:22;
        then A107:sqrt(1+(p1`2/p1`1)^2)>0 by SQUARE_1:93;
        A108:now assume
A109:        px`1=0 & px`2=0;
          then A110:p1`1=0 by A106,A107,XCMPLX_1:6;
            p1`2=0 by A106,A107,A109,XCMPLX_1:6;
         hence contradiction by A104,A110,EUCLID:57,58;
        end;
        A111: now
          p1`2<=p1`1 & -p1`1<=p1`2 or p1`2>=p1`1 &
        p1`2*sqrt(1+(p1`2/p1`1)^2) <= (-p1`1)*sqrt(1+(p1`2/p1`1)^2)
                             by A104,A107,AXIOMS:25;
     then p1`2<=p1`1 & (-p1`1)*sqrt(1+(p1`2/p1`1)^2) <= p1`2*sqrt(1+(p1`2/p1`1
)^2)
           or px`2>=px`1 & px`2<=-px`1 by A106,A107,AXIOMS:25,XCMPLX_1:175;
       hence
         p1`2*sqrt(1+(p1`2/p1`1)^2) <= p1`1*sqrt(1+(p1`2/p1`1)^2) & -px`1<=px`2
           or px`2>=px`1 & px`2<=-px`1 by A106,A107,AXIOMS:25,XCMPLX_1:175;
        end;
        A112:p1=Sq_Circ.px by A100,A101,FUNCT_1:54,JGRAPH_3:32,54;
        A113:Sq_Circ.q=|[q`1/sqrt(1+(q`2/q`1)^2),q`2/sqrt(1+(q`2/q`1)^2)]|
                                 by A106,A108,A111,JGRAPH_2:11,JGRAPH_3:def 1;
        A114: (|[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;
        A115:(px`1)^2 >=0 by SQUARE_1:72;
          (q`2/q`1)^2 >=0 by SQUARE_1:72;
        then 1+(q`2/q`1)^2>=1+0 by REAL_1:55;
        then A116:1+(q`2/q`1)^2>0 by AXIOMS:22;
          (|.p1.|)^2>=|.p1.| by A101,JGRAPH_2:2;
        then A117: (|.p1.|)^2>=1 by A101,AXIOMS:22;
         px`1<>0 by A106,A108,A111;
        then A118: (px`1)^2<>0 by SQUARE_1:73;
          now
          (|.p1.|)^2= (q`1/sqrt(1+(q`2/q`1)^2))^2+(q`2/sqrt(1+(q`2/q`1)^2))^2
                              by A112,A113,A114,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 A116,SQUARE_1:def 4
               .= (q`1)^2/(1+(q`2/q`1)^2)+(q`2)^2/(1+(q`2/q`1)^2)
                              by A116,SQUARE_1:def 4
               .= ((q`1)^2+(q`2)^2)/(1+(q`2/q`1)^2)
                              by XCMPLX_1:63;
        then ((q`1)^2+(q`2)^2)/(1+(q`2/q`1)^2)*(1+(q`2/q`1)^2)>=(1)*(1+(q`2/q
`1)^2)
                              by A116,A117,AXIOMS:25;
        then ((q`1)^2+(q`2)^2)>=(1+(q`2/q`1)^2)
                              by A116,XCMPLX_1:88;
         then (px`1)^2+(px`2)^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 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 A115,AXIOMS:25;
         then ((px`1)^2+(px`2)^2-1)*(px`1)^2>=(px`2)^2
                            by A118,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>=(px`2)^2-(px`2
)^2
                                   by REAL_1:49;
         then (px`1)^2*(px`1)^2+(px`1)^2*((px`2)^2-1)-(px`2)^2>=0 by XCMPLX_1:
14
;
         then (px`1)^2*(px`1)^2+((px`1)^2*(px`2)^2-(px`1)^2*1)-(px`2)^2>=0
                                   by XCMPLX_1:40;
         then (px`1)^2*(px`1)^2+(px`1)^2*(px`2)^2-(px`1)^2*1-(px`2)^2>=0
                                   by XCMPLX_1:29;
         then (px`1)^2*(px`1)^2-(px`1)^2*1+(px`1)^2*(px`2)^2-(1)*(px`2)^2>=0
                                   by XCMPLX_1:29;
         then (px`1)^2*((px`1)^2-1)+(px`1)^2*(px`2)^2-(1)*(px`2)^2>=0
                                   by XCMPLX_1:40;
         then (px`1)^2*((px`1)^2-1)+((px`1)^2*(px`2)^2-(1)*(px`2)^2)>=0
                                   by XCMPLX_1:29;hence
           ((px`1)^2-1)*(px`1)^2+((px`1)^2-1)*(px`2)^2>=0 by XCMPLX_1:40;
         end;
         then A119:((px`1)^2-1)*((px`1)^2+(px`2)^2)>=0 by XCMPLX_1:8;
         A120: ((px`1)^2+(px`2)^2)<>0 by A108,COMPLEX1:2;
           (px`2)^2>=0 by SQUARE_1:72;
         then ((px`1)^2+(px`2)^2)>=0+0 by A115,REAL_1:55;
         then ((px`1)^2-1)>=0 by A119,A120,SQUARE_1:25;
         then (px`1-1)*(px`1+1)>=0 by SQUARE_1:59,67;
        hence -1>=(ff.r)`1 or (ff.r)`1>=1 or
         -1 >=(ff.r)`2 or (ff.r)`2>=1 by Th1;
       case
        A121:p1<>0.REAL 2 &
          not(p1`2<=p1`1 & -p1`1<=p1`2 or p1`2>=p1`1 & p1`2<=-p1`1);

    then A122:Sq_Circ".p1=|[p1`1*sqrt(1+(p1`1/p1`2)^2),p1`2*sqrt(1+(p1`1/p1`2)
^2)]|
                                 by JGRAPH_3:38;
        A123:p1<>0.REAL 2 &
          (p1`1<=p1`2 & -p1`2<=p1`1 or p1`1>=p1`2 & p1`1<=-p1`2)
                      by A121,JGRAPH_2:23;
        reconsider pz=ff.r as Point of TOP-REAL 2;
        A124: pz`2 = p1`2*sqrt(1+(p1`1/p1`2)^2) &
         pz`1 = p1`1*sqrt(1+(p1`1/p1`2)^2) by A100,A101,A122,EUCLID:56;
          (p1`1/p1`2)^2 >=0 by SQUARE_1:72;
        then 1+(p1`1/p1`2)^2>=1+0 by REAL_1:55;
        then 1+(p1`1/p1`2)^2>0 by AXIOMS:22;
        then A125:sqrt(1+(p1`1/p1`2)^2)>0 by SQUARE_1:93;
        A126:now assume
A127:        pz`2=0 & pz`1=0;
          then p1`2=0 by A124,A125,XCMPLX_1:6;
         hence contradiction by A121,A124,A125,A127,XCMPLX_1:6;
        end;
        A128: now
          p1`1<=p1`2 & -p1`2<=p1`1 or p1`1>=p1`2 &
        p1`1*sqrt(1+(p1`1/p1`2)^2) <= (-p1`2)*sqrt(1+(p1`1/p1`2)^2)
                             by A123,A125,AXIOMS:25;
     then p1`1<=p1`2 & (-p1`2)*sqrt(1+(p1`1/p1`2)^2) <= p1`1*sqrt(1+(p1`1/p1`2
)^2)
           or pz`1>=pz`2 & pz`1<=-pz`2 by A124,A125,AXIOMS:25,XCMPLX_1:175;
           hence
         p1`1*sqrt(1+(p1`1/p1`2)^2) <= p1`2*sqrt(1+(p1`1/p1`2)^2) & -pz`2<=pz`1
           or pz`1>=pz`2 & pz`1<=-pz`2 by A124,A125,AXIOMS:25,XCMPLX_1:175;
        end;
        A129:p1=Sq_Circ.pz by A100,A101,FUNCT_1:54,JGRAPH_3:32,54;
     A130:Sq_Circ.pz=|[pz`1/sqrt(1+(pz`1/pz`2)^2),pz`2/sqrt(1+(pz`1/pz`2)^2)]|
                                 by A124,A126,A128,JGRAPH_2:11,JGRAPH_3:14;
        A131: (|[pz`1/sqrt(1+(pz`1/pz`2)^2),pz`2/sqrt(1+(pz`1/pz`2)^2)]|)`2
         = pz`2/sqrt(1+(pz`1/pz`2)^2) &
         (|[pz`1/sqrt(1+(pz`1/pz`2)^2),pz`2/sqrt(1+(pz`1/pz`2)^2)]|)`1
         = pz`1/sqrt(1+(pz`1/pz`2)^2) by EUCLID:56;
        A132:(pz`2)^2 >=0 by SQUARE_1:72;
          (pz`1/pz`2)^2 >=0 by SQUARE_1:72;
        then 1+(pz`1/pz`2)^2>=1+0 by REAL_1:55;
        then A133:1+(pz`1/pz`2)^2>0 by AXIOMS:22;
          (|.p1.|)^2>=|.p1.| by A101,JGRAPH_2:2;
        then A134: (|.p1.|)^2>=1 by A101,AXIOMS:22;
         pz`2<>0 by A124,A126,A128;
        then A135: (pz`2)^2<>0 by SQUARE_1:73;
 A136:(|.p1.|)^2= (pz`2/sqrt(1+(pz`1/pz`2)^2))^2+(pz`1/sqrt(1+(pz`1/pz`2)^2))^2
                              by A129,A130,A131,JGRAPH_3:10
        .= (pz`2)^2/(sqrt(1+(pz`1/pz`2)^2))^2+(pz`1/sqrt(1+(pz`1/pz`2)^2))^2
                              by SQUARE_1:69
               .= (pz`2)^2/(sqrt(1+(pz`1/pz`2)^2))^2
                      +(pz`1)^2/(sqrt(1+(pz`1/pz`2)^2))^2
                              by SQUARE_1:69
               .= (pz`2)^2/(1+(pz`1/pz`2)^2)+(pz`1)^2/(sqrt(1+(pz`1/pz`2)^2))^2
                              by A133,SQUARE_1:def 4
               .= (pz`2)^2/(1+(pz`1/pz`2)^2)+(pz`1)^2/(1+(pz`1/pz`2)^2)
                              by A133,SQUARE_1:def 4
               .= ((pz`2)^2+(pz`1)^2)/(1+(pz`1/pz`2)^2)
                              by XCMPLX_1:63;
          now
          ((pz`2)^2+(pz`1)^2)/(1+(pz`1/pz`2)^2)*(1+(pz`1/pz`2)^2)
                           >=(1)*(1+(pz`1/pz`2)^2)
                              by A133,A134,A136,AXIOMS:25;
        then ((pz`2)^2+(pz`1)^2)>=(1+(pz`1/pz`2)^2)
                              by A133,XCMPLX_1:88;
         then (pz`2)^2+(pz`1)^2>=1+(pz`1)^2/(pz`2)^2 by SQUARE_1:69;
         then (pz`2)^2+(pz`1)^2-1>=1+(pz`1)^2/(pz`2)^2-1 by REAL_1:49;
         then (pz`2)^2+(pz`1)^2-1>=(pz`1)^2/(pz`2)^2 by XCMPLX_1:26;
         then ((pz`2)^2+(pz`1)^2-1)*(pz`2)^2>=(pz`1)^2/(pz`2)^2*(pz`2)^2
                              by A132,AXIOMS:25;
         then ((pz`2)^2+(pz`1)^2-1)*(pz`2)^2>=(pz`1)^2
                            by A135,XCMPLX_1:88;
         then ((pz`2)^2+((pz`1)^2-1))*(pz`2)^2>=(pz`1)^2 by XCMPLX_1:29;
         then (pz`2)^2*(pz`2)^2+((pz`1)^2-1)*(pz`2)^2>=(pz`1)^2 by XCMPLX_1:8;
         then (pz`2)^2*(pz`2)^2+(pz`2)^2*((pz`1)^2-1)-(pz`1)^2
                          >=(pz`1)^2-(pz`1)^2 by REAL_1:49;
         then (pz`2)^2*(pz`2)^2+(pz`2)^2*((pz`1)^2-1)-(pz`1)^2>=0 by XCMPLX_1:
14
;
         then (pz`2)^2*(pz`2)^2+((pz`2)^2*(pz`1)^2-(pz`2)^2*1)-(pz`1)^2>=0
                                   by XCMPLX_1:40;
         then (pz`2)^2*(pz`2)^2+(pz`2)^2*(pz`1)^2-(pz`2)^2*1-(pz`1)^2>=0
                                   by XCMPLX_1:29;
         then (pz`2)^2*(pz`2)^2-(pz`2)^2*1+(pz`2)^2*(pz`1)^2-(1)*(pz`1)^2>=0
                                   by XCMPLX_1:29;
         hence (pz`2)^2*((pz`2)^2-1)+(pz`2)^2*(pz`1)^2-(1)*(pz`1)^2>=0
                                   by XCMPLX_1:40;
         end;
         then (pz`2)^2*((pz`2)^2-1)+((pz`2)^2*(pz`1)^2-(1)*(pz`1)^2)>=0
                                   by XCMPLX_1:29;
         then (pz`2)^2*((pz`2)^2-1)+((pz`2)^2-1)*(pz`1)^2>=0
                                   by XCMPLX_1:40;
         then A137:((pz`2)^2-1)*((pz`2)^2+(pz`1)^2)>=0 by XCMPLX_1:8;
         A138: ((pz`2)^2+(pz`1)^2)<>0 by A126,COMPLEX1:2;
           (pz`1)^2>=0 by SQUARE_1:72;
         then (pz`2)^2+(pz`1)^2>=0+0 by A132,REAL_1:55;
         then (pz`2)^2-1>=0 by A137,A138,SQUARE_1:25;
         then (pz`2-1)*(pz`2+1)>=0 by SQUARE_1:59,67;
        hence -1>=(ff.r)`1 or (ff.r)`1>=1 or
         -1 >=(ff.r)`2 or (ff.r)`2>=1 by Th1;
       end;
        A139: (gg.r)=(Sq_Circ").(g.r) by A4,FUNCT_1:22;
          g.r in rng g by A6,FUNCT_1:12;
        then g.r in C0 by A1;
        then consider p2 being Point of TOP-REAL 2 such that
        A140:  g.r=p2 & |.p2.|>=1 by A1;
         now per cases;
       case A141: p2=0.REAL 2;
           |.0.REAL 2.|=0 by TOPRNS_1:24;
        hence contradiction by A140,A141;
       case
        A142:p2<>0.REAL 2 &
          (p2`2<=p2`1 & -p2`1<=p2`2 or p2`2>=p2`1 & p2`2<=-p2`1);
   then A143:Sq_Circ".p2=|[p2`1*sqrt(1+(p2`2/p2`1)^2),p2`2*sqrt(1+(p2`2/p2`1)^2
)]|
                                 by JGRAPH_3:38;
        reconsider px=gg.r as Point of TOP-REAL 2;
        set q=px;
        A144:Sq_Circ".p2=q by A4,A140,FUNCT_1:22;
        A145: px`1 = p2`1*sqrt(1+(p2`2/p2`1)^2) &
         px`2 = p2`2*sqrt(1+(p2`2/p2`1)^2) by A139,A140,A143,EUCLID:56;
          (p2`2/p2`1)^2 >=0 by SQUARE_1:72;
        then 1+(p2`2/p2`1)^2>=1+0 by REAL_1:55;
        then 1+(p2`2/p2`1)^2>0 by AXIOMS:22;
        then A146:sqrt(1+(p2`2/p2`1)^2)>0 by SQUARE_1:93;
        A147:now assume
A148:        px`1=0 & px`2=0;
          then A149:p2`1=0 by A145,A146,XCMPLX_1:6;
            p2`2=0 by A145,A146,A148,XCMPLX_1:6;
         hence contradiction by A142,A149,EUCLID:57,58;
        end;
        A150: now
          p2`2<=p2`1 & -p2`1<=p2`2 or p2`2>=p2`1 &
        p2`2*sqrt(1+(p2`2/p2`1)^2) <= (-p2`1)*sqrt(1+(p2`2/p2`1)^2)
                             by A142,A146,AXIOMS:25;
     then p2`2<=p2`1 & (-p2`1)*sqrt(1+(p2`2/p2`1)^2) <= p2`2*sqrt(1+(p2`2/p2`1
)^2)
           or px`2>=px`1 & px`2<=-px`1 by A145,A146,AXIOMS:25,XCMPLX_1:175
;hence
          p2`2*sqrt(1+(p2`2/p2`1)^2) <= p2`1*sqrt(1+(p2`2/p2`1)^2) & -px`1<=px
`2
           or px`2>=px`1 & px`2<=-px`1 by A145,A146,AXIOMS:25,XCMPLX_1:175;
        end;
        A151:p2=Sq_Circ.px by A144,FUNCT_1:54,JGRAPH_3:32,54;
        A152:Sq_Circ.q=|[q`1/sqrt(1+(q`2/q`1)^2),q`2/sqrt(1+(q`2/q`1)^2)]|
                                 by A145,A147,A150,JGRAPH_2:11,JGRAPH_3:def 1;
        A153: (|[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;
        A154:(px`1)^2 >=0 by SQUARE_1:72;
          (q`2/q`1)^2 >=0 by SQUARE_1:72;
        then 1+(q`2/q`1)^2>=1+0 by REAL_1:55;
        then A155:1+(q`2/q`1)^2>0 by AXIOMS:22;
          (|.p2.|)^2>=|.p2.| by A140,JGRAPH_2:2;
        then A156: (|.p2.|)^2>=1 by A140,AXIOMS:22;
         px`1<>0 by A145,A147,A150;
        then A157: (px`1)^2<>0 by SQUARE_1:73;
          now
          (|.p2.|)^2= (q`1/sqrt(1+(q`2/q`1)^2))^2+(q`2/sqrt(1+(q`2/q`1)^2))^2
                              by A151,A152,A153,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 A155,SQUARE_1:def 4
               .= (q`1)^2/(1+(q`2/q`1)^2)+(q`2)^2/(1+(q`2/q`1)^2)
                              by A155,SQUARE_1:def 4
               .= ((q`1)^2+(q`2)^2)/(1+(q`2/q`1)^2)
                              by XCMPLX_1:63;
        then ((q`1)^2+(q`2)^2)/(1+(q`2/q`1)^2)*(1+(q`2/q`1)^2)>=(1)*(1+(q`2/q
`1)^2)
                              by A155,A156,AXIOMS:25;
        then ((q`1)^2+(q`2)^2)>=(1+(q`2/q`1)^2)
                              by A155,XCMPLX_1:88;
         then (px`1)^2+(px`2)^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 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 A154,AXIOMS:25;
         then ((px`1)^2+(px`2)^2-1)*(px`1)^2>=(px`2)^2
                            by A157,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>=(px`2)^2-(px`2
)^2
                                   by REAL_1:49;
         then (px`1)^2*(px`1)^2+(px`1)^2*((px`2)^2-1)-(px`2)^2>=0 by XCMPLX_1:
14
;
         then (px`1)^2*(px`1)^2+((px`1)^2*(px`2)^2-(px`1)^2*1)-(px`2)^2>=0
                                   by XCMPLX_1:40;
         then (px`1)^2*(px`1)^2+(px`1)^2*(px`2)^2-(px`1)^2*1-(px`2)^2>=0
                                   by XCMPLX_1:29;
         then (px`1)^2*(px`1)^2-(px`1)^2*1+(px`1)^2*(px`2)^2-(1)*(px`2)^2>=0
                                   by XCMPLX_1:29;
         then (px`1)^2*((px`1)^2-1)+(px`1)^2*(px`2)^2-(1)*(px`2)^2>=0
                                   by XCMPLX_1:40;
         then (px`1)^2*((px`1)^2-1)+((px`1)^2*(px`2)^2-(1)*(px`2)^2)>=0
                                   by XCMPLX_1:29;hence
           ((px`1)^2-1)*(px`1)^2+((px`1)^2-1)*(px`2)^2>=0 by XCMPLX_1:40;
         end;
         then A158:((px`1)^2-1)*((px`1)^2+(px`2)^2)>=0 by XCMPLX_1:8;
         A159: ((px`1)^2+(px`2)^2)<>0 by A147,COMPLEX1:2;
           (px`2)^2>=0 by SQUARE_1:72;
         then ((px`1)^2+(px`2)^2)>=0+0 by A154,REAL_1:55;
         then ((px`1)^2-1)>=0 by A158,A159,SQUARE_1:25;
         then (px`1-1)*(px`1+1)>=0 by SQUARE_1:59,67;
        hence -1>=(gg.r)`1 or (gg.r)`1>=1 or -1 >=(gg.r)`2 or (gg.r)`2>=1
                    by Th1;
       case
        A160:p2<>0.REAL 2 &
          not(p2`2<=p2`1 & -p2`1<=p2`2 or p2`2>=p2`1 & p2`2<=-p2`1);
   then A161:Sq_Circ".p2=|[p2`1*sqrt(1+(p2`1/p2`2)^2),p2`2*sqrt(1+(p2`1/p2`2)^2
)]|
                                 by JGRAPH_3:38;
        A162:p2<>0.REAL 2 &
          (p2`1<=p2`2 & -p2`2<=p2`1 or p2`1>=p2`2 & p2`1<=-p2`2)
                      by A160,JGRAPH_2:23;
        reconsider pz=gg.r as Point of TOP-REAL 2;
        A163: pz`2 = p2`2*sqrt(1+(p2`1/p2`2)^2) &
         pz`1 = p2`1*sqrt(1+(p2`1/p2`2)^2) by A139,A140,A161,EUCLID:56;
          (p2`1/p2`2)^2 >=0 by SQUARE_1:72;
        then 1+(p2`1/p2`2)^2>=1+0 by REAL_1:55;
        then 1+(p2`1/p2`2)^2>0 by AXIOMS:22;
        then A164:sqrt(1+(p2`1/p2`2)^2)>0 by SQUARE_1:93;
        A165:now assume
A166:        pz`2=0 & pz`1=0;
          then p2`2=0 by A163,A164,XCMPLX_1:6;
         hence contradiction by A160,A163,A164,A166,XCMPLX_1:6;
        end;
        A167: now
          p2`1<=p2`2 & -p2`2<=p2`1 or p2`1>=p2`2 &
        p2`1*sqrt(1+(p2`1/p2`2)^2) <= (-p2`2)*sqrt(1+(p2`1/p2`2)^2)
                             by A162,A164,AXIOMS:25;
    then p2`1<=p2`2 & (-p2`2)*sqrt(1+(p2`1/p2`2)^2) <= p2`1*sqrt(1+(p2`1/p2`2)
^2)
           or pz`1>=pz`2 & pz`1<=-pz`2 by A163,A164,AXIOMS:25,XCMPLX_1:175;
           hence
         p2`1*sqrt(1+(p2`1/p2`2)^2) <= p2`2*sqrt(1+(p2`1/p2`2)^2) & -pz`2<=pz`1
           or pz`1>=pz`2 & pz`1<=-pz`2 by A163,A164,AXIOMS:25,XCMPLX_1:175;
        end;
        A168:p2=Sq_Circ.pz by A139,A140,FUNCT_1:54,JGRAPH_3:32,54;
    A169:Sq_Circ.pz=|[pz`1/sqrt(1+(pz`1/pz`2)^2),pz`2/sqrt(1+(pz`1/pz`2)^2)]|
                                 by A163,A165,A167,JGRAPH_2:11,JGRAPH_3:14;
        A170: (|[pz`1/sqrt(1+(pz`1/pz`2)^2),pz`2/sqrt(1+(pz`1/pz`2)^2)]|)`2
         = pz`2/sqrt(1+(pz`1/pz`2)^2) &
         (|[pz`1/sqrt(1+(pz`1/pz`2)^2),pz`2/sqrt(1+(pz`1/pz`2)^2)]|)`1
         = pz`1/sqrt(1+(pz`1/pz`2)^2) by EUCLID:56;
        A171:(pz`2)^2 >=0 by SQUARE_1:72;
          (pz`1/pz`2)^2 >=0 by SQUARE_1:72;
        then 1+(pz`1/pz`2)^2>=1+0 by REAL_1:55;
        then A172:1+(pz`1/pz`2)^2>0 by AXIOMS:22;
          (|.p2.|)^2>=|.p2.| by A140,JGRAPH_2:2;
        then A173: (|.p2.|)^2>=1 by A140,AXIOMS:22;
         pz`2<>0 by A163,A165,A167;
        then A174: (pz`2)^2<>0 by SQUARE_1:73;
 A175:(|.p2.|)^2= (pz`2/sqrt(1+(pz`1/pz`2)^2))^2+(pz`1/sqrt(1+(pz`1/pz`2)^2))^2
                              by A168,A169,A170,JGRAPH_3:10
        .= (pz`2)^2/(sqrt(1+(pz`1/pz`2)^2))^2+(pz`1/sqrt(1+(pz`1/pz`2)^2))^2
                              by SQUARE_1:69
               .= (pz`2)^2/(sqrt(1+(pz`1/pz`2)^2))^2
                      +(pz`1)^2/(sqrt(1+(pz`1/pz`2)^2))^2
                              by SQUARE_1:69
               .= (pz`2)^2/(1+(pz`1/pz`2)^2)+(pz`1)^2/(sqrt(1+(pz`1/pz`2)^2))^2
                              by A172,SQUARE_1:def 4
               .= (pz`2)^2/(1+(pz`1/pz`2)^2)+(pz`1)^2/(1+(pz`1/pz`2)^2)
                              by A172,SQUARE_1:def 4
               .= ((pz`2)^2+(pz`1)^2)/(1+(pz`1/pz`2)^2)
                              by XCMPLX_1:63;
          now
          ((pz`2)^2+(pz`1)^2)/(1+(pz`1/pz`2)^2)*(1+(pz`1/pz`2)^2)
                           >=(1)*(1+(pz`1/pz`2)^2)
                              by A172,A173,A175,AXIOMS:25;
        then ((pz`2)^2+(pz`1)^2)>=(1+(pz`1/pz`2)^2) by A172,XCMPLX_1:88;
         then (pz`2)^2+(pz`1)^2>=1+(pz`1)^2/(pz`2)^2 by SQUARE_1:69;
         then (pz`2)^2+(pz`1)^2-1>=1+(pz`1)^2/(pz`2)^2-1 by REAL_1:49;
         then (pz`2)^2+(pz`1)^2-1>=(pz`1)^2/(pz`2)^2 by XCMPLX_1:26;
         then ((pz`2)^2+(pz`1)^2-1)*(pz`2)^2>=(pz`1)^2/(pz`2)^2*(pz`2)^2
                              by A171,AXIOMS:25;
         then ((pz`2)^2+(pz`1)^2-1)*(pz`2)^2>=(pz`1)^2 by A174,XCMPLX_1:88;
         then ((pz`2)^2+((pz`1)^2-1))*(pz`2)^2>=(pz`1)^2 by XCMPLX_1:29;
         then (pz`2)^2*(pz`2)^2+((pz`1)^2-1)*(pz`2)^2>=(pz`1)^2 by XCMPLX_1:8;
         then (pz`2)^2*(pz`2)^2+(pz`2)^2*((pz`1)^2-1)-(pz`1)^2
                          >=(pz`1)^2-(pz`1)^2 by REAL_1:49;
         then (pz`2)^2*(pz`2)^2+(pz`2)^2*((pz`1)^2-1)-(pz`1)^2>=0 by XCMPLX_1:
14
;
         then (pz`2)^2*(pz`2)^2+((pz`2)^2*(pz`1)^2-(pz`2)^2*1)-(pz`1)^2>=0
                                   by XCMPLX_1:40;
         then (pz`2)^2*(pz`2)^2+(pz`2)^2*(pz`1)^2-(pz`2)^2*1-(pz`1)^2>=0
                                   by XCMPLX_1:29;
         then (pz`2)^2*(pz`2)^2-(pz`2)^2*1+(pz`2)^2*(pz`1)^2-(1)*(pz`1)^2>=0
                                   by XCMPLX_1:29;hence
           (pz`2)^2*((pz`2)^2-1)+(pz`2)^2*(pz`1)^2-(1)*(pz`1)^2>=0
                                   by XCMPLX_1:40;
         end;
         then (pz`2)^2*((pz`2)^2-1)+((pz`2)^2*(pz`1)^2-(1)*(pz`1)^2)>=0
                                   by XCMPLX_1:29;
         then (pz`2)^2*((pz`2)^2-1)+((pz`2)^2-1)*(pz`1)^2>=0
                                   by XCMPLX_1:40;
         then A176:((pz`2)^2-1)*((pz`2)^2+(pz`1)^2)>=0 by XCMPLX_1:8;
         A177: ((pz`2)^2+(pz`1)^2)<>0 by A165,COMPLEX1:2;
           (pz`1)^2>=0 by SQUARE_1:72;
         then ((pz`2)^2+(pz`1)^2)>=0+0 by A171,REAL_1:55;
         then ((pz`2)^2-1)>=0 by A176,A177,SQUARE_1:25;
         then (pz`2-1)*(pz`2+1)>=0 by SQUARE_1:59,67;
        hence -1>=(gg.r)`1 or (gg.r)`1>=1 or
         -1 >=(gg.r)`2 or (gg.r)`2>=1 by Th1;
       end;
       hence
          (-1>=(ff.r)`1 or (ff.r)`1>=1 or -1 >=(ff.r)`2 or (ff.r)`2>=1)
        & (-1>=(gg.r)`1 or (gg.r)`1>=1 or -1 >=(gg.r)`2 or (gg.r)`2>=1)
                                     by A102;
      end;
        (-1 <=(ff.O)`2 & (ff.O)`2 <= 1) &
      (-1 <=(ff.I)`2 & (ff.I)`2 <= 1) &
      (-1 <=(gg.O)`1 & (gg.O)`1 <= 1) &
      (-1 <=(gg.I)`1 & (gg.I)`1 <= 1)
    proof
     A178: (ff.O)=(Sq_Circ").(f.O) by A3,FUNCT_1:22;
     consider p1 being Point of TOP-REAL 2 such that
     A179:  f.O=p1 &( |.p1.|=1 & p1`2>=p1`1 & p1`2<=-p1`1) by A1;
     A180:p1<>0.REAL 2 &
          (p1`2<=p1`1 & -p1`1<=p1`2 or p1`2>=p1`1 & p1`2<=-p1`1)
                            by A179,TOPRNS_1:24;
    then A181:Sq_Circ".p1=|[p1`1*sqrt(1+(p1`2/p1`1)^2),p1`2*sqrt(1+(p1`2/p1`1)
^2)]|
                                 by JGRAPH_3:38;
        reconsider px=ff.O as Point of TOP-REAL 2;
        set q=px;
        A182: px`1 = p1`1*sqrt(1+(p1`2/p1`1)^2) &
         px`2 = p1`2*sqrt(1+(p1`2/p1`1)^2) by A178,A179,A181,EUCLID:56;
          (p1`2/p1`1)^2 >=0 by SQUARE_1:72;
        then 1+(p1`2/p1`1)^2>=1+0 by REAL_1:55;
        then 1+(p1`2/p1`1)^2>0 by AXIOMS:22;
        then A183:sqrt(1+(p1`2/p1`1)^2)>0 by SQUARE_1:93;
        A184:now assume
A185:        px`1=0 & px`2=0;
          then A186:p1`1=0 by A182,A183,XCMPLX_1:6;
            p1`2=0 by A182,A183,A185,XCMPLX_1:6;
         hence contradiction by A180,A186,EUCLID:57,58;
        end;
          p1`2<=p1`1 & -p1`1<=p1`2 or p1`2>=p1`1 &
        p1`2*sqrt(1+(p1`2/p1`1)^2) <= (-p1`1)*sqrt(1+(p1`2/p1`1)^2)
                             by A179,A183,AXIOMS:25;
     then p1`2<=p1`1 & (-p1`1)*sqrt(1+(p1`2/p1`1)^2) <= p1`2*sqrt(1+(p1`2/p1`1
)^2)
           or px`2>=px`1 & px`2<=-px`1 by A182,A183,AXIOMS:25,XCMPLX_1:175;
        then A187:px`2<=px`1 & -px`1<=px`2 or px`2>=px`1 & px`2<=-px`1
                                by A182,A183,AXIOMS:25,XCMPLX_1:175;
        A188:p1=Sq_Circ.px by A178,A179,FUNCT_1:54,JGRAPH_3:32,54;
        A189:Sq_Circ.q=|[q`1/sqrt(1+(q`2/q`1)^2),q`2/sqrt(1+(q`2/q`1)^2)]|
                                 by A184,A187,JGRAPH_2:11,JGRAPH_3:def 1;
        A190: (|[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;
          (q`2/q`1)^2 >=0 by SQUARE_1:72;
        then 1+(q`2/q`1)^2>=1+0 by REAL_1:55;
        then A191:1+(q`2/q`1)^2>0 by AXIOMS:22;
         px`1<>0 by A184,A187;
        then A192: (px`1)^2<>0 by SQUARE_1:73;
          now
          (|.p1.|)^2= (q`1/sqrt(1+(q`2/q`1)^2))^2+(q`2/sqrt(1+(q`2/q`1)^2))^2
                              by A188,A189,A190,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 A191,SQUARE_1:def 4
               .= (q`1)^2/(1+(q`2/q`1)^2)+(q`2)^2/(1+(q`2/q`1)^2)
                              by A191,SQUARE_1:def 4
               .= ((q`1)^2+(q`2)^2)/(1+(q`2/q`1)^2) by XCMPLX_1:63;
        then ((q`1)^2+(q`2)^2)/(1+(q`2/q`1)^2)*(1+(q`2/q`1)^2)=(1)*(1+(q`2/q`1
)^2)
                              by A179,SQUARE_1:59;
        then ((q`1)^2+(q`2)^2)=(1+(q`2/q`1)^2) by A191,XCMPLX_1:88;
         then (px`1)^2+(px`2)^2=1+(px`2)^2/(px`1)^2 by SQUARE_1:69;
         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 by A192,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 XCMPLX_1:14
;
         then (px`1)^2*(px`1)^2+((px`1)^2*(px`2)^2-(px`1)^2*1)-(px`2)^2=0
                                   by XCMPLX_1:40;
         then (px`1)^2*(px`1)^2+(px`1)^2*(px`2)^2-(px`1)^2*1-(px`2)^2=0
                                   by XCMPLX_1:29;
         then (px`1)^2*(px`1)^2-(px`1)^2*1+(px`1)^2*(px`2)^2-(1)*(px`2)^2=0
                                   by XCMPLX_1:29;
         then (px`1)^2*((px`1)^2-1)+(px`1)^2*(px`2)^2-(1)*(px`2)^2=0
                                   by XCMPLX_1:40;
         then (px`1)^2*((px`1)^2-1)+((px`1)^2*(px`2)^2-(1)*(px`2)^2)=0
                                   by XCMPLX_1:29;
         hence
           ((px`1)^2-1)*(px`1)^2+((px`1)^2-1)*(px`2)^2=0 by XCMPLX_1:40;
         end;
         then A193:((px`1)^2-1)*((px`1)^2+(px`2)^2)=0 by XCMPLX_1:8;
           ((px`1)^2+(px`2)^2)<>0 by A184,COMPLEX1:2;
         then ((px`1)^2-1)=0 by A193,XCMPLX_1:6;
         then (px`1-1)*(px`1+1)=0 by SQUARE_1:59,67;
         then px`1-1=0 or px`1+1=0 by XCMPLX_1:6;
         then px`1=0+1 or px`1+1=0 by XCMPLX_1:27;
         then A194: px`1=1 or px`1=0-1 by XCMPLX_1:26;
    A195: (ff.I)=(Sq_Circ").(f.I) by A3,FUNCT_1:22;
    consider p2 being Point of TOP-REAL 2 such that
    A196:  f.I=p2 &(
    |.p2.|=1 & p2`2<=p2`1 & p2`2>=-p2`1) by A1;
    A197:p2<>0.REAL 2 &
          (p2`2<=p2`1 & -p2`1<=p2`2 or p2`2>=p2`1 & p2`2<=-p2`1)
                            by A196,TOPRNS_1:24;
    then A198:Sq_Circ".p2=|[p2`1*sqrt(1+(p2`2/p2`1)^2),p2`2*sqrt(1+(p2`2/p2`1)
^2)]|
                                 by JGRAPH_3:38;
        reconsider py=ff.I as Point of TOP-REAL 2;
        A199: py`1 = p2`1*sqrt(1+(p2`2/p2`1)^2) &
         py`2 = p2`2*sqrt(1+(p2`2/p2`1)^2) by A195,A196,A198,EUCLID:56;
          (p2`2/p2`1)^2 >=0 by SQUARE_1:72;
        then 1+(p2`2/p2`1)^2>=1+0 by REAL_1:55;
        then 1+(p2`2/p2`1)^2>0 by AXIOMS:22;
        then A200:sqrt(1+(p2`2/p2`1)^2)>0 by SQUARE_1:93;
        A201:now assume
A202:        py`1=0 & py`2=0;
          then A203:p2`1=0 by A199,A200,XCMPLX_1:6;
            p2`2=0 by A199,A200,A202,XCMPLX_1:6;
         hence contradiction by A197,A203,EUCLID:57,58;
        end;
        A204: now
       p2`2<=p2`1 & (-p2`1)*sqrt(1+(p2`2/p2`1)^2) <= p2`2*sqrt(1+(p2`2/p2`1)^2)
           or py`2>=py`1 & py`2<=-py`1 by A196,A200,AXIOMS:25;
           hence
          p2`2*sqrt(1+(p2`2/p2`1)^2) <= p2`1*sqrt(1+(p2`2/p2`1)^2) & -py`1<=py
`2
           or py`2>=py`1 & py`2<=-py`1 by A199,A200,AXIOMS:25,XCMPLX_1:175;
        end;
        A205:p2=Sq_Circ.py by A195,A196,FUNCT_1:54,JGRAPH_3:32,54;
     A206:Sq_Circ.py=|[py`1/sqrt(1+(py`2/py`1)^2),py`2/sqrt(1+(py`2/py`1)^2)]|
                                 by A199,A201,A204,JGRAPH_2:11,JGRAPH_3:def 1;
        A207: (|[py`1/sqrt(1+(py`2/py`1)^2),py`2/sqrt(1+(py`2/py`1)^2)]|)`1
         = py`1/sqrt(1+(py`2/py`1)^2) &
         (|[py`1/sqrt(1+(py`2/py`1)^2),py`2/sqrt(1+(py`2/py`1)^2)]|)`2
         = py`2/sqrt(1+(py`2/py`1)^2) by EUCLID:56;
          (py`2/py`1)^2 >=0 by SQUARE_1:72;
        then 1+(py`2/py`1)^2>=1+0 by REAL_1:55;
        then A208:1+(py`2/py`1)^2>0 by AXIOMS:22;
         py`1<>0 by A199,A201,A204;
        then A209: (py`1)^2<>0 by SQUARE_1:73;
          now
      (|.p2.|)^2= (py`1/sqrt(1+(py`2/py`1)^2))^2+(py`2/sqrt(1+(py`2/py`1)^2))^2
                              by A205,A206,A207,JGRAPH_3:10
         .= (py`1)^2/(sqrt(1+(py`2/py`1)^2))^2+(py`2/sqrt(1+(py`2/py`1)^2))^2
                              by SQUARE_1:69
               .= (py`1)^2/(sqrt(1+(py`2/py`1)^2))^2
                             +(py`2)^2/(sqrt(1+(py`2/py`1)^2))^2
                              by SQUARE_1:69
               .= (py`1)^2/(1+(py`2/py`1)^2)+(py`2)^2/(sqrt(1+(py`2/py`1)^2))^2
                              by A208,SQUARE_1:def 4
               .= (py`1)^2/(1+(py`2/py`1)^2)+(py`2)^2/(1+(py`2/py`1)^2)
                              by A208,SQUARE_1:def 4
               .= ((py`1)^2+(py`2)^2)/(1+(py`2/py`1)^2)
                              by XCMPLX_1:63;
       then ((py`1)^2+(py`2)^2)/(1+(py`2/py`1)^2)*(1+(py`2/py`1)^2)
                =(1)*(1+(py`2/py`1)^2) by A196,SQUARE_1:59;
        then ((py`1)^2+(py`2)^2)=(1+(py`2/py`1)^2)
                              by A208,XCMPLX_1:88;
         then (py`1)^2+(py`2)^2=1+(py`2)^2/(py`1)^2 by SQUARE_1:69;
         then (py`1)^2+(py`2)^2-1=(py`2)^2/(py`1)^2 by XCMPLX_1:26;
         then ((py`1)^2+(py`2)^2-1)*(py`1)^2=(py`2)^2
                            by A209,XCMPLX_1:88;
         then ((py`1)^2+((py`2)^2-1))*(py`1)^2=(py`2)^2 by XCMPLX_1:29;
         then (py`1)^2*(py`1)^2+((py`2)^2-1)*(py`1)^2=(py`2)^2 by XCMPLX_1:8;
         then (py`1)^2*(py`1)^2+(py`1)^2*((py`2)^2-1)-(py`2)^2=0 by XCMPLX_1:14
;
         then (py`1)^2*(py`1)^2+((py`1)^2*(py`2)^2-(py`1)^2*1)-(py`2)^2=0
                                   by XCMPLX_1:40;
         then (py`1)^2*(py`1)^2+(py`1)^2*(py`2)^2-(py`1)^2*1-(py`2)^2=0
                                   by XCMPLX_1:29;
         then (py`1)^2*(py`1)^2-(py`1)^2*1+(py`1)^2*(py`2)^2-(1)*(py`2)^2=0
                                   by XCMPLX_1:29;
         then (py`1)^2*((py`1)^2-1)+(py`1)^2*(py`2)^2-(1)*(py`2)^2=0
                                   by XCMPLX_1:40;
         then (py`1)^2*((py`1)^2-1)+((py`1)^2*(py`2)^2-(1)*(py`2)^2)=0
                                   by XCMPLX_1:29;
          hence ((py`1)^2-1)*(py`1)^2+((py`1)^2-1)*(py`2)^2=0 by XCMPLX_1:40;
         end;
         then A210:((py`1)^2-1)*((py`1)^2+(py`2)^2)=0 by XCMPLX_1:8;
           ((py`1)^2+(py`2)^2)<>0 by A201,COMPLEX1:2;
         then ((py`1)^2-1)=0 by A210,XCMPLX_1:6;
         then (py`1-1)*(py`1+1)=0 by SQUARE_1:59,67;
         then py`1-1=0 or py`1+1=0 by XCMPLX_1:6;
         then py`1=0+1 or py`1+1=0 by XCMPLX_1:27;
         then A211: py`1=1 or py`1=0-1 by XCMPLX_1:26;
    A212: gg.O=(Sq_Circ").(g.O) by A4,FUNCT_1:22;
    consider p3 being Point of TOP-REAL 2 such that
    A213:  g.O=p3 &( |.p3.|=1 & p3`2<=p3`1 & p3`2<=-p3`1) by A1;
    A214: -p3`2>=--p3`1 by A213,REAL_1:50;
    then A215:p3<>0.REAL 2 &
          (p3`1<=p3`2 & -p3`2<=p3`1 or p3`1>=p3`2 & p3`1<=-p3`2)
                                 by A213,TOPRNS_1:24;
   then A216:Sq_Circ".p3=|[p3`1*sqrt(1+(p3`1/p3`2)^2),p3`2*sqrt(1+(p3`1/p3`2)^2
)]|
                                 by JGRAPH_3:40;
        reconsider pz=gg.O as Point of TOP-REAL 2;
        A217: pz`2 = p3`2*sqrt(1+(p3`1/p3`2)^2) &
         pz`1 = p3`1*sqrt(1+(p3`1/p3`2)^2) by A212,A213,A216,EUCLID:56;
          (p3`1/p3`2)^2 >=0 by SQUARE_1:72;
        then 1+(p3`1/p3`2)^2>=1+0 by REAL_1:55;
        then 1+(p3`1/p3`2)^2>0 by AXIOMS:22;
        then A218:sqrt(1+(p3`1/p3`2)^2)>0 by SQUARE_1:93;
        A219:now assume
A220:        pz`2=0 & pz`1=0;
          then A221:p3`2=0 by A217,A218,XCMPLX_1:6;
            p3`1=0 by A217,A218,A220,XCMPLX_1:6;
         hence contradiction by A215,A221,EUCLID:57,58;
        end;
        A222: now
          p3`1<=p3`2 & -p3`2<=p3`1 or p3`1>=p3`2 &
        p3`1*sqrt(1+(p3`1/p3`2)^2) <= (-p3`2)*sqrt(1+(p3`1/p3`2)^2)
                             by A213,A214,A218,AXIOMS:25;
     then p3`1<=p3`2 & (-p3`2)*sqrt(1+(p3`1/p3`2)^2) <= p3`1*sqrt(1+(p3`1/p3`2
)^2)
           or pz`1>=pz`2 & pz`1<=-pz`2 by A217,A218,AXIOMS:25,XCMPLX_1:175;
           hence
         p3`1*sqrt(1+(p3`1/p3`2)^2) <= p3`2*sqrt(1+(p3`1/p3`2)^2) & -pz`2<=pz`1
           or pz`1>=pz`2 & pz`1<=-pz`2 by A217,A218,AXIOMS:25,XCMPLX_1:175;
        end;
        A223:p3=Sq_Circ.pz by A212,A213,FUNCT_1:54,JGRAPH_3:32,54;
     A224:Sq_Circ.pz=|[pz`1/sqrt(1+(pz`1/pz`2)^2),pz`2/sqrt(1+(pz`1/pz`2)^2)]|
                                 by A217,A219,A222,JGRAPH_2:11,JGRAPH_3:14;
        A225: (|[pz`1/sqrt(1+(pz`1/pz`2)^2),pz`2/sqrt(1+(pz`1/pz`2)^2)]|)`2
         = pz`2/sqrt(1+(pz`1/pz`2)^2) &
         (|[pz`1/sqrt(1+(pz`1/pz`2)^2),pz`2/sqrt(1+(pz`1/pz`2)^2)]|)`1
         = pz`1/sqrt(1+(pz`1/pz`2)^2) by EUCLID:56;
          (pz`1/pz`2)^2 >=0 by SQUARE_1:72;
        then 1+(pz`1/pz`2)^2>=1+0 by REAL_1:55;
        then A226:1+(pz`1/pz`2)^2>0 by AXIOMS:22;
         pz`2<>0 by A217,A219,A222;
        then A227: (pz`2)^2<>0 by SQUARE_1:73;
 A228:(|.p3.|)^2= (pz`2/sqrt(1+(pz`1/pz`2)^2))^2+(pz`1/sqrt(1+(pz`1/pz`2)^2))^2
                              by A223,A224,A225,JGRAPH_3:10
        .= (pz`2)^2/(sqrt(1+(pz`1/pz`2)^2))^2+(pz`1/sqrt(1+(pz`1/pz`2)^2))^2
                              by SQUARE_1:69
               .= (pz`2)^2/(sqrt(1+(pz`1/pz`2)^2))^2
                      +(pz`1)^2/(sqrt(1+(pz`1/pz`2)^2))^2
                              by SQUARE_1:69
               .= (pz`2)^2/(1+(pz`1/pz`2)^2)+(pz`1)^2/(sqrt(1+(pz`1/pz`2)^2))^2
                              by A226,SQUARE_1:def 4
               .= (pz`2)^2/(1+(pz`1/pz`2)^2)+(pz`1)^2/(1+(pz`1/pz`2)^2)
                              by A226,SQUARE_1:def 4
               .= ((pz`2)^2+(pz`1)^2)/(1+(pz`1/pz`2)^2)
                              by XCMPLX_1:63;
          now
   ((pz`2)^2+(pz`1)^2)/(1+(pz`1/pz`2)^2)*(1+(pz`1/pz`2)^2)=(1)*(1+(pz`1/pz`2)^2
)
                           by A213,A228,SQUARE_1:59;
        then ((pz`2)^2+(pz`1)^2)=(1+(pz`1/pz`2)^2)
                              by A226,XCMPLX_1:88;
         then (pz`2)^2+(pz`1)^2=1+(pz`1)^2/(pz`2)^2 by SQUARE_1:69;
         then (pz`2)^2+(pz`1)^2-1=(pz`1)^2/(pz`2)^2 by XCMPLX_1:26;
         then ((pz`2)^2+(pz`1)^2-1)*(pz`2)^2=(pz`1)^2
                            by A227,XCMPLX_1:88;
         then ((pz`2)^2+((pz`1)^2-1))*(pz`2)^2=(pz`1)^2 by XCMPLX_1:29;
         then (pz`2)^2*(pz`2)^2+((pz`1)^2-1)*(pz`2)^2=(pz`1)^2 by XCMPLX_1:8;
         then (pz`2)^2*(pz`2)^2+(pz`2)^2*((pz`1)^2-1)-(pz`1)^2=0 by XCMPLX_1:14
;
         then (pz`2)^2*(pz`2)^2+((pz`2)^2*(pz`1)^2-(pz`2)^2*1)-(pz`1)^2=0
                                   by XCMPLX_1:40;
         then (pz`2)^2*(pz`2)^2+(pz`2)^2*(pz`1)^2-(pz`2)^2*1-(pz`1)^2=0
                                   by XCMPLX_1:29;
         then (pz`2)^2*(pz`2)^2-(pz`2)^2*1+(pz`2)^2*(pz`1)^2-(1)*(pz`1)^2=0
                                   by XCMPLX_1:29;hence
           (pz`2)^2*((pz`2)^2-1)+(pz`2)^2*(pz`1)^2-(1)*(pz`1)^2=0
                                   by XCMPLX_1:40;
         end;
         then (pz`2)^2*((pz`2)^2-1)+((pz`2)^2*(pz`1)^2-(1)*(pz`1)^2)=0
                                   by XCMPLX_1:29;
         then (pz`2)^2*((pz`2)^2-1)+((pz`2)^2-1)*(pz`1)^2=0
                                   by XCMPLX_1:40;
         then A229:((pz`2)^2-1)*((pz`2)^2+(pz`1)^2)=0 by XCMPLX_1:8;
           ((pz`2)^2+(pz`1)^2)<>0 by A219,COMPLEX1:2;
         then ((pz`2)^2-1)=0 by A229,XCMPLX_1:6;
         then (pz`2-1)*(pz`2+1)=0 by SQUARE_1:59,67;
         then pz`2-1=0 or pz`2+1=0 by XCMPLX_1:6;
         then pz`2=0+1 or pz`2+1=0 by XCMPLX_1:27;
         then A230: pz`2=1 or pz`2=0-1 by XCMPLX_1:26;
    A231: (gg.I)=(Sq_Circ").(g.I) by A4,FUNCT_1:22;
    consider p4 being Point of TOP-REAL 2 such that
    A232:  g.I=p4 &(
    |.p4.|=1 & p4`2>=p4`1 & p4`2>=-p4`1) by A1;
    A233: -p4`2<=--p4`1 by A232,REAL_1:50;
    then A234:p4<>0.REAL 2 &
          (p4`1<=p4`2 & -p4`2<=p4`1 or p4`1>=p4`2 & p4`1<=-p4`2)
                      by A232,TOPRNS_1:24;
    then A235:Sq_Circ".p4=|[p4`1*sqrt(1+(p4`1/p4`2)^2),p4`2*sqrt(1+(p4`1/p4`2)
^2)]|
                                 by JGRAPH_3:40;
        reconsider pu=gg.I as Point of TOP-REAL 2;
        A236: pu`2 = p4`2*sqrt(1+(p4`1/p4`2)^2) &
         pu`1 = p4`1*sqrt(1+(p4`1/p4`2)^2) by A231,A232,A235,EUCLID:56;
          (p4`1/p4`2)^2 >=0 by SQUARE_1:72;
        then 1+(p4`1/p4`2)^2>=1+0 by REAL_1:55;
        then 1+(p4`1/p4`2)^2>0 by AXIOMS:22;
        then A237:sqrt(1+(p4`1/p4`2)^2)>0 by SQUARE_1:93;
        A238:now assume
A239:        pu`2=0 & pu`1=0;
          then A240:p4`2=0 by A236,A237,XCMPLX_1:6;
            p4`1=0 by A236,A237,A239,XCMPLX_1:6;
         hence contradiction by A234,A240,EUCLID:57,58;
        end;
        A241: now
       p4`1<=p4`2 & (-p4`2)*sqrt(1+(p4`1/p4`2)^2) <= p4`1*sqrt(1+(p4`1/p4`2)^2)
           or pu`1>=pu`2 & pu`1<=-pu`2 by A232,A233,A237,AXIOMS:25;hence
         p4`1*sqrt(1+(p4`1/p4`2)^2) <= p4`2*sqrt(1+(p4`1/p4`2)^2) & -pu`2<=pu`1
           or pu`1>=pu`2 & pu`1<=-pu`2 by A236,A237,AXIOMS:25,XCMPLX_1:175;
        end;
        A242:p4=Sq_Circ.pu by A231,A232,FUNCT_1:54,JGRAPH_3:32,54;
    A243:Sq_Circ.pu=|[pu`1/sqrt(1+(pu`1/pu`2)^2),pu`2/sqrt(1+(pu`1/pu`2)^2)]|
                                 by A236,A238,A241,JGRAPH_2:11,JGRAPH_3:14;
        A244: (|[pu`1/sqrt(1+(pu`1/pu`2)^2),pu`2/sqrt(1+(pu`1/pu`2)^2)]|)`2
         = pu`2/sqrt(1+(pu`1/pu`2)^2) &
         (|[pu`1/sqrt(1+(pu`1/pu`2)^2),pu`2/sqrt(1+(pu`1/pu`2)^2)]|)`1
         = pu`1/sqrt(1+(pu`1/pu`2)^2) by EUCLID:56;
          (pu`1/pu`2)^2 >=0 by SQUARE_1:72;
        then 1+(pu`1/pu`2)^2>=1+0 by REAL_1:55;
        then A245:1+(pu`1/pu`2)^2>0 by AXIOMS:22;
         pu`2<>0 by A236,A238,A241;
        then A246: (pu`2)^2<>0 by SQUARE_1:73;
          now
      (|.p4.|)^2= (pu`2/sqrt(1+(pu`1/pu`2)^2))^2+(pu`1/sqrt(1+(pu`1/pu`2)^2))^2
                              by A242,A243,A244,JGRAPH_3:10
        .= (pu`2)^2/(sqrt(1+(pu`1/pu`2)^2))^2+(pu`1/sqrt(1+(pu`1/pu`2)^2))^2
                              by SQUARE_1:69
               .= (pu`2)^2/(sqrt(1+(pu`1/pu`2)^2))^2
                             +(pu`1)^2/(sqrt(1+(pu`1/pu`2)^2))^2
                              by SQUARE_1:69
            .= (pu`2)^2/(1+(pu`1/pu`2)^2)+(pu`1)^2/(sqrt(1+(pu`1/pu`2)^2))^2
                              by A245,SQUARE_1:def 4
               .= (pu`2)^2/(1+(pu`1/pu`2)^2)+(pu`1)^2/(1+(pu`1/pu`2)^2)
                              by A245,SQUARE_1:def 4
               .= ((pu`2)^2+(pu`1)^2)/(1+(pu`1/pu`2)^2) by XCMPLX_1:63;
  then ((pu`2)^2+(pu`1)^2)/(1+(pu`1/pu`2)^2)*(1+(pu`1/pu`2)^2)=(1)*(1+(pu`1/pu
`2)^2)
                               by A232,SQUARE_1:59;
        then ((pu`2)^2+(pu`1)^2)=(1+(pu`1/pu`2)^2) by A245,XCMPLX_1:88;
         then (pu`2)^2+(pu`1)^2=1+(pu`1)^2/(pu`2)^2 by SQUARE_1:69;
         then (pu`2)^2+(pu`1)^2-1=(pu`1)^2/(pu`2)^2 by XCMPLX_1:26;
         then ((pu`2)^2+(pu`1)^2-1)*(pu`2)^2=(pu`1)^2 by A246,XCMPLX_1:88;
         then ((pu`2)^2+((pu`1)^2-1))*(pu`2)^2=(pu`1)^2 by XCMPLX_1:29;
         then (pu`2)^2*(pu`2)^2+((pu`1)^2-1)*(pu`2)^2=(pu`1)^2 by XCMPLX_1:8;
         then (pu`2)^2*(pu`2)^2+(pu`2)^2*((pu`1)^2-1)-(pu`1)^2=0 by XCMPLX_1:14
;
         then (pu`2)^2*(pu`2)^2+((pu`2)^2*(pu`1)^2-(pu`2)^2*1)-(pu`1)^2=0
                                   by XCMPLX_1:40;
         then (pu`2)^2*(pu`2)^2+(pu`2)^2*(pu`1)^2-(pu`2)^2*1-(pu`1)^2=0
                                   by XCMPLX_1:29;
         then (pu`2)^2*(pu`2)^2-(pu`2)^2*1+(pu`2)^2*(pu`1)^2-(1)*(pu`1)^2=0
                                   by XCMPLX_1:29;
         then (pu`2)^2*((pu`2)^2-1)+(pu`2)^2*(pu`1)^2-(1)*(pu`1)^2=0
                                   by XCMPLX_1:40;
         then (pu`2)^2*((pu`2)^2-1)+((pu`2)^2*(pu`1)^2-(1)*(pu`1)^2)=0
                                   by XCMPLX_1:29;
          hence ((pu`2)^2-1)*(pu`2)^2+((pu`2)^2-1)*(pu`1)^2=0 by XCMPLX_1:40;
         end;
         then A247:((pu`2)^2-1)*((pu`2)^2+(pu`1)^2)=0 by XCMPLX_1:8;
           ((pu`2)^2+(pu`1)^2)<>0 by A238,COMPLEX1:2;
         then (pu`2)^2-1=0 by A247,XCMPLX_1:6;
         then (pu`2-1)*(pu`2+1)=0 by SQUARE_1:59,67;
         then pu`2-1=0 or pu`2+1=0 by XCMPLX_1:6;
         then pu`2=0+1 or pu`2+1=0 by XCMPLX_1:27;
         then A248: pu`2=1 or pu`2=0-1 by XCMPLX_1:26;
        thus -1 <=(ff.O)`2 & (ff.O)`2 <= 1 by A187,A194,AXIOMS:22;
        thus -1 <=(ff.I)`2 & (ff.I)`2 <= 1 by A199,A204,A211,AXIOMS:22;
        thus -1 <=(gg.O)`1 & (gg.O)`1 <= 1 by A217,A222,A230,AXIOMS:22;
        thus -1 <=(gg.I)`1 & (gg.I)`1 <= 1 by A236,A241,A248,AXIOMS:22;
      end;
    then rng ff meets rng gg by A1,A7,A9,A10,A11,A12,A99,Th14;
    then consider y being set such that
    A249: y in rng ff & y in rng gg by XBOOLE_0:3;
    consider x1 being set such that
    A250: x1 in dom ff & y=ff.x1 by A249,FUNCT_1:def 5;
    consider x2 being set such that
    A251: x2 in dom gg & y=gg.x2 by A249,FUNCT_1:def 5;
    A252:ff.x1=Sq_Circ".(f.x1) by A250,FUNCT_1:22;
    A253:dom (Sq_Circ")=the carrier of TOP-REAL 2
                         by FUNCT_2:def 1,JGRAPH_3:39;
      x1 in dom f by A250,FUNCT_1:21;
    then A254:f.x1 in rng f by FUNCT_1:def 5;
      x2 in dom g by A251,FUNCT_1:21;
    then A255:g.x2 in rng g by FUNCT_1:def 5;
    A256: Sq_Circ" is one-to-one by FUNCT_1:62,JGRAPH_3:32;
      gg.x2=Sq_Circ".(g.x2) by A251,FUNCT_1:22;
    then f.x1=g.x2 by A250,A251,A252,A253,A254,A255,A256,FUNCT_1:def 8;
  hence thesis by A254,A255,XBOOLE_0:3;
end;

theorem Th18: 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 KXN & f.I in KXP &
  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 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 KXN & f.I in KXP &
  g.O in KYP & g.I in KYN &
  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 one-to-one by Th15;
 thus rng f meets rng g by A1,A2,Th17;
end;

theorem Th19:
 for f,g being map of I[01],TOP-REAL 2,
  C0 being Subset of TOP-REAL 2
  st C0={q: |.q.|>=1} &
  f is continuous one-to-one &
  g is continuous one-to-one &
  f.0=|[-1,0]| & f.1=|[1,0]| & g.1=|[0,1]| & g.0=|[0,-1]|
  & 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 be Subset of TOP-REAL 2;
  assume A1: C0={q: |.q.|>=1} &
  f is continuous one-to-one & g is continuous one-to-one &
  f.0=|[-1,0]| & f.1=|[1,0]| & g.1=|[0,1]| & g.0=|[0,-1]|
  & rng f c= C0 & rng g c= C0;
  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 as Point of I[01] by BORSUK_1:83,RCOMP_1:15;
  reconsider I=1 as Point of I[01] by BORSUK_1:83,RCOMP_1:15;
  A2: (|[-1,0]|)`1=-1 & (|[-1,0]|)`2=0 by EUCLID:56;
  then A3: |. (|[-1,0]|).|=sqrt((-1)^2+0^2) by JGRAPH_3:10
                 .=1 by SQUARE_1:59,60,61,83;
    (|[-1,0]|)`2 <=-((|[-1,0]|)`1) by A2;
  then A4: f.O in KXN by A1,A2,A3;
  A5: (|[1,0]|)`1=1 & (|[1,0]|)`2=0 by EUCLID:56;
  then |.(|[1,0]|).|=sqrt(1^2+0) by JGRAPH_3:10,SQUARE_1:60
                 .=1 by SQUARE_1:59,83;
  then A6: f.I in KXP by A1,A5;
  A7: (|[0,-1]|)`2=-1 & (|[0,-1]|)`1=0 by EUCLID:56;
  then A8: |.(|[0,-1]|).|=sqrt(0^2+(-1)^2) by JGRAPH_3:10
                 .=1 by SQUARE_1:59,60,61,83;
    (|[0,-1]|)`2 <=-((|[0,-1]|)`1) by A7;
  then A9: g.O in KYN by A1,A7,A8;
  A10: (|[0,1]|)`2=1 & (|[0,1]|)`1=0 by EUCLID:56;
  then A11: |. (|[0,1]|).|=sqrt(0+1^2) by JGRAPH_3:10,SQUARE_1:60
                 .=1 by SQUARE_1:59,83;
    (|[0,1]|)`2 >=-((|[0,1]|)`1) by A10;
  then g.I in KYP by A1,A10,A11;
 hence rng f meets rng g by A1,A4,A6,A9,Th17;
end;

theorem for p1,p2,p3,p4 being Point of TOP-REAL 2,
 C0 being Subset of TOP-REAL 2
 st C0={p: |.p.|>=1}
 & |.p1.|=1 & |.p2.|=1 & |.p3.|=1 & |.p4.|=1 &
 (ex h being map of TOP-REAL 2,TOP-REAL 2 st h is_homeomorphism
 & h.:C0 c= C0 &
 h.p1=|[-1,0]| & h.p2=|[0,1]| & h.p3=|[1,0]| & h.p4=|[0,-1]|)
 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 &
  f.0=p1 & f.1=p3 & 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,
 C0 be Subset of TOP-REAL 2;
 assume A1: C0={p: |.p.|>=1}
 & |.p1.|=1 & |.p2.|=1 & |.p3.|=1 & |.p4.|=1 &
 (ex h being map of TOP-REAL 2,TOP-REAL 2 st h is_homeomorphism
 & h.:C0 c= C0 &
 h.p1=(|[-1,0]|) & h.p2=(|[0,1]|) & h.p3=(|[1,0]|) & h.p4=(|[0,-1]|));
 then consider h being map of TOP-REAL 2,TOP-REAL 2 such that
 A2:  h is_homeomorphism
    & h.:C0 c= C0 &
 h.p1=(|[-1,0]|) & h.p2=(|[0,1]|) & h.p3=(|[1,0]|) & h.p4=(|[0,-1]|);
  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 &
    f.0=p1 & f.1=p3 & g.0=p4 & g.1=p2 &
    rng f c= C0 & rng g c= C0;
    reconsider f2=h*f as map of I[01],TOP-REAL 2;
    reconsider g2=h*g as map of I[01],TOP-REAL 2;
    A4: h is continuous by A2,TOPS_2:def 5;
    then A5: f2 is continuous by A3,TOPS_2:58;
    A6: g2 is continuous by A3,A4,TOPS_2:58;
    A7: h is one-to-one by A2,TOPS_2:def 5;
    then A8: f2 is one-to-one by A3,FUNCT_1:46;
    A9: g2 is one-to-one by A3,A7,FUNCT_1:46;
    A10: 0 in dom f2 &1 in dom f2 by Lm1,BORSUK_1:83,FUNCT_2:def 1;
    then A11: f2.0=|[-1,0]| by A2,A3,FUNCT_1:22;
    A12: f2.1=|[1,0]| by A2,A3,A10,FUNCT_1:22;
    A13: 0 in dom g2 &1 in dom g2 by Lm1,BORSUK_1:83,FUNCT_2:def 1;
    then A14: g2.0=|[0,-1]| by A2,A3,FUNCT_1:22;
    A15: g2.1=|[0,1]| by A2,A3,A13,FUNCT_1:22;
    A16: rng f2 c= C0
    proof let y be set;assume y in rng f2;
      then consider x being set such that
      A17: x in dom f2 & y=f2.x by FUNCT_1:def 5;
      A18: y=h.(f.x) by A17,FUNCT_1:22;
        x in dom f by A17,FUNCT_1:21;
      then A19: f.x in rng f by FUNCT_1:def 5;
        dom h=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
      then y in h.:C0 by A3,A18,A19,FUNCT_1:def 12;
     hence y in C0 by A2;
    end;
      rng g2 c= C0
    proof let y be set;assume y in rng g2;
      then consider x being set such that
      A20: x in dom g2 & y=g2.x by FUNCT_1:def 5;
      A21: y=h.(g.x) by A20,FUNCT_1:22;
        x in dom g by A20,FUNCT_1:21;
      then A22: g.x in rng g by FUNCT_1:def 5;
        dom h=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
      then y in h.:C0 by A3,A21,A22,FUNCT_1:def 12;
     hence y in C0 by A2;
    end;
    then rng f2 meets rng g2 by A1,A5,A6,A8,A9,A11,A12,A14,A15,A16,Th19;
    then consider q5 being set such that
    A23: q5 in rng f2 & q5 in rng g2 by XBOOLE_0:3;
    consider x being set such that
    A24: x in dom f2 & q5=f2.x by A23,FUNCT_1:def 5;
    A25: q5=h.(f.x) by A24,FUNCT_1:22;
    consider u being set such that
    A26: u in dom g2 & q5=g2.u by A23,FUNCT_1:def 5;
    A27: q5=h.(g.u) by A26,FUNCT_1:22;
    A28: h is one-to-one by A2,TOPS_2:def 5;
    A29: f.x in dom h by A24,FUNCT_1:21;
      g.u in dom h by A26,FUNCT_1:21;
    then A30: f.x=g.u by A25,A27,A28,A29,FUNCT_1:def 8;
    A31: x in dom f by A24,FUNCT_1:21;
    A32: u in dom g by A26,FUNCT_1:21;
    A33: f.x in rng f by A31,FUNCT_1:def 5;
      g.u in rng g by A32,FUNCT_1:def 5;
   hence thesis by A30,A33,XBOOLE_0:3;
end;

begin :: Properties of Fan Morphisms

theorem Th21:
 for cn being Real,q being Point of TOP-REAL 2 st -1<cn & cn<1 & q`2>0
 holds (for p being Point of TOP-REAL 2 st p=(cn-FanMorphN).q holds p`2>0)
 proof let cn be Real,q be Point of TOP-REAL 2;
  assume A1: -1<cn & cn<1 & q`2>0;
    now per cases;
  case q`1/|.q.|>=cn;
   hence (for p being Point of TOP-REAL 2 st
    p=(cn-FanMorphN).q holds p`2>0) by A1,JGRAPH_4:82;
  case q`1/|.q.|<cn;
   hence (for p being Point of TOP-REAL 2 st
    p=(cn-FanMorphN).q holds p`2>0) by A1,JGRAPH_4:83;
  end;
 hence thesis;
end;

theorem
   for cn being Real,q being Point of TOP-REAL 2 st -1<cn & cn<1 & q`2>=0
 holds (for p being Point of TOP-REAL 2 st
 p=(cn-FanMorphN).q holds p`2>=0)
 proof let cn be Real,q be Point of TOP-REAL 2;
  assume A1: -1<cn & cn<1 & q`2>=0;
     now per cases by A1;
   case q`2>0;
    hence (for p being Point of TOP-REAL 2 st
      p=(cn-FanMorphN).q holds p`2>=0) by A1,Th21;
   case q`2=0;
    hence (for p being Point of TOP-REAL 2 st
     p=(cn-FanMorphN).q holds p`2>=0) by JGRAPH_4:56;
   end;
 hence thesis;
 end;

theorem Th23:
 for cn being Real,q being Point of TOP-REAL 2 st -1<cn & cn<1 & q`2>=0
 & q`1/|.q.|<cn & |.q.|<>0 holds (for p being Point of TOP-REAL 2 st
 p=(cn-FanMorphN).q holds p`2>=0 & p`1<0)
 proof let cn be Real,q be Point of TOP-REAL 2;
  assume A1: -1<cn & cn<1 & q`2>=0
   & q`1/|.q.|<cn & |.q.|<>0;
   let p be Point of TOP-REAL 2;
    assume A2: p=(cn-FanMorphN).q;
       now per cases;
     case A3: q`2=0;
       then A4: q=p by A2,JGRAPH_4:56;
         |.q.|^2=(q`1)^2+0 by A3,JGRAPH_3:10,SQUARE_1:60 .=(q`1)^2;
       then |.q.|=q`1 or |.q.|=-(q`1) by JGRAPH_3:1;
       then (-(q`1))/|.q.|=1 by A1,XCMPLX_1:60;
       then -(q`1/|.q.|)=1 by XCMPLX_1:188;
       then A5: q`1=(-1)*|.q.| by A1,XCMPLX_1:88;
         |.q.|>=0 by TOPRNS_1:26;
      hence p`2>=0 & p`1<0 by A1,A4,A5,SQUARE_1:24;
     case q`2<>0;
      hence p`2>=0 & p`1<0 by A1,A2,JGRAPH_4:83;
     end;
    hence p`2>=0 & p`1<0;
 end;

theorem Th24:
 for cn being Real,q1,q2 being Point of TOP-REAL 2 st -1<cn & cn<1 & q1`2>=0
 & q2`2>=0 & |.q1.|<>0 & |.q2.|<>0 & q1`1/|.q1.|<q2`1/|.q2.|
 holds (for p1,p2 being Point of TOP-REAL 2 st
 p1=(cn-FanMorphN).q1 & p2=(cn-FanMorphN).q2 holds p1`1/|.p1.|<p2`1/|.p2.|)
 proof let cn be Real,q1,q2 be Point of TOP-REAL 2;
   assume A1: -1<cn & cn<1 & q1`2>=0 & q2`2>=0 & |.q1.|<>0 & |.q2.|<>0
     & q1`1/|.q1.|<q2`1/|.q2.|;
      now per cases by A1;
    case A2: q1`2>0;
       now per cases by A1;
     case q2`2>0;
      hence (for p1,p2 being Point of TOP-REAL 2 st
       p1=(cn-FanMorphN).q1 & p2=(cn-FanMorphN).q2 holds
       p1`1/|.p1.|<p2`1/|.p2.|) by A1,A2,JGRAPH_4:86;
     case A3: q2`2=0;
       then |.q2.|^2=(q2`1)^2+0 by JGRAPH_3:10,SQUARE_1:60 .=(q2`1)^2;
       then A4: |.q2.|=q2`1 or |.q2.|=-(q2`1) by JGRAPH_3:1;
         now assume |.q2.|=-(q2`1);
         then 1=(-(q2`1))/|.q2.| by A1,XCMPLX_1:60;
         then A5: q1`1/|.q1.|< -1 by A1,XCMPLX_1:191;
         A6: |.q1.|>=0 by TOPRNS_1:26;
           (|.q1.|)^2=(q1`1)^2+(q1`2)^2 by JGRAPH_3:10;
         then (|.q1.|)^2-(q1`1)^2=(q1`2)^2 by XCMPLX_1:26;
         then (|.q1.|)^2-(q1`1)^2>=0 by SQUARE_1:72;
         then (|.q1.|)^2-(q1`1)^2+(q1`1)^2>=0+(q1`1)^2 by REAL_1:55;
         then (|.q1.|)^2>=(q1`1)^2 by XCMPLX_1:27;
         then -|.q1.|<=q1`1 & q1`1<=|.q1.| by A6,JGRAPH_2:5;
         then (-|.q1.|)/|.q1.|<=q1`1/|.q1.| by A1,A6,REAL_1:73;
        hence contradiction by A1,A5,XCMPLX_1:198;
       end;
       then A7: q2`1/|.q2.|=1 by A1,A4,XCMPLX_1:60;
       thus for p1,p2 being Point of TOP-REAL 2 st
       p1=(cn-FanMorphN).q1 & p2=(cn-FanMorphN).q2 holds
       p1`1/|.p1.|<p2`1/|.p2.|
       proof let p1,p2 be Point of TOP-REAL 2;
        assume A8: p1=(cn-FanMorphN).q1 & p2=(cn-FanMorphN).q2;
         then A9: p2=q2 by A3,JGRAPH_4:56;
         A10: |.p1.|=|.q1.| by A8,JGRAPH_4:73;
         A11: |.p1.|>=0 by TOPRNS_1:26;
           (|.p1.|)^2=(p1`1)^2+(p1`2)^2 by JGRAPH_3:10;
         then A12: (|.p1.|)^2-(p1`1)^2=(p1`2)^2 by XCMPLX_1:26;
         then (|.p1.|)^2-(p1`1)^2>=0 by SQUARE_1:72;
         then (|.p1.|)^2-(p1`1)^2+(p1`1)^2>=0+(p1`1)^2 by REAL_1:55;
         then (|.p1.|)^2>=(p1`1)^2 by XCMPLX_1:27;
         then -|.p1.|<=p1`1 & p1`1<=|.p1.| by A11,JGRAPH_2:5;
         then (|.p1.|)/|.p1.|>=p1`1/|.p1.| by A1,A10,A11,REAL_1:73;
         then A13: 1>= p1`1/|.p1.| by A1,A10,XCMPLX_1:60;
         A14: p1`2>0 by A1,A2,A8,Th21;
           now assume 1= p1`1/|.p1.|; then (1)*|.p1.|=p1`1 by A1,A10,XCMPLX_1:
88
;
           then (|.p1.|)^2-(p1`1)^2=0 by XCMPLX_1:14;
          hence contradiction by A12,A14,SQUARE_1:73;
         end;
        hence p1`1/|.p1.|<p2`1/|.p2.| by A7,A9,A13,REAL_1:def 5;
       end;
     end;
       hence (for p1,p2 being Point of TOP-REAL 2 st
        p1=(cn-FanMorphN).q1 & p2=(cn-FanMorphN).q2 holds
        p1`1/|.p1.|<p2`1/|.p2.|);
    case A15: q1`2=0;
       then |.q1.|^2=(q1`1)^2+0 by JGRAPH_3:10,SQUARE_1:60 .=(q1`1)^2;
       then A16: |.q1.|=q1`1 or |.q1.|=-(q1`1) by JGRAPH_3:1;
         now assume |.q1.|=(q1`1);
         then A17: q2`1/|.q2.|> 1 by A1,XCMPLX_1:60;
         A18: |.q2.|>=0 by TOPRNS_1:26;
           (|.q2.|)^2=(q2`1)^2+(q2`2)^2 by JGRAPH_3:10;
         then (|.q2.|)^2-(q2`1)^2=(q2`2)^2 by XCMPLX_1:26;
         then (|.q2.|)^2-(q2`1)^2>=0 by SQUARE_1:72;
         then (|.q2.|)^2-(q2`1)^2+(q2`1)^2>=0+(q2`1)^2 by REAL_1:55;
         then (|.q2.|)^2>=(q2`1)^2 by XCMPLX_1:27;
         then -|.q2.|<=q2`1 & q2`1<=|.q2.| by A18,JGRAPH_2:5;
         then (|.q2.|)/|.q2.|>=q2`1/|.q2.| by A1,A18,REAL_1:73;
        hence contradiction by A1,A17,XCMPLX_1:60;
       end;
       then (-(q1`1))/|.q1.|=1 by A1,A16,XCMPLX_1:60;
       then A19: -(q1`1/|.q1.|)=1 by XCMPLX_1:188;
     thus (for p1,p2 being Point of TOP-REAL 2 st
       p1=(cn-FanMorphN).q1 & p2=(cn-FanMorphN).q2
       holds p1`1/|.p1.|<p2`1/|.p2.|)
       proof let p1,p2 be Point of TOP-REAL 2;
        assume A20: p1=(cn-FanMorphN).q1 & p2=(cn-FanMorphN).q2;
         then A21: p1=q1 by A15,JGRAPH_4:56;
         A22: |.p2.|=|.q2.| by A20,JGRAPH_4:73;
         A23: |.p2.|>=0 by TOPRNS_1:26;
           (|.p2.|)^2=(p2`1)^2+(p2`2)^2 by JGRAPH_3:10;
         then A24: (|.p2.|)^2-(p2`1)^2=(p2`2)^2 by XCMPLX_1:26;
         then (|.p2.|)^2-(p2`1)^2>=0 by SQUARE_1:72;
         then (|.p2.|)^2-(p2`1)^2+(p2`1)^2>=0+(p2`1)^2 by REAL_1:55;
         then (|.p2.|)^2>=(p2`1)^2 by XCMPLX_1:27;
         then -|.p2.|<=p2`1 & p2`1<=|.p2.| by A23,JGRAPH_2:5;
         then (-|.p2.|)/|.p2.|<=p2`1/|.p2.| by A1,A22,A23,REAL_1:73;
         then A25: -1 <= p2`1/|.p2.| by A1,A22,XCMPLX_1:198;
           now per cases;
         case q2`2=0;
           then p2=q2 by A20,JGRAPH_4:56;
          hence p2`1/|.p2.|>-1 by A1,A19;
         case q2`2<>0;
         then A26: p2`2>0 by A1,A20,Th21;
           now assume -1= p2`1/|.p2.|; then (-1)*|.p2.|=p2`1 by A1,A22,XCMPLX_1
:88;
           then (-|.p2.|)^2=(p2`1)^2 by XCMPLX_1:180;
           then (|.p2.|)^2=(p2`1)^2 by SQUARE_1:61;
           then (|.p2.|)^2-(p2`1)^2=0 by XCMPLX_1:14;
          hence contradiction by A24,A26,SQUARE_1:73;
         end;
         hence p2`1/|.p2.|>-1 by A25,REAL_1:def 5;
         end;
        hence p1`1/|.p1.|<p2`1/|.p2.| by A19,A21;
       end;
    end;
  hence thesis;
 end;

theorem Th25:
 for sn being Real,q being Point of TOP-REAL 2 st -1<sn & sn<1 & q`1>0
 holds (for p being Point of TOP-REAL 2 st p=(sn-FanMorphE).q holds p`1>0)
 proof let sn be Real,q be Point of TOP-REAL 2;
  assume A1: -1<sn & sn<1 & q`1>0;
    now per cases;
  case q`2/|.q.|>=sn;
   hence (for p being Point of TOP-REAL 2 st
    p=(sn-FanMorphE).q holds p`1>0) by A1,JGRAPH_4:113;
  case q`2/|.q.|<sn;
   hence (for p being Point of TOP-REAL 2 st
    p=(sn-FanMorphE).q holds p`1>0) by A1,JGRAPH_4:114;
  end;
 hence thesis;
 end;

theorem
   for sn being Real,q being Point of TOP-REAL 2 st -1<sn & sn<1 & q`1>=0
 & q`2/|.q.|<sn & |.q.|<>0 holds (for p being Point of TOP-REAL 2 st
 p=(sn-FanMorphE).q holds p`1>=0 & p`2<0)
 proof let sn be Real,q be Point of TOP-REAL 2;
  assume A1: -1<sn & sn<1 & q`1>=0 & q`2/|.q.|<sn & |.q.|<>0;
   let p be Point of TOP-REAL 2;
    assume A2: p=(sn-FanMorphE).q;
       now per cases;
     case A3: q`1=0;
       then A4: q=p by A2,JGRAPH_4:89;
         |.q.|^2=(q`2)^2+0 by A3,JGRAPH_3:10,SQUARE_1:60 .=(q`2)^2;
       then |.q.|=q`2 or |.q.|=-(q`2) by JGRAPH_3:1;
       then (-(q`2))/|.q.|=1 by A1,XCMPLX_1:60;
       then -(q`2/|.q.|)=1 by XCMPLX_1:188;
       then A5: q`2=(-1)*|.q.| by A1,XCMPLX_1:88;
         |.q.|>=0 by TOPRNS_1:26;
      hence p`1>=0 & p`2<0 by A1,A4,A5,SQUARE_1:24;
     case q`1<>0;
      hence p`1>=0 & p`2<0 by A1,A2,JGRAPH_4:114;
     end;
    hence p`1>=0 & p`2<0;
 end;

theorem Th27:
 for sn being Real,q1,q2 being Point of TOP-REAL 2 st -1<sn & sn<1 & q1`1>=0
 & q2`1>=0 & |.q1.|<>0 & |.q2.|<>0 & q1`2/|.q1.|<q2`2/|.q2.|
 holds (for p1,p2 being Point of TOP-REAL 2 st
 p1=(sn-FanMorphE).q1 & p2=(sn-FanMorphE).q2 holds p1`2/|.p1.|<p2`2/|.p2.|)
 proof let sn be Real,q1,q2 be Point of TOP-REAL 2;
   assume A1: -1<sn & sn<1 & q1`1>=0 & q2`1>=0 & |.q1.|<>0 & |.q2.|<>0
     & q1`2/|.q1.|<q2`2/|.q2.|;
      now per cases by A1;
    case A2: q1`1>0;
       now per cases by A1;
     case q2`1>0;
      hence (for p1,p2 being Point of TOP-REAL 2 st
       p1=(sn-FanMorphE).q1 & p2=(sn-FanMorphE).q2 holds
       p1`2/|.p1.|<p2`2/|.p2.|) by A1,A2,JGRAPH_4:117;
     case A3: q2`1=0;
       then |.q2.|^2=(q2`2)^2+0 by JGRAPH_3:10,SQUARE_1:60 .=(q2`2)^2;
       then A4: |.q2.|=q2`2 or |.q2.|=-(q2`2) by JGRAPH_3:1;
         now assume |.q2.|=-(q2`2);
         then 1=(-(q2`2))/|.q2.| by A1,XCMPLX_1:60;
         then A5: q1`2/|.q1.|< -1 by A1,XCMPLX_1:191;
         A6: |.q1.|>=0 by TOPRNS_1:26;
           (|.q1.|)^2=(q1`2)^2+(q1`1)^2 by JGRAPH_3:10;
         then (|.q1.|)^2-(q1`2)^2=(q1`1)^2 by XCMPLX_1:26;
         then (|.q1.|)^2-(q1`2)^2>=0 by SQUARE_1:72;
         then (|.q1.|)^2-(q1`2)^2+(q1`2)^2>=0+(q1`2)^2 by REAL_1:55;
         then (|.q1.|)^2>=(q1`2)^2 by XCMPLX_1:27;
         then -|.q1.|<=q1`2 & q1`2<=|.q1.| by A6,JGRAPH_2:5;
         then (-|.q1.|)/|.q1.|<=q1`2/|.q1.| by A1,A6,REAL_1:73;
        hence contradiction by A1,A5,XCMPLX_1:198;
       end;
       then A7: q2`2/|.q2.|=1 by A1,A4,XCMPLX_1:60;
       thus for p1,p2 being Point of TOP-REAL 2 st
       p1=(sn-FanMorphE).q1 & p2=(sn-FanMorphE).q2 holds
       p1`2/|.p1.|<p2`2/|.p2.|
       proof let p1,p2 be Point of TOP-REAL 2;
        assume A8: p1=(sn-FanMorphE).q1 & p2=(sn-FanMorphE).q2;
         then A9: p2=q2 by A3,JGRAPH_4:89;
         A10: |.p1.|=|.q1.| by A8,JGRAPH_4:104;
         A11: |.p1.|>=0 by TOPRNS_1:26;
           (|.p1.|)^2=(p1`2)^2+(p1`1)^2 by JGRAPH_3:10;
         then A12: (|.p1.|)^2-(p1`2)^2=(p1`1)^2 by XCMPLX_1:26;
         then (|.p1.|)^2-(p1`2)^2>=0 by SQUARE_1:72;
         then (|.p1.|)^2-(p1`2)^2+(p1`2)^2>=0+(p1`2)^2 by REAL_1:55;
         then (|.p1.|)^2>=(p1`2)^2 by XCMPLX_1:27;
         then -|.p1.|<=p1`2 & p1`2<=|.p1.| by A11,JGRAPH_2:5;
         then (|.p1.|)/|.p1.|>=p1`2/|.p1.| by A1,A10,A11,REAL_1:73;
         then A13: 1>= p1`2/|.p1.| by A1,A10,XCMPLX_1:60;

         A14: p1`1>0 by A1,A2,A8,Th25;
           now assume 1= p1`2/|.p1.|; then (1)*|.p1.|=p1`2 by A1,A10,XCMPLX_1:
88
;
           then (|.p1.|)^2-(p1`2)^2=0 by XCMPLX_1:14;
          hence contradiction by A12,A14,SQUARE_1:73;
         end;
        hence p1`2/|.p1.|<p2`2/|.p2.| by A7,A9,A13,REAL_1:def 5;
       end;
     end;
       hence (for p1,p2 being Point of TOP-REAL 2 st
        p1=(sn-FanMorphE).q1 & p2=(sn-FanMorphE).q2 holds
        p1`2/|.p1.|<p2`2/|.p2.|);
    case A15: q1`1=0;
       then |.q1.|^2=(q1`2)^2+0 by JGRAPH_3:10,SQUARE_1:60 .=(q1`2)^2;
       then A16: |.q1.|=q1`2 or |.q1.|=-(q1`2) by JGRAPH_3:1;
         now assume |.q1.|=(q1`2);
         then A17: q2`2/|.q2.|> 1 by A1,XCMPLX_1:60;
         A18: |.q2.|>=0 by TOPRNS_1:26;
           (|.q2.|)^2=(q2`2)^2+(q2`1)^2 by JGRAPH_3:10;
         then (|.q2.|)^2-(q2`2)^2=(q2`1)^2 by XCMPLX_1:26;
         then (|.q2.|)^2-(q2`2)^2>=0 by SQUARE_1:72;
         then (|.q2.|)^2-(q2`2)^2+(q2`2)^2>=0+(q2`2)^2 by REAL_1:55;
         then (|.q2.|)^2>=(q2`2)^2 by XCMPLX_1:27;
         then -|.q2.|<=q2`2 & q2`2<=|.q2.| by A18,JGRAPH_2:5;
         then (|.q2.|)/|.q2.|>=q2`2/|.q2.| by A1,A18,REAL_1:73;
        hence contradiction by A1,A17,XCMPLX_1:60;
       end;
       then (-(q1`2))/|.q1.|=1 by A1,A16,XCMPLX_1:60;
       then A19: -(q1`2/|.q1.|)=1 by XCMPLX_1:188;
     thus (for p1,p2 being Point of TOP-REAL 2 st
       p1=(sn-FanMorphE).q1 & p2=(sn-FanMorphE).q2
       holds p1`2/|.p1.|<p2`2/|.p2.|)
       proof let p1,p2 be Point of TOP-REAL 2;
        assume A20: p1=(sn-FanMorphE).q1 & p2=(sn-FanMorphE).q2;
         then A21: p1=q1 by A15,JGRAPH_4:89;
         A22: |.p2.|=|.q2.| by A20,JGRAPH_4:104;
         A23: |.p2.|>=0 by TOPRNS_1:26;
           (|.p2.|)^2=(p2`2)^2+(p2`1)^2 by JGRAPH_3:10;
         then A24: (|.p2.|)^2-(p2`2)^2=(p2`1)^2 by XCMPLX_1:26;
         then (|.p2.|)^2-(p2`2)^2>=0 by SQUARE_1:72;
         then (|.p2.|)^2-(p2`2)^2+(p2`2)^2>=0+(p2`2)^2 by REAL_1:55;
         then (|.p2.|)^2>=(p2`2)^2 by XCMPLX_1:27;
         then -|.p2.|<=p2`2 & p2`2<=|.p2.| by A23,JGRAPH_2:5;
         then (-|.p2.|)/|.p2.|<=p2`2/|.p2.| by A1,A22,A23,REAL_1:73;
         then A25: -1 <= p2`2/|.p2.| by A1,A22,XCMPLX_1:198;
           now per cases;
         case q2`1=0;
           then p2=q2 by A20,JGRAPH_4:89;
          hence p2`2/|.p2.|>-1 by A1,A19;
         case q2`1<>0;
         then A26: p2`1>0 by A1,A20,Th25;
           now assume -1= p2`2/|.p2.|; then (-1)*|.p2.|=p2`2 by A1,A22,XCMPLX_1
:88;
           then (-|.p2.|)^2=(p2`2)^2 by XCMPLX_1:180;
           then (|.p2.|)^2=(p2`2)^2 by SQUARE_1:61;
           then (|.p2.|)^2-(p2`2)^2=0 by XCMPLX_1:14;
          hence contradiction by A24,A26,SQUARE_1:73;
         end;
         hence
           p2`2/|.p2.|>-1 by A25,REAL_1:def 5;
         end;
        hence p1`2/|.p1.|<p2`2/|.p2.| by A19,A21;
       end;
    end;
  hence thesis;
 end;

theorem Th28:
 for cn being Real,q being Point of TOP-REAL 2 st -1<cn & cn<1 & q`2<0
 holds (for p being Point of TOP-REAL 2 st
 p=(cn-FanMorphS).q holds p`2<0)
 proof let cn be Real,q be Point of TOP-REAL 2;
  assume A1: -1<cn & cn<1 & q`2<0;
    now per cases;
  case q`1/|.q.|>=cn;
   hence (for p being Point of TOP-REAL 2 st
    p=(cn-FanMorphS).q holds p`2<0) by A1,JGRAPH_4:144;
  case q`1/|.q.|<cn;
   hence (for p being Point of TOP-REAL 2 st
    p=(cn-FanMorphS).q holds p`2<0) by A1,JGRAPH_4:145;
  end;
 hence thesis;
 end;

theorem Th29:
 for cn being Real,q being Point of TOP-REAL 2 st -1<cn & cn<1 & q`2<0
 & q`1/|.q.|>cn holds (for p being Point of TOP-REAL 2 st
 p=(cn-FanMorphS).q holds p`2<0 & p`1>0)
proof let cn be Real,q be Point of TOP-REAL 2;
 assume A1: -1<cn & cn<1 & q`2<0 & q`1/|.q.|>cn;
 let p be Point of TOP-REAL 2;
  assume A2: p=(cn-FanMorphS).q;
   then A3: p`2<0 & p`1>=0 by A1,JGRAPH_4:144;
     now assume A4: p`1=0;
     then (|.p.|)^2=(p`2)^2+0 by JGRAPH_3:10,SQUARE_1:60 .=(p`2)^2;
     then A5: p`2=|.p.| or p`2= - |.p.| by JGRAPH_3:1;
     then A6: |.p.| <> 0 by A1,A2,JGRAPH_4:144;
     set p1=(1/|.p.|)*p;
     A7: |.p.|*p1=(|.p.|*(1/|.p.|))*p by EUCLID:34
      .=(1)*p by A6,XCMPLX_1:107 .=p by EUCLID:33;
     A8: p1=|[(1/|.p.|)*p`1,(1/|.p.|)*p`2]| by EUCLID:61;
     then p1`2=(-|.p.|)*(1/|.p.|) by A3,A5,EUCLID:56,TOPRNS_1:26 .=-(|.p.|*(1/
|.p.|)) by XCMPLX_1:175
       .=-1 by A6,XCMPLX_1:107;
     then A9: p=|.p.|*(|[0,-1]|) by A4,A7,A8,EUCLID:56;
     set q1=(|.p.|)*|[cn,-sqrt(1-cn^2)]|;
     A10:(|[cn,-sqrt(1-cn^2)]|)`1=cn by EUCLID:56;
     A11:(|[cn,-sqrt(1-cn^2)]|)`2=-sqrt(1-cn^2) by EUCLID:56;
     then A12: q1=|[|.p.|*cn,|.p.|*(-sqrt(1-cn^2))]| by A10,EUCLID:61;
     then A13: q1`1=(|.p.|)*cn by EUCLID:56;
     A14: q1`2= (-sqrt(1-cn^2))*(|.p.|) by A12,EUCLID:56
     .=-(sqrt(1-cn^2)*(|.p.|)) by XCMPLX_1:175;
     A15: |.p.|>=0 by TOPRNS_1:26;
       1^2>cn^2 by A1,JGRAPH_2:8;
     then A16: 1-cn^2>0 by SQUARE_1:11,59;
     then sqrt(1-cn^2)>0 by SQUARE_1:93;
     then --sqrt(1-cn^2)*(|.p.|)>0 by A6,A15,SQUARE_1:21;
     then A17: q1`2<0 by A14,REAL_1:66;
     A18: |.q1.|=abs(|.p.|)*|.(|[cn,-sqrt(1-cn^2)]|).|
          by TOPRNS_1:8
           .=abs(|.p.|)*sqrt((cn)^2+(-sqrt(1-cn^2))^2)
                             by A10,A11,JGRAPH_3:10
           .=abs(|.p.|)*sqrt(cn^2+(sqrt(1-cn^2))^2) by SQUARE_1:61
           .=abs(|.p.|)*sqrt(cn^2+(1-cn^2)) by A16,SQUARE_1:def 4
           .=abs(|.p.|)*1 by SQUARE_1:83,XCMPLX_1:27
           .=|.p.| by A15,ABSVALUE:def 1;
     then A19: q1`1/|.q1.|=cn by A6,A13,XCMPLX_1:90;
     set p2=(cn-FanMorphS).q1;
     A20: p2`2<0 & p2`1=0 by A1,A17,A19,JGRAPH_4:149;
     then (|.p2.|)^2=(p2`2)^2+0 by JGRAPH_3:10,SQUARE_1:60 .=(p2`2)^2;
     then A21: p2`2=|.p2.| or p2`2= - |.p2.| by JGRAPH_3:1;
       |.p2.|=|.p.| by A18,JGRAPH_4:135;
     then A22: p2=|[0,-(|.p.|)]| by A20,A21,EUCLID:57,TOPRNS_1:26;
       (|[0,-1]|)`1=0 & (|[0,-1]|)`2=-1 by EUCLID:56;
     then A23: |.p.|*(|[0,-1]|)=|[|.p.|*0,|.p.|*(-1)]| by EUCLID:61
     .=|[0,-(|.p.|)]| by XCMPLX_1:180;
     A24: (cn-FanMorphS) is one-to-one by A1,JGRAPH_4:140;
       dom (cn-FanMorphS)=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
     then q1=q by A2,A9,A22,A23,A24,FUNCT_1:def 8;
    hence contradiction by A1,A6,A13,A18,XCMPLX_1:90;
   end;
  hence p`2<0 & p`1>0 by A1,A2,JGRAPH_4:144;
end;

theorem Th30:
 for cn being Real,q1,q2 being Point of TOP-REAL 2 st -1<cn & cn<1 & q1`2<=0
 & q2`2<=0 & |.q1.|<>0 & |.q2.|<>0 & q1`1/|.q1.|<q2`1/|.q2.|
 holds (for p1,p2 being Point of TOP-REAL 2 st
 p1=(cn-FanMorphS).q1 & p2=(cn-FanMorphS).q2 holds p1`1/|.p1.|<p2`1/|.p2.|)
 proof let cn be Real,q1,q2 be Point of TOP-REAL 2;
   assume A1: -1<cn & cn<1 & q1`2<=0 & q2`2<=0 & |.q1.|<>0 & |.q2.|<>0
     & q1`1/|.q1.|<q2`1/|.q2.|;
      now per cases by A1;
    case A2: q1`2<0;
       now per cases by A1;
     case q2`2<0;
      hence (for p1,p2 being Point of TOP-REAL 2 st
       p1=(cn-FanMorphS).q1 & p2=(cn-FanMorphS).q2 holds
       p1`1/|.p1.|<p2`1/|.p2.|) by A1,A2,JGRAPH_4:148;
     case A3: q2`2=0;
       then |.q2.|^2=(q2`1)^2+0 by JGRAPH_3:10,SQUARE_1:60 .=(q2`1)^2;
       then A4: |.q2.|=q2`1 or |.q2.|=-(q2`1) by JGRAPH_3:1;
         now assume |.q2.|=-(q2`1);
         then 1=(-(q2`1))/|.q2.| by A1,XCMPLX_1:60;
         then A5: q1`1/|.q1.|< -1 by A1,XCMPLX_1:191;
         A6: |.q1.|>=0 by TOPRNS_1:26;
           (|.q1.|)^2=(q1`1)^2+(q1`2)^2 by JGRAPH_3:10;
         then (|.q1.|)^2-(q1`1)^2=(q1`2)^2 by XCMPLX_1:26;
         then (|.q1.|)^2-(q1`1)^2>=0 by SQUARE_1:72;
         then (|.q1.|)^2-(q1`1)^2+(q1`1)^2>=0+(q1`1)^2 by REAL_1:55;
         then (|.q1.|)^2>=(q1`1)^2 by XCMPLX_1:27;
         then -|.q1.|<=q1`1 & q1`1<=|.q1.| by A6,JGRAPH_2:5;
         then (-|.q1.|)/|.q1.|<=q1`1/|.q1.| by A1,A6,REAL_1:73;
        hence contradiction by A1,A5,XCMPLX_1:198;
       end;
       then A7: q2`1/|.q2.|=1 by A1,A4,XCMPLX_1:60;
       thus for p1,p2 being Point of TOP-REAL 2 st
       p1=(cn-FanMorphS).q1 & p2=(cn-FanMorphS).q2 holds
       p1`1/|.p1.|<p2`1/|.p2.|
       proof let p1,p2 be Point of TOP-REAL 2;
        assume A8: p1=(cn-FanMorphS).q1 & p2=(cn-FanMorphS).q2;
         then A9: p2=q2 by A3,JGRAPH_4:120;
         A10: |.p1.|=|.q1.| by A8,JGRAPH_4:135;
         A11: |.p1.|>=0 by TOPRNS_1:26;
           (|.p1.|)^2=(p1`1)^2+(p1`2)^2 by JGRAPH_3:10;
         then A12: (|.p1.|)^2-(p1`1)^2=(p1`2)^2 by XCMPLX_1:26;
         then (|.p1.|)^2-(p1`1)^2>=0 by SQUARE_1:72;
         then (|.p1.|)^2-(p1`1)^2+(p1`1)^2>=0+(p1`1)^2 by REAL_1:55;
         then (|.p1.|)^2>=(p1`1)^2 by XCMPLX_1:27;
         then -|.p1.|<=p1`1 & p1`1<=|.p1.| by A11,JGRAPH_2:5;
         then (|.p1.|)/|.p1.|>=p1`1/|.p1.| by A1,A10,A11,REAL_1:73;
         then A13: 1>= p1`1/|.p1.| by A1,A10,XCMPLX_1:60;
         A14: p1`2<0 by A1,A2,A8,Th28;
           now assume 1= p1`1/|.p1.|; then (1)*|.p1.|=p1`1 by A1,A10,XCMPLX_1:
88
;
           then (|.p1.|)^2-(p1`1)^2=0 by XCMPLX_1:14;
          hence contradiction by A12,A14,SQUARE_1:73;
         end;
        hence p1`1/|.p1.|<p2`1/|.p2.| by A7,A9,A13,REAL_1:def 5;
       end;
     end;
       hence (for p1,p2 being Point of TOP-REAL 2 st
        p1=(cn-FanMorphS).q1 & p2=(cn-FanMorphS).q2 holds
        p1`1/|.p1.|<p2`1/|.p2.|);
    case A15: q1`2=0;
       then |.q1.|^2=(q1`1)^2+0 by JGRAPH_3:10,SQUARE_1:60 .=(q1`1)^2;
       then A16: |.q1.|=q1`1 or |.q1.|=-(q1`1) by JGRAPH_3:1;
         now assume |.q1.|=(q1`1);
         then A17: q2`1/|.q2.|> 1 by A1,XCMPLX_1:60;
         A18: |.q2.|>=0 by TOPRNS_1:26;
           (|.q2.|)^2=(q2`1)^2+(q2`2)^2 by JGRAPH_3:10;
         then (|.q2.|)^2-(q2`1)^2=(q2`2)^2 by XCMPLX_1:26;
         then (|.q2.|)^2-(q2`1)^2>=0 by SQUARE_1:72;
         then (|.q2.|)^2-(q2`1)^2+(q2`1)^2>=0+(q2`1)^2 by REAL_1:55;
         then (|.q2.|)^2>=(q2`1)^2 by XCMPLX_1:27;
         then -|.q2.|<=q2`1 & q2`1<=|.q2.| by A18,JGRAPH_2:5;
         then (|.q2.|)/|.q2.|>=q2`1/|.q2.| by A1,A18,REAL_1:73;
        hence contradiction by A1,A17,XCMPLX_1:60;
       end;
       then (-(q1`1))/|.q1.|=1 by A1,A16,XCMPLX_1:60;
       then A19: -(q1`1/|.q1.|)=1 by XCMPLX_1:188;
     thus (for p1,p2 being Point of TOP-REAL 2 st
       p1=(cn-FanMorphS).q1 & p2=(cn-FanMorphS).q2
       holds p1`1/|.p1.|<p2`1/|.p2.|)
       proof let p1,p2 be Point of TOP-REAL 2;
        assume A20: p1=(cn-FanMorphS).q1 & p2=(cn-FanMorphS).q2;
         then A21: p1=q1 by A15,JGRAPH_4:120;
         A22: |.p2.|=|.q2.| by A20,JGRAPH_4:135;
         A23: |.p2.|>=0 by TOPRNS_1:26;
           (|.p2.|)^2=(p2`1)^2+(p2`2)^2 by JGRAPH_3:10;
         then A24: (|.p2.|)^2-(p2`1)^2=(p2`2)^2 by XCMPLX_1:26;
         then (|.p2.|)^2-(p2`1)^2>=0 by SQUARE_1:72;
         then (|.p2.|)^2-(p2`1)^2+(p2`1)^2>=0+(p2`1)^2 by REAL_1:55;
         then (|.p2.|)^2>=(p2`1)^2 by XCMPLX_1:27;
         then -|.p2.|<=p2`1 & p2`1<=|.p2.| by A23,JGRAPH_2:5;
         then (-|.p2.|)/|.p2.|<=p2`1/|.p2.| by A1,A22,A23,REAL_1:73;
         then A25: -1 <= p2`1/|.p2.| by A1,A22,XCMPLX_1:198;
           now per cases;
         case q2`2=0;
           then p2=q2 by A20,JGRAPH_4:120;
          hence p2`1/|.p2.|>-1 by A1,A19;
         case q2`2<>0;
         then A26: p2`2<0 by A1,A20,Th28;
           now assume -1= p2`1/|.p2.|; then (-1)*|.p2.|=p2`1 by A1,A22,XCMPLX_1
:88;
           then (-|.p2.|)^2=(p2`1)^2 by XCMPLX_1:180;
           then (|.p2.|)^2=(p2`1)^2 by SQUARE_1:61;
           then (|.p2.|)^2-(p2`1)^2=0 by XCMPLX_1:14;
          hence contradiction by A24,A26,SQUARE_1:73;
         end;
         hence
           p2`1/|.p2.|>-1 by A25,REAL_1:def 5;
         end;
        hence p1`1/|.p1.|<p2`1/|.p2.| by A19,A21;
       end;
    end;
  hence thesis;
 end;

begin :: Order of Points on Circle

Lm2: now let P be compact non empty Subset of TOP-REAL 2;
  assume A1: P={q: |.q.|=1};
  A2: proj1.:P c= [.-1,1.]
   proof let y be set;assume y in proj1.:P;
     then consider x being set such that
     A3: x in dom proj1 & x in P & y=proj1.x by FUNCT_1:def 12;
     reconsider q=x as Point of TOP-REAL 2 by A3;
     A4: y=q`1 by A3,PSCOMP_1:def 28;
     consider q2 being Point of TOP-REAL 2 such that
     A5: q2=x & |.q2.|=1 by A1,A3;
       (q`1)^2+(q`2)^2=1 by A5,JGRAPH_3:10,SQUARE_1:59;
     then A6: (q`2)^2=1-(q`1)^2 by XCMPLX_1:26;
       0<=(q`2)^2 by SQUARE_1:72;
     then 1-(q`1)^2+(q`1)^2 >=0+(q`1)^2 by A6,REAL_1:55;
     then 1>=(q`1)^2 by XCMPLX_1:27;
     then -1<=q`1 & q`1<=1 by JGRAPH_4:4;
    hence y in [.-1,1.] by A4,TOPREAL5:1;
   end;
    [.-1,1.] c= proj1.:P
   proof let y be set;assume
       y in [.-1,1.];
     then y in {r where r is Real: -1<=r & r<=1 } by RCOMP_1:def 1;
     then consider r being Real such that
     A7: y=r & -1<=r & r<=1;
     set q=|[r,sqrt(1-r^2)]|;
     A8: dom proj1=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
     A9: q`1=r & q`2=sqrt(1-r^2) by EUCLID:56;
       1^2>=r^2 by A7,JGRAPH_2:7;
     then A10: 1-r^2>=0 by SQUARE_1:12,59;
       |.q.|=sqrt(r^2+(sqrt(1-r^2))^2) by A9,JGRAPH_3:10
           .=sqrt(r^2+(1-r^2)) by A10,SQUARE_1:def 4
           .=1 by SQUARE_1:83,XCMPLX_1:27;
     then A11: q in P by A1;
       proj1.q=q`1 by PSCOMP_1:def 28 .=r by EUCLID:56;
    hence y in proj1.:P by A7,A8,A11,FUNCT_1:def 12;
   end;
  hence proj1.:P=[.-1,1.] by A2,XBOOLE_0:def 10;
  A12: proj2.:P c= [.-1,1.]
   proof let y be set;assume
       y in proj2.:P;
     then consider x being set such that
     A13: x in dom proj2 & x in P & y=proj2.x by FUNCT_1:def 12;
     reconsider q=x as Point of TOP-REAL 2 by A13;
     A14: y=q`2 by A13,PSCOMP_1:def 29;
     consider q2 being Point of TOP-REAL 2 such that
     A15: q2=x & |.q2.|=1 by A1,A13;
       (q`1)^2+(q`2)^2=1 by A15,JGRAPH_3:10,SQUARE_1:59;
     then A16: (q`1)^2=1-(q`2)^2 by XCMPLX_1:26;
       0<=(q`1)^2 by SQUARE_1:72;
     then 1-(q`2)^2+(q`2)^2 >=0+(q`2)^2 by A16,REAL_1:55;
     then 1>=(q`2)^2 by XCMPLX_1:27;
     then -1<=q`2 & q`2<=1 by JGRAPH_4:4;
    hence y in [.-1,1.] by A14,TOPREAL5:1;
   end;
    [.-1,1.] c= proj2.:P
   proof let y be set;assume
       y in [.-1,1.];
     then y in {r where r is Real: -1<=r & r<=1 } by RCOMP_1:def 1;
     then consider r being Real such that
     A17: y=r & -1<=r & r<=1;
     set q=|[sqrt(1-r^2),r]|;
     A18: dom proj2=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
     A19: q`2=r & q`1=sqrt(1-r^2) by EUCLID:56;
       1^2>=r^2 by A17,JGRAPH_2:7;
     then A20: 1-r^2>=0 by SQUARE_1:12,59;
       |.q.|=sqrt((sqrt(1-r^2))^2+r^2) by A19,JGRAPH_3:10
           .=sqrt((1-r^2)+r^2) by A20,SQUARE_1:def 4
           .=1 by SQUARE_1:83,XCMPLX_1:27;
     then A21: q in P by A1;
       proj2.q=q`2 by PSCOMP_1:def 29 .=r by EUCLID:56;
    hence y in proj2.:P by A17,A18,A21,FUNCT_1:def 12;
   end;
  hence proj2.:P=[.-1,1.] by A12,XBOOLE_0:def 10;
end;

Lm3: for P being compact non empty Subset of TOP-REAL 2 st
 P={q: |.q.|=1} holds W-bound(P)=-1
proof let P be compact non empty Subset of TOP-REAL 2;
  assume A1: P={q: |.q.|=1};
  A2: the carrier of ((TOP-REAL 2)|P) = P by JORDAN1:1;
    proj1.:P=[.-1,1.] by A1,Lm2;
  then (proj1|P).:P=[.-1,1.] by RFUNCT_2:5;
  then (proj1||P).:P=[.-1,1.] by PSCOMP_1:def 26;
  then inf ((proj1||P).:P)=-1 by JORDAN5A:20;
  then inf (proj1||P)=-1 by A2,PSCOMP_1:def 20;
 hence W-bound P=-1 by PSCOMP_1:def 30;
end;

theorem Th31: for P being compact non empty Subset of TOP-REAL 2 st
 P={q: |.q.|=1}
 holds W-bound(P)=-1 & E-bound(P)=1 & S-bound(P)=-1 & N-bound(P)=1
proof let P be compact non empty Subset of TOP-REAL 2;
  assume A1: P={q: |.q.|=1};
   hence W-bound(P)=-1 by Lm3;
  A2: the carrier of ((TOP-REAL 2)|P) =P by JORDAN1:1;
    proj1.:P=[.-1,1.] by A1,Lm2;
  then (proj1|P).:P=[.-1,1.] by RFUNCT_2:5;
  then (proj1||P).:P=[.-1,1.] by PSCOMP_1:def 26;
  then sup ((proj1||P).:the carrier of ((TOP-REAL 2)|P))=1 by A2,JORDAN5A:20;
  then sup (proj1||P)=1 by PSCOMP_1:def 21;
 hence E-bound P=1 by PSCOMP_1:def 32;
    proj2.:P=[.-1,1.] by A1,Lm2;
  then (proj2|P).:P=[.-1,1.] by RFUNCT_2:5;
  then A3: (proj2||P).:P=[.-1,1.] by PSCOMP_1:def 26;
  then inf ((proj2||P).:P)=-1 by JORDAN5A:20;
  then inf (proj2||P)=-1 by A2,PSCOMP_1:def 20;
 hence S-bound P=-1 by PSCOMP_1:def 33;
    sup ((proj2||P).:P)=1 by A3,JORDAN5A:20;
  then sup (proj2||P)=1 by A2,PSCOMP_1:def 21;
 hence N-bound P=1 by PSCOMP_1:def 31;
end;

theorem Th32: for P being compact non empty Subset of TOP-REAL 2 st
 P={q: |.q.|=1}
 holds W-min(P)=|[-1,0]|
proof let P be compact non empty Subset of TOP-REAL 2;
  assume A1: P={q: |.q.|=1};
  A2: the carrier of ((TOP-REAL 2)|P) = P by JORDAN1:1;
  A3: W-bound P=-1 by A1,Lm3;
    proj2.:P=[.-1,1.] by A1,Lm2;
  then (proj2|P).:P=[.-1,1.] by RFUNCT_2:5;
  then A4: (proj2||P).:P=[.-1,1.] by PSCOMP_1:def 26;
  then inf ((proj2||P).:P)=-1 by JORDAN5A:20;
  then inf (proj2||P)=-1 by A2,PSCOMP_1:def 20;
  then S-bound P=-1 by PSCOMP_1:def 33;
  then A5: SW-corner P=|[-1,-1]| by A3,PSCOMP_1:def 34;
    sup ((proj2||P).:P)=1 by A4,JORDAN5A:20;
  then sup (proj2||P)=1 by A2,PSCOMP_1:def 21;
  then N-bound P=1 by PSCOMP_1:def 31;
  then A6: NW-corner P=|[-1,1]| by A3,PSCOMP_1:def 35;
  A7: {|[-1,0]|} c= LSeg(SW-corner P, NW-corner P)/\P
   proof let x be set;assume
       x in {|[-1,0]|};
     then A8: x=|[-1,0]| by TARSKI:def 1;
     set q=|[-1,0]|;
       q`2=0 & q`1=-1 by EUCLID:56;
     then |.q.|=sqrt((-1)^2+0^2) by JGRAPH_3:10
           .=1 by SQUARE_1:59,60,61,83;
     then A9: x in P by A1,A8;
       q=|[(1/2)*(-1)+(1/2)*(-1),(1/2)*(-1)+(1/2)*1]|;
     then q=|[(1/2)*(-1),(1/2)*(-1)]|+|[(1/2)*(-1),(1/2)*1]| by EUCLID:60;
     then q=|[(1/2)*(-1),(1/2)*(-1)]|+(1/2)*|[-1,1]| by EUCLID:62;
     then q=(1/2)*|[-1,-1]|+(1-(1/2))*|[-1,1]| by EUCLID:62;
     then q in { (1-l)*(SW-corner P) + l*(NW-corner P) where l is Real:
             0 <= l & l <= 1 } by A5,A6;
     then x in LSeg(SW-corner P, NW-corner P) by A8,TOPREAL1:def 4;
    hence x in LSeg(SW-corner P, NW-corner P)/\P by A9,XBOOLE_0:def 3;
   end;
    LSeg(SW-corner P, NW-corner P)/\P c= {|[-1,0]|}
   proof let x be set;assume x in LSeg(SW-corner P, NW-corner P)/\P;
     then A10: x in LSeg(SW-corner P, NW-corner P) & x in P by XBOOLE_0:def 3;
     then x in { (1-l)*(SW-corner P) + l*(NW-corner P) where l is Real:
             0 <= l & l <= 1 } by TOPREAL1:def 4;
     then consider l being Real such that
     A11: x=(1-l)*(SW-corner P)+l*(NW-corner P)
       & 0<=l & l<=1;
       x=|[(1-l)*(-1),(1-l)*(-1)]|+(l)*|[-1,1]| by A5,A6,A11,EUCLID:62;
     then x=|[(1-l)*(-1),(1-l)*(-1)]|+|[(l)*(-1),(l)*1]| by EUCLID:62;
     then x=|[(1-l)*(-1)+(l)*(-1),(1-l)*(-1)+(l)*1]| by EUCLID:60;
     then x=|[((1-l)+l)*(-1),(1-l)*(-1)+(l)*1]| by XCMPLX_1:8;
then A12:  x=|[(1)*(-1),(1-l)*(-1)+(l)*1]| by XCMPLX_1:27;
     reconsider q3=x as Point of TOP-REAL 2 by A11;
     A13: q3`1=-1 & q3`2=(1-l)*(-1)+l by A12,EUCLID:56;
     consider q2 being Point of TOP-REAL 2 such that
     A14: q2=x & |.q2.|=1 by A1,A10;
     A15: 1=sqrt((-1)^2+(q3`2)^2) by A13,A14,JGRAPH_3:10
     .=sqrt(1+(q3`2)^2) by SQUARE_1:59,61;
      now assume (q3`2)^2>0; then 1<1+(q3`2)^2 by REAL_1:69;
      hence contradiction by A15,SQUARE_1:83,95;
     end;
     then (q3`2)^2=0 by SQUARE_1:72;
     then q3`2=0 by SQUARE_1:73;
    hence x in {|[-1,0]|} by A12,A13,TARSKI:def 1;
   end;
  then LSeg(SW-corner P, NW-corner P)/\P={|[-1,0]|} by A7,XBOOLE_0:def 10;
  then A16: W-most P={|[-1,0]|} by PSCOMP_1:def 38;
    (proj2||W-most P).:the carrier of ((TOP-REAL 2)|(W-most P))
    =(proj2||W-most P).:(W-most P) by JORDAN1:1
   .=(proj2|(W-most P)).:(W-most P) by PSCOMP_1:def 26
   .=proj2.:(W-most P) by RFUNCT_2:5
   .={proj2.(|[-1,0]|)} by A16,SETWISEO:13
   .={(|[-1,0]|)`2} by PSCOMP_1:def 29
   .={0} by EUCLID:56;
  then inf ((proj2||W-most P).:the carrier of ((TOP-REAL 2)|(W-most P)))
    =0 by SEQ_4:22;
  then inf (proj2||W-most P)=0 by PSCOMP_1:def 20;
 hence W-min(P)=|[-1,0]| by A3,PSCOMP_1:def 42;
end;

theorem Th33: for P being compact non empty Subset of TOP-REAL 2 st
 P={q: |.q.|=1}
 holds E-max(P)=|[1,0]|
proof let P be compact non empty Subset of TOP-REAL 2;
  assume A1: P={q: |.q.|=1};
  A2: the carrier of ((TOP-REAL 2)|P) =P by JORDAN1:1;
  A3: E-bound P=1 by A1,Th31;
    proj2.:P=[.-1,1.] by A1,Lm2;
  then (proj2|P).:P=[.-1,1.] by RFUNCT_2:5;
  then A4: (proj2||P).:P=[.-1,1.] by PSCOMP_1:def 26;
  then inf ((proj2||P).:P)=-1 by JORDAN5A:20;
  then inf (proj2||P)=-1 by A2,PSCOMP_1:def 20;
  then S-bound P=-1 by PSCOMP_1:def 33;
  then A5: SE-corner P=|[1,-1]| by A3,PSCOMP_1:def 37;
    sup ((proj2||P).:P)=1 by A4,JORDAN5A:20;
  then sup (proj2||P)=1 by A2,PSCOMP_1:def 21;
  then N-bound P=1 by PSCOMP_1:def 31;
  then A6: NE-corner P=|[1,1]| by A3,PSCOMP_1:def 36;
  A7: {|[1,0]|} c= LSeg(SE-corner P, NE-corner P)/\P
   proof let x be set;assume
       x in {|[1,0]|};
     then A8: x=|[1,0]| by TARSKI:def 1;
     set q=|[1,0]|;
       q`2=0 & q`1=1 by EUCLID:56;
     then |.q.|=sqrt((1)^2+0^2) by JGRAPH_3:10
           .=1 by SQUARE_1:59,60,83;
     then A9: x in P by A1,A8;
       q=|[(1/2)*(1)+(1/2)*(1),(1/2)*(-1)+(1/2)*1]|;
     then q=|[(1/2)*(1),(1/2)*(-1)]|+|[(1/2)*(1),(1/2)*1]| by EUCLID:60;
     then q=|[(1/2)*(1),(1/2)*(-1)]|+(1/2)*|[1,1]| by EUCLID:62;
     then q=(1/2)*|[1,-1]|+(1-(1/2))*|[1,1]| by EUCLID:62;
     then q in { (1-l)*(SE-corner P) + l*(NE-corner P) where l is Real:
             0 <= l & l <= 1 } by A5,A6;
     then x in LSeg(SE-corner P, NE-corner P) by A8,TOPREAL1:def 4;
    hence x in LSeg(SE-corner P, NE-corner P)/\P by A9,XBOOLE_0:def 3;
   end;
    LSeg(SE-corner P, NE-corner P)/\P c= {|[1,0]|}
   proof let x be set;assume x in LSeg(SE-corner P, NE-corner P)/\P;
     then A10: x in LSeg(SE-corner P, NE-corner P) & x in P by XBOOLE_0:def 3;
     then x in { (1-l)*(SE-corner P) + l*(NE-corner P) where l is Real:
             0 <= l & l <= 1 } by TOPREAL1:def 4;
     then consider l being Real such that
     A11: x=(1-l)*(SE-corner P)+l*(NE-corner P)
       & 0<=l & l<=1;
       x=|[(1-l)*(1),(1-l)*(-1)]|+(l)*|[1,1]| by A5,A6,A11,EUCLID:62;
     then x=|[(1-l)*(1),(1-l)*(-1)]|+|[(l)*(1),(l)*1]| by EUCLID:62;
     then x=|[((1-l)+l)*(1),(1-l)*(-1)+(l)*1]| by EUCLID:60;
     then A12: x=|[1,(1-l)*(-1)+l]| by XCMPLX_1:27;
     reconsider q3=x as Point of TOP-REAL 2 by A11;
     A13: q3`1=1 & q3`2=(1-l)*(-1)+l by A12,EUCLID:56;
     consider q2 being Point of TOP-REAL 2 such that
     A14: q2=x & |.q2.|=1 by A1,A10;
      now assume (q3`2)^2>0; then 1<1+(q3`2)^2 by REAL_1:69;
      hence contradiction by A13,A14,JGRAPH_3:10,SQUARE_1:59;
     end;
     then (q3`2)^2=0 by SQUARE_1:72;
     then q3`2=0 by SQUARE_1:73;
    hence x in {|[1,0]|} by A12,A13,TARSKI:def 1;
   end;
  then LSeg(SE-corner P, NE-corner P)/\P={|[1,0]|} by A7,XBOOLE_0:def 10;
  then A15: E-most P={|[1,0]|} by PSCOMP_1:def 40;
    (proj2||E-most P).:the carrier of ((TOP-REAL 2)|(E-most P))
    =(proj2||E-most P).:(E-most P) by JORDAN1:1
   .=(proj2|(E-most P)).:(E-most P) by PSCOMP_1:def 26
   .=proj2.:(E-most P) by RFUNCT_2:5
   .={proj2.(|[1,0]|)} by A15,SETWISEO:13
   .={(|[1,0]|)`2} by PSCOMP_1:def 29
   .={0} by EUCLID:56;
  then sup ((proj2||E-most P).:the carrier of ((TOP-REAL 2)|(E-most P)))
    =0 by SEQ_4:22;
  then sup (proj2||E-most P)=0 by PSCOMP_1:def 21;
 hence E-max(P)=|[1,0]| by A3,PSCOMP_1:def 46;
end;

theorem
  for f being map of TOP-REAL 2,R^1 st
(for p being Point of TOP-REAL 2 holds f.p=proj1.p) holds f is continuous
proof let f be map of TOP-REAL 2,R^1;
 assume A1:for p being Point of TOP-REAL 2 holds f.p=proj1.p;
   (TOP-REAL 2)|([#](TOP-REAL 2))=TOP-REAL 2 by TSEP_1:3;
 hence f is continuous by A1,JGRAPH_2:39;
end;

theorem Th35:
for f being map of TOP-REAL 2,R^1 st
(for p being Point of TOP-REAL 2 holds f.p=proj2.p) holds f is continuous
proof let f be map of TOP-REAL 2,R^1;
 assume A1:for p being Point of TOP-REAL 2 holds f.p=proj2.p;
   (TOP-REAL 2)|([#](TOP-REAL 2))=TOP-REAL 2 by TSEP_1:3;
 hence f is continuous by A1,JGRAPH_2:40;
end;

theorem Th36: for P being compact non empty Subset of TOP-REAL 2 st
 P={q where q is Point of TOP-REAL 2: |.q.|=1} holds
 Upper_Arc(P) c= P & Lower_Arc(P) c= P
proof let P be compact non empty Subset of TOP-REAL 2;
 assume P={q where q is Point of TOP-REAL 2: |.q.|=1};
 then P is_simple_closed_curve by JGRAPH_3:36;
 hence thesis by JORDAN1A:16;
end;

theorem Th37: for P being compact non empty Subset of TOP-REAL 2 st
P={q where q is Point of TOP-REAL 2: |.q.|=1} holds
Upper_Arc(P)={p where p is Point of TOP-REAL 2:p in P & p`2>=0}
proof let P be compact non empty Subset of TOP-REAL 2;
 assume A1: P={q where q is Point of TOP-REAL 2: |.q.|=1};
  then A2: P is_simple_closed_curve by JGRAPH_3:36;
  then A3: Upper_Arc(P) is_an_arc_of W-min(P),E-max(P) by JORDAN6:def 8;
  A4: Lower_Arc(P) is_an_arc_of E-max(P),W-min(P) by A2,JORDAN6:def 9;
  consider P2 being non empty Subset of TOP-REAL 2
   such that
   A5: P2 is_an_arc_of E-max(P),W-min(P)
     & Upper_Arc(P) /\ P2={W-min(P),E-max(P)}
     & Upper_Arc(P) \/ P2=P
     & First_Point(Upper_Arc(P),W-min(P),E-max(P),
        Vertical_Line((W-bound(P)+E-bound(P))/2))`2>
       Last_Point(P2,E-max(P),W-min(P),
        Vertical_Line((W-bound(P)+E-bound(P))/2))`2 by A2,JORDAN6:def 8;
   set P4=Lower_Arc(P);
   A6: P4 is_an_arc_of E-max(P),W-min(P)
     & Upper_Arc(P) /\ P4={W-min(P),E-max(P)}
       & Upper_Arc(P) \/ P4=P
         & First_Point(Upper_Arc(P),W-min(P),E-max(P),
         Vertical_Line((W-bound(P)+E-bound(P))/2))`2>
              Last_Point(P4,E-max(P),W-min(P),
         Vertical_Line((W-bound(P)+E-bound(P))/2))`2 by A2,JORDAN6:def 9;
   A7: W-bound(P)=-1 by A1,Th31;
   A8: E-bound(P)=1 by A1,Th31;
   A9: Vertical_Line(0)={p where p is Point of TOP-REAL 2: p`1=0}
                         by JORDAN6:def 6;
     set P1=Upper_Arc(P), P2=Lower_Arc(P), Q=Vertical_Line(0);
     set p11=W-min(P), p22=E-max(P);
     set p8= First_Point(Upper_Arc(P),W-min(P),E-max(P),
         Vertical_Line(0));
     set pj= Last_Point(Lower_Arc(P),E-max(P),W-min(P),
         Vertical_Line(0));
        A10: W-bound P=-1 by A1,Th31;
        A11: E-bound P=1 by A1,Th31;
        A12: S-bound P=-1 by A1,Th31;
        A13: N-bound P=1 by A1,Th31;
        then A14: LSeg(|[0,-1]|,|[0,1]|) meets Upper_Arc P
                     by A2,A10,A11,A12,JORDAN1B:26;
         A15: LSeg(|[0,-1]|,|[0,1]|) c= Q
         proof let x be set;assume x in LSeg(|[0,-1]|,|[0,1]|);
           then x in {(1-l)*(|[0,-1]|) +l*(|[0,1]|) where l is Real: 0<=l & l
<=1}
                                     by TOPREAL1:def 4;
           then consider l being Real such that
           A16: x=(1-l)*(|[0,-1]|) +l*(|[0,1]|) & 0<=l & l<=1;
             ((1-l)*(|[0,-1]|) +l*(|[0,1]|))`1
             = ((1-l)*(|[0,-1]|))`1 +(l*(|[0,1]|))`1 by TOPREAL3:7
             .=(1-l)*(|[0,-1]|)`1+(l*(|[0,1]|))`1 by TOPREAL3:9
             .=(1-l)*(|[0,-1]|)`1+l*((|[0,1]|))`1 by TOPREAL3:9
             .=(1-l)*0+l*((|[0,1]|))`1 by EUCLID:56
             .=(1-l)*0+l*0 by EUCLID:56
             .=0;
          hence x in Q by A9,A16;
         end;
        then A17: P1 meets Q by A14,XBOOLE_1:64;
        A18: Upper_Arc(P) is closed by A3,JORDAN6:12;
          Vertical_Line(0) is closed by JORDAN6:33;
        then P1 /\ Q is closed by A18,TOPS_1:35;
        then A19: p8 in P1 /\ Q &
       for g being map of I[01], (TOP-REAL 2)|P1, s2 being Real st
           g is_homeomorphism & g.0 = p11 & g.1 = p22
           & g.s2 = p8 & 0 <= s2 & s2 <= 1 holds
       (for t being Real st 0 <= t & t < s2 holds not g.t in Q)
                       by A3,A17,JORDAN5C:def 1;
         P1 /\ Q c= {|[0,-1]|,|[0,1]|}
        proof let x be set;assume x in P1 /\ Q;
          then A20: x in P1 & x in Q by XBOOLE_0:def 3;
          then consider p being Point of TOP-REAL 2 such that
          A21: p=x & p`1=0 by A9;
            x in P by A5,A20,XBOOLE_0:def 2;
          then consider q being Point of TOP-REAL 2 such that
          A22: q=x & |.q.|=1 by A1;
             0+(q`2)^2 =1 by A21,A22,JGRAPH_3:10,SQUARE_1:59,60;
          then q`2=1 or q`2=-1 by JGRAPH_3:2;
          then x=|[0,-1]| or x=|[0,1]| by A21,A22,EUCLID:57;
         hence x in {|[0,-1]|,|[0,1]|} by TARSKI:def 2;
        end;
       then p8=|[0,-1]| or p8=|[0,1]| by A19,TARSKI:def 2;
       then A23: p8`2=-1 or p8`2=1 by EUCLID:56;
          LSeg(|[0,-1]|,|[0,1]|) meets Lower_Arc P
                     by A2,A10,A11,A12,A13,JORDAN1B:27;
        then A24: P2 meets Q by A15,XBOOLE_1:64;
        A25: Lower_Arc(P) is closed by A4,JORDAN6:12;
          Vertical_Line(0) is closed by JORDAN6:33;
        then P2 /\ Q is closed by A25,TOPS_1:35;
        then A26: pj in P2 /\ Q &
       for g being map of I[01], (TOP-REAL 2)|P2, s2 being Real st
           g is_homeomorphism & g.0 = p22 & g.1 = p11
           & g.s2 = pj & 0 <= s2 & s2 <= 1 holds
       for t being Real st 1 >= t & t > s2 holds not g.t in Q
                       by A4,A24,JORDAN5C:def 2;
         P2 /\ Q c= {|[0,-1]|,|[0,1]|}
        proof let x be set;assume x in P2 /\ Q;
          then A27: x in P2 & x in Q by XBOOLE_0:def 3;
          then consider p being Point of TOP-REAL 2 such that
          A28: p=x & p`1=0 by A9;
            x in P by A6,A27,XBOOLE_0:def 2;
          then consider q being Point of TOP-REAL 2 such that
          A29: q=x & |.q.|=1 by A1;
             0+(q`2)^2 =1 by A28,A29,JGRAPH_3:10,SQUARE_1:59,60;
          then q`2=1 or q`2=-1 by JGRAPH_3:2;
          then x=|[0,-1]| or x=|[0,1]| by A28,A29,EUCLID:57;
         hence x in {|[0,-1]|,|[0,1]|} by TARSKI:def 2;
        end;
       then pj=|[0,-1]| or pj=|[0,1]| by A26,TARSKI:def 2;
       then A30: pj`2=-1 or pj`2=1 by EUCLID:56;
  A31: p8 in P1 by A19,XBOOLE_0:def 3;
  A32: Upper_Arc(P) c= P by A5,XBOOLE_1:7;
  A33: Lower_Arc(P) c= P by A6,XBOOLE_1:7;
    E-max(P) in {W-min(P),E-max(P)} by TARSKI:def 2;
  then A34: E-max(P) in Upper_Arc(P) by A5,XBOOLE_0:def 3;
    W-min(P) in {W-min(P),E-max(P)} by TARSKI:def 2;
  then A35: W-min(P) in Upper_Arc(P) by A5,XBOOLE_0:def 3;
  reconsider R=Upper_Arc(P) as non empty Subset of TOP-REAL 2;
  consider f being map of I[01], (TOP-REAL 2)|R such that
   A36: f is_homeomorphism & f.0 =W-min(P) &
      f.1 =E-max(P) by A3,TOPREAL1:def 2;
         rng f =[#]((TOP-REAL 2)|R) by A36,TOPS_2:def 5
            .=R by PRE_TOPC:def 10;
       then consider x8 being set such that
       A37: x8 in dom f & p8=f.x8 by A31,FUNCT_1:def 5;
         dom f= [.0,1.]
                                  by BORSUK_1:83,FUNCT_2:def 1;
       then x8 in {r where r is Real: 0<=r & r<=1 } by A37,RCOMP_1:def 1;
       then consider r8 being Real such that
       A38: x8=r8 & 0<=r8 & r8<=1;
        A39: now assume r8=0;
         then p8=|[-1,0]| by A1,A36,A37,A38,Th32;
        hence contradiction by A23,EUCLID:56;
       end;
         now assume r8=1;
         then p8=|[1,0]| by A1,A36,A37,A38,Th33;
        hence contradiction by A23,EUCLID:56;
       end;
       then A40: 1>r8 by A38,REAL_1:def 5;
   reconsider h2=proj2 as map of TOP-REAL 2,R^1 by JGRAPH_2:17;
   A41: f is continuous by A36,TOPS_2:def 5;
   A42: f is one-to-one by A36,TOPS_2:def 5;
     for p being Point of (TOP-REAL 2) holds h2.p=proj2.p;
   then A43: h2 is continuous by Th35;
   A44: dom f=the carrier of I[01] by FUNCT_2:def 1;
   A45: the carrier of ((TOP-REAL 2)|R)=R by JORDAN1:1;
   then A46: rng f c= the carrier of TOP-REAL 2 by XBOOLE_1:1;
     dom h2=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
   then A47: dom (h2*f)=the carrier of I[01] by A44,A46,RELAT_1:46;
     rng (h2*f) c= rng h2 by RELAT_1:45;
   then rng (h2*f) c= the carrier of R^1 by XBOOLE_1:1;
   then h2*f is Function of the carrier of I[01],the carrier of R^1
                   by A47,FUNCT_2:4;
   then reconsider g0=h2*f as map of I[01],R^1;
  A48: (ex p being Point of TOP-REAL 2,t being Real st
  0<t & t<1 & f.t=p & p`2>0) implies for q being Point of TOP-REAL 2
  st q in Upper_Arc(P) holds q`2>=0
   proof assume
      (ex p being Point of TOP-REAL 2,t being Real st
     0<t & t<1 & f.t=p & p`2>0);
     then consider p being Point of TOP-REAL 2,t being Real such that
     A49: 0<t & t<1 & f.t=p & p`2>0;
       now assume ex q being Point of TOP-REAL 2
       st q in Upper_Arc(P) & q`2<0;
       then consider q being Point of TOP-REAL 2 such that
       A50: q in Upper_Arc(P) & q`2<0;
         rng f =[#]((TOP-REAL 2)|R) by A36,TOPS_2:def 5
            .=R by PRE_TOPC:def 10;
       then consider x being set such that
       A51: x in dom f & q=f.x by A50,FUNCT_1:def 5;
       A52: dom f= [.0,1.]
                                  by BORSUK_1:83,FUNCT_2:def 1;
       then x in {r where r is Real: 0<=r & r<=1 } by A51,RCOMP_1:def 1;
       then consider r being Real such that
       A53: x=r & 0<=r & r<=1;
       A54: (h2*f).r=h2.q by A51,A53,FUNCT_1:23
       .=q`2 by PSCOMP_1:def 29;
         t in {v where v is Real: 0<=v & v<=1 } by A49;
       then A55: t in [.0,1.] by RCOMP_1:def 1;
       then A56: (h2*f).t=h2.p by A49,A52,FUNCT_1:23
       .=p`2 by PSCOMP_1:def 29;
        now per cases by REAL_1:def 5;
      case A57: r<t;
         [.r,t.] c= [.0,1.] by A51,A52,A53,A55,RCOMP_1:16;
       then reconsider B=[.r,t.] as non empty Subset of I[01]
                               by A57,BORSUK_1:83,RCOMP_1:15;
       reconsider B0=B as Subset of I[01];
       reconsider g=g0|B0 as map of (I[01]|B0),R^1 by JGRAPH_3:12;
         g0 is continuous by A41,A43,Th10;
       then A58: g is continuous by TOPMETR:10;
       A59: Closed-Interval-TSpace(r,t)=I[01]|B by A49,A53,A57,Th6,TOPMETR:27;
          r in {r4 where r4 is Real: r<=r4 & r4<=t} by A57;
        then A60: r in B by RCOMP_1:def 1;
          t in {r4 where r4 is Real: r<=r4 & r4<=t} by A57;
        then t in B by RCOMP_1:def 1;
        then (q`2)*(p`2)<0 & q`2=g.r & p`2=g.t by A49,A50,A54,A56,A60,FUNCT_1:
72,SQUARE_1:24;
        then consider r1 being Real such that
        A61: g.r1=0 & r<r1 & r1<t by A57,A58,A59,TOPREAL5:14;
          r1 in {r4 where r4 is Real: r<=r4 & r4<=t} by A61;
        then A62: r1 in B by RCOMP_1:def 1;
        A63: 0<r1 by A53,A61;
          r1<1 by A49,A61,AXIOMS:22;
        then r1 in {r2 where r2 is Real: 0<=r2 & r2<=1} by A63;
        then A64: r1 in dom f by A52,RCOMP_1:def 1;
        then f.r1 in rng f by FUNCT_1:def 5;
        then f.r1 in R by A45;
        then f.r1 in P by A32;
        then consider q3 being Point of TOP-REAL 2 such that
        A65: q3=f.r1 & |.q3.|=1 by A1;
        A66: q3`2=h2.(f.r1) by A65,PSCOMP_1:def 29
          .=g0.r1 by A64,FUNCT_1:23
          .=0 by A61,A62,FUNCT_1:72;
        then A67: 1=(q3`1)^2 +0^2 by A65,JGRAPH_3:10,SQUARE_1:59
        .=(q3`1)^2 by SQUARE_1:60;
          now per cases by A67,JGRAPH_3:2;
        case q3`1=1;
           then A68: q3=|[1,0]| by A66,EUCLID:57
           .=E-max(P) by A1,Th33;
             1 in dom f by A52,RCOMP_1:15;
         hence contradiction
                  by A36,A42,A49,A61,A64,A65,A68,FUNCT_1:def 8;
        case q3`1=-1;
           then A69: q3=|[-1,0]| by A66,EUCLID:57
           .=W-min(P) by A1,Th32;
             0 in dom f by A52,RCOMP_1:15;
         hence contradiction
              by A36,A42,A53,A61,A64,A65,A69,FUNCT_1:def 8;
        end;
      hence contradiction;
      case A70: t<r;
         [.t,r.] c= [.0,1.] by A51,A52,A53,A55,RCOMP_1:16;
       then reconsider B=[.t,r.] as non empty Subset of I[01]
                               by A70,BORSUK_1:83,RCOMP_1:15;
       reconsider B0=B as Subset of I[01];
       reconsider g=g0|B0 as map of (I[01]|B0),R^1 by JGRAPH_3:12;
         g0 is continuous by A41,A43,Th10;
       then A71: g is continuous by TOPMETR:10;
       A72: Closed-Interval-TSpace(t,r)=I[01]|B by A49,A53,A70,Th6,TOPMETR:27;
          r in {r4 where r4 is Real: t<=r4 & r4<=r} by A70;
        then A73: r in B by RCOMP_1:def 1;
          t in {r4 where r4 is Real: t<=r4 & r4<=r} by A70;
        then t in B by RCOMP_1:def 1;
        then (q`2)*(p`2)<0 & q`2=g.r & p`2=g.t by A49,A50,A54,A56,A73,FUNCT_1:
72,SQUARE_1:24;
        then consider r1 being Real such that
        A74: g.r1=0 & t<r1 & r1<r by A70,A71,A72,TOPREAL5:14;
          r1 in {r4 where r4 is Real: t<=r4 & r4<=r} by A74;
        then A75: r1 in B by RCOMP_1:def 1;
        A76: 0<r1 by A49,A74,AXIOMS:22;
          r1<1 by A53,A74,AXIOMS:22;
        then r1 in {r2 where r2 is Real: 0<=r2 & r2<=1} by A76;
        then A77: r1 in dom f by A52,RCOMP_1:def 1;
        then f.r1 in rng f by FUNCT_1:def 5;
        then f.r1 in R by A45;
        then f.r1 in P by A32;
        then consider q3 being Point of TOP-REAL 2 such that
        A78: q3=f.r1 & |.q3.|=1 by A1;
        A79: q3`2=h2.(f.r1) by A78,PSCOMP_1:def 29
          .=(h2*f).r1 by A77,FUNCT_1:23
          .=0 by A74,A75,FUNCT_1:72;
        then A80: 1=(q3`1)^2 +0^2 by A78,JGRAPH_3:10,SQUARE_1:59
        .=(q3`1)^2 by SQUARE_1:60;
          now per cases by A80,JGRAPH_3:2;
        case q3`1=1;
           then A81: q3=|[1,0]| by A79,EUCLID:57
           .=E-max(P) by A1,Th33;
             1 in dom f by A52,RCOMP_1:15;
         hence contradiction
                by A36,A42,A53,A74,A77,A78,A81,FUNCT_1:def 8;
        case q3`1=-1;
           then A82: q3=|[-1,0]| by A79,EUCLID:57
           .=W-min(P) by A1,Th32;
             0 in dom f by A52,RCOMP_1:15;
         hence contradiction
               by A36,A42,A49,A74,A77,A78,A82,FUNCT_1:def 8;
        end;
      hence contradiction;
      case t=r;
      hence contradiction by A49,A50,A54,A56;
      end;
      hence contradiction;
     end;
     hence for q being Point of TOP-REAL 2
       st q in Upper_Arc(P) holds q`2>=0;
   end;
  reconsider R=Lower_Arc(P) as non empty Subset of TOP-REAL 2;
  consider f2 being map of I[01], (TOP-REAL 2)|R such that
   A83: f2 is_homeomorphism & f2.0 =E-max(P) &
      f2.1 =W-min(P) by A4,TOPREAL1:def 2;
   A84: f2 is continuous by A83,TOPS_2:def 5;
   A85: f2 is one-to-one by A83,TOPS_2:def 5;
     for p being Point of (TOP-REAL 2) holds
    h2.p=proj2.p;
   then A86: h2 is continuous by Th35;
   A87: dom f2=the carrier of I[01] by FUNCT_2:def 1;
   A88: the carrier of ((TOP-REAL 2)|R)=R by JORDAN1:1;
   then A89: rng f2 c= the carrier of TOP-REAL 2 by XBOOLE_1:1;
     dom h2=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
   then A90: dom (h2*f2)=the carrier of I[01] by A87,A89,RELAT_1:46;
     rng (h2*f2) c= rng h2 by RELAT_1:45;
   then rng (h2*f2) c= the carrier of R^1 by XBOOLE_1:1;
   then h2*f2 is Function of the carrier of I[01],the carrier of R^1
                   by A90,FUNCT_2:4;
   then reconsider g1=h2*f2 as map of I[01],R^1;
  A91: (ex p being Point of TOP-REAL 2,t being Real st
  0<t & t<1 & f2.t=p & p`2>0) implies for q being Point of TOP-REAL 2
  st q in Lower_Arc(P) holds q`2>=0
   proof assume
      (ex p being Point of TOP-REAL 2,t being Real st
     0<t & t<1 & f2.t=p & p`2>0);
     then consider p being Point of TOP-REAL 2,t being Real such that
     A92: 0<t & t<1 & f2.t=p & p`2>0;
       now assume ex q being Point of TOP-REAL 2
       st q in Lower_Arc(P) & q`2<0;
       then consider q being Point of TOP-REAL 2 such that
       A93: q in Lower_Arc(P) & q`2<0;
         rng f2 =[#]((TOP-REAL 2)|R) by A83,TOPS_2:def 5
             .=R by PRE_TOPC:def 10;
       then consider x being set such that
       A94: x in dom f2 & q=f2.x by A93,FUNCT_1:def 5;
       A95: dom f2= [.0,1.]
                                  by BORSUK_1:83,FUNCT_2:def 1;
       then x in {r where r is Real: 0<=r & r<=1 } by A94,RCOMP_1:def 1;
       then consider r being Real such that
       A96: x=r & 0<=r & r<=1;
       A97: (h2*f2).r=h2.q by A94,A96,FUNCT_1:23
       .=q`2 by PSCOMP_1:def 29;
         t in {v where v is Real: 0<=v & v<=1 } by A92;
       then A98: t in [.0,1.] by RCOMP_1:def 1;
       then A99: (h2*f2).t=h2.p by A92,A95,FUNCT_1:23
       .=p`2 by PSCOMP_1:def 29;
        now per cases by REAL_1:def 5;
      case A100: r<t;
         [.r,t.] c= [.0,1.] by A94,A95,A96,A98,RCOMP_1:16;
       then reconsider B=[.r,t.] as non empty Subset of I[01]
                               by A100,BORSUK_1:83,RCOMP_1:15;
       reconsider B0=B as Subset of I[01];
       reconsider g=g1|B0 as map of (I[01]|B0),R^1 by JGRAPH_3:12;
         g1 is continuous by A84,A86,Th10;
       then A101: g is continuous by TOPMETR:10;
       A102: Closed-Interval-TSpace(r,t)=I[01]|B by A92,A96,A100,Th6,TOPMETR:27
;
          r in {r4 where r4 is Real: r<=r4 & r4<=t} by A100;
        then A103: r in B by RCOMP_1:def 1;
          t in {r4 where r4 is Real: r<=r4 & r4<=t} by A100;
        then t in B by RCOMP_1:def 1;
        then (q`2)*(p`2)<0 & q`2=g.r & p`2=g.t by A92,A93,A97,A99,A103,FUNCT_1:
72,SQUARE_1:24;
        then consider r1 being Real such that
        A104: g.r1=0 & r<r1 & r1<t by A100,A101,A102,TOPREAL5:14;
          r1 in {r4 where r4 is Real: r<=r4 & r4<=t} by A104;
        then A105: r1 in B by RCOMP_1:def 1;
        A106: 0<r1 by A96,A104;
          r1<1 by A92,A104,AXIOMS:22;
        then r1 in {r2 where r2 is Real: 0<=r2 & r2<=1} by A106;
        then A107: r1 in dom f2 by A95,RCOMP_1:def 1;
        then f2.r1 in rng f2 by FUNCT_1:def 5;
        then f2.r1 in R by A88;
        then f2.r1 in P by A33;
        then consider q3 being Point of TOP-REAL 2 such that
        A108: q3=f2.r1 & |.q3.|=1 by A1;
        A109: q3`2=h2.(f2.r1) by A108,PSCOMP_1:def 29
          .=(h2*f2).r1 by A107,FUNCT_1:23
          .=0 by A104,A105,FUNCT_1:72;
        then A110: 1=(q3`1)^2 +0^2 by A108,JGRAPH_3:10,SQUARE_1:59
        .=(q3`1)^2 by SQUARE_1:60;
          now per cases by A110,JGRAPH_3:2;
        case q3`1=1;
           then A111: q3=|[1,0]| by A109,EUCLID:57
           .=E-max(P) by A1,Th33;
             0 in dom f2 by A95,RCOMP_1:15;
         hence contradiction
               by A83,A85,A96,A104,A107,A108,A111,FUNCT_1:def 8;
        case q3`1=-1;
           then A112: q3=|[-1,0]| by A109,EUCLID:57
           .=W-min(P) by A1,Th32;
             1 in dom f2 by A95,RCOMP_1:15;
         hence contradiction
               by A83,A85,A92,A104,A107,A108,A112,FUNCT_1:def 8;
        end;
      hence contradiction;
      case A113: t<r;
         [.t,r.] c= [.0,1.] by A94,A95,A96,A98,RCOMP_1:16;
       then reconsider B=[.t,r.] as non empty Subset of I[01]
                               by A113,BORSUK_1:83,RCOMP_1:15;
       reconsider B0=B as Subset of I[01];
       reconsider g=g1|B0 as map of (I[01]|B0),R^1 by JGRAPH_3:12;
         g1 is continuous by A84,A86,Th10;
       then A114: g is continuous by TOPMETR:10;
       A115: Closed-Interval-TSpace(t,r)=I[01]|B by A92,A96,A113,Th6,TOPMETR:27
;
          r in {r4 where r4 is Real: t<=r4 & r4<=r} by A113;
        then A116: r in B by RCOMP_1:def 1;
          t in {r4 where r4 is Real: t<=r4 & r4<=r} by A113;
        then t in B by RCOMP_1:def 1;
        then (q`2)*(p`2)<0 & q`2=g.r & p`2=g.t by A92,A93,A97,A99,A116,FUNCT_1:
72,SQUARE_1:24;
        then consider r1 being Real such that
        A117: g.r1=0 & t<r1 & r1<r by A113,A114,A115,TOPREAL5:14;
          r1 in {r4 where r4 is Real: t<=r4 & r4<=r} by A117;
        then A118: r1 in B by RCOMP_1:def 1;
        A119: 0<r1 by A92,A117,AXIOMS:22;
          r1<1 by A96,A117,AXIOMS:22;
        then r1 in {r2 where r2 is Real: 0<=r2 & r2<=1} by A119;
        then A120: r1 in dom f2 by A95,RCOMP_1:def 1;
        then f2.r1 in rng f2 by FUNCT_1:def 5;
        then f2.r1 in R by A88;
        then f2.r1 in P by A33;
        then consider q3 being Point of TOP-REAL 2 such that
        A121: q3=f2.r1 & |.q3.|=1 by A1;
        A122: q3`2=h2.(f2.r1) by A121,PSCOMP_1:def 29
          .=g1.r1 by A120,FUNCT_1:23
          .=0 by A117,A118,FUNCT_1:72;
        then A123: 1=(q3`1)^2 +0^2 by A121,JGRAPH_3:10,SQUARE_1:59
        .=(q3`1)^2 by SQUARE_1:60;
          now per cases by A123,JGRAPH_3:2;
        case q3`1=1;
           then A124: q3=|[1,0]| by A122,EUCLID:57
           .=E-max(P) by A1,Th33;
             0 in dom f2 by A95,RCOMP_1:15;
         hence contradiction
                by A83,A85,A92,A117,A120,A121,A124,FUNCT_1:def 8;
        case q3`1=-1;
           then A125: q3=|[-1,0]| by A122,EUCLID:57
           .=W-min(P) by A1,Th32;
             1 in dom f2 by A95,RCOMP_1:15;
         hence contradiction
               by A83,A85,A96,A117,A120,A121,A125,FUNCT_1:def 8;
        end;
      hence contradiction;
      case t=r;
      hence contradiction by A92,A93,A97,A99;
      end;
      hence contradiction;
     end;
     hence for q being Point of TOP-REAL 2
       st q in Lower_Arc(P) holds q`2>=0;
   end;

   A126: Upper_Arc(P)
     c= {p where p is Point of TOP-REAL 2:p in P & p`2>=0}
    proof let x2 be set;assume
      A127: x2 in Upper_Arc(P);
      then reconsider q3=x2 as Point of TOP-REAL 2;
        q3`2>=0 by A6,A7,A8,A23,A30,A37,A38,A39,A40,A48,A127;
     hence x2 in {p where p is Point of TOP-REAL 2:p in P & p`2>=0}
                           by A32,A127;
    end;
     {p where p is Point of TOP-REAL 2:p in P & p`2>=0} c= Upper_Arc(P)
    proof let x be set;assume
        x in {p where p is Point of TOP-REAL 2:p in P & p`2>=0};
      then consider p being Point of TOP-REAL 2 such that
      A128: p=x & p in P & p`2>=0;
        now per cases by A128;
      case A129: p`2=0;
        consider p8 being Point of TOP-REAL 2 such that
        A130: p8=p & |.p8.|=1 by A1,A128;
        A131: p=|[p`1,p`2]| by EUCLID:57;
          1=sqrt((p`1)^2+(p`2)^2) by A130,JGRAPH_3:10
           .=abs(p`1) by A129,SQUARE_1:60,91;
        then (p`1)^2=1 by SQUARE_1:59,62;
        then p=|[1,0]| or p=|[-1,0]| by A129,A131,JGRAPH_3:2;
       hence x in Upper_Arc(P) by A1,A34,A35,A128,Th32,Th33;
      case A132: p`2>0;
         now assume not x in Upper_Arc(P);
         then A133: x in Lower_Arc(P) by A6,A128,XBOOLE_0:def 2;
           rng f2 =[#]((TOP-REAL 2)|R) by A83,TOPS_2:def 5
               .=R by PRE_TOPC:def 10;
         then consider x2 being set such that
         A134: x2 in dom f2 & p=f2.x2 by A128,A133,FUNCT_1:def 5;
           dom f2= [.0,1.]
                                  by BORSUK_1:83,FUNCT_2:def 1;
         then x2 in {r where r is Real: 0<=r & r<=1 } by A134,RCOMP_1:def 1;
         then consider t2 being Real such that
         A135: x2=t2 & 0<=t2 & t2<=1;
         A136: now assume t2=0;
           then p=|[1,0]| by A1,A83,A134,A135,Th33;
          hence contradiction by A132,EUCLID:56;
         end;
           now assume t2=1;
           then p=|[-1,0]| by A1,A83,A134,A135,Th32;
          hence contradiction by A132,EUCLID:56;
         end;
         then A137: 0<t2 & t2<1 & f2.t2=p & p`2>0 by A132,A134,A135,A136,REAL_1
:def 5;
         A138: (|[0,-1]|)`1=0 by EUCLID:56;
         A139: (|[0,-1]|)`2=-1 by EUCLID:56;
         then |.|[0,-1]|.|=sqrt((0)^2+(-1)^2) by A138,JGRAPH_3:10
           .=1 by SQUARE_1:59,60,61,83;
         then A140: |[0,-1]| in {q where q is Point of TOP-REAL 2: |.q.|=1};
           now per cases by A1,A6,A140,XBOOLE_0:def 2;
         case |[0,-1]| in Upper_Arc(P);
          hence contradiction by A6,A7,A8,A23,A30,A37,A38,A39,A40,A48,A139;
         case |[0,-1]| in Lower_Arc(P);
          hence contradiction by A91,A137,A139;
         end;
        hence contradiction;
       end;
       hence x in Upper_Arc(P);
      end;
     hence x in Upper_Arc(P);
    end;
 hence thesis by A126,XBOOLE_0:def 10;
end;

theorem Th38: for P being compact non empty Subset of TOP-REAL 2 st
P={q where q is Point of TOP-REAL 2: |.q.|=1} holds
Lower_Arc(P)={p where p is Point of TOP-REAL 2:p in P & p`2<=0}
proof let P be compact non empty Subset of TOP-REAL 2;
 assume A1: P={q where q is Point of TOP-REAL 2: |.q.|=1};
  then A2: P is_simple_closed_curve by JGRAPH_3:36;
  then A3: Upper_Arc(P) is_an_arc_of W-min(P),E-max(P) by JORDAN6:def 8;
  A4: Lower_Arc(P) is_an_arc_of E-max(P),W-min(P) by A2,JORDAN6:def 9;
   set P4=Lower_Arc(P);
   A5: P4 is_an_arc_of E-max(P),W-min(P)
     & Upper_Arc(P) /\ P4={W-min(P),E-max(P)}
       & Upper_Arc(P) \/ P4=P
         & First_Point(Upper_Arc(P),W-min(P),E-max(P),
         Vertical_Line((W-bound(P)+E-bound(P))/2))`2>
              Last_Point(P4,E-max(P),W-min(P),
         Vertical_Line((W-bound(P)+E-bound(P))/2))`2 by A2,JORDAN6:def 9;
   A6: W-bound(P)=-1 by A1,Th31;
   A7: E-bound(P)=1 by A1,Th31;
   A8: Vertical_Line(0)={p where p is Point of TOP-REAL 2: p`1=0}
                         by JORDAN6:def 6;
     reconsider P1=Lower_Arc(P) as Subset of TOP-REAL 2;
     reconsider P2=Upper_Arc(P) as Subset of TOP-REAL 2;
     reconsider Q=Vertical_Line(0) as Subset of TOP-REAL 2;
     set p11=W-min(P);
     set p22=E-max(P);
     set pj= First_Point(Upper_Arc(P),W-min(P),E-max(P),
         Vertical_Line(0));
     set p8= Last_Point(Lower_Arc(P),E-max(P),W-min(P),
         Vertical_Line(0));
        A9: W-bound P=-1 by A1,Th31;
        A10: E-bound P=1 by A1,Th31;
        A11: S-bound P=-1 by A1,Th31;
        A12: N-bound P=1 by A1,Th31;
        then A13: LSeg(|[0,-1]|,|[0,1]|) meets Lower_Arc P
                     by A2,A9,A10,A11,JORDAN1B:27;
         A14: LSeg(|[0,-1]|,|[0,1]|) c= Q
         proof let x be set;assume x in LSeg(|[0,-1]|,|[0,1]|);
           then x in {(1-l)*(|[0,-1]|) +l*(|[0,1]|) where l is Real: 0<=l & l
<=1}
                                     by TOPREAL1:def 4;
           then consider l being Real such that
           A15: x=(1-l)*(|[0,-1]|) +l*(|[0,1]|) & 0<=l & l<=1;
             ((1-l)*(|[0,-1]|) +l*(|[0,1]|))`1
             = ((1-l)*(|[0,-1]|))`1 +(l*(|[0,1]|))`1 by TOPREAL3:7
             .=(1-l)*(|[0,-1]|)`1+(l*(|[0,1]|))`1 by TOPREAL3:9
             .=(1-l)*(|[0,-1]|)`1+l*((|[0,1]|))`1 by TOPREAL3:9
             .=(1-l)*0+l*((|[0,1]|))`1 by EUCLID:56
             .=(1-l)*0+l*0 by EUCLID:56
             .=0;
          hence x in Q by A8,A15;
         end;
        then A16: P1 meets Q by A13,XBOOLE_1:64;
        A17: Lower_Arc(P) is closed by A4,JORDAN6:12;
          Vertical_Line(0) is closed by JORDAN6:33;
        then P1 /\ Q is closed by A17,TOPS_1:35;
        then A18: p8 in P1 /\ Q &
       for g being map of I[01], (TOP-REAL 2)|P1, s2 being Real st
           g is_homeomorphism & g.0 = p22 & g.1 = p11
           & g.s2 = p8 & 0 <= s2 & s2 <= 1 holds
       (for t being Real st 1 >= t & t > s2 holds not g.t in Q)
                       by A4,A16,JORDAN5C:def 2;
         P1 /\ Q c= {|[0,-1]|,|[0,1]|}
        proof let x be set;assume x in P1 /\ Q;
          then A19: x in P1 & x in Q by XBOOLE_0:def 3;
          then consider p being Point of TOP-REAL 2 such that
          A20: p=x & p`1=0 by A8;
            x in P by A5,A19,XBOOLE_0:def 2;
          then consider q being Point of TOP-REAL 2 such that
          A21: q=x & |.q.|=1 by A1;
             0+(q`2)^2 =1 by A20,A21,JGRAPH_3:10,SQUARE_1:59,60;
          then q`2=1 or q`2=-1 by JGRAPH_3:2;
          then x=|[0,-1]| or x=|[0,1]| by A20,A21,EUCLID:57;
         hence x in {|[0,-1]|,|[0,1]|} by TARSKI:def 2;
        end;
       then p8=|[0,-1]| or p8=|[0,1]| by A18,TARSKI:def 2;
       then A22: p8`2=-1 or p8`2=1 by EUCLID:56;
          LSeg(|[0,-1]|,|[0,1]|) meets Upper_Arc P
                     by A2,A9,A10,A11,A12,JORDAN1B:26;
        then A23: P2 meets Q by A14,XBOOLE_1:64;
        A24: Upper_Arc(P) is closed by A3,JORDAN6:12;
          Vertical_Line(0) is closed by JORDAN6:33;
        then P2 /\ Q is closed by A24,TOPS_1:35;
        then A25: pj in P2 /\ Q &
       for g being map of I[01], (TOP-REAL 2)|P2, s2 being Real st
           g is_homeomorphism & g.0 = p11 & g.1 = p22
           & g.s2 = pj & 0 <= s2 & s2 <= 1 holds
       for t being Real st 0 <= t & t < s2 holds not g.t in Q
                       by A3,A23,JORDAN5C:def 1;
         P2 /\ Q c= {|[0,-1]|,|[0,1]|}
        proof let x be set;assume x in P2 /\ Q;
          then A26: x in P2 & x in Q by XBOOLE_0:def 3;
          then consider p being Point of TOP-REAL 2 such that
          A27: p=x & p`1=0 by A8;
            x in P by A5,A26,XBOOLE_0:def 2;
          then consider q being Point of TOP-REAL 2 such that
          A28: q=x & |.q.|=1 by A1;
             0+(q`2)^2 =1 by A27,A28,JGRAPH_3:10,SQUARE_1:59,60;
          then q`2=1 or q`2=-1 by JGRAPH_3:2;
          then x=|[0,-1]| or x=|[0,1]| by A27,A28,EUCLID:57;
         hence x in {|[0,-1]|,|[0,1]|} by TARSKI:def 2;
        end;
       then pj=|[0,-1]| or pj=|[0,1]| by A25,TARSKI:def 2;
       then A29: pj`2=-1 or pj`2=1 by EUCLID:56;
  A30: p8 in P1 by A18,XBOOLE_0:def 3;
  A31: Lower_Arc(P) c= P by A5,XBOOLE_1:7;
  A32: Upper_Arc(P) c= P by A5,XBOOLE_1:7;
    E-max(P) in {W-min(P),E-max(P)} by TARSKI:def 2;
  then A33: E-max(P) in Lower_Arc(P) by A5,XBOOLE_0:def 3;
    W-min(P) in {W-min(P),E-max(P)} by TARSKI:def 2;
  then A34: W-min(P) in Lower_Arc(P) by A5,XBOOLE_0:def 3;
  reconsider R=Lower_Arc(P)
            as non empty Subset of TOP-REAL 2;
  consider f being map of I[01], (TOP-REAL 2)|R such that
   A35: f is_homeomorphism & f.0 =E-max(P) &
      f.1 =W-min(P) by A4,TOPREAL1:def 2;
         rng f =[#]((TOP-REAL 2)|R) by A35,TOPS_2:def 5
            .=R by PRE_TOPC:def 10;
       then consider x8 being set such that
       A36: x8 in dom f & p8=f.x8 by A30,FUNCT_1:def 5;
         dom f= [.0,1.]
                                  by BORSUK_1:83,FUNCT_2:def 1;
       then x8 in {r where r is Real: 0<=r & r<=1 } by A36,RCOMP_1:def 1;
       then consider r8 being Real such that
       A37: x8=r8 & 0<=r8 & r8<=1;
        A38: now assume r8=0;
         then p8=|[1,0]| by A1,A35,A36,A37,Th33;
        hence contradiction by A22,EUCLID:56;
       end;
         now assume r8=1;
         then p8=|[-1,0]| by A1,A35,A36,A37,Th32;
        hence contradiction by A22,EUCLID:56;
       end;
       then A39: 1>r8 by A37,REAL_1:def 5;
   reconsider h2=proj2 as map of TOP-REAL 2,R^1 by JGRAPH_2:17;
   A40: f is continuous by A35,TOPS_2:def 5;
   A41: f is one-to-one by A35,TOPS_2:def 5;
     for p being Point of (TOP-REAL 2) holds
    h2.p=proj2.p;
   then A42: h2 is continuous by Th35;
   A43: dom f=the carrier of I[01] by FUNCT_2:def 1;
   A44: the carrier of ((TOP-REAL 2)|R)=R by JORDAN1:1;
   then A45: rng f c= the carrier of TOP-REAL 2 by XBOOLE_1:1;
     dom h2=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
   then A46: dom (h2*f)=the carrier of I[01] by A43,A45,RELAT_1:46;
     rng (h2*f) c= rng h2 by RELAT_1:45;
   then rng (h2*f) c= the carrier of R^1 by XBOOLE_1:1;
   then h2*f is Function of the carrier of I[01],the carrier of R^1
                   by A46,FUNCT_2:4;
   then reconsider g0=h2*f as map of I[01],R^1;
  A47: (ex p being Point of TOP-REAL 2,t being Real st
  0<t & t<1 & f.t=p & p`2<0) implies for q being Point of TOP-REAL 2
  st q in Lower_Arc(P) holds q`2<=0
   proof assume
      (ex p being Point of TOP-REAL 2,t being Real st
     0<t & t<1 & f.t=p & p`2<0);
     then consider p being Point of TOP-REAL 2,t being Real such that
     A48: 0<t & t<1 & f.t=p & p`2<0;
       now assume ex q being Point of TOP-REAL 2
       st q in Lower_Arc(P) & q`2>0;
       then consider q being Point of TOP-REAL 2 such that
       A49: q in Lower_Arc(P) & q`2>0;
         rng f =[#]((TOP-REAL 2)|R) by A35,TOPS_2:def 5
            .=R by PRE_TOPC:def 10;
       then consider x being set such that
       A50: x in dom f & q=f.x by A49,FUNCT_1:def 5;
       A51: dom f= [.0,1.]
                                  by BORSUK_1:83,FUNCT_2:def 1;
       then x in {r where r is Real: 0<=r & r<=1 } by A50,RCOMP_1:def 1;
       then consider r being Real such that
       A52: x=r & 0<=r & r<=1;
       A53: (h2*f).r=h2.q by A50,A52,FUNCT_1:23
       .=q`2 by PSCOMP_1:def 29;
         t in {v where v is Real: 0<=v & v<=1 } by A48;
       then A54: t in [.0,1.] by RCOMP_1:def 1;
       then A55: (h2*f).t=h2.p by A48,A51,FUNCT_1:23
       .=p`2 by PSCOMP_1:def 29;
        now per cases by REAL_1:def 5;
      case A56: r<t;
         [.r,t.] c= [.0,1.] by A50,A51,A52,A54,RCOMP_1:16;
       then reconsider B=[.r,t.] as non empty Subset of I[01]
                               by A56,BORSUK_1:83,RCOMP_1:15;
       reconsider B0=B as Subset of I[01];
       reconsider g=g0|B0 as map of (I[01]|B0),R^1 by JGRAPH_3:12;
         g0 is continuous by A40,A42,Th10;
       then A57: g is continuous by TOPMETR:10;
       A58: Closed-Interval-TSpace(r,t)=I[01]|B by A48,A52,A56,Th6,TOPMETR:27;
          r in {r4 where r4 is Real: r<=r4 & r4<=t} by A56;
        then A59: r in B by RCOMP_1:def 1;
          t in {r4 where r4 is Real: r<=r4 & r4<=t} by A56;
        then t in B by RCOMP_1:def 1;
        then (q`2)*(p`2)<0 & q`2=g.r & p`2=g.t by A48,A49,A53,A55,A59,FUNCT_1:
72,SQUARE_1:24;
        then consider r1 being Real such that
        A60: g.r1=0 & r<r1 & r1<t by A56,A57,A58,TOPREAL5:14;
          r1 in {r4 where r4 is Real: r<=r4 & r4<=t} by A60;
        then A61: r1 in B by RCOMP_1:def 1;
        A62: 0<r1 by A52,A60;
          r1<1 by A48,A60,AXIOMS:22;
        then r1 in {r2 where r2 is Real: 0<=r2 & r2<=1} by A62;
        then A63: r1 in dom f by A51,RCOMP_1:def 1;
        then f.r1 in rng f by FUNCT_1:def 5;
        then f.r1 in R by A44;
        then f.r1 in P by A31;
        then consider q3 being Point of TOP-REAL 2 such that
        A64: q3=f.r1 & |.q3.|=1 by A1;
        A65: q3`2=h2.(f.r1) by A64,PSCOMP_1:def 29
          .=(h2*f).r1 by A63,FUNCT_1:23
          .=0 by A60,A61,FUNCT_1:72;
        then A66: 1=(q3`1)^2 +0^2 by A64,JGRAPH_3:10,SQUARE_1:59
        .=(q3`1)^2 by SQUARE_1:60;
          now per cases by A66,JGRAPH_3:2;
        case q3`1=1;
           then A67: q3=|[1,0]| by A65,EUCLID:57
           .=E-max(P) by A1,Th33;
             0 in dom f by A51,RCOMP_1:15;
         hence contradiction
               by A35,A41,A52,A60,A63,A64,A67,FUNCT_1:def 8;
        case q3`1=-1;
           then A68: q3=|[-1,0]| by A65,EUCLID:57
           .=W-min(P) by A1,Th32;
             1 in dom f by A51,RCOMP_1:15;
         hence contradiction
               by A35,A41,A48,A60,A63,A64,A68,FUNCT_1:def 8;
        end;
      hence contradiction;
      case A69: t<r;
         [.t,r.] c= [.0,1.] by A50,A51,A52,A54,RCOMP_1:16;
       then reconsider B=[.t,r.] as non empty Subset of I[01]
                               by A69,BORSUK_1:83,RCOMP_1:15;
       reconsider B0=B as Subset of I[01];
       reconsider g=g0|B0 as map of (I[01]|B0),R^1 by JGRAPH_3:12;
         g0 is continuous by A40,A42,Th10;
       then A70: g is continuous by TOPMETR:10;
       A71: Closed-Interval-TSpace(t,r)=I[01]|B by A48,A52,A69,Th6,TOPMETR:27;
          r in {r4 where r4 is Real: t<=r4 & r4<=r} by A69;
        then A72: r in B by RCOMP_1:def 1;
          t in {r4 where r4 is Real: t<=r4 & r4<=r} by A69;
        then t in B by RCOMP_1:def 1;
        then (q`2)*(p`2)<0 & q`2=g.r & p`2=g.t by A48,A49,A53,A55,A72,FUNCT_1:
72,SQUARE_1:24;
        then consider r1 being Real such that
        A73: g.r1=0 & t<r1 & r1<r by A69,A70,A71,TOPREAL5:14;
          r1 in {r4 where r4 is Real: t<=r4 & r4<=r} by A73;
        then A74: r1 in B by RCOMP_1:def 1;
        A75: 0<r1 by A48,A73,AXIOMS:22;
          r1<1 by A52,A73,AXIOMS:22;
        then r1 in {r2 where r2 is Real: 0<=r2 & r2<=1} by A75;
        then A76: r1 in dom f by A51,RCOMP_1:def 1;
        then f.r1 in rng f by FUNCT_1:def 5;
        then f.r1 in R by A44;
        then f.r1 in P by A31;
        then consider q3 being Point of TOP-REAL 2 such that
        A77: q3=f.r1 & |.q3.|=1 by A1;
        A78: q3`2=h2.(f.r1) by A77,PSCOMP_1:def 29
          .=(h2*f).r1 by A76,FUNCT_1:23
          .=0 by A73,A74,FUNCT_1:72;
        then A79: 1=(q3`1)^2 +0^2 by A77,JGRAPH_3:10,SQUARE_1:59
        .=(q3`1)^2 by SQUARE_1:60;
          now per cases by A79,JGRAPH_3:2;
        case q3`1=1;
           then A80: q3=|[1,0]| by A78,EUCLID:57
           .=E-max(P) by A1,Th33;
             0 in dom f by A51,RCOMP_1:15;
         hence contradiction
             by A35,A41,A48,A73,A76,A77,A80,FUNCT_1:def 8;
        case q3`1=-1;
           then A81: q3=|[-1,0]| by A78,EUCLID:57
           .=W-min(P) by A1,Th32;
             1 in dom f by A51,RCOMP_1:15;
         hence contradiction
              by A35,A41,A52,A73,A76,A77,A81,FUNCT_1:def 8;
        end;
      hence contradiction;
      case t=r;
      hence contradiction by A48,A49,A53,A55;
      end;
      hence contradiction;
     end;
     hence for q being Point of TOP-REAL 2
       st q in Lower_Arc(P) holds q`2<=0;
   end;
  reconsider R=Upper_Arc(P) as non empty Subset of TOP-REAL 2;
  consider f2 being map of I[01], (TOP-REAL 2)|R such that
   A82: f2 is_homeomorphism & f2.0 =W-min(P) &
      f2.1 =E-max(P) by A3,TOPREAL1:def 2;
   A83: f2 is continuous by A82,TOPS_2:def 5;
   A84: f2 is one-to-one by A82,TOPS_2:def 5;
     for p being Point of (TOP-REAL 2) holds h2.p=proj2.p;
   then A85: h2 is continuous by Th35;
   A86: dom f2=the carrier of I[01] by FUNCT_2:def 1;
   A87: the carrier of ((TOP-REAL 2)|R)=R by JORDAN1:1;
   then A88: rng f2 c= the carrier of TOP-REAL 2 by XBOOLE_1:1;
     dom h2=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
   then A89: dom (h2*f2)=the carrier of I[01] by A86,A88,RELAT_1:46;
     rng (h2*f2) c= rng h2 by RELAT_1:45;
   then rng (h2*f2) c= the carrier of R^1 by XBOOLE_1:1;
   then h2*f2 is Function of the carrier of I[01],the carrier of R^1
                   by A89,FUNCT_2:4;
   then reconsider g1=h2*f2 as map of I[01],R^1;
  A90: (ex p being Point of TOP-REAL 2,t being Real st
  0<t & t<1 & f2.t=p & p`2<0) implies for q being Point of TOP-REAL 2
  st q in Upper_Arc(P) holds q`2<=0
   proof assume
      (ex p being Point of TOP-REAL 2,t being Real st
     0<t & t<1 & f2.t=p & p`2<0);
     then consider p being Point of TOP-REAL 2,t being Real such that
     A91: 0<t & t<1 & f2.t=p & p`2<0;
       now assume ex q being Point of TOP-REAL 2
       st q in Upper_Arc(P) & q`2>0;
       then consider q being Point of TOP-REAL 2 such that
       A92: q in Upper_Arc(P) & q`2>0;
         rng f2 =[#]((TOP-REAL 2)|R) by A82,TOPS_2:def 5
             .=R by PRE_TOPC:def 10;
       then consider x being set such that
       A93: x in dom f2 & q=f2.x by A92,FUNCT_1:def 5;
       A94: dom f2= [.0,1.]
                                  by BORSUK_1:83,FUNCT_2:def 1;
       then x in {r where r is Real: 0<=r & r<=1 } by A93,RCOMP_1:def 1;
       then consider r being Real such that
       A95: x=r & 0<=r & r<=1;
       A96: (h2*f2).r=h2.q by A93,A95,FUNCT_1:23
       .=q`2 by PSCOMP_1:def 29;
         t in {v where v is Real: 0<=v & v<=1 } by A91;
       then A97: t in [.0,1.] by RCOMP_1:def 1;
       then A98: (h2*f2).t=h2.p by A91,A94,FUNCT_1:23
       .=p`2 by PSCOMP_1:def 29;
        now per cases by REAL_1:def 5;
      case A99: r<t;
         [.r,t.] c= [.0,1.] by A93,A94,A95,A97,RCOMP_1:16;
       then reconsider B=[.r,t.] as non empty Subset of I[01]
                               by A99,BORSUK_1:83,RCOMP_1:15;
       reconsider B0=B as Subset of I[01];
       reconsider g=g1|B0 as map of (I[01]|B0),R^1 by JGRAPH_3:12;
         g1 is continuous by A83,A85,Th10;
       then A100: g is continuous by TOPMETR:10;
       A101: Closed-Interval-TSpace(r,t)=I[01]|B by A91,A95,A99,Th6,TOPMETR:27;
          r in {r4 where r4 is Real: r<=r4 & r4<=t} by A99;
        then A102: r in B by RCOMP_1:def 1;
          t in {r4 where r4 is Real: r<=r4 & r4<=t} by A99;
        then t in B by RCOMP_1:def 1;
        then (q`2)*(p`2)<0 & q`2=g.r & p`2=g.t by A91,A92,A96,A98,A102,FUNCT_1:
72,SQUARE_1:24;
        then consider r1 being Real such that
        A103: g.r1=0 & r<r1 & r1<t by A99,A100,A101,TOPREAL5:14;
          r1 in {r4 where r4 is Real: r<=r4 & r4<=t} by A103;
        then A104: r1 in B by RCOMP_1:def 1;
        A105: 0<r1 by A95,A103;
          r1<1 by A91,A103,AXIOMS:22;
        then r1 in {r2 where r2 is Real: 0<=r2 & r2<=1} by A105;
        then A106: r1 in dom f2 by A94,RCOMP_1:def 1;
        then f2.r1 in rng f2 by FUNCT_1:def 5;
        then f2.r1 in R by A87;
        then f2.r1 in P by A32;
        then consider q3 being Point of TOP-REAL 2 such that
        A107: q3=f2.r1 & |.q3.|=1 by A1;
        A108: q3`2=h2.(f2.r1) by A107,PSCOMP_1:def 29
          .=(h2*f2).r1 by A106,FUNCT_1:23
          .=0 by A103,A104,FUNCT_1:72;
        then A109: 1=(q3`1)^2 +0^2 by A107,JGRAPH_3:10,SQUARE_1:59
        .=(q3`1)^2 by SQUARE_1:60;
          now per cases by A109,JGRAPH_3:2;
        case q3`1=1;
           then A110: q3=|[1,0]| by A108,EUCLID:57
           .=E-max(P) by A1,Th33;
             1 in dom f2 by A94,RCOMP_1:15;
         hence contradiction
               by A82,A84,A91,A103,A106,A107,A110,FUNCT_1:def 8;
        case q3`1=-1;
           then A111: q3=|[-1,0]| by A108,EUCLID:57
           .=W-min(P) by A1,Th32;
             0 in dom f2 by A94,RCOMP_1:15;
         hence contradiction
               by A82,A84,A95,A103,A106,A107,A111,FUNCT_1:def 8;
        end;
      hence contradiction;
      case A112: t<r;
         [.t,r.] c= [.0,1.] by A93,A94,A95,A97,RCOMP_1:16;
       then reconsider B=[.t,r.] as non empty Subset of I[01]
                               by A112,BORSUK_1:83,RCOMP_1:15;
       reconsider B0=B as Subset of I[01];
       reconsider g=g1|B0 as map of (I[01]|B0),R^1 by JGRAPH_3:12;
         g1 is continuous by A83,A85,Th10;
       then A113: g is continuous by TOPMETR:10;
       A114: Closed-Interval-TSpace(t,r)=I[01]|B by A91,A95,A112,Th6,TOPMETR:27
;
          r in {r4 where r4 is Real: t<=r4 & r4<=r} by A112;
        then A115: r in B by RCOMP_1:def 1;
          t in {r4 where r4 is Real: t<=r4 & r4<=r} by A112;
        then t in B by RCOMP_1:def 1;
        then (q`2)*(p`2)<0 & q`2=g.r & p`2=g.t by A91,A92,A96,A98,A115,FUNCT_1:
72,SQUARE_1:24;
        then consider r1 being Real such that
        A116: g.r1=0 & t<r1 & r1<r by A112,A113,A114,TOPREAL5:14;
          r1 in {r4 where r4 is Real: t<=r4 & r4<=r} by A116;
        then A117: r1 in B by RCOMP_1:def 1;
        A118: 0<r1 by A91,A116,AXIOMS:22;
          r1<1 by A95,A116,AXIOMS:22;
        then r1 in {r2 where r2 is Real: 0<=r2 & r2<=1} by A118;
        then A119: r1 in dom f2 by A94,RCOMP_1:def 1;
        then f2.r1 in rng f2 by FUNCT_1:def 5;
        then f2.r1 in R by A87;
        then f2.r1 in P by A32;
        then consider q3 being Point of TOP-REAL 2 such that
        A120: q3=f2.r1 & |.q3.|=1 by A1;
        A121: q3`2=h2.(f2.r1) by A120,PSCOMP_1:def 29
          .=(h2*f2).r1 by A119,FUNCT_1:23
          .=0 by A116,A117,FUNCT_1:72;
        then A122: 1=(q3`1)^2 +0^2 by A120,JGRAPH_3:10,SQUARE_1:59
        .=(q3`1)^2 by SQUARE_1:60;
          now per cases by A122,JGRAPH_3:2;
        case q3`1=1;
           then A123: q3=|[1,0]| by A121,EUCLID:57
           .=E-max(P) by A1,Th33;
             1 in dom f2 by A94,RCOMP_1:15;
         hence contradiction
               by A82,A84,A95,A116,A119,A120,A123,FUNCT_1:def 8;
        case q3`1=-1;
           then A124: q3=|[-1,0]| by A121,EUCLID:57
           .=W-min(P) by A1,Th32;
             0 in dom f2 by A94,RCOMP_1:15;
         hence contradiction
                by A82,A84,A91,A116,A119,A120,A124,FUNCT_1:def 8;
        end;
      hence contradiction;
      case t=r;
      hence contradiction by A91,A92,A96,A98;
      end;
      hence contradiction;
     end;
     hence for q being Point of TOP-REAL 2
       st q in Upper_Arc(P) holds q`2<=0;
   end;
   A125: Lower_Arc(P)
     c= {p where p is Point of TOP-REAL 2:p in P & p`2<=0}
    proof let x2 be set;assume
      A126: x2 in Lower_Arc(P);
      then reconsider q3=x2 as Point of TOP-REAL 2;
        q3`2<=0 by A5,A6,A7,A22,A29,A36,A37,A38,A39,A47,A126;
     hence x2 in {p where p is Point of TOP-REAL 2:p in P & p`2<=0}
                           by A31,A126;
    end;
     {p where p is Point of TOP-REAL 2:p in P & p`2<=0} c= Lower_Arc(P)
    proof let x be set;assume
        x in {p where p is Point of TOP-REAL 2:p in P & p`2<=0};
      then consider p being Point of TOP-REAL 2 such that
      A127: p=x & p in P & p`2<=0;
        now per cases by A127;
      case A128: p`2=0;
        consider p8 being Point of TOP-REAL 2 such that
        A129: p8=p & |.p8.|=1 by A1,A127;
        A130: p=|[p`1,p`2]| by EUCLID:57;
          1=sqrt((p`1)^2+(p`2)^2) by A129,JGRAPH_3:10
           .=abs(p`1) by A128,SQUARE_1:60,91;
        then (p`1)^2=1 by SQUARE_1:59,62;
        then p=|[1,0]| or p=|[-1,0]| by A128,A130,JGRAPH_3:2;
       hence x in Lower_Arc(P) by A1,A33,A34,A127,Th32,Th33;
      case A131: p`2<0;
         now assume not x in Lower_Arc(P);
         then A132: x in Upper_Arc(P) by A5,A127,XBOOLE_0:def 2;
           rng f2 =[#]((TOP-REAL 2)|R) by A82,TOPS_2:def 5
               .=R by PRE_TOPC:def 10;
         then consider x2 being set such that
         A133: x2 in dom f2 & p=f2.x2 by A127,A132,FUNCT_1:def 5;
           dom f2= [.0,1.]
                                  by BORSUK_1:83,FUNCT_2:def 1;
         then x2 in {r where r is Real: 0<=r & r<=1 } by A133,RCOMP_1:def 1;
         then consider t2 being Real such that
         A134: x2=t2 & 0<=t2 & t2<=1;
         A135: now assume t2=1;
           then p=|[1,0]| by A1,A82,A133,A134,Th33;
          hence contradiction by A131,EUCLID:56;
         end;
           now assume t2=0;
           then p=|[-1,0]| by A1,A82,A133,A134,Th32;
          hence contradiction by A131,EUCLID:56;
         end;
         then A136: 0<t2 & t2<1 & f2.t2=p & p`2<0 by A131,A133,A134,A135,REAL_1
:def 5;
         A137: (|[0,1]|)`1=0 by EUCLID:56;
         A138: (|[0,1]|)`2=1 by EUCLID:56;
         then |.|[0,1]|.|=sqrt((0)^2+(1)^2) by A137,JGRAPH_3:10
           .=1 by SQUARE_1:59,60,83;
         then A139: |[0,1]| in {q where q is Point of TOP-REAL 2: |.q.|=1};
           now per cases by A1,A5,A139,XBOOLE_0:def 2;
         case |[0,1]| in Lower_Arc(P);
          hence contradiction by A5,A6,A7,A22,A29,A36,A37,A38,A39,A47,A138;
         case |[0,1]| in Upper_Arc(P);
          hence contradiction by A90,A136,A138;
         end;
        hence contradiction;
       end;
       hence x in Lower_Arc(P);
      end;
     hence x in Lower_Arc(P);
    end;
 hence thesis by A125,XBOOLE_0:def 10;
end;

theorem Th39: for a,b,d,e being Real st a<=b & e>0
  ex f being map of
     Closed-Interval-TSpace(a,b),Closed-Interval-TSpace(e*a+d,e*b+d)
  st f is_homeomorphism & for r being Real st r in [.a,b.] holds
  f.r=e*r+d
proof let a,b,d,e be Real;assume A1: a<=b & e>0;
     then e*a<=e*b by AXIOMS:25;
  then A2: e*a+d<=e*b+d by REAL_1:55;
  set X=the carrier of Closed-Interval-TSpace(a,b);
  set Y=the carrier of Closed-Interval-TSpace(e*a+d,e*b+d);
  defpred P[set,set] means (for r being Real st $1=r holds $2=e*r+d);
  A3: X=[.a,b.] by A1,TOPMETR:25;
  A4: Y=[.e*a+d,e*b+d.] by A2,TOPMETR:25;
  A5: for x being set st x in X ex y being set st y in Y & P[x,y]
   proof let x be set;assume A6: x in X;
     then reconsider r1=x as Real by A3;
     A7: a<=r1 & r1<=b by A3,A6,TOPREAL5:1;
     then A8: e*a<=e*r1 by A1,AXIOMS:25;
     A9: e*r1<=e*b by A1,A7,AXIOMS:25;
     set y1=e*r1+d;
     A10: for r being Real st x=r holds y1=e*r+d;
       e*a+d<=y1 & y1 <=e*b+d by A8,A9,REAL_1:55;
     then y1 in Y by A4,TOPREAL5:1;
    hence ex y being set st y in Y & P[x,y] by A10;
   end;
     ex f being Function of X,Y st for x being set st x in X holds P[x,f.x]
                    from FuncEx1(A5);
   then consider f1 being Function of X,Y such that
   A11: for x being set st x in X holds P[x,f1.x];
   reconsider f2=f1 as map of
     Closed-Interval-TSpace(a,b),Closed-Interval-TSpace(e*a+d,e*b+d);
   A12: for r being Real st r in [.a,b.] holds f2.r=e*r+d by A3,A11;
    set S=Closed-Interval-TSpace(a,b);
    set T=Closed-Interval-TSpace(e*a+d,e*b+d);
    A13: dom f2=the carrier of S by FUNCT_2:def 1;
    then A14: dom f2=[#]S by PRE_TOPC:12;
      for x1,x2 being set st x1 in dom f2 & x2 in dom f2 & f2.x1=f2.x2
        holds x1=x2
     proof let x1,x2 be set;assume
       A15: x1 in dom f2 & x2 in dom f2 & f2.x1=f2.x2;
       then reconsider r1=x1 as Real by A3,A13;
       reconsider r2=x2 as Real by A3,A13,A15;
         f2.x1=e*r1+d by A11,A13,A15;
       then e*r1+d-d=e*r2+d-d by A11,A13,A15 .=e*r2 by XCMPLX_1:26;
       then e*r1=e*r2 by XCMPLX_1:26;
       then r1*e/e=r2 by A1,XCMPLX_1:90;
      hence x1=x2 by A1,XCMPLX_1:90;
     end;
    then A16: f2 is one-to-one by FUNCT_1:def 8;
      rng f2 c= the carrier of T;
    then A17: rng f2 c= [#]T by PRE_TOPC:12;
      [#]T c= rng f2
     proof let y be set;assume A18: y in [#]T;
       then y in [.e*a+d,e*b+d.] by A4;
       then reconsider ry=y as Real;
       A19: e*a+d <= ry & ry<=e*b+d by A4,A18,TOPREAL5:1;
       then e*a+d-d<=ry-d by REAL_1:49;
       then e*a<=ry-d by XCMPLX_1:26;
       then a*e/e<=(ry-d)/e by A1,REAL_1:73;
       then A20: a<=(ry-d)/e by A1,XCMPLX_1:90;
         e*b+d-d>=ry-d by A19,REAL_1:49;
       then e*b>=ry-d by XCMPLX_1:26;
       then b*e/e>=(ry-d)/e by A1,REAL_1:73;
       then b>=(ry-d)/e by A1,XCMPLX_1:90;
       then A21: (ry-d)/e in [.a,b.] by A20,TOPREAL5:1;
       then f2.((ry-d)/e)=e*((ry-d)/e)+d by A3,A11
       .=ry-d+d by A1,XCMPLX_1:88.=ry by XCMPLX_1:27;
      hence y in rng f2 by A3,A13,A21,FUNCT_1:12;
     end;
    then A22: rng f2 = [#]T by A17,XBOOLE_0:def 10;
    then rng f2=Y by PRE_TOPC:12;
    then f1 is Function of the carrier of S,the carrier of R^1
                              by A4,A13,FUNCT_2:4,TOPMETR:24;
    then reconsider f3=f1 as map of S,R^1;
    defpred P1[set,set] means for r being Real st r=$1 holds $2=e*r+d;
    A23: for x being set st x in the carrier of R^1
         ex y being set st y in the carrier of R^1 & P1[x,y]
      proof let x be set;assume x in the carrier of R^1;
        then reconsider rx=x as Real by TOPMETR:24;
        reconsider ry=e*rx+d as Real;
          for r being Real st r=x holds ry=e*r+d;
       hence ex y being set st y in the carrier of R^1 & P1[x,y] by TOPMETR:24
;
      end;
       ex f4 being Function of the carrier of R^1,the carrier of R^1 st
        for x being set st x in the carrier of R^1 holds P1[x,f4.x]
            from FuncEx1(A23);
     then consider f4 being Function of the carrier of R^1,the carrier of R^1
      such that
      A24: for x being set st x in the carrier of R^1 holds P1[x,f4.x];
     reconsider f5=f4 as map of R^1,R^1;
       for x being Real holds f5.x = e*x + d by A24,TOPMETR:24;
     then A25: f5 is continuous by TOPMETR:28;
     reconsider B=the carrier of S as Subset of R^1
          by A3,TOPMETR:24;
     A26: R^1|B= S by A1,A3,TOPMETR:26;
     A27: dom f3=B by FUNCT_2:def 1;
     A28: (dom f5) /\ B
       =REAL /\ B by FUNCT_2:def 1,TOPMETR:24
      .=B by TOPMETR:24,XBOOLE_1:28;
       for x being set st x in dom f3 holds f3.x=f5.x
      proof let x be set;assume
        A29: x in dom f3;
        then A30: x in the carrier of S by FUNCT_2:def 1;
        reconsider rx=x as Real by A3,A13,A29;
          f4.x=e*rx+d by A24,TOPMETR:24;
       hence f3.x=f5.x by A11,A30;
      end;
    then f3=f5|B by A27,A28,FUNCT_1:68;
    then A31: f3 is continuous by A25,A26,TOPMETR:10;
    reconsider C=the carrier of T as Subset of R^1
                  by A4,TOPMETR:24;
      R^1|C=T by A2,A4,TOPMETR:26;
    then A32: f2 is continuous by A31,TOPMETR:9;
    A33: S is compact by A1,HEINE:11;
      T=TopSpaceMetr(Closed-Interval-MSpace(e*a+d,e*b+d)) by TOPMETR:def 8;
    then T is_T2 by PCOMPS_1:38;
    then f2 is_homeomorphism by A14,A16,A22,A32,A33,COMPTS_1:26;
 hence thesis by A12;
end;

theorem Th40: for a,b,d,e being Real st a<=b & e<0
  ex f being map of
     Closed-Interval-TSpace(a,b),Closed-Interval-TSpace(e*b+d,e*a+d)
  st f is_homeomorphism & for r being Real st r in [.a,b.] holds
  f.r=e*r+d
proof let a,b,d,e be Real;assume A1: a<=b & e<0;
     then e*a>=e*b by REAL_1:52;
  then A2: e*a+d>=e*b+d by REAL_1:55;
  set X=the carrier of Closed-Interval-TSpace(a,b);
  set Y=the carrier of Closed-Interval-TSpace(e*b+d,e*a+d);
  defpred P[set,set] means (for r being Real st $1=r holds $2=e*r+d);
  A3: X=[.a,b.] by A1,TOPMETR:25;
  A4: Y=[.e*b+d,e*a+d.] by A2,TOPMETR:25;
  A5: for x being set st x in X ex y being set st y in Y & P[x,y]
   proof let x be set;assume A6: x in X;
     then reconsider r1=x as Real by A3;
     A7: a<=r1 & r1<=b by A3,A6,TOPREAL5:1;
     then A8: e*a>=e*r1 by A1,REAL_1:52;
     A9: e*r1>=e*b by A1,A7,REAL_1:52;
     set y1=e*r1+d;
     A10: for r being Real st x=r holds y1=e*r+d;
       e*a+d>=y1 & y1 >=e*b+d by A8,A9,REAL_1:55;
     then y1 in Y by A4,TOPREAL5:1;
    hence ex y being set st y in Y & P[x,y] by A10;
   end;
     ex f being Function of X,Y st for x being set st x in X holds P[x,f.x]
                    from FuncEx1(A5);
   then consider f1 being Function of X,Y such that
   A11: for x being set st x in X holds P[x,f1.x];
   reconsider f2=f1 as map of
     Closed-Interval-TSpace(a,b),Closed-Interval-TSpace(e*b+d,e*a+d);
   A12: for r being Real st r in [.a,b.] holds f2.r=e*r+d by A3,A11;
    set S=Closed-Interval-TSpace(a,b);
    set T=Closed-Interval-TSpace(e*b+d,e*a+d);
    A13: dom f2=the carrier of S by FUNCT_2:def 1;
    then A14: dom f2=[#]S by PRE_TOPC:12;
      for x1,x2 being set st x1 in dom f2 & x2 in dom f2 & f2.x1=f2.x2
        holds x1=x2
     proof let x1,x2 be set;assume
       A15: x1 in dom f2 & x2 in dom f2 & f2.x1=f2.x2;
       then reconsider r1=x1 as Real by A3,A13;
       reconsider r2=x2 as Real by A3,A13,A15;
         f2.x1=e*r1+d by A11,A13,A15;
       then e*r1+d-d=e*r2+d-d by A11,A13,A15 .=e*r2 by XCMPLX_1:26;
       then e*r1=e*r2 by XCMPLX_1:26;
       then r1*e/e=r2 by A1,XCMPLX_1:90;
      hence x1=x2 by A1,XCMPLX_1:90;
     end;
    then A16: f2 is one-to-one by FUNCT_1:def 8;
      rng f2 c= the carrier of T;
    then A17: rng f2 c= [#]T by PRE_TOPC:12;
      [#]T c= rng f2
     proof let y be set;assume A18: y in [#]T;
       then y in [.e*b+d,e*a+d.] by A4;
       then reconsider ry=y as Real;
       A19: e*b+d <= ry & ry<=e*a+d by A4,A18,TOPREAL5:1;
       then e*a+d-d>=ry-d by REAL_1:49;
       then e*a>=ry-d by XCMPLX_1:26;
       then a*e/e<=(ry-d)/e by A1,REAL_1:74;
       then A20: a<=(ry-d)/e by A1,XCMPLX_1:90;
         e*b+d-d<=ry-d by A19,REAL_1:49;
       then e*b<=ry-d by XCMPLX_1:26;
       then b*e/e>=(ry-d)/e by A1,REAL_1:74;
       then b>=(ry-d)/e by A1,XCMPLX_1:90;
       then A21: (ry-d)/e in [.a,b.] by A20,TOPREAL5:1;
       then f2.((ry-d)/e)=e*((ry-d)/e)+d by A3,A11
       .=ry-d+d by A1,XCMPLX_1:88.=ry by XCMPLX_1:27;
      hence y in rng f2 by A3,A13,A21,FUNCT_1:12;
     end;
    then A22: rng f2 = [#]T by A17,XBOOLE_0:def 10;
    then rng f2=Y by PRE_TOPC:12;
    then f1 is Function of the carrier of S,the carrier of R^1
                              by A4,A13,FUNCT_2:4,TOPMETR:24;
    then reconsider f3=f1 as map of S,R^1;
    defpred P1[set,set] means for r being Real st r=$1 holds $2=e*r+d;
    A23: for x being set st x in the carrier of R^1
         ex y being set st y in the carrier of R^1 & P1[x,y]
      proof let x be set;assume x in the carrier of R^1;
        then reconsider rx=x as Real by TOPMETR:24;
        reconsider ry=e*rx+d as Real;
          for r being Real st r=x holds ry=e*r+d;
       hence ex y being set st y in the carrier of R^1 & P1[x,y]
          by TOPMETR:24;
      end;
       ex f4 being Function of the carrier of R^1,the carrier of R^1 st
        for x being set st x in the carrier of R^1 holds P1[x,f4.x]
            from FuncEx1(A23);
     then consider f4 being Function of the carrier of R^1,the carrier of R^1
      such that
      A24: for x being set st x in the carrier of R^1 holds P1[x,f4.x];
     reconsider f5=f4 as map of R^1,R^1;
       for x being Real holds f5.x = e*x + d by A24,TOPMETR:24;
     then A25: f5 is continuous by TOPMETR:28;
     reconsider B=the carrier of S as Subset of R^1
                   by A3,TOPMETR:24;
     A26: R^1|B= S by A1,A3,TOPMETR:26;
     A27: dom f3=B by FUNCT_2:def 1;
     A28:(dom f5) /\ B =REAL /\ B by FUNCT_2:def 1,TOPMETR:24
      .=B by TOPMETR:24,XBOOLE_1:28;
       for x being set st x in dom f3 holds f3.x=f5.x
      proof let x be set;assume
        A29: x in dom f3;
        then A30: x in the carrier of S by FUNCT_2:def 1;
        reconsider rx=x as Real by A3,A13,A29;
          f4.x=e*rx+d by A24,TOPMETR:24;
       hence f3.x=f5.x by A11,A30;
      end;
     then f3=f5|B by A27,A28,FUNCT_1:68;
    then A31: f3 is continuous by A25,A26,TOPMETR:10;
    reconsider C=the carrier of T as Subset of R^1
                by A4,TOPMETR:24;
      R^1|C=T by A2,A4,TOPMETR:26;
    then A32: f2 is continuous by A31,TOPMETR:9;
    A33: S is compact by A1,HEINE:11;
      T=TopSpaceMetr(Closed-Interval-MSpace(e*b+d,e*a+d))
                                by TOPMETR:def 8;
    then T is_T2 by PCOMPS_1:38;
    then f2 is_homeomorphism by A14,A16,A22,A32,A33,COMPTS_1:26;
 hence thesis by A12;
end;

theorem Th41:
  ex f being map of I[01],Closed-Interval-TSpace(-1,1)
  st f is_homeomorphism & (for r being Real st r in [.0,1.] holds
  f.r=(-2)*r+1) & f.0=1 & f.1=-1
proof
  consider f being map of I[01],
      Closed-Interval-TSpace((-2)*1+1,(-2)*0+1) such that
  A1: f is_homeomorphism & (for r being Real st r in [.0,1.] holds
            f.r=(-2)*r+1) by Th40,TOPMETR:27;
  A2: f.0=(-2)*0+1 by A1,Lm1;
    1 in [.0,1.] by TOPREAL5:1;
  then f.1=-1 by A1;
 hence thesis by A1,A2;
end;

theorem Th42:
  ex f being map of I[01],Closed-Interval-TSpace(-1,1)
  st f is_homeomorphism & (for r being Real st r in [.0,1.] holds
  f.r=2*r-1) & f.0=-1 & f.1=1
proof
  consider f being map of I[01],
      Closed-Interval-TSpace(2*0+-1,2*1+-1) such that
  A1: f is_homeomorphism & (for r being Real st r in [.0,1.] holds
            f.r=2*r+-1) by Th39,TOPMETR:27;
  A2: for r being Real st r in [.0,1.] holds f.r=2*r-1
     proof let r be Real;assume r in [.0,1.];
      hence f.r=2*r+-1 by A1 .=2*r-1 by XCMPLX_0:def 8;
     end;
  then A3: f.0=2*0-1 by Lm1 .=-1;
    1 in [.0,1.] by TOPREAL5:1;
  then f.1=2*1-1 by A2 .=1;
 hence thesis by A1,A2,A3;
end;

Lm4:
now let P be compact non empty Subset of TOP-REAL 2;
   assume A1: P={p where p is Point of TOP-REAL 2: |.p.|=1};
   reconsider g=proj1 as map of TOP-REAL 2,R^1 by JGRAPH_2:16;
   set K0=Lower_Arc(P);
   reconsider g2=g|K0 as map of (TOP-REAL 2)|K0,R^1 by JGRAPH_3:12;
   A2: for p being Point of (TOP-REAL 2)|K0 holds g2.p=proj1.p
    proof let p be Point of (TOP-REAL 2)|K0;
        p in the carrier of (TOP-REAL 2)|K0;
      then p in K0 by JORDAN1:1;
     hence g2.p=proj1.p by FUNCT_1:72;
    end;
   then A3: g2 is continuous by JGRAPH_2:39;
   A4: dom g2=the carrier of ((TOP-REAL 2)|K0) by FUNCT_2:def 1;
   then A5: dom g2=K0 by JORDAN1:1;
   A6: K0 c= P by A1,Th36;
   A7: rng g2 c= the carrier of Closed-Interval-TSpace(-1,1)
    proof let x be set;assume x in rng g2;
      then consider z being set such that
      A8: z in dom g2 & x=g2.z by FUNCT_1:def 5;
        z in P by A5,A6,A8;
      then consider p being Point of TOP-REAL 2 such that
      A9: p=z & |.p.|=1 by A1;
      A10: x=proj1.p by A2,A4,A8,A9 .=p`1 by PSCOMP_1:def 28;
        1=(p`1)^2+(p`2)^2 by A9,JGRAPH_3:10,SQUARE_1:59;
      then 1-(p`1)^2=(p`2)^2 by XCMPLX_1:26;
      then 1-(p`1)^2>=0 by SQUARE_1:72;
      then -(1-(p`1)^2)<=0 by REAL_1:66;
      then (p`1)^2-1<=0 by XCMPLX_1:143;
      then -1<=p`1 & p`1<=1 by JGRAPH_3:5;
      then x in [.-1,1.] by A10,TOPREAL5:1;
     hence x in the carrier of Closed-Interval-TSpace(-1,1) by TOPMETR:25;
    end;
   then g2 is Function of the carrier of ((TOP-REAL 2)|K0),
    the carrier of Closed-Interval-TSpace(-1,1) by A4,FUNCT_2:4;
   then reconsider g3=g2 as map of (TOP-REAL 2)|K0,
                Closed-Interval-TSpace(-1,1);
     dom g3=the carrier of ((TOP-REAL 2)|K0) by FUNCT_2:def 1;
   then dom g3=[#]((TOP-REAL 2)|K0) by PRE_TOPC:12;
   then A11: dom g3=K0 by PRE_TOPC:def 10;
   A12: rng g2 c= [#](Closed-Interval-TSpace(-1,1)) by A7,PRE_TOPC:12;
   A13: [#](Closed-Interval-TSpace(-1,1)) c= rng g3
    proof let x be set;assume x in [#](Closed-Interval-TSpace(-1,1));
       then x in the carrier of (Closed-Interval-TSpace(-1,1));
       then A14: x in [.-1,1.] by TOPMETR:25;
       then reconsider r=x as Real;
       set q=|[r,-sqrt(1-r^2)]|;
       A15: |.q.|=sqrt((q`1)^2+(q`2)^2) by JGRAPH_3:10
           .=sqrt(r^2+(q`2)^2) by EUCLID:56
           .=sqrt(r^2+(-sqrt(1-r^2))^2) by EUCLID:56
           .=sqrt(r^2+(sqrt(1-r^2))^2) by SQUARE_1:61;
         -1<=r & r<=1 by A14,TOPREAL5:1;
       then 1^2>=r^2 by JGRAPH_2:7;
       then A16: 1-r^2>=0 by SQUARE_1:12,59;
       then 0<=sqrt(1-r^2) by SQUARE_1:def 4;
       then A17: -sqrt(1-r^2)<=0 by REAL_1:66;
         |.q.|=sqrt(r^2+(1-r^2)) by A15,A16,SQUARE_1:def 4
           .=1 by SQUARE_1:83,XCMPLX_1:27;
       then A18: q in P by A1;
         q`2=-sqrt(1-r^2) by EUCLID:56;
       then q in {p where p is Point of TOP-REAL 2:p in P & p`2<=0} by A17,A18
;
       then A19: q in dom g3 by A1,A11,Th38;
       then g3.q=proj1.q by A2,A4 .=q`1 by PSCOMP_1:def 28.=r by EUCLID:56;
     hence x in rng g3 by A19,FUNCT_1:def 5;
    end;
   reconsider B=[.-1,1.] as non empty Subset of R^1
                               by TOPMETR:24,TOPREAL5:1;
   A20: Closed-Interval-TSpace(-1,1)=R^1|B by TOPMETR:26;
     for x,y being set st x in dom g3 & y in dom g3 & g3.x=g3.y holds x=y
    proof let x,y be set;assume
      A21: x in dom g3 & y in dom g3 & g3.x=g3.y;
      then reconsider p1=x as Point of TOP-REAL 2 by A11;
      reconsider p2=y as Point of TOP-REAL 2 by A11,A21;
      A22: g3.x=proj1.p1 by A2,A4,A21 .=p1`1 by PSCOMP_1:def 28;
      A23: g3.y=proj1.p2 by A2,A4,A21 .=p2`1 by PSCOMP_1:def 28;
      A24: p1 in P by A6,A11,A21;
        p2 in P by A6,A11,A21;
      then consider p22 being Point of TOP-REAL 2 such that
      A25: p2=p22 & |.p22.|=1 by A1;
        1^2= (p2`1)^2+(p2`2)^2 by A25,JGRAPH_3:10;
      then A26: 1^2-(p2`1)^2= (p2`2)^2 by XCMPLX_1:26;
      consider p11 being Point of TOP-REAL 2 such that
      A27: p1=p11 & |.p11.|=1 by A1,A24;
        1^2= (p1`1)^2+(p1`2)^2 by A27,JGRAPH_3:10;
      then 1^2-(p1`1)^2= (p1`2)^2 by XCMPLX_1:26;
      then (-(p1`2))^2 =(p2`2)^2 by A21,A22,A23,A26,SQUARE_1:61;
      then A28: (-(p1`2))^2 =(-(p2`2))^2 by SQUARE_1:61;
        p1 in {p3 where p3 is Point of TOP-REAL 2:p3 in P & p3`2<=0}
                     by A1,A11,A21,Th38;
      then consider p33 being Point of TOP-REAL 2 such that
      A29: p1=p33 & p33 in P & p33`2<=0;
        --(p1`2)<=0 by A29;
      then -(p1`2)>=0 by REAL_1:66;
      then A30: -(p1`2)=sqrt((-(p2`2))^2) by A28,SQUARE_1:89;
        p2 in {p3 where p3 is Point of TOP-REAL 2:p3 in P & p3`2<=0}
                     by A1,A11,A21,Th38;
      then consider p44 being Point of TOP-REAL 2 such that
      A31: p2=p44 & p44 in P & p44`2<=0;
        --(p2`2)<=0 by A31;
      then -(p2`2)>=0 by REAL_1:66;
      then -(p1`2)=-(p2`2) by A30,SQUARE_1:89;
      then --(p1`2)=(p2`2);
      then p1=|[p2`1,p2`2]| by A21,A22,A23,EUCLID:57
      .=p2 by EUCLID:57;
     hence x=y;
    end;
   hence proj1|K0 is continuous map of (TOP-REAL 2)|K0,
     Closed-Interval-TSpace(-1,1) &
    proj1|K0 is one-to-one &
    rng (proj1|K0)=[#](Closed-Interval-TSpace(-1,1)) by A3,A12,A13,A20,FUNCT_1:
def 8,JGRAPH_1:63,XBOOLE_0:def 10;
end;

Lm5:
now let P be compact non empty Subset of TOP-REAL 2;
   assume A1: P={p where p is Point of TOP-REAL 2: |.p.|=1};
   reconsider g=proj1 as map of TOP-REAL 2,R^1 by JGRAPH_2:16;
   set K0=Upper_Arc(P);
   reconsider g2=g|K0 as map of (TOP-REAL 2)|K0,R^1 by JGRAPH_3:12;
   A2: for p being Point of (TOP-REAL 2)|K0 holds g2.p=proj1.p
    proof let p be Point of (TOP-REAL 2)|K0;
        p in the carrier of (TOP-REAL 2)|K0;
      then p in K0 by JORDAN1:1;
     hence g2.p=proj1.p by FUNCT_1:72;
    end;
   then A3: g2 is continuous by JGRAPH_2:39;
   A4: dom g2=the carrier of ((TOP-REAL 2)|K0) by FUNCT_2:def 1;
   then A5: dom g2=K0 by JORDAN1:1;
   A6: K0 c= P by A1,Th36;
   A7: rng g2 c= the carrier of Closed-Interval-TSpace(-1,1)
    proof let x be set;assume x in rng g2;
      then consider z being set such that
      A8: z in dom g2 & x=g2.z by FUNCT_1:def 5;
        z in P by A5,A6,A8;
      then consider p being Point of TOP-REAL 2 such that
      A9: p=z & |.p.|=1 by A1;
      A10: x=proj1.p by A2,A4,A8,A9 .=p`1 by PSCOMP_1:def 28;
        1=(p`1)^2+(p`2)^2 by A9,JGRAPH_3:10,SQUARE_1:59;
      then 1-(p`1)^2=(p`2)^2 by XCMPLX_1:26;
      then 1-(p`1)^2>=0 by SQUARE_1:72;
      then -(1-(p`1)^2)<=0 by REAL_1:66;
      then (p`1)^2-1<=0 by XCMPLX_1:143;
      then -1<=p`1 & p`1<=1 by JGRAPH_3:5;
      then x in [.-1,1.] by A10,TOPREAL5:1;
     hence x in the carrier of Closed-Interval-TSpace(-1,1)
                      by TOPMETR:25;
    end;
   then g2 is Function of the carrier of ((TOP-REAL 2)|K0),
    the carrier of Closed-Interval-TSpace(-1,1) by A4,FUNCT_2:4;
   then reconsider g3=g2 as map of (TOP-REAL 2)|K0,
                Closed-Interval-TSpace(-1,1);
     dom g3=the carrier of ((TOP-REAL 2)|K0) by FUNCT_2:def 1;
   then dom g3=[#]((TOP-REAL 2)|K0) by PRE_TOPC:12;
   then A11: dom g3=K0 by PRE_TOPC:def 10;
   A12: rng g2 c= [#](Closed-Interval-TSpace(-1,1)) by A7,PRE_TOPC:12;
   A13: [#](Closed-Interval-TSpace(-1,1)) c= rng g3
    proof let x be set;assume x in
       [#](Closed-Interval-TSpace(-1,1));
       then x in the carrier of (Closed-Interval-TSpace(-1,1));
       then A14: x in [.-1,1.] by TOPMETR:25;
       then reconsider r=x as Real;
       set q=|[r,sqrt(1-r^2)]|;
       A15: |.q.|=sqrt((q`1)^2+(q`2)^2) by JGRAPH_3:10
           .=sqrt(r^2+(q`2)^2) by EUCLID:56
           .=sqrt(r^2+(sqrt(1-r^2))^2) by EUCLID:56;
         -1<=r & r<=1 by A14,TOPREAL5:1;
       then 1^2>=r^2 by JGRAPH_2:7;
       then A16: 1-r^2>=0 by SQUARE_1:12,59;
       then A17: 0<=sqrt(1-r^2) by SQUARE_1:def 4;
         |.q.|=sqrt(r^2+(1-r^2)) by A15,A16,SQUARE_1:def 4
           .=1 by SQUARE_1:83,XCMPLX_1:27;
       then A18: q in P by A1;
         q`2=sqrt(1-r^2) by EUCLID:56;
       then q in {p where p is Point of TOP-REAL 2:p in P & p`2>=0} by A17,A18
;
       then A19: q in dom g3 by A1,A11,Th37;
       then g3.q=proj1.q by A2,A4 .=q`1 by PSCOMP_1:def 28.=r by EUCLID:56;
     hence x in rng g3 by A19,FUNCT_1:def 5;
    end;
   reconsider B=[.-1,1.] as non empty Subset of R^1
                               by TOPMETR:24,TOPREAL5:1;
   A20: Closed-Interval-TSpace(-1,1)=R^1|B by TOPMETR:26;
     for x,y being set st x in dom g3 & y in dom g3 & g3.x=g3.y holds x=y
    proof let x,y be set;assume
      A21: x in dom g3 & y in dom g3 & g3.x=g3.y;
      then reconsider p1=x as Point of TOP-REAL 2 by A11;
      reconsider p2=y as Point of TOP-REAL 2 by A11,A21;
      A22: g3.x=proj1.p1 by A2,A4,A21 .=p1`1 by PSCOMP_1:def 28;
      A23: g3.y=proj1.p2 by A2,A4,A21 .=p2`1 by PSCOMP_1:def 28;
      A24: p1 in P by A6,A11,A21;
        p2 in P by A6,A11,A21;
      then consider p22 being Point of TOP-REAL 2 such that
      A25: p2=p22 & |.p22.|=1 by A1;
        1^2= (p2`1)^2+(p2`2)^2 by A25,JGRAPH_3:10;
      then A26: 1^2-(p2`1)^2= (p2`2)^2 by XCMPLX_1:26;
      consider p11 being Point of TOP-REAL 2 such that
      A27: p1=p11 & |.p11.|=1 by A1,A24;
        1^2= (p1`1)^2+(p1`2)^2 by A27,JGRAPH_3:10;
      then A28: (p1`2)^2 =(p2`2)^2 by A21,A22,A23,A26,XCMPLX_1:26;
        p1 in {p3 where p3 is Point of TOP-REAL 2:p3 in P & p3`2>=0}
                     by A1,A11,A21,Th37;
      then consider p33 being Point of TOP-REAL 2 such that
      A29: p1=p33 & p33 in P & p33`2>=0;
      A30: p1`2=sqrt(((p2`2))^2) by A28,A29,SQUARE_1:89;
        p2 in {p3 where p3 is Point of TOP-REAL 2:p3 in P & p3`2>=0}
                     by A1,A11,A21,Th37;
      then consider p44 being Point of TOP-REAL 2 such that
      A31: p2=p44 & p44 in P & p44`2>=0;
        (p1`2)=(p2`2) by A30,A31,SQUARE_1:89;
      then p1=|[p2`1,p2`2]| by A21,A22,A23,EUCLID:57
      .=p2 by EUCLID:57;
     hence x=y;
    end;
   hence proj1|K0 is continuous map of (TOP-REAL 2)|K0,
     Closed-Interval-TSpace(-1,1) &
    proj1|K0 is one-to-one &
    rng (proj1|K0)=[#](Closed-Interval-TSpace(-1,1)) by A3,A12,A13,A20,FUNCT_1:
def 8,JGRAPH_1:63,XBOOLE_0:def 10;
end;

theorem Th43: for P being compact non empty Subset of TOP-REAL 2
 st P={p where p is Point of TOP-REAL 2: |.p.|=1}
  ex f being map of
     Closed-Interval-TSpace(-1,1),(TOP-REAL 2)|Lower_Arc(P)
  st f is_homeomorphism & (for q being Point of TOP-REAL 2 st
  q in Lower_Arc(P) holds f.(q`1)=q) & f.(-1)=W-min(P) & f.1=E-max(P)
 proof let P be compact non empty Subset of TOP-REAL 2;
   assume A1: P={p where p is Point of TOP-REAL 2: |.p.|=1};
  then A2: P is_simple_closed_curve by JGRAPH_3:36;
  then A3: Lower_Arc(P) is_an_arc_of E-max(P),W-min(P) by JORDAN6:def 9;
   set P4=Lower_Arc(P);
   A4: P4 is_an_arc_of E-max(P),W-min(P)
     & Upper_Arc(P) /\ P4={W-min(P),E-max(P)}
       & Upper_Arc(P) \/ P4=P
         & First_Point(Upper_Arc(P),W-min(P),E-max(P),
         Vertical_Line((W-bound(P)+E-bound(P))/2))`2>
              Last_Point(P4,E-max(P),W-min(P),
         Vertical_Line((W-bound(P)+E-bound(P))/2))`2 by A2,JORDAN6:def 9;
   reconsider g=proj1 as map of TOP-REAL 2,R^1 by JGRAPH_2:16;
   set K0=Lower_Arc(P);
   reconsider g2=g|K0 as map of (TOP-REAL 2)|K0,R^1 by JGRAPH_3:12;
   A5: for p being Point of (TOP-REAL 2)|K0 holds g2.p=proj1.p
    proof let p be Point of (TOP-REAL 2)|K0;
        p in the carrier of (TOP-REAL 2)|K0;
      then p in K0 by JORDAN1:1;
     hence g2.p=proj1.p by FUNCT_1:72;
    end;
   reconsider g3=g2 as continuous map of (TOP-REAL 2)|K0,
                Closed-Interval-TSpace(-1,1) by A1,Lm4;
     E-max(P) in {W-min(P),E-max(P)} by TARSKI:def 2;
   then A6: E-max(P) in Lower_Arc(P) by A4,XBOOLE_0:def 3;
     W-min(P) in {W-min(P),E-max(P)} by TARSKI:def 2;
   then A7: W-min(P) in Lower_Arc(P) by A4,XBOOLE_0:def 3;
   A8: dom g3=the carrier of ((TOP-REAL 2)|K0) by FUNCT_2:def 1;
   then A9: dom g3=[#]((TOP-REAL 2)|K0) by PRE_TOPC:12;
   then A10: dom g3=K0 by PRE_TOPC:def 10;
   A11: rng g3=[#](Closed-Interval-TSpace(-1,1)) by A1,Lm4;
   A12: g3 is one-to-one by A1,Lm4;
      K0 is non empty compact by A3,JORDAN5A:1;
    then A13: (TOP-REAL 2)|K0 is compact by COMPTS_1:12;
      Closed-Interval-TSpace(-1,1)
      =TopSpaceMetr(Closed-Interval-MSpace(-1,1)) by TOPMETR:def 8;
    then Closed-Interval-TSpace(-1,1) is_T2 by PCOMPS_1:38;
    then g3 is_homeomorphism by A9,A11,A12,A13,COMPTS_1:26;
    then A14: g3/" is_homeomorphism by TOPS_2:70;
    A15: for q be Point of TOP-REAL 2 st
    q in Lower_Arc(P) holds (g3/").(q`1)=q
    proof let q be Point of TOP-REAL 2;
      assume A16: q in Lower_Arc(P);
      reconsider g4=g3 as Function;
      A17: q in dom g4 implies q = (g4").(g4.q) & q = (g4"*g4).q
                       by A12,FUNCT_1:56;
        g3.q=proj1.q by A5,A8,A10,A16 .=q`1 by PSCOMP_1:def 28;
     hence (g3/").(q`1)=q by A9,A11,A12,A16,A17,PRE_TOPC:def 10,TOPS_2:def 4;
    end;
    A18: W-min(P)=|[-1,0]| by A1,Th32;
    A19: g3/".(-1)=g3/".((|[-1,0]|)`1) by EUCLID:56 .=W-min(P)
                                by A7,A15,A18;
    A20: E-max(P)=|[1,0]| by A1,Th33;
      g3/".1=g3/".((|[1,0]|)`1) by EUCLID:56 .=E-max(P) by A6,A15,A20;
  hence thesis by A14,A15,A19;
 end;

theorem Th44: for P being compact non empty Subset of TOP-REAL 2
 st P={p where p is Point of TOP-REAL 2: |.p.|=1}
  ex f being map of
     Closed-Interval-TSpace(-1,1),(TOP-REAL 2)|Upper_Arc(P)
  st f is_homeomorphism & (for q being Point of TOP-REAL 2 st
  q in Upper_Arc(P) holds f.(q`1)=q) & f.(-1)=W-min(P) & f.1=E-max(P)
 proof let P be compact non empty Subset of TOP-REAL 2;
   assume A1: P={p where p is Point of TOP-REAL 2: |.p.|=1};
  then A2: P is_simple_closed_curve by JGRAPH_3:36;
  then A3: Upper_Arc(P) is_an_arc_of W-min(P),E-max(P) by JORDAN6:def 8;
   set P4=Lower_Arc(P);
   A4: P4 is_an_arc_of E-max(P),W-min(P)
     & Upper_Arc(P) /\ P4={W-min(P),E-max(P)}
       & Upper_Arc(P) \/ P4=P
         & First_Point(Upper_Arc(P),W-min(P),E-max(P),
         Vertical_Line((W-bound(P)+E-bound(P))/2))`2>
              Last_Point(P4,E-max(P),W-min(P),
         Vertical_Line((W-bound(P)+E-bound(P))/2))`2 by A2,JORDAN6:def 9;
   reconsider g=proj1 as map of TOP-REAL 2,R^1 by JGRAPH_2:16;
   set K0=Upper_Arc(P);
   reconsider g2=g|K0 as map of (TOP-REAL 2)|K0,R^1 by JGRAPH_3:12;
   A5: for p being Point of (TOP-REAL 2)|K0 holds g2.p=proj1.p
    proof let p be Point of (TOP-REAL 2)|K0;
        p in the carrier of (TOP-REAL 2)|K0;
      then p in K0 by JORDAN1:1;
     hence g2.p=proj1.p by FUNCT_1:72;
    end;
   reconsider g3=g2 as continuous map of (TOP-REAL 2)|K0,
                Closed-Interval-TSpace(-1,1) by A1,Lm5;
     E-max(P) in {W-min(P),E-max(P)} by TARSKI:def 2;
   then A6: E-max(P) in Upper_Arc(P) by A4,XBOOLE_0:def 3;
     W-min(P) in {W-min(P),E-max(P)} by TARSKI:def 2;
   then A7: W-min(P) in Upper_Arc(P) by A4,XBOOLE_0:def 3;
   A8: dom g3=the carrier of ((TOP-REAL 2)|K0) by FUNCT_2:def 1;
   then A9: dom g3=[#]((TOP-REAL 2)|K0) by PRE_TOPC:12;
   then A10: dom g3=K0 by PRE_TOPC:def 10;
   A11: rng g3=[#](Closed-Interval-TSpace(-1,1)) by A1,Lm5;
   A12: g3 is one-to-one by A1,Lm5;
      K0 is non empty compact by A3,JORDAN5A:1;
    then A13: (TOP-REAL 2)|K0 is compact by COMPTS_1:12;
      Closed-Interval-TSpace(-1,1)
      =TopSpaceMetr(Closed-Interval-MSpace(-1,1)) by TOPMETR:def 8;
    then Closed-Interval-TSpace(-1,1) is_T2 by PCOMPS_1:38;
    then g3 is_homeomorphism by A9,A11,A12,A13,COMPTS_1:26;
   then A14: g3/" is_homeomorphism by TOPS_2:70;
   A15: for q be Point of TOP-REAL 2 st
    q in Upper_Arc(P) holds (g3/").(q`1)=q
    proof let q be Point of TOP-REAL 2;
      assume A16: q in Upper_Arc(P);
      reconsider g4=g3 as Function;
      A17: q in dom g4 implies q = (g4").(g4.q) & q = (g4"*g4).q
                       by A12,FUNCT_1:56;
        g3.q=proj1.q by A5,A8,A10,A16 .=q`1 by PSCOMP_1:def 28;
     hence (g3/").(q`1)=q by A9,A11,A12,A16,A17,PRE_TOPC:def 10,TOPS_2:def 4;
    end;
    A18: W-min(P)=|[-1,0]| by A1,Th32;
    A19: g3/".(-1)=g3/".((|[-1,0]|)`1) by EUCLID:56 .=W-min(P)
                                by A7,A15,A18;
    A20: E-max(P)=|[1,0]| by A1,Th33;
      g3/".(1)=g3/".((|[1,0]|)`1) by EUCLID:56 .=E-max(P) by A6,A15,A20;
  hence thesis by A14,A15,A19;
 end;

theorem Th45: for P being compact non empty Subset of TOP-REAL 2
 st P={p where p is Point of TOP-REAL 2: |.p.|=1}
  ex f being map of I[01],(TOP-REAL 2)|Lower_Arc(P)
  st f is_homeomorphism & (for q1,q2 being Point of TOP-REAL 2,
  r1,r2 being Real st f.r1=q1 & f.r2=q2 &
  r1 in [.0,1.] & r2 in [.0,1.] holds r1<r2 iff q1`1>q2`1)&
  f.0 = E-max(P) & f.1 = W-min(P)
proof let P be compact non empty Subset of TOP-REAL 2;
 assume A1: P={p where p is Point of TOP-REAL 2: |.p.|=1};
  then consider f1 being map of
     Closed-Interval-TSpace(-1,1),(TOP-REAL 2)|Lower_Arc(P)
  such that
  A2: f1 is_homeomorphism & (for q being Point of TOP-REAL 2 st
  q in Lower_Arc(P) holds f1.(q`1)=q)& f1.(-1)=W-min(P) & f1.1=E-max(P)
                        by Th43;
  consider g being map of
     I[01],Closed-Interval-TSpace(-1,1) such that
  A3: g is_homeomorphism & (for r being Real st r in [.0,1.] holds
  g.r=(-2)*r+1)& g.0=1 & g.1=-1 by Th41;
  reconsider T= (TOP-REAL 2)|Lower_Arc(P) as non empty TopSpace;
  reconsider f2=f1 as map of Closed-Interval-TSpace(-1,1),T;
  A4: f2*g is_homeomorphism by A2,A3,TOPS_2:71;
  reconsider h=f1*g as map of I[01],(TOP-REAL 2)|Lower_Arc(P);
  A5: for q1,q2 being Point of TOP-REAL 2,
   r1,r2 being Real st h.r1=q1 & h.r2=q2 &
   r1 in [.0,1.] & r2 in [.0,1.] holds r1<r2 iff q1`1>q2`1
   proof let q1,q2 be Point of TOP-REAL 2,r1,r2 be Real;
    assume A6: h.r1=q1 & h.r2=q2 &
       r1 in [.0,1.] & r2 in [.0,1.];
     A7: now assume A8: r1<r2;
       set s1=(-2)*r1+1,s2=(-2)*r2+1;
       set p1=|[s1,-sqrt(1-s1^2)]|,p2=|[s2,-sqrt(1-s2^2)]|;
       A9: g.r1=(-2)*r1+1 by A3,A6;
       A10: g.r2=(-2)*r2+1 by A3,A6;
         (-2)*r1 > (-2)*r2 by A8,REAL_1:71;
       then A11: (-2)*r1 +1 > (-2)*r2 +1 by REAL_1:67;
         r1<=1 by A6,TOPREAL5:1;
       then (-2)*r1>=(-2)*1 by REAL_1:52;
       then (-2)*r1+1>=(-2)*1+1 by REAL_1:55;
       then A12: -1<=s1;
         r1>=0 by A6,TOPREAL5:1;
       then (-2)*r1<=(-2)*0 by REAL_1:52;
       then (-2)*r1+1<=(-2)*0+1 by REAL_1:55;
       then s1^2<=1^2 by A12,JGRAPH_2:7;
       then A13: 1-s1^2>=0 by SQUARE_1:12,59;
       A14: (|[s1,-sqrt(1-s1^2)]|)`1=s1 by EUCLID:56;
       A15: (|[s1,-sqrt(1-s1^2)]|)`2=-sqrt(1-s1^2) by EUCLID:56;
         sqrt(1-s1^2)>=0 by A13,SQUARE_1:def 4;
       then A16: -sqrt(1-s1^2)<=0 by REAL_1:66;
         |.p1.|=sqrt((p1`1)^2+(p1`2)^2) by JGRAPH_3:10
           .=sqrt((s1)^2+(sqrt(1-s1^2))^2) by A14,A15,SQUARE_1:61
           .=sqrt((s1)^2+(1-s1^2)) by A13,SQUARE_1:def 4
           .=1 by SQUARE_1:83,XCMPLX_1:27;
       then p1 in P by A1;
       then p1 in {p3 where p3 is Point of TOP-REAL 2: p3 in P & p3`2<=0}
                         by A15,A16;
       then A17: |[s1,-sqrt(1-s1^2)]| in Lower_Arc(P) by A1,Th38;
         dom h=[.0,1.] by BORSUK_1:83,FUNCT_2:def 1;
       then h.r1=f1.s1 by A6,A9,FUNCT_1:22 .=p1 by A2,A14,A17;
       then A18: q1`1=s1 by A6,EUCLID:56;
         r2<=1 by A6,TOPREAL5:1;
       then (-2)*r2>=(-2)*1 by REAL_1:52;
       then (-2)*r2+1>=(-2)*1+1 by REAL_1:55;
       then A19: -1<=s2;
         r2>=0 by A6,TOPREAL5:1;
       then (-2)*r2<=(-2)*0 by REAL_1:52;
       then (-2)*r2+1<=(-2)*0+1 by REAL_1:55;
       then s2^2<=1^2 by A19,JGRAPH_2:7;
       then A20: 1-s2^2>=0 by SQUARE_1:12,59;
       A21: (|[s2,-sqrt(1-s2^2)]|)`1=s2 by EUCLID:56;
       A22: (|[s2,-sqrt(1-s2^2)]|)`2=-sqrt(1-s2^2) by EUCLID:56;
         sqrt(1-s2^2)>=0 by A20,SQUARE_1:def 4;
       then A23: -sqrt(1-s2^2)<=0 by REAL_1:66;
         |.p2.|=sqrt((p2`1)^2+(p2`2)^2) by JGRAPH_3:10
           .=sqrt((s2)^2+(sqrt(1-s2^2))^2) by A21,A22,SQUARE_1:61
           .=sqrt((s2)^2+(1-s2^2)) by A20,SQUARE_1:def 4
           .=1 by SQUARE_1:83,XCMPLX_1:27;
       then p2 in P by A1;
       then p2 in {p3 where p3 is Point of TOP-REAL 2: p3 in P & p3`2<=0}
                         by A22,A23;
       then A24: |[s2,-sqrt(1-s2^2)]| in Lower_Arc(P) by A1,Th38;
         dom h=[.0,1.] by BORSUK_1:83,FUNCT_2:def 1;
       then h.r2=f1.s2 by A6,A10,FUNCT_1:22 .=p2 by A2,A21,A24;
      hence q1`1>q2`1 by A6,A11,A18,EUCLID:56;
     end;
     A25: now assume A26: r2<r1;
       set s1=(-2)*r2+1,s2=(-2)*r1+1;
       set p1=|[s1,-sqrt(1-s1^2)]|,p2=|[s2,-sqrt(1-s2^2)]|;
       A27: g.r2=(-2)*r2+1 by A3,A6;
       A28: g.r1=(-2)*r1+1 by A3,A6;
       A29: (-2)*r2 > (-2)*r1 by A26,REAL_1:71;
         r2<=1 by A6,TOPREAL5:1;
       then (-2)*r2>=(-2)*1 by REAL_1:52;
       then (-2)*r2+1>=(-2)*1+1 by REAL_1:55;
       then A30: -1<=s1;
         r2>=0 by A6,TOPREAL5:1;
       then (-2)*r2<=(-2)*0 by REAL_1:52;
       then (-2)*r2+1<=(-2)*0+1 by REAL_1:55;
       then s1^2<=1^2 by A30,JGRAPH_2:7;
       then A31: 1-s1^2>=0 by SQUARE_1:12,59;
       A32: (|[s1,-sqrt(1-s1^2)]|)`1=s1 by EUCLID:56;
       A33: (|[s1,-sqrt(1-s1^2)]|)`2=-sqrt(1-s1^2) by EUCLID:56;
         sqrt(1-s1^2)>=0 by A31,SQUARE_1:def 4;
       then A34: -sqrt(1-s1^2)<=0 by REAL_1:66;
         |.p1.|=sqrt((p1`1)^2+(p1`2)^2) by JGRAPH_3:10
           .=sqrt((s1)^2+(sqrt(1-s1^2))^2) by A32,A33,SQUARE_1:61
           .=sqrt((s1)^2+(1-s1^2)) by A31,SQUARE_1:def 4
           .=1 by SQUARE_1:83,XCMPLX_1:27;
       then p1 in P by A1;
       then p1 in {p3 where p3 is Point of TOP-REAL 2: p3 in P & p3`2<=0}
                         by A33,A34;
       then A35: |[s1,-sqrt(1-s1^2)]| in Lower_Arc(P) by A1,Th38;
         dom h=[.0,1.] by BORSUK_1:83,FUNCT_2:def 1;
       then h.r2=f1.s1 by A6,A27,FUNCT_1:22 .=p1 by A2,A32,A35;
       then A36: q2`1=s1 by A6,EUCLID:56;
         r1<=1 by A6,TOPREAL5:1;
       then (-2)*r1>=(-2)*1 by REAL_1:52;
       then (-2)*r1+1>=(-2)*1+1 by REAL_1:55;
       then A37: -1<=s2;
         r1>=0 by A6,TOPREAL5:1;
       then (-2)*r1<=(-2)*0 by REAL_1:52;
       then (-2)*r1+1<=(-2)*0+1 by REAL_1:55;
       then s2^2<=1^2 by A37,JGRAPH_2:7;
       then A38: 1-s2^2>=0 by SQUARE_1:12,59;
       A39: (|[s2,-sqrt(1-s2^2)]|)`1=s2 by EUCLID:56;
       A40: (|[s2,-sqrt(1-s2^2)]|)`2=-sqrt(1-s2^2) by EUCLID:56;
         sqrt(1-s2^2)>=0 by A38,SQUARE_1:def 4;
       then A41: -sqrt(1-s2^2)<=0 by REAL_1:66;
         |.p2.|=sqrt((p2`1)^2+(p2`2)^2) by JGRAPH_3:10
           .=sqrt((s2)^2+(sqrt(1-s2^2))^2) by A39,A40,SQUARE_1:61
           .=sqrt((s2)^2+(1-s2^2)) by A38,SQUARE_1:def 4
           .=1 by SQUARE_1:83,XCMPLX_1:27;
       then p2 in P by A1;
       then p2 in {p3 where p3 is Point of TOP-REAL 2: p3 in P & p3`2<=0}
                          by A40,A41;
       then A42: |[s2,-sqrt(1-s2^2)]| in Lower_Arc(P) by A1,Th38;
         dom h=[.0,1.] by BORSUK_1:83,FUNCT_2:def 1;
       then h.r1=f1.s2 by A6,A28,FUNCT_1:22 .=p2 by A2,A39,A42;
      hence q2`1>q1`1 by A6,A29,A36,A39,REAL_1:67;
     end;
       now assume A43: q1`1>q2`1;
         now assume A44: r1>=r2;
           now per cases by A44,REAL_1:def 5;
         case r1>r2;
          hence contradiction by A25,A43;
         case r1=r2;
          hence contradiction by A6,A43;
         end;
        hence contradiction;
       end;
      hence r1<r2;
     end;
    hence r1<r2 iff q1`1>q2`1 by A7;
   end;
  A45: dom h=[.0,1.] by BORSUK_1:83,FUNCT_2:def 1;
  then 0 in dom h by TOPREAL5:1;
  then A46: h.0=E-max(P) by A2,A3,FUNCT_1:22;
    1 in dom h by A45,TOPREAL5:1;
  then h.1=W-min(P) by A2,A3,FUNCT_1:22;
 hence thesis by A4,A5,A46;
end;

theorem Th46: for P being compact non empty Subset of TOP-REAL 2
 st P={p where p is Point of TOP-REAL 2: |.p.|=1}
  ex f being map of I[01],(TOP-REAL 2)|Upper_Arc(P)
  st f is_homeomorphism & (for q1,q2 being Point of TOP-REAL 2,
  r1,r2 being Real st f.r1=q1 & f.r2=q2 &
  r1 in [.0,1.] & r2 in [.0,1.] holds r1<r2 iff q1`1<q2`1)&
  f.0 = W-min(P) & f.1 = E-max(P)
proof let P be compact non empty Subset of TOP-REAL 2;
 assume A1: P={p where p is Point of TOP-REAL 2: |.p.|=1};
  then consider f1 being map of
     Closed-Interval-TSpace(-1,1),(TOP-REAL 2)|Upper_Arc(P) such that
  A2: f1 is_homeomorphism & (for q being Point of TOP-REAL 2 st
  q in Upper_Arc(P) holds f1.(q`1)=q)& f1.(-1)=W-min(P) & f1.1=E-max(P)
                        by Th44;
  consider g being map of
     I[01],Closed-Interval-TSpace(-1,1) such that
  A3: g is_homeomorphism & (for r being Real st r in [.0,1.] holds
  g.r=2*r-1)& g.0=-1 & g.1=1 by Th42;
  reconsider T= (TOP-REAL 2)|Upper_Arc(P) as non empty TopSpace;
  reconsider f2=f1 as map of Closed-Interval-TSpace(-1,1),T;
  A4: f2*g is_homeomorphism by A2,A3,TOPS_2:71;
  reconsider h=f1*g as map of I[01],(TOP-REAL 2)|Upper_Arc(P);
  A5: for q1,q2 being Point of TOP-REAL 2,
   r1,r2 being Real st h.r1=q1 & h.r2=q2 &
   r1 in [.0,1.] & r2 in [.0,1.] holds r1<r2 iff q1`1<q2`1
   proof let q1,q2 be Point of TOP-REAL 2,r1,r2 be Real;
    assume A6: h.r1=q1 & h.r2=q2 &
       r1 in [.0,1.] & r2 in [.0,1.];
     A7: now assume A8: r1>r2;
       set s1=2*r1-1,s2=2*r2-1;
       set p1=|[s1,sqrt(1-s1^2)]|,p2=|[s2,sqrt(1-s2^2)]|;
       A9: g.r1=2*r1-1 by A3,A6;
       A10: g.r2=2*r2-1 by A3,A6;
       A11: 2*r1 > 2*r2 by A8,REAL_1:70;
         r1<=1 by A6,TOPREAL5:1;
       then 2*r1<=2*1 by AXIOMS:25;
       then A12: 2*r1-1<=2*1-1 by REAL_1:49;
         r1>=0 by A6,TOPREAL5:1;
       then 2*r1>=2*0 by AXIOMS:25;
       then A13: 2*r1-1>=2*0-1 by REAL_1:49;
         2*0-1=-1;
       then s1^2<=1^2 by A12,A13,JGRAPH_2:7;
       then A14: 1-s1^2>=0 by SQUARE_1:12,59;
       A15: (|[s1,sqrt(1-s1^2)]|)`1=s1 by EUCLID:56;
       A16: (|[s1,sqrt(1-s1^2)]|)`2=sqrt(1-s1^2) by EUCLID:56;
       A17: sqrt(1-s1^2)>=0 by A14,SQUARE_1:def 4;
         |.p1.|=sqrt((s1)^2+(sqrt(1-s1^2))^2) by A15,A16,JGRAPH_3:10
           .=sqrt((s1)^2+(1-s1^2)) by A14,SQUARE_1:def 4
           .=1 by SQUARE_1:83,XCMPLX_1:27;
       then p1 in P by A1;
       then p1 in {p3 where p3 is Point of TOP-REAL 2: p3 in P & p3`2>=0}
                         by A16,A17;
       then A18: |[s1,sqrt(1-s1^2)]| in Upper_Arc(P) by A1,Th37;
         dom h=[.0,1.] by BORSUK_1:83,FUNCT_2:def 1;
       then h.r1=f1.s1 by A6,A9,FUNCT_1:22 .=p1 by A2,A15,A18;
       then A19: q1`1=s1 by A6,EUCLID:56;
         r2<=1 by A6,TOPREAL5:1;
       then 2*r2<=2*1 by AXIOMS:25;
       then A20: 2*r2-1<=2*1-1 by REAL_1:49;
         r2>=0 by A6,TOPREAL5:1;
       then 2*r2>=2*0 by AXIOMS:25;
       then A21: 2*r2-1>=2*0-1 by REAL_1:49;
         2*0-1=-1;
       then s2^2<=1^2 by A20,A21,JGRAPH_2:7;
       then A22: 1-s2^2>=0 by SQUARE_1:12,59;
       A23: (|[s2,sqrt(1-s2^2)]|)`1=s2 by EUCLID:56;
       A24: (|[s2,sqrt(1-s2^2)]|)`2=sqrt(1-s2^2) by EUCLID:56;
       A25: sqrt(1-s2^2)>=0 by A22,SQUARE_1:def 4;
         |.p2.|=sqrt((s2)^2+(sqrt(1-s2^2))^2) by A23,A24,JGRAPH_3:10
           .=sqrt((s2)^2+(1-s2^2)) by A22,SQUARE_1:def 4
           .=1 by SQUARE_1:83,XCMPLX_1:27;
       then p2 in P by A1;
       then p2 in {p3 where p3 is Point of TOP-REAL 2: p3 in P & p3`2>=0}
                          by A24,A25;
       then A26: |[s2,sqrt(1-s2^2)]| in Upper_Arc(P) by A1,Th37;
         dom h=[.0,1.] by BORSUK_1:83,FUNCT_2:def 1;
       then h.r2=f1.s2 by A6,A10,FUNCT_1:22 .=p2 by A2,A23,A26;
      hence q1`1>q2`1 by A6,A11,A19,A23,REAL_1:92;
     end;
     A27: now assume A28: r2>r1;
       set s1=2*r2-1,s2=2*r1-1;
       set p1=|[s1,sqrt(1-s1^2)]|,p2=|[s2,sqrt(1-s2^2)]|;
       A29: g.r2=2*r2-1 by A3,A6;
       A30: g.r1=2*r1-1 by A3,A6;
       A31: 2*r2 > 2*r1 by A28,REAL_1:70;
         r2<=1 by A6,TOPREAL5:1;
       then 2*r2<=2*1 by AXIOMS:25;
       then A32: 2*r2-1<=2*1-1 by REAL_1:49;
         r2>=0 by A6,TOPREAL5:1;
       then 2*r2>=2*0 by AXIOMS:25;
       then 2*r2-1>=2*0-1 by REAL_1:49;
       then -1<=s1;
       then s1^2<=1^2 by A32,JGRAPH_2:7;
       then A33: 1-s1^2>=0 by SQUARE_1:12,59;
       A34: (|[s1,sqrt(1-s1^2)]|)`1=s1 by EUCLID:56;
       A35: (|[s1,sqrt(1-s1^2)]|)`2=sqrt(1-s1^2) by EUCLID:56;
       A36: sqrt(1-s1^2)>=0 by A33,SQUARE_1:def 4;
         |.p1.|=sqrt((s1)^2+(sqrt(1-s1^2))^2) by A34,A35,JGRAPH_3:10
           .=sqrt((s1)^2+(1-s1^2)) by A33,SQUARE_1:def 4
           .=1 by SQUARE_1:83,XCMPLX_1:27;
       then p1 in P by A1;
       then p1 in {p3 where p3 is Point of TOP-REAL 2: p3 in P & p3`2>=0}
                         by A35,A36;
       then A37: |[s1,sqrt(1-s1^2)]| in Upper_Arc(P) by A1,Th37;
         dom h=[.0,1.] by BORSUK_1:83,FUNCT_2:def 1;
       then h.r2=f1.s1 by A6,A29,FUNCT_1:22 .=p1 by A2,A34,A37;
       then A38: q2`1=s1 by A6,EUCLID:56;
         r1<=1 by A6,TOPREAL5:1;
       then 2*r1<=2*1 by AXIOMS:25;
       then A39: 2*r1-1<=2*1-1 by REAL_1:49;
         r1>=0 by A6,TOPREAL5:1;
       then 2*r1>=2*0 by AXIOMS:25;
       then 2*r1-1>=2*0-1 by REAL_1:49;
       then -1<=s2;
       then s2^2<=1^2 by A39,JGRAPH_2:7;
       then A40: 1-s2^2>=0 by SQUARE_1:12,59;
       A41: (|[s2,sqrt(1-s2^2)]|)`1=s2 by EUCLID:56;
       A42: (|[s2,sqrt(1-s2^2)]|)`2=sqrt(1-s2^2) by EUCLID:56;
       A43: sqrt(1-s2^2)>=0 by A40,SQUARE_1:def 4;
         |.p2.|=sqrt((s2)^2+(sqrt(1-s2^2))^2) by A41,A42,JGRAPH_3:10
           .=sqrt((s2)^2+(1-s2^2)) by A40,SQUARE_1:def 4
           .=1 by SQUARE_1:83,XCMPLX_1:27;
       then p2 in P by A1;
       then p2 in {p3 where p3 is Point of TOP-REAL 2: p3 in P & p3`2>=0}
                         by A42,A43;
       then A44: |[s2,sqrt(1-s2^2)]| in Upper_Arc(P) by A1,Th37;
         dom h=[.0,1.] by BORSUK_1:83,FUNCT_2:def 1;
       then h.r1=f1.s2 by A6,A30,FUNCT_1:22 .=p2 by A2,A41,A44;
      hence q2`1>q1`1 by A6,A31,A38,A41,REAL_1:92;
     end;
       now assume A45: q1`1<q2`1;
         now assume A46: r1>=r2;
           now per cases by A46,REAL_1:def 5;
         case r1>r2;
          hence contradiction by A7,A45;
         case r1=r2;
          hence contradiction by A6,A45;
         end;
        hence contradiction;
       end;
      hence r1<r2;
     end;
    hence r1<r2 iff q1`1<q2`1 by A27;
   end;
  A47: dom h=[.0,1.] by BORSUK_1:83,FUNCT_2:def 1;
  then 0 in dom h by TOPREAL5:1;
  then A48: h.0=W-min(P) by A2,A3,FUNCT_1:22;
    1 in dom h by A47,TOPREAL5:1;
  then h.1=E-max(P) by A2,A3,FUNCT_1:22;
 hence thesis by A4,A5,A48;
end;

theorem Th47: for p1,p2 being Point of TOP-REAL 2,
P being compact non empty Subset of TOP-REAL 2
 st P={p where p is Point of TOP-REAL 2: |.p.|=1}
  & p2 in Upper_Arc(P) & LE p1,p2,P
holds p1 in Upper_Arc(P)
proof let p1,p2 be Point of TOP-REAL 2,
  P be compact non empty Subset of TOP-REAL 2;
  assume A1: P={p where p is Point of TOP-REAL 2: |.p.|=1}
   & p2 in Upper_Arc(P) & LE p1,p2,P;
     then A2: P is_simple_closed_curve by JGRAPH_3:36;
     then A3: Lower_Arc(P) is_an_arc_of E-max(P),W-min(P) by JORDAN6:def 9;
   set P4b=Lower_Arc(P);
   A4: P4b is_an_arc_of E-max(P),W-min(P)
     & Upper_Arc(P) /\ P4b={W-min(P),E-max(P)}
       & Upper_Arc(P) \/ P4b=P
         & First_Point(Upper_Arc(P),W-min(P),E-max(P),
         Vertical_Line((W-bound(P)+E-bound(P))/2))`2>
              Last_Point(P4b,E-max(P),W-min(P),
         Vertical_Line((W-bound(P)+E-bound(P))/2))`2
                   by A2,JORDAN6:def 9;
   then E-max(P) in Upper_Arc(P) /\ P4b by TARSKI:def 2;
   then A5: E-max(P) in Upper_Arc(P) by XBOOLE_0:def 3;
   A6: p1 in Upper_Arc(P) & p2 in Lower_Arc(P)& not p2=W-min(P) or
         p1 in Upper_Arc(P) & p2 in Upper_Arc(P) &
         LE p1,p2,Upper_Arc(P),W-min(P),E-max(P) or
         p1 in Lower_Arc(P) & p2 in Lower_Arc(P)& not p2=W-min(P) &
      LE p1,p2,Lower_Arc(P),E-max(P),W-min(P) by A1,JORDAN6:def 10;
    now assume A7: not p1 in Upper_Arc(P);
         then p2 in Upper_Arc(P) /\ P4b by A1,A6,XBOOLE_0:def 3;
         then A8: p2=E-max(P) by A4,A6,A7,TARSKI:def 2;
         then LE p2,p1,Lower_Arc(P),E-max(P),W-min(P) by A3,A6,A7,JORDAN5C:10;
      hence contradiction by A3,A5,A6,A7,A8,JORDAN5C:12;
  end;
 hence p1 in Upper_Arc(P);
end;

theorem Th48: for p1,p2 being Point of TOP-REAL 2,
P being compact non empty Subset of TOP-REAL 2
 st P={p where p is Point of TOP-REAL 2: |.p.|=1}
  & LE p1,p2,P & p1<>p2 & p1`1<0 & p2`1<0 & p1`2<0 & p2`2<0
 holds p1`1>p2`1 & p1`2<p2`2
proof let p1,p2 be Point of TOP-REAL 2,
 P be compact non empty Subset of TOP-REAL 2;
 assume A1: P={p where p is Point of TOP-REAL 2: |.p.|=1}
  & LE p1,p2,P & p1<>p2 & p1`1<0 & p2`1<0 & p1`2<0 & p2`2<0;
  then P is_simple_closed_curve by JGRAPH_3:36;
  then A2: p1 in P & p2 in P by A1,JORDAN7:5;
   A3: Upper_Arc(P)={p where p is Point of TOP-REAL 2:p in P & p`2>=0}
                              by A1,Th37;
     now assume p1 in Upper_Arc(P);
      then consider p being Point of TOP-REAL 2 such that
      A4: p1=p & p in P & p`2>=0 by A3;
     thus contradiction by A1,A4;
    end;
   then A5: p1 in Lower_Arc(P) & p2 in Lower_Arc(P)& not p2=W-min(P) &
      LE p1,p2,Lower_Arc(P),E-max(P),W-min(P)
                           by A1,JORDAN6:def 10;
  consider f being map of I[01],(TOP-REAL 2)|Lower_Arc(P)
  such that
  A6: f is_homeomorphism & (for q1,q2 being Point of TOP-REAL 2,
  r1,r2 being Real st f.r1=q1 & f.r2=q2 &
  r1 in [.0,1.] & r2 in [.0,1.] holds r1<r2 iff q1`1>q2`1)
  & f.0=E-max(P) & f.1=W-min(P) by A1,Th45;
  A7: rng f=[#]((TOP-REAL 2)|Lower_Arc(P)) by A6,TOPS_2:def 5
       .=Lower_Arc(P) by PRE_TOPC:def 10;
  then consider x1 being set such that
  A8: x1 in dom f & p1=f.x1 by A5,FUNCT_1:def 5;
  consider x2 being set such that
  A9: x2 in dom f & p2=f.x2 by A5,A7,FUNCT_1:def 5;
  A10: dom f=[#](I[01]) by A6,TOPS_2:def 5
            .=[.0,1.] by BORSUK_1:83,PRE_TOPC:12;
  then reconsider r11=x1 as Real by A8;
  reconsider r22=x2 as Real by A9,A10;
  A11: 0<=r11 & r11<=1 by A8,A10,TOPREAL5:1;
  A12: 0<=r22 & r22<=1 by A9,A10,TOPREAL5:1;
  A13: r11<r22 iff p1`1>p2`1 by A6,A8,A9,A10;
  A14: r11<=r22 by A5,A6,A8,A9,A11,A12,JORDAN5C:def 3;
  consider p3 being Point of TOP-REAL 2 such that
  A15: p3=p1 & |.p3.|=1 by A1,A2;
    1^2=(p1`1)^2+(p1`2)^2 by A15,JGRAPH_3:10;
  then 1^2-(p1`1)^2=(p1`2)^2 by XCMPLX_1:26.=(-(p1`2))^2
              by SQUARE_1:61;
  then A16: 1^2-(-(p1`1))^2=(-(p1`2))^2 by SQUARE_1:61;
  A17: -(p1`1)>0 by A1,REAL_1:66;
    -(p1`2)>0 by A1,REAL_1:66;
  then -(p1`2)=sqrt(1^2-(-(p1`1))^2) by A16,SQUARE_1:89;
  then A18: (p1`2)=-sqrt(1^2-(-(p1`1))^2);
  consider p4 being Point of TOP-REAL 2 such that
  A19: p4=p2 & |.p4.|=1 by A1,A2;
    1^2=(p2`1)^2+(p2`2)^2 by A19,JGRAPH_3:10;
  then 1^2-(p2`1)^2=(p2`2)^2 by XCMPLX_1:26.=(-(p2`2))^2
              by SQUARE_1:61;
  then A20: 1^2-(-(p2`1))^2=(-(p2`2))^2 by SQUARE_1:61;
    -(p2`2)>0 by A1,REAL_1:66;
  then -(p2`2)=sqrt(1^2-(-(p2`1))^2) by A20,SQUARE_1:89;
  then A21: (p2`2)=-sqrt(1^2-(-(p2`1))^2);
    -(p1`1)< -(p2`1) by A1,A8,A9,A13,A14,REAL_1:50,def 5;
  then (-(p1`1))^2 < (-(p2`1))^2 by A17,SQUARE_1:78;
  then A22: 1^2- (-(p1`1))^2 > 1^2-(-(p2`1))^2 by REAL_1:92;
    1^2-(-(p2`1))^2>=0 by A20,SQUARE_1:72;
  then sqrt(1^2- (-(p1`1))^2) > sqrt(1^2-(-(p2`1))^2) by A22,SQUARE_1:95;
 hence p1`1>p2`1 & p1`2<p2`2 by A8,A9,A13,A14,A18,A21,REAL_1:50,def 5;
end;

theorem Th49: for p1,p2 being Point of TOP-REAL 2,
P being compact non empty Subset of TOP-REAL 2
 st P={p where p is Point of TOP-REAL 2: |.p.|=1}
 & LE p1,p2,P & p1<>p2 & p1`1<0 & p2`1<0 & p1`2>=0 & p2`2>=0
 holds p1`1<p2`1 & p1`2<p2`2
proof let p1,p2 be Point of TOP-REAL 2,
 P be compact non empty Subset of TOP-REAL 2;
 assume A1: P={p where p is Point of TOP-REAL 2: |.p.|=1}
 & LE p1,p2,P & p1<>p2 & p1`1<0 & p2`1<0 & p1`2>=0 & p2`2>=0;
  then A2: P is_simple_closed_curve by JGRAPH_3:36;
  then A3: P={p where p is Point of TOP-REAL 2: |.p.|=1}
    & p1 in P & p2 in P &
  LE p1,p2,P & p1<>p2 & p1`1<0 & p2`1<0 & p1`2>=0 & p2`2>=0 by A1,JORDAN7:5;
   set P4=Lower_Arc(P);
   A4: P4 is_an_arc_of E-max(P),W-min(P)
     & Upper_Arc(P) /\ P4={W-min(P),E-max(P)}
       & Upper_Arc(P) \/ P4=P
         & First_Point(Upper_Arc(P),W-min(P),E-max(P),
         Vertical_Line((W-bound(P)+E-bound(P))/2))`2>
              Last_Point(P4,E-max(P),W-min(P),
         Vertical_Line((W-bound(P)+E-bound(P))/2))`2 by A2,JORDAN6:def 9;
   A5: Upper_Arc(P)={p where p is Point of TOP-REAL 2:p in P & p`2>=0}
                              by A1,Th37;
   A6: now assume p2=W-min(P);
      then LE p2,p1,P by A2,A3,JORDAN7:3;
     hence contradiction by A1,A2,JORDAN6:72;
   end;
     now assume A7: p2 in Lower_Arc(P);
        p2 in Upper_Arc(P) by A3,A5;
      then p2 in {W-min(P),E-max(P)} by A4,A7,XBOOLE_0:def 3;
      then A8: p2=W-min(P) or p2=E-max(P) by TARSKI:def 2;
        E-max(P)=|[1,0]| by A1,Th33;
      then (E-max(P))`1=1 by EUCLID:56;
     hence contradiction by A1,A6,A8;
   end;
   then A9: p1 in Upper_Arc(P) & p2 in Upper_Arc(P)& not p2=W-min(P) &
      LE p1,p2,Upper_Arc(P),W-min(P),E-max(P)
                           by A1,A6,JORDAN6:def 10;
  consider f being map of
     I[01],(TOP-REAL 2)|Upper_Arc(P)
  such that
  A10: f is_homeomorphism & (for q1,q2 being Point of TOP-REAL 2,
  r1,r2 being Real st f.r1=q1 & f.r2=q2 &
  r1 in [.0,1.] & r2 in [.0,1.] holds r1<r2 iff q1`1<q2`1)
  & f.0=W-min(P) & f.1=E-max(P) by A1,Th46;
  A11: rng f=[#]((TOP-REAL 2)|Upper_Arc(P)) by A10,TOPS_2:def 5
       .=Upper_Arc(P) by PRE_TOPC:def 10;
  then consider x1 being set such that
  A12: x1 in dom f & p1=f.x1 by A9,FUNCT_1:def 5;
  consider x2 being set such that
  A13: x2 in dom f & p2=f.x2 by A9,A11,FUNCT_1:def 5;
  A14: dom f=[#](I[01]) by A10,TOPS_2:def 5
            .=[.0,1.] by BORSUK_1:83,PRE_TOPC:12;
  then reconsider r11=x1 as Real by A12;
  reconsider r22=x2 as Real by A13,A14;
  A15: 0<=r11 & r11<=1 by A12,A14,TOPREAL5:1;
  A16: 0<=r22 & r22<=1 by A13,A14,TOPREAL5:1;
  A17: r11<r22 iff p1`1<p2`1 by A10,A12,A13,A14;
  A18: r11<=r22 by A9,A10,A12,A13,A15,A16,JORDAN5C:def 3;
  consider p3 being Point of TOP-REAL 2 such that
  A19: p3=p1 & |.p3.|=1 by A3;
    1^2=(p1`1)^2+(p1`2)^2 by A19,JGRAPH_3:10;
  then 1^2-(p1`1)^2=(p1`2)^2 by XCMPLX_1:26;
  then A20: 1^2-(-(p1`1))^2=((p1`2))^2 by SQUARE_1:61;
  A21: -(p2`1)>0 by A1,REAL_1:66;
  A22: (p1`2)=sqrt(1^2-(-(p1`1))^2) by A1,A20,SQUARE_1:89;
  consider p4 being Point of TOP-REAL 2 such that
  A23: p4=p2 & |.p4.|=1 by A3;
    1^2=(p2`1)^2+(p2`2)^2 by A23,JGRAPH_3:10;
  then 1^2-(p2`1)^2=(p2`2)^2 by XCMPLX_1:26;
  then 1^2-(-(p2`1))^2=((p2`2))^2 by SQUARE_1:61;
  then A24: (p2`2)=sqrt(1^2-(-(p2`1))^2) by A1,SQUARE_1:89;
    -(p1`1)> -(p2`1) by A1,A12,A13,A17,A18,REAL_1:50,def 5;
  then (-(p1`1))^2 > (-(p2`1))^2 by A21,SQUARE_1:78;
  then A25: 1^2- (-(p1`1))^2 < 1^2-(-(p2`1))^2 by REAL_1:92;
    1^2=(p1`1)^2+(p1`2)^2 by A19,JGRAPH_3:10;
  then 1^2-(p1`1)^2=(p1`2)^2 by XCMPLX_1:26;
  then 1^2-(-(p1`1))^2=((p1`2))^2 by SQUARE_1:61;
  then 1^2-(-(p1`1))^2>=0 by SQUARE_1:72;
 hence p1`1<p2`1 & p1`2<p2`2 by A12,A13,A17,A18,A22,A24,A25,REAL_1:def 5,
SQUARE_1:95;
end;

theorem Th50: for p1,p2 being Point of TOP-REAL 2,
P being compact non empty Subset of TOP-REAL 2
 st P={p where p is Point of TOP-REAL 2: |.p.|=1}
 & LE p1,p2,P & p1<>p2 & p1`2>=0 & p2`2>=0
 holds p1`1<p2`1
proof let p1,p2 be Point of TOP-REAL 2,
 P be compact non empty Subset of TOP-REAL 2;
 assume A1: P={p where p is Point of TOP-REAL 2: |.p.|=1}
 & LE p1,p2,P & p1<>p2 & p1`2>=0 & p2`2>=0;
  then A2: P is_simple_closed_curve by JGRAPH_3:36;
  then A3: P={p where p is Point of TOP-REAL 2: |.p.|=1}
   & p1 in P & p2 in P &
   LE p1,p2,P & p1<>p2 & p1`2>=0 & p2`2>=0 by A1,JORDAN7:5;
   set P4=Lower_Arc(P);
   A4: P4 is_an_arc_of E-max(P),W-min(P)
     & Upper_Arc(P) /\ P4={W-min(P),E-max(P)}
       & Upper_Arc(P) \/ P4=P
         & First_Point(Upper_Arc(P),W-min(P),E-max(P),
         Vertical_Line((W-bound(P)+E-bound(P))/2))`2>
              Last_Point(P4,E-max(P),W-min(P),
         Vertical_Line((W-bound(P)+E-bound(P))/2))`2 by A2,JORDAN6:def 9;
   A5: Upper_Arc(P)={p where p is Point of TOP-REAL 2:p in P & p`2>=0}
                              by A1,Th37;
   A6: now assume p2=W-min(P);
      then LE p2,p1,P by A2,A3,JORDAN7:3;
     hence contradiction by A1,A2,JORDAN6:72;
   end;
     now assume A7: p2 in Lower_Arc(P);
        p2 in Upper_Arc(P) by A3,A5;
      then p2 in {W-min(P),E-max(P)} by A4,A7,XBOOLE_0:def 3;
      then A8: p2=W-min(P) or p2=E-max(P) by TARSKI:def 2;
      consider p8 being Point of TOP-REAL 2 such that
      A9: p8=p1 & |.p8.|=1 by A3;
        now assume p2=W-min(P);
        then LE p2,p1,P by A2,A3,JORDAN7:3;
       hence contradiction by A1,A2,JORDAN6:72;
      end;
      then A10: p2= |[1,0]| by A1,A8,Th33;
      then A11: p2`1=1 by EUCLID:56;
      A12: now assume A13: p1`1=1;
        1^2=(p1`1)^2+(p1`2)^2 by A9,JGRAPH_3:10;
      then 1^2-(p1`1)^2=(p1`2)^2 by XCMPLX_1:26;
      then 0=((p1`2))^2 by A13,XCMPLX_1:14;
        then p1`2=0 by SQUARE_1:73;
       hence contradiction by A1,A10,A13,EUCLID:57;
      end;
        p1`1<=1 by A9,Th3;
     hence p1`1<p2`1 by A11,A12,REAL_1:def 5;
   end;
   then A14: p1 in Upper_Arc(P) & p2 in Upper_Arc(P)& not p2=W-min(P) &
      LE p1,p2,Upper_Arc(P),W-min(P),E-max(P) or p1`1<p2`1
                           by A1,A6,JORDAN6:def 10;
  consider f being map of I[01],(TOP-REAL 2)|Upper_Arc(P)
  such that
  A15: f is_homeomorphism & (for q1,q2 being Point of TOP-REAL 2,
  r1,r2 being Real st f.r1=q1 & f.r2=q2 &
  r1 in [.0,1.] & r2 in [.0,1.] holds r1<r2 iff q1`1<q2`1)
  & f.0=W-min(P) & f.1=E-max(P) by A1,Th46;
  A16: rng f=[#]((TOP-REAL 2)|Upper_Arc(P)) by A15,TOPS_2:def 5
       .=Upper_Arc(P) by PRE_TOPC:def 10;
   now per cases;
 case A17: not p1`1 < p2`1;
  then consider x1 being set such that
  A18: x1 in dom f & p1=f.x1 by A14,A16,FUNCT_1:def 5;
  consider x2 being set such that
  A19: x2 in dom f & p2=f.x2 by A14,A16,A17,FUNCT_1:def 5;
  A20: dom f=[#](I[01]) by A15,TOPS_2:def 5
            .=[.0,1.] by BORSUK_1:83,PRE_TOPC:12;
  then reconsider r11=x1 as Real by A18;
  reconsider r22=x2 as Real by A19,A20;
  A21: 0<=r11 & r11<=1 by A18,A20,TOPREAL5:1;
  A22: 0<=r22 & r22<=1 by A19,A20,TOPREAL5:1;
  A23: r11<r22 iff p1`1<p2`1 by A15,A18,A19,A20;
    r11<=r22 or p1`1<p2`1 by A14,A15,A18,A19,A21,A22,JORDAN5C:def 3;
  hence p1`1<p2`1 by A1,A18,A19,A23,REAL_1:def 5;
 case p1`1<p2`1;
  hence p1`1<p2`1;
 end;
 hence p1`1<p2`1;
end;

theorem Th51: for p1,p2 being Point of TOP-REAL 2,
P being compact non empty Subset of TOP-REAL 2
 st P={p where p is Point of TOP-REAL 2: |.p.|=1}
 & LE p1,p2,P & p1<>p2 & p1`2<=0 & p2`2<=0 & p1<>W-min(P)
 holds p1`1>p2`1
proof let p1,p2 be Point of TOP-REAL 2,
 P be compact non empty Subset of TOP-REAL 2;
 assume A1: P={p where p is Point of TOP-REAL 2: |.p.|=1}
 & LE p1,p2,P & p1<>p2 & p1`2<=0 & p2`2<=0 & p1<>W-min(P);
  then A2: P is_simple_closed_curve by JGRAPH_3:36;
  then A3: P={p where p is Point of TOP-REAL 2: |.p.|=1}
   & p1 in P & p2 in P &
  LE p1,p2,P & p1<>p2 & p1`2<=0 & p2`2<=0 & p1<>W-min(P) by A1,JORDAN7:5;
   set P4=Lower_Arc(P);
   A4: P4 is_an_arc_of E-max(P),W-min(P)
     & Upper_Arc(P) /\ P4={W-min(P),E-max(P)}
       & Upper_Arc(P) \/ P4=P
         & First_Point(Upper_Arc(P),W-min(P),E-max(P),
         Vertical_Line((W-bound(P)+E-bound(P))/2))`2>
              Last_Point(P4,E-max(P),W-min(P),
         Vertical_Line((W-bound(P)+E-bound(P))/2))`2 by A2,JORDAN6:def 9;
   A5: Lower_Arc(P)={p where p is Point of TOP-REAL 2:p in P & p`2<=0}
                              by A1,Th38;
     now assume A6: p1 in Upper_Arc(P);
        p1 in Lower_Arc(P) by A3,A5;
      then p1 in {W-min(P),E-max(P)} by A4,A6,XBOOLE_0:def 3;
      then A7: p1=W-min(P) or p1=E-max(P) by TARSKI:def 2;
      consider p9 being Point of TOP-REAL 2 such that
      A8: p9=p2 & |.p9.|=1 by A3;
      A9: p1= |[1,0]| by A1,A7,Th33;
      then A10: p1`1=1 by EUCLID:56;
      A11: now assume A12: p2`1=1;
        1^2 =(p2`1)^2+(p2`2)^2 by A8,JGRAPH_3:10;
      then (1)^2-(p2`1)^2=(p2`2)^2 by XCMPLX_1:26;
       then 0=((p2`2))^2 by A12,XCMPLX_1:14;
        then p2`2=0 by SQUARE_1:73;
       hence contradiction by A1,A9,A12,EUCLID:57;
      end;
        p2`1<=1 by A8,Th3;
     hence p1`1>p2`1 by A10,A11,REAL_1:def 5;
   end;
   then A13: p1 in Lower_Arc(P) & p2 in Lower_Arc(P)& not p2=W-min(P) &
      LE p1,p2,Lower_Arc(P),E-max(P),W-min(P) or p1`1>p2`1
                           by A1,JORDAN6:def 10;
  consider f being map of I[01],(TOP-REAL 2)|Lower_Arc(P) such that
  A14: f is_homeomorphism & (for q1,q2 being Point of TOP-REAL 2,
  r1,r2 being Real st f.r1=q1 & f.r2=q2 &
  r1 in [.0,1.] & r2 in [.0,1.] holds r1<r2 iff q1`1>q2`1)
  & f.0=E-max(P) & f.1=W-min(P) by A1,Th45;
  A15: rng f=[#]((TOP-REAL 2)|Lower_Arc(P)) by A14,TOPS_2:def 5
       .=Lower_Arc(P) by PRE_TOPC:def 10;
   now per cases;
 case A16: not p1`1 > p2`1;
  then consider x1 being set such that
  A17: x1 in dom f & p1=f.x1 by A13,A15,FUNCT_1:def 5;
  consider x2 being set such that
  A18: x2 in dom f & p2=f.x2 by A13,A15,A16,FUNCT_1:def 5;
  A19: dom f=[#](I[01]) by A14,TOPS_2:def 5
            .=[.0,1.] by BORSUK_1:83,PRE_TOPC:12;
  then reconsider r11=x1 as Real by A17;
  reconsider r22=x2 as Real by A18,A19;
  A20: 0<=r11 & r11<=1 by A17,A19,TOPREAL5:1;
  A21: 0<=r22 & r22<=1 by A18,A19,TOPREAL5:1;
  A22: r11<r22 iff p1`1>p2`1 by A14,A17,A18,A19;
    r11<=r22 or p1`1>p2`1 by A13,A14,A17,A18,A20,A21,JORDAN5C:def 3;
  hence p1`1>p2`1 by A1,A17,A18,A22,REAL_1:def 5;
 case p1`1>p2`1;
  hence p1`1>p2`1;
 end;
 hence p1`1>p2`1;
end;

theorem Th52: for p1,p2 being Point of TOP-REAL 2,
P being compact non empty Subset of TOP-REAL 2
 st P={p where p is Point of TOP-REAL 2: |.p.|=1}
  & (p2`2>=0 or p2`1>=0) & LE p1,p2,P
holds p1`2>=0 or p1`1>=0
proof let p1,p2 be Point of TOP-REAL 2,
  P be compact non empty Subset of TOP-REAL 2;
  assume A1: P={p where p is Point of TOP-REAL 2: |.p.|=1}
     & (p2`2>=0 or p2`1>=0) & LE p1,p2,P;
     then A2: P is_simple_closed_curve by JGRAPH_3:36;
     then A3: Lower_Arc(P) is_an_arc_of E-max(P),W-min(P) by JORDAN6:def 9;
     A4: P={p where p is Point of TOP-REAL 2: |.p.|=1}
     & p1 in P & p2 in P
     & (p2`2>=0 or p2`1>=0) & LE p1,p2,P by A1,A2,JORDAN7:5;
   A5: Upper_Arc(P)={p where p is Point of TOP-REAL 2:p in P & p`2>=0}
                              by A1,Th37;
   A6: Lower_Arc(P)={p where p is Point of TOP-REAL 2:p in P & p`2<=0}
                              by A1,Th38;
    per cases by A1;
    suppose p2`2>=0;
      then p2 in Upper_Arc(P) by A4,A5;
      then p1 in Upper_Arc(P) by A1,Th47;
      then consider p8 being Point of TOP-REAL 2 such that
      A7: p8=p1 & p8 in P & p8`2>=0 by A5;
     thus p1`2>=0 or p1`1>=0 by A7;
    suppose A8: p2`2<0 & p2`1>=0;
      then not ex p8 being Point of TOP-REAL 2 st p8=p2 & p8 in P & p8`2>=0;
      then A9: not p2 in Upper_Arc(P) by A5;
           now per cases by A1,A9,JORDAN6:def 10;
         case p1 in Upper_Arc(P) & p2 in Lower_Arc(P)& not p2=W-min(P);
           then consider p8 being Point of TOP-REAL 2 such that
           A10: p8=p1 & p8 in P & p8`2>=0 by A5;
          thus p1`2>=0 or p1`1>=0 by A10;
         case A11: p1 in Lower_Arc(P) & p2 in Lower_Arc(P)& not p2=W-min(P)&
           LE p1,p2,Lower_Arc(P),E-max(P),W-min(P);
           then consider p8 being Point of TOP-REAL 2 such that
           A12: p8=p1 & p8 in P & p8`2<=0 by A6;
             now assume A13: p1=W-min(P);
             then LE p2,p1,Lower_Arc(P),E-max(P),W-min(P) by A3,A11,JORDAN5C:10
;
            hence contradiction by A3,A11,A13,JORDAN5C:12;
           end;
           then p1`1>=p2`1 by A1,A8,A12,Th51;
          hence p1`2>=0 or p1`1>=0 by A8;
         end;
     hence p1`2>=0 or p1`1>=0;
end;

theorem Th53: for p1,p2 being Point of TOP-REAL 2,
P being compact non empty Subset of TOP-REAL 2
 st P={p where p is Point of TOP-REAL 2: |.p.|=1}
 & LE p1,p2,P & p1<>p2 & p1`1>=0 & p2`1>=0
 holds p1`2>p2`2
proof let p1,p2 be Point of TOP-REAL 2,
 P be compact non empty Subset of TOP-REAL 2;
 assume A1: P={p where p is Point of TOP-REAL 2: |.p.|=1}
 & LE p1,p2,P & p1<>p2 & p1`1>=0 & p2`1>=0;
  then A2: P is_simple_closed_curve by JGRAPH_3:36;
  then A3: P={p where p is Point of TOP-REAL 2: |.p.|=1} & p1 in P & p2 in P &
  LE p1,p2,P & p1<>p2 & p1`1>=0 & p2`1>=0 by A1,JORDAN7:5;
  then consider p4 being Point of TOP-REAL 2 such that
  A4: p4=p1 & |.p4.|=1;
  consider p3 being Point of TOP-REAL 2 such that
  A5: p3=p2 & |.p3.|=1 by A3;
   A6: Upper_Arc(P)={p where p is Point of TOP-REAL 2:p in P & p`2>=0}
                              by A1,Th37;
   A7: Lower_Arc(P)={p where p is Point of TOP-REAL 2:p in P & p`2<=0}
                              by A1,Th38;
     W-min(P)=|[-1,0]| by A1,Th32;
   then A8:(W-min(P))`2=0 by EUCLID:56;
   now per cases;
 case A9: p1`2>=0 & p2`2>=0;
   then p1`1<p2`1 by A1,Th50;
  then (p1`1)^2 < ((p2`1))^2 by A1,SQUARE_1:78;
  then A10: 1^2- ((p1`1))^2 > 1^2-((p2`1))^2 by REAL_1:92;
    1^2=(p1`1)^2+(p1`2)^2 by A4,JGRAPH_3:10;
  then 1^2-((p1`1))^2=((p1`2))^2 by XCMPLX_1:26;
  then A11: p1`2=sqrt(1^2-((p1`1))^2) by A9,SQUARE_1:89;
    1^2=(p2`1)^2+(p2`2)^2 by A5,JGRAPH_3:10;
  then A12: 1^2-((p2`1))^2=((p2`2))^2 by XCMPLX_1:26;
  then A13: (p2`2)=sqrt(1^2-((p2`1))^2) by A9,SQUARE_1:89;
    1^2-((p2`1))^2>=0 by A12,SQUARE_1:72;
  hence p1`2>p2`2 by A10,A11,A13,SQUARE_1:95;
 case p1`2>=0 & p2`2<0;
  hence p1`2>p2`2;
 case A14: p1`2<0 & p2`2>=0;
   then A15: p1 in Lower_Arc(P) by A3,A7;
     p2 in Upper_Arc(P) by A3,A6,A14;
   then LE p2,p1,P by A8,A14,A15,JORDAN6:def 10;
  hence contradiction by A1,A2,JORDAN6:72;
 case A16: p1`2<0 & p2`2<0;
   then not ex p being Point of TOP-REAL 2 st p=p1 & p in P & p`2>=0;
   then not p1 in Upper_Arc(P) by A6;
   then A17: p1 in Lower_Arc(P) & p2 in Lower_Arc(P)& not p2=W-min(P) &
       LE p1,p2,Lower_Arc(P),E-max(P),W-min(P) by A1,JORDAN6:def 10;
  consider f being map of I[01],(TOP-REAL 2)|Lower_Arc(P) such that
  A18: f is_homeomorphism & (for q1,q2 being Point of TOP-REAL 2,
  r1,r2 being Real st f.r1=q1 & f.r2=q2 &
  r1 in [.0,1.] & r2 in [.0,1.] holds r1<r2 iff q1`1>q2`1)
  & f.0=E-max(P) & f.1=W-min(P) by A1,Th45;
  A19: rng f=[#]((TOP-REAL 2)|Lower_Arc(P)) by A18,TOPS_2:def 5
       .=Lower_Arc(P) by PRE_TOPC:def 10;
  then consider x1 being set such that
  A20: x1 in dom f & p1=f.x1 by A17,FUNCT_1:def 5;
  consider x2 being set such that
  A21: x2 in dom f & p2=f.x2 by A17,A19,FUNCT_1:def 5;
  A22: dom f=[#](I[01]) by A18,TOPS_2:def 5
            .=[.0,1.] by BORSUK_1:83,PRE_TOPC:12;
  then reconsider r11=x1 as Real by A20;
  reconsider r22=x2 as Real by A21,A22;
  A23: 0<=r11 & r11<=1 by A20,A22,TOPREAL5:1;
  A24: 0<=r22 & r22<=1 by A21,A22,TOPREAL5:1;
  A25: r11<r22 iff p1`1>p2`1 by A18,A20,A21,A22;
  A26: r11<=r22 by A17,A18,A20,A21,A23,A24,JORDAN5C:def 3;
  consider p3 being Point of TOP-REAL 2 such that
  A27: p3=p1 & |.p3.|=1 by A3;
    1^2=(p1`1)^2+(p1`2)^2 by A27,JGRAPH_3:10;
  then A28: 1^2-(p1`1)^2=(p1`2)^2 by XCMPLX_1:26;
  then A29: 1^2-((p1`1))^2=(-(p1`2))^2 by SQUARE_1:61;
    -(p1`2)>0 by A16,REAL_1:66;
  then A30: -(p1`2)=sqrt(1^2-((p1`1))^2) by A29,SQUARE_1:89;
  consider p4 being Point of TOP-REAL 2 such that
  A31: p4=p2 & |.p4.|=1 by A3;
    1^2=(p2`1)^2+(p2`2)^2 by A31,JGRAPH_3:10;
  then 1^2-(p2`1)^2=(p2`2)^2 by XCMPLX_1:26;
  then A32: 1^2-((p2`1))^2=(-(p2`2))^2 by SQUARE_1:61;
    -(p2`2)>0 by A16,REAL_1:66;
  then A33: -(p2`2)=sqrt(1^2-((p2`1))^2) by A32,SQUARE_1:89;
    ((p1`1))^2 > ((p2`1))^2 by A1,A20,A21,A25,A26,REAL_1:def 5,SQUARE_1:78;
  then A34: 1^2- ((p1`1))^2 < 1^2-((p2`1))^2 by REAL_1:92;
    1^2-((p1`1))^2>=0 by A28,SQUARE_1:72;
  then sqrt(1^2- ((p1`1))^2) < sqrt(1^2-((p2`1))^2) by A34,SQUARE_1:95;
 hence p1`2>p2`2 by A30,A33,REAL_1:50;
 end;
 hence p1`2>p2`2;
end;

theorem Th54: for p1,p2 being Point of TOP-REAL 2,
 P being compact non empty Subset of TOP-REAL 2
 st P={p where p is Point of TOP-REAL 2: |.p.|=1}
 & p1 in P & p2 in P
  & p1`1<0 & p2`1<0 & p1`2<0 & p2`2<0 &
 (p1`1>=p2`1 or p1`2<=p2`2) holds LE p1,p2,P
proof let p1,p2 be Point of TOP-REAL 2,
 P be compact non empty Subset of TOP-REAL 2;
 assume A1: P={p where p is Point of TOP-REAL 2: |.p.|=1}
  & p1 in P & p2 in P & p1`1<0 & p2`1<0 & p1`2<0 & p2`2<0 &
  (p1`1>=p2`1 or p1`2<=p2`2);
  then consider p3 being Point of TOP-REAL 2 such that
  A2: p3=p1 & |.p3.|=1;
  consider p3 being Point of TOP-REAL 2 such that
  A3: p3=p2 & |.p3.|=1 by A1;
  A4: -p2`2>0 by A1,REAL_1:66;
  A5: now assume p1`2<=p2`2;
  then -p1`2>=-p2`2 by REAL_1:50;
  then (-(p1`2))^2 >= (-(p2`2))^2 by A4,SQUARE_1:77;
  then A6: 1^2- (-(p1`2))^2 <= 1^2-(-(p2`2))^2 by REAL_1:92;
    1^2=(p1`1)^2+(p1`2)^2 by A2,JGRAPH_3:10;
  then 1^2-(p1`2)^2=(p1`1)^2 by XCMPLX_1:26;
  then A7: 1^2-(-(p1`2))^2=((p1`1))^2 by SQUARE_1:61;
  then A8: 1^2-(-(p1`2))^2=(-(p1`1))^2 by SQUARE_1:61;
    -(p1`1)>=0 by A1,REAL_1:66;
  then A9: -(p1`1)=sqrt(1^2-(-(p1`2))^2) by A8,SQUARE_1:89;
    1^2=(p2`1)^2+(p2`2)^2 by A3,JGRAPH_3:10;
  then 1^2-(p2`2)^2=(p2`1)^2 by XCMPLX_1:26;
  then 1^2-(-(p2`2))^2=((p2`1))^2 by SQUARE_1:61;
  then A10: 1^2-(-(p2`2))^2=(-(p2`1))^2 by SQUARE_1:61;
    -(p2`1)>=0 by A1,REAL_1:66;
  then A11: -(p2`1)=sqrt(1^2-(-(p2`2))^2) by A10,SQUARE_1:89;
    1^2-(-(p1`2))^2>=0 by A7,SQUARE_1:72;
  then sqrt(1^2- (-(p1`2))^2) <= sqrt(1^2-(-(p2`2))^2) by A6,SQUARE_1:94;
   hence p1`1>=p2`1 by A9,A11,REAL_1:50;
  end;
  A12: P is_simple_closed_curve by A1,JGRAPH_3:36;
  then A13: Lower_Arc(P) is_an_arc_of E-max(P),W-min(P) by JORDAN6:def 9;
   set P4=Lower_Arc(P);
   A14: P4 is_an_arc_of E-max(P),W-min(P)
     & Upper_Arc(P) /\ P4={W-min(P),E-max(P)}
       & Upper_Arc(P) \/ P4=P
         & First_Point(Upper_Arc(P),W-min(P),E-max(P),
         Vertical_Line((W-bound(P)+E-bound(P))/2))`2>
              Last_Point(P4,E-max(P),W-min(P),
         Vertical_Line((W-bound(P)+E-bound(P))/2))`2 by A12,JORDAN6:def 9;
   A15: Upper_Arc(P)={p where p is Point of TOP-REAL 2:p in P & p`2>=0}
                              by A1,Th37;
   A16: now assume not p1 in Lower_Arc(P);
      then p1 in Upper_Arc(P) by A1,A14,XBOOLE_0:def 2;
      then consider p being Point of TOP-REAL 2 such that
      A17: p1=p & p in P & p`2>=0 by A15;
     thus contradiction by A1,A17;
    end;
   A18: now assume not p2 in Lower_Arc(P);
      then p2 in Upper_Arc(P) by A1,A14,XBOOLE_0:def 2;
      then consider p being Point of TOP-REAL 2 such that
      A19: p2=p & p in P & p`2>=0 by A15;
     thus contradiction by A1,A19;
    end;
   A20: W-min(P)=|[-1,0]| by A1,Th32;
   A21: E-max(P)=|[1,0]| by A1,Th33;
   A22: now assume A23: p2=W-min(P);
        W-min(P)=|[-1,0]| by A1,Th32;
    hence contradiction by A1,A23,EUCLID:56;
   end;
     for g being map of I[01], (TOP-REAL 2)|P4,
       s1, s2 being Real st g is_homeomorphism
         & g.0 = E-max(P) & g.1 = W-min(P)
         & 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)|P4,
       s1, s2 be Real;
     assume A24: g is_homeomorphism
         & g.0 = E-max(P) & g.1 =W-min(P)
         & g.s1 = p1 & 0 <= s1 & s1 <= 1
         & g.s2 = p2 & 0 <= s2 & s2 <= 1;
   then A25: dom g=[#](I[01]) by TOPS_2:def 5
   .=[.0,1.] by BORSUK_1:83,PRE_TOPC:12;
   reconsider g0=proj1 as map of TOP-REAL 2,R^1 by JGRAPH_2:16;
   set K0=Lower_Arc(P);
   reconsider g2=g0|K0 as map of (TOP-REAL 2)|K0,R^1 by JGRAPH_3:12;
   reconsider g3=g2 as continuous map of (TOP-REAL 2)|K0,
                Closed-Interval-TSpace(-1,1) by A1,Lm4;
     E-max(P) in {W-min(P),E-max(P)} by TARSKI:def 2;
   then A26: E-max(P) in Lower_Arc(P) by A14,XBOOLE_0:def 3;
     W-min(P) in {W-min(P),E-max(P)} by TARSKI:def 2;
   then A27: W-min(P) in Lower_Arc(P) by A14,XBOOLE_0:def 3;
     dom g3=the carrier of ((TOP-REAL 2)|K0) by FUNCT_2:def 1;
   then A28: dom g3=[#]((TOP-REAL 2)|K0) by PRE_TOPC:12;
   A29: rng g3=[#](Closed-Interval-TSpace(-1,1)) by A1,Lm4;
   A30: g3 is one-to-one by A1,Lm4;
        K0 is non empty compact by A13,JORDAN5A:1;
      then A31: (TOP-REAL 2)|K0 is compact by COMPTS_1:12;
        Closed-Interval-TSpace(-1,1)
       =TopSpaceMetr(Closed-Interval-MSpace(-1,1)) by TOPMETR:def 8;
      then Closed-Interval-TSpace(-1,1) is_T2 by PCOMPS_1:38;
      then A32: g3 is_homeomorphism by A28,A29,A30,A31,COMPTS_1:26;
      reconsider h=g3*g as map
          of Closed-Interval-TSpace(0,1),Closed-Interval-TSpace(-1,1)
                  by TOPMETR:27;
      A33: h is_homeomorphism by A24,A32,TOPMETR:27,TOPS_2:71;
      A34: 0 in dom g by A25,TOPREAL5:1;
      A35: 1 in dom g by A25,TOPREAL5:1;
      A36: s1 in [.0,1.] by A24,TOPREAL5:1;
      A37: s2 in [.0,1.] by A24,TOPREAL5:1;
      A38: -1=(|[-1,0]|)`1 by EUCLID:56.=proj1.(|[-1,0]|) by PSCOMP_1:def 28
      .=g3.(g.1) by A20,A24,A27,FUNCT_1:72
      .= h.1 by A35,FUNCT_1:23;
      A39: 1=(|[1,0]|)`1 by EUCLID:56.=g0.(|[1,0]|) by PSCOMP_1:def 28
      .=g3.(|[1,0]|) by A21,A26,FUNCT_1:72 .= h.0 by A21,A24,A34,FUNCT_1:23;
      A40: p1`1=g0.p1 by PSCOMP_1:def 28
      .=g3.(g.s1) by A16,A24,FUNCT_1:72
      .= h.s1 by A25,A36,FUNCT_1:23;
        p2`1=proj1.p2 by PSCOMP_1:def 28
      .=g3.(g.s2) by A18,A24,FUNCT_1:72
      .= h.s2 by A25,A37,FUNCT_1:23;
     hence s1 <= s2 by A1,A5,A33,A36,A37,A38,A39,A40,Th12;
    end;
    then p1 in Lower_Arc(P) & p2 in Lower_Arc(P) & not p2=W-min(P) &
       LE p1,p2,Lower_Arc(P),E-max(P),W-min(P)
                       by A16,A18,A22,JORDAN5C:def 3;
 hence LE p1,p2,P by JORDAN6:def 10;
end;

theorem for p1,p2 being Point of TOP-REAL 2,
 P being compact non empty Subset of TOP-REAL 2
 st P={p where p is Point of TOP-REAL 2: |.p.|=1}
 & p1 in P & p2 in P
 & p1`1>0 & p2`1>0 & p1`2<0 & p2`2<0 &
 (p1`1>=p2`1 or p1`2>=p2`2) holds LE p1,p2,P
proof let p1,p2 be Point of TOP-REAL 2,
 P be compact non empty Subset of TOP-REAL 2;
 assume A1: P={p where p is Point of TOP-REAL 2: |.p.|=1}
  & p1 in P & p2 in P
  & p1`1>0 & p2`1>0 & p1`2<0 & p2`2<0 &
  (p1`1>=p2`1 or p1`2>=p2`2);
  then consider p3 being Point of TOP-REAL 2 such that
  A2: p3=p1 & |.p3.|=1;
  consider p3 being Point of TOP-REAL 2 such that
  A3: p3=p2 & |.p3.|=1 by A1;
  A4: -p1`2>0 by A1,REAL_1:66;
  A5: now assume p1`2>=p2`2;
  then -p1`2<=-p2`2 by REAL_1:50;
  then (-(p1`2))^2 <= (-(p2`2))^2 by A4,SQUARE_1:77;
  then A6: 1^2- (-(p1`2))^2 >= 1^2-(-(p2`2))^2 by REAL_1:92;
    1^2=(p1`1)^2+(p1`2)^2 by A2,JGRAPH_3:10;
  then 1^2-(p1`2)^2=(p1`1)^2 by XCMPLX_1:26;
  then 1^2-(-(p1`2))^2=((p1`1))^2 by SQUARE_1:61;
  then A7: (p1`1)=sqrt(1^2-(-(p1`2))^2) by A1,SQUARE_1:89;
    1^2=(p2`1)^2+(p2`2)^2 by A3,JGRAPH_3:10;
  then 1^2-(p2`2)^2=(p2`1)^2 by XCMPLX_1:26;
  then A8: 1^2-(-(p2`2))^2=((p2`1))^2 by SQUARE_1:61;
  then A9: (p2`1)=sqrt(1^2-(-(p2`2))^2) by A1,SQUARE_1:89;
    1^2-(-(p2`2))^2>=0 by A8,SQUARE_1:72;
  hence p1`1>=p2`1 by A6,A7,A9,SQUARE_1:94;
  end;
  A10: P is_simple_closed_curve by A1,JGRAPH_3:36;
  then A11: Lower_Arc(P) is_an_arc_of E-max(P),W-min(P) by JORDAN6:def 9;
   set P4=Lower_Arc(P);
   A12: P4 is_an_arc_of E-max(P),W-min(P)
     & Upper_Arc(P) /\ P4={W-min(P),E-max(P)}
       & Upper_Arc(P) \/ P4=P
         & First_Point(Upper_Arc(P),W-min(P),E-max(P),
         Vertical_Line((W-bound(P)+E-bound(P))/2))`2>
              Last_Point(P4,E-max(P),W-min(P),
         Vertical_Line((W-bound(P)+E-bound(P))/2))`2 by A10,JORDAN6:def 9;
   A13: Upper_Arc(P)={p where p is Point of TOP-REAL 2:p in P & p`2>=0}
                              by A1,Th37;
   A14: now assume not p1 in Lower_Arc(P);
      then p1 in Upper_Arc(P) by A1,A12,XBOOLE_0:def 2;
      then consider p being Point of TOP-REAL 2 such that
      A15: p1=p & p in P & p`2>=0 by A13;
     thus contradiction by A1,A15;
    end;
   A16: now assume not p2 in Lower_Arc(P);
      then p2 in Upper_Arc(P) by A1,A12,XBOOLE_0:def 2;
      then consider p being Point of TOP-REAL 2 such that
      A17: p2=p & p in P & p`2>=0 by A13;
     thus contradiction by A1,A17;
    end;
   A18: W-min(P)=|[-1,0]| by A1,Th32;
   A19: E-max(P)=|[1,0]| by A1,Th33;
   A20: now assume A21: p2=W-min(P);
        W-min(P)=|[-1,0]| by A1,Th32;
    hence contradiction by A1,A21,EUCLID:56;
   end;
     for g being map of I[01], (TOP-REAL 2)|P4,
       s1, s2 being Real st
       g is_homeomorphism
         & g.0 = E-max(P) & g.1 = W-min(P)
         & 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)|P4,
       s1, s2 be Real;
     assume A22: g is_homeomorphism
         & g.0 = E-max(P) & g.1 =W-min(P)
         & g.s1 = p1 & 0 <= s1 & s1 <= 1
         & g.s2 = p2 & 0 <= s2 & s2 <= 1;
   then A23: dom g=[#](I[01]) by TOPS_2:def 5
   .=[.0,1.] by BORSUK_1:83,PRE_TOPC:12;
   reconsider g0=proj1 as map of TOP-REAL 2,R^1 by JGRAPH_2:16;
   set K0=Lower_Arc(P);
   reconsider g2=g0|K0 as map of (TOP-REAL 2)|K0,R^1 by JGRAPH_3:12;
   reconsider g3=g2 as continuous map of (TOP-REAL 2)|K0,
                Closed-Interval-TSpace(-1,1) by A1,Lm4;
     E-max(P) in {W-min(P),E-max(P)} by TARSKI:def 2;
   then A24: E-max(P) in Lower_Arc(P) by A12,XBOOLE_0:def 3;
     W-min(P) in {W-min(P),E-max(P)} by TARSKI:def 2;
   then A25: W-min(P) in Lower_Arc(P) by A12,XBOOLE_0:def 3;
     dom g3=the carrier of ((TOP-REAL 2)|K0) by FUNCT_2:def 1;
   then A26: dom g3=[#]((TOP-REAL 2)|K0) by PRE_TOPC:12;
   A27: rng g3=[#](Closed-Interval-TSpace(-1,1)) by A1,Lm4;
   A28: g3 is one-to-one by A1,Lm4;
        K0 is non empty compact by A11,JORDAN5A:1;
      then A29: (TOP-REAL 2)|K0 is compact by COMPTS_1:12;
        Closed-Interval-TSpace(-1,1)
       =TopSpaceMetr(Closed-Interval-MSpace(-1,1)) by TOPMETR:def 8;
      then Closed-Interval-TSpace(-1,1) is_T2 by PCOMPS_1:38;
      then A30: g3 is_homeomorphism by A26,A27,A28,A29,COMPTS_1:26;
      reconsider h=g3*g as map
          of Closed-Interval-TSpace(0,1),Closed-Interval-TSpace(-1,1)
                              by TOPMETR:27;
      A31: h is_homeomorphism by A22,A30,TOPMETR:27,TOPS_2:71;
      A32: 0 in dom g by A23,TOPREAL5:1;
      A33: 1 in dom g by A23,TOPREAL5:1;
      A34: s1 in [.0,1.] by A22,TOPREAL5:1;
      A35: s2 in [.0,1.] by A22,TOPREAL5:1;
      A36: -1=(|[-1,0]|)`1 by EUCLID:56.=proj1.(|[-1,0]|) by PSCOMP_1:def 28
      .=g3.(g.1) by A18,A22,A25,FUNCT_1:72
      .= h.1 by A33,FUNCT_1:23;
      A37: 1=(|[1,0]|)`1 by EUCLID:56.=proj1.(|[1,0]|) by PSCOMP_1:def 28
      .=g3.(g.0) by A19,A22,A24,FUNCT_1:72
      .= h.0 by A32,FUNCT_1:23;
      A38: p1`1=g0.p1 by PSCOMP_1:def 28
      .=g3.(g.s1) by A14,A22,FUNCT_1:72
      .= h.s1 by A23,A34,FUNCT_1:23;
        p2`1=proj1.p2 by PSCOMP_1:def 28
      .=g3.p2 by A16,FUNCT_1:72 .= h.s2 by A22,A23,A35,FUNCT_1:23;
     hence s1 <= s2 by A1,A5,A31,A34,A35,A36,A37,A38,Th12;
    end;
    then p1 in Lower_Arc(P) & p2 in Lower_Arc(P) & not p2=W-min(P) &
       LE p1,p2,Lower_Arc(P),E-max(P),W-min(P)
                    by A14,A16,A20,JORDAN5C:def 3;
 hence LE p1,p2,P by JORDAN6:def 10;
end;

theorem Th56: for p1,p2 being Point of TOP-REAL 2,
 P being compact non empty Subset of TOP-REAL 2
 st P={p where p is Point of TOP-REAL 2: |.p.|=1}
 & p1 in P & p2 in P
 & p1`1<0 & p2`1<0 & p1`2>=0 & p2`2>=0 &
 (p1`1<=p2`1 or p1`2<=p2`2) holds LE p1,p2,P
proof let p1,p2 be Point of TOP-REAL 2,
 P be compact non empty Subset of TOP-REAL 2;
 assume A1: P={p where p is Point of TOP-REAL 2: |.p.|=1}
  & p1 in P & p2 in P
  & p1`1<0 & p2`1<0 & p1`2>=0 & p2`2>=0 &
  (p1`1<=p2`1 or p1`2<=p2`2);
  then consider p3 being Point of TOP-REAL 2 such that
  A2: p3=p1 & |.p3.|=1;
  consider p3 being Point of TOP-REAL 2 such that
  A3: p3=p2 & |.p3.|=1 by A1;
  A4: now assume p1`2<=p2`2;
  then ((p1`2))^2 <= ((p2`2))^2 by A1,SQUARE_1:77;
  then A5: 1^2- ((p1`2))^2 >= 1^2-((p2`2))^2 by REAL_1:92;
    1^2=(p1`1)^2+(p1`2)^2 by A2,JGRAPH_3:10;
  then 1^2-(p1`2)^2=(p1`1)^2 by XCMPLX_1:26;
  then A6: 1^2-(p1`2)^2=(-(p1`1))^2 by SQUARE_1:61;
    -(p1`1)>=0 by A1,REAL_1:66;
  then A7: -(p1`1)=sqrt(1^2-((p1`2))^2) by A6,SQUARE_1:89;
    1^2=(p2`1)^2+(p2`2)^2 by A3,JGRAPH_3:10;
  then A8: 1^2-(p2`2)^2=(p2`1)^2 by XCMPLX_1:26;
  then A9: 1^2-(p2`2)^2=(-(p2`1))^2 by SQUARE_1:61;
    -(p2`1)>=0 by A1,REAL_1:66;
  then A10: -(p2`1)=sqrt(1^2-((p2`2))^2) by A9,SQUARE_1:89;
    1^2-((p2`2))^2>=0 by A8,SQUARE_1:72;
  then sqrt(1^2- ((p1`2))^2) >= sqrt(1^2-((p2`2))^2) by A5,SQUARE_1:94;
   hence p1`1<=p2`1 by A7,A10,REAL_1:50;
  end;
  A11: P is_simple_closed_curve by A1,JGRAPH_3:36;
  then A12: Upper_Arc(P) is_an_arc_of W-min(P),E-max(P) by JORDAN6:def 8;
   set P4=Lower_Arc(P);
   set P4b=Upper_Arc(P);
   A13: P4 is_an_arc_of E-max(P),W-min(P)
     & Upper_Arc(P) /\ P4={W-min(P),E-max(P)}
       & Upper_Arc(P) \/ P4=P
         & First_Point(Upper_Arc(P),W-min(P),E-max(P),
         Vertical_Line((W-bound(P)+E-bound(P))/2))`2>
              Last_Point(P4,E-max(P),W-min(P),
         Vertical_Line((W-bound(P)+E-bound(P))/2))`2 by A11,JORDAN6:def 9;
   A14: Upper_Arc(P)={p where p is Point of TOP-REAL 2:p in P & p`2>=0}
                              by A1,Th37;
   then A15: p1 in Upper_Arc(P) by A1;
   A16: p2 in Upper_Arc(P) by A1,A14;
   A17: W-min(P)=|[-1,0]| by A1,Th32;
   A18: E-max(P)=|[1,0]| by A1,Th33;
     for g being map of I[01], (TOP-REAL 2)|P4b, s1, s2 being Real st
       g is_homeomorphism
         & g.0 = W-min(P) & g.1 = E-max(P)
         & 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)|P4b,
       s1, s2 be Real;
     assume A19: g is_homeomorphism
         & g.0 = W-min(P) & g.1 = E-max(P)
         & g.s1 = p1 & 0 <= s1 & s1 <= 1
         & g.s2 = p2 & 0 <= s2 & s2 <= 1;
   then A20: dom g=[#](I[01]) by TOPS_2:def 5
   .=[.0,1.] by BORSUK_1:83,PRE_TOPC:12;
   reconsider g0=proj1 as map of TOP-REAL 2,R^1 by JGRAPH_2:16;
   set K0=Upper_Arc(P);
   reconsider g2=g0|K0 as map of (TOP-REAL 2)|K0,R^1 by JGRAPH_3:12;
   reconsider g3=g2 as continuous map of (TOP-REAL 2)|K0,
                Closed-Interval-TSpace(-1,1) by A1,Lm5;
     E-max(P) in {W-min(P),E-max(P)} by TARSKI:def 2;
   then A21: E-max(P) in Upper_Arc(P) by A13,XBOOLE_0:def 3;
     W-min(P) in {W-min(P),E-max(P)} by TARSKI:def 2;
   then A22: W-min(P) in Upper_Arc(P) by A13,XBOOLE_0:def 3;
     dom g3=the carrier of ((TOP-REAL 2)|K0) by FUNCT_2:def 1;
   then A23: dom g3=[#]((TOP-REAL 2)|K0) by PRE_TOPC:12;
   A24: rng g3=[#](Closed-Interval-TSpace(-1,1)) by A1,Lm5;
   A25: g3 is one-to-one by A1,Lm5;
        K0 is non empty compact by A12,JORDAN5A:1;
      then A26: (TOP-REAL 2)|K0 is compact by COMPTS_1:12;
        Closed-Interval-TSpace(-1,1)
       =TopSpaceMetr(Closed-Interval-MSpace(-1,1)) by TOPMETR:def 8;
      then Closed-Interval-TSpace(-1,1) is_T2 by PCOMPS_1:38;
      then A27: g3 is_homeomorphism by A23,A24,A25,A26,COMPTS_1:26;
      reconsider h=g3*g as map
          of Closed-Interval-TSpace(0,1),Closed-Interval-TSpace(-1,1)
                                by TOPMETR:27;
      A28: h is_homeomorphism by A19,A27,TOPMETR:27,TOPS_2:71;
      A29: 0 in dom g by A20,TOPREAL5:1;
      A30: 1 in dom g by A20,TOPREAL5:1;
      A31: s1 in [.0,1.] by A19,TOPREAL5:1;
      A32: s2 in [.0,1.] by A19,TOPREAL5:1;
      A33: -1=(|[-1,0]|)`1 by EUCLID:56.=proj1.(|[-1,0]|) by PSCOMP_1:def 28
      .=g3.(g.0) by A17,A19,A22,FUNCT_1:72
      .= h.0 by A29,FUNCT_1:23;
      A34: 1=(|[1,0]|)`1 by EUCLID:56.=g0.(|[1,0]|) by PSCOMP_1:def 28
      .=g3.(|[1,0]|) by A18,A21,FUNCT_1:72 .= h.1 by A18,A19,A30,FUNCT_1:23;
      A35: p1`1=g0.p1 by PSCOMP_1:def 28
      .=g3.(g.s1) by A15,A19,FUNCT_1:72
      .= h.s1 by A20,A31,FUNCT_1:23;
        p2`1=g0.p2 by PSCOMP_1:def 28
      .=g3.p2 by A16,FUNCT_1:72 .= h.s2 by A19,A20,A32,FUNCT_1:23;
     hence s1 <= s2 by A1,A4,A28,A31,A32,A33,A34,A35,Th11;
    end;
    then p1 in Upper_Arc(P) & p2 in Upper_Arc(P) &
       LE p1,p2,Upper_Arc(P),W-min(P),E-max(P) by A15,A16,JORDAN5C:def 3;
 hence LE p1,p2,P by JORDAN6:def 10;
end;

theorem Th57: for p1,p2 being Point of TOP-REAL 2,
 P being compact non empty Subset of TOP-REAL 2
 st P={p where p is Point of TOP-REAL 2: |.p.|=1}
 & p1 in P & p2 in P
 & p1`2>=0 & p2`2>=0 &
 p1`1<=p2`1 holds LE p1,p2,P
proof let p1,p2 be Point of TOP-REAL 2,
 P be compact non empty Subset of TOP-REAL 2;
 assume A1: P={p where p is Point of TOP-REAL 2: |.p.|=1}
  & p1 in P & p2 in P & p1`2>=0 & p2`2>=0 & p1`1<=p2`1;
  then A2: P is_simple_closed_curve by JGRAPH_3:36;
  then A3: Upper_Arc(P) is_an_arc_of W-min(P),E-max(P) by JORDAN6:def 8;
   set P4=Lower_Arc(P);
   set P4b=Upper_Arc(P);
   A4: P4 is_an_arc_of E-max(P),W-min(P)
     & Upper_Arc(P) /\ P4={W-min(P),E-max(P)}
       & Upper_Arc(P) \/ P4=P
         & First_Point(Upper_Arc(P),W-min(P),E-max(P),
         Vertical_Line((W-bound(P)+E-bound(P))/2))`2>
              Last_Point(P4,E-max(P),W-min(P),
         Vertical_Line((W-bound(P)+E-bound(P))/2))`2 by A2,JORDAN6:def 9;
   A5: Upper_Arc(P)={p where p is Point of TOP-REAL 2:p in P & p`2>=0}
                              by A1,Th37;
   then A6: p1 in Upper_Arc(P) by A1;
   A7: p2 in Upper_Arc(P) by A1,A5;
   A8: W-min(P)=|[-1,0]| by A1,Th32;
   A9: E-max(P)=|[1,0]| by A1,Th33;
     for g being map of I[01], (TOP-REAL 2)|P4b,
       s1, s2 being Real st
       g is_homeomorphism
         & g.0 = W-min(P) & g.1 = E-max(P)
         & 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)|P4b,
       s1, s2 be Real;
     assume A10: g is_homeomorphism
         & g.0 = W-min(P) & g.1 = E-max(P)
         & g.s1 = p1 & 0 <= s1 & s1 <= 1
         & g.s2 = p2 & 0 <= s2 & s2 <= 1;
   then A11: dom g=[#](I[01]) by TOPS_2:def 5
   .=[.0,1.] by BORSUK_1:83,PRE_TOPC:12;
   reconsider g0=proj1 as map of TOP-REAL 2,R^1 by JGRAPH_2:16;
   set K0=Upper_Arc(P);
   reconsider g2=g0|K0 as map of (TOP-REAL 2)|K0,R^1 by JGRAPH_3:12;
   reconsider g3=g2 as continuous map of (TOP-REAL 2)|K0,
                Closed-Interval-TSpace(-1,1) by A1,Lm5;
     E-max(P) in {W-min(P),E-max(P)} by TARSKI:def 2;
   then A12: E-max(P) in Upper_Arc(P) by A4,XBOOLE_0:def 3;
     W-min(P) in {W-min(P),E-max(P)} by TARSKI:def 2;
   then A13: W-min(P) in Upper_Arc(P) by A4,XBOOLE_0:def 3;
     dom g3=the carrier of ((TOP-REAL 2)|K0) by FUNCT_2:def 1;
   then A14: dom g3=[#]((TOP-REAL 2)|K0) by PRE_TOPC:12;
   A15: rng g3=[#](Closed-Interval-TSpace(-1,1)) by A1,Lm5;
   A16: g3 is one-to-one by A1,Lm5;
        K0 is non empty compact by A3,JORDAN5A:1;
      then A17: (TOP-REAL 2)|K0 is compact by COMPTS_1:12;
        Closed-Interval-TSpace(-1,1)
       =TopSpaceMetr(Closed-Interval-MSpace(-1,1)) by TOPMETR:def 8;
      then Closed-Interval-TSpace(-1,1) is_T2 by PCOMPS_1:38;
      then A18: g3 is_homeomorphism by A14,A15,A16,A17,COMPTS_1:26;
      reconsider h=g3*g as map
          of Closed-Interval-TSpace(0,1),Closed-Interval-TSpace(-1,1)
                   by TOPMETR:27;
      A19: h is_homeomorphism by A10,A18,TOPMETR:27,TOPS_2:71;
      A20: 0 in dom g by A11,TOPREAL5:1;
      A21: 1 in dom g by A11,TOPREAL5:1;
      A22: s1 in [.0,1.] by A10,TOPREAL5:1;
      A23: s2 in [.0,1.] by A10,TOPREAL5:1;
      A24: -1=(|[-1,0]|)`1 by EUCLID:56.=proj1.(|[-1,0]|) by PSCOMP_1:def 28
      .=g3.(|[-1,0]|) by A8,A13,FUNCT_1:72.= h.0 by A8,A10,A20,FUNCT_1:23;
      A25: 1=(|[1,0]|)`1 by EUCLID:56.=g0.(|[1,0]|) by PSCOMP_1:def 28
      .=g3.(|[1,0]|) by A9,A12,FUNCT_1:72
      .= h.1 by A9,A10,A21,FUNCT_1:23;
      A26: p1`1=g0.p1 by PSCOMP_1:def 28
      .=g3.p1 by A6,FUNCT_1:72 .= h.s1 by A10,A11,A22,FUNCT_1:23;
        p2`1=g0.p2 by PSCOMP_1:def 28
      .=g3.p2 by A7,FUNCT_1:72 .= h.s2 by A10,A11,A23,FUNCT_1:23;
     hence s1 <= s2 by A1,A19,A22,A23,A24,A25,A26,Th11;
    end;
    then p1 in Upper_Arc(P) & p2 in Upper_Arc(P) &
       LE p1,p2,Upper_Arc(P),W-min(P),E-max(P) by A6,A7,JORDAN5C:def 3;
 hence LE p1,p2,P by JORDAN6:def 10;
end;

theorem Th58: for p1,p2 being Point of TOP-REAL 2,
 P being compact non empty Subset of TOP-REAL 2
 st P={p where p is Point of TOP-REAL 2: |.p.|=1}
 & p1 in P & p2 in P & p1`1>=0 & p2`1>=0 &
 p1`2>=p2`2 holds LE p1,p2,P
proof let p1,p2 be Point of TOP-REAL 2,
 P be compact non empty Subset of TOP-REAL 2;
 assume A1: P={p where p is Point of TOP-REAL 2: |.p.|=1}
  & p1 in P & p2 in P & p1`1>=0 & p2`1>=0 & p1`2>=p2`2;
  then consider p3 being Point of TOP-REAL 2 such that
  A2: p3=p1 & |.p3.|=1;
  consider p3 being Point of TOP-REAL 2 such that
  A3: p3=p2 & |.p3.|=1 by A1;
  A4: P is_simple_closed_curve by A1,JGRAPH_3:36;
  then A5: Lower_Arc(P) is_an_arc_of E-max(P),W-min(P) by JORDAN6:def 9;
   set P4b=Lower_Arc(P);
   A6: P4b is_an_arc_of E-max(P),W-min(P)
     & Upper_Arc(P) /\ P4b={W-min(P),E-max(P)}
       & Upper_Arc(P) \/ P4b=P
         & First_Point(Upper_Arc(P),W-min(P),E-max(P),
         Vertical_Line((W-bound(P)+E-bound(P))/2))`2>
              Last_Point(P4b,E-max(P),W-min(P),
         Vertical_Line((W-bound(P)+E-bound(P))/2))`2 by A4,JORDAN6:def 9;
   A7: Upper_Arc(P)={p where p is Point of TOP-REAL 2:p in P & p`2>=0}
                              by A1,Th37;
   A8: W-min(P)=|[-1,0]| by A1,Th32;
   A9: E-max(P)=|[1,0]| by A1,Th33;
     now per cases;
   case p1 in Upper_Arc(P) & p2 in Upper_Arc(P);
     then consider p22 being Point of TOP-REAL 2 such that
     A10:  p2=p22 & p22 in P & p22`2>=0 by A7;
    ((p1`2))^2 >= ((p2`2))^2 by A1,A10,SQUARE_1:77;
  then A11: 1^2- ((p1`2))^2 <= 1^2-((p2`2))^2 by REAL_1:92;
    1^2=(p1`1)^2+(p1`2)^2 by A2,JGRAPH_3:10;
  then A12: 1^2-((p1`2))^2=((p1`1))^2 by XCMPLX_1:26;
  then A13: (p1`1)=sqrt(1^2-((p1`2))^2) by A1,SQUARE_1:89;
    1^2=(p2`1)^2+(p2`2)^2 by A3,JGRAPH_3:10;
  then 1^2-((p2`2))^2=((p2`1))^2 by XCMPLX_1:26;
  then A14: (p2`1)=sqrt(1^2-((p2`2))^2) by A1,SQUARE_1:89;
    1^2-((p1`2))^2>=0 by A12,SQUARE_1:72;
  then sqrt(1^2- ((p1`2))^2) <= sqrt(1^2-((p2`2))^2) by A11,SQUARE_1:94;
    hence LE p1,p2,P by A1,A10,A13,A14,Th57;
   case A15: p1 in Upper_Arc(P) & not p2 in Upper_Arc(P);
     then A16: p2 in Lower_Arc(P) by A1,A6,XBOOLE_0:def 2;
       now assume A17: p2=W-min(P);
         W-min(P)=|[-1,0]| by A1,Th32;
       then p2`2=0 by A17,EUCLID:56;
      hence contradiction by A1,A7,A15;
     end;
    hence LE p1,p2,P by A15,A16,JORDAN6:def 10;
   case A18:not p1 in Upper_Arc(P) & p2 in Upper_Arc(P);
      then consider p9 being Point of TOP-REAL 2 such that
      A19: p2=p9 & p9 in P & p9`2>=0 by A7;
     thus contradiction by A1,A7,A18,A19;
   case A20: not p1 in Upper_Arc(P) & not p2 in Upper_Arc(P);
     then A21: p1 in Lower_Arc(P) by A1,A6,XBOOLE_0:def 2;
     A22: p2 in Lower_Arc(P) by A1,A6,A20,XBOOLE_0:def 2;
       p1`2<0 by A1,A7,A20;
     then A23: -p1`2>0 by REAL_1:66;
    -p1`2<=-p2`2 by A1,REAL_1:50;
  then (-(p1`2))^2 <= (-(p2`2))^2 by A23,SQUARE_1:77;
  then A24: 1^2- (-(p1`2))^2 >= 1^2-(-(p2`2))^2 by REAL_1:92;
    1^2=(p1`1)^2+(p1`2)^2 by A2,JGRAPH_3:10;
  then 1^2-(p1`2)^2=(p1`1)^2 by XCMPLX_1:26;
  then 1^2-(-(p1`2))^2=((p1`1))^2 by SQUARE_1:61;
  then A25: p1`1=sqrt(1^2-(-(p1`2))^2) by A1,SQUARE_1:89;
    1^2=(p2`1)^2+(p2`2)^2 by A3,JGRAPH_3:10;
  then 1^2-(p2`2)^2=(p2`1)^2 by XCMPLX_1:26;
  then A26: 1^2-(-(p2`2))^2=((p2`1))^2 by SQUARE_1:61;
  then A27: (p2`1)=sqrt(1^2-(-(p2`2))^2) by A1,SQUARE_1:89;
    1^2-(-(p2`2))^2>=0 by A26,SQUARE_1:72;
  then A28: p1`1>=p2`1 by A24,A25,A27,SQUARE_1:94;
     A29: for g being map of I[01], (TOP-REAL 2)|P4b,
       s1, s2 being Real st
       g is_homeomorphism
         & g.0 = E-max(P) & g.1 = W-min(P)
         & 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)|P4b, s1, s2 be Real;
     assume A30: g is_homeomorphism
         & g.0 = E-max(P) & g.1 = W-min(P)
         & g.s1 = p1 & 0 <= s1 & s1 <= 1
         & g.s2 = p2 & 0 <= s2 & s2 <= 1;
       then A31: dom g=[#](I[01]) by TOPS_2:def 5
       .=[.0,1.] by BORSUK_1:83,PRE_TOPC:12;
       reconsider g0=proj1 as map of TOP-REAL 2,R^1 by JGRAPH_2:16;
       set K0=Lower_Arc(P);
       reconsider g2=g0|K0 as map of (TOP-REAL 2)|K0,R^1 by JGRAPH_3:12;
       reconsider g3=g2 as continuous map of (TOP-REAL 2)|K0,
                Closed-Interval-TSpace(-1,1) by A1,Lm4;
         E-max(P) in {W-min(P),E-max(P)} by TARSKI:def 2;
       then A32: E-max(P) in Lower_Arc(P) by A6,XBOOLE_0:def 3;
         W-min(P) in {W-min(P),E-max(P)} by TARSKI:def 2;
       then A33: W-min(P) in Lower_Arc(P) by A6,XBOOLE_0:def 3;
         dom g3=the carrier of ((TOP-REAL 2)|K0) by FUNCT_2:def 1;
       then A34: dom g3=[#]((TOP-REAL 2)|K0) by PRE_TOPC:12;
       A35: rng g3=[#](Closed-Interval-TSpace(-1,1)) by A1,Lm4;
       A36: g3 is one-to-one by A1,Lm4;
         K0 is non empty compact by A5,JORDAN5A:1;
       then A37: (TOP-REAL 2)|K0 is compact by COMPTS_1:12;
         Closed-Interval-TSpace(-1,1)
        =TopSpaceMetr(Closed-Interval-MSpace(-1,1)) by TOPMETR:def 8;
       then Closed-Interval-TSpace(-1,1) is_T2 by PCOMPS_1:38;
       then A38: g3 is_homeomorphism by A34,A35,A36,A37,COMPTS_1:26;
      reconsider h=g3*g as map
       of Closed-Interval-TSpace(0,1),Closed-Interval-TSpace(-1,1)
               by TOPMETR:27;
      A39: h is_homeomorphism by A30,A38,TOPMETR:27,TOPS_2:71;
      A40: 0 in dom g by A31,TOPREAL5:1;
      A41: 1 in dom g by A31,TOPREAL5:1;
      A42: s1 in [.0,1.] by A30,TOPREAL5:1;
      A43: s2 in [.0,1.] by A30,TOPREAL5:1;
      A44: -1=(|[-1,0]|)`1 by EUCLID:56.=proj1.(|[-1,0]|) by PSCOMP_1:def 28
      .=g3.(|[-1,0]|) by A8,A33,FUNCT_1:72 .= h.1 by A8,A30,A41,FUNCT_1:23
;
      A45: 1=(|[1,0]|)`1 by EUCLID:56.=proj1.(|[1,0]|) by PSCOMP_1:def 28
      .=g3.(|[1,0]|) by A9,A32,FUNCT_1:72 .= h.0 by A9,A30,A40,FUNCT_1:23;
      A46: p1`1=g0.p1 by PSCOMP_1:def 28
      .=g3.p1 by A21,FUNCT_1:72 .= h.s1 by A30,A31,A42,FUNCT_1:23;
        p2`1=g0.p2 by PSCOMP_1:def 28
      .=g3.p2 by A22,FUNCT_1:72 .= h.s2 by A30,A31,A43,FUNCT_1:23;
     hence s1 <= s2 by A28,A39,A42,A43,A44,A45,A46,Th12;
     end;
     A47: now assume A48: p2=W-min(P);
         W-min(P)=|[-1,0]| by A1,Th32;
       then p2`2=0 by A48,EUCLID:56;
      hence contradiction by A1,A7,A20;
     end;
       p1 in Lower_Arc(P) & p2 in Lower_Arc(P) &
       LE p1,p2,Lower_Arc(P),E-max(P),W-min(P) by A21,A22,A29,JORDAN5C:def 3;
    hence LE p1,p2,P by A47,JORDAN6:def 10;
   end;
 hence LE p1,p2,P;
end;

theorem Th59: for p1,p2 being Point of TOP-REAL 2,
 P being compact non empty Subset of TOP-REAL 2
 st P={p where p is Point of TOP-REAL 2: |.p.|=1}
 & p1 in P & p2 in P & p1`2<=0 & p2`2<=0 & p2<>W-min(P) &
 p1`1>=p2`1 holds LE p1,p2,P
proof let p1,p2 be Point of TOP-REAL 2,
 P be compact non empty Subset of TOP-REAL 2;
 assume A1: P={p where p is Point of TOP-REAL 2: |.p.|=1}
  & p1 in P & p2 in P & p1`2<=0 & p2`2<=0 & p2<>W-min(P) &
  p1`1>=p2`1;
  then A2: P is_simple_closed_curve by JGRAPH_3:36;
  then A3: Lower_Arc(P) is_an_arc_of E-max(P),W-min(P) by JORDAN6:def 9;
   set P4=Lower_Arc(P);
   A4: P4 is_an_arc_of E-max(P),W-min(P)
     & Upper_Arc(P) /\ P4={W-min(P),E-max(P)}
       & Upper_Arc(P) \/ P4=P
         & First_Point(Upper_Arc(P),W-min(P),E-max(P),
         Vertical_Line((W-bound(P)+E-bound(P))/2))`2>
              Last_Point(P4,E-max(P),W-min(P),
         Vertical_Line((W-bound(P)+E-bound(P))/2))`2 by A2,JORDAN6:def 9;
   A5: Lower_Arc(P)={p where p is Point of TOP-REAL 2:p in P & p`2<=0}
                              by A1,Th38;
   then A6: p1 in Lower_Arc(P) by A1;
   A7: p2 in Lower_Arc(P) by A1,A5;
   A8: W-min(P)=|[-1,0]| by A1,Th32;
   A9: E-max(P)=|[1,0]| by A1,Th33;
     for g being map of I[01], (TOP-REAL 2)|P4, s1, s2 being Real st
       g is_homeomorphism
         & g.0 = E-max(P) & g.1 = W-min(P)
         & 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)|P4, s1, s2 be Real;
     assume A10: g is_homeomorphism
         & g.0 = E-max(P) & g.1 = W-min(P)
         & g.s1 = p1 & 0 <= s1 & s1 <= 1
         & g.s2 = p2 & 0 <= s2 & s2 <= 1;
   then A11: dom g=[#](I[01]) by TOPS_2:def 5
   .=[.0,1.] by BORSUK_1:83,PRE_TOPC:12;
   reconsider g0=proj1 as map of TOP-REAL 2,R^1 by JGRAPH_2:16;
   set K0=Lower_Arc(P);
   reconsider g2=g0|K0 as map of (TOP-REAL 2)|K0,R^1 by JGRAPH_3:12;
   reconsider g3=g2 as continuous map of (TOP-REAL 2)|K0,
                Closed-Interval-TSpace(-1,1) by A1,Lm4;
     E-max(P) in {W-min(P),E-max(P)} by TARSKI:def 2;
   then A12: E-max(P) in Lower_Arc(P) by A4,XBOOLE_0:def 3;
     W-min(P) in {W-min(P),E-max(P)} by TARSKI:def 2;
   then A13: W-min(P) in Lower_Arc(P) by A4,XBOOLE_0:def 3;
     dom g3=the carrier of ((TOP-REAL 2)|K0) by FUNCT_2:def 1;
   then A14: dom g3=[#]((TOP-REAL 2)|K0) by PRE_TOPC:12;
   A15: rng g3=[#](Closed-Interval-TSpace(-1,1)) by A1,Lm4;
   A16: g3 is one-to-one by A1,Lm4;
        K0 is non empty compact by A3,JORDAN5A:1;
      then A17: (TOP-REAL 2)|K0 is compact by COMPTS_1:12;
        Closed-Interval-TSpace(-1,1)
       =TopSpaceMetr(Closed-Interval-MSpace(-1,1)) by TOPMETR:def 8;
      then Closed-Interval-TSpace(-1,1) is_T2 by PCOMPS_1:38;
      then A18: g3 is_homeomorphism by A14,A15,A16,A17,COMPTS_1:26;
      reconsider h=g3*g as map
          of Closed-Interval-TSpace(0,1),Closed-Interval-TSpace(-1,1)
            by TOPMETR:27;
      A19: h is_homeomorphism by A10,A18,TOPMETR:27,TOPS_2:71;
      A20: 0 in dom g by A11,TOPREAL5:1;
      A21: 1 in dom g by A11,TOPREAL5:1;
      A22: s1 in [.0,1.] by A10,TOPREAL5:1;
      A23: s2 in [.0,1.] by A10,TOPREAL5:1;
      A24: -1=(|[-1,0]|)`1 by EUCLID:56.=proj1.(|[-1,0]|) by PSCOMP_1:def 28
      .=g3.(|[-1,0]|) by A8,A13,FUNCT_1:72 .= h.1 by A8,A10,A21,FUNCT_1:23
;
      A25: 1=(|[1,0]|)`1 by EUCLID:56.=proj1.(|[1,0]|) by PSCOMP_1:def 28
      .=g3.(|[1,0]|) by A9,A12,FUNCT_1:72 .= h.0 by A9,A10,A20,FUNCT_1:23;
      A26: p1`1=g0.p1 by PSCOMP_1:def 28
      .=g3.p1 by A6,FUNCT_1:72 .= h.s1 by A10,A11,A22,FUNCT_1:23;
        p2`1=g0.p2 by PSCOMP_1:def 28
      .=g3.p2 by A7,FUNCT_1:72 .= h.s2 by A10,A11,A23,FUNCT_1:23;
     hence s1 <= s2 by A1,A19,A22,A23,A24,A25,A26,Th12;
    end;
    then p1 in Lower_Arc(P) & p2 in Lower_Arc(P) & not p2=W-min(P) &
    LE p1,p2,Lower_Arc(P),E-max(P),W-min(P) by A1,A6,A7,JORDAN5C:def 3;
 hence LE p1,p2,P by JORDAN6:def 10;
end;

theorem Th60: for cn being Real,q being Point of TOP-REAL 2 st
  -1<cn & cn<1 & q`2<=0 holds (for p being Point of TOP-REAL 2 st
 p=(cn-FanMorphS).q holds p`2<=0)
 proof let cn be Real,q be Point of TOP-REAL 2;
  assume A1: -1<cn & cn<1 & q`2<=0;
   let p be Point of TOP-REAL 2;
     assume A2: p=(cn-FanMorphS).q;
      per cases by A1;
      suppose A3: q`2<0;
         now per cases;
       case q`1/|.q.|<cn;
        hence p`2<=0 by A1,A2,A3,JGRAPH_4:145;
       case q`1/|.q.|>=cn;
        hence p`2<=0 by A1,A2,A3,JGRAPH_4:144;
       end;
      hence p`2<=0;
      suppose q`2=0;
       hence p`2<=0 by A2,JGRAPH_4:120;
 end;

theorem Th61: for cn being Real,p1,p2,q1,q2 being Point of TOP-REAL 2,
 P being compact non empty Subset of TOP-REAL 2
 st -1<cn & cn<1 & P={p where p is Point of TOP-REAL 2: |.p.|=1}
 & LE p1,p2,P & q1=(cn-FanMorphS).p1 & q2=(cn-FanMorphS).p2
 holds LE q1,q2,P
proof let cn be Real,p1,p2,q1,q2 be Point of TOP-REAL 2,
 P be compact non empty Subset of TOP-REAL 2;
 assume A1: -1<cn & cn<1 & P={p where p is Point of TOP-REAL 2: |.p.|=1}
  & LE p1,p2,P & q1=(cn-FanMorphS).p1 & q2=(cn-FanMorphS).p2;
   then A2: Upper_Arc(P)={p where p is Point of TOP-REAL 2:p in P & p`2>=0}
                              by Th37;
   A3: Lower_Arc(P)={p where p is Point of TOP-REAL 2:p in P & p`2<=0}
                              by A1,Th38;
  A4: P is_simple_closed_curve by A1,JGRAPH_3:36;
      W-min(P)=|[-1,0]| by A1,Th32;
    then A5: (W-min(P))`2=0 by EUCLID:56;
    then A6: (cn-FanMorphS).(W-min(P))=W-min(P) by JGRAPH_4:120;
      W-min(P) in the carrier of TOP-REAL 2;
    then A7: W-min(P) in dom ((cn-FanMorphS)) by FUNCT_2:def 1;
      p2 in the carrier of TOP-REAL 2;
    then A8: p2 in dom ((cn-FanMorphS)) by FUNCT_2:def 1;
    A9: (cn-FanMorphS) is one-to-one by A1,JGRAPH_4:140;
    A10: Upper_Arc(P) c= P by A1,Th36;
    A11: Lower_Arc(P) c= P by A1,Th36;
    A12: now per cases by A1,JORDAN6:def 10;
    case p1 in Upper_Arc(P);
     hence p1 in P by A10;
    case p1 in Lower_Arc(P);
     hence p1 in P by A11;
    end;
    A13: now assume A14: q2=W-min(P);
      then p2=W-min(P) by A1,A6,A7,A8,A9,FUNCT_1:def 8;
      then LE p2,p1,P by A4,A12,JORDAN7:3;
      then A15: q1=q2 by A1,A4,JORDAN6:72;
        W-min(P) in Lower_Arc(P) by A4,JORDAN7:1;
      then LE q1,q2,P by A4,A11,A14,A15,JORDAN6:71;
   hence q1 in Upper_Arc(P) & q2 in Lower_Arc(P)& not q2=W-min(P) or
     q1 in Upper_Arc(P) & q2 in Upper_Arc(P) &
       LE q1,q2,Upper_Arc(P),W-min(P),E-max(P) or
     q1 in Lower_Arc(P) & q2 in Lower_Arc(P)& not q2=W-min(P) &
       LE q1,q2,Lower_Arc(P),E-max(P),W-min(P) by JORDAN6:def 10;
    end;
  per cases by A1,JORDAN6:def 10;
  suppose A16: p1 in Upper_Arc(P) & p2 in Lower_Arc(P)& not p2=W-min(P);
    then consider p8 being Point of TOP-REAL 2 such that
    A17: p8=p1 & p8 in P & p8`2>=0 by A2;
    consider p9 being Point of TOP-REAL 2 such that
    A18: p9=p2 & p9 in P & p9`2<=0 by A3,A16;
    A19: |.q2.|=|.p2.| by A1,JGRAPH_4:135;
    consider p10 being Point of TOP-REAL 2 such that
    A20: p10=p2 & |.p10.|=1 by A1,A18;
    A21: q2 in P by A1,A19,A20;
      q2`2<=0 by A1,A18,Th60;
   hence q1 in Upper_Arc(P) & q2 in Lower_Arc(P)& not q2=W-min(P) or
     q1 in Upper_Arc(P) & q2 in Upper_Arc(P) &
       LE q1,q2,Upper_Arc(P),W-min(P),E-max(P) or
     q1 in Lower_Arc(P) & q2 in Lower_Arc(P)& not q2=W-min(P) &
       LE q1,q2,Lower_Arc(P),E-max(P),W-min(P) by A1,A3,A13,A16,A17,A21,
JGRAPH_4:120;
  suppose A22: p1 in Upper_Arc(P) & p2 in Upper_Arc(P) &
       LE p1,p2,Upper_Arc(P),W-min(P),E-max(P);
    then consider p8 being Point of TOP-REAL 2 such that
    A23: p8=p1 & p8 in P & p8`2>=0 by A2;
    consider p9 being Point of TOP-REAL 2 such that
    A24: p9=p2 & p9 in P & p9`2>=0 by A2,A22;
      p1=(cn-FanMorphS).p1 by A23,JGRAPH_4:120;
   hence q1 in Upper_Arc(P) & q2 in Lower_Arc(P)& not q2=W-min(P) or
     q1 in Upper_Arc(P) & q2 in Upper_Arc(P) &
       LE q1,q2,Upper_Arc(P),W-min(P),E-max(P) or
     q1 in Lower_Arc(P) & q2 in Lower_Arc(P)& not q2=W-min(P) &
       LE q1,q2,Lower_Arc(P),E-max(P),W-min(P) by A1,A22,A24,JGRAPH_4:120;
  suppose A25: p1 in Lower_Arc(P) & p2 in Lower_Arc(P)& not p2=W-min(P) &
       LE p1,p2,Lower_Arc(P),E-max(P),W-min(P) & not p1 in Upper_Arc(P);
    A26: |.q1.|=|.p1.| by A1,JGRAPH_4:135;
    A27: |.q2.|=|.p2.| by A1,JGRAPH_4:135;
    consider p8 being Point of TOP-REAL 2 such that
    A28: p8=p1 & p8 in P & p8`2<=0 by A3,A25;
    A29: q1`2<=0 by A1,A28,Th60;
    consider p9 being Point of TOP-REAL 2 such that
    A30: p9=p2 & p9 in P & p9`2<=0 by A3,A25;
    A31: q2`2<=0 by A1,A30,Th60;
    consider p10 being Point of TOP-REAL 2 such that
    A32: p10=p1 & |.p10.|=1 by A1,A28;
    consider p11 being Point of TOP-REAL 2 such that
    A33: p11=p2 & |.p11.|=1 by A1,A30;
    A34: q1 in P by A1,A26,A32;
    A35: q2 in P by A1,A27,A33;
      now per cases;
    case A36: p1=W-min(P); then p1=(cn-FanMorphS).p1 by A5,JGRAPH_4:120;
      then LE q1,q2,P by A1,A4,A35,A36,JORDAN7:3;
     hence q1 in Upper_Arc(P) & q2 in Lower_Arc(P)& not q2=W-min(P) or
       q1 in Upper_Arc(P) & q2 in Upper_Arc(P) &
       LE q1,q2,Upper_Arc(P),W-min(P),E-max(P) or
       q1 in Lower_Arc(P) & q2 in Lower_Arc(P)& not q2=W-min(P) &
       LE q1,q2,Lower_Arc(P),E-max(P),W-min(P) by JORDAN6:def 10;
    case A37: p1<>W-min(P);
      now per cases by A1,A28,A30,A37,Th51;
    case A38: p1`1=p2`1;
           (p1`1)^2+(p1`2)^2=1^2 by A32,JGRAPH_3:10
         .=(p1`1)^2+(p2`2)^2 by A33,A38,JGRAPH_3:10;
         then (p1`2)^2=(p1`1)^2+(p2`2)^2 -(p1`1)^2 by XCMPLX_1:26
         .=(p2`2)^2 by XCMPLX_1:26;
         then A39: p1`2=p2`2 or p1`2=-p2`2 by JGRAPH_3:1;
           A40: p1=|[p1`1,p1`2]| by EUCLID:57;
           A41: p2=|[p2`1,p2`2]| by EUCLID:57;
           now assume A42: p1`2=-p2`2; then p2`2=0 by A28,A30,REAL_1:66;
          hence p1=p2 by A38,A41,A42,EUCLID:57;
         end;
         then LE q1,q2,P by A1,A4,A35,A38,A39,A40,A41,JORDAN6:71;
     hence q1 in Upper_Arc(P) & q2 in Lower_Arc(P)& not q2=W-min(P) or
       q1 in Upper_Arc(P) & q2 in Upper_Arc(P) &
       LE q1,q2,Upper_Arc(P),W-min(P),E-max(P) or
       q1 in Lower_Arc(P) & q2 in Lower_Arc(P)& not q2=W-min(P) &
       LE q1,q2,Lower_Arc(P),E-max(P),W-min(P) by JORDAN6:def 10;
    case p1`1>p2`1;
      then A43: p1`1/|.p1.|>p2`1/|.p2.| by A32,A33;
      A44: q2<> W-min(P) by A1,A6,A7,A8,A9,A25,FUNCT_1:def 8;
        q1`1/|.q1.|>=q2`1/|.q2.| by A1,A28,A30,A32,A33,A43,Th30;
      then LE q1,q2,P by A1,A26,A27,A29,A31,A32,A33,A34,A35,A44,Th59;
     hence
        q1 in Upper_Arc(P) & q2 in Lower_Arc(P)& not q2=W-min(P) or
       q1 in Upper_Arc(P) & q2 in Upper_Arc(P) &
       LE q1,q2,Upper_Arc(P),W-min(P),E-max(P) or
       q1 in Lower_Arc(P) & q2 in Lower_Arc(P)& not q2=W-min(P) &
       LE q1,q2,Lower_Arc(P),E-max(P),W-min(P) by JORDAN6:def 10;
    end;
   hence q1 in Upper_Arc(P) & q2 in Lower_Arc(P)& not q2=W-min(P) or
     q1 in Upper_Arc(P) & q2 in Upper_Arc(P) &
       LE q1,q2,Upper_Arc(P),W-min(P),E-max(P) or
     q1 in Lower_Arc(P) & q2 in Lower_Arc(P)& not q2=W-min(P) &
       LE q1,q2,Lower_Arc(P),E-max(P),W-min(P);
    end;
   hence
      q1 in Upper_Arc(P) & q2 in Lower_Arc(P)& not q2=W-min(P) or
     q1 in Upper_Arc(P) & q2 in Upper_Arc(P) &
       LE q1,q2,Upper_Arc(P),W-min(P),E-max(P) or
     q1 in Lower_Arc(P) & q2 in Lower_Arc(P)& not q2=W-min(P) &
       LE q1,q2,Lower_Arc(P),E-max(P),W-min(P);
end;

theorem Th62: for p1,p2,p3,p4 being Point of TOP-REAL 2,
P being compact non empty 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
  & p1`1<0 & p1`2>=0 & p2`1<0 & p2`2>=0
  & p3`1<0 & p3`2>=0 & p4`1<0 & p4`2>=0
 ex f being map of TOP-REAL 2,TOP-REAL 2,
   q1,q2,q3,q4 being Point of TOP-REAL 2 st
   f is_homeomorphism &
   (for q being Point of TOP-REAL 2 holds |.(f.q).|=|.q.|)&
   q1=f.p1 & q2=f.p2 & q3=f.p3 & q4=f.p4 &
   (q1`1<0 & q1`2<0)&(q2`1<0 & q2`2<0)&(q3`1<0 & q3`2<0)&(q4`1<0 & q4`2<0)&
   LE q1,q2,P & LE q2,q3,P & LE q3,q4,P
 proof let p1,p2,p3,p4 be Point of TOP-REAL 2,
   P be compact non empty 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
     & p1`1<0 & p1`2>=0 & p2`1<0 & p2`2>=0
     & p3`1<0 & p3`2>=0 & p4`1<0 & p4`2>=0;
     then P is_simple_closed_curve by JGRAPH_3:36;
     then A2: P={p where p is Point of TOP-REAL 2: |.p.|=1}
     & p1 in P & p2 in P & p3 in P & p4 in P &
     LE p1,p2,P & LE p2,p3,P & LE p3,p4,P
     & p1`1<0 & p1`2>=0 & p2`1<0 & p2`2>=0
     & p3`1<0 & p3`2>=0 & p4`1<0 & p4`2>=0 by A1,JORDAN7:5;
     consider r being real number such that
     A3: p4`1<r & r<0 by A1,REAL_1:75;
     reconsider r1=r as Real by XREAL_0:def 1;
     set s=sqrt(1-r1^2);
       (p4`1)^2>r1^2 by A3,Th2;
     then A4: 1-(p4`1)^2<1-r1^2 by REAL_1:92;
     consider p being Point of TOP-REAL 2 such that
     A5: p=p4 & |.p.|=1 by A2;
     consider p11 being Point of TOP-REAL 2 such that
     A6: p11=p1 & |.p11.|=1 by A2;
     consider p22 being Point of TOP-REAL 2 such that
     A7: p22=p2 & |.p22.|=1 by A2;
     consider p33 being Point of TOP-REAL 2 such that
     A8: p33=p3 & |.p33.|=1 by A2;
       -1<=p4`1 by A5,Th3;
     then A9: -1<=r1 by A3,AXIOMS:22;
       r1<=1 by A3,AXIOMS:22;
     then r1^2<=1^2 by A9,JGRAPH_2:7;
     then A10: 1-r1^2>=0 by SQUARE_1:12,59;
     then A11: s^2=1-r1^2 by SQUARE_1:def 4;
     then 1-s^2=1-1+r1^2 by XCMPLX_1:37 .=r1^2;
     then 1-s^2>0 by A3,SQUARE_1:74;
     then 1-s^2+s^2>0+s^2 by REAL_1:67;
     then 1>0+s^2 by XCMPLX_1:27;
     then A12: -1<s & s<1 by JGRAPH_4:5;
     then consider f1 being map of TOP-REAL 2,TOP-REAL 2 such that
     A13: f1=s-FanMorphW & f1 is_homeomorphism by JGRAPH_4:48;
     A14: for q being Point of TOP-REAL 2 holds |.(f1.q).|=|.q.|
                             by A13,JGRAPH_4:40;
     A15: s>=0 by A10,SQUARE_1:def 4;
     set q11=f1.p1, q22=f1.p2, q33=f1.p3, q44=f1.p4;
     A16: |.q11.|=1 by A6,A13,JGRAPH_4:40;
     then A17: q11 in P by A1;
     A18: |.q22.|=1 by A7,A13,JGRAPH_4:40;
     then A19: q22 in P by A1;
     A20: |.q33.|=1 by A8,A13,JGRAPH_4:40;
     then A21: q33 in P by A1;
     A22: |.q44.|=1 by A5,A13,JGRAPH_4:40;
     then A23: q44 in P by A1;
     A24: -(p4`1)>0 by A1,REAL_1:66;
     A25: p1`1<p2`1 or p1=p2 by A1,Th49;
     A26: p2`1<p3`1 or p2=p3 by A1,Th49;
     A27: p3`1<p4`1 or p3=p4 by A1,Th49;
     then A28: p2`1<=p4`1 by A26,AXIOMS:22;
     then p1`1<=p4`1 by A25,AXIOMS:22;
     then -(p1`1)>= -(p4`1) by REAL_1:50;
     then (-(p1`1))^2>=(-(p4`1))^2 by A24,SQUARE_1:77;
     then (-(p1`1))^2>=((p4`1))^2 by SQUARE_1:61;
     then ((p1`1))^2>=((p4`1))^2 by SQUARE_1:61;
     then 1-((p1`1))^2<=1-((p4`1))^2 by REAL_2:106;
     then A29: 1-(p1`1)^2< s^2 by A4,A11,AXIOMS:22;
       1^2=(p1`1)^2+(p1`2)^2 by A6,JGRAPH_3:10;
     then 1-((p1`1))^2=((p1`2))^2 by SQUARE_1:59,XCMPLX_1:26;
     then A30: p1`2/|.p1.|<s by A6,A15,A29,JGRAPH_2:6;
     then A31: q11`1<0 & q11`2<0 by A1,A12,A13,JGRAPH_4:50;
       -(p2`1)>= -(p4`1) by A28,REAL_1:50;
     then (-(p2`1))^2>=(-(p4`1))^2 by A24,SQUARE_1:77;
     then (-(p2`1))^2>=((p4`1))^2 by SQUARE_1:61;
     then ((p2`1))^2>=((p4`1))^2 by SQUARE_1:61;
     then 1-((p2`1))^2<=1-((p4`1))^2 by REAL_2:106;
     then A32: 1-(p2`1)^2< s^2 by A4,A11,AXIOMS:22;
       1^2=(p2`1)^2+(p2`2)^2 by A7,JGRAPH_3:10;
     then 1-((p2`1))^2=((p2`2))^2 by SQUARE_1:59,XCMPLX_1:26;
     then A33: p2`2/|.p2.|<s by A7,A15,A32,JGRAPH_2:6;
     then A34: q22`1<0 & q22`2<0 by A1,A12,A13,JGRAPH_4:50;
       -(p3`1)>= -(p4`1) by A27,REAL_1:50;
     then (-(p3`1))^2>=(-(p4`1))^2 by A24,SQUARE_1:77;
     then (-(p3`1))^2>=((p4`1))^2 by SQUARE_1:61;
     then ((p3`1))^2>=((p4`1))^2 by SQUARE_1:61;
     then 1-((p3`1))^2<=1-((p4`1))^2 by REAL_2:106;
     then A35: 1-(p3`1)^2< s^2 by A4,A11,AXIOMS:22;
       1^2=(p3`1)^2+(p3`2)^2 by A8,JGRAPH_3:10;
     then 1-((p3`1))^2=((p3`2))^2 by SQUARE_1:59,XCMPLX_1:26;
     then A36: p3`2/|.p3.|<s by A8,A15,A35,JGRAPH_2:6;
     then A37: q33`1<0 & q33`2<0 by A1,A12,A13,JGRAPH_4:50;
       1^2=(p4`1)^2+(p4`2)^2 by A5,JGRAPH_3:10;
     then 1-((p4`1))^2=((p4`2))^2 by SQUARE_1:59,XCMPLX_1:26;
     then p4`2/|.p4.|<s by A4,A5,A11,A15,JGRAPH_2:6;
     then A38: (q11`1<0 & q11`2<0)&(q22`1<0 & q22`2<0)&(q33`1<0 & q33`2<0)
     &(q44`1<0 & q44`2<0) by A1,A12,A13,A30,A33,A36,JGRAPH_4:50;
       p1`2/|.p1.|<p2`2/|.p2.| or p1=p2 by A1,A6,A7,Th49;
     then q11`2/|.q11.|<q22`2/|.q22.| or p1=p2 by A1,A12,A13,JGRAPH_4:53;
     then A39: LE q11,q22,P by A1,A16,A17,A18,A19,A31,A34,Th54;
       p2`2/|.p2.|<p3`2/|.p3.| or p2=p3 by A1,A7,A8,Th49;
     then q22`2/|.q22.|<q33`2/|.q33.| or p2=p3 by A1,A12,A13,JGRAPH_4:53;
     then A40: LE q22,q33,P by A1,A18,A19,A20,A21,A34,A37,Th54;
       p3`2/|.p3.|<p4`2/|.p4.| or p3=p4 by A1,A5,A8,Th49;
     then q33`2/|.q33.|<q44`2/|.q44.| or p3=p4 by A1,A12,A13,JGRAPH_4:53;
     then LE q33,q44,P by A1,A20,A21,A22,A23,A38,Th54;
    hence thesis by A13,A14,A38,A39,A40;
 end;

theorem Th63: for p1,p2,p3,p4 being Point of TOP-REAL 2,
P being compact non empty 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
  & p1`2>=0 & p2`2>=0
  & p3`2>=0 & p4`2>0
 ex f being map of TOP-REAL 2,TOP-REAL 2,
   q1,q2,q3,q4 being Point of TOP-REAL 2 st
   f is_homeomorphism &
   (for q being Point of TOP-REAL 2 holds |.(f.q).|=|.q.|)&
   q1=f.p1 & q2=f.p2 & q3=f.p3 & q4=f.p4 &
   (q1`1<0 & q1`2>=0)&(q2`1<0 & q2`2>=0)&
   (q3`1<0 & q3`2>=0)&(q4`1<0 & q4`2>=0)&
   LE q1,q2,P & LE q2,q3,P & LE q3,q4,P
 proof let p1,p2,p3,p4 be Point of TOP-REAL 2,
   P be compact non empty 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
     & p1`2>=0 & p2`2>=0
     & p3`2>=0 & p4`2>0;
     then P is_simple_closed_curve by JGRAPH_3:36;
     then A2: P={p where p is Point of TOP-REAL 2: |.p.|=1}
     & p1 in P & p2 in P & p3 in P & p4 in P &
     LE p1,p2,P & LE p2,p3,P & LE p3,p4,P
     & p1`2>=0 & p2`2>=0
     & p3`2>=0 & p4`2>0 by A1,JORDAN7:5;
     then consider p being Point of TOP-REAL 2 such that
     A3: p=p4 & |.p.|=1;
     consider p11 being Point of TOP-REAL 2 such that
     A4: p11=p1 & |.p11.|=1 by A2;
     consider p22 being Point of TOP-REAL 2 such that
     A5: p22=p2 & |.p22.|=1 by A2;
     consider p33 being Point of TOP-REAL 2 such that
     A6: p33=p3 & |.p33.|=1 by A2;
     A7: -1<=p4`1 & p4`1<=1 by A3,Th3;
       now assume p4`1=1;
       then 1=1+(p4`2)^2 by A3,JGRAPH_3:10,SQUARE_1:59;
       then 1-1=(p4`2)^2 by XCMPLX_1:26;
      hence contradiction by A1,SQUARE_1:73;
     end;
     then p4`1<1 by A7,REAL_1:def 5;
     then consider r being real number such that
     A8: p4`1<r & r<1 by REAL_1:75;
     reconsider r1=r as Real by XREAL_0:def 1;
     A9: -1<r1 & r1<1 by A7,A8,AXIOMS:22;
     then consider f1 being map of TOP-REAL 2,TOP-REAL 2 such that
     A10: f1=r1-FanMorphN & f1 is_homeomorphism by JGRAPH_4:81;
     A11: for q being Point of TOP-REAL 2 holds |.(f1.q).|=|.q.|
                             by A10,JGRAPH_4:73;
     set q11=f1.p1, q22=f1.p2, q33=f1.p3, q44=f1.p4;
     A12: |.q11.|=1 by A4,A10,JGRAPH_4:73;
     then A13: q11 in P by A1;
     A14: |.q22.|=1 by A5,A10,JGRAPH_4:73;
     then A15: q22 in P by A1;
     A16: |.q33.|=1 by A6,A10,JGRAPH_4:73;
     then A17: q33 in P by A1;
     A18: |.q44.|=1 by A3,A10,JGRAPH_4:73;
     then A19: q44 in P by A1;
     A20: p1`1<p2`1 or p1=p2 by A1,Th50;
     A21: p2`1<p3`1 or p2=p3 by A1,Th50;
     A22: p3`1<p4`1 or p3=p4 by A1,Th50;
     then A23: p3`1<r1 by A8,AXIOMS:22;
     then A24: p2`1<r1 by A21,AXIOMS:22;
     A25: p2`1/|.p2.|<r1 by A5,A21,A23,AXIOMS:22;
     then A26: q22`2>=0 & q22`1<0 by A1,A5,A9,A10,Th23;
     A27: p1`1/|.p1.|<r1 by A4,A20,A24,AXIOMS:22;
     then A28: q11`2>=0 & q11`1<0 by A1,A4,A9,A10,Th23;
       p4`1/|.p4.|<r1 by A3,A8;
     then A29: q44`1<0 & q44`2>0 by A1,A9,A10,JGRAPH_4:83;
     A30: (q11`1<0 & q11`2>=0 or q11`1<0 & q11`2=0)&(q22`1<0 & q22`2>=0)
                    by A1,A4,A5,A9,A10,A25,A27,Th23;
       p1`1/|.p1.|<p2`1/|.p2.| or p1=p2 by A1,A4,A5,Th50;
     then q11`1/|.q11.|<q22`1/|.q22.| or p1=p2
                     by A1,A4,A5,A9,A10,Th24;
     then A31: LE q11,q22,P by A1,A12,A13,A14,A15,A26,A28,Th56;
       p3`1/|.p3.|<r1 by A6,A8,A22,AXIOMS:22;
     then A32: q33`2>=0 & q33`1<0 by A1,A6,A9,A10,Th23;
       p2`1/|.p2.|<p3`1/|.p3.| or p2=p3 by A1,A5,A6,Th50;
     then q22`1/|.q22.|<q33`1/|.q33.| or p2=p3
                        by A1,A5,A6,A9,A10,Th24;
     then A33: LE q22,q33,P by A1,A14,A15,A16,A17,A26,A32,Th56;
       p3`1/|.p3.|<p4`1/|.p4.| or p3=p4 by A1,A3,A6,Th50;
     then q33`1/|.q33.|<q44`1/|.q44.| or p3=p4
                          by A1,A3,A6,A9,A10,Th24;
     then LE q33,q44,P by A1,A16,A17,A18,A19,A29,A32,Th56;
    hence thesis by A10,A11,A29,A30,A31,A32,A33;
 end;

theorem Th64: for p1,p2,p3,p4 being Point of TOP-REAL 2,
P being compact non empty 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
  & p1`2>=0 & p2`2>=0
  & p3`2>=0 & p4`2>0
 ex f being map of TOP-REAL 2,TOP-REAL 2,
   q1,q2,q3,q4 being Point of TOP-REAL 2 st
   f is_homeomorphism &
   (for q being Point of TOP-REAL 2 holds |.(f.q).|=|.q.|)&
   q1=f.p1 & q2=f.p2 & q3=f.p3 & q4=f.p4 &
   (q1`1<0 & q1`2<0)&(q2`1<0 & q2`2<0)&(q3`1<0 & q3`2<0)&(q4`1<0 & q4`2<0)&
   LE q1,q2,P & LE q2,q3,P & LE q3,q4,P
proof let p1,p2,p3,p4 be Point of TOP-REAL 2,
  P be compact non empty 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
   & p1`2>=0 & p2`2>=0
   & p3`2>=0 & p4`2>0;
   then consider f being map of TOP-REAL 2,TOP-REAL 2,
     q1,q2,q3,q4 being Point of TOP-REAL 2 such that
   A2: f is_homeomorphism &
    (for q being Point of TOP-REAL 2 holds |.(f.q).|=|.q.|)&
    q1=f.p1 & q2=f.p2 & q3=f.p3 & q4=f.p4 &
    (q1`1<0 & q1`2>=0)&(q2`1<0 & q2`2>=0)&
    (q3`1<0 & q3`2>=0)&(q4`1<0 & q4`2>=0)&
    LE q1,q2,P & LE q2,q3,P & LE q3,q4,P by Th63;
  consider f2 being map of TOP-REAL 2,TOP-REAL 2,
   q1b,q2b,q3b,q4b being Point of TOP-REAL 2 such that
   A3: f2 is_homeomorphism &
   (for q being Point of TOP-REAL 2 holds |.(f2.q).|=|.q.|)&
   q1b=f2.q1 & q2b=f2.q2 & q3b=f2.q3 & q4b=f2.q4 &
   (q1b`1<0 & q1b`2<0)&(q2b`1<0 & q2b`2<0)&
   (q3b`1<0 & q3b`2<0)&(q4b`1<0 & q4b`2<0)&
   LE q1b,q2b,P & LE q2b,q3b,P & LE q3b,q4b,P by A1,A2,Th62;
  reconsider f3=f2*f as map of TOP-REAL 2,TOP-REAL 2;
  A4: f3 is_homeomorphism by A2,A3,TOPS_2:71;
  A5: for q being Point of TOP-REAL 2 holds |.(f3.q).|=|.q.|
   proof let q be Point of TOP-REAL 2;
       dom f=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
     then f3.q=f2.(f.q) by FUNCT_1:23;
    hence |.f3.q.|=|.(f.q).| by A3 .=|.q.| by A2;
   end;
  A6: dom f=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
  then A7: f3.p1=q1b by A2,A3,FUNCT_1:23;
  A8: f3.p2=q2b by A2,A3,A6,FUNCT_1:23;
  A9: f3.p3=q3b by A2,A3,A6,FUNCT_1:23;
    f3.p4=q4b by A2,A3,A6,FUNCT_1:23;
 hence thesis by A3,A4,A5,A7,A8,A9;
end;

theorem Th65: for p1,p2,p3,p4 being Point of TOP-REAL 2,
P being compact non empty 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
  & (p1`2>=0 or p1`1>=0)& (p2`2>=0 or p2`1>=0)
  & (p3`2>=0 or p3`1>=0)& (p4`2>0 or p4`1>0)
 ex f being map of TOP-REAL 2,TOP-REAL 2,
   q1,q2,q3,q4 being Point of TOP-REAL 2 st
   f is_homeomorphism &
   (for q being Point of TOP-REAL 2 holds |.(f.q).|=|.q.|)&
   q1=f.p1 & q2=f.p2 & q3=f.p3 & q4=f.p4 &
   q1`2>=0 & q2`2>=0 &
   q3`2>=0 & q4`2>0 &
   LE q1,q2,P & LE q2,q3,P & LE q3,q4,P
 proof let p1,p2,p3,p4 be Point of TOP-REAL 2,
   P be compact non empty 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
     & (p1`2>=0 or p1`1>=0)& (p2`2>=0 or p2`1>=0)
     & (p3`2>=0 or p3`1>=0)& (p4`2>0 or p4`1>0);
     then A2: P is_simple_closed_curve by JGRAPH_3:36;
     then A3: P={p where p is Point of TOP-REAL 2: |.p.|=1}
     & p1 in P & p2 in P & p3 in P & p4 in P &
     LE p1,p2,P & LE p2,p3,P & LE p3,p4,P
     & (p1`2>=0 or p1`1>=0)& (p2`2>=0 or p2`1>=0)
     & (p3`2>=0 or p3`1>=0)& (p4`2>0 or p4`1>0) by A1,JORDAN7:5;
     then consider p44 being Point of TOP-REAL 2 such that
     A4: p44=p4 & |.p44.|=1;
     consider p11 being Point of TOP-REAL 2 such that
     A5: p11=p1 & |.p11.|=1 by A3;
     consider p22 being Point of TOP-REAL 2 such that
     A6: p22=p2 & |.p22.|=1 by A3;
     consider p33 being Point of TOP-REAL 2 such that
     A7: p33=p3 & |.p33.|=1 by A3;
     A8: -1<=p4`2 & p4`2<=1 by A4,Th3;
       now assume A9: p4`2=-1;
         1=(p4`1)^2+(p4`2)^2 by A4,JGRAPH_3:10,SQUARE_1:59
       .=(p4`1)^2+1 by A9,SQUARE_1:59,61;
       then 1-1=(p4`1)^2 by XCMPLX_1:26;
      hence contradiction by A1,A9,SQUARE_1:73;
     end;
     then p4`2> -1 by A8,REAL_1:def 5;
     then consider r being real number such that
     A10: -1<r & r<p4`2 by REAL_1:75;
     reconsider r1=r as Real by XREAL_0:def 1;
     A11: -1<r1 & r1<1 by A8,A10,AXIOMS:22;
     then consider f1 being map of TOP-REAL 2,TOP-REAL 2 such that
     A12: f1=r1-FanMorphE & f1 is_homeomorphism by JGRAPH_4:112;
     set q11=f1.p1, q22=f1.p2, q33=f1.p3, q44=f1.p4;
     A13: |.q11.|=1 by A5,A12,JGRAPH_4:104;
     then A14: q11 in P by A1;
     A15: |.q22.|=1 by A6,A12,JGRAPH_4:104;
     then A16: q22 in P by A1;
     A17: |.q33.|=1 by A7,A12,JGRAPH_4:104;
     then A18: q33 in P by A1;
     A19: |.q44.|=1 by A4,A12,JGRAPH_4:104;
     then A20: q44 in P by A1;
      now per cases;
    case A21: p4`2<=0;
     A22: p4`2/|.p4.|>r1 by A4,A10;
     then A23: q44`1>0 & q44`2>=0 by A1,A11,A12,A21,JGRAPH_4:113;
     A24: now assume A25: q44`2=0;
         1^2=(q44`1)^2+(q44`2)^2 by A19,JGRAPH_3:10
       .=(q44`1)^2 by A25,SQUARE_1:60;
       then q44`1=-1 or q44`1=1 by JGRAPH_3:2,SQUARE_1:59;
       then A26: q44=|[1,0]| by A1,A11,A12,A21,A22,A25,EUCLID:57,JGRAPH_4:113;
       set q8=|[sqrt(1-r1^2),r1]|;
       A27: q8`1=sqrt(1-r1^2) & q8`2=r1 by EUCLID:56;
       then A28: |.q8.|=sqrt((sqrt(1-r1^2))^2+r1^2)by JGRAPH_3:10;
         1^2>r1^2 by A11,JGRAPH_2:8;
       then A29: 1-r1^2>0 by SQUARE_1:11,59;
       then A30: |.q8.|=sqrt((1-r1^2)+r1^2) by A28,SQUARE_1:def 4
           .=1 by SQUARE_1:83,XCMPLX_1:27;
       then A31: q8`2/|.q8.|=r1 by EUCLID:56;
       A32: q8`1>0 by A27,A29,SQUARE_1:93;
       set r8=f1.q8;
       A33: r8`1>0 & r8`2=0 by A11,A12,A31,A32,JGRAPH_4:118;
         |.r8.|=1 by A12,A30,JGRAPH_4:104;
       then 1^2=(r8`1)^2+(r8`2)^2 by JGRAPH_3:10
       .=(r8`1)^2 by A33,SQUARE_1:60;
       then r8`1=-1 or r8`1=1 by JGRAPH_3:2,SQUARE_1:59;
       then A34: f1.(|[sqrt(1-r1^2),r1]|)=|[1,0]| by A33,EUCLID:57;
       A35: f1 is one-to-one by A11,A12,JGRAPH_4:109;
         dom f1=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
       then p4=|[sqrt(1-r1^2),r1]| by A26,A34,A35,FUNCT_1:def 8;
      hence contradiction by A10,EUCLID:56;
     end;
     A36: Lower_Arc(P)
         ={p7 where p7 is Point of TOP-REAL 2:p7 in P & p7`2<=0}
                          by A1,Th38;
     A37: Upper_Arc(P)
         ={p7 where p7 is Point of TOP-REAL 2:p7 in P & p7`2>=0}
                          by A1,Th37;
     A38: now per cases;
     case A39: p3`1<=0;
      then A40: q33=p3 by A12,JGRAPH_4:89;
      A41: now per cases by A39;
      case p3`1=0;
        then 1=0+(q33`2)^2 by A7,A40,JGRAPH_3:10,SQUARE_1:59,60
            .=(q33`2)^2;
        then A42: q33`2=-1 or q33`2=1 by JGRAPH_3:2;
          now assume q33`2=-1;
          then -1>=p4`2 by A1,A21,A40,Th53;
         hence contradiction by A10,AXIOMS:22;
        end;
       hence q33`2>=0 by A42;
      case p3`1<0;
       hence q33`2>=0 by A1,A12,JGRAPH_4:89;
      end;
A43:  q33`1<q44`1 by A1,A11,A12,A21,A22,A39,A40,JGRAPH_4:113;
        now per cases;
      case A44: p2<> W-min(P);
        A45: now assume A46: p2`2<0;
          then A47: p2 in Lower_Arc(P) by A3,A36;
            p3 in Upper_Arc(P) by A3,A37,A40,A41;
          then LE p3,p2,P by A44,A47,JORDAN6:def 10;
         hence contradiction by A1,A2,A40,A41,A46,JORDAN6:72;
        end;
          then A48: p2`1<=p3`1 by A1,A40,A41,Th50;
        then A49: q22=p2 by A12,A39,JGRAPH_4:89;
          now per cases;
        case A50: p1<> W-min(P);
          A51: now assume A52: p1`2<0;
            then A53: p1 in Lower_Arc(P) by A3,A36;
              p2 in Upper_Arc(P) by A3,A37,A45;
             then LE p2,p1,P by A50,A53,JORDAN6:def 10;
           hence contradiction by A1,A2,A45,A52,JORDAN6:72;
          end;
            then p1`1<=p2`1 by A1,A45,Th50;
         hence q11`2>=0 & q22`2>=0 &
          q33`2>=0 & q44`2>0 &
          LE q11,q22,P & LE q22,q33,P & LE q33,q44,P
        by A1,A12,A18,A20,A23,A24,A39,A41,A43,A45,A48,A49,A51,Th57,JGRAPH_4:89
;
        case A54: p1=W-min(P);
          A55: W-min(P)=|[-1,0]| by A1,Th32;
          then p1`1=-1 & p1`2=0 by A54,EUCLID:56;
          then p1=q11 by A12,JGRAPH_4:89;
         hence q11`2>=0 & q22`2>=0 &
          q33`2>=0 & q44`2>0 &
          LE q11,q22,P & LE q22,q33,P & LE q33,q44,P
    by A3,A12,A20,A23,A24,A39,A40,A41,A43,A45,A48,A54,A55,Th57,EUCLID:56,
JGRAPH_4:89;
        end;
        hence q11`2>=0 & q22`2>=0 &
         q33`2>=0 & q44`2>0 &
         LE q11,q22,P & LE q22,q33,P & LE q33,q44,P;
      case A56: p2=W-min(P);
        A57: W-min(P)=|[-1,0]| by A1,Th32;
        then A58: p2`1=-1 & p2`2=0 by A56,EUCLID:56;
        then A59: p2=q22 by A12,JGRAPH_4:89;
        A60: now assume A61: p1`2<0;
          then A62: p1 in Lower_Arc(P) by A3,A36;
          A63: p2 in Upper_Arc(P) by A16,A37,A58,A59;
          A64: p1<>W-min(P) by A57,A61,EUCLID:56;
             LE p2,p1,P by A56,A58,A61,A62,A63,JORDAN6:def 10;
         hence contradiction by A1,A2,A56,A64,JORDAN6:72;
        end;
        then p1`1<=p2`1 by A1,A58,Th50;
        then p1`1<=0 by A58,AXIOMS:22;
       hence q11`2>=0 & q22`2>=0 &
        q33`2>=0 & q44`2>0 &
        LE q11,q22,P & LE q22,q33,P & LE q33,q44,P
          by A1,A12,A18,A20,A23,A24,A41,A43,A56,A57,A59,A60,Th57,EUCLID:56,
JGRAPH_4:89;
      end;
      hence q11`2>=0 & q22`2>=0 & q33`2>=0 & q44`2>0 &
       LE q11,q22,P & LE q22,q33,P & LE q33,q44,P;

     case A65: p3`1>0;
       A66: now per cases;
       case A67: p3<>p4;
       then p3`2>p4`2 by A1,A21,A65,Th53;
       then A68: p3`2/|.p3.|>=r1 by A7,A10,AXIOMS:22;
       then A69: q33`1>0 & q33`2>=0 by A11,A12,A65,JGRAPH_4:113;
         p3`2/|.p3.|>p4`2/|.p4.| by A1,A4,A7,A21,A65,A67,Th53;
       then A70: q33`2/|.q33.|>q44`2/|.q44.| by A1,A4,A7,A11,A12,A21,A65,Th27;
         1^2=(q33`1)^2+(q33`2)^2 by A17,JGRAPH_3:10;
       then A71: 1^2-(q33`2)^2=(q33`1)^2 by XCMPLX_1:26;
       then A72: (q33`1)=sqrt(1^2-((q33`2))^2) by A69,SQUARE_1:89;
         1^2=(q44`1)^2+(q44`2)^2 by A19,JGRAPH_3:10;
       then 1^2-(q44`2)^2=(q44`1)^2 by XCMPLX_1:26;
       then A73: (q44`1)=sqrt(1^2-((q44`2))^2) by A23,SQUARE_1:89;
         ((q33`2))^2 > ((q44`2))^2 by A17,A19,A23,A70,SQUARE_1:78;
       then A74: 1^2- ((q33`2))^2 < 1^2-((q44`2))^2 by REAL_1:92;
         1^2-((q33`2))^2>=0 by A71,SQUARE_1:72;
       then A75: q33`1< q44`1 by A72,A73,A74,SQUARE_1:95;
        A76: now assume p2`1=0;
          then 1^2=0+(p2`2)^2 by A6,JGRAPH_3:10,SQUARE_1:60;
         hence p2`2=1 or p2`2=-1 by JGRAPH_3:1;
        end;
        A77: now assume A78: p2`1=0 & p2`2=-1;
          then p2`2<=p4`2 by A4,Th3;
          then A79: LE p4,p2,P by A3,A21,A78,Th58;
            LE p2,p4,P by A1,A2,JORDAN6:73;
          then p2=p4 by A2,A79,JORDAN6:72;
         hence contradiction by A1,A78;
        end;
          now per cases by A1,A76,A77;
        case A80: p2`1<=0 & p2`2>=0;
          then A81: q22=p2 by A12,JGRAPH_4:89;
          A82: q33`2>=0 by A11,A12,A65,A68,JGRAPH_4:113;
            q22`1<=q33`1 by A69,A80,A81,AXIOMS:22;
         hence q22`2>=0 & LE q22,q33,P by A3,A18,A80,A81,A82,Th57;
        case A83: p2`1>0;
          then A84: q22`1>0 by A11,A12,Th25;
            now per cases;
          case p2=p3;
           hence q22`2>=0 & LE q22,q33,P by A2,A11,A12,A18,A65,A68,JGRAPH_4:113
,JORDAN6:71;
          case p2<>p3;
          then p2`2/|.p2.|>p3`2/|.p3.| by A1,A6,A7,A65,A83,Th53;
          then q22`2/|.q22.|>q33`2/|.q33.|
                 by A6,A7,A11,A12,A65,A83,Th27;
         hence q22`2>=0 & LE q22,q33,P
                    by A1,A15,A16,A17,A18,A69,A84,Th58,AXIOMS:22;
        end;
         hence q22`2>=0 & LE q22,q33,P;
        end;
       hence q22`2>=0 & LE q22,q33,P & q33`2>=0 & LE q33,q44,P
                   by A1,A18,A20,A23,A69,A75,Th57;
       case A85: p3=p4;
         then A86: p3`2/|.p3.|>=r1 by A7,A10;
         then A87: q33`1>0 & q33`2>=0 by A11,A12,A65,JGRAPH_4:113;
        A88: now assume p2`1=0;
          then 1^2=0+(p2`2)^2 by A6,JGRAPH_3:10,SQUARE_1:60;
         hence p2`2=1 or p2`2=-1 by JGRAPH_3:1;
        end;
        A89: now assume A90: p2`1=0 & p2`2=-1;
          then A91: LE p4,p2,P by A3,A8,A21,Th58;
            LE p2,p4,P by A1,A2,JORDAN6:73;
          then p2=p4 by A2,A91,JORDAN6:72;
         hence contradiction by A1,A90;
        end;
          now per cases by A1,A88,A89;
        case A92: p2`1<=0 & p2`2>=0;
          then A93: q22=p2 by A12,JGRAPH_4:89;
          A94: q33`2>=0 by A11,A12,A65,A86,JGRAPH_4:113;
            q22`1<=q33`1 by A87,A92,A93,AXIOMS:22;
         hence q22`2>=0 & LE q22,q33,P by A3,A18,A92,A93,A94,Th57;
        case A95: p2`1>0;
          then A96: q22`1>0 by A11,A12,Th25;
            now per cases;
          case p2=p3;
           hence q22`2>=0 & LE q22,q33,P by A2,A11,A12,A18,A65,A86,JGRAPH_4:113
,JORDAN6:71;
          case p2<>p3;
          then p2`2/|.p2.|>p3`2/|.p3.| by A1,A6,A7,A65,A95,Th53;
          then q22`2/|.q22.|>q33`2/|.q33.|
                 by A6,A7,A11,A12,A65,A95,Th27;
         hence q22`2>=0 & LE q22,q33,P
               by A1,A15,A16,A17,A18,A87,A96,Th58,AXIOMS:22;
        end;
         hence q22`2>=0 & LE q22,q33,P;
        end;
       hence q22`2>=0 & LE q22,q33,P & q33`2>=0 & LE q33,q44,P
                   by A1,A18,A23,A85,Th57;
       end;
        A97: now assume p1`1=0;
          then 1^2=0+(p1`2)^2 by A5,JGRAPH_3:10,SQUARE_1:60;
         hence p1`2=1 or p1`2=-1 by JGRAPH_3:1;
        end;
        A98: now assume A99: p1`1=0 & p1`2=-1;
          then A100: LE p4,p1,P by A3,A8,A21,Th58;
            LE p1,p3,P by A1,A2,JORDAN6:73;
          then LE p1,p4,P by A1,A2,JORDAN6:73;
          then p1=p4 by A2,A100,JORDAN6:72;
         hence contradiction by A1,A99;
        end;
        A101: now assume p2`1=0;
          then 1^2=0+(p2`2)^2 by A6,JGRAPH_3:10,SQUARE_1:60;
         hence p2`2=1 or p2`2=-1 by JGRAPH_3:1;
        end;
        A102: now assume A103: p2`1=0 & p2`2=-1;
          then p2`2<=p4`2 by A4,Th3;
          then A104: LE p4,p2,P by A3,A21,A103,Th58;
            LE p2,p4,P by A1,A2,JORDAN6:73;
          then p2=p4 by A2,A104,JORDAN6:72;
         hence contradiction by A1,A103;
        end;
          now per cases by A1,A97,A98;
        case A105: p1`1<=0 & p1`2>=0;
          then A106: p1=q11 by A12,JGRAPH_4:89;
          A107: q11`2>=0 by A12,A105,JGRAPH_4:89;
            now per cases by A1,A101,A102;
          case p2`1<=0 & p2`2>=0;
           hence q11`2>=0 & LE q11,q22,P by A1,A12,A105,A106,JGRAPH_4:89;
          case p2`1>0;
            then q11`1<q22`1 by A11,A12,A105,A106,Th25;
           hence q11`2>=0 & LE q11,q22,P by A1,A14,A16,A66,A107,Th57;
          end;
         hence q11`2>=0 & LE q11,q22,P;
        case A108: p1`1>0;
          then A109: q11`1>0 by A11,A12,Th25;
            now per cases by A1,A101,A102;
          case A110: p2`1<=0 & p2`2>=0;
            then A111: p2`1<p1`1 by A108;
              now assume A112: p1`2<0;
              then A113: p1 in Lower_Arc(P) by A3,A36;
              A114: p2 in Upper_Arc(P) by A3,A37,A110;
                W-min(P)=|[-1,0]| by A1,Th32;
              then p1<>W-min(P) by A112,EUCLID:56;
              then LE p2,p1,P by A113,A114,JORDAN6:def 10;
             hence contradiction by A1,A2,A108,A110,JORDAN6:72;
            end;
            then LE p2,p1,P by A3,A110,A111,Th57;
            then q11=q22 by A1,A2,JORDAN6:72;
           hence q11`2>=0 & LE q11,q22,P by A2,A12,A14,A110,JGRAPH_4:89,JORDAN6
:71;
          case A115: p2`1>0;
            then A116: q22`1>0 by A11,A12,Th25;
            now per cases;
          case p1=p2;
           hence q11`2>=0 & LE q11,q22,P by A2,A14,A66,JORDAN6:71;
          case p1<>p2;
            then p1`2/|.p1.|>p2`2/|.p2.| by A1,A5,A6,A108,A115,Th53;
            then q11`2/|.q11.|>q22`2/|.q22.|
                 by A5,A6,A11,A12,A108,A115,Th27;
           hence q11`2>=0 & LE q11,q22,P
               by A1,A13,A14,A15,A16,A66,A109,A116,Th58,AXIOMS:22;
          end;
           hence q11`2>=0 & LE q11,q22,P;
          end;
         hence q11`2>=0 & LE q11,q22,P;
        end;
      hence q11`2>=0 & q22`2>=0 &
       q33`2>=0 & q44`2>0 &
       LE q11,q22,P & LE q22,q33,P & LE q33,q44,P by A1,A11,A12,A21,A22,A24,A66
,JGRAPH_4:113;
     end;
        for q being Point of TOP-REAL 2 holds |.(f1.q).|=|.q.|
                             by A12,JGRAPH_4:104;
    hence ex f being map of TOP-REAL 2,TOP-REAL 2,
     q1,q2,q3,q4 being Point of TOP-REAL 2 st
     f is_homeomorphism &
     (for q being Point of TOP-REAL 2 holds |.(f.q).|=|.q.|)&
     q1=f.p1 & q2=f.p2 & q3=f.p3 & q4=f.p4 &
     q1`2>=0 & q2`2>=0 &
     q3`2>=0 & q4`2>0 &
     LE q1,q2,P & LE q2,q3,P & LE q3,q4,P by A12,A38;
    case A117: p4`2>0;
      A118: p3 in Upper_Arc(P) & p4 in Lower_Arc(P)& not p4=W-min(P) or
      p3 in Upper_Arc(P) & p4 in Upper_Arc(P) &
       LE p3,p4,Upper_Arc(P),W-min(P),E-max(P) or
      p3 in Lower_Arc(P) & p4 in Lower_Arc(P)& not p4=W-min(P) &
       LE p3,p4,Lower_Arc(P),E-max(P),W-min(P) by A1,JORDAN6:def 10;
      A119: Upper_Arc(P)={p7 where p7 is Point of TOP-REAL 2:p7 in P & p7`2>=0}
                      by A1,Th37;
      A120: Lower_Arc(P)={p8 where p8 is Point of TOP-REAL 2:p8 in P & p8`2<=0}
                      by A1,Th38;
      A121: now assume p4 in Lower_Arc(P);
        then consider p9 being Point of TOP-REAL 2 such that
        A122: p9=p4 & p9 in P & p9`2<=0 by A120;
       thus contradiction by A117,A122;
      end;
      then consider p33 being Point of TOP-REAL 2 such that
      A123: p33=p3 & p33 in P & p33`2>=0 by A118,A119;
      A124: LE p2,p4,P by A1,A2,JORDAN6:73;
      then p2 in Upper_Arc(P) & p4 in Lower_Arc(P)& not p4=W-min(P) or
      p2 in Upper_Arc(P) & p4 in Upper_Arc(P) &
       LE p2,p4,Upper_Arc(P),W-min(P),E-max(P) or
      p2 in Lower_Arc(P) & p4 in Lower_Arc(P)& not p4=W-min(P) &
       LE p2,p4,Lower_Arc(P),E-max(P),W-min(P)
                 by JORDAN6:def 10;
      then consider p22 being Point of TOP-REAL 2 such that
      A125: p22=p2 & p22 in P & p22`2>=0 by A119,A121;
        LE p1,p4,P by A1,A2,A124,JORDAN6:73;
      then p1 in Upper_Arc(P) & p4 in Lower_Arc(P)& not p4=W-min(P) or
      p1 in Upper_Arc(P) & p4 in Upper_Arc(P) &
       LE p1,p4,Upper_Arc(P),W-min(P),E-max(P) or
      p1 in Lower_Arc(P) & p4 in Lower_Arc(P)& not p4=W-min(P) &
       LE p1,p4,Lower_Arc(P),E-max(P),W-min(P)
                 by JORDAN6:def 10;
      then consider p11 being Point of TOP-REAL 2 such that
      A126: p11=p1 & p11 in P & p11`2>=0 by A119,A121;
      set f4=id (TOP-REAL 2);
      A127: for q being Point of TOP-REAL 2 holds f4.q=q
       proof let q be Point of TOP-REAL 2;
           f4=id (the carrier of TOP-REAL 2) by GRCAT_1:def 11;
        hence f4.q=q by FUNCT_1:35;
       end;
      then A128: f4.p1=p1 & f4.p2=p2 & f4.p3=p3 & f4.p4=p4;
        (for q being Point of TOP-REAL 2 holds |.(f4.q).|=|.q.|) by A127;
    hence ex f being map of TOP-REAL 2,TOP-REAL 2,
     q1,q2,q3,q4 being Point of TOP-REAL 2 st
     f is_homeomorphism &
     (for q being Point of TOP-REAL 2 holds |.(f.q).|=|.q.|)&
     q1=f.p1 & q2=f.p2 & q3=f.p3 & q4=f.p4 &
     q1`2>=0 & q2`2>=0 &
     q3`2>=0 & q4`2>0 &
     LE q1,q2,P & LE q2,q3,P & LE q3,q4,P by A1,A117,A123,A125,A126,A128;
    end;
    hence thesis;
 end;

theorem Th66: for p1,p2,p3,p4 being Point of TOP-REAL 2,
P being compact non empty 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
  & (p1`2>=0 or p1`1>=0)& (p2`2>=0 or p2`1>=0)
  & (p3`2>=0 or p3`1>=0)& (p4`2>0 or p4`1>0)
 ex f being map of TOP-REAL 2,TOP-REAL 2,
   q1,q2,q3,q4 being Point of TOP-REAL 2 st
   f is_homeomorphism &
   (for q being Point of TOP-REAL 2 holds |.(f.q).|=|.q.|)&
   q1=f.p1 & q2=f.p2 & q3=f.p3 & q4=f.p4 &
   (q1`1<0 & q1`2<0)&(q2`1<0 & q2`2<0)&(q3`1<0 & q3`2<0)&(q4`1<0 & q4`2<0)&
   LE q1,q2,P & LE q2,q3,P & LE q3,q4,P
proof let p1,p2,p3,p4 be Point of TOP-REAL 2,
  P be compact non empty 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
  & (p1`2>=0 or p1`1>=0)& (p2`2>=0 or p2`1>=0)
  & (p3`2>=0 or p3`1>=0)& (p4`2>0 or p4`1>0);
   then consider f being map of TOP-REAL 2,TOP-REAL 2,
     q1,q2,q3,q4 being Point of TOP-REAL 2 such that
   A2: f is_homeomorphism &
    (for q being Point of TOP-REAL 2 holds |.(f.q).|=|.q.|)&
    q1=f.p1 & q2=f.p2 & q3=f.p3 & q4=f.p4 &
    q1`2>=0 & q2`2>=0 &
    q3`2>=0 & q4`2>0 &
    LE q1,q2,P & LE q2,q3,P & LE q3,q4,P by Th65;
  consider f2 being map of TOP-REAL 2,TOP-REAL 2,
   q1b,q2b,q3b,q4b being Point of TOP-REAL 2 such that
   A3: f2 is_homeomorphism &
   (for q being Point of TOP-REAL 2 holds |.(f2.q).|=|.q.|)&
   q1b=f2.q1 & q2b=f2.q2 & q3b=f2.q3 & q4b=f2.q4 &
   (q1b`1<0 & q1b`2<0)&(q2b`1<0 & q2b`2<0)&
   (q3b`1<0 & q3b`2<0)&(q4b`1<0 & q4b`2<0)&
   LE q1b,q2b,P & LE q2b,q3b,P & LE q3b,q4b,P by A1,A2,Th64;
  reconsider f3=f2*f as map of TOP-REAL 2,TOP-REAL 2;
  A4: f3 is_homeomorphism by A2,A3,TOPS_2:71;
  A5: for q being Point of TOP-REAL 2 holds |.(f3.q).|=|.q.|
   proof let q be Point of TOP-REAL 2;
       dom f=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
     then f3.q=f2.(f.q) by FUNCT_1:23;
    hence |.f3.q.|=|.(f.q).| by A3 .=|.q.| by A2;
   end;
  A6: dom f=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
  then A7: f3.p1=q1b by A2,A3,FUNCT_1:23;
  A8: f3.p2=q2b by A2,A3,A6,FUNCT_1:23;
  A9: f3.p3=q3b by A2,A3,A6,FUNCT_1:23;
    f3.p4=q4b by A2,A3,A6,FUNCT_1:23;
 hence thesis by A3,A4,A5,A7,A8,A9;
end;

theorem for p1,p2,p3,p4 being Point of TOP-REAL 2,
P being compact non empty Subset of TOP-REAL 2
 st P={p where p is Point of TOP-REAL 2: |.p.|=1}
 & p4=W-min(P) & LE p1,p2,P & LE p2,p3,P & LE p3,p4,P
 ex f being map of TOP-REAL 2,TOP-REAL 2,
   q1,q2,q3,q4 being Point of TOP-REAL 2 st
   f is_homeomorphism &
   (for q being Point of TOP-REAL 2 holds |.(f.q).|=|.q.|)&
   q1=f.p1 & q2=f.p2 & q3=f.p3 & q4=f.p4 &
   (q1`1<0 & q1`2<0)&(q2`1<0 & q2`2<0)&(q3`1<0 & q3`2<0)&(q4`1<0 & q4`2<0)&
   LE q1,q2,P & LE q2,q3,P & LE q3,q4,P
proof
  let p1,p2,p3,p4 be Point of TOP-REAL 2,
   P be compact non empty Subset of TOP-REAL 2;
   assume A1: P={p where p is Point of TOP-REAL 2: |.p.|=1}
    & p4=W-min(P) &
    LE p1,p2,P & LE p2,p3,P & LE p3,p4,P;
     then A2: P is_simple_closed_curve by JGRAPH_3:36;
     then A3: Upper_Arc(P) is_an_arc_of W-min(P),E-max(P) by JORDAN6:def 8;
    A4: P={p where p is Point of TOP-REAL 2: |.p.|=1}
    & p1 in P & p2 in P & p3 in P & p4 in P
    & p4=W-min(P) &
    LE p1,p2,P & LE p2,p3,P & LE p3,p4,P by A1,A2,JORDAN7:5;
     A5: Upper_Arc(P)
         ={p7 where p7 is Point of TOP-REAL 2:p7 in P & p7`2>=0}
                          by A1,Th37;
       W-min(P)=|[-1,0]| by A1,Th32;
     then A6: (W-min(P))`1=-1 & (W-min(P))`2=0 by EUCLID:56;
     then A7: p4 in Upper_Arc(P) by A4,A5;
     then A8: p3 in Upper_Arc(P) by A1,Th47;
     then LE p4,p3,Upper_Arc(P),W-min(P),E-max(P) by A1,A3,JORDAN5C:10;
     then LE p4,p3,P by A7,A8,JORDAN6:def 10;
     then A9: p3=p4 by A1,A2,JORDAN6:72;
     A10: p2 in Upper_Arc(P) by A1,A8,Th47;
     then LE p4,p2,Upper_Arc(P),W-min(P),E-max(P) by A1,A3,JORDAN5C:10;
     then A11: LE p4,p2,P by A7,A10,JORDAN6:def 10;
       LE p2,p4,P by A1,A2,JORDAN6:73;
     then A12: p2=p4 by A2,A11,JORDAN6:72;
     A13: p1 in Upper_Arc(P) by A1,A10,Th47;
     then LE p4,p1,Upper_Arc(P),W-min(P),E-max(P) by A1,A3,JORDAN5C:10;
     then LE p4,p1,P by A7,A13,JORDAN6:def 10;
     then p1=p4 by A1,A2,A12,JORDAN6:72;
 hence thesis by A1,A6,A9,A12,Th62;
end;

theorem Th68: for p1,p2,p3,p4 being Point of TOP-REAL 2,
P being compact non empty 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
 ex f being map of TOP-REAL 2,TOP-REAL 2,
   q1,q2,q3,q4 being Point of TOP-REAL 2 st
   f is_homeomorphism &
   (for q being Point of TOP-REAL 2 holds |.(f.q).|=|.q.|)&
   q1=f.p1 & q2=f.p2 & q3=f.p3 & q4=f.p4 &
   (q1`1<0 & q1`2<0)&(q2`1<0 & q2`2<0)&(q3`1<0 & q3`2<0)&(q4`1<0 & q4`2<0)&
   LE q1,q2,P & LE q2,q3,P & LE q3,q4,P
 proof let p1,p2,p3,p4 be Point of TOP-REAL 2,
   P be compact non empty 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;
     then A2: P is_simple_closed_curve by JGRAPH_3:36;
     then A3: Upper_Arc(P) is_an_arc_of W-min(P),E-max(P) by JORDAN6:def 8;
     A4: P={p where p is Point of TOP-REAL 2: |.p.|=1}
     & p1 in P & p2 in P & p3 in P & p4 in P &
     LE p1,p2,P & LE p2,p3,P & LE p3,p4,P by A1,A2,JORDAN7:5;
     then consider p44 being Point of TOP-REAL 2 such that
     A5: p44=p4 & |.p44.|=1;
     A6: -1<=p4`1 & p4`1<=1 by A5,Th3;
     A7: Lower_Arc(P)
         ={p7 where p7 is Point of TOP-REAL 2:p7 in P & p7`2<=0} by A1,Th38;
     A8: Upper_Arc(P)
         ={p7 where p7 is Point of TOP-REAL 2:p7 in P & p7`2>=0}
                          by A1,Th37;
     A9: W-min(P)=|[-1,0]| by A1,Th32;
     then A10: (W-min(P))`1=-1 & (W-min(P))`2=0 by EUCLID:56;
       now per cases;
     case A11: p4`1=-1;
         1=(p4`1)^2+(p4`2)^2 by A5,JGRAPH_3:10,SQUARE_1:59
       .=(p4`2)^2+1 by A11,SQUARE_1:59,61;
       then 1-1=(p4`2)^2 by XCMPLX_1:26;
       then A12: p4`2=0 by SQUARE_1:73;
       then A13: p4 in Upper_Arc(P) by A4,A8;
       A14: p4=W-min(P) by A9,A11,A12,EUCLID:57;
       A15: now per cases;
       case A16: p1 in Upper_Arc(P);
         then LE p4,p1,Upper_Arc(P),W-min(P),E-max(P)
                     by A3,A14,JORDAN5C:10;
        hence LE p4,p1,P by A13,A16,JORDAN6:def 10;
       case not p1 in Upper_Arc(P);
         then A17: p1`2<0 by A4,A8;
         then p1 in Lower_Arc(P) by A4,A7;
        hence LE p4,p1,P by A10,A13,A17,JORDAN6:def 10;
       end;
          LE p1,p3,P by A1,A2,JORDAN6:73;
        then LE p1,p4,P by A1,A2,JORDAN6:73;
        then A18: p4=p1 by A2,A15,JORDAN6:72;
        A19: LE p4,p2,P by A1,A2,A15,JORDAN6:73;
          LE p2,p4,P by A1,A2,JORDAN6:73;
        then A20: p2=p4 by A2,A19,JORDAN6:72;
          LE p4,p3,P by A1,A2,A19,JORDAN6:73;
        then p3=p4 by A1,A2,JORDAN6:72;
       hence ex f being map of TOP-REAL 2,TOP-REAL 2,
        q1,q2,q3,q4 being Point of TOP-REAL 2 st
        f is_homeomorphism &
       (for q being Point of TOP-REAL 2 holds |.(f.q).|=|.q.|)&
        q1=f.p1 & q2=f.p2 & q3=f.p3 & q4=f.p4 &
       (q1`1<0 & q1`2<0)&(q2`1<0 & q2`2<0)&(q3`1<0 & q3`2<0)&(q4`1<0 & q4`2<0)&
        LE q1,q2,P & LE q2,q3,P & LE q3,q4,P by A1,A11,A12,A18,A20,Th62;
     case A21: p4`1<>-1;
     then p4`1> -1 by A6,REAL_1:def 5;
     then consider r being real number such that
     A22: -1<r & r<p4`1 by REAL_1:75;
     reconsider r1=r as Real by XREAL_0:def 1;
     A23: -1<r1 & r1<1 by A6,A22,AXIOMS:22;
     then consider f1 being map of TOP-REAL 2,TOP-REAL 2 such that
     A24: f1=r1-FanMorphS & f1 is_homeomorphism by JGRAPH_4:143;
     set q11=f1.p1, q22=f1.p2, q33=f1.p3, q44=f1.p4;
       now per cases;
     case A25: (p4`1>0 or p4`2>=0);
       then A26: p3`1>=0 or p3`2>=0 by A1,Th52;
       then A27: p2`1>=0 or p2`2>=0 by A1,Th52;
       then A28: p1`1>=0 or p1`2>=0 by A1,Th52;
         now assume A29: p4`2=0 & p4`1<=0;
           1^2 =(p4`1)^2+(p4`2)^2 by A5,JGRAPH_3:10
         .=(p4`1)^2 by A29,SQUARE_1:60;
         then p4`1=1 or p4`1=-1 by JGRAPH_3:1;
        hence contradiction by A21,A29;
       end;
      hence ex f being map of TOP-REAL 2,TOP-REAL 2,
       q1,q2,q3,q4 being Point of TOP-REAL 2 st
       f is_homeomorphism &
       (for q being Point of TOP-REAL 2 holds |.(f.q).|=|.q.|)&
       q1=f.p1 & q2=f.p2 & q3=f.p3 & q4=f.p4 &
     (q1`1<0 & q1`2<0)&(q2`1<0 & q2`2<0)&(q3`1<0 & q3`2<0)&(q4`1<0 & q4`2<0)&
      LE q1,q2,P & LE q2,q3,P & LE q3,q4,P by A1,A25,A26,A27,A28,Th66;
     case A30: p4`1<=0 & p4`2<0;
         p4`1/|.p4.|>r1 by A5,A22;
       then A31: q44`1>0 & q44`2<0 by A23,A24,A30,Th29;
       A32: LE q33,q44,P by A1,A23,A24,Th61;
       A33: LE q22,q33,P by A1,A23,A24,Th61;
       then A34: LE q22,q44,P by A2,A32,JORDAN6:73;
       A35: LE q11,q22,P by A1,A23,A24,Th61;
       then A36: LE q11,q44,P by A2,A34,JORDAN6:73;
         W-min(P)=|[-1,0]| by A1,Th32;
       then A37: (W-min(P))`2=0 by EUCLID:56;
       A38: now per cases;
       case q33`2>=0;
        hence q33`2>=0 or q33`1>=0;
       case q33`2<0;
         then q33`1>=q44`1 by A1,A31,A32,A37,Th51;
        hence q33`2>=0 or q33`1>=0 by A31,AXIOMS:22;
       end;
       A39: now per cases;
       case q22`2>=0;
        hence q22`2>=0 or q22`1>=0;
       case q22`2<0;
         then q22`1>=q44`1 by A1,A31,A34,A37,Th51;
        hence q22`2>=0 or q22`1>=0 by A31,AXIOMS:22;
       end;
         now per cases;
       case q11`2>=0;
        hence q11`2>=0 or q11`1>=0;
       case q11`2<0;
         then q11`1>=q44`1 by A1,A31,A36,A37,Th51;
        hence q11`2>=0 or q11`1>=0 by A31,AXIOMS:22;
       end;
       then consider f2 being map of TOP-REAL 2,TOP-REAL 2,
         q81,q82,q83,q84 being Point of TOP-REAL 2 such that
       A40: f2 is_homeomorphism &
        (for q being Point of TOP-REAL 2 holds |.(f2.q).|=|.q.|)&
        q81=f2.q11 & q82=f2.q22 & q83=f2.q33 & q84=f2.q44 &
        (q81`1<0 & q81`2<0)&(q82`1<0 & q82`2<0)&(q83`1<0 & q83`2<0)
        &(q84`1<0 & q84`2<0)&
        LE q81,q82,P & LE q82,q83,P & LE q83,q84,P
             by A1,A31,A32,A33,A35,A38,A39,Th66;
        reconsider f3=f2*f1 as map of TOP-REAL 2,TOP-REAL 2;
        A41: f3 is_homeomorphism by A24,A40,TOPS_2:71;
        A42: for q being Point of TOP-REAL 2 holds |.(f3.q).|=|.q.|
         proof let q be Point of TOP-REAL 2;
             dom f1=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
           then f3.q=f2.(f1.q) by FUNCT_1:23;
          hence |.f3.q.|=|.(f1.q).| by A40 .=|.q.| by A24,JGRAPH_4:135;
         end;
        A43: dom f1=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
        then A44: f3.p1=q81 by A40,FUNCT_1:23;
        A45: f3.p2=q82 by A40,A43,FUNCT_1:23;
        A46: f3.p3=q83 by A40,A43,FUNCT_1:23;
          f3.p4=q84 by A40,A43,FUNCT_1:23;
      hence ex f being map of TOP-REAL 2,TOP-REAL 2,
       q1,q2,q3,q4 being Point of TOP-REAL 2 st
       f is_homeomorphism &
       (for q being Point of TOP-REAL 2 holds |.(f.q).|=|.q.|)&
       q1=f.p1 & q2=f.p2 & q3=f.p3 & q4=f.p4 &
    (q1`1<0 & q1`2<0)&(q2`1<0 & q2`2<0)&(q3`1<0 & q3`2<0)&(q4`1<0 & q4`2<0)&
       LE q1,q2,P & LE q2,q3,P & LE q3,q4,P
                  by A40,A41,A42,A44,A45,A46;
     end;
   hence ex f being map of TOP-REAL 2,TOP-REAL 2,
    q1,q2,q3,q4 being Point of TOP-REAL 2 st
    f is_homeomorphism &
    (for q being Point of TOP-REAL 2 holds |.(f.q).|=|.q.|)&
    q1=f.p1 & q2=f.p2 & q3=f.p3 & q4=f.p4 &
    (q1`1<0 & q1`2<0)&(q2`1<0 & q2`2<0)&(q3`1<0 & q3`2<0)&(q4`1<0 & q4`2<0)&
    LE q1,q2,P & LE q2,q3,P & LE q3,q4,P;
     end;
   hence thesis;
 end;

begin :: General Fashoda Meet Theorems

theorem Th69: for p1,p2,p3,p4 being Point of TOP-REAL 2,
P being compact non empty 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 &
  p1<>p2 & p2<>p3 & p3<>p4 &
  p1`1<0 & p2`1<0 & p3`1<0 & p4`1<0 &
  p1`2<0 & p2`2<0 & p3`2<0 & p4`2<0
 ex f being map of TOP-REAL 2,TOP-REAL 2 st
   f is_homeomorphism &
   (for q being Point of TOP-REAL 2 holds |.(f.q).|=|.q.|)&
   |[-1,0]|=f.p1 & |[0,1]|=f.p2 & |[1,0]|=f.p3 & |[0,-1]|=f.p4
 proof let p1,p2,p3,p4 be Point of TOP-REAL 2,
    P be compact non empty 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 &
     p1<>p2 & p2<>p3 & p3<>p4 &
     p1`1<0 & p2`1<0 & p3`1<0 & p4`1<0 &
     p1`2<0 & p2`2<0 & p3`2<0 & p4`2<0;
     then A2: p1`1>p2`1 & p1`2<p2`2 by Th48;
     A3: p2`1>p3`1 & p2`2<p3`2 by A1,Th48;
     A4: p3`1>p4`1 & p3`2<p4`2 by A1,Th48;
       P is_simple_closed_curve by A1,JGRAPH_3:36;
     then A5: P={p where p is Point of TOP-REAL 2: |.p.|=1}
     & p1 in P & p2 in P & p3 in P & p4 in P &
     LE p1,p2,P & LE p2,p3,P & LE p3,p4,P &
     p1<>p2 & p2<>p3 & p3<>p4 &
     p1`1<0 & p2`1<0 & p3`1<0 & p4`1<0 &
     p1`2<0 & p2`2<0 & p3`2<0 & p4`2<0 by A1,JORDAN7:5;
     then consider p11 being Point of TOP-REAL 2 such that
         A6: p11=p1 & |.p11.|=1;
     consider p22 being Point of TOP-REAL 2 such that
         A7: p22=p2 & |.p22.|=1 by A5;
     consider p33 being Point of TOP-REAL 2 such that
         A8: p33=p3 & |.p33.|=1 by A5;
     consider p44 being Point of TOP-REAL 2 such that
         A9: p44=p4 & |.p44.|=1 by A5;
     A10: -1<p1`2 & p1`2<1 by A1,A6,Th4;
    then consider f1 being map of TOP-REAL 2,TOP-REAL 2 such that
    A11: f1=((p1`2)-FanMorphW) & f1 is_homeomorphism by JGRAPH_4:48;
     A12: p1`2/|.p1.|=p1`2 by A6;
     A13: p2`2/|.p2.|=p2`2 by A7;
     set q1=((p1`2)-FanMorphW).p1;
     set q2=((p1`2)-FanMorphW).p2;
     set q3=((p1`2)-FanMorphW).p3;
     set q4=((p1`2)-FanMorphW).p4;
     A14: |.q1.|=1 by A6,JGRAPH_4:40;
     then A15: q1`2/|.q1.|=q1`2;
     A16: |.q2.|=1 by A7,JGRAPH_4:40;
     then A17: q2`2/|.q2.|=q2`2;
     A18: |.q3.|=1 by A8,JGRAPH_4:40;
     then A19: q3`2/|.q3.|=q3`2;
     A20: |.q4.|=1 by A9,JGRAPH_4:40;
     then A21: q4`2/|.q4.|=q4`2;
     A22: q1`1<0 & q1`2=0 by A1,A10,A12,JGRAPH_4:54;
     A23: q2`1<0 & q2`2>=0 by A1,A2,A10,A13,JGRAPH_4:49;
     A24: p3`2/|.p3.|>p2`2 by A1,A8,Th48;
     A25: p3`2/|.p3.|>p1`2 by A2,A3,A8,AXIOMS:22;
     then A26: q3`1<0 & q3`2>=0 by A1,A10,JGRAPH_4:49;
     A27: p4`2/|.p4.|>p3`2 by A1,A9,Th48;
       p4`2/|.p4.|>p2`2 by A3,A4,A9,AXIOMS:22;
     then A28: p4`2/|.p4.|>p1`2 by A2,AXIOMS:22;
     A29: q1`2<q2`2 by A1,A2,A10,A12,A13,A15,A17,JGRAPH_4:51;
     A30: q1`2<q2`2 & q2`2<q3`2 & q3`2<q4`2
            by A1,A2,A8,A10,A12,A13,A15,A17,A19,A21,A24,A25,A27,A28,JGRAPH_4:51
;
     A31: 0<q3`2 by A1,A10,A12,A15,A19,A22,A25,JGRAPH_4:51;
     A32: 0<q4`2 by A1,A10,A12,A15,A21,A22,A28,JGRAPH_4:51;
     A33: 1^2=(q2`1)^2+(q2`2)^2 by A16,JGRAPH_3:10;
     A34: -q2`1>0 by A23,REAL_1:66;
     A35: 1^2=(q3`1)^2+(q3`2)^2 by A18,JGRAPH_3:10;
     then (q2`1)^2=(q3`1)^2+(q3`2)^2-(q2`2)^2 by A33,XCMPLX_1:26
             .=(q3`1)^2+((q3`2)^2-(q2`2)^2) by XCMPLX_1:29;
     then A36: (q2`1)^2-(q3`1)^2=(q3`2)^2-(q2`2)^2 by XCMPLX_1:26;
       (q3`2)^2>(q2`2)^2 by A22,A30,SQUARE_1:78;
     then (q3`2)^2-(q2`2)^2>0 by SQUARE_1:11;
     then (q2`1)^2-(q3`1)^2+(q3`1)^2>0+(q3`1)^2 by A36,REAL_1:67;
     then (q2`1)^2>(q3`1)^2 by XCMPLX_1:27;
     then (-(q2`1))^2>(q3`1)^2 by SQUARE_1:61;
     then A37: --(q2`1)<(q3`1) & q3`1<-(q2`1) by A34,JGRAPH_2:6;
     A38: -q3`1>0 by A26,REAL_1:66;
       1^2=(q4`1)^2+(q4`2)^2 by A20,JGRAPH_3:10;
     then (q3`1)^2=(q4`1)^2+(q4`2)^2-(q3`2)^2 by A35,XCMPLX_1:26
             .=(q4`1)^2+((q4`2)^2-(q3`2)^2) by XCMPLX_1:29;
     then A39: (q3`1)^2-(q4`1)^2=(q4`2)^2-(q3`2)^2 by XCMPLX_1:26;
       (q4`2)^2>(q3`2)^2 by A30,A31,SQUARE_1:78;
     then (q4`2)^2-(q3`2)^2>0 by SQUARE_1:11;
     then (q3`1)^2-(q4`1)^2+(q4`1)^2>0+(q4`1)^2 by A39,REAL_1:67;
     then (q3`1)^2>(q4`1)^2 by XCMPLX_1:27;
     then (-(q3`1))^2>(q4`1)^2 by SQUARE_1:61;
     then A40: --(q3`1)<(q4`1) & q4`1<-(q3`1) by A38,JGRAPH_2:6;
        (|.q1.|)^2 =(q1`1)^2+(q1`2)^2 by JGRAPH_3:10;
      then A41: q1`1=-1 or q1`1=1 by A14,A22,JGRAPH_3:1,SQUARE_1:60;
     then A42: q1=|[-1,0]| by A22,EUCLID:57;
     A43: -1<q2`1 & q2`1<1 by A16,A22,A23,A30,Th4;
    then consider f2 being map of TOP-REAL 2,TOP-REAL 2 such that
    A44: f2=((q2`1)-FanMorphN) & f2 is_homeomorphism by JGRAPH_4:81;
     A45: q2`1/|.q2.|=q2`1 by A16;
     A46: q3`1/|.q3.|=q3`1 by A18;
     set r1=((q2`1)-FanMorphN).q1;
     set r2=((q2`1)-FanMorphN).q2;
     set r3=((q2`1)-FanMorphN).q3;
     set r4=((q2`1)-FanMorphN).q4;
     A47: |.r2.|=1 by A16,JGRAPH_4:73;
     then A48: r2`1/|.r2.|=r2`1;
     A49: |.r3.|=1 by A18,JGRAPH_4:73;
     then A50: r3`1/|.r3.|=r3`1;
     A51: |.r4.|=1 by A20,JGRAPH_4:73;
     then A52: r4`1/|.r4.|=r4`1;
     A53: r2`2>0 & r2`1=0 by A22,A30,A43,A45,JGRAPH_4:87;
     A54: r3`2>0 & r3`1>=0 by A31,A37,A43,A46,JGRAPH_4:82;
     A55: q4`1/|.q4.|>q3`1 by A20,A40;
     A56: q4`1/|.q4.|>q2`1 by A20,A37,A40,AXIOMS:22;
     A57: r1=|[-1,0]| by A22,A42,JGRAPH_4:56;
     A58: r1`1=-1 & r1`2=0 by A22,A41,JGRAPH_4:56;
        (|.r2.|)^2 =(r2`1)^2+(r2`2)^2 by JGRAPH_3:10;
      then A59: r2`2=-1 or r2`2=1 by A47,A53,JGRAPH_3:1,SQUARE_1:60;
     then A60: r2=|[0,1]| by A53,EUCLID:57;
     A61: r2`1<r3`1 by A22,A30,A31,A37,A43,A45,A46,A48,A50,JGRAPH_4:86;
     A62: r2`1<r3`1 & r3`1<r4`1
           by A22,A29,A31,A32,A37,A43,A45,A46,A48,A50,A52,A55,JGRAPH_4:86;
     then A63: 0<r4`1 by A53,AXIOMS:22;
     A64:  r2`2>0 & r3`2>0 & r4`2>0 by A22,A29,A31,A32,A37,A43,A45,A46,A56,
JGRAPH_4:82;
     A65: 1^2=(r2`1)^2+(r2`2)^2 by A47,JGRAPH_3:10;
     A66: 1^2=(r3`1)^2+(r3`2)^2 by A49,JGRAPH_3:10;
     then (r2`2)^2=(r3`2)^2+(r3`1)^2-(r2`1)^2 by A65,XCMPLX_1:26
             .=(r3`2)^2+((r3`1)^2-(r2`1)^2) by XCMPLX_1:29;
     then A67: (r2`2)^2-(r3`2)^2=(r3`1)^2-(r2`1)^2 by XCMPLX_1:26;
       (r3`1)^2>(r2`1)^2 by A53,A62,SQUARE_1:78;
     then (r3`1)^2-(r2`1)^2>0 by SQUARE_1:11;
     then (r2`2)^2-(r3`2)^2+(r3`2)^2>0+(r3`2)^2 by A67,REAL_1:67;
     then A68: (r2`2)^2>(r3`2)^2 by XCMPLX_1:27;
       1^2=(r4`1)^2+(r4`2)^2 by A51,JGRAPH_3:10;
     then (r3`2)^2=(r4`2)^2+(r4`1)^2-(r3`1)^2 by A66,XCMPLX_1:26
             .=(r4`2)^2+((r4`1)^2-(r3`1)^2) by XCMPLX_1:29;
     then A69: (r3`2)^2-(r4`2)^2=(r4`1)^2-(r3`1)^2 by XCMPLX_1:26;
       (r4`1)^2>(r3`1)^2 by A53,A62,SQUARE_1:78;
     then (r4`1)^2-(r3`1)^2>0 by SQUARE_1:11;
     then (r3`2)^2-(r4`2)^2+(r4`2)^2>0+(r4`2)^2 by A69,REAL_1:67;
     then ((r3`2))^2>(r4`2)^2 by XCMPLX_1:27;
     then A70: r2`2>r3`2 & r3`2>r4`2 by A64,A68,JGRAPH_2:6;
     A71: -1<r3`2 & r3`2<1 by A49,A53,A54,A61,Th4;
     then consider f3 being map of TOP-REAL 2,TOP-REAL 2 such that
     A72: f3=((r3`2)-FanMorphE) & f3 is_homeomorphism by JGRAPH_4:112;
     A73: r3`2/|.r3.|=r3`2 by A49;
     A74: r4`2/|.r4.|=r4`2 by A51;
     set s1=((r3`2)-FanMorphE).r1;
     set s2=((r3`2)-FanMorphE).r2;
     set s3=((r3`2)-FanMorphE).r3;
     set s4=((r3`2)-FanMorphE).r4;
     A75: |.s3.|=1 by A49,JGRAPH_4:104;
     A76: |.s4.|=1 by A51,JGRAPH_4:104;
     A77: s3`1>0 & s3`2=0 by A53,A62,A71,A73,JGRAPH_4:118;
     A78: s4`1>0 & s4`2<0 by A63,A70,A71,A74,JGRAPH_4:114;
     A79: s1=|[-1,0]| by A57,A58,JGRAPH_4:89;
     A80: s1`1=-1 & s1`2=0 by A58,JGRAPH_4:89;
     A81: s2=|[0,1]| by A53,A60,JGRAPH_4:89;
     A82: s2`1=0 & s2`2=1 by A53,A59,JGRAPH_4:89;
        (|.s3.|)^2 =(s3`1)^2+(s3`2)^2 by JGRAPH_3:10;
      then s3`1=-1 or s3`1=1 by A75,A77,JGRAPH_3:1,SQUARE_1:60;
     then A83: s3=|[1,0]| by A77,EUCLID:57;
     A84: s3`2/|.s3.|>s4`2/|.s4.| by A53,A62,A63,A70,A71,A73,A74,JGRAPH_4:117;
     A85: -1<s4`1 & s4`1<1 by A76,A78,Th4;
     then consider f4 being map of TOP-REAL 2,TOP-REAL 2 such that
     A86: f4=((s4`1)-FanMorphS) & f4 is_homeomorphism by JGRAPH_4:143;
     A87: s4`1/|.s4.|=s4`1 by A76;
     set t4=((s4`1)-FanMorphS).s4;
     A88: |.t4.|=1 by A76,JGRAPH_4:135;
     A89: t4`2<0 & t4`1=0 by A76,A77,A84,A85,A87,JGRAPH_4:149;
     A90: ((s4`1)-FanMorphS).s1=|[-1,0]| by A79,A80,JGRAPH_4:120;
     A91: ((s4`1)-FanMorphS).s2=|[0,1]| by A81,A82,JGRAPH_4:120;
     A92: ((s4`1)-FanMorphS).s3=|[1,0]| by A77,A83,JGRAPH_4:120;
        (|.t4.|)^2 =(t4`1)^2+(t4`2)^2 by JGRAPH_3:10;
      then t4`2=-1 or t4`2=1 by A88,A89,JGRAPH_3:1,SQUARE_1:60;
     then A93: t4=|[0,-1]| by A89,EUCLID:57;
     reconsider g=f4*(f3*(f2*f1)) as map of TOP-REAL 2,TOP-REAL 2;
       f2*f1 is_homeomorphism by A11,A44,TOPS_2:71;
     then f3*(f2*f1) is_homeomorphism by A72,TOPS_2:71;
     then A94: g is_homeomorphism by A86,TOPS_2:71;
     A95: dom g=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
     A96: dom (f2*f1)=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
     A97: dom (f3*(f2*f1))=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
     A98: for q being Point of TOP-REAL 2 holds |.(g.q).|=|.q.|
      proof let q be Point of TOP-REAL 2;
        A99: |.((f2*f1).q).|=|.(f2.(f1.q)).| by A96,FUNCT_1:22
        .=|.(f1.q).| by A44,JGRAPH_4:73 .=|.q.| by A11,JGRAPH_4:40;
        A100: |.((f3*(f2*f1)).q).|=|.(f3.((f2*f1).q)).| by A97,FUNCT_1:22
        .=|.q.| by A72,A99,JGRAPH_4:104;
       thus |.(g.q).|=|.(f4.((f3*(f2*f1)).q)).| by A95,FUNCT_1:22
        .=|.q.| by A86,A100,JGRAPH_4:135;
      end;
     A101: g.p1=(f4.((f3*(f2*f1)).p1)) by A95,FUNCT_1:22
     .=f4.((f3.((f2*f1).p1))) by A97,FUNCT_1:22
     .=|[-1,0]| by A11,A44,A72,A86,A90,A96,FUNCT_1:22;
     A102: g.p2=(f4.((f3*(f2*f1)).p2)) by A95,FUNCT_1:22
     .=f4.((f3.((f2*f1).p2))) by A97,FUNCT_1:22
     .=|[0,1]| by A11,A44,A72,A86,A91,A96,FUNCT_1:22;
     A103: g.p3= (f4.((f3*(f2*f1)).p3)) by A95,FUNCT_1:22
     .=f4.((f3.((f2*f1).p3))) by A97,FUNCT_1:22
     .=|[1,0]| by A11,A44,A72,A86,A92,A96,FUNCT_1:22;
       g.p4= (f4.((f3*(f2*f1)).p4)) by A95,FUNCT_1:22
     .=f4.((f3.((f2*f1).p4))) by A97,FUNCT_1:22
     .=|[0,-1]| by A11,A44,A72,A86,A93,A96,FUNCT_1:22;
  hence thesis by A94,A98,A101,A102,A103;
 end;

theorem Th70: for p1,p2,p3,p4 being Point of TOP-REAL 2,
P being compact non empty 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 &
  p1<>p2 & p2<>p3 & p3<>p4
 ex f being map of TOP-REAL 2,TOP-REAL 2 st
   f is_homeomorphism &
   (for q being Point of TOP-REAL 2 holds |.(f.q).|=|.q.|)&
   |[-1,0]|=f.p1 & |[0,1]|=f.p2 & |[1,0]|=f.p3 & |[0,-1]|=f.p4
 proof let p1,p2,p3,p4 be Point of TOP-REAL 2,
   P be compact non empty 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 &
     p1<>p2 & p2<>p3 & p3<>p4;
       then consider f being map of TOP-REAL 2,TOP-REAL 2,
        q1,q2,q3,q4 being Point of TOP-REAL 2 such that
       A2: f is_homeomorphism &
        (for q being Point of TOP-REAL 2 holds |.(f.q).|=|.q.|)&
        q1=f.p1 & q2=f.p2 & q3=f.p3 & q4=f.p4 &
        (q1`1<0 & q1`2<0)&(q2`1<0 & q2`2<0)&
        (q3`1<0 & q3`2<0)&(q4`1<0 & q4`2<0)&
        LE q1,q2,P & LE q2,q3,P & LE q3,q4,P by Th68;
        A3:  dom f=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
        A4: f is one-to-one by A2,TOPS_2:def 5;
        then A5: q1<>q2 by A1,A2,A3,FUNCT_1:def 8;
        A6: q2<>q3 by A1,A2,A3,A4,FUNCT_1:def 8;
          q3<>q4 by A1,A2,A3,A4,FUNCT_1:def 8;
       then consider f2 being map of TOP-REAL 2,TOP-REAL 2 such that
       A7: f2 is_homeomorphism &
        (for q being Point of TOP-REAL 2 holds |.(f2.q).|=|.q.|)&
        |[-1,0]|=f2.q1 & |[0,1]|=f2.q2 & |[1,0]|=f2.q3 & |[0,-1]|=f2.q4
                     by A1,A2,A5,A6,Th69;
       reconsider f3=f2*f as map of TOP-REAL 2,TOP-REAL 2;
       A8: f3 is_homeomorphism by A2,A7,TOPS_2:71;
       A9:  dom f3=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
       A10: for q being Point of TOP-REAL 2 holds |.(f3.q).|=|.q.|
       proof let q be Point of TOP-REAL 2;
           |.(f3.q).|=|.f2.(f.q).| by A9,FUNCT_1:22.=|.(f.q).| by A7
         .=|.q.| by A2;
        hence |.(f3.q).|=|.q.|;
       end;
       A11: f3.p1=|[-1,0]| by A2,A7,A9,FUNCT_1:22;
       A12: f3.p2=|[0,1]| by A2,A7,A9,FUNCT_1:22;
       A13: f3.p3=|[1,0]| by A2,A7,A9,FUNCT_1:22;
         f3.p4=|[0,-1]| by A2,A7,A9,FUNCT_1:22;
      hence thesis by A8,A10,A11,A12,A13;
 end;

Lm6: (|[-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;
Lm7: now thus |.(|[-1,0]|).|=sqrt((-1)^2+0^2) by Lm6,JGRAPH_3:10
            .=1 by SQUARE_1:59,60,61,83;
          thus |.(|[1,0]|).|=sqrt(1+0) by Lm6,JGRAPH_3:10,SQUARE_1:59,60
            .=1 by SQUARE_1:83;
          thus |.(|[0,-1]|).|=sqrt(0^2+(-1)^2) by Lm6,JGRAPH_3:10
            .=1 by SQUARE_1:59,60,61,83;
          thus |.(|[0,1]|).|=sqrt(0^2+1^2) by Lm6,JGRAPH_3:10
            .=1 by SQUARE_1:59,60,83;
      end;
Lm8: 0 in [.0,1.] by TOPREAL5:1;
Lm9: 1 in [.0,1.] by TOPREAL5:1;

theorem 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={p: |.p.|<=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}
     & 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={p: |.p.|<=1}& f.0=p1 & f.1=p3 & 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,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,Lm9,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,Lm9,BORSUK_1:83,FUNCT_1:def 5;
          p4 in rng g by A2,A4,Lm9,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,Th70;
    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.0= |[-1,0]| by A2,A12,A14,Lm8,BORSUK_1:83,FUNCT_1:22;
    A17: g2.0= |[0,1]| by A2,A12,A15,Lm8,BORSUK_1:83,FUNCT_1:22;
    A18: f2.1= |[1,0]| by A2,A12,A14,Lm9,BORSUK_1:83,FUNCT_1:22;
    A19: g2.1= |[0,-1]| by A2,A12,A15,Lm9,BORSUK_1:83,FUNCT_1:22;
    A20: f2 is continuous one-to-one &
     g2 is continuous one-to-one &
     f2.0=|[-1,0]| & f2.1=|[1,0]| & g2.0=|[0,1]| & g2.1= |[0,-1]|
                           by A2,A12,A14,A15,Lm8,Lm9,Th8,Th9,BORSUK_1:83,
FUNCT_1:22;
    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 Lm6;
    then A33: f2.O in KXN by A16,Lm6,Lm7;
    A34: f2.I in KXP by A18,Lm6,Lm7;
       -(|[0,-1]|)`1= 0 by Lm6;
    then A35: g2.I in KYN by A19,Lm6,Lm7;
       -(|[0,1]|)`1= 0 by Lm6;
    then g2.O in KYP by A17,Lm6,Lm7;
    then rng f2 meets rng g2 by A2,A20,A21,A27,A33,A34,A35,Th16;
    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
  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={p: |.p.|<=1}& f.0=p1 & f.1=p3 & 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={p: |.p.|<=1}& f.0=p1 & f.1=p3 & 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,Lm9,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,Lm9,BORSUK_1:83,FUNCT_1:def 5;
          p2 in rng g by A2,A4,Lm9,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,Lm9,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,Th70;
    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.0= |[-1,0]| by A2,A12,A14,Lm8,BORSUK_1:83,FUNCT_1:22;
    A17: g2.0= |[0,-1]| by A2,A12,A15,Lm8,BORSUK_1:83,FUNCT_1:22;
    A18: f2.1= |[1,0]| by A2,A12,A14,Lm9,BORSUK_1:83,FUNCT_1:22;
    A19: g2.1= |[0,1]| by A2,A12,A15,Lm9,BORSUK_1:83,FUNCT_1:22;
    A20: f2 is continuous one-to-one &
     g2 is continuous one-to-one &
     f2.0=|[-1,0]| & f2.1=|[1,0]| & g2.0=|[0,-1]| & g2.1= |[0,1]|
                           by A2,A12,A14,A15,Lm8,Lm9,Th8,Th9,BORSUK_1:83,
FUNCT_1:22;
    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 Lm6;
    then A33: f2.O in KXN by A16,Lm6,Lm7;
    A34: f2.I in KXP by A18,Lm6,Lm7;
       -(|[0,-1]|)`1= 0 by Lm6;
    then A35: g2.O in KYN by A17,Lm6,Lm7;
       -(|[0,1]|)`1= 0 by Lm6;
    then g2.I in KYP by A19,Lm6,Lm7;
    then rng f2 meets rng g2 by A2,A20,A21,A27,A33,A34,A35,JGRAPH_3:55;
    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 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={p: |.p.|>=1}& f.0=p1 & f.1=p3 & 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={p: |.p.|>=1}& f.0=p1 & f.1=p3 & 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,Lm9,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,Lm9,BORSUK_1:83,FUNCT_1:def 5;
          p2 in rng g by A2,A4,Lm9,BORSUK_1:83,FUNCT_1:def 5;
        hence rng f meets rng g 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,Lm9,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,Th70;
    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.0= |[-1,0]| by A2,A12,A14,Lm8,BORSUK_1:83,FUNCT_1:22;
    A17: g2.0= |[0,-1]| by A2,A12,A15,Lm8,BORSUK_1:83,FUNCT_1:22;
    A18: f2.1= |[1,0]| by A2,A12,A14,Lm9,BORSUK_1:83,FUNCT_1:22;
    A19: g2.1= |[0,1]| by A2,A12,A15,Lm9,BORSUK_1:83,FUNCT_1:22;
    A20: f2 is continuous one-to-one & g2 is continuous one-to-one &
     f2.0=|[-1,0]| & f2.1=|[1,0]| & g2.0=|[0,-1]| & g2.1= |[0,1]|
                           by A2,A12,A14,A15,Lm8,Lm9,Th8,Th9,BORSUK_1:83,
FUNCT_1:22;
    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 Lm6;
    then A33: f2.O in KXN by A16,Lm6,Lm7;
    A34: f2.I in KXP by A18,Lm6,Lm7;
       -(|[0,-1]|)`1= 0 by Lm6;
    then A35: g2.O in KYN by A17,Lm6,Lm7;
       -(|[0,1]|)`1= 0 by Lm6;
    then g2.I in KYP by A19,Lm6,Lm7;
    then rng f2 meets rng g2 by A2,A20,A21,A27,A33,A34,A35,Th17;
    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 rng f meets rng g by XBOOLE_0:3;
 end;

theorem 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={p: |.p.|>=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}
    & 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={p: |.p.|>=1}& f.0=p1 & f.1=p3 & 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,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,Lm9,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,Lm9,BORSUK_1:83,FUNCT_1:def 5;
          p4 in rng g by A2,A4,Lm9,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,Th70;
    A13: h is one-to-one by A12,TOPS_2:def 5;
    reconsider f2=h*f,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.0 = |[-1,0]| by A2,A12,A14,Lm8,BORSUK_1:83,FUNCT_1:22;
    A17: g2.0 = |[0,1]| by A2,A12,A15,Lm8,BORSUK_1:83,FUNCT_1:22;
    A18: f2.1 = |[1,0]| by A2,A12,A14,Lm9,BORSUK_1:83,FUNCT_1:22;
    A19: g2.1 = |[0,-1]| by A2,A12,A15,Lm9,BORSUK_1:83,FUNCT_1:22;
    A20: f2 is continuous one-to-one & g2 is continuous one-to-one &
     f2.0=|[-1,0]| & f2.1=|[1,0]| & g2.0=|[0,1]| & g2.1= |[0,-1]|
                           by A2,A12,A14,A15,Lm8,Lm9,Th8,Th9,BORSUK_1:83,
FUNCT_1:22;
    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 Lm6;
    then A33: f2.O in KXN by A16,Lm6,Lm7;
    A34: f2.I in KXP by A18,Lm6,Lm7;
       -(|[0,-1]|)`1= 0 by Lm6;
    then A35: g2.I in KYN by A19,Lm6,Lm7;
       -(|[0,1]|)`1= 0 by Lm6;
    then g2.O in KYP by A17,Lm6,Lm7;
    then rng f2 meets rng g2 by A2,A20,A21,A27,A33,A34,A35,Th18;
    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 rng f meets rng g by XBOOLE_0:3;
 end;

Back to top