The Mizar article:

Fan Homeomorphisms in the Plane

by
Yatsuka Nakamura

Received January 8, 2002

Copyright (c) 2002 Association of Mizar Users

MML identifier: JGRAPH_4
[ MML identifier index ]


environ

 vocabulary FUNCT_1, BOOLE, ABSVALUE, EUCLID, PRE_TOPC, JORDAN2C, SQUARE_1,
      RELAT_1, SUBSET_1, ARYTM_3, METRIC_1, RCOMP_1, FUNCT_4, FUNCT_5, TOPMETR,
      COMPTS_1, JGRAPH_4, ORDINAL2, TOPS_2, ARYTM_1, CARD_3, COMPLEX1, MCART_1,
      PARTFUN1, PCOMPS_1, LATTICES, ARYTM;
 notation TARSKI, XBOOLE_0, SUBSET_1, ORDINAL1, XCMPLX_0, XREAL_0, REAL_1,
      ABSVALUE, RELAT_1, FUNCT_1, FUNCT_2, NAT_1, STRUCT_0, PARTFUN1, TOPMETR,
      PCOMPS_1, EUCLID, COMPTS_1, TOPS_2, METRIC_1, SQUARE_1, TBSP_1, RCOMP_1,
      PSCOMP_1, PRE_TOPC, JGRAPH_1, JORDAN2C, FUNCT_4;
 constructors REAL_1, TOPS_2, TBSP_1, RCOMP_1, PSCOMP_1, JORDAN2C, COMPTS_1,
      TOPGRP_1, TOPMETR, WELLFND1, TOPRNS_1, MEMBERED;
 clusters XREAL_0, STRUCT_0, RELSET_1, EUCLID, PRE_TOPC, SQUARE_1, PSCOMP_1,
      METRIC_1, TOPREAL6, BORSUK_2, JORDAN6, MEMBERED, ZFMISC_1;
 requirements REAL, NUMERALS, BOOLE, SUBSET, ARITHM;
 definitions TARSKI;
 theorems TARSKI, AXIOMS, RELAT_1, SUBSET_1, FUNCT_1, FUNCT_2, FUNCT_4, TOPS_1,
      TOPS_2, PARTFUN1, PRE_TOPC, FRECHET, TOPMETR, JORDAN6, EUCLID, RELSET_1,
      REAL_1, REAL_2, JGRAPH_1, SQUARE_1, PSCOMP_1, METRIC_1, JGRAPH_2,
      GOBOARD6, GOBOARD9, TBSP_1, TOPREAL3, TOPREAL5, TOPREAL6, JGRAPH_3,
      ABSVALUE, JORDAN2C, TSEP_1, EXTENS_1, XREAL_0, JORDAN1, XBOOLE_0,
      XBOOLE_1, TOPRNS_1, SEQ_2, XCMPLX_0, XCMPLX_1;
 schemes FUNCT_2, DOMAIN_1, JGRAPH_2, COMPLSP1, SUPINF_2, JGRAPH_3;

begin :: Preliminaries

 reserve x,a for real number;
 reserve p,q for Point of TOP-REAL 2;
Lm1:p <> 0.REAL 2 implies |.p.| > 0
proof
  assume p <> 0.REAL 2;
then |.p.| <> 0 by TOPRNS_1:25;
  hence thesis by TOPRNS_1:26;
end;

canceled;

theorem Th2: a>=0 & (x-a)*(x+a)<0 implies -a<x & x<a
 proof assume A1:a>=0 & (x-a)*(x+a)<0;
   then A2:x-a>0 & x+a<0 or x-a<0 & x+a>0 by REAL_2:132;
   A3:x+0<=x+a by A1,REAL_1:55;
   then x-a<=x+a-a by REAL_1:49;
   then x-a<=x by XCMPLX_1:26;
   then x-a<=x+a by A3,AXIOMS:22;
   then x<a+0 & x+a>0 by A2,AXIOMS:22,REAL_1:84;
   then x<a & x>0-a by REAL_1:84;
  hence -a<x & x<a by XCMPLX_1:150;
 end;

theorem Th3: for sn being real number st -1<sn & sn<1 holds 1+sn>0 & 1-sn>0
proof let sn be real number;
 assume A1: -1<sn & sn<1;
  then -1+1<sn+1 by REAL_1:67;
 hence 1+sn>0;
 thus 1-sn>0 by A1,SQUARE_1:11;
end;

theorem Th4: for a being real number st a^2<=1 holds -1<=a & a<=1
proof let a be real number;assume a^2<=1;
  then a^2-1^2<=1^2-1^2 by REAL_1:49,SQUARE_1:59;
  then a^2-1^2<=0 by XCMPLX_1:14;
  then (a-1)*(a+1)<=0 by SQUARE_1:67;
 hence -1<=a & a<=1 by JGRAPH_3:4;
end;

theorem Th5: for a being real number st a^2<1 holds -1<a & a<1
proof let a be real number;assume a^2<1;
  then a^2-1^2<1^2-1^2 by REAL_1:54,SQUARE_1:59;
  then a^2-1^2<0 by XCMPLX_1:14;
  then (a-1)*(a+1)<0 by SQUARE_1:67;
 hence -1<a & a<1 by Th2;
end;

theorem Th6:
for X being non empty TopStruct,
    g being map of X,R^1, B being Subset of X,
    a being real number st g is continuous &
B = {p where p is Point of X: pi(g,p) > a } holds B is open
proof let X be non empty TopStruct,g be map of X,R^1,B be Subset of X,
          a be real number;
 assume A1: g is continuous & B={p where p is Point of X:pi(g,p) > a };
    {r where r is Real: r>a} c= the carrier of R^1
   proof let x be set;assume x in {r where r is Real: r>a};
     then consider r being Real such that
     A2: r=x & r>a;
    thus x in the carrier of R^1 by A2,TOPMETR:24;
   end;
  then reconsider D={r where r is Real: r>a} as Subset of R^1;
  A3: D is open by TOPREAL5:7;
  A4: B c= g"D
   proof let x be set;assume
       x in B;
     then consider p being Point of X such that
     A5:  p=x & pi(g,p) > a by A1;
     A6: dom g=the carrier of X by FUNCT_2:def 1;
A7:   pi(g,p) is Real by XREAL_0:def 1;
       pi(g,p)=g.p by JORDAN2C:def 10;
     then x in dom g & g.x in D by A5,A6,A7;
    hence x in g"D by FUNCT_1:def 13;
   end;
    g"D c= B
   proof let x be set;assume
     A8: x in g"D;
     then x in dom g & g.x in D by FUNCT_1:def 13;
     then consider r being Real such that
     A9: r=g.x & r>a;
     reconsider p=x as Point of X by A8;
       pi(g,p)=g.p by JORDAN2C:def 10;
    hence x in B by A1,A9;
   end;
  then B=g"D by A4,XBOOLE_0:def 10;
 hence B is open by A1,A3,TOPS_2:55;
end;

theorem Th7:
 for X being non empty TopStruct,
     g being map of X,R^1,B being Subset of X,
     a being Real st g is continuous &
 B={p where p is Point of X: pi(g,p) < a } holds B is open
proof let X be non empty TopStruct,g be map of X,R^1,
B be Subset of X,a be Real;
 assume A1: g is continuous & B={p where p is Point of X:  pi(g,p) < a };
    {r where r is Real: r<a} c= the carrier of R^1
   proof let x be set;assume x in {r where r is Real: r<a};
     then consider r being Real such that
     A2: r=x & r<a;
    thus x in the carrier of R^1 by A2,TOPMETR:24;
   end;
  then reconsider D={r where r is Real: r<a} as Subset of R^1;
  A3: D is open by TOPREAL5:8;
  A4: B c= g"D
   proof let x be set;assume
       x in B;
     then consider p being Point of X such that
     A5:  p=x & pi(g,p) < a by A1;
     A6: dom g=the carrier of X by FUNCT_2:def 1;
A7:   pi(g,p) is Real by XREAL_0:def 1;
       pi(g,p)=g.p by JORDAN2C:def 10;
     then x in dom g & g.x in D by A5,A6,A7;
    hence x in g"D by FUNCT_1:def 13;
   end;
    g"D c= B
   proof let x be set;assume
     A8: x in g"D;
     then x in dom g & g.x in D by FUNCT_1:def 13;
     then consider r being Real such that
     A9: r=g.x & r<a;
     reconsider p=x as Point of X by A8;
       pi(g,p)=g.p by JORDAN2C:def 10;
    hence x in B by A1,A9;
   end;
  then B=g"D by A4,XBOOLE_0:def 10;
 hence B is open by A1,A3,TOPS_2:55;
end;

theorem Th8: for f being map of TOP-REAL 2,TOP-REAL 2 st
 f is continuous one-to-one & rng f=[#](TOP-REAL 2) &
   (for p2 being Point of TOP-REAL 2
   ex K being non empty compact Subset of TOP-REAL 2 st K = f.:K &
  (ex V2 being Subset of TOP-REAL 2 st
   p2 in V2 & V2 is open & V2 c= K & f.p2 in V2))
holds f is_homeomorphism
proof let f be map of TOP-REAL 2,TOP-REAL 2;
  assume A1: f is continuous one-to-one & rng f=[#](TOP-REAL 2)
  & (for p2 being Point of TOP-REAL 2
  ex K being non empty compact Subset of TOP-REAL 2 st K = f.:K &
 (ex V2 being Subset of TOP-REAL 2 st
  p2 in V2 & V2 is open & V2 c= K & f.p2 in V2));
  A2: dom f=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
  then A3: dom f=[#](TOP-REAL 2) by PRE_TOPC:12;
    rng f=the carrier of TOP-REAL 2 by A1,PRE_TOPC:12;
  then (f qua Function)" is Function of the carrier of TOP-REAL 2,
                  the carrier of TOP-REAL 2 by A1,FUNCT_2:31;
  then reconsider g=(f qua Function)"
          as map of TOP-REAL 2,TOP-REAL 2 ;
    for p being Point of TOP-REAL 2, V being Subset of
  TOP-REAL 2 st g.p in V & V is open ex W being Subset of TOP-REAL 2 st
  p in W & W is open & g.:W c= V
   proof let p be Point of TOP-REAL 2, V be Subset of TOP-REAL 2;
    assume A4: g.p in V & V is open;
  consider K being non empty compact Subset of TOP-REAL 2 such that
  A5: K = f.:K & (ex V2 being Subset of TOP-REAL 2 st
  g.p in V2 & V2 is open & V2 c= K & f.(g.p) in V2) by A1;
  consider V2 being Subset of TOP-REAL 2 such that
  A6: g.p in V2 & V2 is open & V2 c= K & f.(g.p) in V2 by A5;
       p in the carrier of TOP-REAL 2;
     then p in rng f by A1,PRE_TOPC:12;
     then A7: g.p in V2 & V2 is open & V2 c= K & p in V2 by A1,A6,FUNCT_1:57;
     A8: dom (f|K)=dom f /\ K by FUNCT_1:68
      .=(the carrier of TOP-REAL 2)/\ K by FUNCT_2:def 1
      .=K by XBOOLE_1:28
      .=the carrier of ((TOP-REAL 2)|K) by JORDAN1:1;
        rng (f|K) c= rng f by FUNCT_1:76;
      then rng (f|K) c= the carrier of TOP-REAL 2 by XBOOLE_1:1;
     then f|K is Function of the carrier of (TOP-REAL 2)|K,
           the carrier of (TOP-REAL 2) by A8,FUNCT_2:4;
     then reconsider h=f|K as map of (TOP-REAL 2)|K,(TOP-REAL 2)
                          ;
     A9: h is continuous by A1,TOPMETR:10;
     A10: h is one-to-one by A1,FUNCT_1:84;
     A11: dom (f|K)=dom f /\ K by RELAT_1:90 .=K by A2,XBOOLE_1:28;
     A12:K=(f|K).:K by A5,EXTENS_1:1 .=rng (f|K)
                      by A11,RELAT_1:146;
     then consider f1 being map of (TOP-REAL 2)|K,(TOP-REAL 2)|K such that
     A13: h=f1 & f1 is_homeomorphism by A9,A10,JGRAPH_1:64;
     A14: rng f1=[#]((TOP-REAL 2)|K) by A13,TOPS_2:def 5;
     A15: the carrier of ((TOP-REAL 2)|K)=K by JORDAN1:1;
       V /\ V2 /\ K c= K by XBOOLE_1:17;
     then reconsider R=V /\ V2 /\ K as Subset of (TOP-REAL 2)|K
                                by A15;
     A16: V /\ V2 is open by A4,A6,TOPS_1:38;
       R=V /\ V2 /\ [#]((TOP-REAL 2)|K) by PRE_TOPC:def 10;
     then A17: R is open by A16,TOPS_2:32;
     A18: g.p in V /\ V2 by A4,A6,XBOOLE_0:def 3;
     A19: f1.((f1 qua Function)".p)=p by A7,A10,A12,A13,FUNCT_1:57;
     A20: f1.(g.p)=f.(g.p) by A6,A13,FUNCT_1:72.=p
                                  by A1,A2,A3,FUNCT_1:57;
     A21: dom f1=dom f /\ K by A13,FUNCT_1:68 .=K
                               by A2,XBOOLE_1:28;
     A22: rng ((f1 qua Function)")=dom f1 by A10,A13,FUNCT_1:55;
       rng f1=dom ((f1 qua Function)") by A10,A13,FUNCT_1:55;
      then ((f1 qua Function)").p in rng ((f1 qua Function)")
                         by A7,A12,A13,FUNCT_1:12;
     then A23: ((f1 qua Function)").p=g.p by A5,A10,A13,A19,A20,A21,A22,FUNCT_1
:def 8;
     reconsider q=p as Point of ((TOP-REAL 2)|K) by A7,JORDAN1:1;
     A24: ((f1 qua Function)").p in R by A6,A18,A23,XBOOLE_0:def 3;
     A25: dom ((f1 qua Function)")=rng f1 by A10,A13,FUNCT_1:54;
     A26: rng ((f1 qua Function)")=dom f1 by A10,A13,FUNCT_1:55;
     A27: dom f1=the carrier of (TOP-REAL 2)|K by FUNCT_2:def 1;
       the carrier of (TOP-REAL 2)|K=[#]((TOP-REAL 2)|K)
                                by PRE_TOPC:12;
     then A28: the carrier of (TOP-REAL 2)|K=K by PRE_TOPC:def 10;
     A29: rng f1 = the carrier of (TOP-REAL 2)|K by A14,PRE_TOPC:12;
     then (f1 qua Function)" is Function of the carrier of (TOP-REAL 2)|K,
      the carrier of (TOP-REAL 2)|K
                      by A21,A25,A26,A28,FUNCT_2:4;
     then reconsider g1=(f1 qua Function)"
            as map of (TOP-REAL 2)|K,(TOP-REAL 2)|K ;
       g1=f1" by A10,A13,A14,TOPS_2:def 4;
     then g1 is continuous by A13,TOPS_2:def 5;
     then consider W3 being Subset of (TOP-REAL 2)|K such that
     A30: q in W3 & W3 is open & ((f1 qua Function)").:W3 c= R
                        by A17,A24,JGRAPH_2:20;
     A31: ((f qua Function)").:W3 c= R
     proof let y be set;assume y in ((f qua Function)").:W3;
       then consider x being set such that
       A32: x in dom ((f qua Function)") & x in W3 & y=((f qua Function)").x
               by FUNCT_1:def 12;
         x in rng f by A1,A32,FUNCT_1:55;
       then A33: y in dom f & f.y=x by A1,A32,FUNCT_1:54;
       A34: dom ((f1 qua Function)")=the carrier of (TOP-REAL 2)|K by A10,A13,
A29,FUNCT_1:55;
         the carrier of (TOP-REAL 2)|K=K by JORDAN1:1;
       then consider z2 being set such that
       A35: z2 in dom f & z2 in K & f.y=f.z2 by A5,A32,A33,FUNCT_1:def 12;
       A36: y in K by A1,A33,A35,FUNCT_1:def 8;
       then A37: y in the carrier of ((TOP-REAL 2)|K) by JORDAN1:1;
         f1.y=x by A13,A33,A36,FUNCT_1:72;
       then x in dom ((f1 qua Function)") & y=((f1 qua Function)").x
                         by A10,A13,A27,A32,A34,A37,FUNCT_1:54;
       then y in ((f1 qua Function)").:W3 by A32,FUNCT_1:def 12;
      hence y in R by A30;
     end;
     consider W5 being Subset of TOP-REAL 2 such that
     A38: W5 is open & W3=W5 /\ [#]((TOP-REAL 2)|K) by A30,TOPS_2:32;
     reconsider W4=W5 /\ V2 as Subset of TOP-REAL 2;
       p in W5 by A30,A38,XBOOLE_0:def 3;
     then A39: p in W4 by A7,XBOOLE_0:def 3;
     A40: W4 is open by A6,A38,TOPS_1:38;
       W4=W5 /\ (V2 /\ K) by A6,XBOOLE_1:28
       .=W5 /\ K /\ V2 by XBOOLE_1:16
       .=W3 /\ V2 by A38,PRE_TOPC:def 10;
     then A41: g.:W4 c= (g).:W3 /\ (g).:V2 by RELAT_1:154;
       g.:W3 /\ (g).:V2 c= (g).:W3 by XBOOLE_1:17;
     then g.:W4 c= g.:W3 by A41,XBOOLE_1:1;
     then A42: (g).:W4 c= R by A31,XBOOLE_1:1;
       R=V /\ (V2 /\ K) by XBOOLE_1:16;
     then R c= V by XBOOLE_1:17;
     then g.:W4 c= V by A42,XBOOLE_1:1;
    hence ex W being Subset of TOP-REAL 2 st
     p in W & W is open & g.:W c= V by A39,A40;
   end;
  then A43: g is continuous by JGRAPH_2:20;
    g=f" by A1,TOPS_2:def 4;
 hence f is_homeomorphism by A1,A3,A43,TOPS_2:def 5;
end;

theorem Th9: for X being non empty TopSpace,
 f1,f2 being map of X,R^1,a,b being real number st f1 is continuous &
 f2 is continuous & b<>0 & (for q being Point of X holds f2.q<>0)
 holds ex g being map of X,R^1
  st (for p being Point of X,r1,r2 being real number st
  f1.p=r1 & f2.p=r2 holds g.p=(r1/r2-a)/b) & g is continuous
proof let X be non empty TopSpace,
  f1,f2 be map of X,R^1,a,b be real number;
 assume A1:f1 is continuous & f2 is continuous & b<>0 &
   (for q being Point of X holds f2.q<>0);
   consider g1 being map of X,R^1
    such that A2: for p being Point of X
    holds g1.p=b & g1 is continuous by JGRAPH_2:30;
   consider g2 being map of X,R^1
    such that A3: for p being Point of X
    holds g2.p=a & g2 is continuous by JGRAPH_2:30;
   consider g3 being map of X,R^1
    such that A4: (for p being Point of X,r1,r0 being real number st
    f1.p=r1 & f2.p=r0 holds g3.p=r1/r0) & g3 is continuous
                          by A1,JGRAPH_2:37;
   consider g4 being map of X,R^1
    such that A5: (for p being Point of X,r1,r0 being real number st
    g3.p=r1 & g2.p=r0 holds g4.p=r1-r0) & g4 is continuous
                          by A3,A4,JGRAPH_2:31;
     for q being Point of X holds g1.q<>0 by A1,A2;
   then consider g5 being map of X,R^1
    such that A6: (for p being Point of X,r1,r0 being real number st
    g4.p=r1 & g1.p=r0 holds g5.p=r1/r0) & g5 is continuous
                          by A2,A5,JGRAPH_2:37;
     for p being Point of X,r1,r2 being real number st
    f1.p=r1 & f2.p=r2 holds g5.p=(r1/r2-a)/b
    proof let p be Point of X,r1,r2 be real number;
     assume A7:f1.p=r1 & f2.p=r2;
      A8: g1.p=b by A2;
      A9: g2.p=a by A3;
      set r8=r1/r2;
        g3.p=r8 by A4,A7;
      then g4.p=r8-a by A5,A9;
     hence g5.p=(r1/r2-a)/b by A6,A8;
    end;
  hence thesis by A6;
end;

theorem Th10: for X being non empty TopSpace,
 f1,f2 being map of X,R^1,a,b being Real st f1 is continuous &
 f2 is continuous & b<>0 & (for q being Point of X holds f2.q<>0)
 holds ex g being map of X,R^1
  st (for p being Point of X,r1,r2 being Real st
  f1.p=r1 & f2.p=r2 holds g.p=r2*((r1/r2-a)/b)) & g is continuous
proof let X be non empty TopSpace,
  f1,f2 be map of X,R^1,a,b be Real;
 assume A1:f1 is continuous & f2 is continuous & b<>0 &
   (for q being Point of X holds f2.q<>0);
   consider g1 being map of X,R^1
    such that A2: for p being Point of X
    holds g1.p=b & g1 is continuous by JGRAPH_2:30;
   consider g2 being map of X,R^1
    such that A3: for p being Point of X
    holds g2.p=a & g2 is continuous by JGRAPH_2:30;
   consider g3 being map of X,R^1
    such that A4: (for p being Point of X,r1,r0 being real number st
    f1.p=r1 & f2.p=r0 holds g3.p=r1/r0) & g3 is continuous
                          by A1,JGRAPH_2:37;
   consider g4 being map of X,R^1
    such that A5: (for p being Point of X,r1,r0 being real number st
    g3.p=r1 & g2.p=r0 holds g4.p=r1-r0) & g4 is continuous
                          by A3,A4,JGRAPH_2:31;
     for q being Point of X holds g1.q<>0 by A1,A2;
   then consider g5 being map of X,R^1
    such that A6: (for p being Point of X,r1,r0 being real number st
    g4.p=r1 & g1.p=r0 holds g5.p=r1/r0) & g5 is continuous
                          by A2,A5,JGRAPH_2:37;
   consider g6 being map of X,R^1
    such that A7: (for p being Point of X,r1,r0 being real number st
    f2.p=r1 & g5.p=r0 holds g6.p=r1*r0) & g6 is continuous
                          by A1,A6,JGRAPH_2:35;
     for p being Point of X,r1,r2 being Real st
    f1.p=r1 & f2.p=r2 holds g6.p=r2*((r1/r2-a)/b)
    proof let p be Point of X,r1,r2 be Real;
     assume A8:f1.p=r1 & f2.p=r2;
      A9: g1.p=b by A2;
      A10: g2.p=a by A3;
      reconsider r8=r1/r2 as Real;
        g3.p=r8 by A4,A8;
      then g4.p=r8-a by A5,A10;
      then g5.p=(r1/r2-a)/b by A6,A9;
     hence g6.p=r2*((r1/r2-a)/b) by A7,A8;
    end;
  hence thesis by A7;
end;

theorem Th11:
for X being non empty TopSpace,
    f1 being map of X,R^1 st f1 is continuous
holds ex g being map of X,R^1
st (for p being Point of X,r1 being real number st f1.p=r1 holds g.p=r1^2) &
 g is continuous
proof let X be non empty TopSpace,
  f1 being map of X,R^1;
  assume f1 is continuous;
  then consider g1 being map of X,R^1
  such that A1:(for p being Point of X,r1 being real number st
  f1.p=r1 holds g1.p=r1*r1) & g1 is continuous by JGRAPH_2:32;
    for p being Point of X,r1 being real number st
  f1.p=r1 holds g1.p=r1^2
  proof let p be Point of X,r1 be real number;
    assume f1.p=r1;
     then g1.p=r1*r1 by A1;
    hence thesis by SQUARE_1:def 3;
  end;
  hence thesis by A1;
end;

theorem Th12:
for X being non empty TopSpace,
    f1 being map of X,R^1 st f1 is continuous
holds ex g being map of X,R^1
st (for p being Point of X,r1 being real number st
f1.p=r1 holds g.p=abs(r1)) & g is continuous
proof let X be non empty TopSpace,
  f1 be map of X,R^1;
  assume f1 is continuous;
   then consider g1 being map of X,R^1 such that
    A1: (for p being Point of X,r1 being real number st
    f1.p=r1 holds g1.p=r1^2) & g1 is continuous by Th11;
     for q being Point of X ex r being real number st g1.q=r & r>=0
   proof let q be Point of X;
    reconsider r11=f1.q as Real by TOPMETR:24;
    A2: g1.q=r11^2 by A1;
      r11^2>=0 by SQUARE_1:72;
   hence ex r being real number st g1.q=r & r>=0 by A2;
   end;
   then consider g2 being map of X,R^1 such that
    A3: (for p being Point of X,r1 being real number st
    g1.p=r1 holds g2.p=sqrt(r1)) & g2 is continuous by A1,JGRAPH_3:15;
     for p being Point of X,r1 being real number st
    f1.p=r1 holds g2.p=abs(r1)
    proof let p be Point of X,r1 be real number;
     assume f1.p=r1;
      then g1.p=r1^2 by A1;
      then g2.p=sqrt(r1^2) by A3 .=abs(r1) by SQUARE_1:91;
     hence g2.p=abs(r1);
    end;
 hence thesis by A3;
end;

theorem Th13:
for X being non empty TopSpace,
f1 being map of X,R^1 st f1 is continuous
holds ex g being map of X,R^1
st (for p being Point of X,r1 being real number st
f1.p=r1 holds g.p=-r1) & g is continuous
proof let X be non empty TopSpace,
f1 be map of X,R^1;
 assume A1: f1 is continuous;
  consider g1 being map of X,R^1 such that
  A2: (for p being Point of X holds g1.p=0) & g1 is continuous
                       by JGRAPH_2:30;
  consider g2 being map of X,R^1
   such that A3: (for p being Point of X,r1,r2 being real number st
   g1.p=r1 & f1.p=r2 holds g2.p=r1-r2) & g2 is continuous
                   by A1,A2,JGRAPH_2:31;
    for p being Point of X,r1 being real number st
   f1.p=r1 holds g2.p=-r1
   proof let p be Point of X,r1 be real number;
    assume A4: f1.p=r1;
       g1.p=0 by A2;
     then g2.p=0-r1 by A3,A4;
    hence g2.p=-r1 by XCMPLX_1:150;
   end;
 hence thesis by A3;
end;

theorem Th14: for X being non empty TopSpace,
 f1,f2 being map of X,R^1,a,b being real number st f1 is continuous &
 f2 is continuous & b<>0 & (for q being Point of X holds f2.q<>0)
 holds ex g being map of X,R^1
  st (for p being Point of X,r1,r2 being real number st
  f1.p=r1 & f2.p=r2 holds g.p= r2*(-sqrt(abs(1-((r1/r2-a)/b)^2))))
  & g is continuous
 proof let X be non empty TopSpace,
   f1,f2 be map of X,R^1,a,b being real number;
   assume A1: f1 is continuous & f2 is continuous & b<>0 &
   (for q being Point of X holds f2.q<>0);
   then consider g1 being map of X,R^1 such that
    A2: (for p being Point of X,r1,r2 being real number st
    f1.p=r1 & f2.p=r2 holds g1.p=(r1/r2-a)/b) & g1 is continuous
               by Th9;
   consider g2 being map of X,R^1 such that
    A3: (for p being Point of X,s being real number st
    g1.p=s holds g2.p=s^2) & g2 is continuous by A2,Th11;
   consider g0 being map of X,R^1 such that
    A4: (for p being Point of X
     holds g0.p=1)& g0 is continuous by JGRAPH_2:30;
   consider g3 being map of X,R^1 such that
    A5: (for p being Point of X,s,t being real number st
    g0.p=s & g2.p=t holds g3.p=s-t)
    & g3 is continuous by A3,A4,JGRAPH_2:31;
   consider g4 being map of X,R^1 such that
    A6: (for p being Point of X,s being real number st
    g3.p=s holds g4.p=abs s)
    & g4 is continuous by A5,Th12;
     for q being Point of X ex r being real number st g4.q=r & r>=0
   proof let q be Point of X;
     reconsider s=g3.q as Real by TOPMETR:24;
     A7: g4.q=abs s by A6;
       abs s>=0 by ABSVALUE:5;
    hence ex r being real number st g4.q=r & r>=0 by A7;
   end;
   then consider g5 being map of X,R^1 such that
    A8: (for p being Point of X,s being real number st
    g4.p=s holds g5.p=sqrt(s))
    & g5 is continuous by A6,JGRAPH_3:15;
   consider g6 being map of X,R^1 such that
    A9: (for p being Point of X,s being real number st
    g5.p=s holds g6.p=-s) & g6 is continuous by A8,Th13;
   consider g7 being map of X,R^1
    such that A10: (for p being Point of X,r1,r0 being real number st
    f2.p=r1 & g6.p=r0 holds g7.p=r1*r0) & g7 is continuous
                          by A1,A9,JGRAPH_2:35;
     for p being Point of X,r1,r2 being real number st
    f1.p=r1 & f2.p=r2 holds g7.p= r2*(-sqrt(abs(1-((r1/r2-a)/b)^2)))
    proof let p be Point of X,r1,r2 be real number;
     assume A11: f1.p=r1 & f2.p=r2;
      A12: g0.p=1 by A4;
        g1.p=(r1/r2-a)/b by A2,A11;
      then g2.p=((r1/r2-a)/b)^2 by A3;
      then g3.p=1-((r1/r2-a)/b)^2 by A5,A12;
      then g4.p= abs(1-((r1/r2-a)/b)^2) by A6;
      then g5.p=sqrt(abs(1-((r1/r2-a)/b)^2)) by A8;
      then g6.p= -sqrt(abs(1-((r1/r2-a)/b)^2)) by A9;
     hence g7.p=r2*(-sqrt(abs(1-((r1/r2-a)/b)^2))) by A10,A11;
    end;
  hence thesis by A10;
 end;

theorem Th15: for X being non empty TopSpace,
 f1,f2 being map of X,R^1,a,b being real number st
 f1 is continuous &
 f2 is continuous & b<>0 & (for q being Point of X holds f2.q<>0)
 holds ex g being map of X,R^1
  st (for p being Point of X,r1,r2 being real number st
  f1.p=r1 & f2.p=r2 holds g.p= r2*( sqrt(abs(1-((r1/r2-a)/b)^2))))
  & g is continuous
 proof let X be non empty TopSpace,
   f1,f2 be map of X,R^1,a,b being real number;
   assume A1: f1 is continuous &
   f2 is continuous & b<>0 & (for q being Point of X holds f2.q<>0);
   then consider g1 being map of X,R^1 such that
    A2: (for p being Point of X,r1,r2 being real number st
    f1.p=r1 & f2.p=r2 holds g1.p=(r1/r2-a)/b) & g1 is continuous
               by Th9;
   consider g2 being map of X,R^1 such that
    A3: (for p being Point of X,s being real number st
    g1.p=s holds g2.p=s^2) & g2 is continuous by A2,Th11;
   consider g0 being map of X,R^1 such that
    A4: (for p being Point of X
     holds g0.p=1)& g0 is continuous by JGRAPH_2:30;
   consider g3 being map of X,R^1 such that
    A5: (for p being Point of X,s,t being real number st
    g0.p=s & g2.p=t holds g3.p=s-t)
    & g3 is continuous by A3,A4,JGRAPH_2:31;
   consider g4 being map of X,R^1 such that
    A6: (for p being Point of X,s being real number st
    g3.p=s holds g4.p=abs s)
    & g4 is continuous by A5,Th12;
     for q being Point of X ex r being real number st g4.q=r & r>=0
   proof let q be Point of X;
     reconsider s=g3.q as Real by TOPMETR:24;
     A7: g4.q=abs s by A6;
       abs s>=0 by ABSVALUE:5;
    hence ex r being real number st g4.q=r & r>=0 by A7;
   end;
   then consider g5 being map of X,R^1 such that
    A8: (for p being Point of X,s being real number st
    g4.p=s holds g5.p=sqrt(s))
    & g5 is continuous by A6,JGRAPH_3:15;
   consider g7 being map of X,R^1
    such that A9: (for p being Point of X,r1,r0 being real number st
    f2.p=r1 & g5.p=r0 holds g7.p=r1*r0) & g7 is continuous
                          by A1,A8,JGRAPH_2:35;
     for p being Point of X,r1,r2 being real number st
    f1.p=r1 & f2.p=r2 holds g7.p= r2*( sqrt(abs(1-((r1/r2-a)/b)^2)))
    proof let p be Point of X,r1,r2 be real number;
     assume A10: f1.p=r1 & f2.p=r2;
      A11: g0.p=1 by A4;
        g1.p=(r1/r2-a)/b by A2,A10;
      then g2.p=((r1/r2-a)/b)^2 by A3;
      then g3.p=1-((r1/r2-a)/b)^2 by A5,A11;
      then g4.p= abs(1-((r1/r2-a)/b)^2) by A6;
      then g5.p=sqrt(abs(1-((r1/r2-a)/b)^2)) by A8;
     hence g7.p=r2*( sqrt(abs(1-((r1/r2-a)/b)^2))) by A9,A10;
    end;
  hence thesis by A9;
 end;

definition let n be Nat;
 func n NormF -> Function of the carrier of TOP-REAL n, the carrier of R^1
 means
  :Def1: for q being Point of TOP-REAL n holds
     it.q=|.q.|;
existence
 proof
   deffunc F(Point of TOP-REAL n)=|.$1.|;
A1:  for x being Element of TOP-REAL n holds
       F(x) in the carrier of R^1 by TOPMETR:24;
   thus ex IT being Function of the carrier of TOP-REAL n, the carrier of R^1
   st for q being Point of TOP-REAL n holds
     IT.q = F(q) from FunctR_ealEx(A1);
 end;
 uniqueness
  proof
   deffunc F(Point of TOP-REAL n)=|.$1.|;
    thus
    for f,g being Function of the carrier of TOP-REAL n, the carrier of R^1
st (for q being Point of TOP-REAL n holds f.q=F(q)) &
   (for q being Point of TOP-REAL n holds g.q=F(q))
   holds f=g from FuncDefUniq;
  end;
end;

theorem Th16: for n being Nat holds
dom (n NormF)=the carrier of TOP-REAL n & dom (n NormF)=REAL n
proof let n be Nat;
  thus dom (n NormF)=the carrier of TOP-REAL n by FUNCT_2:def 1;
 hence thesis by EUCLID:25;
end;

canceled 2;

theorem Th19:for n being Nat,f being map of TOP-REAL n,R^1 st f=n NormF
holds f is continuous
proof let n be Nat,f be map of TOP-REAL n,R^1;
 assume A1: f=n NormF;
    for p being Point of TOP-REAL n,V being Subset of R^1
  st f.p in V & V is open holds
  ex W being Subset of TOP-REAL n st p in W & W is open & f.:W c= V
  proof let p be Point of TOP-REAL n,V be Subset of R^1;
    assume A2: f.p in V & V is open;
     reconsider v=p as Point of Euclid n by TOPREAL3:13;
     reconsider u=f.p as Point of RealSpace by METRIC_1:def 14,TOPMETR:24;
     reconsider u'=f.p as Real by TOPMETR:24;
     A3: f.p=|.p.| by A1,Def1;
     consider r being real number such that
     A4: r>0 & Ball(u,r) c= V by A2,TOPMETR:22,def 7;
     reconsider r as Real by XREAL_0:def 1;
     A5: Ball(v,r)={q where q is Point of TOP-REAL n:
                          |. p-q .| < r} by JGRAPH_2:10;
     defpred P[Point of TOP-REAL n] means |. p-$1 .| < r;
       {q where q is Point of TOP-REAL n: P[q]}
       is Subset of TOP-REAL n from SubsetD;
     then reconsider W1=Ball(v,r) as Subset of TOP-REAL n by A5;
     A6: v in W1 by A4,GOBOARD6:4;
     A7: W1 is open by GOBOARD6:6;
       f.:W1 c= Ball(u,r)
      proof let y be set;assume y in f.:W1;
        then consider x being set such that
        A8: x in dom f & x in W1 & y=f.x by FUNCT_1:def 12;
        reconsider q=x as Point of TOP-REAL n by A8;
        consider q1 being Point of TOP-REAL n such that
        A9: q=q1 & |. p-q1 .| < r by A5,A8;
        A10: f.x=|.q.| by A1,Def1;
        A11: Ball(u,r)=].u'-r,u'+r.[ by A4,FRECHET:7;
          abs(|.p.|-(|.q.|))<= |. p-q .| by JORDAN2C:11;
        then abs(u'-(|.q.|))<r by A3,A9,AXIOMS:22;
        then A12: -r < u'-(|.q.|) & u'-(|.q.|)<r by SEQ_2:9;
        then A13: u'-(u'-(|.q.|))>u'-r by REAL_1:92;
          u'>|.q.|+-r by A12,REAL_1:86;
        then u' --r > |.q.| +-r--r by REAL_1:54;
        then u' +--r > |.q.| +-r--r by XCMPLX_0:def 8;
        then u'-r<|.q.| & |.q.|<u'+r by A13,XCMPLX_1:18,26;
       hence y in Ball(u,r) by A8,A10,A11,JORDAN6:45;
      end;
       then f.:W1 c= V by A4,XBOOLE_1:1;
    hence thesis by A6,A7;
  end;
 hence f is continuous by JGRAPH_2:20;
end;

theorem Th20: for n being Nat,K0 being Subset of TOP-REAL n,
f being map of (TOP-REAL n)|K0,R^1 st
(for p being Point of (TOP-REAL n)|K0 holds
f.p=(n NormF).p) holds f is continuous
proof let n be Nat,K0 be Subset of TOP-REAL n,
f be map of (TOP-REAL n)|K0,R^1;
assume A1: (for p being Point of (TOP-REAL n)|K0 holds f.p=(n NormF).p);
  A2:dom f= the carrier of (TOP-REAL n)|K0 by FUNCT_2:def 1;
  A3: the carrier of (TOP-REAL n)|K0=K0 by JORDAN1:1;
   (the carrier of TOP-REAL n)/\K0=K0 by XBOOLE_1:28;
  then A4:dom f=dom (n NormF) /\ K0 by A2,A3,Th16;
  A5:for x being set st x in dom f holds f.x=(n NormF).x by A1,A2;
  reconsider g=(n NormF) as map of TOP-REAL n,R^1 ;
  A6:f=g|K0 by A4,A5,FUNCT_1:68;
    g is continuous by Th19;
 hence f is continuous by A6,TOPMETR:10;
end;

theorem Th21: for n being Nat,p being Point of Euclid n,r being Real,
 B being Subset of TOP-REAL n st B=cl_Ball(p,r)
 holds B is Bounded & B is closed
 proof let n be Nat,p be Point of Euclid n,r be Real,
   B be Subset of TOP-REAL n;
   assume A1: B=cl_Ball(p,r);
   A2: TOP-REAL n=TopSpaceMetr(Euclid n) by EUCLID:def 8;
   A3: cl_Ball(p,r) c= Ball(p,r+1)
    proof let x be set;assume A4: x in cl_Ball(p,r);
      then reconsider q=x as Point of Euclid n;
      A5: dist(p,q)<=r by A4,METRIC_1:13;
        r<r+1 by REAL_1:69;
      then dist(p,q)<r+1 by A5,AXIOMS:22;
     hence x in Ball(p,r+1) by METRIC_1:12;
    end;
     Ball(p,r+1) is bounded by TBSP_1:19;
   then cl_Ball(p,r) is bounded by A3,TBSP_1:21;
  hence B is Bounded by A1,JORDAN2C:def 2;
  thus B is closed by A1,A2,TOPREAL6:65;
 end;

theorem Th22: for p being Point of Euclid 2,r being Real,
 B being Subset of TOP-REAL 2 st B=cl_Ball(p,r)
 holds B is compact
 proof let p be Point of Euclid 2,r be Real,
   B be Subset of TOP-REAL 2;
   assume B=cl_Ball(p,r);
   then B is Bounded & B is closed by Th21;
  hence B is compact by TOPREAL6:88;
 end;

begin :: Fan Morphism for West

definition let s be real number, q be Point of TOP-REAL 2;
  func FanW(s,q) -> Point of TOP-REAL 2 equals
  :Def2: |.q.|*|[-sqrt(1-((q`2/|.q.|-s)/(1-s))^2),
          (q`2/|.q.|-s)/(1-s)]| if q`2/|.q.|>=s & q`1<0,
       |.q.|*|[-sqrt(1-((q`2/|.q.|-s)/(1+s))^2),
          (q`2/|.q.|-s)/(1+s)]| if q`2/|.q.|<s & q`1<0
   otherwise q;
 correctness;
end;

definition let s be real number;
 func s-FanMorphW -> Function of the carrier of TOP-REAL 2,
 the carrier of TOP-REAL 2 means
  :Def3: for q being Point of TOP-REAL 2 holds
       it.q=FanW(s,q);
 existence
  proof
    deffunc F(Point of TOP-REAL 2)=FanW(s,$1);
    thus ex IT being Function of the carrier of TOP-REAL 2,
 the carrier of TOP-REAL 2 st
  for q being Point of TOP-REAL 2 holds
       IT.q=F(q) from LambdaD;
  end;
 uniqueness
   proof
    deffunc F(Point of TOP-REAL 2)=FanW(s,$1);
     thus for f,g being Function of the carrier of TOP-REAL 2,
 the carrier of TOP-REAL 2 st
  (for q being Point of TOP-REAL 2 holds f.q=F(q)) &
  (for q being Point of TOP-REAL 2 holds g.q=F(q)) holds f=g
  from FuncDefUniq;
   end;
end;

theorem Th23: for sn being real number holds
  (q`2/|.q.|>=sn & q`1<0 implies
    sn-FanMorphW.q= |[ |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1-sn))]|) &
  (q`1>=0 implies sn-FanMorphW.q=q)
proof let sn be real number;
  A1: sn-FanMorphW.q=FanW(sn,q) by Def3;
   hereby assume q`2/|.q.|>=sn & q`1<0;
    then FanW(sn,q)= |.q.|*|[-sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2),
          (q`2/|.q.|-sn)/(1-sn)]| by Def2
          .= |[ |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1-sn))]| by EUCLID:62;
     hence sn-FanMorphW.q= |[ |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1-sn))]| by Def3;
   end;
   assume q`1>=0; hence thesis by A1,Def2;
end;

theorem Th24: for sn being Real holds
  (q`2/|.q.|<=sn & q`1<0 implies
    sn-FanMorphW.q= |[ |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)),
           |.q.|*((q`2/|.q.|-sn)/(1+sn))]|)
proof let sn be Real;
   assume A1: q`2/|.q.|<=sn & q`1<0;
     now per cases by A1,REAL_1:def 5;
   case q`2/|.q.|<sn;
    then FanW(sn,q)= |.q.|*|[-sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2),
          (q`2/|.q.|-sn)/(1+sn)]| by A1,Def2
          .= |[ |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1+sn))]| by EUCLID:62;
    hence sn-FanMorphW.q= |[ |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)),
           |.q.|*((q`2/|.q.|-sn)/(1+sn))]| by Def3;
   case A2: q`2/|.q.|=sn;
    then A3: q`2/|.q.|-sn=0 by XCMPLX_1:14;
    then (q`2/|.q.|-sn)/(1-sn)=0;
    hence sn-FanMorphW.q= |[ |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)),
           |.q.|*((q`2/|.q.|-sn)/(1+sn))]| by A1,A2,A3,Th23;
   end;
   hence thesis;
end;

theorem Th25: for sn being Real st -1<sn & sn<1 holds
  (q`2/|.q.|>=sn & q`1<=0 & q<>0.REAL 2 implies
    sn-FanMorphW.q= |[ |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1-sn))]|) &
  (q`2/|.q.|<=sn & q`1<=0 & q<>0.REAL 2 implies
    sn-FanMorphW.q= |[ |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)),
           |.q.|*((q`2/|.q.|-sn)/(1+sn))]|)
proof let sn be Real;
 assume A1: -1<sn & sn<1;
     now per cases;
   case A2: q`2/|.q.|>=sn & q`1<=0 & q<>0.REAL 2;
      now per cases;
    case A3: q`1<0;
      then FanW(sn,q)= |.q.|*|[-sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2),
          (q`2/|.q.|-sn)/(1-sn)]| by A2,Def2
          .= |[ |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1-sn))]| by EUCLID:62;
     hence thesis by A3,Def3,Th24;
    case A4: q`1>=0;
     then A5: sn-FanMorphW.q=q by Th23;
     A6: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
     A7: |.q.|<>0 by A2,TOPRNS_1:25;
     then A8: |.q.|^2>0 by SQUARE_1:74;
     A9: q`1=0 by A2,A4;
       |.q.|>=0 by TOPRNS_1:26;
     then A10: sqrt((|.q.|)^2)=|.q.| by SQUARE_1:89;
     A11: 1-sn>0 by A1,Th3;
     A12: |.q.|>0 by A2,Lm1;
       (q`2)^2/|.q.|^2=1^2 by A6,A8,A9,SQUARE_1:59,60,XCMPLX_1:60;
      then ((q`2)/|.q.|)^2=1^2 by SQUARE_1:69;
      then A13: sqrt(((q`2)/|.q.|)^2)=1 by SQUARE_1:89;
       now assume q`2<0;
       then q`2/|.q.|<0 by A12,REAL_2:128;
       then -((q`2)/|.q.|)=1 by A13,SQUARE_1:90;
      hence contradiction by A1,A2;
     end;
     then A14: |.q.|=q`2 by A6,A9,A10,SQUARE_1:60,89;
     then 1=q`2/|.q.| by A7,XCMPLX_1:60;
     then (q`2/|.q.|-sn)/(1-sn)=1 by A11,XCMPLX_1:60;
     hence thesis by A1,A5,A7,A9,A14,EUCLID:57,SQUARE_1:59,82,XCMPLX_1:60;
    end;
    hence thesis;
   case A15: q`2/|.q.|<=sn & q`1<=0 & q<>0.REAL 2;
      now per cases;
    case q`1<0;
     hence thesis by Th23,Th24;
    case A16: q`1>=0;
     A17: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
     A18: |.q.|<>0 by A15,TOPRNS_1:25;
     A19: q`1=0 by A15,A16;
     A20: |.q.|>0 by A15,Lm1;
     A21: 1+sn>0 by A1,Th3;
       1-sn>0 by A1,Th3;
     then 1-sn+sn>0+sn by REAL_1:53;
     then 1>sn by XCMPLX_1:27;
     then 1>q`2/|.q.| by A15,AXIOMS:22;
     then (1)*(|.q.|)>q`2/|.q.|*(|.q.|) by A20,REAL_1:70;
     then A22: (|.q.|)>q`2 by A18,XCMPLX_1:88;
     then A23: |.q.|=-q`2 by A17,A19,JGRAPH_3:1,SQUARE_1:60;
     A24: q`2= -(|.q.|) by A17,A19,A22,JGRAPH_3:1,SQUARE_1:60;
     then -1=q`2/|.q.| by A18,XCMPLX_1:198;
     then (q`2/|.q.|-sn)/(1+sn)
     =(-(1+sn))/(1+sn) by XCMPLX_1:161 .=-1 by A21,XCMPLX_1:198;
     then |[ |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1+sn))]|
      =|[ |.q.|*(-sqrt(1-1)),|.q.|*(-1)]| by SQUARE_1:59,61
     .=|[ |.q.|*(0),-(|.q.|)]| by SQUARE_1:82,XCMPLX_1:180
     .=q by A19,A23,EUCLID:57;
     hence thesis by A1,A16,A18,A24,Th23,XCMPLX_1:198;
    end;
     hence thesis;
   case q`1>0 or q=0.REAL 2;
    hence thesis;
   end;
  hence thesis;
end;

Lm2: for K being non empty Subset of TOP-REAL 2 holds
  proj1|K is continuous map of (TOP-REAL 2)|K,R^1 &
  for q being Point of (TOP-REAL 2)|K holds (proj1|K).q=proj1.q
proof
  let K be non empty Subset of TOP-REAL 2;
  A1: the carrier of (TOP-REAL 2)|K=K by JORDAN1:1;
    proj1|K is Function of K,the carrier of R^1 by FUNCT_2:38,TOPMETR:24;
  then reconsider g2=proj1|K as map of (TOP-REAL 2)|K,R^1 by A1;
    for q being Point of (TOP-REAL 2)|K holds g2.q=proj1.q
  proof let q be Point of (TOP-REAL 2)|K;
    A2:q in the carrier of (TOP-REAL 2)|K;
     dom proj1=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
    then q in dom proj1 /\ K by A1,A2,XBOOLE_0:def 3;
   hence g2.q=proj1.q by FUNCT_1:71;
  end;
  hence thesis by JGRAPH_2:39;
end;

Lm3: for K being non empty Subset of TOP-REAL 2 holds
  proj2|K is continuous map of (TOP-REAL 2)|K,R^1 &
  for q being Point of (TOP-REAL 2)|K holds (proj2|K).q=proj2.q
proof
  let K be non empty Subset of TOP-REAL 2;
  A1: the carrier of (TOP-REAL 2)|K=K by JORDAN1:1;
    proj2|K is Function of K,the carrier of R^1 by FUNCT_2:38,TOPMETR:24;
  then reconsider g2=proj2|K as map of (TOP-REAL 2)|K,R^1 by A1;
    for q being Point of (TOP-REAL 2)|K holds g2.q=proj2.q
  proof let q be Point of (TOP-REAL 2)|K;
    A2:q in the carrier of (TOP-REAL 2)|K;
     dom proj2=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
    then q in dom proj2 /\ K by A1,A2,XBOOLE_0:def 3;
   hence g2.q=proj2.q by FUNCT_1:71;
  end;
  hence thesis by JGRAPH_2:40;
end;

Lm4: dom (2 NormF)=the carrier of TOP-REAL 2 by FUNCT_2:def 1;

Lm5: for K being non empty Subset of TOP-REAL 2 holds
  (2 NormF)|K is continuous map of (TOP-REAL 2)|K,R^1 &
  (for q being Point of (TOP-REAL 2)|K holds ((2 NormF)|K).q=(2 NormF).q)
proof
  let K1 be non empty Subset of TOP-REAL 2;
A1:the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
    (2 NormF)|K1 is Function of K1,the carrier of R^1 by FUNCT_2:38;
  then reconsider g1=(2 NormF)|K1 as map of (TOP-REAL 2)|K1,R^1 by A1;
    for q being Point of (TOP-REAL 2)|K1 holds g1.q=(2 NormF).q
  proof let q be Point of (TOP-REAL 2)|K1;
     q in the carrier of (TOP-REAL 2)|K1;
    then q in dom (2 NormF) /\ K1 by A1,Lm4,XBOOLE_0:def 3;
   hence g1.q=(2 NormF).q by FUNCT_1:71;
  end;
  hence thesis by Th20;
end;

Lm6:
for K1 be non empty Subset of TOP-REAL 2,
    g1 be map of (TOP-REAL 2)|K1,R^1 st g1=(2 NormF)|K1 &
(for q being Point of TOP-REAL 2 st q in
     the carrier of (TOP-REAL 2)|K1 holds q<>0.REAL 2) holds
for q being Point of (TOP-REAL 2)|K1 holds g1.q<>0
  proof
    let K1 be non empty Subset of TOP-REAL 2,
        g1 be map of (TOP-REAL 2)|K1,R^1;
A1: the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
    assume A2: g1=(2 NormF)|K1 & (for q being Point of TOP-REAL 2 st q in
     the carrier of (TOP-REAL 2)|K1 holds q<>0.REAL 2);
    let q be Point of (TOP-REAL 2)|K1;
      q in the carrier of (TOP-REAL 2)|K1;
    then reconsider q2=q as Point of TOP-REAL 2 by A1;
    A3: g1.q=(2 NormF).q by A2,Lm5 .= |.q2.| by Def1;
      now assume |.q2.|=0; then q2=0.REAL 2 by TOPRNS_1:25;
     hence contradiction by A2;
    end;
   hence g1.q<>0 by A3;
  end;

theorem Th26:
 for sn being Real,K1 being non empty Subset of TOP-REAL 2,
f being map of (TOP-REAL 2)|K1,R^1 st -1<sn & sn<1 &
(for p being Point of (TOP-REAL 2) st p in the carrier of (TOP-REAL 2)|K1
  holds f.p=|.p.|* ((p`2/|.p.|-sn)/(1-sn)))
  & (for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds q`1<=0 & q<>0.REAL 2) holds
f is continuous
proof
 let sn be Real,K1 be non empty Subset of TOP-REAL 2,
f be map of (TOP-REAL 2)|K1,R^1;
 assume A1: -1<sn & sn<1 &
(for p being Point of (TOP-REAL 2) st p in the carrier of (TOP-REAL 2)|K1
  holds f.p=|.p.|* ((p`2/|.p.|-sn)/(1-sn)))
  & (for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds q`1<=0 & q<>0.REAL 2);
  set a=sn, b=(1-sn);
  A2: b>0 by A1,Th3;
  reconsider g2=proj2|K1 as continuous map of (TOP-REAL 2)|K1,R^1 by Lm3;
  reconsider g1=(2 NormF)|K1 as continuous map of (TOP-REAL 2)|K1,R^1
    by Lm5;
    for q being Point of TOP-REAL 2 st q in
    the carrier of (TOP-REAL 2)|K1 holds q<>0.REAL 2 by A1;
  then for q being Point of (TOP-REAL 2)|K1 holds g1.q<>0 by Lm6;
  then consider g3 being map of (TOP-REAL 2)|K1,R^1 such that
  A3: (for q being Point of
  (TOP-REAL 2)|K1,r1,r2 being Real st g2.q=r1 & g1.q=r2
  holds g3.q=r2*((r1/r2-a)/b)) & g3 is continuous by A2,Th10;
  A4:dom g3=the carrier of (TOP-REAL 2)|K1 by FUNCT_2:def 1;
  then A5:dom f=dom g3 by FUNCT_2:def 1;
    for x being set st x in dom f holds f.x=g3.x
   proof let x be set;assume
     A6: x in dom f;
     then x in K1 by A4,A5,JORDAN1:1;
     then reconsider r=x as Point of TOP-REAL 2;
     reconsider s=x as Point of (TOP-REAL 2)|K1 by A6,FUNCT_2:def 1;
     A7:f.r=(|.r.|)* ((r`2/|.r.|-sn)/(1-sn)) by A1,A4,A5,A6;
     A8:g2.s=proj2.s by Lm3;
     A9:g1.s=(2 NormF).s by Lm5;
     A10:proj2.r=r`2 by PSCOMP_1:def 29;
       (2 NormF).r=|.r.| by Def1;
    hence f.x=g3.x by A3,A7,A8,A9,A10;
   end;
hence f is continuous by A3,A5,FUNCT_1:9;
end;

theorem Th27:
 for sn being Real,K1 being non empty Subset of TOP-REAL 2,
f being map of (TOP-REAL 2)|K1,R^1 st -1<sn & sn<1 &
(for p being Point of (TOP-REAL 2) st p in the carrier of (TOP-REAL 2)|K1
  holds f.p=|.p.|* ((p`2/|.p.|-sn)/(1+sn)))
  & (for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds q`1<=0 & q<>0.REAL 2) holds
f is continuous
proof
 let sn be Real,K1 be non empty Subset of TOP-REAL 2,
f be map of (TOP-REAL 2)|K1,R^1;
 assume A1:
  -1<sn & sn<1 &
(for p being Point of (TOP-REAL 2) st p in the carrier of (TOP-REAL 2)|K1
  holds f.p=|.p.|* ((p`2/|.p.|-sn)/(1+sn)))
  & (for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds q`1<=0 & q<>0.REAL 2);
  set a=sn, b=(1+sn);
  A2: 1+sn>0 by A1,Th3;
  reconsider g2=proj2|K1 as continuous map of (TOP-REAL 2)|K1,R^1 by Lm3;
  reconsider g1=(2 NormF)|K1 as continuous map of (TOP-REAL 2)|K1,R^1
     by Lm5;
    for q being Point of TOP-REAL 2 st q in
    the carrier of (TOP-REAL 2)|K1 holds q<>0.REAL 2 by A1;
  then for q being Point of (TOP-REAL 2)|K1 holds g1.q<>0 by Lm6;
  then consider g3 being map of (TOP-REAL 2)|K1,R^1 such that
  A3: (for q being Point of
  (TOP-REAL 2)|K1,r1,r2 being Real st g2.q=r1 & g1.q=r2
  holds g3.q=r2*((r1/r2-a)/b)) & g3 is continuous by A2,Th10;
  A4:dom g3=the carrier of (TOP-REAL 2)|K1 by FUNCT_2:def 1;
  then A5:dom f=dom g3 by FUNCT_2:def 1;
    for x being set st x in dom f holds f.x=g3.x
   proof let x be set;assume
     A6: x in dom f; then A7: x in dom g3 by A4,FUNCT_2:def 1;
     then x in K1 by A4,JORDAN1:1;
     then reconsider r=x as Point of (TOP-REAL 2);
     reconsider s=x as Point of (TOP-REAL 2)|K1 by A6,FUNCT_2:def 1;
     A8:f.r=(|.r.|)* ((r`2/|.r.|-sn)/(1+sn)) by A1,A4,A7;
     A9:g2.s=proj2.s by Lm3;
     A10:g1.s=(2 NormF).s by Lm5;
     A11:proj2.r=r`2 by PSCOMP_1:def 29;
       (2 NormF).r=|.r.| by Def1;
    hence f.x=g3.x by A3,A8,A9,A10,A11;
   end;
hence f is continuous by A3,A5,FUNCT_1:9;
end;

theorem Th28:
 for sn being Real,K1 being non empty Subset of TOP-REAL 2,
f being map of (TOP-REAL 2)|K1,R^1 st -1<sn & sn<1 &
(for p being Point of (TOP-REAL 2) st p in the carrier of (TOP-REAL 2)|K1
  holds f.p=|.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)))
  & (for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds
        q`1<=0 & q`2/|.q.|>=sn & q<>0.REAL 2) holds
f is continuous
proof
 let sn be Real,K1 be non empty Subset of TOP-REAL 2,
f be map of (TOP-REAL 2)|K1,R^1;
 assume A1: -1<sn & sn<1 &
(for p being Point of (TOP-REAL 2) st p in the carrier of (TOP-REAL 2)|K1
  holds f.p=|.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)))
  & (for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds
            q`1<=0 & q`2/|.q.|>=sn & q<>0.REAL 2);
  set a=sn, b=(1-sn);
  A2: b>0 by A1,Th3;
  reconsider g2=proj2|K1 as continuous map of (TOP-REAL 2)|K1,R^1 by Lm3;
  reconsider g1=(2 NormF)|K1 as continuous map of (TOP-REAL 2)|K1,R^1
                          by Lm5;
    for q being Point of TOP-REAL 2 st q in
    the carrier of (TOP-REAL 2)|K1 holds q<>0.REAL 2 by A1;
  then for q being Point of (TOP-REAL 2)|K1 holds g1.q<>0 by Lm6;
  then consider g3 being map of (TOP-REAL 2)|K1,R^1 such that
  A3: (for q being Point of
  (TOP-REAL 2)|K1,r1,r2 being real number st g2.q=r1 & g1.q=r2
  holds g3.q=r2*(-sqrt(abs(1-((r1/r2-a)/b)^2))))
          & g3 is continuous by A2,Th14;
  A4:dom g3=the carrier of (TOP-REAL 2)|K1 by FUNCT_2:def 1;
  then A5:dom f=dom g3 by FUNCT_2:def 1;
    for x being set st x in dom f holds f.x=g3.x
   proof let x be set;assume
     A6: x in dom f;
     then x in K1 by A4,A5,JORDAN1:1;
     then reconsider r=x as Point of (TOP-REAL 2);
     reconsider s=x as Point of (TOP-REAL 2)|K1 by A6,FUNCT_2:def 1;
     A7: now assume |.r.|=0; then r=0.REAL 2 by TOPRNS_1:25;
      hence contradiction by A1,A4,A5,A6;
     end;
     A8: |.r.|>=0 by TOPRNS_1:26;
       |.r.|^2=(r`1)^2+(r`2)^2 by JGRAPH_3:10;
     then |.r.|^2 -(r`2)^2 =(r`1)^2 by XCMPLX_1:26;
     then (r`2)^2 -(|.r.|)^2 =-(r`1)^2 by XCMPLX_1:143;
     then A9: ((r`2) -(|.r.|))*((r`2)+|.r.|) =-(r`1)^2 by SQUARE_1:67;
       (r`1)^2>=0 by SQUARE_1:72;
     then -(r`1)^2 <=0 by REAL_1:66;
     then -(|.r.|)<=r`2 & r`2<= |.r.| by A8,A9,JGRAPH_3:4;
     then r`2/|.r.| <= |.r.|/|.r.| by A7,A8,REAL_1:73;
     then r`2/|.r.|<=1 by A7,XCMPLX_1:60;
     then A10: r`2/|.r.|-sn<=(1-sn) by REAL_1:49;
     A11: now assume (1-sn)^2=0;
       then 1-sn+sn=0+sn by SQUARE_1:73;
      hence contradiction by A1,XCMPLX_1:27;
     end;
     A12: 1-sn>0 by A1,Th3;
      A13: (1-sn)^2>=0 by SQUARE_1:72;
       sn<=r`2/(|.r.|) by A1,A4,A5,A6;
     then sn-r`2/|.r.|<=0 by REAL_2:106;
     then -(sn- r`2/|.r.|)>=-(1-sn) by A12,REAL_1:50;
     then -(1-sn)<= r`2/|.r.|-sn by XCMPLX_1:143;
     then (r`2/|.r.|-sn)^2<=(1-sn)^2 by A10,JGRAPH_2:7;
     then (r`2/|.r.|-sn)^2/(1-sn)^2<=(1-sn)^2/(1-sn)^2 by A11,A13,REAL_1:73;
     then (r`2/|.r.|-sn)^2/(1-sn)^2<=1 by A11,XCMPLX_1:60;
     then ((r`2/|.r.|-sn)/(1-sn))^2<=1 by SQUARE_1:69;
     then 1-((r`2/|.r.|-sn)/(1-sn))^2>=0 by SQUARE_1:12;
     then abs(1-((r`2/|.r.|-sn)/(1-sn))^2)
          =1-((r`2/|.r.|-sn)/(1-sn))^2 by ABSVALUE:def 1;
     then A14:f.r=(|.r.|)*(-sqrt(abs(1-((r`2/|.r.|-sn)/(1-sn))^2))) by A1,A4,A5
,A6;
     A15:g2.s=proj2.s by Lm3;
     A16:g1.s=(2 NormF).s by Lm5;
     A17:proj2.r=r`2 by PSCOMP_1:def 29;
       (2 NormF).r=|.r.| by Def1;
    hence f.x=g3.x by A3,A14,A15,A16,A17;
   end;
hence f is continuous by A3,A5,FUNCT_1:9;
end;

theorem Th29:
 for sn being Real,K1 being non empty Subset of TOP-REAL 2,
f being map of (TOP-REAL 2)|K1,R^1 st -1<sn & sn<1 &
(for p being Point of (TOP-REAL 2) st p in the carrier of (TOP-REAL 2)|K1
  holds f.p=|.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)))
  & (for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds
        q`1<=0 & q`2/|.q.|<=sn & q<>0.REAL 2) holds
f is continuous
proof
 let sn be Real,K1 be non empty Subset of TOP-REAL 2,
f be map of (TOP-REAL 2)|K1,R^1;
 assume A1: -1<sn & sn<1 &
(for p being Point of (TOP-REAL 2) st p in the carrier of (TOP-REAL 2)|K1
  holds f.p=|.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)))
  & (for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds
            q`1<=0 & q`2/|.q.|<=sn & q<>0.REAL 2);
  set a=sn, b=(1+sn);
  A2: 1+sn>0 by A1,Th3;
  reconsider g2=proj2|K1 as continuous map of (TOP-REAL 2)|K1,R^1 by Lm3;
  reconsider g1=(2 NormF)|K1 as continuous map of (TOP-REAL 2)|K1,R^1
                          by Lm5;
    for q being Point of TOP-REAL 2 st q in
    the carrier of (TOP-REAL 2)|K1 holds q<>0.REAL 2 by A1;
  then for q being Point of (TOP-REAL 2)|K1 holds g1.q<>0 by Lm6;
  then consider g3 being map of (TOP-REAL 2)|K1,R^1 such that
  A3: (for q being Point of
  (TOP-REAL 2)|K1,r1,r2 being real number st g2.q=r1 & g1.q=r2
  holds g3.q=r2*(-sqrt(abs(1-((r1/r2-a)/b)^2))))
          & g3 is continuous by A2,Th14;
  A4:dom g3=the carrier of (TOP-REAL 2)|K1 by FUNCT_2:def 1;
  then A5:dom f=dom g3 by FUNCT_2:def 1;
    for x being set st x in dom f holds f.x=g3.x
   proof let x be set;assume
     A6:x in dom f;
     then x in K1 by A4,A5,JORDAN1:1;
     then reconsider r=x as Point of (TOP-REAL 2);
     reconsider s=x as Point of (TOP-REAL 2)|K1 by A6,FUNCT_2:def 1;
     A7: now assume |.r.|=0; then r=0.REAL 2 by TOPRNS_1:25;
      hence contradiction by A1,A4,A5,A6;
     end;
     A8: |.r.|>=0 by TOPRNS_1:26;
       |.r.|^2=(r`1)^2+(r`2)^2 by JGRAPH_3:10;
     then |.r.|^2 -(r`2)^2 =(r`1)^2 by XCMPLX_1:26;
     then (r`2)^2 -(|.r.|)^2 =-(r`1)^2 by XCMPLX_1:143;
     then A9: ((r`2) -(|.r.|))*((r`2)+|.r.|) =-(r`1)^2 by SQUARE_1:67;
       (r`1)^2>=0 by SQUARE_1:72;
     then -(r`1)^2 <=0 by REAL_1:66;
     then -(|.r.|)<=r`2 & r`2<= |.r.| by A8,A9,JGRAPH_3:4;
     then r`2/|.r.| >= (-(|.r.|))/|.r.| by A7,A8,REAL_1:73;
     then r`2/|.r.|>= -1 by A7,XCMPLX_1:198;
     then r`2/|.r.|-sn>=-1-sn by REAL_1:49;
     then A10: r`2/|.r.|-sn>=-(1+sn) by XCMPLX_1:161;
     A11: (1+sn)^2>0 by A2,SQUARE_1:74;
       sn>=r`2/(|.r.|) by A1,A4,A5,A6;
     then sn-r`2/|.r.|>=0 by SQUARE_1:12;
     then -(sn-r`2/|.r.|)<=-0 by REAL_1:50;
     then r`2/|.r.|-sn<=-0 by XCMPLX_1:143;
     then r`2/|.r.|-sn<=1+sn by A2,AXIOMS:22;
     then (r`2/|.r.|-sn)^2<=(1+sn)^2 by A10,JGRAPH_2:7;
     then (r`2/|.r.|-sn)^2/(1+sn)^2<=(1+sn)^2/(1+sn)^2 by A11,REAL_1:73;
     then (r`2/|.r.|-sn)^2/(1+sn)^2<=1 by A11,XCMPLX_1:60;
     then ((r`2/|.r.|-sn)/(1+sn))^2<=1 by SQUARE_1:69;
     then 1-((r`2/|.r.|-sn)/(1+sn))^2>=0 by SQUARE_1:12;
     then abs(1-((r`2/|.r.|-sn)/(1+sn))^2)
          =1-((r`2/|.r.|-sn)/(1+sn))^2 by ABSVALUE:def 1;
     then A12:f.r=(|.r.|)*(-sqrt(abs(1-((r`2/|.r.|-sn)/(1+sn))^2))) by A1,A4,A5
,A6;
     A13:g2.s=proj2.s by Lm3;
     A14:g1.s=(2 NormF).s by Lm5;
     A15:proj2.r=r`2 by PSCOMP_1:def 29;
       (2 NormF).r=|.r.| by Def1;
    hence f.x=g3.x by A3,A12,A13,A14,A15;
   end;
hence f is continuous by A3,A5,FUNCT_1:9;
end;

theorem Th30: for sn being Real,
K0,B0 being Subset of TOP-REAL 2,f being
map of (TOP-REAL 2)|K0,(TOP-REAL 2)|B0
st -1<sn & sn<1 & f=(sn-FanMorphW)|K0 &
B0={q where q is Point of TOP-REAL 2: q`1<=0 & q<>0.REAL 2}
& K0={p: p`2/|.p.|>=sn & p`1<=0 & p<>0.REAL 2}
holds f is continuous
proof let sn be Real,K0,B0 be Subset of TOP-REAL 2,f be
map of (TOP-REAL 2)|K0,(TOP-REAL 2)|B0;
assume A1: -1<sn & sn<1 & f=(sn-FanMorphW)|K0 &
B0={q where q is Point of TOP-REAL 2: q`1<=0 & q<>0.REAL 2} &
K0={p: p`2/|.p.|>=sn & p`1<=0 & p<>0.REAL 2};
 set cn=sqrt(1-sn^2);
 set p0=|[-cn,sn]|;
 A2: p0`2=sn by EUCLID:56;
 A3: p0`1=-cn by EUCLID:56;
     sn^2<1^2 by A1,JGRAPH_2:8;
 then A4: 1-sn^2>0 by SQUARE_1:11,59;
 A5: now assume p0=0.REAL 2;
   then --cn=-0 by EUCLID:56,JGRAPH_2:11;
  hence contradiction by A4,SQUARE_1:93;
 end;
   --cn>0 by A4,SQUARE_1:93;
 then A6: p0`1<0 by A3,REAL_1:66;
 A7: |.p0.|=sqrt((-cn)^2+sn^2) by A2,A3,JGRAPH_3:10
       .=sqrt((cn)^2+sn^2) by SQUARE_1:61;
   cn^2=1-sn^2 by A4,SQUARE_1:def 4;
 then |.p0.|=1 by A7,SQUARE_1:83,XCMPLX_1:27;
 then p0`2/|.p0.|=sn by EUCLID:56;
 then A8: p0 in K0 by A1,A5,A6;
then reconsider K1=K0 as non empty Subset of TOP-REAL 2;
  A9: K0 c= B0
   proof let x be set;assume x in K0;
      then consider p8 being Point of TOP-REAL 2 such that
      A10: x=p8 & (p8`2/|.p8.|>=sn & p8`1<=0 & p8<>0.REAL 2) by A1;
    thus x in B0 by A1,A10;
   end;
A11:dom ((proj2)*((sn-FanMorphW)|K1)) c= dom ((sn-FanMorphW)|K1)
                                  by RELAT_1:44;
 dom ((sn-FanMorphW)|K1) c= dom ((proj2)*((sn-FanMorphW)|K1))
proof let x be set;assume A12:x in dom ((sn-FanMorphW)|K1);
   then A13:x in dom (sn-FanMorphW) /\ K1 by FUNCT_1:68;
   A14:((sn-FanMorphW)|K1).x=(sn-FanMorphW).x by A12,FUNCT_1:68;
   A15: dom proj2 = (the carrier of TOP-REAL 2) by FUNCT_2:def 1;
     x in dom (sn-FanMorphW) by A13,XBOOLE_0:def 3;
   then (sn-FanMorphW).x in rng (sn-FanMorphW) by FUNCT_1:12;
  hence x in dom ((proj2)*((sn-FanMorphW)|K1)) by A12,A14,A15,FUNCT_1:21;
end;
then A16:dom ((proj2)*((sn-FanMorphW)|K1))
=dom ((sn-FanMorphW)|K1) by A11,XBOOLE_0:def 10
.=dom (sn-FanMorphW) /\ K1 by FUNCT_1:68
.=(the carrier of TOP-REAL 2)/\ K1 by FUNCT_2:def 1
.=K1 by XBOOLE_1:28
.=the carrier of (TOP-REAL 2)|K1 by JORDAN1:1;
A17:rng ((proj2)*((sn-FanMorphW)|K1)) c= rng (proj2) by RELAT_1:45;
  rng proj2 c= the carrier of R^1 by JGRAPH_2:17,RELSET_1:12;
then rng ((proj2)*((sn-FanMorphW)|K1)) c= the carrier of R^1 by A17,XBOOLE_1:1
;
then (proj2)*((sn-FanMorphW)|K1) is Function of the carrier of (TOP-REAL 2)|K1
,
the carrier of R^1 by A16,FUNCT_2:4;
then reconsider g2=(proj2)*((sn-FanMorphW)|K1) as map of (TOP-REAL 2)|K1,R^1
                                ;
A18:dom ((proj1)*((sn-FanMorphW)|K1))
                     c= dom ((sn-FanMorphW)|K1) by RELAT_1:44;
 dom ((sn-FanMorphW)|K1) c= dom ((proj1)*((sn-FanMorphW)|K1))
proof let x be set;assume A19:x in dom ((sn-FanMorphW)|K1);
   then A20:x in dom (sn-FanMorphW) /\ K1 by FUNCT_1:68;
   A21:((sn-FanMorphW)|K1).x=(sn-FanMorphW).x by A19,FUNCT_1:68;
   A22: dom proj1 = (the carrier of TOP-REAL 2) by FUNCT_2:def 1;
     x in dom (sn-FanMorphW) by A20,XBOOLE_0:def 3;
   then (sn-FanMorphW).x in rng (sn-FanMorphW) by FUNCT_1:12;
  hence x in dom ((proj1)*((sn-FanMorphW)|K1)) by A19,A21,A22,FUNCT_1:21;
end;
then A23:dom ((proj1)*((sn-FanMorphW)|K1))
=dom ((sn-FanMorphW)|K1) by A18,XBOOLE_0:def 10
.=dom (sn-FanMorphW) /\ K1 by FUNCT_1:68
.=((the carrier of TOP-REAL 2))/\ K1 by FUNCT_2:def 1
.=K1 by XBOOLE_1:28
.=the carrier of (TOP-REAL 2)|K1 by JORDAN1:1;
A24:rng ((proj1)*((sn-FanMorphW)|K1)) c= rng (proj1) by RELAT_1:45;
  rng proj1 c= the carrier of R^1 by JGRAPH_2:16,RELSET_1:12;
then rng ((proj1)*((sn-FanMorphW)|K1)) c= the carrier of R^1 by A24,XBOOLE_1:1
;
then (proj1)*((sn-FanMorphW)|K1) is Function of the carrier of (TOP-REAL 2)|K1
,
the carrier of R^1 by A23,FUNCT_2:4;
then reconsider g1=(proj1)*((sn-FanMorphW)|K1) as map of (TOP-REAL 2)|K1,R^1
                                ;
A25: for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds q`1<=0 & q<>0.REAL 2
   proof let q be Point of TOP-REAL 2;
    assume A26:q in the carrier of (TOP-REAL 2)|K1;
        the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
     then consider p3 being Point of TOP-REAL 2 such that
     A27: q=p3 & (p3`2/|.p3.|>=sn & p3`1<=0 & p3<>0.REAL 2) by A1,A26;
    thus q`1<=0 & q<>0.REAL 2 by A27;
   end;
    for p being Point of TOP-REAL 2 st p in the carrier of (TOP-REAL 2)|K1
  holds g2.p=|.p.|* ((p`2/|.p.|-sn)/(1-sn))
   proof let p be Point of TOP-REAL 2;
    assume A28: p in the carrier of (TOP-REAL 2)|K1;
     A29:dom ((sn-FanMorphW)|K1)=dom (sn-FanMorphW) /\ K1 by FUNCT_1:68
     .=((the carrier of TOP-REAL 2))/\ K1 by FUNCT_2:def 1
     .=K1 by XBOOLE_1:28;
     A30:the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
     then consider p3 being Point of TOP-REAL 2 such that
     A31: p=p3 & (p3`2/|.p3.|>=sn & p3`1<=0 & p3<>0.REAL 2) by A1,A28;
     A32:(sn-FanMorphW).p
     =|[|.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)),
      |.p.|* ((p`2/|.p.|-sn)/(1-sn))]| by A1,A31,Th25;
      ((sn-FanMorphW)|K1).p=(sn-FanMorphW).p by A28,A30,FUNCT_1:72;
     then g2.p=(proj2).(|[|.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)),
            |.p.|* ((p`2/|.p.|-sn)/(1-sn))]|)
                             by A28,A29,A30,A32,FUNCT_1:23
      .=(|[|.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)),
        |.p.|* ((p`2/|.p.|-sn)/(1-sn))]|)`2 by PSCOMP_1:def 29
      .=|.p.|* ((p`2/|.p.|-sn)/(1-sn)) by EUCLID:56;
    hence g2.p=|.p.|* ((p`2/|.p.|-sn)/(1-sn));
   end;
then consider f2 being map of (TOP-REAL 2)|K1,R^1 such that
  A33:for p being Point of TOP-REAL 2 st p in the carrier of (TOP-REAL 2)|K1
  holds f2.p=|.p.|* ((p`2/|.p.|-sn)/(1-sn));
A34:f2 is continuous by A1,A25,A33,Th26;
    for p being Point of TOP-REAL 2 st p in the carrier of (TOP-REAL 2)|K1
  holds g1.p=|.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2))
   proof let p be Point of TOP-REAL 2;
    assume A35: p in the carrier of (TOP-REAL 2)|K1;
     A36:dom ((sn-FanMorphW)|K1)=dom (sn-FanMorphW) /\ K1 by FUNCT_1:68
     .=((the carrier of TOP-REAL 2))/\ K1 by FUNCT_2:def 1
     .=K1 by XBOOLE_1:28;
     A37:the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
     then consider p3 being Point of TOP-REAL 2 such that
     A38: p=p3 & (p3`2/|.p3.|>=sn & p3`1<=0 & p3<>0.REAL 2) by A1,A35;
     A39:(sn-FanMorphW).p=|[|.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)),
      |.p.|* ((p`2/|.p.|-sn)/(1-sn))]| by A1,A38,Th25;
      ((sn-FanMorphW)|K1).p=(sn-FanMorphW).p by A35,A37,FUNCT_1:72;
     then g1.p=(proj1).(|[|.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)),
          |.p.|* ((p`2/|.p.|-sn)/(1-sn))]|)
                             by A35,A36,A37,A39,FUNCT_1:23
        .=(|[|.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)),
          |.p.|* ((p`2/|.p.|-sn)/(1-sn)) ]|)`1 by PSCOMP_1:def 28
        .= |.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)) by EUCLID:56;
    hence g1.p=|.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2));
   end;
then consider f1 being map of (TOP-REAL 2)|K1,R^1 such that
  A40:for p being Point of TOP-REAL 2 st p in the carrier of (TOP-REAL 2)|K1
  holds f1.p=|.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2));
    for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds q`1<=0 & q`2/|.q.|>=sn & q<>0.REAL 2
   proof let q be Point of TOP-REAL 2;
    assume A41:q in the carrier of (TOP-REAL 2)|K1;
        the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
     then consider p3 being Point of TOP-REAL 2 such that
     A42: q=p3 & (p3`2/|.p3.|>=sn & p3`1<=0 & p3<>0.REAL 2) by A1,A41;
    thus thesis by A42;
   end;
then A43:f1 is continuous by A1,A40,Th28;
  for x,y,r,s being real number st |[x,y]| in K1 &
  r=f1.(|[x,y]|) & s=f2.(|[x,y]|) holds
  f.(|[x,y]|)=|[r,s]|
  proof let x,y,r,s be real number;
   assume A44: |[x,y]| in K1 & r=f1.(|[x,y]|) & s=f2.(|[x,y]|);
     set p99=|[x,y]|;
     A45:the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
     consider p3 being Point of TOP-REAL 2 such that
     A46: p99=p3 & (p3`2/|.p3.|>=sn & p3`1<=0 & p3<>0.REAL 2) by A1,A44;
     A47:f1.p99=|.p99.|*(-sqrt(1-((p99`2/|.p99.|-sn)/(1-sn))^2))
                          by A40,A44,A45;
      ((sn-FanMorphW)|K0).(|[x,y]|)=((sn-FanMorphW)).(|[x,y]|) by A44,FUNCT_1:
72
    .= |[|.p99.|*(-sqrt(1-((p99`2/|.p99.|-sn)/(1-sn))^2)),
      |.p99.|* ((p99`2/|.p99.|-sn)/(1-sn))]| by A1,A46,Th25
    .=|[r,s]| by A33,A44,A45,A47;
   hence f.(|[x,y]|)=|[r,s]| by A1;
  end;
hence f is continuous by A8,A9,A34,A43,JGRAPH_2:45;
end;

theorem Th31: for sn being Real,
K0,B0 being Subset of TOP-REAL 2,f being
map of (TOP-REAL 2)|K0,(TOP-REAL 2)|B0
st -1<sn & sn<1 & f=(sn-FanMorphW)|K0 &
B0={q where q is Point of TOP-REAL 2: q`1<=0 & q<>0.REAL 2}
& K0={p: p`2/|.p.|<=sn & p`1<=0 & p<>0.REAL 2}
holds f is continuous
proof let sn be Real,K0,B0 be Subset of TOP-REAL 2,f be
map of (TOP-REAL 2)|K0,(TOP-REAL 2)|B0;
assume A1: -1<sn & sn<1 & f=(sn-FanMorphW)|K0 &
B0={q where q is Point of TOP-REAL 2: q`1<=0 & q<>0.REAL 2} &
K0={p: p`2/|.p.|<=sn & p`1<=0 & p<>0.REAL 2};
 set cn=sqrt(1-sn^2);
 set p0=|[-cn,sn]|;
 A2: p0`2=sn by EUCLID:56;
 A3: p0`1=-cn by EUCLID:56;
     sn^2<1^2 by A1,JGRAPH_2:8;
 then A4: 1-sn^2>0 by SQUARE_1:11,59;
 A5: now assume p0=0.REAL 2;
   then --cn=-0 by EUCLID:56,JGRAPH_2:11;
  hence contradiction by A4,SQUARE_1:93;
 end;
   --cn>0 by A4,SQUARE_1:93;
 then A6: p0`1<0 by A3,REAL_1:66;
 A7: |.p0.|=sqrt((-cn)^2+sn^2) by A2,A3,JGRAPH_3:10
       .=sqrt((cn)^2+sn^2) by SQUARE_1:61;
   cn^2=1-sn^2 by A4,SQUARE_1:def 4;
 then |.p0.|=1 by A7,SQUARE_1:83,XCMPLX_1:27;
 then p0`2/|.p0.|=sn by EUCLID:56;
 then A8: p0 in K0 by A1,A5,A6;
 then reconsider K1=K0 as non empty Subset of TOP-REAL 2;
  A9: K0 c= B0
   proof let x be set;assume x in K0;
      then consider p8 being Point of TOP-REAL 2 such that
      A10: x=p8 & (
      p8`2/|.p8.|<=sn & p8`1<=0 & p8<>0.REAL 2) by A1;
    thus x in B0 by A1,A10;
   end;
A11:dom ((proj2)*((sn-FanMorphW)|K1)) c= dom ((sn-FanMorphW)|K1)
                                  by RELAT_1:44;
 dom ((sn-FanMorphW)|K1) c= dom ((proj2)*((sn-FanMorphW)|K1))
proof let x be set;assume A12:x in dom ((sn-FanMorphW)|K1);
   then A13:x in dom (sn-FanMorphW) /\ K1 by FUNCT_1:68;
   A14:((sn-FanMorphW)|K1).x=(sn-FanMorphW).x by A12,FUNCT_1:68;
   A15: dom proj2 = (the carrier of TOP-REAL 2) by FUNCT_2:def 1;
     x in dom (sn-FanMorphW) by A13,XBOOLE_0:def 3;
   then (sn-FanMorphW).x in rng (sn-FanMorphW) by FUNCT_1:12;
  hence x in dom ((proj2)*((sn-FanMorphW)|K1)) by A12,A14,A15,FUNCT_1:21;
end;
then A16:dom ((proj2)*((sn-FanMorphW)|K1))
=dom ((sn-FanMorphW)|K1) by A11,XBOOLE_0:def 10
.=dom (sn-FanMorphW) /\ K1 by FUNCT_1:68
.=(the carrier of TOP-REAL 2)/\ K1 by FUNCT_2:def 1
.=K1 by XBOOLE_1:28
.=the carrier of (TOP-REAL 2)|K1 by JORDAN1:1;
A17:rng ((proj2)*((sn-FanMorphW)|K1)) c= rng (proj2) by RELAT_1:45;
  rng proj2 c= the carrier of R^1 by JGRAPH_2:17,RELSET_1:12;
then rng ((proj2)*((sn-FanMorphW)|K1)) c= the carrier of R^1 by A17,XBOOLE_1:1
;
then (proj2)*((sn-FanMorphW)|K1) is Function of the carrier of (TOP-REAL 2)|K1
,
the carrier of R^1 by A16,FUNCT_2:4;
then reconsider g2=(proj2)*((sn-FanMorphW)|K1) as map of (TOP-REAL 2)|K1,R^1
                                ;
A18:dom ((proj1)*((sn-FanMorphW)|K1))
                     c= dom ((sn-FanMorphW)|K1) by RELAT_1:44;
 dom ((sn-FanMorphW)|K1) c= dom ((proj1)*((sn-FanMorphW)|K1))
proof let x be set;assume A19:x in dom ((sn-FanMorphW)|K1);
   then A20:x in dom (sn-FanMorphW) /\ K1 by FUNCT_1:68;
   A21:((sn-FanMorphW)|K1).x=(sn-FanMorphW).x by A19,FUNCT_1:68;
   A22: dom proj1 = (the carrier of TOP-REAL 2) by FUNCT_2:def 1;
     x in dom (sn-FanMorphW) by A20,XBOOLE_0:def 3;
   then (sn-FanMorphW).x in rng (sn-FanMorphW) by FUNCT_1:12;
  hence x in dom ((proj1)*((sn-FanMorphW)|K1)) by A19,A21,A22,FUNCT_1:21;
end;
then A23:dom ((proj1)*((sn-FanMorphW)|K1))
=dom ((sn-FanMorphW)|K1) by A18,XBOOLE_0:def 10
.=dom (sn-FanMorphW) /\ K1 by FUNCT_1:68
.=((the carrier of TOP-REAL 2))/\ K1 by FUNCT_2:def 1
.=K1 by XBOOLE_1:28
.=the carrier of (TOP-REAL 2)|K1 by JORDAN1:1;
A24:rng ((proj1)*((sn-FanMorphW)|K1)) c= rng (proj1) by RELAT_1:45;
  rng proj1 c= the carrier of R^1 by JGRAPH_2:16,RELSET_1:12;
then rng ((proj1)*((sn-FanMorphW)|K1)) c= the carrier of R^1 by A24,XBOOLE_1:1
;
then (proj1)*((sn-FanMorphW)|K1) is Function of the carrier of (TOP-REAL 2)|K1
,
the carrier of R^1 by A23,FUNCT_2:4;
then reconsider g1=(proj1)*((sn-FanMorphW)|K1) as map of (TOP-REAL 2)|K1,R^1
                                ;
A25: for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds q`1<=0 & q<>0.REAL 2
   proof let q be Point of TOP-REAL 2;
    assume A26:q in the carrier of (TOP-REAL 2)|K1;
        the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
     then consider p3 being Point of TOP-REAL 2 such that
     A27: q=p3 & (p3`2/|.p3.|<=sn & p3`1<=0 & p3<>0.REAL 2) by A1,A26;
    thus q`1<=0 & q<>0.REAL 2 by A27;
   end;
    for p being Point of TOP-REAL 2 st p in the carrier of (TOP-REAL 2)|K1
  holds g2.p=|.p.|* ((p`2/|.p.|-sn)/(1+sn))
   proof let p be Point of TOP-REAL 2;
    assume A28: p in the carrier of (TOP-REAL 2)|K1;
     A29:dom ((sn-FanMorphW)|K1)=dom (sn-FanMorphW) /\ K1 by FUNCT_1:68
     .=((the carrier of TOP-REAL 2))/\ K1 by FUNCT_2:def 1
     .=K1 by XBOOLE_1:28;
     A30:the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
     then consider p3 being Point of TOP-REAL 2 such that
     A31: p=p3 & (p3`2/|.p3.|<=sn & p3`1<=0 & p3<>0.REAL 2) by A1,A28;
     A32:(sn-FanMorphW).p
     =|[|.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)),
      |.p.|* ((p`2/|.p.|-sn)/(1+sn))]| by A1,A31,Th25;
      ((sn-FanMorphW)|K1).p=(sn-FanMorphW).p by A28,A30,FUNCT_1:72;
     then g2.p=(proj2).(|[|.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)),
            |.p.|* ((p`2/|.p.|-sn)/(1+sn))]|)
                             by A28,A29,A30,A32,FUNCT_1:23
      .=(|[|.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)),
        |.p.|* ((p`2/|.p.|-sn)/(1+sn))]|)`2 by PSCOMP_1:def 29
      .=|.p.|* ((p`2/|.p.|-sn)/(1+sn)) by EUCLID:56;
    hence g2.p=|.p.|* ((p`2/|.p.|-sn)/(1+sn));
   end;
then consider f2 being map of (TOP-REAL 2)|K1,R^1 such that
  A33:for p being Point of TOP-REAL 2 st p in the carrier of (TOP-REAL 2)|K1
  holds f2.p=|.p.|* ((p`2/|.p.|-sn)/(1+sn));
A34:f2 is continuous by A1,A25,A33,Th27;
    for p being Point of TOP-REAL 2 st p in the carrier of (TOP-REAL 2)|K1
  holds g1.p=|.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2))
   proof let p be Point of TOP-REAL 2;
    assume A35: p in the carrier of (TOP-REAL 2)|K1;
     A36:dom ((sn-FanMorphW)|K1)=dom (sn-FanMorphW) /\ K1 by FUNCT_1:68
     .=((the carrier of TOP-REAL 2))/\ K1 by FUNCT_2:def 1
     .=K1 by XBOOLE_1:28;
     A37:the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
     then consider p3 being Point of TOP-REAL 2 such that
     A38: p=p3 & (p3`2/|.p3.|<=sn & p3`1<=0 & p3<>0.REAL 2) by A1,A35;
     A39:(sn-FanMorphW).p=|[|.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)),
      |.p.|* ((p`2/|.p.|-sn)/(1+sn))]| by A1,A38,Th25;
      ((sn-FanMorphW)|K1).p=(sn-FanMorphW).p by A35,A37,FUNCT_1:72;
     then g1.p=(proj1).(|[|.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)),
          |.p.|* ((p`2/|.p.|-sn)/(1+sn))]|)
                             by A35,A36,A37,A39,FUNCT_1:23
        .=(|[|.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)),
          |.p.|* ((p`2/|.p.|-sn)/(1+sn)) ]|)`1 by PSCOMP_1:def 28
        .= |.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)) by EUCLID:56;
    hence g1.p=|.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2));
   end;
then consider f1 being map of (TOP-REAL 2)|K1,R^1 such that
  A40:for p being Point of TOP-REAL 2 st p in the carrier of (TOP-REAL 2)|K1
  holds f1.p=|.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2));
    for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds q`1<=0 & q`2/|.q.|<=sn & q<>0.REAL 2
   proof let q be Point of TOP-REAL 2;
    assume A41:q in the carrier of (TOP-REAL 2)|K1;
        the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
     then consider p3 being Point of TOP-REAL 2 such that
     A42: q=p3 & (p3`2/|.p3.|<=sn & p3`1<=0 & p3<>0.REAL 2) by A1,A41;
    thus thesis by A42;
   end;
then A43:f1 is continuous by A1,A40,Th29;
  for x,y,r,s being real number st |[x,y]| in K1 &
  r=f1.(|[x,y]|) & s=f2.(|[x,y]|) holds
  f.(|[x,y]|)=|[r,s]|
  proof let x,y,r,s be real number;
   assume A44: |[x,y]| in K1 & r=f1.(|[x,y]|) & s=f2.(|[x,y]|);
     set p99=|[x,y]|;
     A45:the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
     consider p3 being Point of TOP-REAL 2 such that
     A46: p99=p3 & (p3`2/|.p3.|<=sn & p3`1<=0 & p3<>0.REAL 2) by A1,A44;
     A47:f1.p99=|.p99.|*(-sqrt(1-((p99`2/|.p99.|-sn)/(1+sn))^2))
                          by A40,A44,A45;
      ((sn-FanMorphW)|K0).(|[x,y]|)=((sn-FanMorphW)).(|[x,y]|) by A44,FUNCT_1:
72
    .= |[|.p99.|*(-sqrt(1-((p99`2/|.p99.|-sn)/(1+sn))^2)),
      |.p99.|* ((p99`2/|.p99.|-sn)/(1+sn))]| by A1,A46,Th25
    .=|[r,s]| by A33,A44,A45,A47;
   hence f.(|[x,y]|)=|[r,s]| by A1;
  end;
hence f is continuous by A8,A9,A34,A43,JGRAPH_2:45;
end;

Lm7: for sn being Real, K1 being Subset of TOP-REAL 2
 st K1={p7 where p7 is Point of TOP-REAL 2:p7`2>=sn*|.p7.|}
 holds K1 is closed
proof let sn be Real, K1 be Subset of TOP-REAL 2;
 defpred P[Point of TOP-REAL 2] means ($1`2>=sn*|.$1.|);
 assume K1={p: p`2>=(sn)*(|.p.|)};
then A1:K1={p7 where p7 is Point of TOP-REAL 2:P[p7]};
A2: K1`={p7 where p7 is Point of TOP-REAL 2:not P[p7]} from TopCompl(A1);
  set K10=[#](TOP-REAL 2);
  reconsider g1=proj2|K10 as continuous map of (TOP-REAL 2)|K10,R^1 by Lm3;
  reconsider g0=(2 NormF)|K10 as continuous map of (TOP-REAL 2)|K10,R^1
                          by Lm5;
  consider g2 being map of (TOP-REAL 2)|K10,R^1 such that
  A3: (for q being Point of
  (TOP-REAL 2)|K10,r1 being real number st g0.q=r1
  holds g2.q=sn*r1) & g2 is continuous by JGRAPH_2:33;
  consider g3 being map of (TOP-REAL 2)|K10,R^1 such that
  A4: (for q being Point of
  (TOP-REAL 2)|K10,r1,r2 being real number st g2.q=r1 & g1.q=r2
  holds g3.q=r1-r2) & g3 is continuous by A3,JGRAPH_2:31;
  A5: (TOP-REAL 2)|K10=TOP-REAL 2 by TSEP_1:3;
   A6: for p being Point of TOP-REAL 2 holds g3.p=sn*|.p.|-p`2
   proof let p be Point of TOP-REAL 2;
       g0.p=(2 NormF).p by A5,Lm5 .=|.p.| by Def1;
     then A7: g2.p=sn*|.p.| by A3,A5;
       g1.p=proj2.p by A5,Lm3 .=p`2 by PSCOMP_1:def 29;
    hence g3.p=sn*|.p.|-p`2 by A4,A5,A7;
   end;
   reconsider g=g3 as map of TOP-REAL 2,R^1 by A5;
  A8: {p7 where p7 is Point of TOP-REAL 2: pi(g,p7)>0} c= K1`
    proof let x be set;assume x in
      {p7 where p7 is Point of TOP-REAL 2: pi(g,p7)>0};
     then consider p7 being Point of TOP-REAL 2 such that
      A9: p7=x & pi(g,p7)>0;
        pi(g,p7)=g.p7 by JORDAN2C:def 10
       .=sn*|.p7.|-p7`2 by A6;
      then sn*|.p7.| - p7`2+p7`2>0+p7`2 by A9,REAL_1:67;
      then sn*|.p7.| >0+p7`2 by XCMPLX_1:27;
     hence x in K1` by A2,A9;
    end;
    K1` c= {p7 where p7 is Point of TOP-REAL 2:pi(g,p7)>0}
   proof let x be set;assume x in K1`;
     then consider p9 being Point of TOP-REAL 2 such that
     A10: x=p9 & p9`2<sn*|.p9.| by A2;
     A11: sn*|.p9.| - p9`2>0 by A10,SQUARE_1:11;
       pi(g,p9)=g.p9 by JORDAN2C:def 10
       .=sn*|.p9.|-p9`2 by A6;
    hence x in {p7 where p7 is Point of TOP-REAL 2:
    pi(g,p7)>0} by A10,A11;
   end;
  then K1`={p7 where p7 is Point of TOP-REAL 2:pi(g,p7)>0} by A8,XBOOLE_0:def
10;
  then K1` is open by A4,A5,Th6;
 hence thesis by TOPS_1:29;
end;

Lm8: for sn being Real, K1 being Subset of TOP-REAL 2
 st K1={p7 where p7 is Point of TOP-REAL 2:p7`1>=sn*|.p7.|}
 holds K1 is closed
proof let sn be Real, K1 be Subset of TOP-REAL 2;
  defpred P[Point of TOP-REAL 2] means ($1`1>=sn*|.$1.|);
 assume K1={p: p`1>=(sn)*(|.p.|)};
then A1:K1={p7 where p7 is Point of TOP-REAL 2:P[p7]};
A2: K1`={p7 where p7 is Point of TOP-REAL 2:not P[p7]} from TopCompl(A1);
  set K10=[#](TOP-REAL 2);
  reconsider g1=proj1|K10 as continuous map of (TOP-REAL 2)|K10,R^1 by Lm2;
  reconsider g0=(2 NormF)|K10 as continuous map of (TOP-REAL 2)|K10,R^1
                          by Lm5;
  consider g2 being map of (TOP-REAL 2)|K10,R^1 such that
  A3: (for q being Point of
  (TOP-REAL 2)|K10,r1 being real number st g0.q=r1
  holds g2.q=sn*r1) & g2 is continuous by JGRAPH_2:33;
  consider g3 being map of (TOP-REAL 2)|K10,R^1 such that
  A4: (for q being Point of
  (TOP-REAL 2)|K10,r1,r2 being real number st g2.q=r1 & g1.q=r2
  holds g3.q=r1-r2) & g3 is continuous by A3,JGRAPH_2:31;
  A5: (TOP-REAL 2)|K10=TOP-REAL 2 by TSEP_1:3;
   A6: for p being Point of TOP-REAL 2 holds g3.p=sn*|.p.|-p`1
   proof let p be Point of TOP-REAL 2;
       g0.p=(2 NormF).p by A5,Lm5 .=|.p.| by Def1;
     then A7: g2.p=sn*|.p.| by A3,A5;
       g1.p=proj1.p by A5,Lm2 .=p`1 by PSCOMP_1:def 28;
    hence g3.p=sn*|.p.|-p`1 by A4,A5,A7;
   end;
   reconsider g=g3 as map of TOP-REAL 2,R^1 by A5;
  A8: {p7 where p7 is Point of TOP-REAL 2: pi(g,p7)>0} c= K1`
    proof let x be set;assume x in
      {p7 where p7 is Point of TOP-REAL 2: pi(g,p7)>0};
     then consider p7 being Point of TOP-REAL 2 such that
      A9: p7=x & pi(g,p7)>0;
        pi(g,p7)=g.p7 by JORDAN2C:def 10
       .=sn*|.p7.|-p7`1 by A6;
      then sn*|.p7.| - p7`1+p7`1>0+p7`1 by A9,REAL_1:67;
      then sn*|.p7.| >0+p7`1 by XCMPLX_1:27;
     hence x in K1` by A2,A9;
    end;
    K1` c= {p7 where p7 is Point of TOP-REAL 2:pi(g,p7)>0}
   proof let x be set;assume x in K1`;
     then consider p9 being Point of TOP-REAL 2 such that
     A10: x=p9 & p9`1<sn*|.p9.| by A2;
     A11: sn*|.p9.| - p9`1>0 by A10,SQUARE_1:11;
       pi(g,p9)=g.p9 by JORDAN2C:def 10
       .=sn*|.p9.|-p9`1 by A6;
    hence x in {p7 where p7 is Point of TOP-REAL 2: pi(g,p7)>0} by A10,A11;
   end;
  then K1`={p7 where p7 is Point of TOP-REAL 2:pi(g,p7)>0} by A8,XBOOLE_0:def
10;
  then K1` is open by A4,A5,Th6;
 hence thesis by TOPS_1:29;
end;

TopInter2 { P,Q[set] } : { p: P[p] & Q[p] } =
   { p1 where p1 is Point of TOP-REAL 2 : P[p1] } /\
   { p2 where p2 is Point of TOP-REAL 2 : Q[p2] }
proof
  set K0 = { p: P[p] & Q[p] };
  set K1 = { p7 where p7 is Point of TOP-REAL 2 : P[p7]};
  set B0 = { p1 where p1 is Point of TOP-REAL 2 : Q[p1]};
  A1:K1 /\ B0 c= K0
  proof let x be set;assume x in K1 /\ B0;
    then A2:x in K1 & x in B0 by XBOOLE_0:def 3;
    then consider p7 being Point of TOP-REAL 2 such that
    A3: p7=x & P[p7];
    consider p6 being Point of TOP-REAL 2 such that
    A4: p6=x & Q[p6] by A2;
   thus x in K0 by A3,A4;
  end;
    K0 c= K1 /\ B0
   proof let x be set;assume x in K0;
    then consider p being Point of TOP-REAL 2 such that
    A5: x=p & P[p] & Q[p];
      x in K1 & x in B0 by A5;
   hence x in K1 /\ B0 by XBOOLE_0:def 3;
   end;
   hence thesis by A1,XBOOLE_0:def 10;
end;

theorem Th32: for sn being Real, K03 being Subset of TOP-REAL 2
 st K03={p: p`2>=(sn)*(|.p.|) & p`1<=0}
 holds K03 is closed
proof
 let sn be Real, K003 be Subset of TOP-REAL 2;
 defpred P[Point of TOP-REAL 2] means ($1`2>=sn*|.$1.|);
 assume A1: K003={p: p`2>=(sn)*(|.p.|) & p`1<=0};
  reconsider K1={p7 where p7 is Point of TOP-REAL 2:P[p7]}
  as Subset of TOP-REAL 2 from TopSubset;
  A2: K1 is closed by Lm7;
  defpred Q[Point of TOP-REAL 2] means $1`1<=0;
  reconsider KX={p where p is Point of TOP-REAL 2: Q[p]}
    as Subset of TOP-REAL 2 from TopSubset;
A3: KX is closed by JORDAN6:6;
     {p : P[p] & Q[p]}
    = {p7 where p7 is Point of TOP-REAL 2:P[p7]} /\
     {p1 where p1 is Point of TOP-REAL 2: Q[p1]} from TopInter2;
  hence thesis by A1,A2,A3,TOPS_1:35;
end;

Lm9: for sn being Real, K1 being Subset of TOP-REAL 2
 st K1={p7 where p7 is Point of TOP-REAL 2:p7`2<=sn*|.p7.|}
 holds K1 is closed
proof let sn be Real, K1 be Subset of TOP-REAL 2;
  defpred P[Point of TOP-REAL 2] means ($1`2<=sn*|.$1.|);
 assume K1={p: p`2<=(sn)*(|.p.|)};
then A1:K1={p7 where p7 is Point of TOP-REAL 2:P[p7]};
A2:K1`={p7 where p7 is Point of TOP-REAL 2:not P[p7]} from TopCompl(A1);
  set K10=[#](TOP-REAL 2);
  reconsider g1=proj2|K10 as continuous map of (TOP-REAL 2)|K10,R^1 by Lm3;
  reconsider g0=(2 NormF)|K10 as continuous map of (TOP-REAL 2)|K10,R^1
                          by Lm5;
  consider g2 being map of (TOP-REAL 2)|K10,R^1 such that
  A3: (for q being Point of
  (TOP-REAL 2)|K10,r1 being real number st g0.q=r1
  holds g2.q=sn*r1) & g2 is continuous by JGRAPH_2:33;
  consider g3 being map of (TOP-REAL 2)|K10,R^1 such that
  A4: (for q being Point of
  (TOP-REAL 2)|K10,r1,r2 being real number st g2.q=r1 & g1.q=r2
  holds g3.q=r1-r2) & g3 is continuous by A3,JGRAPH_2:31;
  A5: (TOP-REAL 2)|K10=TOP-REAL 2 by TSEP_1:3;
   A6: for p being Point of TOP-REAL 2 holds g3.p=sn*|.p.|-p`2
   proof let p be Point of TOP-REAL 2;
       g0.p=(2 NormF).p by A5,Lm5 .=|.p.| by Def1;
     then A7: g2.p=sn*|.p.| by A3,A5;
       g1.p=proj2.p by A5,Lm3 .=p`2 by PSCOMP_1:def 29;
    hence g3.p=sn*|.p.|-p`2 by A4,A5,A7;
   end;
   reconsider g=g3 as map of TOP-REAL 2,R^1 by A5;
  A8: {p7 where p7 is Point of TOP-REAL 2: pi(g,p7)<0} c= K1`
    proof let x be set;assume x in
      {p7 where p7 is Point of TOP-REAL 2: pi(g,p7)<0};
     then consider p7 being Point of TOP-REAL 2 such that
      A9: p7=x & pi(g,p7)<0;
        pi(g,p7)=g.p7 by JORDAN2C:def 10
       .=sn*|.p7.|-p7`2 by A6;
      then sn*|.p7.| - p7`2+p7`2<0+p7`2 by A9,REAL_1:67;
      then sn*|.p7.| <0+p7`2 by XCMPLX_1:27;
     hence x in K1` by A2,A9;
    end;
    K1` c= {p7 where p7 is Point of TOP-REAL 2:pi(g,p7)<0}
   proof let x be set;assume x in K1`;
     then consider p9 being Point of TOP-REAL 2 such that
     A10: x=p9 & p9`2>sn*|.p9.| by A2;
     A11: sn*|.p9.| - p9`2<0 by A10,REAL_2:105;
       pi(g,p9)=g.p9 by JORDAN2C:def 10
       .=sn*|.p9.|-p9`2 by A6;
    hence x in {p7 where p7 is Point of TOP-REAL 2:
    pi(g,p7)<0} by A10,A11;
   end;
  then K1`={p7 where p7 is Point of TOP-REAL 2:pi(g,p7)<0} by A8,XBOOLE_0:def
10;
  then K1` is open by A4,A5,Th7;
 hence thesis by TOPS_1:29;
end;

Lm10: for sn being Real, K1 being Subset of TOP-REAL 2
 st K1={p7 where p7 is Point of TOP-REAL 2:p7`1<=sn*|.p7.|}
 holds K1 is closed
proof let sn be Real, K1 be Subset of TOP-REAL 2;
  defpred P[Point of TOP-REAL 2] means ($1`1<=sn*|.$1.|);
 assume K1={p: p`1<=(sn)*(|.p.|)};
then A1:K1={p7 where p7 is Point of TOP-REAL 2:P[p7]};
A2:K1`={p7 where p7 is Point of TOP-REAL 2:not P[p7]} from TopCompl(A1);
  set K10=[#](TOP-REAL 2);
  reconsider g1=proj1|K10 as continuous map of (TOP-REAL 2)|K10,R^1 by Lm2;
  reconsider g0=(2 NormF)|K10 as continuous map of (TOP-REAL 2)|K10,R^1
                          by Lm5;
  consider g2 being map of (TOP-REAL 2)|K10,R^1 such that
  A3: (for q being Point of
  (TOP-REAL 2)|K10,r1 being real number st g0.q=r1
  holds g2.q=sn*r1) & g2 is continuous by JGRAPH_2:33;
  consider g3 being map of (TOP-REAL 2)|K10,R^1 such that
  A4: (for q being Point of
  (TOP-REAL 2)|K10,r1,r2 being real number st g2.q=r1 & g1.q=r2
  holds g3.q=r1-r2) & g3 is continuous by A3,JGRAPH_2:31;
  A5: (TOP-REAL 2)|K10=TOP-REAL 2 by TSEP_1:3;
   A6: for p being Point of TOP-REAL 2 holds g3.p=sn*|.p.|-p`1
   proof let p be Point of TOP-REAL 2;
       g0.p=(2 NormF).p by A5,Lm5 .=|.p.| by Def1;
     then A7: g2.p=sn*|.p.| by A3,A5;
       g1.p=proj1.p by A5,Lm2 .=p`1 by PSCOMP_1:def 28;
    hence g3.p=sn*|.p.|-p`1 by A4,A5,A7;
   end;
   reconsider g=g3 as map of TOP-REAL 2,R^1 by A5;
  A8: {p7 where p7 is Point of TOP-REAL 2: pi(g,p7)<0} c= K1`
    proof let x be set;assume x in
      {p7 where p7 is Point of TOP-REAL 2: pi(g,p7)<0};
     then consider p7 being Point of TOP-REAL 2 such that
      A9: p7=x & pi(g,p7)<0;
        pi(g,p7)=g.p7 by JORDAN2C:def 10
       .=sn*|.p7.|-p7`1 by A6;
      then sn*|.p7.| - p7`1+p7`1<0+p7`1 by A9,REAL_1:67;
      then sn*|.p7.| <0+p7`1 by XCMPLX_1:27;
     hence x in K1` by A2,A9;
    end;
    K1` c= {p7 where p7 is Point of TOP-REAL 2:pi(g,p7)<0}
   proof let x be set;assume x in K1`;
     then consider p9 being Point of TOP-REAL 2 such that
     A10: x=p9 & p9`1>sn*|.p9.| by A2;
     A11: sn*|.p9.| - p9`1<0 by A10,REAL_2:105;
       pi(g,p9)=g.p9 by JORDAN2C:def 10
       .=sn*|.p9.|-p9`1 by A6;
    hence x in {p7 where p7 is Point of TOP-REAL 2: pi(g,p7)<0} by A10,A11;
   end;
  then K1`={p7 where p7 is Point of TOP-REAL 2:pi(g,p7)<0} by A8,XBOOLE_0:def
10;
  then K1` is open by A4,A5,Th7;
 hence thesis by TOPS_1:29;
end;

theorem Th33: for sn being Real, K03 being Subset of TOP-REAL 2
 st K03={p: p`2<=(sn)*(|.p.|) & p`1<=0}
 holds K03 is closed
proof let sn be Real, K003 be Subset of TOP-REAL 2;
  defpred P[Point of TOP-REAL 2] means ($1`2<=sn*|.$1.|);
  defpred Q[Point of TOP-REAL 2] means $1`1<=0;
 assume A1: K003={p: P[p] & Q[p]};
  reconsider K1={p7 where p7 is Point of TOP-REAL 2:P[p7]}
  as Subset of TOP-REAL 2 from TopSubset;
  A2: K1 is closed by Lm9;
  reconsider KX={p where p is Point of TOP-REAL 2: Q[p]}
    as Subset of TOP-REAL 2 from TopSubset;
A3: KX is closed by JORDAN6:6;
     {p : P[p] & Q[p]}
    = {p7 where p7 is Point of TOP-REAL 2:P[p7]} /\
     {p1 where p1 is Point of TOP-REAL 2: Q[p1]} from TopInter2;
  hence thesis by A1,A2,A3,TOPS_1:35;
end;

theorem Th34: for sn being Real,
K0,B0 being Subset of TOP-REAL 2,f being
map of (TOP-REAL 2)|K0,(TOP-REAL 2)|B0
st -1<sn & sn<1 & f=(sn-FanMorphW)|K0 &
B0=(the carrier of TOP-REAL 2) \ {0.REAL 2}
& K0={p: p`1<=0 & p<>0.REAL 2}
holds f is continuous
proof let sn be Real,K0,B0 be Subset of TOP-REAL 2,f be
map of (TOP-REAL 2)|K0,(TOP-REAL 2)|B0;
assume A1: -1<sn & sn<1 & f=(sn-FanMorphW)|K0 &
B0=(the carrier of TOP-REAL 2) \ {0.REAL 2} & K0={p: p`1<=0 & p<>0.REAL 2};
 set cn=sqrt(1-sn^2);
 set p0=|[-cn,sn]|;
 A2: p0`2=sn by EUCLID:56;
 A3: p0`1=-cn by EUCLID:56;
     sn^2<1 by A1,JGRAPH_2:8,SQUARE_1:59;
 then A4: 1-sn^2>0 by SQUARE_1:11;
 A5: now assume p0=0.REAL 2;
   then --cn=-0 by EUCLID:56,JGRAPH_2:11;
  hence contradiction by A4,SQUARE_1:93;
 end;
   --cn>0 by A4,SQUARE_1:93;
 then A6: p0`1<0 by A3,REAL_1:66;
 A7: |.p0.|=sqrt((-cn)^2+sn^2) by A2,A3,JGRAPH_3:10
       .=sqrt((cn)^2+sn^2) by SQUARE_1:61;
   cn^2=1-sn^2 by A4,SQUARE_1:def 4;
 then |.p0.|=1 by A7,SQUARE_1:83,XCMPLX_1:27;
 then A8: p0`2/|.p0.|=sn by EUCLID:56;
   p0 in K0 by A1,A5,A6;
then reconsider K1=K0 as non empty Subset of TOP-REAL 2;
    p0 in the carrier of TOP-REAL 2 & not p0 in {0.REAL 2}
               by A5,TARSKI:def 1;
  then reconsider D=B0 as non empty Subset of TOP-REAL 2 by A1,XBOOLE_0:def 4;
   A9: p0 in {p: p`2/|.p.|>=sn & p`1<=0 & p<>0.REAL 2} by A5,A6,A8;
   defpred P[Point of TOP-REAL 2] means
   $1`2/|.$1.|>=sn & $1`1<=0 & $1<>0.REAL 2;
   A10: {p: P[p]}
     is Subset of TOP-REAL 2 from SubsetD;
   A11: the carrier of ((TOP-REAL 2)|K1)=K1 by JORDAN1:1;
   A12: {p: p`2/|.p.|>=sn & p`1<=0 & p<>0.REAL 2} c= K1
    proof let x be set;assume x in
      {p: p`2/|.p.|>=sn & p`1<=0 & p<>0.REAL 2};
      then consider p such that
        A13:p=x & p`2/|.p.|>=sn & p`1<=0 & p<>0.REAL 2;
     thus x in K1 by A1,A13;
    end;
   then reconsider K00={p: p`2/|.p.|>=sn & p`1<=0 & p<>0.REAL 2}
      as non empty Subset of ((TOP-REAL 2)|K1)
                 by A9,A11;
   reconsider K001={p: p`2/|.p.|>=sn & p`1<=0 & p<>0.REAL 2}
            as non empty Subset of (TOP-REAL 2)
                       by A9,A10;
     defpred P[Point of TOP-REAL 2] means $1`2>=(sn)*(|.$1.|) & $1`1<=0;
     {p: P[p]} is Subset of
   TOP-REAL 2 from SubsetD;
   then reconsider K003={p: p`2>=(sn)*(|.p.|) & p`1<=0}
            as Subset of (TOP-REAL 2);
   A14: p0 in {p: p`2/|.p.|<=sn & p`1<=0 & p<>0.REAL 2} by A5,A6,A8;
   A15: {p: p`2/|.p.|<=sn & p`1<=0 & p<>0.REAL 2} c= K1
    proof let x be set;assume x in
      {p: p`2/|.p.|<=sn & p`1<=0 & p<>0.REAL 2};
      then consider p such that
        A16: p=x & p`2/|.p.|<=sn & p`1<=0 & p<>0.REAL 2;
     thus x in K1 by A1,A16;
    end;
   then reconsider K11={p: p`2/|.p.|<=sn & p`1<=0 & p<>0.REAL 2}
            as non empty Subset of ((TOP-REAL 2)|K1)
            by A11,A14;
    A17: p0 in {p: p`2/|.p.|<=sn & p`1<=0 & p<>0.REAL 2} by A5,A6,A8;
    defpred P[Point of TOP-REAL 2] means
    $1`2/|.$1.|<=sn & $1`1<=0 & $1<>0.REAL 2;
     {p: P[p]} is Subset of
   TOP-REAL 2 from SubsetD;
   then reconsider K111={p: p`2/|.p.|<=sn & p`1<=0 & p<>0.REAL 2}
            as non empty Subset of (TOP-REAL 2)
                       by A17;
     defpred P[Point of TOP-REAL 2] means $1`2<=(sn)*(|.$1.|) & $1`1<=0;
     {p: P[p]} is Subset of
   TOP-REAL 2 from SubsetD;
   then reconsider K004={p: p`2<=(sn)*(|.p.|) & p`1<=0}
            as Subset of (TOP-REAL 2);
    the carrier of (TOP-REAL 2)|B0=the carrier of (TOP-REAL 2)|D;
  then A18: dom f=the carrier of (TOP-REAL 2)|K1 by FUNCT_2:def 1
    .=K1 by JORDAN1:1;
  then A19: dom (f|K00)=K00 by A12,RELAT_1:91
  .= the carrier of ((TOP-REAL 2)|K1)|K00 by JORDAN1:1;
  A20: the carrier of (TOP-REAL 2)|D=D by JORDAN1:1;
    rng (f|K00) c= rng f by RELAT_1:99;
  then rng (f|K00) c=D by A20,XBOOLE_1:1;
  then f|K00 is Function of the carrier of ((TOP-REAL 2)|K1)|K00,
   the carrier of (TOP-REAL 2)|D by A19,A20,FUNCT_2:4;
  then reconsider f1=f|K00 as map of ((TOP-REAL 2)|K1)|K00,(TOP-REAL 2)|D
                           ;
  A21: dom f1=the carrier of ((TOP-REAL 2)|K1)|K00 by FUNCT_2:def 1
  .=K00 by JORDAN1:1;
  A22: dom (sn-FanMorphW)=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
  then A23: dom ((sn-FanMorphW)|K001)=K001 by RELAT_1:91
  .= the carrier of (TOP-REAL 2)|K001 by JORDAN1:1;
  A24: the carrier of (TOP-REAL 2)|K1 = K1 by JORDAN1:1;
    rng ((sn-FanMorphW)|K001) c= K1
   proof let y be set;assume y in rng ((sn-FanMorphW)|K001);
     then consider x being set such that
     A25: x in dom ((sn-FanMorphW)|K001) & y=((sn-FanMorphW)|K001).x
        by FUNCT_1:def 5;
     A26: dom ((sn-FanMorphW)|K001)=(dom (sn-FanMorphW))/\ K001 by FUNCT_1:68
     .=(the carrier of TOP-REAL 2)/\ K001 by FUNCT_2:def 1
     .=K001 by XBOOLE_1:28;
     then reconsider q=x as Point of TOP-REAL 2 by A25;
     A27: y=(sn-FanMorphW).q by A25,FUNCT_1:70;
     consider p2 being Point of TOP-REAL 2 such that
     A28: p2=q & p2`2/|.p2.|>=sn & p2`1<=0 & p2<>0.REAL 2 by A25,A26;
     A29: sn-FanMorphW.q= |[ |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1-sn))]| by A1,A28,Th25;
     set q4= |[ |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1-sn))]|;
     A30: |.q.|>=0 by TOPRNS_1:26;
       |.q.|<>0 by A28,TOPRNS_1:25;
     then A31: (|.q.|)^2>0 by SQUARE_1:74;
     A32: 1-sn>0 by A1,Th3;
        A33: q4`1= |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2))&
        q4`2= |.q.|* ((q`2/|.q.|-sn)/(1-sn))
                     by EUCLID:56;
        A34: (q`2/|.q.|-sn)>= 0 by A28,SQUARE_1:12;
          0<=(q`1)^2 by SQUARE_1:72;
        then 0+(q`2)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
        then (q`2)^2 <= (|.q.|)^2 by JGRAPH_3:10;
        then (q`2)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A31,REAL_1:73;
        then (q`2)^2/(|.q.|)^2 <= 1 by A31,XCMPLX_1:60;
        then ((q`2)/|.q.|)^2 <= 1 by SQUARE_1:69;
        then 1>=q`2/|.q.| by Th4;
        then A35: 1-sn>=q`2/|.q.|-sn by REAL_1:49;
        A36: -(1-sn)<= -0 by A32,REAL_1:50;
           -(1-sn)<= -( q`2/|.q.|-sn) by A35,REAL_1:50;
         then (-(1-sn))/(1-sn)<=(-( q`2/|.q.|-sn))/(1-sn)
                        by A32,REAL_1:73;
        then A37: -1<=(-( q`2/|.q.|-sn))/(1-sn)
                        by A32,XCMPLX_1:198;
         --(1-sn)>= -(q`2/|.q.|-sn) by A34,A36,REAL_1:50;
       then (-(q`2/|.q.|-sn))/(1-sn)<=1 by A32,REAL_2:117;
       then ((-(q`2/|.q.|-sn))/(1-sn))^2<=1 by A37,JGRAPH_2:7,SQUARE_1:59;
       then A38: 1-((-(q`2/|.q.|-sn))/(1-sn))^2>=0 by SQUARE_1:12;
       then 1-(-((q`2/|.q.|-sn))/(1-sn))^2>=0 by XCMPLX_1:188;
       then A39: 1-((q`2/|.q.|-sn)/(1-sn))^2>=0 by SQUARE_1:61;
         sqrt(1-((-(q`2/|.q.|-sn))/(1-sn))^2)>=0 by A38,SQUARE_1:def 4;
       then sqrt(1-(-(q`2/|.q.|-sn))^2/(1-sn)^2)>=0 by SQUARE_1:69;
       then sqrt(1-(q`2/|.q.|-sn)^2/(1-sn)^2)>=0 by SQUARE_1:61;
       then sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)>=0 by SQUARE_1:69;
       then A40: -sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)<= -0 by REAL_1:50;
        A41: (q4`1)^2= (|.q.|)^2*(-sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2))^2
                         by A33,SQUARE_1:68
         .= (|.q.|)^2*(sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2))^2
                     by SQUARE_1:61
         .= (|.q.|)^2*(1-((q`2/|.q.|-sn)/(1-sn))^2) by A39,SQUARE_1:def 4;
        A42: (q4`2)^2= (|.q.|)^2*((q`2/|.q.|-sn)/(1-sn))^2
                         by A33,SQUARE_1:68;
          (|.q4.|)^2=(q4`1)^2+(q4`2)^2 by JGRAPH_3:10
        .=(|.q.|)^2*(1-((q`2/|.q.|-sn)/(1-sn))^2
           +((q`2/|.q.|-sn)/(1-sn))^2) by A41,A42,XCMPLX_1:8
        .=(|.q.|)^2*1 by XCMPLX_1:27 .=(|.q.|)^2;
     then q4`1<=0 & q4<>0.REAL 2 by A30,A31,A33,A40,SQUARE_1:23,60,TOPRNS_1:24
;
    hence y in K1 by A1,A27,A29;
   end;
  then (sn-FanMorphW)|K001 is Function of the carrier of (TOP-REAL 2)|K001,
   the carrier of (TOP-REAL 2)|K1 by A23,A24,FUNCT_2:4;
  then reconsider f3=(sn-FanMorphW)|K001
   as map of (TOP-REAL 2)|K001,(TOP-REAL 2)|K1 ;
  A43: dom (f|K11)=K11 by A15,A18,RELAT_1:91
  .= the carrier of ((TOP-REAL 2)|K1)|K11 by JORDAN1:1;
  A44: the carrier of (TOP-REAL 2)|D=D by JORDAN1:1;
    rng (f|K11) c= rng f by RELAT_1:99;
  then rng (f|K11) c=D by A44,XBOOLE_1:1;
  then f|K11 is Function of the carrier of ((TOP-REAL 2)|K1)|K11,
   the carrier of (TOP-REAL 2)|D by A43,A44,FUNCT_2:4;
  then reconsider f2=f|K11 as map of ((TOP-REAL 2)|K1)|K11,(TOP-REAL 2)|D
                   ;
  A45: dom f2=the carrier of ((TOP-REAL 2)|K1)|K11 by FUNCT_2:def 1
  .=K11 by JORDAN1:1;
  A46: dom ((sn-FanMorphW)|K111)=K111 by A22,RELAT_1:91
  .= the carrier of (TOP-REAL 2)|K111 by JORDAN1:1;
  A47: the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
    rng ((sn-FanMorphW)|K111) c= K1
   proof let y be set;assume y in rng ((sn-FanMorphW)|K111);
     then consider x being set such that
     A48: x in dom ((sn-FanMorphW)|K111) & y=((sn-FanMorphW)|K111).x
        by FUNCT_1:def 5;
     A49: dom ((sn-FanMorphW)|K111)=(dom (sn-FanMorphW))/\ K111 by FUNCT_1:68
     .=(the carrier of TOP-REAL 2)/\ K111 by FUNCT_2:def 1
     .=K111 by XBOOLE_1:28;
     then reconsider q=x as Point of TOP-REAL 2 by A48;
     A50: y=(sn-FanMorphW).q by A48,FUNCT_1:70;
     consider p2 being Point of TOP-REAL 2 such that
     A51: p2=q & p2`2/|.p2.|<=sn & p2`1<=0 & p2<>0.REAL 2 by A48,A49;
     A52: sn-FanMorphW.q= |[ |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1+sn))]| by A1,A51,Th25;
     set q4= |[ |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1+sn))]|;
     A53: |.q.|>=0 by TOPRNS_1:26;
       |.q.|<>0 by A51,TOPRNS_1:25;
     then A54: (|.q.|)^2>0 by SQUARE_1:74;
   A55: 1+sn>0 by A1,Th3;
        A56: q4`1= |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2))&
        q4`2= |.q.|* ((q`2/|.q.|-sn)/(1+sn))
                     by EUCLID:56;
        A57: (q`2/|.q.|-sn)<=0 by A51,REAL_2:106;
        A58: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
          0<=(q`1)^2 by SQUARE_1:72;
        then 0+(q`2)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
        then (q`2)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A54,A58,REAL_1:73;
        then (q`2)^2/(|.q.|)^2 <= 1 by A54,XCMPLX_1:60;
        then ((q`2)/|.q.|)^2 <= 1 by SQUARE_1:69;
        then -1<=q`2/|.q.| by Th4;
        then -1-sn<=q`2/|.q.|-sn by REAL_1:49;
        then A59: -(1+sn)<=q`2/|.q.|-sn by XCMPLX_1:161;
         (1+sn)/(1+sn)>=(q`2/|.q.|-sn)/(1+sn) by A55,A57,REAL_1:73;
        then A60: 1>=(q`2/|.q.|-sn)/(1+sn) by A55,XCMPLX_1:60;
           (-(1+sn))/(1+sn)<=(( q`2/|.q.|-sn))/(1+sn)
                        by A55,A59,REAL_1:73;
        then -1<=(( q`2/|.q.|-sn))/(1+sn)
                        by A55,XCMPLX_1:198;
       then (((q`2/|.q.|-sn))/(1+sn))^2<=1^2 by A60,JGRAPH_2:7;
       then A61: 1-((q`2/|.q.|-sn)/(1+sn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((q`2/|.q.|-sn)/(1+sn)))^2>=0 by SQUARE_1:61;
       then 1-((-(q`2/|.q.|-sn))/(1+sn))^2>=0 by XCMPLX_1:188;
       then sqrt(1-((-(q`2/|.q.|-sn))/(1+sn))^2)>=0 by SQUARE_1:def 4;
       then sqrt(1-(-(q`2/|.q.|-sn))^2/(1+sn)^2)>=0 by SQUARE_1:69;
       then sqrt(1-(q`2/|.q.|-sn)^2/(1+sn)^2)>=0 by SQUARE_1:61;
       then sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)>=0 by SQUARE_1:69;
       then -sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)<= -0 by REAL_1:50;
       then A62: q4`1<= 0 by A53,A56,SQUARE_1:23;
        A63: (q4`1)^2= (|.q.|)^2*(-sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2))^2
                         by A56,SQUARE_1:68
         .= (|.q.|)^2*(sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2))^2
                     by SQUARE_1:61
         .= (|.q.|)^2*(1-((q`2/|.q.|-sn)/(1+sn))^2) by A61,SQUARE_1:def 4;
        A64: (q4`2)^2= (|.q.|)^2*((q`2/|.q.|-sn)/(1+sn))^2
                         by A56,SQUARE_1:68;
          (|.q4.|)^2=(q4`1)^2+(q4`2)^2 by JGRAPH_3:10
        .=(|.q.|)^2*(1-((q`2/|.q.|-sn)/(1+sn))^2
           +((q`2/|.q.|-sn)/(1+sn))^2) by A63,A64,XCMPLX_1:8
        .=(|.q.|)^2*1 by XCMPLX_1:27 .=(|.q.|)^2;
        then q4<>0.REAL 2 by A54,SQUARE_1:60,TOPRNS_1:24;
    hence y in K1 by A1,A50,A52,A62;
   end;
  then (sn-FanMorphW)|K111 is Function of the carrier of (TOP-REAL 2)|K111,
   the carrier of (TOP-REAL 2)|K1 by A46,A47,FUNCT_2:4;
  then reconsider f4=(sn-FanMorphW)|K111
   as map of (TOP-REAL 2)|K111,(TOP-REAL 2)|K1 ;
  set T1= ((TOP-REAL 2)|K1)|K00,T2=((TOP-REAL 2)|K1)|K11;
  A65: [#](((TOP-REAL 2)|K1)|K00)=K00 by PRE_TOPC:def 10;
  A66: [#](((TOP-REAL 2)|K1)|K11)=K11 by PRE_TOPC:def 10;
  A67: [#]((TOP-REAL 2)|K1)=K1 by PRE_TOPC:def 10;
  A68: K00 \/ K11 c= K1
   proof let x be set;assume
     A69: x in K00 \/ K11;
       now per cases by A69,XBOOLE_0:def 2;
     case x in K00;
       then consider p8 being Point of TOP-REAL 2 such that
       A70: p8=x & p8`2/|.p8.|>=sn & p8`1<=0 & p8<>0.REAL 2;
      thus x in K1 by A1,A70;
     case x in K11;
       then consider p8 being Point of TOP-REAL 2 such that
       A71: p8=x & p8`2/|.p8.|<=sn & p8`1<=0 & p8<>0.REAL 2;
      thus x in K1 by A1,A71;
     end;
    hence x in K1;
   end;
  A72: K1 c= K00 \/ K11
   proof let x be set;assume x in K1;
     then consider p such that
     A73: p=x & (p`1<=0 & p<>0.REAL 2) by A1;
       now per cases;
     case p`2/|.p.|>=sn;
       then x in K00 by A73;
      hence x in K00 \/ K11 by XBOOLE_0:def 2;
     case p`2/|.p.|<sn;
       then x in K11 by A73;
      hence x in K00 \/ K11 by XBOOLE_0:def 2;
     end;
    hence x in K00 \/ K11;
   end;
  then A74: [#](((TOP-REAL 2)|K1)|K00) \/ [#](((TOP-REAL 2)|K1)|K11)
        =[#]((TOP-REAL 2)|K1) by A65,A66,A67,A68,XBOOLE_0:def 10;
  A75: K003 is closed by Th32;
  A76: K003 /\ K1 c= K00
   proof let x be set;assume x in K003 /\ K1;
     then A77: x in K003 & x in K1 by XBOOLE_0:def 3;
     then consider q1 being Point of TOP-REAL 2 such that
     A78: q1=x & q1`2>=(sn)*(|.q1.|) & q1`1<=0;
     consider q2 being Point of TOP-REAL 2 such that
     A79: q2=x & q2`1<=0 & q2<>0.REAL 2 by A1,A77;
     A80: |.q2.|<>0 by A79,TOPRNS_1:25;
       |.q1.|>0 by A78,A79,Lm1;
     then q1`2/|.q1.|>=(sn)*(|.q1.|)/|.q1.| by A78,REAL_1:73;
     then q1`2/|.q1.|>=(sn) by A78,A79,A80,XCMPLX_1:90;
    hence x in K00 by A78,A79;
   end;
    K00 c= K003 /\ K1
   proof let x be set;assume x in K00;
     then consider p such that
     A81: p=x & (p`2/|.p.|>=sn & p`1<=0 & p<>0.REAL 2);
     A82: |.p.|<>0 by A81,TOPRNS_1:25;
       |.p.|>=0 by TOPRNS_1:26;
     then p`2/|.p.|*|.p.|>=(sn)*(|.p.|) by A81,AXIOMS:25;
     then p`2>=(sn)*(|.p.|) by A82,XCMPLX_1:88;
     then A83: x in K003 by A81;
       x in K1 by A1,A81;
    hence x in K003 /\ K1 by A83,XBOOLE_0:def 3;
   end;
  then K00=K003 /\ [#]((TOP-REAL 2)|K1) by A67,A76,XBOOLE_0:def 10;
  then A84: K00 is closed by A75,PRE_TOPC:43;
  A85: K004 is closed by Th33;
  A86: K004 /\ K1 c= K11
   proof let x be set;assume x in K004 /\ K1;
     then A87: x in K004 & x in K1 by XBOOLE_0:def 3;
     then consider q1 being Point of TOP-REAL 2 such that
     A88: q1=x & q1`2<=(sn)*(|.q1.|) & q1`1<=0;
     consider q2 being Point of TOP-REAL 2 such that
     A89: q2=x & q2`1<=0 & q2<>0.REAL 2 by A1,A87;
     A90: |.q2.|<>0 by A89,TOPRNS_1:25;
       |.q1.|>0 by A88,A89,Lm1;
     then q1`2/|.q1.|<=(sn)*(|.q1.|)/|.q1.| by A88,REAL_1:73;
     then q1`2/|.q1.|<=(sn) by A88,A89,A90,XCMPLX_1:90;
    hence x in K11 by A88,A89;
   end;
    K11 c= K004 /\ K1
   proof let x be set;assume x in K11;
     then consider p such that
     A91: p=x & (p`2/|.p.|<=sn & p`1<=0 & p<>0.REAL 2);
     A92: |.p.|<>0 by A91,TOPRNS_1:25;
       |.p.|>=0 by TOPRNS_1:26;
     then p`2/|.p.|*|.p.|<=(sn)*(|.p.|) by A91,AXIOMS:25;
     then p`2<=(sn)*(|.p.|) by A92,XCMPLX_1:88;
     then A93: x in K004 by A91;
       x in K1 by A1,A91;
    hence x in K004 /\ K1 by A93,XBOOLE_0:def 3;
   end;
  then K11=K004 /\ [#]((TOP-REAL 2)|K1) by A67,A86,XBOOLE_0:def 10;
  then A94: K11 is closed by A85,PRE_TOPC:43;
  A95: ((TOP-REAL 2)|K1)|K00=(TOP-REAL 2)|K001 by GOBOARD9:4;
    K1 c= D
   proof let x be set;assume x in K1;
     then consider p6 being Point of TOP-REAL 2 such that
     A96: p6=x & p6`1<=0 & p6<>0.REAL 2 by A1;
       x in the carrier of TOP-REAL 2 & not x in {0.REAL 2} by A96,TARSKI:def 1
;
    hence x in D by A1,XBOOLE_0:def 4;
   end;
  then D=K1 \/ D by XBOOLE_1:12;
  then A97: (TOP-REAL 2)|K1 is SubSpace of (TOP-REAL 2)|D by TOPMETR:5;
    the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
  then A98: f1= f3 by A1,FUNCT_1:82;
    f3 is continuous by A1,Th30;
  then A99: f1 is continuous by A95,A97,A98,TOPMETR:7;
  A100: ((TOP-REAL 2)|K1)|K11=(TOP-REAL 2)|K111 by GOBOARD9:4;
    the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
  then A101: f2= f4 by A1,FUNCT_1:82;
    f4 is continuous by A1,Th31;
  then A102: f2 is continuous by A97,A100,A101,TOPMETR:7;
  A103: D<>{};
  A104: the carrier of (TOP-REAL 2)|B0=B0 by JORDAN1:1;
    for p being set st p in ([#]T1)/\([#]T2) holds f1.p = f2.p
   proof let p be set;
    assume p in ([#]T1)/\([#]T2);
     then A105: p in K00 & p in K11 by A65,A66,XBOOLE_0:def 3;
    hence f1.p=f.p by FUNCT_1:72 .=f2.p by A105,FUNCT_1:72;
   end;
  then consider h being map of (TOP-REAL 2)|K1,(TOP-REAL 2)|D
  such that
  A106: h=f1+*f2 & h is continuous by A65,A66,A74,A84,A94,A99,A102,JGRAPH_2:9;
  A107: dom h=the carrier of ((TOP-REAL 2)|K1) by FUNCT_2:def 1;
  A108: the carrier of ((TOP-REAL 2)|K1)=K0 by JORDAN1:1;
  A109: K0=the carrier of ((TOP-REAL 2)|K0) by JORDAN1:1
  .=dom f by A103,A104,FUNCT_2:def 1;
    for y being set st y in dom h holds h.y=f.y
   proof let y be set;assume
    A110: y in dom h;
      now per cases by A72,A107,A108,A110,XBOOLE_0:def 2;
    case A111: y in K00 & not y in K11;
      then y in dom f1 \/ dom f2 by A21,XBOOLE_0:def 2;
     hence h.y=f1.y by A45,A106,A111,FUNCT_4:def 1 .=f.y by A111,FUNCT_1:72;
    case A112: y in K11;
      then y in dom f1 \/ dom f2 by A45,XBOOLE_0:def 2;
     hence h.y=f2.y by A45,A106,A112,FUNCT_4:def 1 .=f.y by A112,FUNCT_1:72;
    end;
   hence h.y=f.y;
   end;
 hence thesis by A106,A107,A108,A109,FUNCT_1:9;
end;

theorem Th35: for sn being Real,
K0,B0 being Subset of TOP-REAL 2,f being
map of (TOP-REAL 2)|K0,(TOP-REAL 2)|B0
st -1<sn & sn<1 & f=(sn-FanMorphW)|K0 &
B0=(the carrier of TOP-REAL 2) \ {0.REAL 2}
& K0={p: p`1>=0 & p<>0.REAL 2}
holds f is continuous
proof let sn be Real,K0,B0 be Subset of TOP-REAL 2,f be
map of (TOP-REAL 2)|K0,(TOP-REAL 2)|B0;
assume A1:  -1<sn & sn<1 & f=(sn-FanMorphW)|K0 &
B0=(the carrier of TOP-REAL 2)\ {0.REAL 2} &
K0={p: p`1>=0 & p<>0.REAL 2};
 set cn=sqrt(1-sn^2);
 set p0=|[cn,-sn]|;
 A2: p0`1=cn by EUCLID:56;
     sn^2<1^2 by A1,JGRAPH_2:8;
 then A3: 1-sn^2>0 by SQUARE_1:11,59;
 then A4: p0<>0.REAL 2 by A2,JGRAPH_2:11,SQUARE_1:93;
   p0`1>0 by A2,A3,SQUARE_1:93;
 then p0 in K0 by A1,JGRAPH_2:11;
then reconsider K1=K0 as non empty Subset of TOP-REAL 2;
   p0 in the carrier of TOP-REAL 2 & not p0 in {0.REAL 2}
               by A4,TARSKI:def 1;
  then reconsider D=B0 as non empty Subset of TOP-REAL 2 by A1,XBOOLE_0:def 4;
   A5: K1 c= D
    proof let x be set;assume x in K1;
      then consider p2 being Point of TOP-REAL 2 such that
      A6: p2=x & p2`1>=0 & p2<>0.REAL 2 by A1;
        not p2 in {0.REAL 2} by A6,TARSKI:def 1;
     hence x in D by A1,A6,XBOOLE_0:def 4;
    end;
     for p being Point of (TOP-REAL 2)|K1,V being Subset of (TOP-REAL 2)|D
      st f.p in V & V is open holds
      ex W being Subset of (TOP-REAL 2)|K1 st
      p in W & W is open & f.:W c= V
    proof let p be Point of (TOP-REAL 2)|K1,V be Subset of (TOP-REAL 2)|D;
      assume A7: f.p in V & V is open;
      then consider V2 being Subset of TOP-REAL 2 such that
      A8: V2 is open & V2 /\ [#]((TOP-REAL 2)|D)=V by TOPS_2:32;
      A9: p in the carrier of (TOP-REAL 2)|K1;
      A10: the carrier of (TOP-REAL 2)|K1=[#]((TOP-REAL 2)|K1) by PRE_TOPC:12;
      A11: [#]((TOP-REAL 2)|K1)=K1 by PRE_TOPC:def 10;
        V2 /\ [#]((TOP-REAL 2)|K1) c= the carrier of (TOP-REAL 2)|K1
       proof let x be set;assume x in V2 /\ [#]((TOP-REAL 2)|K1);
         then x in [#]((TOP-REAL 2)|K1) by XBOOLE_0:def 3;
        hence x in the carrier of (TOP-REAL 2)|K1;
       end;
      then reconsider W2=V2 /\ [#]((TOP-REAL 2)|K1) as Subset of (TOP-REAL 2)|
K1
                               ;
      A12: W2 is open by A8,TOPS_2:32;
      A13: f.p=(sn-FanMorphW).p by A1,A10,A11,FUNCT_1:72;
      consider q being Point of TOP-REAL 2 such that
      A14: q=p & q`1>=0 & q <>0.REAL 2 by A1,A9,A10,A11;
        (sn-FanMorphW).q=q by A14,Th23;
      then p in V2 & p in [#]((TOP-REAL 2)|D) by A7,A8,A13,A14,XBOOLE_0:def 3;
      then A15: p in W2 by A10,XBOOLE_0:def 3;
        f.:W2 c= V
       proof let y be set;assume y in f.:W2;
         then consider x being set such that
         A16: x in dom f & x in W2 & y=f.x by FUNCT_1:def 12;
           f is Function of the carrier of (TOP-REAL 2)|K1,
           the carrier of (TOP-REAL 2)|D;
         then dom f= K1 by A10,A11,FUNCT_2:def 1;
         then consider p4 being Point of TOP-REAL 2 such that
         A17: x=p4 & p4`1>=0 & p4<>0.REAL 2 by A1,A16;
         A18: f.p4=(sn-FanMorphW).p4 by A1,A10,A11,A16,A17,FUNCT_1:72
         .=p4 by A17,Th23;
         A19: p4 in V2 & p4 in [#]((TOP-REAL 2)|K1) by A16,A17,XBOOLE_0:def 3
;
         then p4 in D by A5,A11;
         then p4 in [#]((TOP-REAL 2)|D) by PRE_TOPC:def 10;
        hence y in V by A8,A16,A17,A18,A19,XBOOLE_0:def 3;
       end;
     hence ex W being Subset of (TOP-REAL 2)|K1 st
      p in W & W is open & f.:W c= V by A12,A15;
    end;
  hence f is continuous by JGRAPH_2:20;
end;

theorem Th36:
 for B0 being Subset of TOP-REAL 2,
     K0 being Subset of (TOP-REAL 2)|B0
st B0=(the carrier of TOP-REAL 2) \ {0.REAL 2}
& K0={p: p`1<=0 & p<>0.REAL 2} holds K0 is closed
proof
 defpred P[Point of TOP-REAL 2] means $1`1<=0;
 set I1 = {p: P[p] & p<>0.REAL 2};
 set J0 = (the carrier of TOP-REAL 2) \ {0.REAL 2};
 let B0 be Subset of TOP-REAL 2,K0 be Subset of (TOP-REAL 2)|B0;
  assume A1: B0=J0 & K0=I1;
  reconsider K1={p7 where p7 is Point of TOP-REAL 2:P[p7]}
  as Subset of TOP-REAL 2 from TopSubset;
  A2:K1 is closed by JORDAN6:6;
    I1 = {p7 where p7 is Point of TOP-REAL 2 : P[p7]} /\ J0
    from TopInter;
  then K0=K1 /\ [#]((TOP-REAL 2)|B0) by A1,PRE_TOPC:def 10;
hence K0 is closed by A2,PRE_TOPC:43;
end;

theorem Th37: for sn being Real,
 B0 being Subset of TOP-REAL 2,K0 being Subset of
(TOP-REAL 2)|B0,f being map of ((TOP-REAL 2)|B0)|K0,((TOP-REAL 2)|B0)
st -1<sn & sn<1 & f=(sn-FanMorphW)|K0 &
B0=(the carrier of TOP-REAL 2) \ {0.REAL 2}
& K0={p: p`1<=0 & p<>0.REAL 2}
holds f is continuous
proof let sn be Real,
 B0 be Subset of TOP-REAL 2,K0 be Subset of
(TOP-REAL 2)|B0,f being map of ((TOP-REAL 2)|B0)|K0,((TOP-REAL 2)|B0);
defpred P[Point of TOP-REAL 2] means $1`1<=0 & $1<>0.REAL 2;
assume A1:  -1<sn & sn<1 & f=(sn-FanMorphW)|K0 &
  B0=(the carrier of TOP-REAL 2) \ {0.REAL 2}
   & K0={p: p`1<=0 & p<>0.REAL 2 };
  reconsider K1={p: P[p] } as Subset of TOP-REAL 2
    from TopSubset;
    K0 c= B0
   proof let x be set;assume x in K0;
      then consider p8 being Point of TOP-REAL 2 such that
      A2: x=p8 & (p8`1<=0 & p8<>0.REAL 2) by A1;
       not x in {0.REAL 2} by A2,TARSKI:def 1;
    hence x in B0 by A1,A2,XBOOLE_0:def 4;
   end;
  then ((TOP-REAL 2)|B0)|K0=(TOP-REAL 2)|K1 by A1,JORDAN6:47;
hence f is continuous by A1,Th34;
end;

theorem Th38:
 for B0 being Subset of TOP-REAL 2,
     K0 being Subset of (TOP-REAL 2)|B0
st B0=(the carrier of TOP-REAL 2) \ {0.REAL 2}
& K0={p: p`1>=0 & p<>0.REAL 2} holds K0 is closed
proof
 defpred P[Point of TOP-REAL 2] means $1`1>=0;
 set I1 = {p: P[p] & p<>0.REAL 2};
 set J0 = (the carrier of TOP-REAL 2) \ {0.REAL 2};
 let B0 be Subset of TOP-REAL 2,K0 be Subset of (TOP-REAL 2)|B0;
 assume A1: B0=J0 & K0=I1;
  reconsider K1={p7 where p7 is Point of TOP-REAL 2:P[p7]}
    as Subset of TOP-REAL 2 from TopSubset;
  A2:K1 is closed by JORDAN6:5;
    I1 = {p7 where p7 is Point of TOP-REAL 2 : P[p7]} /\ J0
    from TopInter;
  then K0=K1 /\ [#]((TOP-REAL 2)|B0) by A1,PRE_TOPC:def 10;
hence K0 is closed by A2,PRE_TOPC:43;
end;

theorem Th39: for sn being Real,
 B0 being Subset of TOP-REAL 2,K0 being Subset of
(TOP-REAL 2)|B0,f being map of ((TOP-REAL 2)|B0)|K0,((TOP-REAL 2)|B0)
st -1<sn & sn<1 & f=(sn-FanMorphW)|K0 &
B0=(the carrier of TOP-REAL 2) \ {0.REAL 2}
& K0={p: p`1>=0 & p<>0.REAL 2}
holds f is continuous
proof let sn be Real,
 B0 be Subset of TOP-REAL 2,K0 be Subset of
(TOP-REAL 2)|B0,f be map of ((TOP-REAL 2)|B0)|K0,((TOP-REAL 2)|B0);
 assume A1:  -1<sn & sn<1 & f=(sn-FanMorphW)|K0 &
  B0=(the carrier of TOP-REAL 2) \ {0.REAL 2}
  & K0={p: p`1>=0 & p<>0.REAL 2};
    the carrier of (TOP-REAL 2)|B0= B0 by JORDAN1:1;
  then K0 c= the carrier of TOP-REAL 2 by XBOOLE_1:1;
  then reconsider K1=K0 as Subset of TOP-REAL 2;
    K0 c= B0
   proof let x be set;assume x in K0;
      then consider p8 being Point of TOP-REAL 2 such that
      A2: x=p8 & (p8`1>=0 & p8<>0.REAL 2) by A1;
       not x in {0.REAL 2} by A2,TARSKI:def 1;
    hence x in B0 by A1,A2,XBOOLE_0:def 4;
   end;
  then ((TOP-REAL 2)|B0)|K0=(TOP-REAL 2)|K1 by JORDAN6:47;
hence thesis by A1,Th35;
end;

theorem Th40: for sn being Real,p being Point of TOP-REAL 2
holds |.(sn-FanMorphW).p.|=|.p.|
proof let sn be Real,p be Point of TOP-REAL 2;
       set z=(sn-FanMorphW).p;
        reconsider q=p as Point of TOP-REAL 2;
        reconsider qz=z as Point of TOP-REAL 2;
        A1: |.q.|>=0 by TOPRNS_1:26;
          now per cases;
        case A2: q`2/|.q.|>=sn & q`1<0;
          then (sn-FanMorphW).q= |[
          |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1-sn))]| by Th23;
          then A3: qz`1= |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2))&
                  qz`2= |.q.|* ((q`2/|.q.|-sn)/(1-sn)) by EUCLID:56;
            |.q.|<>0 by A2,JGRAPH_2:11,TOPRNS_1:25;
          then A4: (|.q.|)^2>0 by SQUARE_1:74;
          A5: (q`2/|.q.|-sn)>=0 by A2,SQUARE_1:12;
          A6: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
            0<=(q`1)^2 by SQUARE_1:72;
          then 0+(q`2)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
          then (q`2)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A4,A6,REAL_1:73;
          then (q`2)^2/(|.q.|)^2 <= 1 by A4,XCMPLX_1:60;
          then ((q`2)/|.q.|)^2 <= 1 by SQUARE_1:69;
          then 1>=q`2/|.q.| by Th4;
          then A7: 1-sn>=q`2/|.q.|-sn by REAL_1:49;
            now per cases;
          case A8: 1-sn=0;
            A9:((q`2/|.q.|-sn)/(1-sn))=(q`2/|.q.|-sn)*(1-sn)" by XCMPLX_0:def 9
             .= (q`2/|.q.|-sn)*0 by A8,XCMPLX_0:def 7 .=0;
             then -sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)=-1 by SQUARE_1:60,83;
             then (sn-FanMorphW).q= |[|.q.|*(-1),|.q.|*0]| by A2,A9,Th23
            .=|[-(|.q.|),0]| by XCMPLX_1:180;
             then ((sn-FanMorphW).q)`1=-(|.q.|) & ((sn-FanMorphW).q)`2=0
                              by EUCLID:56;
            then |.(sn-FanMorphW).p.|=sqrt((-(|.q.|))^2+0) by JGRAPH_3:10,
SQUARE_1:60
            .=sqrt((|.q.|)^2) by SQUARE_1:61
            .=|.q.| by A1,SQUARE_1:89;
           hence |.(sn-FanMorphW).p.|=|.p.|;
          case A10: 1-sn<>0;
              now per cases by A10;
            case A11: 1-sn>0;
              then A12: -(1-sn)<= -0 by REAL_1:50;
                -(1-sn)<= -( q`2/|.q.|-sn) by A7,REAL_1:50;
              then (-(1-sn))/(1-sn)<=(-( q`2/|.q.|-sn))/(1-sn)
                        by A11,REAL_1:73;
              then A13: -1<=(-( q`2/|.q.|-sn))/(1-sn)
                        by A11,XCMPLX_1:198;
               --(1-sn)>= -(q`2/|.q.|-sn) by A5,A12,REAL_1:50;
              then (-(q`2/|.q.|-sn))/(1-sn)<=1 by A11,REAL_2:117;
              then ((-(q`2/|.q.|-sn))/(1-sn))^2<=1^2 by A13,JGRAPH_2:7;
              then 1-((-(q`2/|.q.|-sn))/(1-sn))^2>=0 by SQUARE_1:12,59;
              then 1-(-((q`2/|.q.|-sn))/(1-sn))^2>=0 by XCMPLX_1:188;
              then A14: 1-((q`2/|.q.|-sn)/(1-sn))^2>=0 by SQUARE_1:61;
              A15: (qz`1)^2= (|.q.|)^2*(-sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2))^2
                         by A3,SQUARE_1:68
                  .= (|.q.|)^2*(sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2))^2
                     by SQUARE_1:61
                  .= (|.q.|)^2*(1-((q`2/|.q.|-sn)/(1-sn))^2)
                            by A14,SQUARE_1:def 4;
              A16: (qz`2)^2= (|.q.|)^2*((q`2/|.q.|-sn)/(1-sn))^2
                         by A3,SQUARE_1:68;
                (|.qz.|)^2=(qz`1)^2+(qz`2)^2 by JGRAPH_3:10
               .=(|.q.|)^2*(1-((q`2/|.q.|-sn)/(1-sn))^2
                    +((q`2/|.q.|-sn)/(1-sn))^2) by A15,A16,XCMPLX_1:8
               .=(|.q.|)^2*1 by XCMPLX_1:27 .=(|.q.|)^2;
              then A17: sqrt((|.qz.|)^2)=|.q.| by A1,SQUARE_1:89;
                |.qz.|>=0 by TOPRNS_1:26;
             hence |.(sn-FanMorphW).p.|=|.p.| by A17,SQUARE_1:89;
            case A18: 1-sn<0;
              A19: q`2/|.q.|-sn>=0 by A2,SQUARE_1:12;
                0<(q`1)^2 by A2,SQUARE_1:74;
              then 0+(q`2)^2<(q`1)^2+(q`2)^2 by REAL_1:67;
              then (q`2)^2/(|.q.|)^2 < (|.q.|)^2/(|.q.|)^2 by A4,A6,REAL_1:73;
              then (q`2)^2/(|.q.|)^2 < 1 by A4,XCMPLX_1:60;
              then ((q`2)/|.q.|)^2 < 1 by SQUARE_1:69;
              then 1 > q`2/|.p.| by Th5;
              then q`2/|.q.|-sn<1-sn by REAL_1:54;
             hence |.(sn-FanMorphW).p.|=|.p.| by A18,A19,AXIOMS:22;
            end;
           hence |.(sn-FanMorphW).p.|=|.p.|;
          end;
         hence |.(sn-FanMorphW).p.|=|.p.|;
        case A20: q`2/|.q.|<sn & q`1<0;
          then A21:(sn-FanMorphW).q= |[
          |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1+sn))]| by Th24;
          then A22: qz`1= |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2))&
                  qz`2= |.q.|* ((q`2/|.q.|-sn)/(1+sn))
                                by EUCLID:56;
            |.q.|<>0 by A20,JGRAPH_2:11,TOPRNS_1:25;
          then A23: (|.q.|)^2>0 by SQUARE_1:74;
          A24: (q`2/|.q.|-sn)<0 by A20,REAL_2:105;
          A25: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
            0<=(q`1)^2 by SQUARE_1:72;
          then 0+(q`2)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
          then (q`2)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A23,A25,REAL_1:73;
          then (q`2)^2/(|.q.|)^2 <= 1 by A23,XCMPLX_1:60;
          then ((q`2)/|.q.|)^2 <= 1 by SQUARE_1:69;
          then -1<=q`2/|.q.| by Th4;
          then A26: -1-sn<=q`2/|.q.|-sn by REAL_1:49;
            now per cases;
          case A27: 1+sn=0;
             ((q`2/|.q.|-sn)/(1+sn))=(q`2/|.q.|-sn)*(1+sn)" by XCMPLX_0:def 9
             .= (q`2/|.q.|-sn)*0 by A27,XCMPLX_0:def 7 .=0;
            then (sn-FanMorphW).q=|[-(|.q.|),0]| by A21,SQUARE_1:60,83,XCMPLX_1
:180;
             then ((sn-FanMorphW).q)`1=-(|.q.|) & ((sn-FanMorphW).q)`2=0
                              by EUCLID:56;
            then |.(sn-FanMorphW).p.|=sqrt((-(|.q.|))^2+0) by JGRAPH_3:10,
SQUARE_1:60
            .=sqrt((|.q.|)^2) by SQUARE_1:61
            .=|.q.| by A1,SQUARE_1:89;
           hence |.(sn-FanMorphW).p.|=|.p.|;
          case A28: 1+sn<>0;
              now per cases by A28;
            case A29: 1+sn>0;
                -(1+sn)<= ( q`2/|.q.|-sn) by A26,XCMPLX_1:161;
              then (-(1+sn))/(1+sn)<=(( q`2/|.q.|-sn))/(1+sn)
                        by A29,REAL_1:73;
              then A30: -1<=(( q`2/|.q.|-sn))/(1+sn) by A29,XCMPLX_1:198;
                (1+sn)>= (q`2/|.q.|-sn) by A24,A29,AXIOMS:22;
              then ((q`2/|.q.|-sn))/(1+sn)<=1 by A29,REAL_2:117;
              then (((q`2/|.q.|-sn))/(1+sn))^2<=1 by A30,JGRAPH_2:7,SQUARE_1:59
;
              then A31: 1-(((q`2/|.q.|-sn))/(1+sn))^2>=0 by SQUARE_1:12;
              A32: (qz`1)^2= (|.q.|)^2*(-sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2))^2
                         by A22,SQUARE_1:68
                  .= (|.q.|)^2*(sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2))^2
                     by SQUARE_1:61
                  .= (|.q.|)^2*(1-((q`2/|.q.|-sn)/(1+sn))^2)
                            by A31,SQUARE_1:def 4;
              A33: (qz`2)^2= (|.q.|)^2*((q`2/|.q.|-sn)/(1+sn))^2
                         by A22,SQUARE_1:68;
                (|.qz.|)^2=(qz`1)^2+(qz`2)^2 by JGRAPH_3:10
               .=(|.q.|)^2*(1-((q`2/|.q.|-sn)/(1+sn))^2
                    +((q`2/|.q.|-sn)/(1+sn))^2) by A32,A33,XCMPLX_1:8
               .=(|.q.|)^2*1 by XCMPLX_1:27 .=(|.q.|)^2;
              then A34: sqrt((|.qz.|)^2)=|.q.| by A1,SQUARE_1:89;
                |.qz.|>=0 by TOPRNS_1:26;
             hence |.(sn-FanMorphW).p.|=|.p.| by A34,SQUARE_1:89;
            case 1+sn<0;
              then A35: -(1+sn)>-0 by REAL_1:50;
                0<(q`1)^2 by A20,SQUARE_1:74;
              then 0+(q`2)^2<(q`1)^2+(q`2)^2 by REAL_1:67;
              then (q`2)^2/(|.q.|)^2 < (|.q.|)^2/(|.q.|)^2 by A23,A25,REAL_1:73
;
              then (q`2)^2/(|.q.|)^2 < 1 by A23,XCMPLX_1:60;
              then ((q`2)/|.q.|)^2 < 1 by SQUARE_1:69;
              then -1 < q`2/|.p.| by Th5;
              then q`2/|.q.|-sn>-1-sn by REAL_1:54;
              then q`2/|.q.|-sn>-(1+sn) by XCMPLX_1:161;
             hence |.(sn-FanMorphW).p.|=|.p.| by A24,A35,AXIOMS:22;
            end;
           hence |.(sn-FanMorphW).p.|=|.p.|;
          end;
         hence |.(sn-FanMorphW).p.|=|.p.|;
        case q`1>=0;
         hence |.(sn-FanMorphW).p.|=|.p.| by Th23;
        end;
 hence |.(sn-FanMorphW).p.|=|.p.|;
end;

theorem Th41: for sn being Real,x,K0 being set
 st -1<sn & sn<1 & x in K0 & K0={p: p`1<=0 & p<>0.REAL 2}
holds (sn-FanMorphW).x in K0
proof let sn be Real,x,K0 be set;
 assume A1:  -1<sn & sn<1 & x in K0 & K0={p: p`1<=0 & p<>0.REAL 2};
  then consider p such that
  A2: p=x & p`1<=0 & p<>0.REAL 2;
  A3:now assume |.p.|<=0;
    then |.p.|=0 by TOPRNS_1:26;
   hence contradiction by A2,TOPRNS_1:25;
  end;
  then A4: (|.p.|)^2>0 by SQUARE_1:74;
    now per cases;
  case A5: p`2/|.p.|<=sn;
   then A6: (sn-FanMorphW).p= |[ |.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)),
           |.p.|*((p`2/|.p.|-sn)/(1+sn))]| by A1,A2,Th25;
   reconsider p9= (sn-FanMorphW).p as Point of TOP-REAL 2;
   A7: p9`1=|.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)) by A6,EUCLID:56;
   A8: 1+sn>0 by A1,Th3;
       A9: (|.p.|)^2 =(p`1)^2+(p`2)^2 by JGRAPH_3:10;
         now per cases;
       case p`1=0;
        hence (sn-FanMorphW).x in K0 by A1,A2,Th23;
       case p`1<>0;
         then 0<(p`1)^2 by SQUARE_1:74;
         then 0+(p`2)^2<(p`1)^2+(p`2)^2 by REAL_1:67;
         then (p`2)^2/(|.p.|)^2 < (|.p.|)^2/(|.p.|)^2 by A4,A9,REAL_1:73;
         then (p`2)^2/(|.p.|)^2 < 1 by A4,XCMPLX_1:60;
         then A10: ((p`2)/|.p.|)^2 < 1 by SQUARE_1:69;
           p`2/|.p.|-sn<=0 by A5,REAL_2:106;
         then A11: (p`2/|.p.|-sn)/(1+sn)<(1+sn)/(1+sn) by A8,REAL_1:73;
           -1 < p`2/|.p.| by A10,Th5;
         then -1-sn< p`2/|.p.|-sn by REAL_1:54;
         then -(1+sn)<p`2/|.p.|-sn by XCMPLX_1:161;
         then (-1)*(1+sn)< p`2/|.p.|-sn by XCMPLX_1:180;
         then (-1)*(1+sn)/(1+sn)< (p`2/|.p.|-sn)/(1+sn)
                       by A8,REAL_1:73;
         then -1< (p`2/|.p.|-sn)/(1+sn) & (p`2/|.p.|-sn)/(1+sn)<1
                           by A8,A11,XCMPLX_1:60,90;
         then 1^2> ((p`2/|.p.|-sn)/(1+sn))^2 by JGRAPH_2:8;
         then 1-((p`2/|.p.|-sn)/(1+sn))^2>0 by SQUARE_1:11,59;
         then --sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)>0 by SQUARE_1:93;
         then -sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)<0 by REAL_1:66;
         then |.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2))<0
                           by A3,SQUARE_1:24;
        hence (sn-FanMorphW).x in K0 by A1,A2,A7,JGRAPH_2:11;
        end;
       hence (sn-FanMorphW).x in K0;
  case A12: p`2/|.p.|>sn;
   then A13: (sn-FanMorphW).p= |[ |.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)),
           |.p.|*((p`2/|.p.|-sn)/(1-sn))]| by A1,A2,Th25;
   reconsider p9= (sn-FanMorphW).p as Point of TOP-REAL 2;
   A14: p9`1=|.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)) by A13,EUCLID:56;
    A15: 1-sn>0 by A1,Th3;
       A16: (|.p.|)^2 =(p`1)^2+(p`2)^2 by JGRAPH_3:10;
         now per cases;
       case p`1=0;
        hence (sn-FanMorphW).x in K0 by A1,A2,Th23;
       case p`1<>0;
         then 0<(p`1)^2 by SQUARE_1:74;
         then 0+(p`2)^2<(p`1)^2+(p`2)^2 by REAL_1:67;
         then (p`2)^2/(|.p.|)^2 < (|.p.|)^2/(|.p.|)^2 by A4,A16,REAL_1:73;
         then (p`2)^2/(|.p.|)^2 < 1 by A4,XCMPLX_1:60;
         then ((p`2)/|.p.|)^2 < 1 by SQUARE_1:69;
         then p`2/|.p.|<1 by Th5;
         then (p`2/|.p.|-sn)<1-sn by REAL_1:54;
         then A17: (p`2/|.p.|-sn)/(1-sn)<(1-sn)/(1-sn) by A15,REAL_1:73;
           -(1-sn)< -0 by A15,REAL_1:50;
         then A18: -(1-sn)<sn-sn by XCMPLX_1:14;
           p`2/|.p.|-sn>=sn-sn by A12,REAL_1:49;
         then -(1-sn)<p`2/|.p.|-sn by A18,AXIOMS:22;
         then (-1)*(1-sn)< p`2/|.p.|-sn by XCMPLX_1:180;
         then (-1)*(1-sn)/(1-sn)< (p`2/|.p.|-sn)/(1-sn)
                       by A15,REAL_1:73;
         then -1< (p`2/|.p.|-sn)/(1-sn) & (p`2/|.p.|-sn)/(1-sn)<1
                           by A15,A17,XCMPLX_1:60,90;
         then 1> ((p`2/|.p.|-sn)/(1-sn))^2 by JGRAPH_2:8,SQUARE_1:59;
         then 1-((p`2/|.p.|-sn)/(1-sn))^2>0 by SQUARE_1:11;
         then --sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)>0 by SQUARE_1:93;
         then -sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)<0 by REAL_1:66;
         then p9`1<0 by A3,A14,SQUARE_1:24;
        hence (sn-FanMorphW).x in K0 by A1,A2,JGRAPH_2:11;
       end;
        hence (sn-FanMorphW).x in K0;
        end;
 hence (sn-FanMorphW).x in K0;
end;

theorem Th42: for sn being Real,x,K0 being set
 st -1<sn & sn<1 & x in K0 & K0={p: p`1>=0 & p<>0.REAL 2}
holds (sn-FanMorphW).x in K0
proof let sn be Real,x,K0 be set;
 assume A1:  -1<sn & sn<1 & x in K0 &
 K0={p: p`1>=0 & p<>0.REAL 2};
 then consider p such that
 A2: p=x & p`1>=0 & p<>0.REAL 2;
 thus (sn-FanMorphW).x in K0 by A1,A2,Th23;
end;

InclSub { D() -> non empty Subset of TOP-REAL 2, P[set] } :
  { p : P[p] & p<>0.REAL 2 } c= the carrier of (TOP-REAL 2)|D()
  provided
A1:  D() = (the carrier of TOP-REAL 2) \ {0.REAL 2}
proof
  let x be set;
A2:D()` = {0.REAL 2} by A1,JGRAPH_3:30;
  assume x in {p: P[p] & p<>0.REAL 2};
  then consider p such that A3: x=p & P[p] & p<>0.REAL 2;
    now assume not x in D();
    then x in (the carrier of TOP-REAL 2) \ D() by A3,XBOOLE_0:def 4;
    then x in D()` by SUBSET_1:def 5;
    hence contradiction by A2,A3,TARSKI:def 1;
  end;
  hence thesis by JORDAN1:1;
end;

theorem Th43:for sn being Real,
 D being non empty Subset of TOP-REAL 2
st -1<sn & sn<1 & D`={0.REAL 2} holds
ex h being map of (TOP-REAL 2)|D,(TOP-REAL 2)|D
st h=(sn-FanMorphW)|D & h is continuous
proof let sn be Real,D be non empty Subset of TOP-REAL 2;
assume A1:  -1<sn & sn<1 & D`={0.REAL 2};
  reconsider B0= {0.REAL 2} as Subset of TOP-REAL 2;
  A2: D =(B0)` by A1
     .=((the carrier of TOP-REAL 2)\ {0.REAL 2}) by SUBSET_1:def 5;
     defpred P[Point of TOP-REAL 2] means $1`1<=0;
    A3:{p: P[p] & p<>0.REAL 2} c= the carrier of (TOP-REAL 2)|D
      from InclSub(A2);
        (|[0,1]|)`1=0 & (|[0,1]|)`2=1 by EUCLID:56;
      then |[0,1]| in {p where p is Point of TOP-REAL 2:
        p`1<=0 & p<>0.REAL 2} by JGRAPH_2:11;
      then reconsider K0={p:p`1<=0 & p<>0.REAL 2}
          as non empty Subset of (TOP-REAL 2)|D by A3;
          defpred P[Point of TOP-REAL 2] means $1`1>=0;
    A4:{p: P[p] & p<>0.REAL 2} c= the carrier of (TOP-REAL 2)|D
       from InclSub(A2);
     set Y1=|[0,1]|;
      Y1`1=0 & Y1`2=1 by EUCLID:56;
     then Y1 in {p where p is Point of TOP-REAL 2:
     p`1>=0 & p<>0.REAL 2} by JGRAPH_2:11;
  then reconsider K1={p: p`1>=0 & p<>0.REAL 2}
       as non empty Subset of (TOP-REAL 2)|D by A4;
   A5:the carrier of ((TOP-REAL 2)|D) =D by JORDAN1:1;
   A6:K0 c= (the carrier of TOP-REAL 2)
   proof let z be set;assume z in K0;
        then consider p8 being Point of TOP-REAL 2 such that
        A7: p8=z & p8`1<=0 & p8<>0.REAL 2;
    thus z in (the carrier of TOP-REAL 2) by A7;
   end;
  A8:dom ((sn-FanMorphW)|K0)= dom ((sn-FanMorphW)) /\ K0 by FUNCT_1:68
       .=((the carrier of TOP-REAL 2)) /\ K0 by FUNCT_2:def 1
       .=K0 by A6,XBOOLE_1:28;
  A9: K0 =the carrier of ((TOP-REAL 2)|D)|K0 by JORDAN1:1;
    rng ((sn-FanMorphW)|K0) c= the carrier of ((TOP-REAL 2)|D)|K0
   proof let y be set;assume y in rng ((sn-FanMorphW)|K0);
     then consider x being set such that
     A10:x in dom ((sn-FanMorphW)|K0)
     & y=((sn-FanMorphW)|K0).x by FUNCT_1:def 5;
       x in (dom (sn-FanMorphW)) /\ K0 by A10,FUNCT_1:68;
     then A11:x in K0 by XBOOLE_0:def 3;
       K0 c= the carrier of TOP-REAL 2 by A5,XBOOLE_1:1;
     then reconsider p=x as Point of TOP-REAL 2 by A11;
      (sn-FanMorphW).p=y by A10,A11,FUNCT_1:72;
     then y in K0 by A1,A11,Th41;
    hence y in the carrier of ((TOP-REAL 2)|D)|K0 by JORDAN1:1;
   end;
  then rng ((sn-FanMorphW)|K0)c= the carrier of ((TOP-REAL 2)|D)
                            by A9,XBOOLE_1:1;
  then (sn-FanMorphW)|K0 is Function of the carrier of ((TOP-REAL 2)|D)|K0,
  the carrier of ((TOP-REAL 2)|D) by A8,A9,FUNCT_2:4;
  then reconsider f=(sn-FanMorphW)|K0 as map of ((TOP-REAL 2)|D)|K0,
  ((TOP-REAL 2)|D) ;
   A12:K1 c= (the carrier of TOP-REAL 2)
   proof let z be set;assume z in K1;
        then consider p8 being Point of TOP-REAL 2 such that
        A13: p8=z & p8`1>=0 & p8<>0.REAL 2;
    thus z in (the carrier of TOP-REAL 2) by A13;
   end;
  A14:dom ((sn-FanMorphW)|K1)= dom ((sn-FanMorphW)) /\ K1 by FUNCT_1:68
       .=((the carrier of TOP-REAL 2)) /\ K1 by FUNCT_2:def 1
       .=K1 by A12,XBOOLE_1:28;
  A15: K1=the carrier of ((TOP-REAL 2)|D)|K1 by JORDAN1:1;
    rng ((sn-FanMorphW)|K1) c= the carrier of ((TOP-REAL 2)|D)|K1
   proof let y be set;assume y in rng ((sn-FanMorphW)|K1);
     then consider x being set such that
     A16:x in dom ((sn-FanMorphW)|K1)
     & y=((sn-FanMorphW)|K1).x by FUNCT_1:def 5;
       x in (dom (sn-FanMorphW)) /\ K1 by A16,FUNCT_1:68;
     then A17:x in K1 by XBOOLE_0:def 3;
       K1 c= the carrier of TOP-REAL 2 by A5,XBOOLE_1:1;
     then reconsider p=x as Point of TOP-REAL 2 by A17;
      (sn-FanMorphW).p=y by A16,A17,FUNCT_1:72;
     then y in K1 by A1,A17,Th42;
    hence y in the carrier of ((TOP-REAL 2)|D)|K1 by JORDAN1:1;
   end;
  then rng ((sn-FanMorphW)|K1)c= the carrier of ((TOP-REAL 2)|D)
                     by A15,XBOOLE_1:1;
  then (sn-FanMorphW)|K1 is Function of the carrier of ((TOP-REAL 2)|D)|K1,
  the carrier of ((TOP-REAL 2)|D) by A14,A15,FUNCT_2:4;
  then reconsider g=(sn-FanMorphW)|K1 as map of ((TOP-REAL 2)|D)|K1,
  ((TOP-REAL 2)|D) ;
  A18:K0=[#](((TOP-REAL 2)|D)|K0) by PRE_TOPC:def 10;
  A19:K1=[#](((TOP-REAL 2)|D)|K1) by PRE_TOPC:def 10;
  A20:D= [#]((TOP-REAL 2)|D) by PRE_TOPC:def 10;
  A21:K0 \/ K1 c= D
   proof let x be set;assume A22:x in K0 \/ K1;
       now per cases by A22,XBOOLE_0:def 2;
     case x in K0;
       then consider p such that A23:p=x &
       p`1<=0 & p<>0.REAL 2;
     thus
       x in the carrier of TOP-REAL 2 & not x=0.REAL 2 by A23;
     case x in K1;
       then consider p such that A24:p=x &
       p`1>=0 & p<>0.REAL 2;
     thus x in the carrier of TOP-REAL 2 & not x=0.REAL 2 by A24;
     end;
     then x in the carrier of TOP-REAL 2 & not x in {0.REAL 2} by TARSKI:def 1
;
    hence x in D by A2,XBOOLE_0:def 4;
   end;
    D c= K0 \/ K1
   proof let x be set;assume A25: x in D;
     then A26:x in (the carrier of TOP-REAL 2) & not x in {0.REAL 2}
                                     by A2,XBOOLE_0:def 4;
     reconsider px=x as Point of TOP-REAL 2 by A25;
       px`1<=0 & px<>0.REAL 2 or
       px`1>=0 & px<>0.REAL 2 by A26,TARSKI:def 1;
     then x in K0 or x in K1;
    hence x in K0 \/ K1 by XBOOLE_0:def 2;
   end;
  then A27:([#](((TOP-REAL 2)|D)|K0)) \/ ([#](((TOP-REAL 2)|D)|K1))
  = [#]((TOP-REAL 2)|D) by A18,A19,A20,A21,XBOOLE_0:def 10;
  A28: f is continuous & K0 is closed by A1,A2,Th36,Th37;
  A29: g is continuous & K1 is closed by A1,A2,Th38,Th39;
  A30: for x be set st x in ([#]((((TOP-REAL 2)|D)|K0)))
        /\ ([#] ((((TOP-REAL 2)|D)|K1))) holds f.x = g.x
   proof let x be set;assume x in ([#]((((TOP-REAL 2)|D)|K0)))
        /\ ([#] ((((TOP-REAL 2)|D)|K1)));
     then A31:x in K0 & x in K1 by A18,A19,XBOOLE_0:def 3;
     then f.x=(sn-FanMorphW).x by FUNCT_1:72;
    hence f.x = g.x by A31,FUNCT_1:72;
   end;
  then consider h being map of (TOP-REAL 2)|D,(TOP-REAL 2)|D
  such that A32: h= f+*g & h is continuous
                          by A18,A19,A27,A28,A29,JGRAPH_2:9;
  A33:dom h=the carrier of ((TOP-REAL 2)|D) by FUNCT_2:def 1;
  A34:the carrier of ((TOP-REAL 2)|D)=D by JORDAN1:1;
  A35:the carrier of ((TOP-REAL 2)|D) =
((the carrier of TOP-REAL 2)\ {0.REAL 2}) by A2,JORDAN1:1;
    dom (sn-FanMorphW)=(the carrier of (TOP-REAL 2))
                    by FUNCT_2:def 1;
  then A36:dom ((sn-FanMorphW)|D)=(the carrier of (TOP-REAL 2))/\ D
                    by FUNCT_1:68
              .=the carrier of ((TOP-REAL 2)|D) by A34,XBOOLE_1:28;
  A37:dom f=K0 by A9,FUNCT_2:def 1;
  A38:K0=[#](((TOP-REAL 2)|D)|K0) by PRE_TOPC:def 10;
  A39:dom g=K1 by A15,FUNCT_2:def 1;
   K1=[#](((TOP-REAL 2)|D)|K1) by PRE_TOPC:def 10;
  then A40: f tolerates g by A30,A37,A38,A39,PARTFUN1:def 6;
    for x being set st x in dom h holds h.x=((sn-FanMorphW)|D).x
   proof let x be set;assume A41: x in dom h;
     then x in (the carrier of TOP-REAL 2) & not x in {0.REAL 2}
                     by A33,A35,XBOOLE_0:def 4;
     then A42:x <>0.REAL 2 by TARSKI:def 1;
     reconsider p=x as Point of TOP-REAL 2 by A33,A35,A41,XBOOLE_0:def 4;
     A43: x in D`` by A33,A41,JORDAN1:1;
       now per cases;
     case A44:x in K0;
       A45:(sn-FanMorphW)|D.p=(sn-FanMorphW).p by A43,FUNCT_1:72
          .=f.p by A44,FUNCT_1:72;
         h.p=(g+*f).p by A32,A40,FUNCT_4:35
       .=f.p by A37,A44,FUNCT_4:14;
      hence h.x=(sn-FanMorphW)|D.x by A45;
     case not x in K0;
       then not (p`1<=0) by A42;
       then A46:x in K1 by A42;
        (sn-FanMorphW)|D.p=(sn-FanMorphW).p by A43,FUNCT_1:72
                           .=g.p by A46,FUNCT_1:72;
      hence h.x=(sn-FanMorphW)|D.x by A32,A39,A46,FUNCT_4:14;
     end;
    hence h.x=(sn-FanMorphW)|D.x;
   end;
  then f+*g=(sn-FanMorphW)|D by A32,A33,A36,FUNCT_1:9;
 hence thesis by A18,A19,A27,A28,A29,A30,JGRAPH_2:9;
end;

theorem Th44: for sn being Real
  st -1<sn & sn<1 holds
  ex h being map of (TOP-REAL 2),(TOP-REAL 2)
  st h=(sn-FanMorphW) & h is continuous
proof let sn be Real;
 assume A1: -1<sn & sn<1;
 reconsider f=(sn-FanMorphW) as map of (TOP-REAL 2),(TOP-REAL 2)
                            ;
  (the carrier of TOP-REAL 2) \{0.REAL 2} c=(the carrier of TOP-REAL 2)
                         by XBOOLE_1:36;
 then reconsider D=(the carrier of TOP-REAL 2) \{0.REAL 2}
   as non empty Subset of TOP-REAL 2 by JGRAPH_2:19;
 A2: f.(0.REAL 2)=0.REAL 2 by Th23,JGRAPH_2:11;
 A3: D`= {0.REAL 2} by JGRAPH_3:30;
 A4: for p being Point of (TOP-REAL 2)|D holds f.p<>f.(0.REAL 2)
  proof let p be Point of (TOP-REAL 2)|D;
    A5:the carrier of (TOP-REAL 2)|D=[#]((TOP-REAL 2)|D) by PRE_TOPC:12;
    A6:[#]((TOP-REAL 2)|D)=D by PRE_TOPC:def 10;
    then A7:p in the carrier of TOP-REAL 2 & not p in {0.REAL 2}
                         by A5,XBOOLE_0:def 4;
    reconsider q=p as Point of TOP-REAL 2 by A5,A6,XBOOLE_0:def 4;
    A8:not p=0.REAL 2 by A7,TARSKI:def 1;
      now per cases;
    case A9: q`2/|.q.|>=sn & q`1<=0;
      set q9= |[ |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1-sn))]|;
      A10:q9`1= |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)) &
      q9`2= |.q.|* ((q`2/|.q.|-sn)/(1-sn)) by EUCLID:56;
        now assume A11: q9=0.REAL 2;
       A12: |.q.|<>0 by A8,TOPRNS_1:25;
       then -sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)=-sqrt(1-0) by A10,A11,JGRAPH_2:
11,SQUARE_1:60,XCMPLX_1:6
       .=-1 by SQUARE_1:83;
      hence contradiction by A10,A11,A12,JGRAPH_2:11,XCMPLX_1:6;
     end;
     hence f.p<>f.(0.REAL 2) by A1,A2,A8,A9,Th25;
    case A13: q`2/|.q.|<sn & q`1<=0;
      set q9=|[ |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)),
              |.q.|* ((q`2/|.q.|-sn)/(1+sn))]|;
      A14:q9`1= |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)) &
          q9`2= |.q.|* ((q`2/|.q.|-sn)/(1+sn)) by EUCLID:56;
        now assume A15: q9=0.REAL 2;
       A16: |.q.|<>0 by A8,TOPRNS_1:25;
       then -sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)=-sqrt(1-0) by A14,A15,JGRAPH_2:
11,SQUARE_1:60,XCMPLX_1:6
       .=-1 by SQUARE_1:83;
      hence contradiction by A14,A15,A16,JGRAPH_2:11,XCMPLX_1:6;
     end;
     hence f.p<>f.(0.REAL 2) by A1,A2,A8,A13,Th25;
     case q`1>0; then f.p=p by Th23;
     hence f.p<>f.(0.REAL 2) by A8,Th23,JGRAPH_2:11;
    end;
   hence f.p<>f.(0.REAL 2);
  end;
 A17:ex h being map of (TOP-REAL 2)|D,(TOP-REAL 2)|D
  st h=(sn-FanMorphW)|D & h is continuous by A1,A3,Th43;
   for V being Subset of (TOP-REAL 2) st f.(0.REAL 2) in V & V is open
 ex W being Subset of (TOP-REAL 2) st 0.REAL 2 in W
 & W is open & f.:W c= V
  proof let V be Subset of (TOP-REAL 2);
    assume A18:f.(0.REAL 2) in V & V is open;
     A19:TOP-REAL 2=TopSpaceMetr(Euclid 2) by EUCLID:def 8;
     A20:the carrier of TOP-REAL 2=REAL 2 by EUCLID:25;
     A21:Euclid 2=MetrStruct(#REAL 2,Pitag_dist 2#) by EUCLID:def 7;
     then reconsider u0=0.REAL 2 as Point of Euclid 2 by EUCLID:25;
     consider r being real number such that A22:r>0 & Ball(u0,r) c= V
                                   by A2,A18,A19,TOPMETR:22;
     reconsider r as Real by XREAL_0:def 1;
     reconsider W1=Ball(u0,r) as Subset of TOP-REAL 2
                       by A20,A21;
     A23:u0 in W1 by A22,GOBOARD6:4;
     A24:W1 is open by GOBOARD6:6;
       f.:W1 c= W1
      proof let z be set;assume z in f.:W1;
        then consider y being set such that
        A25: y in dom f & y in W1 & z=f.y by FUNCT_1:def 12;
        reconsider q=y as Point of TOP-REAL 2 by A25;
        reconsider qy=q as Point of Euclid 2 by A21,EUCLID:25;
          z in rng f by A25,FUNCT_1:def 5;
        then reconsider qz=z as Point of TOP-REAL 2;
        reconsider pz=qz as Point of Euclid 2 by A21,EUCLID:25;
          dist(u0,qy)<r by A25,METRIC_1:12;
        then A26: |.(0.REAL 2) - q.|<r by JGRAPH_1:45;
        A27: |.q.|>=0 by TOPRNS_1:26;
          now per cases by JGRAPH_2:11;
        case q`1>=0;
         hence z in W1 by A25,Th23;
        case A28: q<>0.REAL 2 &
          (q`2/|.q.|>=sn & q`1<=0);
          then A29:(sn-FanMorphW).q= |[
          |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1-sn))]| by A1,Th25;
            |.q.|<>0 by A28,TOPRNS_1:25;
          then A30: (|.q.|)^2>0 by SQUARE_1:74;
        A31: 1-sn>0 by A1,Th3;
        A32: qz`1= |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2))&
        qz`2= |.q.|* ((q`2/|.q.|-sn)/(1-sn))
                     by A25,A29,EUCLID:56;
        A33: (q`2/|.q.|-sn)>= 0 by A28,SQUARE_1:12;
        A34: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
          0<=(q`1)^2 by SQUARE_1:72;
        then 0+(q`2)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
        then (q`2)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A30,A34,REAL_1:73;
        then (q`2)^2/(|.q.|)^2 <= 1 by A30,XCMPLX_1:60;
        then ((q`2)/|.q.|)^2 <= 1 by SQUARE_1:69;
        then 1>=q`2/|.q.| by Th4;
        then A35: 1-sn>=q`2/|.q.|-sn by REAL_1:49;
        A36: -(1-sn)<= -0 by A31,REAL_1:50;
           -(1-sn)<= -( q`2/|.q.|-sn) by A35,REAL_1:50;
         then (-(1-sn))/(1-sn)<=(-( q`2/|.q.|-sn))/(1-sn)
                        by A31,REAL_1:73;
        then A37: -1<=(-( q`2/|.q.|-sn))/(1-sn)
                        by A31,XCMPLX_1:198;
         --(1-sn)>= -(q`2/|.q.|-sn) by A33,A36,REAL_1:50;
       then (-(q`2/|.q.|-sn))/(1-sn)<=1 by A31,REAL_2:117;
       then ((-(q`2/|.q.|-sn))/(1-sn))^2<=1^2 by A37,JGRAPH_2:7;
       then 1-((-(q`2/|.q.|-sn))/(1-sn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((q`2/|.q.|-sn))/(1-sn))^2>=0 by XCMPLX_1:188;
       then A38: 1-((q`2/|.q.|-sn)/(1-sn))^2>=0 by SQUARE_1:61;
        A39: (qz`1)^2= (|.q.|)^2*(-sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2))^2
                         by A32,SQUARE_1:68
         .= (|.q.|)^2*(sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2))^2
                     by SQUARE_1:61
         .= (|.q.|)^2*(1-((q`2/|.q.|-sn)/(1-sn))^2) by A38,SQUARE_1:def 4;
        A40: (qz`2)^2= (|.q.|)^2*((q`2/|.q.|-sn)/(1-sn))^2
                         by A32,SQUARE_1:68;
          (|.qz.|)^2=(qz`1)^2+(qz`2)^2 by JGRAPH_3:10
        .=(|.q.|)^2*(1-((q`2/|.q.|-sn)/(1-sn))^2
           +((q`2/|.q.|-sn)/(1-sn))^2) by A39,A40,XCMPLX_1:8
        .=(|.q.|)^2*1 by XCMPLX_1:27 .=(|.q.|)^2;
        then A41: sqrt((|.qz.|)^2)=|.q.| by A27,SQUARE_1:89;
            |.qz.|>=0 by TOPRNS_1:26;
          then A42: |.qz.|=|.q.| by A41,SQUARE_1:89;
            |.(0.REAL 2) +- q.|<r by A26,EUCLID:45;
          then |.- q.|<r by EUCLID:31;
          then |.q.|<r by TOPRNS_1:27;
          then |.- qz.|<r by A42,TOPRNS_1:27;
          then |.(0.REAL 2) +- qz.|<r by EUCLID:31;
          then |.(0.REAL 2) - qz.|<r by EUCLID:45;
          then dist(u0,pz)<r by JGRAPH_1:45;
         hence z in W1 by METRIC_1:12;
        case A43: q<>0.REAL 2 &
          q`2/|.q.|<sn & q`1<=0;
          then A44:(sn-FanMorphW).q= |[
          |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1+sn))]| by A1,Th25;
            |.q.|<>0 by A43,TOPRNS_1:25;
          then A45: (|.q.|)^2>0 by SQUARE_1:74;
        A46: (sn-q`2/|.q.|)>=0 by A43,SQUARE_1:12;
        A47: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
        A48: 1+sn>0 by A1,Th3;
        A49: qz`1= |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2))&
        qz`2= |.q.|* ((q`2/|.q.|-sn)/(1+sn))
                     by A25,A44,EUCLID:56;
          0<=(q`1)^2 by SQUARE_1:72;
        then 0+(q`2)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
        then (q`2)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A45,A47,REAL_1:73;
        then (q`2)^2/(|.q.|)^2 <= 1 by A45,XCMPLX_1:60;
        then ((q`2)/|.q.|)^2 <= 1 by SQUARE_1:69;
        then -1<=q`2/|.q.| by Th4;
        then --1>=-q`2/|.q.| by REAL_1:50;
        then A50: 1+sn>=-q`2/|.q.|+sn by REAL_1:55;
          -(1+sn)<= -0 by A48,REAL_1:50;
         then -(1+sn)<= -( q`2/|.q.|-sn) by A46,XCMPLX_1:143;
         then (-(1+sn))/(1+sn)<=(-( q`2/|.q.|-sn))/(1+sn)
                        by A48,REAL_1:73;
        then A51: -1<=(-( q`2/|.q.|-sn))/(1+sn)
                        by A48,XCMPLX_1:198;
          -(1+sn)<=-(-q`2/|.q.|+sn) by A50,REAL_1:50;
        then -(1+sn)<=--(q`2/|.q.|)-sn by XCMPLX_1:161;
        then --(1+sn)>= -(q`2/|.q.|-sn) by REAL_1:50;
       then (-(q`2/|.q.|-sn))/(1+sn)<=1 by A48,REAL_2:117;
       then ((-(q`2/|.q.|-sn))/(1+sn))^2<=1^2 by A51,JGRAPH_2:7;
       then 1-((-(q`2/|.q.|-sn))/(1+sn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((q`2/|.q.|-sn))/(1+sn))^2>=0 by XCMPLX_1:188;
       then A52: 1-((q`2/|.q.|-sn)/(1+sn))^2>=0 by SQUARE_1:61;
        A53: (qz`1)^2= (|.q.|)^2*(-sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2))^2
                         by A49,SQUARE_1:68
         .= (|.q.|)^2*(sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2))^2
                     by SQUARE_1:61
         .= (|.q.|)^2*(1-((q`2/|.q.|-sn)/(1+sn))^2) by A52,SQUARE_1:def 4;
        A54: (qz`2)^2= (|.q.|)^2*((q`2/|.q.|-sn)/(1+sn))^2
                         by A49,SQUARE_1:68;
          (|.qz.|)^2=(qz`1)^2+(qz`2)^2 by JGRAPH_3:10
        .=(|.q.|)^2*(1-((q`2/|.q.|-sn)/(1+sn))^2
           +((q`2/|.q.|-sn)/(1+sn))^2) by A53,A54,XCMPLX_1:8
        .=(|.q.|)^2*1 by XCMPLX_1:27 .=(|.q.|)^2;
        then A55: sqrt((|.qz.|)^2)=|.q.| by A27,SQUARE_1:89;
            |.qz.|>=0 by TOPRNS_1:26;
          then A56: |.qz.|=|.q.| by A55,SQUARE_1:89;
            |.(0.REAL 2) +- q.|<r by A26,EUCLID:45;
          then |.- q.|<r by EUCLID:31;
          then |.q.|<r by TOPRNS_1:27;
          then |.- qz.|<r by A56,TOPRNS_1:27;
          then |.(0.REAL 2) +- qz.|<r by EUCLID:31;
          then |.(0.REAL 2) - qz.|<r by EUCLID:45;
          then dist(u0,pz)<r by JGRAPH_1:45;
         hence z in W1 by METRIC_1:12;
        end;
       hence z in W1;
      end;
     then f.:W1 c= V by A22,XBOOLE_1:1;
   hence ex W being Subset of (TOP-REAL 2) st 0.REAL 2 in W
       & W is open & f.:W c= V by A23,A24;
  end;
  then f is continuous by A2,A3,A4,A17,JGRAPH_3:13;
 hence thesis;
end;

theorem Th45: for sn being Real
  st -1<sn & sn<1 holds
  (sn-FanMorphW) is one-to-one
proof let sn be Real;
 assume A1:  -1<sn & sn<1;
    for x1,x2 being set st x1 in dom (sn-FanMorphW) & x2 in dom (sn-FanMorphW)
   & (sn-FanMorphW).x1=(sn-FanMorphW).x2 holds x1=x2
  proof let x1,x2 be set;
   assume A2: x1 in dom (sn-FanMorphW) & x2 in dom (sn-FanMorphW)
    & (sn-FanMorphW).x1=(sn-FanMorphW).x2;
    then reconsider p1=x1 as Point of TOP-REAL 2 by FUNCT_2:def 1;
    reconsider p2=x2 as Point of TOP-REAL 2 by A2,FUNCT_2:def 1;
    set q=p1,p=p2;
    A3: 1-sn>0 by A1,Th3;
      now per cases by JGRAPH_2:11;
    case A4: q`1>=0;
      then A5:(sn-FanMorphW).q=q by Th23;
        now per cases by JGRAPH_2:11;
      case p`1>=0;
       hence x1=x2 by A2,A5,Th23;
      case A6:p<>0.REAL 2 &
          (p`2/|.p.|>=sn & p`1<=0);
        then A7:
         sn-FanMorphW.p= |[ |.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)),
          |.p.|* ((p`2/|.p.|-sn)/(1-sn))]| by A1,Th25;
         set p4= |[ |.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)),
          |.p.|* ((p`2/|.p.|-sn)/(1-sn))]|;
        A8: p4`1= |.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)) &
          p4`2= |.p.|* ((p`2/|.p.|-sn)/(1-sn)) by EUCLID:56;
     A9: |.p.|>=0 by TOPRNS_1:26;
     A10: |.p.|>0 by A6,Lm1;
     then A11: (|.p.|)^2>0 by SQUARE_1:74;
        A12: (p`2/|.p.|-sn)>=0 by A6,SQUARE_1:12;
        A13: (p`2/|.p.|-sn)>= 0 by A6,SQUARE_1:12;
        A14: (p`2/|.p.|-sn)/(1-sn)>=0 by A3,A12,REAL_2:125;
        A15: (|.p.|)^2 =(p`1)^2+(p`2)^2 by JGRAPH_3:10;
          0<=(p`1)^2 by SQUARE_1:72;
        then 0+(p`2)^2<=(p`1)^2+(p`2)^2 by REAL_1:55;
        then (p`2)^2/(|.p.|)^2 <= (|.p.|)^2/(|.p.|)^2 by A11,A15,REAL_1:73;
        then (p`2)^2/(|.p.|)^2 <= 1 by A11,XCMPLX_1:60;
        then ((p`2)/|.p.|)^2 <= 1 by SQUARE_1:69;
        then 1>=p`2/|.p.| by Th4;
        then A16: 1-sn>=p`2/|.p.|-sn by REAL_1:49;
        A17: -(1-sn)<= -0 by A3,REAL_1:50;
           -(1-sn)<= -( p`2/|.p.|-sn) by A16,REAL_1:50;
         then (-(1-sn))/(1-sn)<=(-( p`2/|.p.|-sn))/(1-sn)
                        by A3,REAL_1:73;
        then A18: -1<=(-( p`2/|.p.|-sn))/(1-sn)
                        by A3,XCMPLX_1:198;
         --(1-sn)>= -(p`2/|.p.|-sn) by A13,A17,REAL_1:50;
       then (-(p`2/|.p.|-sn))/(1-sn)<=1 by A3,REAL_2:117;
       then ((-(p`2/|.p.|-sn))/(1-sn))^2<=1^2 by A18,JGRAPH_2:7;
       then A19: 1-((-(p`2/|.p.|-sn))/(1-sn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((p`2/|.p.|-sn))/(1-sn))^2>=0 by XCMPLX_1:188;
       then A20: 1-((p`2/|.p.|-sn)/(1-sn))^2>=0 by SQUARE_1:61;
         sqrt(1-((-(p`2/|.p.|-sn))/(1-sn))^2)>=0 by A19,SQUARE_1:def 4;
       then sqrt(1-(-(p`2/|.p.|-sn))^2/(1-sn)^2)>=0 by SQUARE_1:69;
       then sqrt(1-(p`2/|.p.|-sn)^2/(1-sn)^2)>=0 by SQUARE_1:61;
       then sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)>=0 by SQUARE_1:69;
       then -sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)<= -0 by REAL_1:50;
        then q`1=0 by A2,A4,A5,A7,A8,A9,REAL_2:123;
        then (-sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2))=0 by A2,A5,A7,A8,A10,XCMPLX_1
:6;
        then sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)=-0 .=0;
        then 1-((p`2/|.p.|-sn)/(1-sn))^2=0 by A20,SQUARE_1:92;
        then 1=0+((p`2/|.p.|-sn)/(1-sn))^2 by XCMPLX_1:27;
        then 1= (p`2/|.p.|-sn)/(1-sn) by A14,SQUARE_1:83,89;
        then (1)*(1-sn)=(p`2/|.p.|-sn) by A3,XCMPLX_1:88;
        then 1-sn+sn=p`2/|.p.| by XCMPLX_1:27;
        then 1=p`2/|.p.| by XCMPLX_1:27;
        then (1)*|.p.|=p`2 by A10,XCMPLX_1:88;
        then (p`2)^2-(p`2)^2 =(p`1)^2 by A15,XCMPLX_1:26;
        then 0= (p`1)^2 by XCMPLX_1:14;
        then p`1=0 by SQUARE_1:73;
       hence x1=x2 by A2,A5,Th23;
      case A21:p<>0.REAL 2 &
        (p`2/|.p.|<sn & p`1<=0);
        then A22:
         sn-FanMorphW.p= |[ |.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)),
          |.p.|* ((p`2/|.p.|-sn)/(1+sn))]| by A1,Th25;
         set p4= |[ |.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)),
          |.p.|* ((p`2/|.p.|-sn)/(1+sn))]|;
        A23: p4`1= |.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)) &
          p4`2= |.p.|* ((p`2/|.p.|-sn)/(1+sn)) by EUCLID:56;
     A24: |.p.|>=0 by TOPRNS_1:26;
     A25: |.p.|>0 by A21,Lm1;
     then A26: (|.p.|)^2>0 by SQUARE_1:74;
     A27: 1+sn>0 by A1,Th3;
        A28: (p`2/|.p.|-sn)<=0 by A21,REAL_2:106;
        then A29: -(p`2/|.p.|-sn)>= -0 by REAL_1:50;
          --(p`2/|.p.|-sn)/(1+sn)<=0 by A27,A28,REAL_2:126;
        then A30: -((p`2/|.p.|-sn)/(1+sn))>=0 by REAL_1:66;
        A31: (|.p.|)^2 =(p`1)^2+(p`2)^2 by JGRAPH_3:10;
          0<=(p`1)^2 by SQUARE_1:72;
        then 0+(p`2)^2<=(p`1)^2+(p`2)^2 by REAL_1:55;
        then (p`2)^2/(|.p.|)^2 <= (|.p.|)^2/(|.p.|)^2 by A26,A31,REAL_1:73;
        then (p`2)^2/(|.p.|)^2 <= 1 by A26,XCMPLX_1:60;
        then ((p`2)/|.p.|)^2 <= 1 by SQUARE_1:69;
        then A32: (-((p`2)/|.p.|))^2 <= 1 by SQUARE_1:61;
          -(1+sn)<= -0 by A27,REAL_1:50;
        then (-(1+sn))/(1+sn)<=(-( p`2/|.p.|-sn))/(1+sn)
                        by A27,A29,REAL_1:73;
        then A33: -1<=(-( p`2/|.p.|-sn))/(1+sn) by A27,XCMPLX_1:198;
         1>= -p`2/|.p.| by A32,Th4;
       then (1+sn)>= -p`2/|.p.|+sn by REAL_1:55;
       then (1+sn)>= -(p`2/|.p.|-sn) by XCMPLX_1:162;
       then (-(p`2/|.p.|-sn))/(1+sn)<=1 by A27,REAL_2:117;
       then ((-(p`2/|.p.|-sn))/(1+sn))^2<=1^2 by A33,JGRAPH_2:7;
       then A34: 1-((-(p`2/|.p.|-sn))/(1+sn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((p`2/|.p.|-sn))/(1+sn))^2>=0 by XCMPLX_1:188;
       then A35: 1-((p`2/|.p.|-sn)/(1+sn))^2>=0 by SQUARE_1:61;
         sqrt(1-((-(p`2/|.p.|-sn))/(1+sn))^2)>=0 by A34,SQUARE_1:def 4;
       then sqrt(1-(-(p`2/|.p.|-sn))^2/(1+sn)^2)>=0 by SQUARE_1:69;
       then sqrt(1-((p`2/|.p.|-sn))^2/(1+sn)^2)>=0 by SQUARE_1:61;
       then sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)>=0 by SQUARE_1:69;
       then -sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)<= -0 by REAL_1:50;
        then q`1=0 by A2,A4,A5,A22,A23,A24,REAL_2:123;
        then (-sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2))=0 by A2,A5,A22,A23,A25,
XCMPLX_1:6;
        then sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)=-0 .=0;
        then 1-((p`2/|.p.|-sn)/(1+sn))^2=0 by A35,SQUARE_1:92;
        then 1=0+((p`2/|.p.|-sn)/(1+sn))^2 by XCMPLX_1:27;
        then 1=sqrt((-((p`2/|.p.|-sn)/(1+sn)))^2) by SQUARE_1:61,83;
        then 1= -((p`2/|.p.|-sn)/(1+sn)) by A30,SQUARE_1:89;
        then 1= ((-(p`2/|.p.|-sn))/(1+sn)) by XCMPLX_1:188;
        then (1)*(1+sn)=-(p`2/|.p.|-sn) by A27,XCMPLX_1:88;
        then (1+sn)=-p`2/|.p.|+sn by XCMPLX_1:162;
        then 1+sn-sn=-p`2/|.p.| by XCMPLX_1:26;
        then 1=-p`2/|.p.| by XCMPLX_1:26;
        then 1=(-p`2)/|.p.| by XCMPLX_1:188;
        then (1)*|.p.|=-p`2 by A25,XCMPLX_1:88;
        then (-p`2)^2-(p`2)^2 =(p`1)^2 by A31,XCMPLX_1:26;
        then (p`2)^2-(p`2)^2 =(p`1)^2 by SQUARE_1:61;
        then 0= (p`1)^2 by XCMPLX_1:14;
        then p`1=0 by SQUARE_1:73;
       hence x1=x2 by A2,A5,Th23;
      end;
     hence x1=x2;
    case A36:
      q`2/|.q.|>=sn & q`1<=0 & q<>0.REAL 2;
        then A37:
         sn-FanMorphW.q= |[ |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1-sn))]| by A1,Th25;
         set q4= |[ |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1-sn))]|;
        A38: q4`1= |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)) &
          q4`2= |.q.|* ((q`2/|.q.|-sn)/(1-sn)) by EUCLID:56;
     A39: |.q.|>0 by A36,Lm1;
     then A40: (|.q.|)^2>0 by SQUARE_1:74;
        now per cases by JGRAPH_2:11;
      case A41: p`1>=0;
        then A42: (sn-FanMorphW).p=p by Th23;
        A43: |.q.|>=0 by TOPRNS_1:26;
        A44: |.q.|>0 by A36,Lm1;
        then A45: (|.q.|)^2>0 by SQUARE_1:74;
        A46: 1-sn>0 by A1,Th3;
        A47: (q`2/|.q.|-sn)>=0 by A36,SQUARE_1:12;
        A48: (q`2/|.q.|-sn)>= 0 by A36,SQUARE_1:12;
        A49: (q`2/|.q.|-sn)/(1-sn)>=0 by A46,A47,REAL_2:125;
        A50: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
          0<=(q`1)^2 by SQUARE_1:72;
        then 0+(q`2)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
        then (q`2)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A45,A50,REAL_1:73;
        then (q`2)^2/(|.q.|)^2 <= 1 by A45,XCMPLX_1:60;
        then ((q`2)/|.q.|)^2 <= 1 by SQUARE_1:69;
        then 1>=q`2/|.q.| by Th4;
        then A51: 1-sn>=q`2/|.q.|-sn by REAL_1:49;
        A52: -(1-sn)<= -0 by A46,REAL_1:50;
           -(1-sn)<= -( q`2/|.q.|-sn) by A51,REAL_1:50;
         then (-(1-sn))/(1-sn)<=(-( q`2/|.q.|-sn))/(1-sn)
                        by A46,REAL_1:73;
        then A53: -1<=(-( q`2/|.q.|-sn))/(1-sn)
                        by A46,XCMPLX_1:198;
         --(1-sn)>= -(q`2/|.q.|-sn) by A48,A52,REAL_1:50;
       then (-(q`2/|.q.|-sn))/(1-sn)<=1 by A46,REAL_2:117;
       then ((-(q`2/|.q.|-sn))/(1-sn))^2<=1^2 by A53,JGRAPH_2:7;
       then A54: 1-((-(q`2/|.q.|-sn))/(1-sn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((q`2/|.q.|-sn))/(1-sn))^2>=0 by XCMPLX_1:188;
       then A55: 1-((q`2/|.q.|-sn)/(1-sn))^2>=0 by SQUARE_1:61;
         sqrt(1-((-(q`2/|.q.|-sn))/(1-sn))^2)>=0 by A54,SQUARE_1:def 4;
       then sqrt(1-(-(q`2/|.q.|-sn))^2/(1-sn)^2)>=0 by SQUARE_1:69;
       then sqrt(1-(q`2/|.q.|-sn)^2/(1-sn)^2)>=0 by SQUARE_1:61;
       then sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)>=0 by SQUARE_1:69;
       then -sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)<= -0 by REAL_1:50;
        then p`1=0 by A2,A37,A38,A41,A42,A43,REAL_2:123;
        then (-sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2))=0 by A2,A37,A38,A42,A44,
XCMPLX_1:6;
        then sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)=-0 .=0;
        then 1-((q`2/|.q.|-sn)/(1-sn))^2=0 by A55,SQUARE_1:92;
        then 1=0+((q`2/|.q.|-sn)/(1-sn))^2 by XCMPLX_1:27;
        then 1= (q`2/|.q.|-sn)/(1-sn) by A49,SQUARE_1:83,89;
        then (1)*(1-sn)=(q`2/|.q.|-sn) by A46,XCMPLX_1:88;
        then 1-sn+sn=q`2/|.q.| by XCMPLX_1:27;
        then 1=q`2/|.q.| by XCMPLX_1:27;
        then (1)*|.q.|=q`2 by A44,XCMPLX_1:88;
        then (q`2)^2-(q`2)^2 =(q`1)^2 by A50,XCMPLX_1:26;
        then 0= (q`1)^2 by XCMPLX_1:14;
        then q`1=0 by SQUARE_1:73;
       hence x1=x2 by A2,A42,Th23;
      case A56:p<>0.REAL 2 &
          (p`2/|.p.|>=sn & p`1<=0);
        then A57:
         sn-FanMorphW.p= |[ |.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)),
          |.p.|* ((p`2/|.p.|-sn)/(1-sn))]| by A1,Th25;
         set p4= |[ |.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)),
          |.p.|* ((p`2/|.p.|-sn)/(1-sn))]|;
        A58: p4`1= |.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)) &
          p4`2= |.p.|* ((p`2/|.p.|-sn)/(1-sn)) by EUCLID:56;
        A59: |.q4.|>=0 by TOPRNS_1:26;
        A60: q4`1= |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2))&
        q4`2= |.q.|* ((q`2/|.q.|-sn)/(1-sn))
                     by EUCLID:56;
        A61: (q`2/|.q.|-sn)>= 0 by A36,SQUARE_1:12;
        A62: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
          0<=(q`1)^2 by SQUARE_1:72;
        then 0+(q`2)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
        then (q`2)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A40,A62,REAL_1:73;
        then (q`2)^2/(|.q.|)^2 <= 1 by A40,XCMPLX_1:60;
        then ((q`2)/|.q.|)^2 <= 1 by SQUARE_1:69;
        then 1>=q`2/|.q.| by Th4;
        then A63: 1-sn>=q`2/|.q.|-sn by REAL_1:49;
        A64: -(1-sn)<= -0 by A3,REAL_1:50;
           -(1-sn)<= -( q`2/|.q.|-sn) by A63,REAL_1:50;
         then (-(1-sn))/(1-sn)<=(-( q`2/|.q.|-sn))/(1-sn)
                        by A3,REAL_1:73;
        then A65: -1<=(-( q`2/|.q.|-sn))/(1-sn)
                        by A3,XCMPLX_1:198;
         --(1-sn)>= -(q`2/|.q.|-sn) by A61,A64,REAL_1:50;
       then (-(q`2/|.q.|-sn))/(1-sn)<=1 by A3,REAL_2:117;
       then ((-(q`2/|.q.|-sn))/(1-sn))^2<=1^2 by A65,JGRAPH_2:7;
       then 1-((-(q`2/|.q.|-sn))/(1-sn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((q`2/|.q.|-sn))/(1-sn))^2>=0 by XCMPLX_1:188;
       then A66: 1-((q`2/|.q.|-sn)/(1-sn))^2>=0 by SQUARE_1:61;
        A67: (q4`1)^2= (|.q.|)^2*(-sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2))^2
                         by A60,SQUARE_1:68
         .= (|.q.|)^2*(sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2))^2
                     by SQUARE_1:61
         .= (|.q.|)^2*(1-((q`2/|.q.|-sn)/(1-sn))^2) by A66,SQUARE_1:def 4;
        A68: (q4`2)^2= (|.q.|)^2*((q`2/|.q.|-sn)/(1-sn))^2
                         by A60,SQUARE_1:68;
          (|.q4.|)^2=(q4`1)^2+(q4`2)^2 by JGRAPH_3:10
        .=(|.q.|)^2*(1-((q`2/|.q.|-sn)/(1-sn))^2
           +((q`2/|.q.|-sn)/(1-sn))^2) by A67,A68,XCMPLX_1:8
        .=(|.q.|)^2*1 by XCMPLX_1:27 .=(|.q.|)^2;
        then sqrt((|.q4.|)^2)=|.q.| by A39,SQUARE_1:89;
        then A69: |.q4.|=|.q.| by A59,SQUARE_1:89;
        A70: |.p4.|>=0 by TOPRNS_1:26;
        A71: |.p.|>0 by A56,Lm1;
        then A72: (|.p.|)^2>0 by SQUARE_1:74;
        A73: (p`2/|.p.|-sn)>= 0 by A56,SQUARE_1:12;
        A74: (|.p.|)^2 =(p`1)^2+(p`2)^2 by JGRAPH_3:10;
          0<=(p`1)^2 by SQUARE_1:72;
        then 0+(p`2)^2<=(p`1)^2+(p`2)^2 by REAL_1:55;
        then (p`2)^2/(|.p.|)^2 <= (|.p.|)^2/(|.p.|)^2 by A72,A74,REAL_1:73;
        then (p`2)^2/(|.p.|)^2 <= 1 by A72,XCMPLX_1:60;
        then ((p`2)/|.p.|)^2 <= 1 by SQUARE_1:69;
        then 1>=p`2/|.p.| by Th4;
        then A75: 1-sn>=p`2/|.p.|-sn by REAL_1:49;
        A76: -(1-sn)<= -0 by A3,REAL_1:50;
           -(1-sn)<= -( p`2/|.p.|-sn) by A75,REAL_1:50;
         then (-(1-sn))/(1-sn)<=(-( p`2/|.p.|-sn))/(1-sn)
                        by A3,REAL_1:73;
        then A77: -1<=(-( p`2/|.p.|-sn))/(1-sn)
                        by A3,XCMPLX_1:198;
         --(1-sn)>= -(p`2/|.p.|-sn) by A73,A76,REAL_1:50;
       then (-(p`2/|.p.|-sn))/(1-sn)<=1 by A3,REAL_2:117;
       then ((-(p`2/|.p.|-sn))/(1-sn))^2<=1^2 by A77,JGRAPH_2:7;
       then 1-((-(p`2/|.p.|-sn))/(1-sn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((p`2/|.p.|-sn))/(1-sn))^2>=0 by XCMPLX_1:188;
       then A78: 1-((p`2/|.p.|-sn)/(1-sn))^2>=0 by SQUARE_1:61;
        A79: (p4`1)^2= (|.p.|)^2*(-sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2))^2
                         by A58,SQUARE_1:68
         .= (|.p.|)^2*(sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2))^2
                     by SQUARE_1:61
         .= (|.p.|)^2*(1-((p`2/|.p.|-sn)/(1-sn))^2) by A78,SQUARE_1:def 4;
        A80: (p4`2)^2= (|.p.|)^2*((p`2/|.p.|-sn)/(1-sn))^2
                         by A58,SQUARE_1:68;
          (|.p4.|)^2=(p4`1)^2+(p4`2)^2 by JGRAPH_3:10
        .=(|.p.|)^2*(1-((p`2/|.p.|-sn)/(1-sn))^2
           +((p`2/|.p.|-sn)/(1-sn))^2) by A79,A80,XCMPLX_1:8
        .=(|.p.|)^2*1 by XCMPLX_1:27 .=(|.p.|)^2;
        then sqrt((|.p4.|)^2)=|.p.| by A71,SQUARE_1:89;
        then A81: |.p4.|=|.p.| by A70,SQUARE_1:89;
        ((p`2/|.p.|-sn)/(1-sn))
        =|.q.|* ((q`2/|.q.|-sn)/(1-sn))/|.p.|
                         by A2,A37,A38,A57,A58,A71,XCMPLX_1:90;
        then (p`2/|.p.|-sn)/(1-sn)=(q`2/|.q.|-sn)/(1-sn)
                              by A2,A37,A57,A69,A71,A81,XCMPLX_1:90;
        then (p`2/|.p.|-sn)/(1-sn)*(1-sn)=q`2/|.q.|-sn
                            by A3,XCMPLX_1:88;
        then p`2/|.p.|-sn=q`2/|.q.|-sn by A3,XCMPLX_1:88;
        then p`2/|.p.|-sn+sn=q`2/|.q.| by XCMPLX_1:27;
        then p`2/|.p.|=q`2/|.q.| by XCMPLX_1:27;
        then p`2/|.p.|*|.p.|=q`2 by A2,A37,A57,A69,A71,A81,XCMPLX_1:88;
        then A82: p`2=q`2 by A71,XCMPLX_1:88;
        A83: |.p.|^2=(p`1)^2+(p`2)^2 by JGRAPH_3:10;
          |.q.|^2=(q`1)^2+(q`2)^2 by JGRAPH_3:10;
        then (p`1)^2+(p`2)^2-(p`2)^2=(q`1)^2 by A2,A37,A57,A69,A81,A82,A83,
XCMPLX_1:26;
        then (p`1)^2=(q`1)^2 by XCMPLX_1:26;
        then (-p`1)^2=(q`1)^2 by SQUARE_1:61;
        then A84: sqrt((-p`1)^2)=sqrt((-q`1)^2) by SQUARE_1:61;
          --p`1<=0 by A56;
        then -p`1>=0 by REAL_1:66;
        then A85: -p`1=sqrt((-q`1)^2) by A84,SQUARE_1:89;
          --q`1<=0 by A36;
        then -q`1>=0 by REAL_1:66;
        then A86: --p`1=--q`1 by A85,SQUARE_1:89;
          p=|[p`1,p`2]| by EUCLID:57;
       hence x1=x2 by A82,A86,EUCLID:57;
      case A87: p<>0.REAL 2 &
        (p`2/|.p.|<sn & p`1<=0);
        then A88:
         sn-FanMorphW.p= |[ |.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)),
          |.p.|* ((p`2/|.p.|-sn)/(1+sn))]| by A1,Th25;
         set p4= |[ |.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)),
          |.p.|* ((p`2/|.p.|-sn)/(1+sn))]|;
        A89: p4`1= |.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)) &
          p4`2= |.p.|* ((p`2/|.p.|-sn)/(1+sn)) by EUCLID:56;
         A90: p`2/|.p.|-sn<0 by A87,REAL_2:105;
           1+sn>0 by A1,Th3;
         then A91: ((p`2/|.p.|-sn)/(1+sn))<0 by A90,REAL_2:128;
           |.p.|>0 by A87,Lm1;
         then A92: |.p.|* ((p`2/|.p.|-sn)/(1+sn))<0 by A91,SQUARE_1:24;
         A93: q`2/|.q.|-sn>=0 by A36,SQUARE_1:12;
           1-sn>0 by A1,Th3;
         then ((q`2/|.q.|-sn)/(1-sn))>=0 by A93,REAL_2:125;
       hence x1=x2 by A2,A37,A38,A39,A88,A89,A92,REAL_2:121;
      end;
     hence x1=x2;
    case A94:
      q`2/|.q.|<sn & q`1<=0 & q<>0.REAL 2;
        then A95:
         sn-FanMorphW.q= |[ |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1+sn))]| by A1,Th25;
         set q4= |[ |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1+sn))]|;
        A96: q4`1= |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)) &
          q4`2= |.q.|* ((q`2/|.q.|-sn)/(1+sn)) by EUCLID:56;
     A97: |.q.|>=0 by TOPRNS_1:26;
     A98: |.q.|>0 by A94,Lm1;
     then A99: (|.q.|)^2>0 by SQUARE_1:74;
        now per cases by JGRAPH_2:11;
      case A100: p`1>=0;
        then A101:(sn-FanMorphW).p=p by Th23;
        A102: 1+sn>0 by A1,Th3;
        A103: (q`2/|.q.|-sn)<=0 by A94,REAL_2:106;
        then A104: -(q`2/|.q.|-sn)>= -0 by REAL_1:50;
          --(q`2/|.q.|-sn)/(1+sn)<=0 by A102,A103,REAL_2:126;
        then A105: -((q`2/|.q.|-sn)/(1+sn))>=0 by REAL_1:66;
        A106: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
          0<=(q`1)^2 by SQUARE_1:72;
        then 0+(q`2)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
        then (q`2)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A99,A106,REAL_1:73;
        then (q`2)^2/(|.q.|)^2 <= 1 by A99,XCMPLX_1:60;
        then ((q`2)/|.q.|)^2 <= 1 by SQUARE_1:69;
        then A107: (-((q`2)/|.q.|))^2 <= 1 by SQUARE_1:61;
          -(1+sn)<= -0 by A102,REAL_1:50;
        then (-(1+sn))/(1+sn)<=(-( q`2/|.q.|-sn))/(1+sn)
                        by A102,A104,REAL_1:73;
        then A108: -1<=(-( q`2/|.q.|-sn))/(1+sn)
                        by A102,XCMPLX_1:198;
         1>= -q`2/|.q.| by A107,Th4;
       then (1+sn)>= -q`2/|.q.|+sn by REAL_1:55;
       then (1+sn)>= -(q`2/|.q.|-sn) by XCMPLX_1:162;
       then (-(q`2/|.q.|-sn))/(1+sn)<=1 by A102,REAL_2:117;
       then ((-(q`2/|.q.|-sn))/(1+sn))^2<=1^2 by A108,JGRAPH_2:7;
       then A109: 1-((-(q`2/|.q.|-sn))/(1+sn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((q`2/|.q.|-sn))/(1+sn))^2>=0 by XCMPLX_1:188;
       then A110: 1-((q`2/|.q.|-sn)/(1+sn))^2>=0 by SQUARE_1:61;
         sqrt(1-((-(q`2/|.q.|-sn))/(1+sn))^2)>=0 by A109,SQUARE_1:def 4;
       then sqrt(1-(-(q`2/|.q.|-sn))^2/(1+sn)^2)>=0 by SQUARE_1:69;
       then sqrt(1-((q`2/|.q.|-sn))^2/(1+sn)^2)>=0 by SQUARE_1:61;
       then sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)>=0 by SQUARE_1:69;
       then -sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)<= -0 by REAL_1:50;
        then p`1=0 by A2,A95,A96,A97,A100,A101,REAL_2:123;
        then (-sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2))=0 by A2,A95,A96,A98,A101,
XCMPLX_1:6;
        then sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)=-0 .=0;
        then 1-((q`2/|.q.|-sn)/(1+sn))^2=0 by A110,SQUARE_1:92;
        then 1=0+((q`2/|.q.|-sn)/(1+sn))^2 by XCMPLX_1:27;
        then 1=sqrt((-((q`2/|.q.|-sn)/(1+sn)))^2) by SQUARE_1:61,83;
        then 1= -((q`2/|.q.|-sn)/(1+sn)) by A105,SQUARE_1:89;
        then 1= ((-(q`2/|.q.|-sn))/(1+sn)) by XCMPLX_1:188;
        then (1)*(1+sn)=-(q`2/|.q.|-sn) by A102,XCMPLX_1:88;
        then (1+sn)=-q`2/|.q.|+sn by XCMPLX_1:162;
        then 1+sn-sn=-q`2/|.q.| by XCMPLX_1:26;
        then 1=-q`2/|.q.| by XCMPLX_1:26;
        then 1=(-q`2)/|.q.| by XCMPLX_1:188;
        then (1)*|.q.|=-q`2 by A98,XCMPLX_1:88;
        then (-q`2)^2-(q`2)^2 =(q`1)^2 by A106,XCMPLX_1:26;
        then (q`2)^2-(q`2)^2 =(q`1)^2 by SQUARE_1:61;
        then 0= (q`1)^2 by XCMPLX_1:14;
        then q`1=0 by SQUARE_1:73;
       hence x1=x2 by A2,A101,Th23;
      case A111:p<>0.REAL 2 &
          (p`2/|.p.|>=sn & p`1<=0);
        then A112:
         sn-FanMorphW.p= |[ |.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)),
          |.p.|* ((p`2/|.p.|-sn)/(1-sn))]| by A1,Th25;
         set p4= |[ |.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)),
          |.p.|* ((p`2/|.p.|-sn)/(1-sn))]|;
        A113: p4`1= |.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)) &
          p4`2= |.p.|* ((p`2/|.p.|-sn)/(1-sn)) by EUCLID:56;
         A114: q`2/|.q.|-sn<0 by A94,REAL_2:105;
           1+sn>0 by A1,Th3;
         then A115: ((q`2/|.q.|-sn)/(1+sn))<0 by A114,REAL_2:128;
           |.q.|>0 by A94,Lm1;
         then A116: |.q.|* ((q`2/|.q.|-sn)/(1+sn))<0 by A115,SQUARE_1:24;
         A117: p`2/|.p.|-sn>=0 by A111,SQUARE_1:12;
           1-sn>0 by A1,Th3;
         then A118: ((p`2/|.p.|-sn)/(1-sn))>=0 by A117,REAL_2:125;
           |.p.|>=0 by TOPRNS_1:26;
       hence x1=x2 by A2,A95,A96,A112,A113,A116,A118,REAL_2:121;
      case A119:p<>0.REAL 2 &
          (p`2/|.p.|<sn & p`1<=0);
        then A120:
         sn-FanMorphW.p= |[ |.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)),
          |.p.|* ((p`2/|.p.|-sn)/(1+sn))]| by A1,Th25;
         set p4= |[ |.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)),
          |.p.|* ((p`2/|.p.|-sn)/(1+sn))]|;
        A121: p4`1= |.p.|*(-sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)) &
          p4`2= |.p.|* ((p`2/|.p.|-sn)/(1+sn)) by EUCLID:56;
        A122: |.q4.|>=0 by TOPRNS_1:26;
        A123: q4`1= |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2))&
        q4`2= |.q.|* ((q`2/|.q.|-sn)/(1+sn)) by EUCLID:56;
          (q`2/|.q.|-sn)<=0 by A94,REAL_2:106;
        then A124: -(q`2/|.q.|-sn)>= -0 by REAL_1:50;
        A125: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
          0<=(q`1)^2 by SQUARE_1:72;
        then 0+(q`2)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
        then (q`2)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A99,A125,REAL_1:73;
        then (q`2)^2/(|.q.|)^2 <= 1 by A99,XCMPLX_1:60;
        then ((q`2)/|.q.|)^2 <= 1 by SQUARE_1:69;
        then -1<=q`2/|.q.| by Th4;
        then A126: -1-sn<=q`2/|.q.|-sn by REAL_1:49;
        A127: 1+sn>0 by A1,Th3;
        then -(1+sn)<= -0 by REAL_1:50;
        then (-(1+sn))/(1+sn)<=(-( q`2/|.q.|-sn))/(1+sn) by A124,A127,REAL_1:73
;
        then A128: -1<=(-( q`2/|.q.|-sn))/(1+sn) by A127,XCMPLX_1:198;
          -(1+sn)<=q`2/|.q.|-sn by A126,XCMPLX_1:161;
        then --(1+sn)>= -(q`2/|.q.|-sn) by REAL_1:50;
       then (-(q`2/|.q.|-sn))/(1+sn)<=1 by A127,REAL_2:117;
       then ((-(q`2/|.q.|-sn))/(1+sn))^2<=1^2 by A128,JGRAPH_2:7;
       then 1-((-(q`2/|.q.|-sn))/(1+sn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((q`2/|.q.|-sn))/(1+sn))^2>=0 by XCMPLX_1:188;
       then A129: 1-((q`2/|.q.|-sn)/(1+sn))^2>=0 by SQUARE_1:61;
        A130: (q4`1)^2= (|.q.|)^2*(-sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2))^2
                         by A123,SQUARE_1:68
         .= (|.q.|)^2*(sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2))^2
                     by SQUARE_1:61
         .= (|.q.|)^2*(1-((q`2/|.q.|-sn)/(1+sn))^2) by A129,SQUARE_1:def 4;
        A131: (q4`2)^2= (|.q.|)^2*((q`2/|.q.|-sn)/(1+sn))^2
                         by A123,SQUARE_1:68;
          (|.q4.|)^2=(q4`1)^2+(q4`2)^2 by JGRAPH_3:10
        .=(|.q.|)^2*(1-((q`2/|.q.|-sn)/(1+sn))^2
           +((q`2/|.q.|-sn)/(1+sn))^2) by A130,A131,XCMPLX_1:8
        .=(|.q.|)^2*1 by XCMPLX_1:27 .=(|.q.|)^2;
        then sqrt((|.q4.|)^2)=|.q.| by A98,SQUARE_1:89;
        then A132: |.q4.|=|.q.| by A122,SQUARE_1:89;
        A133: |.p4.|>=0 by TOPRNS_1:26;
        A134: |.p.|>0 by A119,Lm1;
        then A135: (|.p.|)^2>0 by SQUARE_1:74;
          (p`2/|.p.|-sn)<=0 by A119,REAL_2:106;
        then A136: -(p`2/|.p.|-sn)>= -0 by REAL_1:50;
        A137: (|.p.|)^2 =(p`1)^2+(p`2)^2 by JGRAPH_3:10;
          0<=(p`1)^2 by SQUARE_1:72;
        then 0+(p`2)^2<=(p`1)^2+(p`2)^2 by REAL_1:55;
        then (p`2)^2/(|.p.|)^2 <= (|.p.|)^2/(|.p.|)^2 by A135,A137,REAL_1:73;
        then (p`2)^2/(|.p.|)^2 <= 1 by A135,XCMPLX_1:60;
        then ((p`2)/|.p.|)^2 <= 1 by SQUARE_1:69;
        then -1<=p`2/|.p.| by Th4;
        then A138: -1-sn<=p`2/|.p.|-sn by REAL_1:49;
          -(1+sn)<= -0 by A127,REAL_1:50;
        then (-(1+sn))/(1+sn)<=(-( p`2/|.p.|-sn))/(1+sn)
                        by A127,A136,REAL_1:73;
        then A139: -1<=(-( p`2/|.p.|-sn))/(1+sn)
                        by A127,XCMPLX_1:198;
          -(1+sn)<=p`2/|.p.|-sn by A138,XCMPLX_1:161;
        then --(1+sn)>= -(p`2/|.p.|-sn) by REAL_1:50;
       then (-(p`2/|.p.|-sn))/(1+sn)<=1 by A127,REAL_2:117;
       then ((-(p`2/|.p.|-sn))/(1+sn))^2<=1^2 by A139,JGRAPH_2:7;
       then 1-((-(p`2/|.p.|-sn))/(1+sn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((p`2/|.p.|-sn))/(1+sn))^2>=0 by XCMPLX_1:188;
       then A140: 1-((p`2/|.p.|-sn)/(1+sn))^2>=0 by SQUARE_1:61;
        A141: (p4`1)^2= (|.p.|)^2*(-sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2))^2
                         by A121,SQUARE_1:68
         .= (|.p.|)^2*(sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2))^2
                     by SQUARE_1:61
         .= (|.p.|)^2*(1-((p`2/|.p.|-sn)/(1+sn))^2)
                      by A140,SQUARE_1:def 4;
        A142: (p4`2)^2= (|.p.|)^2*((p`2/|.p.|-sn)/(1+sn))^2
                         by A121,SQUARE_1:68;
          (|.p4.|)^2=(p4`1)^2+(p4`2)^2 by JGRAPH_3:10
        .=(|.p.|)^2*(1-((p`2/|.p.|-sn)/(1+sn))^2
           +((p`2/|.p.|-sn)/(1+sn))^2) by A141,A142,XCMPLX_1:8
        .=(|.p.|)^2*1 by XCMPLX_1:27 .=(|.p.|)^2;
        then sqrt((|.p4.|)^2)=|.p.| by A134,SQUARE_1:89;
        then A143: |.p4.|=|.p.| by A133,SQUARE_1:89;
        ((p`2/|.p.|-sn)/(1+sn))
        =|.q.|* ((q`2/|.q.|-sn)/(1+sn))/|.p.|
                         by A2,A95,A96,A120,A121,A134,XCMPLX_1:90;
        then (p`2/|.p.|-sn)/(1+sn)=(q`2/|.q.|-sn)/(1+sn)
                              by A2,A95,A120,A132,A134,A143,XCMPLX_1:90
;
        then (p`2/|.p.|-sn)/(1+sn)*(1+sn)=q`2/|.q.|-sn
                            by A127,XCMPLX_1:88;
        then p`2/|.p.|-sn=q`2/|.q.|-sn by A127,XCMPLX_1:88;
        then p`2/|.p.|-sn+sn=q`2/|.q.| by XCMPLX_1:27;
        then p`2/|.p.|=q`2/|.q.| by XCMPLX_1:27;
        then p`2/|.p.|*|.p.|=q`2 by A2,A95,A120,A132,A134,A143,XCMPLX_1:88;
        then A144: p`2=q`2 by A134,XCMPLX_1:88;
        A145: |.p.|^2=(p`1)^2+(p`2)^2 by JGRAPH_3:10;
          |.q.|^2=(q`1)^2+(q`2)^2 by JGRAPH_3:10;
        then (p`1)^2+(p`2)^2-(p`2)^2=(q`1)^2 by A2,A95,A120,A132,A143,A144,A145
,XCMPLX_1:26;
        then (p`1)^2=(q`1)^2 by XCMPLX_1:26;
        then (-p`1)^2=(q`1)^2 by SQUARE_1:61;
        then A146: sqrt((-p`1)^2)=sqrt((-q`1)^2) by SQUARE_1:61;
          --p`1<=0 by A119;
        then -p`1>=0 by REAL_1:66;
        then A147: -p`1=sqrt((-q`1)^2) by A146,SQUARE_1:89;
          --q`1<=0 by A94;
        then -q`1>=0 by REAL_1:66;
        then A148: --p`1=--q`1 by A147,SQUARE_1:89;
          p=|[p`1,p`2]| by EUCLID:57;
       hence x1=x2 by A144,A148,EUCLID:57;
      end;
     hence x1=x2;
    end;
   hence x1=x2;
  end;
 hence thesis by FUNCT_1:def 8;
end;

theorem Th46: for sn being Real
  st -1<sn & sn<1 holds
  (sn-FanMorphW) is map of TOP-REAL 2,TOP-REAL 2 &
 rng (sn-FanMorphW) = the carrier of TOP-REAL 2
proof let sn be Real;
  assume A1:  -1<sn & sn<1;
thus (sn-FanMorphW) is map of TOP-REAL 2,TOP-REAL 2 ;
  A2:for f being map of TOP-REAL 2,TOP-REAL 2 st f=(sn-FanMorphW) holds
    rng (sn-FanMorphW)=the carrier of TOP-REAL 2
  proof let f be map of TOP-REAL 2,TOP-REAL 2;
    assume A3:f=(sn-FanMorphW);
    A4:dom f=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
      the carrier of TOP-REAL 2 c= rng f
    proof let y be set;assume y in the carrier of TOP-REAL 2;
     then reconsider p2=y as Point of TOP-REAL 2;
     set q=p2;
       now per cases by JGRAPH_2:11;
     case q`1>=0;
       then q in dom (sn-FanMorphW) & y=(sn-FanMorphW).q by A3,A4,Th23;
      hence ex x being set st x in dom (sn-FanMorphW) & y=(sn-FanMorphW).x;
     case A5: q`2/|.q.|>=0 & q`1<=0 & q<>0.REAL 2;
          set px=|[ -(|.q.|)*sqrt(1-(q`2/|.q.|*(1-sn)+sn)^2),
                   |.q.|*(q`2/|.q.|*(1-sn)+sn)]|;
        A6: px`1 = -(|.q.|)*sqrt(1-(q`2/|.q.|*(1-sn)+sn)^2) &
         px`2 = |.q.|*(q`2/|.q.|*(1-sn)+sn) by EUCLID:56;
        then A7: |.px.|^2=(-(|.q.|)*sqrt(1-(q`2/|.q.|*(1-sn)+sn)^2))^2
        +(|.q.|*(q`2/|.q.|*(1-sn)+sn))^2 by JGRAPH_3:10
        .=(|.q.|*sqrt(1-(q`2/|.q.|*(1-sn)+sn)^2))^2
        +(|.q.|*(q`2/|.q.|*(1-sn)+sn))^2 by SQUARE_1:61
        .=(|.q.|)^2*(sqrt(1-(q`2/|.q.|*(1-sn)+sn)^2))^2
        +(|.q.|*(q`2/|.q.|*(1-sn)+sn))^2 by SQUARE_1:68
        .=(|.q.|)^2*(sqrt(1-(q`2/|.q.|*(1-sn)+sn)^2))^2
        +(|.q.|)^2*((q`2/|.q.|*(1-sn)+sn))^2 by SQUARE_1:68;
       A8: |.q.|<>0 by A5,TOPRNS_1:25;
         1-sn>=0 by A1,Th3;
       then A9: q`2/|.q.|*(1-sn)>=0 by A5,REAL_2:121;
       then A10: (q`2/|.q.|*(1-sn)+sn)>=0+sn by REAL_1:55;
       A11: |.q.|>=0 by TOPRNS_1:26;
       A12: |.q.|>0 by A5,Lm1;
       A13: |.q.|^2>0 by A8,SQUARE_1:74;
         --(1+sn)>0 by A1,Th3;
       then -(1+sn)<0 by REAL_1:66;
       then -1-sn <0 by XCMPLX_1:161;
       then -1-sn<= q`2/|.q.|*(1-sn) by A9,AXIOMS:22;
       then -1-sn+sn<= q`2/|.q.|*(1-sn)+sn by REAL_1:55;
       then A14: -1<= (q`2/|.q.|*(1-sn)+sn) by XCMPLX_1:27;
       A15: 1-sn>0 by A1,Th3;
       A16: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
         0<=(q`1)^2 by SQUARE_1:72;
       then 0+(q`2)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
       then (q`2)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A13,A16,REAL_1:73;
       then (q`2)^2/(|.q.|)^2 <= 1 by A13,XCMPLX_1:60;
       then ((q`2)/|.q.|)^2 <= 1 by SQUARE_1:69;
       then q`2/|.q.|<=1 by Th4;
       then q`2/|.q.|*(1-sn) <=(1)*(1-sn) by A15,AXIOMS:25;
       then q`2/|.q.|*(1-sn)+sn-sn <=1-sn by XCMPLX_1:26;
       then (q`2/|.q.|*(1-sn)+sn) <=1 by REAL_1:54;
       then 1^2>=(q`2/|.q.|*(1-sn)+sn)^2 by A14,JGRAPH_2:7;
       then A17: 1-(q`2/|.q.|*(1-sn)+sn)^2>=0 by SQUARE_1:12,59;
       then A18: |.px.|^2=(|.q.|)^2*(1-(q`2/|.q.|*(1-sn)+sn)^2)
        +(|.q.|)^2*((q`2/|.q.|*(1-sn)+sn))^2 by A7,SQUARE_1:def 4
        .= (|.q.|)^2*((1-(q`2/|.q.|*(1-sn)+sn)^2)
        +((q`2/|.q.|*(1-sn)+sn))^2) by XCMPLX_1:8
        .= (|.q.|)^2*(1) by XCMPLX_1:27
        .= (|.q.|)^2;
          |.px.|>=0 by TOPRNS_1:26;
        then A19: |.px.|=sqrt(|.q.|^2) by A18,SQUARE_1:89
        .=|.q.| by A11,SQUARE_1:89;
        then A20: |.px.|<>0 by A5,TOPRNS_1:25;
         sqrt(1-(q`2/|.q.|*(1-sn)+sn)^2)>=0 by A17,SQUARE_1:def 4;
       then |.q.|*sqrt(1-(q`2/|.q.|*(1-sn)+sn)^2)>=0 by A11,REAL_2:121;
        then px`2/|.px.| >=sn & px`1<=0 & px<>0.REAL 2
                           by A6,A8,A10,A19,REAL_1:66,TOPRNS_1:24,XCMPLX_1:90;
    then A21: (sn-FanMorphW).px
       =|[ |.px.|*(-sqrt(1-((px`2/|.px.|-sn)/(1-sn))^2)),
               |.px.|* ((px`2/|.px.|-sn)/(1-sn))]| by A1,Th25;
        A22: |.px.|* ((px`2/|.px.|-sn)/(1-sn))
        =|.q.|* (( ((q`2/|.q.|*(1-sn)+sn))-sn)/(1-sn))
                            by A6,A8,A19,XCMPLX_1:90
        .=|.q.|* (( q`2/|.q.|*(1-sn) )/(1-sn))
                            by XCMPLX_1:26
        .=|.q.|* ( q`2/|.q.|) by A15,XCMPLX_1:90
        .= q`2 by A8,XCMPLX_1:88;
        then A23: |.px.|*(-sqrt(1-((px`2/|.px.|-sn)/(1-sn))^2))
           = |.px.|*(-sqrt(1-(q`2/|.px.|)^2)) by A20,XCMPLX_1:90
           .= |.px.|*(-sqrt(1-(q`2)^2/(|.px.|)^2)) by SQUARE_1:69
           .= |.px.|*(-sqrt( (|.px.|)^2/(|.px.|)^2-(q`2)^2/(|.px.|)^2))
                            by A13,A18,XCMPLX_1:60
           .= |.px.|*(-sqrt( ((|.px.|)^2-(q`2)^2)/(|.px.|)^2))
                            by XCMPLX_1:121
           .= |.px.|*(-sqrt( ((q`1)^2+(q`2)^2-(q`2)^2)/(|.px.|)^2))
                            by A18,JGRAPH_3:10
           .= |.px.|*(-sqrt( ((q`1)^2)/(|.px.|)^2))
                            by XCMPLX_1:26
           .= |.px.|*(-sqrt((q`1/|.q.|)^2))
                            by A19,SQUARE_1:69;
           q`1/|.q.|<=0 by A5,A12,REAL_2:126;
         then |.px.|*(-sqrt((q`1/|.q.|)^2))=|.px.|*(--(q`1/|.q.|))
                            by SQUARE_1:90
         .=q`1 by A12,A19,XCMPLX_1:88;
        then A24:q=(sn-FanMorphW).px by A21,A22,A23,EUCLID:57;
          dom (sn-FanMorphW)=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
      hence ex x being set st x in dom (sn-FanMorphW) & y=(sn-FanMorphW).x
                               by A24;
     case A25: q`2/|.q.|<0 & q`1<=0 & q<>0.REAL 2;
          set px=|[ -(|.q.|)*sqrt(1-(q`2/|.q.|*(1+sn)+sn)^2),
                   |.q.|*(q`2/|.q.|*(1+sn)+sn)]|;
        A26: px`1 = -(|.q.|)*sqrt(1-(q`2/|.q.|*(1+sn)+sn)^2) &
         px`2 = |.q.|*(q`2/|.q.|*(1+sn)+sn) by EUCLID:56;
        then A27: |.px.|^2=(-(|.q.|)*sqrt(1-(q`2/|.q.|*(1+sn)+sn)^2))^2
        +(|.q.|*(q`2/|.q.|*(1+sn)+sn))^2 by JGRAPH_3:10
        .=(|.q.|*sqrt(1-(q`2/|.q.|*(1+sn)+sn)^2))^2
        +(|.q.|*(q`2/|.q.|*(1+sn)+sn))^2 by SQUARE_1:61
        .=(|.q.|)^2*(sqrt(1-(q`2/|.q.|*(1+sn)+sn)^2))^2
        +(|.q.|*(q`2/|.q.|*(1+sn)+sn))^2 by SQUARE_1:68
        .=(|.q.|)^2*(sqrt(1-(q`2/|.q.|*(1+sn)+sn)^2))^2
        +(|.q.|)^2*((q`2/|.q.|*(1+sn)+sn))^2 by SQUARE_1:68;
       A28: |.q.|<>0 by A25,TOPRNS_1:25;
         1+sn>=0 by A1,Th3;
       then A29: q`2/|.q.|*(1+sn)<=0 by A25,REAL_2:123;
       then A30: (q`2/|.q.|*(1+sn)+sn)<=0+sn by REAL_1:55;
       A31: |.q.|>=0 by TOPRNS_1:26;
       A32: |.q.|>0 by A25,Lm1;
       A33: |.q.|^2>0 by A28,SQUARE_1:74;
         (1-sn)>0 by A1,Th3;
       then 1-sn>= q`2/|.q.|*(1+sn) by A29,AXIOMS:22;
       then 1-sn+sn>= q`2/|.q.|*(1+sn)+sn by REAL_1:55;
       then A34: 1>= (q`2/|.q.|*(1+sn)+sn) by XCMPLX_1:27;
       A35: 1+sn>0 by A1,Th3;
       A36: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
         0<=(q`1)^2 by SQUARE_1:72;
       then 0+(q`2)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
       then (q`2)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A33,A36,REAL_1:73;
       then (q`2)^2/(|.q.|)^2 <= 1 by A33,XCMPLX_1:60;
       then ((q`2)/|.q.|)^2 <= 1 by SQUARE_1:69;
       then q`2/|.q.|>=-1 by Th4;
       then q`2/|.q.|*(1+sn) >=(-1)*(1+sn) by A35,AXIOMS:25;
       then q`2/|.q.|*(1+sn) >=-(1+sn) by XCMPLX_1:180;
       then q`2/|.q.|*(1+sn) >=-1-sn by XCMPLX_1:161;
       then q`2/|.q.|*(1+sn)+sn-sn >=-1-sn by XCMPLX_1:26;
       then (q`2/|.q.|*(1+sn)+sn) >=-1 by REAL_1:54;
       then 1^2>=(q`2/|.q.|*(1+sn)+sn)^2 by A34,JGRAPH_2:7;
       then A37: 1-(q`2/|.q.|*(1+sn)+sn)^2>=0 by SQUARE_1:12,59;
       then A38: |.px.|^2=(|.q.|)^2*(1-(q`2/|.q.|*(1+sn)+sn)^2)
        +(|.q.|)^2*((q`2/|.q.|*(1+sn)+sn))^2 by A27,SQUARE_1:def 4
        .= (|.q.|)^2*((1-(q`2/|.q.|*(1+sn)+sn)^2)
        +((q`2/|.q.|*(1+sn)+sn))^2) by XCMPLX_1:8
        .= (|.q.|)^2*(1) by XCMPLX_1:27
        .= (|.q.|)^2;
          |.px.|>=0 by TOPRNS_1:26;
        then A39: |.px.|=sqrt(|.q.|^2) by A38,SQUARE_1:89
        .=|.q.| by A31,SQUARE_1:89;
        then A40: |.px.|<>0 by A25,TOPRNS_1:25;
         sqrt(1-(q`2/|.q.|*(1+sn)+sn)^2)>=0 by A37,SQUARE_1:def 4;
       then |.q.|*sqrt(1-(q`2/|.q.|*(1+sn)+sn)^2)>=0 by A31,REAL_2:121;
        then px`2/|.px.| <=sn & px`1<=0 & px<>0.REAL 2
                           by A26,A28,A30,A39,REAL_1:66,TOPRNS_1:24,XCMPLX_1:90
;
    then A41: (sn-FanMorphW).px
       =|[ |.px.|*(-sqrt(1-((px`2/|.px.|-sn)/(1+sn))^2)),
               |.px.|* ((px`2/|.px.|-sn)/(1+sn))]|
                            by A1,Th25;
        A42: |.px.|* ((px`2/|.px.|-sn)/(1+sn))
        =|.q.|* (( ((q`2/|.q.|*(1+sn)+sn))-sn)/(1+sn)) by A26,A28,A39,XCMPLX_1:
90
        .=|.q.|* (( q`2/|.q.|*(1+sn) )/(1+sn)) by XCMPLX_1:26
        .=|.q.|* ( q`2/|.q.|) by A35,XCMPLX_1:90
        .= q`2 by A28,XCMPLX_1:88;
        then A43: |.px.|*(-sqrt(1-((px`2/|.px.|-sn)/(1+sn))^2))
           = |.px.|*(-sqrt(1-(q`2/|.px.|)^2)) by A40,XCMPLX_1:90
           .= |.px.|*(-sqrt(1-(q`2)^2/(|.px.|)^2)) by SQUARE_1:69
           .= |.px.|*(-sqrt( (|.px.|)^2/(|.px.|)^2-(q`2)^2/(|.px.|)^2))
                            by A33,A38,XCMPLX_1:60
           .= |.px.|*(-sqrt( ((|.px.|)^2-(q`2)^2)/(|.px.|)^2))
                            by XCMPLX_1:121
           .= |.px.|*(-sqrt( ((q`1)^2+(q`2)^2-(q`2)^2)/(|.px.|)^2))
                            by A38,JGRAPH_3:10
           .= |.px.|*(-sqrt( ((q`1)^2)/(|.px.|)^2))
                            by XCMPLX_1:26
           .= |.px.|*(-sqrt((q`1/|.q.|)^2))
                            by A39,SQUARE_1:69;
           q`1/|.q.|<=0 by A25,A32,REAL_2:126;
         then |.px.|*(-sqrt((q`1/|.q.|)^2))=|.px.|*(--(q`1/|.q.|))
                            by SQUARE_1:90
         .=q`1 by A32,A39,XCMPLX_1:88;
        then A44:q=(sn-FanMorphW).px by A41,A42,A43,EUCLID:57;
          dom (sn-FanMorphW)=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
      hence ex x being set st x in dom (sn-FanMorphW) & y=(sn-FanMorphW).x
                               by A44;
     end;
    hence y in rng f by A3,FUNCT_1:def 5;
   end;
   hence rng (sn-FanMorphW)=the carrier of TOP-REAL 2
                 by A3,XBOOLE_0:def 10;
  end;
 thus rng (sn-FanMorphW)=the carrier of TOP-REAL 2 by A2;
end;

Lm11:now let K0 be Subset of TOP-REAL 2, q4,q,p2 be Point of TOP-REAL 2,
        O,u,uq be Point of Euclid 2;
  assume A1: u in cl_Ball(O,|.p2.|+1);
  assume A2: q = uq & q4 = u & O=0.REAL 2;
  assume A3: |.q4.|=|.q.|;
    now assume not q in cl_Ball(O,|.p2.|+1);
    then not dist(O,uq)<=|.p2.|+1 by A2,METRIC_1:13;
    then |.0.REAL 2 - q.|> |.p2.|+1 by A2,JGRAPH_1:45;
    then |.0.REAL 2 +- q.|> |.p2.|+1 by EUCLID:45;
    then |. -q.|> |.p2.|+1 by EUCLID:31;
    then |.q.|> |.p2.|+1 by TOPRNS_1:27;
    then |. -q4.|> |.p2.|+1 by A3,TOPRNS_1:27;
    then |.0.REAL 2 +- q4.|> |.p2.|+1 by EUCLID:31;
    then |.0.REAL 2 - q4.|> |.p2.|+1 by EUCLID:45;
    then dist(O,u)> |.p2.|+1 by A2,JGRAPH_1:45;
   hence contradiction by A1,METRIC_1:13;
  end; hence q in cl_Ball(O,|.p2.|+1);
end;

theorem Th47: for sn being Real,p2 being Point of TOP-REAL 2
  st -1<sn & sn<1
  ex K being non empty compact Subset of TOP-REAL 2 st
  K = (sn-FanMorphW).:K &
 (ex V2 being Subset of TOP-REAL 2 st
  p2 in V2 & V2 is open & V2 c= K & (sn-FanMorphW).p2 in V2)
 proof let sn be Real,p2 be Point of TOP-REAL 2;
  assume A1: -1<sn & sn<1;
    A2: the carrier of TOP-REAL 2=REAL 2 by EUCLID:25;
A3: Euclid 2=MetrStruct(#REAL 2,Pitag_dist 2#) by EUCLID:def 7;
    then reconsider O=0.REAL 2 as Point of Euclid 2 by EUCLID:25;
    A4: TOP-REAL 2=TopSpaceMetr(Euclid 2) by EUCLID:def 8;
    reconsider V0=Ball(O,|.p2.|+1) as Subset of TOP-REAL 2
                              by A2,A3;
    A5: |. p2.|<|.p2.|+1 by REAL_1:69;
      0<=|.p2.| by TOPRNS_1:26;
    then A6: O in V0 by A5,GOBOARD6:4;
    A7: V0 c= cl_Ball(O,|.p2.|+1) by METRIC_1:15;
    reconsider K01=cl_Ball(O,|.p2.|+1) as Subset of TOP-REAL 2
                            by A2,A3;
      K01 is compact by Th22;
    then reconsider K0=cl_Ball(O,|.p2.|+1)
          as non empty compact Subset of TOP-REAL 2 by A6,A7;
    reconsider u2=p2 as Point of Euclid 2 by A3,EUCLID:25;
    A8: (sn-FanMorphW).:K0 c= K0
     proof let y be set;assume y in (sn-FanMorphW).:K0;
       then consider x being set such that
       A9: x in dom (sn-FanMorphW) & x in K0 & y=(sn-FanMorphW).x
                     by FUNCT_1:def 12;
       reconsider q=x as Point of TOP-REAL 2 by A9;
       reconsider uq=q as Point of Euclid 2 by A3,EUCLID:25;
       A10: y in rng (sn-FanMorphW) by A9,FUNCT_1:def 5;
       then reconsider u=y as Point of Euclid 2 by A3,EUCLID:25;
       reconsider q4=y as Point of TOP-REAL 2 by A10;
       A11: |.q4.|=|.q.| by A9,Th40;
         dist(O,uq)<= |.p2.|+1 by A9,METRIC_1:13;
       then |.0.REAL 2 - q.|<= |.p2.|+1 by JGRAPH_1:45;
       then |.0.REAL 2 +- q.|<= |.p2.|+1 by EUCLID:45;
       then |. -q.|<= |.p2.|+1 by EUCLID:31;
       then |.q.|<= |.p2.|+1 by TOPRNS_1:27;
       then |. -q4.|<= |.p2.|+1 by A11,TOPRNS_1:27;
       then |.0.REAL 2 +- q4.|<= |.p2.|+1 by EUCLID:31;
       then |.0.REAL 2 - q4.|<= |.p2.|+1 by EUCLID:45;
       then dist(O,u)<= |.p2.|+1 by JGRAPH_1:45;
      hence y in K0 by METRIC_1:13;
     end;
      K0 c= (sn-FanMorphW).:K0
     proof let y be set;assume A12: y in K0;
       then reconsider u=y as Point of Euclid 2;
       reconsider q4=y as Point of TOP-REAL 2 by A12;
         the carrier of TOP-REAL 2 c= rng (sn-FanMorphW) by A1,Th46;
       then q4 in rng (sn-FanMorphW) by TARSKI:def 3;
       then consider x being set such that
       A13: x in dom (sn-FanMorphW) & y=(sn-FanMorphW).x by FUNCT_1:def 5;
       reconsider q=x as Point of TOP-REAL 2 by A13,FUNCT_2:def 1;
       reconsider uq=q as Point of Euclid 2 by A3,EUCLID:25;
       A14: |.q4.|=|.q.| by A13,Th40;
         q = uq & q4 = u & O=0.REAL 2;
       then q in K0 by A12,A14,Lm11;
      hence y in (sn-FanMorphW).:K0 by A13,FUNCT_1:def 12;
     end;
    then A15: K0= (sn-FanMorphW).:K0 by A8,XBOOLE_0:def 10;
      |. -p2.|<|.p2.|+1 by A5,TOPRNS_1:27;
    then |.0.REAL 2 +- p2.|<|.p2.|+1 by EUCLID:31;
    then |.0.REAL 2 - p2.|<|.p2.|+1 by EUCLID:45;
    then dist(O,u2)<|.p2.|+1 by JGRAPH_1:45;
    then A16: p2 in V0 by METRIC_1:12;
    A17: V0 is open by A4,TOPMETR:21;
    set q3= (sn-FanMorphW).p2;
    reconsider u3=q3 as Point of Euclid 2 by A3,EUCLID:25;
      |.q3.|=|.p2.| by Th40;
    then |. -q3.|<|.p2.|+1 by A5,TOPRNS_1:27;
    then |.0.REAL 2 +- q3.|<|.p2.|+1 by EUCLID:31;
    then |.0.REAL 2 - q3.|<|.p2.|+1 by EUCLID:45;
    then dist(O,u3)< |.p2.|+1 by JGRAPH_1:45;
    then (sn-FanMorphW).p2 in V0 by METRIC_1:12;
   hence thesis by A7,A15,A16,A17;
 end;

theorem for sn being Real st -1<sn & sn<1
 ex f being map of TOP-REAL 2,TOP-REAL 2 st f=(sn-FanMorphW)
   & f is_homeomorphism
proof let sn be Real;assume A1:  -1<sn & sn<1;
  then A2: (sn-FanMorphW) is map of TOP-REAL 2,TOP-REAL 2 &
  rng (sn-FanMorphW) = the carrier of TOP-REAL 2 by Th46;
  reconsider f=(sn-FanMorphW) as map of TOP-REAL 2,TOP-REAL 2;
  consider h being map of (TOP-REAL 2),(TOP-REAL 2)
  such that A3: h=(sn-FanMorphW) & h is continuous by A1,Th44;
  A4: f is one-to-one by A1,Th45;
  A5: rng f=[#](TOP-REAL 2) by A2,PRE_TOPC:12;
   for p2 being Point of TOP-REAL 2
  ex K being non empty compact Subset of TOP-REAL 2 st K = f.:K &
 (ex V2 being Subset of TOP-REAL 2 st
  p2 in V2 & V2 is open & V2 c= K & f.p2 in V2) by A1,Th47;
  then f is_homeomorphism by A3,A4,A5,Th8;
 hence thesis;
end;

Lm12:now let q be Point of TOP-REAL 2, sn,t be Real;
  assume ((-(t/|.q.|-sn))/(1-sn))^2<1^2;
  then 1-((-(t/|.q.|-sn))/(1-sn))^2>0 by SQUARE_1:11,59;
  then sqrt(1-((-(t/|.q.|-sn))/(1-sn))^2)>0 by SQUARE_1:93;
  then sqrt(1-(-(t/|.q.|-sn))^2/(1-sn)^2)> 0 by SQUARE_1:69;
  then sqrt(1-(t/|.q.|-sn)^2/(1-sn)^2)> 0 by SQUARE_1:61;
  then sqrt(1-((t/|.q.|-sn)/(1-sn))^2)> 0 by SQUARE_1:69;
  hence -sqrt(1-((t/|.q.|-sn)/(1-sn))^2)< -0 by REAL_1:50;
end;

theorem Th49:
 for sn being Real,q being Point of TOP-REAL 2 st -1<sn & sn<1 & q`1<0
 & q`2/|.q.|>=sn holds (for p being Point of TOP-REAL 2 st
 p=(sn-FanMorphW).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;
  let p be Point of TOP-REAL 2;
    assume p=(sn-FanMorphW).q;
    then A2: p=|[ |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1-sn))]| by A1,Th23;
      set qz=p;
      A3: |.q.|>0 by A1,Lm1,JGRAPH_2:11;
      then A4: (|.q.|)^2>0 by SQUARE_1:74;
     A5: 1-sn>0 by A1,Th3;
        A6: qz`1= |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2))&
        qz`2= |.q.|* ((q`2/|.q.|-sn)/(1-sn)) by A2,EUCLID:56;
        A7: (q`2/|.q.|-sn)>= 0 by A1,SQUARE_1:12;
        A8: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
          0<(q`1)^2 by A1,SQUARE_1:74;
        then 0+(q`2)^2<(q`1)^2+(q`2)^2 by REAL_1:67;
        then (q`2)^2/(|.q.|)^2 < (|.q.|)^2/(|.q.|)^2 by A4,A8,REAL_1:73;
        then (q`2)^2/(|.q.|)^2 < 1 by A4,XCMPLX_1:60;
        then ((q`2)/|.q.|)^2 < 1 by SQUARE_1:69;
        then 1>q`2/|.q.| by Th5;
        then A9: 1-sn>q`2/|.q.|-sn by REAL_1:54;
        A10: -(1-sn)< -0 by A5,REAL_1:50;
           -(1-sn)< -( q`2/|.q.|-sn) by A9,REAL_1:50;
         then (-(1-sn))/(1-sn)<(-( q`2/|.q.|-sn))/(1-sn)
                        by A5,REAL_1:73;
        then A11: -1<(-( q`2/|.q.|-sn))/(1-sn) by A5,XCMPLX_1:198;
         --(1-sn)> -(q`2/|.q.|-sn) by A7,A10,REAL_1:50;
       then (-(q`2/|.q.|-sn))/(1-sn)<1 by A5,REAL_2:118;
       then ((-(q`2/|.q.|-sn))/(1-sn))^2<1^2 by A11,JGRAPH_2:8;
       then A12: -sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)< -0 by Lm12;
         ((q`2/|.q.|-sn)/(1-sn))>=0 by A5,A7,SQUARE_1:27;
   hence p`1<0 & p`2>=0 by A3,A6,A12,SQUARE_1:19,24;
 end;

theorem Th50:
 for sn being Real,q being Point of TOP-REAL 2 st -1<sn & sn<1 & q`1<0
 & q`2/|.q.|<sn holds (for p being Point of TOP-REAL 2 st
 p=(sn-FanMorphW).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;
  let p be Point of TOP-REAL 2;
    assume p=(sn-FanMorphW).q;
    then A2: p=|[ |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1+sn))]| by A1,Th24;
      set qz=p;
      A3: |.q.|>0 by A1,Lm1,JGRAPH_2:11;
      then A4: (|.q.|)^2>0 by SQUARE_1:74;
      A5: (q`2/|.q.|-sn)< 0 by A1,REAL_2:105;
     A6: 1+sn>0 by A1,Th3;
        A7: qz`1= |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2))&
        qz`2= |.q.|* ((q`2/|.q.|-sn)/(1+sn)) by A2,EUCLID:56;
        A8: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
          0<(q`1)^2 by A1,SQUARE_1:74;
        then 0+(q`2)^2<(q`1)^2+(q`2)^2 by REAL_1:67;
        then (q`2)^2/(|.q.|)^2 < (|.q.|)^2/(|.q.|)^2 by A4,A8,REAL_1:73;
        then (q`2)^2/(|.q.|)^2 < 1 by A4,XCMPLX_1:60;
        then ((q`2)/|.q.|)^2 < 1 by SQUARE_1:69;
        then -1<q`2/|.q.| by Th5;
        then -1-sn<q`2/|.q.|-sn by REAL_1:54;
        then -(1+sn)<q`2/|.q.|-sn by XCMPLX_1:161;
        then A9: --(1+sn)> -(q`2/|.q.|-sn) by REAL_1:50;
        A10: -( q`2/|.q.|-sn)>0 by A5,REAL_1:66;
          -(1+sn)< -0 by A6,REAL_1:50;
         then -(1+sn)< -( q`2/|.q.|-sn) by A10,AXIOMS:22;
         then (-(1+sn))/(1+sn)<(-( q`2/|.q.|-sn))/(1+sn) by A6,REAL_1:73;
        then A11: -1<(-( q`2/|.q.|-sn))/(1+sn) by A6,XCMPLX_1:198;
         (-(q`2/|.q.|-sn))/(1+sn)<1 by A6,A9,REAL_2:118;
       then ((-(q`2/|.q.|-sn))/(1+sn))^2<1^2 by A11,JGRAPH_2:8;
       then 1-((-(q`2/|.q.|-sn))/(1+sn))^2>0 by SQUARE_1:11,59;
       then sqrt(1-((-(q`2/|.q.|-sn))/(1+sn))^2)>0 by SQUARE_1:93;
       then sqrt(1-(-(q`2/|.q.|-sn))^2/(1+sn)^2)> 0 by SQUARE_1:69;
       then sqrt(1-(q`2/|.q.|-sn)^2/(1+sn)^2)> 0 by SQUARE_1:61;
       then sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)> 0 by SQUARE_1:69;
       then A12: -sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)< -0 by REAL_1:50;
         ((q`2/|.q.|-sn)/(1+sn))<0 by A5,A6,REAL_2:128;
   hence p`1<0 & p`2<0 by A3,A7,A12,SQUARE_1:24;
 end;

theorem Th51:
 for sn being Real,q1,q2 being Point of TOP-REAL 2 st -1<sn & sn<1 & q1`1<0
 & q1`2/|.q1.|>=sn & q2`1<0 & q2`2/|.q2.|>=sn & q1`2/|.q1.|<q2`2/|.q2.|
 holds (for p1,p2 being Point of TOP-REAL 2 st
 p1=(sn-FanMorphW).q1 & p2=(sn-FanMorphW).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
   & q1`2/|.q1.|>=sn & q2`1<0 & q2`2/|.q2.|>=sn & q1`2/|.q1.|<q2`2/|.q2.|;
  let p1,p2 be Point of TOP-REAL 2;
   assume A2: p1=(sn-FanMorphW).q1 & p2=(sn-FanMorphW).q2;
    then p1=|[ |.q1.|*(-sqrt(1-((q1`2/|.q1.|-sn)/(1-sn))^2)),
          |.q1.|* ((q1`2/|.q1.|-sn)/(1-sn))]| by A1,Th23;
    then A3: p1`1= |.q1.|*(-sqrt(1-((q1`2/|.q1.|-sn)/(1-sn))^2))&
        p1`2= |.q1.|* ((q1`2/|.q1.|-sn)/(1-sn)) by EUCLID:56;
      p2=|[ |.q2.|*(-sqrt(1-((q2`2/|.q2.|-sn)/(1-sn))^2)),
          |.q2.|* ((q2`2/|.q2.|-sn)/(1-sn))]| by A1,A2,Th23;
    then A4: p2`1= |.q2.|*(-sqrt(1-((q2`2/|.q2.|-sn)/(1-sn))^2))&
        p2`2= |.q2.|* ((q2`2/|.q2.|-sn)/(1-sn)) by EUCLID:56;
    A5: |.q1.|>0 & |.q2.|>0 by A1,Lm1,JGRAPH_2:11;
    A6: |.p1.|=|.q1.| & |.p2.|=|.q2.| by A2,Th40;
    then A7: p1`2/|.p1.|= ((q1`2/|.q1.|-sn)/(1-sn)) by A3,A5,XCMPLX_1:90
;
    A8: p2`2/|.p2.|= ((q2`2/|.q2.|-sn)/(1-sn)) by A4,A5,A6,XCMPLX_1:90
;
    A9: q1`2/|.q1.|-sn< q2`2/|.q2.|-sn by A1,REAL_1:54;
      1-sn>0 by A1,Th3;
   hence thesis by A7,A8,A9,REAL_1:73;
 end;

theorem Th52:
 for sn being Real,q1,q2 being Point of TOP-REAL 2 st -1<sn & sn<1 & q1`1<0
 & q1`2/|.q1.|<sn & q2`1<0 & q2`2/|.q2.|<sn & q1`2/|.q1.|<q2`2/|.q2.|
 holds (for p1,p2 being Point of TOP-REAL 2 st
 p1=(sn-FanMorphW).q1 & p2=(sn-FanMorphW).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
   & q1`2/|.q1.|<sn & q2`1<0 & q2`2/|.q2.|<sn & q1`2/|.q1.|<q2`2/|.q2.|;
  let p1,p2 be Point of TOP-REAL 2;
   assume A2: p1=(sn-FanMorphW).q1 & p2=(sn-FanMorphW).q2;
    then p1=|[ |.q1.|*(-sqrt(1-((q1`2/|.q1.|-sn)/(1+sn))^2)),
          |.q1.|* ((q1`2/|.q1.|-sn)/(1+sn))]| by A1,Th24;
    then A3: p1`1= |.q1.|*(-sqrt(1-((q1`2/|.q1.|-sn)/(1+sn))^2))&
        p1`2= |.q1.|* ((q1`2/|.q1.|-sn)/(1+sn)) by EUCLID:56;
      p2=|[ |.q2.|*(-sqrt(1-((q2`2/|.q2.|-sn)/(1+sn))^2)),
          |.q2.|* ((q2`2/|.q2.|-sn)/(1+sn))]| by A1,A2,Th24;
    then A4: p2`1= |.q2.|*(-sqrt(1-((q2`2/|.q2.|-sn)/(1+sn))^2))&
        p2`2= |.q2.|* ((q2`2/|.q2.|-sn)/(1+sn)) by EUCLID:56;
    A5: |.q1.|>0 & |.q2.|>0 by A1,Lm1,JGRAPH_2:11;
    A6: |.p1.|=|.q1.| & |.p2.|=|.q2.| by A2,Th40;
    then A7: p1`2/|.p1.|= ((q1`2/|.q1.|-sn)/(1+sn)) by A3,A5,XCMPLX_1:90
;
    A8: p2`2/|.p2.|= ((q2`2/|.q2.|-sn)/(1+sn)) by A4,A5,A6,XCMPLX_1:90
;
    A9: q1`2/|.q1.|-sn< q2`2/|.q2.|-sn by A1,REAL_1:54;
      1+sn>0 by A1,Th3;
   hence p1`2/|.p1.|<p2`2/|.p2.| by A7,A8,A9,REAL_1:73;
 end;

theorem
   for sn being Real,q1,q2 being Point of TOP-REAL 2 st -1<sn & sn<1 & q1`1<0
 & q2`1<0 & q1`2/|.q1.|<q2`2/|.q2.|
 holds (for p1,p2 being Point of TOP-REAL 2 st
 p1=(sn-FanMorphW).q1 & p2=(sn-FanMorphW).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`2/|.q1.|<q2`2/|.q2.|;
  let p1,p2 be Point of TOP-REAL 2;
   assume A2: p1=(sn-FanMorphW).q1 & p2=(sn-FanMorphW).q2;
      now per cases;
    case q1`2/|.q1.|>=sn & q2`2/|.q2.|>=sn;
     hence p1`2/|.p1.|<p2`2/|.p2.| by A1,A2,Th51;
    case q1`2/|.q1.|>=sn & q2`2/|.q2.|<sn;
     hence p1`2/|.p1.|<p2`2/|.p2.| by A1,AXIOMS:22;
    case A3: q1`2/|.q1.|<sn & q2`2/|.q2.|>=sn;
       then A4: p1`1<0 & p1`2<0 by A1,A2,Th50;
       then A5: |.p1.|>0 by Lm1,JGRAPH_2:11;
       A6: p2`1<0 & p2`2>=0 by A1,A2,A3,Th49;
       then |.p2.|>0 by Lm1,JGRAPH_2:11;
       then p2`2/|.p2.|>=0 by A6,REAL_2:125;
     hence p1`2/|.p1.|<p2`2/|.p2.| by A4,A5,REAL_2:128;
    case q1`2/|.q1.|<sn & q2`2/|.q2.|<sn;
     hence p1`2/|.p1.|<p2`2/|.p2.| by A1,A2,Th52;
    end;
   hence p1`2/|.p1.|<p2`2/|.p2.|;
 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 holds (for p being Point of TOP-REAL 2 st
 p=(sn-FanMorphW).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;
  let p be Point of TOP-REAL 2;
   assume p=(sn-FanMorphW).q;
    then A2: p=|[ |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1-sn))]| by A1,Th23;
    A3: |.q.|>0 by A1,Lm1,JGRAPH_2:11;
     A4: (q`2/|.q.|-sn)=0 by A1,XCMPLX_1:14;
     A5: p`1= |.q.|*(-sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2))&
        p`2= |.q.|* ((q`2/|.q.|-sn)/(1-sn)) by A2,EUCLID:56;
         ((-(q`2/|.q.|-sn))/(1-sn))^2<1^2 by A4,JGRAPH_2:8,XCMPLX_1:198;
       then -sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)< -0 by Lm12;
   hence thesis by A3,A4,A5,SQUARE_1:24;
 end;

theorem
   for sn being real number holds 0.REAL 2=(sn-FanMorphW).(0.REAL 2)
   by Th23,JGRAPH_2:11;

begin :: Fan Morphism for North

definition let s be real number, q be Point of TOP-REAL 2;
  func FanN(s,q) -> Point of TOP-REAL 2 equals
  :Def4: |.q.|*|[(q`1/|.q.|-s)/(1-s),
     sqrt(1-((q`1/|.q.|-s)/(1-s))^2)]| if q`1/|.q.|>=s & q`2>0,
       |.q.|*|[(q`1/|.q.|-s)/(1+s),
     sqrt(1-((q`1/|.q.|-s)/(1+s))^2)]| if q`1/|.q.|<s & q`2>0
   otherwise q;
 correctness;
end;

definition let c be real number;
 func c-FanMorphN -> Function of the carrier of TOP-REAL 2,
 the carrier of TOP-REAL 2 means
  :Def5: for q being Point of TOP-REAL 2 holds
       it.q=FanN(c,q);
 existence
  proof
    deffunc F(Point of TOP-REAL 2)=FanN(c,$1);
    thus ex IT being Function of the carrier of TOP-REAL 2,
 the carrier of TOP-REAL 2 st for q being Point of TOP-REAL 2 holds
       IT.q=F(q) from LambdaD;
  end;
 uniqueness
  proof
deffunc F(Point of TOP-REAL 2)=FanN(c,$1);
    thus for a,b being Function of the carrier of TOP-REAL 2,
 the carrier of TOP-REAL 2 st
 (for q being Point of TOP-REAL 2 holds a.q=F(q)) &
 (for q being Point of TOP-REAL 2 holds b.q=F(q)) holds
 a = b from FuncDefUniq;
  end;
end;

theorem Th56: for cn being real number holds
  (q`1/|.q.|>=cn & q`2>0 implies
    cn-FanMorphN.q= |[ |.q.|* ((q`1/|.q.|-cn)/(1-cn)),
    |.q.|*(sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))]|)&
  (q`2<=0 implies cn-FanMorphN.q=q)
proof let cn be real number;
  A1: cn-FanMorphN.q=FanN(cn,q) by Def5;
   hereby assume q`1/|.q.|>=cn & q`2>0;
    then FanN(cn,q)= |.q.|*|[(q`1/|.q.|-cn)/(1-cn),
         sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2)]| by Def4
          .= |[ |.q.|* ((q`1/|.q.|-cn)/(1-cn)),
          |.q.|*(sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))]| by EUCLID:62;
     hence cn-FanMorphN.q= |[ |.q.|* ((q`1/|.q.|-cn)/(1-cn)),
    |.q.|*(sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))]| by Def5;
   end;
   assume q`2<=0; hence thesis by A1,Def4;
end;

theorem Th57: for cn being Real holds
  (q`1/|.q.|<=cn & q`2>0 implies
    cn-FanMorphN.q= |[ |.q.|*((q`1/|.q.|-cn)/(1+cn)),
       |.q.|*( sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))]|)
proof let cn be Real;
   assume A1: q`1/|.q.|<=cn & q`2>0;
     now per cases by A1,REAL_1:def 5;
   case q`1/|.q.|<cn;
    then FanN(cn,q)= |.q.|*|[(q`1/|.q.|-cn)/(1+cn),
        sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2)]| by A1,Def4
          .= |[ |.q.|* ((q`1/|.q.|-cn)/(1+cn)),
           |.q.|*( sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))]| by EUCLID:62;
    hence cn-FanMorphN.q= |[ |.q.|*((q`1/|.q.|-cn)/(1+cn)),
      |.q.|*( sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))]| by Def5;
   case A2: q`1/|.q.|=cn;
    then A3: q`1/|.q.|-cn=0 by XCMPLX_1:14;
    then (q`1/|.q.|-cn)/(1-cn)=0;
    hence cn-FanMorphN.q= |[ |.q.|*((q`1/|.q.|-cn)/(1+cn)),
        |.q.|*( sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))]| by A1,A2,A3,Th56;
   end;
   hence thesis;
end;

theorem Th58: for cn being Real st -1<cn & cn<1 holds
  (q`1/|.q.|>=cn & q`2>=0 & q<>0.REAL 2 implies
    cn-FanMorphN.q= |[ |.q.|* ((q`1/|.q.|-cn)/(1-cn)),
        |.q.|*( sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))]|) &
  (q`1/|.q.|<=cn & q`2>=0 & q<>0.REAL 2 implies
    cn-FanMorphN.q= |[ |.q.|*((q`1/|.q.|-cn)/(1+cn)),
        |.q.|*( sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))]|)
proof let cn be Real;
 assume A1: -1<cn & cn<1;
     now per cases;
   case A2: q`1/|.q.|>=cn & q`2>=0 & q<>0.REAL 2;
      now per cases;
    case A3: q`2>0;
      then FanN(cn,q)= |.q.|*|[(q`1/|.q.|-cn)/(1-cn),
          sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2)]| by A2,Def4
          .= |[|.q.|* ((q`1/|.q.|-cn)/(1-cn)),
          |.q.|*( sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))]| by EUCLID:62;
     hence thesis by A3,Def5,Th57;
    case A4: q`2<=0;
     then A5: cn-FanMorphN.q=q by Th56;
     A6: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
     A7: |.q.|<>0 by A2,TOPRNS_1:25;
     then A8: |.q.|^2>0 by SQUARE_1:74;
     A9: q`2=0 by A2,A4;
       |.q.|>=0 by TOPRNS_1:26;
     then A10: sqrt((|.q.|)^2)=|.q.| by SQUARE_1:89;
     A11: 1-cn>0 by A1,Th3;
     A12: |.q.|>0 by A2,Lm1;
       (q`1)^2/|.q.|^2=1^2 by A6,A8,A9,SQUARE_1:59,60,XCMPLX_1:60;
      then ((q`1)/|.q.|)^2=1^2 by SQUARE_1:69;
      then A13: sqrt(((q`1)/|.q.|)^2)=1 by SQUARE_1:89;
       now assume q`1<0;
       then q`1/|.q.|<0 by A12,REAL_2:128;
       then -((q`1)/|.q.|)=1 by A13,SQUARE_1:90;
      hence contradiction by A1,A2;
     end;
     then A14: |.q.|=q`1 by A6,A9,A10,SQUARE_1:60,89;
     then 1=q`1/|.q.| by A7,XCMPLX_1:60;
     then (q`1/|.q.|-cn)/(1-cn)=1 by A11,XCMPLX_1:60;
     hence thesis by A1,A5,A7,A9,A14,EUCLID:57,SQUARE_1:59,82,XCMPLX_1:60;
    end;
    hence thesis;
   case A15: q`1/|.q.|<=cn & q`2>=0 & q<>0.REAL 2;
      now per cases;
    case q`2>0;
     hence thesis by Th56,Th57;
    case A16: q`2<=0;
     A17: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
     A18: |.q.|<>0 by A15,TOPRNS_1:25;
     A19: q`2=0 by A15,A16;
     A20: |.q.|>0 by A15,Lm1;
     A21: 1+cn>0 by A1,Th3;
       1-cn>0 by A1,Th3;
     then 1-cn+cn>0+cn by REAL_1:53;
     then 1>cn by XCMPLX_1:27;
     then 1>q`1/|.q.| by A15,AXIOMS:22;
     then (1)*(|.q.|)>q`1/|.q.|*(|.q.|) by A20,REAL_1:70;
     then A22: (|.q.|)>q`1 by A18,XCMPLX_1:88;
     then A23: |.q.|=-q`1 by A17,A19,JGRAPH_3:1,SQUARE_1:60;
     A24: q`1= -(|.q.|) by A17,A19,A22,JGRAPH_3:1,SQUARE_1:60;
     then -1=q`1/|.q.| by A18,XCMPLX_1:198;
     then (q`1/|.q.|-cn)/(1+cn)
     =(-(1+cn))/(1+cn) by XCMPLX_1:161 .=-1 by A21,XCMPLX_1:198;
     then |[ |.q.|* ((q`1/|.q.|-cn)/(1+cn)),
       |.q.|*( sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))]|
      =|[ |.q.|*(-1), |.q.|*( sqrt(1-1))]| by SQUARE_1:59,61
     .=|[ -(|.q.|), |.q.|*(0)]| by SQUARE_1:82,XCMPLX_1:180
     .=q by A19,A23,EUCLID:57;
     hence thesis by A1,A16,A18,A24,Th56,XCMPLX_1:198;
    end;
     hence thesis;
   case q`2<0 or q=0.REAL 2;
    hence thesis;
   end;
  hence thesis;
end;

theorem Th59:
 for cn being Real,K1 being non empty Subset of TOP-REAL 2,
f being map of (TOP-REAL 2)|K1,R^1 st -1<cn & cn<1 &
(for p being Point of (TOP-REAL 2) st p in the carrier of (TOP-REAL 2)|K1
  holds f.p=|.p.|* ((p`1/|.p.|-cn)/(1-cn)))
  & (for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds q`2>=0 & q<>0.REAL 2) holds
f is continuous
proof
 let cn be Real,K1 be non empty Subset of TOP-REAL 2,
f be map of (TOP-REAL 2)|K1,R^1;
 assume A1: -1<cn & cn<1 &
(for p being Point of (TOP-REAL 2) st p in the carrier of (TOP-REAL 2)|K1
  holds f.p=|.p.|* ((p`1/|.p.|-cn)/(1-cn)))
  & (for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds q`2>=0 & q<>0.REAL 2);
  set a=cn, b=(1-cn);
  A2: b>0 by A1,Th3;
  reconsider g2=proj1|K1 as continuous map of (TOP-REAL 2)|K1,R^1 by Lm2;
  reconsider g1=(2 NormF)|K1 as continuous map of (TOP-REAL 2)|K1,R^1
                          by Lm5;
    for q being Point of TOP-REAL 2 st q in
    the carrier of (TOP-REAL 2)|K1 holds q<>0.REAL 2 by A1;
  then for q being Point of (TOP-REAL 2)|K1 holds g1.q<>0 by Lm6;
  then consider g3 being map of (TOP-REAL 2)|K1,R^1 such that
  A3: (for q being Point of
  (TOP-REAL 2)|K1,r1,r2 being Real st g2.q=r1 & g1.q=r2
  holds g3.q=r2*((r1/r2-a)/b)) & g3 is continuous by A2,Th10;
  A4:dom g3=the carrier of (TOP-REAL 2)|K1 by FUNCT_2:def 1;
  then A5:dom f=dom g3 by FUNCT_2:def 1;
    for x being set st x in dom f holds f.x=g3.x
   proof let x be set;assume
     A6: x in dom f;
     then x in K1 by A4,A5,JORDAN1:1;
     then reconsider r=x as Point of (TOP-REAL 2);
     reconsider s=x as Point of (TOP-REAL 2)|K1 by A6,FUNCT_2:def 1;
     A7:f.r=(|.r.|)* ((r`1/|.r.|-cn)/(1-cn)) by A1,A4,A5,A6;
     A8:g2.s=proj1.s by Lm2;
     A9:g1.s=(2 NormF).s by Lm5;
     A10:proj1.r=r`1 by PSCOMP_1:def 28;
       (2 NormF).r=|.r.| by Def1;
    hence f.x=g3.x by A3,A7,A8,A9,A10;
   end;
hence f is continuous by A3,A5,FUNCT_1:9;
end;

theorem Th60:
 for cn being Real,K1 being non empty Subset of TOP-REAL 2,
f being map of (TOP-REAL 2)|K1,R^1 st
  -1<cn & cn<1 &
(for p being Point of (TOP-REAL 2) st p in the carrier of (TOP-REAL 2)|K1
  holds f.p=|.p.|* ((p`1/|.p.|-cn)/(1+cn)))
  & (for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds q`2>=0 & q<>0.REAL 2) holds
f is continuous
proof
 let cn be Real,K1 be non empty Subset of TOP-REAL 2,
f be map of (TOP-REAL 2)|K1,R^1;
 assume A1:  -1<cn & cn<1 &
(for p being Point of (TOP-REAL 2) st p in the carrier of (TOP-REAL 2)|K1
  holds f.p=|.p.|* ((p`1/|.p.|-cn)/(1+cn)))
  & (for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds q`2>=0 & q<>0.REAL 2);
  set a=cn, b=(1+cn);
  A2: 1+cn>0 by A1,Th3;
  reconsider g2=proj1|K1 as continuous map of (TOP-REAL 2)|K1,R^1 by Lm2;
  reconsider g1=(2 NormF)|K1 as continuous map of (TOP-REAL 2)|K1,R^1
                          by Lm5;
    for q being Point of TOP-REAL 2 st q in
    the carrier of (TOP-REAL 2)|K1 holds q<>0.REAL 2 by A1;
  then for q being Point of (TOP-REAL 2)|K1 holds g1.q<>0 by Lm6;
  then consider g3 being map of (TOP-REAL 2)|K1,R^1 such that
  A3: (for q being Point of
  (TOP-REAL 2)|K1,r1,r2 being Real st g2.q=r1 & g1.q=r2
  holds g3.q=r2*((r1/r2-a)/b)) & g3 is continuous by A2,Th10;
  A4:dom g3=the carrier of (TOP-REAL 2)|K1 by FUNCT_2:def 1;
  then A5:dom f=dom g3 by FUNCT_2:def 1;
    for x being set st x in dom f holds f.x=g3.x
   proof let x be set;assume
     A6: x in dom f; then A7: x in dom g3 by A4,FUNCT_2:def 1;
     then x in K1 by A4,JORDAN1:1;
     then reconsider r=x as Point of (TOP-REAL 2);
     reconsider s=x as Point of (TOP-REAL 2)|K1 by A6,FUNCT_2:def 1;
     A8:f.r=(|.r.|)* ((r`1/|.r.|-cn)/(1+cn)) by A1,A4,A7;
     A9:g2.s=proj1.s by Lm2;
     A10:g1.s=(2 NormF).s by Lm5;
     A11:proj1.r=r`1 by PSCOMP_1:def 28;
       (2 NormF).r=|.r.| by Def1;
    hence f.x=g3.x by A3,A8,A9,A10,A11;
   end;
hence f is continuous by A3,A5,FUNCT_1:9;
end;

theorem Th61:
 for cn being Real,K1 being non empty Subset of TOP-REAL 2,
f being map of (TOP-REAL 2)|K1,R^1 st -1<cn & cn<1 &
(for p being Point of (TOP-REAL 2) st p in the carrier of (TOP-REAL 2)|K1
  holds f.p=|.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2)))
  & (for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds
        q`2>=0 & q`1/|.q.|>=cn & q<>0.REAL 2) holds
f is continuous
proof
 let cn be Real,K1 be non empty Subset of TOP-REAL 2,
f be map of (TOP-REAL 2)|K1,R^1;
 assume A1: -1<cn & cn<1 &
(for p being Point of (TOP-REAL 2) st p in the carrier of (TOP-REAL 2)|K1
  holds f.p=|.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2)))
  & (for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds
            q`2>=0 & q`1/|.q.|>=cn & q<>0.REAL 2);
  set a=cn, b=(1-cn);
  A2: b>0 by A1,Th3;
  reconsider g2=proj1|K1 as continuous map of (TOP-REAL 2)|K1,R^1 by Lm2;
  reconsider g1=(2 NormF)|K1 as continuous map of (TOP-REAL 2)|K1,R^1
                          by Lm5;
    for q being Point of TOP-REAL 2 st q in
    the carrier of (TOP-REAL 2)|K1 holds q<>0.REAL 2 by A1;
  then for q being Point of (TOP-REAL 2)|K1 holds g1.q<>0 by Lm6;
  then consider g3 being map of (TOP-REAL 2)|K1,R^1 such that
  A3: (for q being Point of
  (TOP-REAL 2)|K1,r1,r2 being real number st g2.q=r1 & g1.q=r2
  holds g3.q=r2*( sqrt(abs(1-((r1/r2-a)/b)^2))))
          & g3 is continuous by A2,Th15;
  A4:dom g3=the carrier of (TOP-REAL 2)|K1 by FUNCT_2:def 1;
  then A5:dom f=dom g3 by FUNCT_2:def 1;
    for x being set st x in dom f holds f.x=g3.x
   proof let x be set;assume
     A6: x in dom f;
     then x in K1 by A4,A5,JORDAN1:1;
     then reconsider r=x as Point of (TOP-REAL 2);
     reconsider s=x as Point of (TOP-REAL 2)|K1 by A6,FUNCT_2:def 1;
     A7: now assume |.r.|=0; then r=0.REAL 2 by TOPRNS_1:25;
      hence contradiction by A1,A4,A5,A6;
     end;
     A8: |.r.|>=0 by TOPRNS_1:26;
       |.r.|^2=(r`1)^2+(r`2)^2 by JGRAPH_3:10;
     then |.r.|^2 -(r`1)^2 =(r`2)^2 by XCMPLX_1:26;
     then (r`1)^2 -(|.r.|)^2 =-(r`2)^2 by XCMPLX_1:143;
     then A9: ((r`1) -(|.r.|))*((r`1)+|.r.|) =-(r`2)^2 by SQUARE_1:67;
       (r`2)^2>=0 by SQUARE_1:72;
     then -(r`2)^2 <=0 by REAL_1:66;
     then -(|.r.|)<=r`1 & r`1<= |.r.| by A8,A9,JGRAPH_3:4;
     then r`1/|.r.| <= |.r.|/|.r.| by A7,A8,REAL_1:73;
     then r`1/|.r.|<=1 by A7,XCMPLX_1:60;
     then A10: r`1/|.r.|-cn<=(1-cn) by REAL_1:49;
     A11: now assume (1-cn)^2=0;
       then 1-cn+cn=0+cn by SQUARE_1:73;
      hence contradiction by A1,XCMPLX_1:27;
     end;
     A12: 1-cn>0 by A1,Th3;
      A13: (1-cn)^2>=0 by SQUARE_1:72;
       cn<=r`1/(|.r.|) by A1,A4,A5,A6;
     then cn-r`1/|.r.|<=0 by REAL_2:106;
     then -(cn- r`1/|.r.|)>=-(1-cn) by A12,REAL_1:50;
     then -(1-cn)<= r`1/|.r.|-cn by XCMPLX_1:143;
     then (r`1/|.r.|-cn)^2<=(1-cn)^2 by A10,JGRAPH_2:7;
     then (r`1/|.r.|-cn)^2/(1-cn)^2<=(1-cn)^2/(1-cn)^2
                         by A11,A13,REAL_1:73;
     then (r`1/|.r.|-cn)^2/(1-cn)^2<=1 by A11,XCMPLX_1:60;
     then ((r`1/|.r.|-cn)/(1-cn))^2<=1 by SQUARE_1:69;
     then 1-((r`1/|.r.|-cn)/(1-cn))^2>=0 by SQUARE_1:12;
     then abs(1-((r`1/|.r.|-cn)/(1-cn))^2)
          =1-((r`1/|.r.|-cn)/(1-cn))^2 by ABSVALUE:def 1;
     then A14:f.r=(|.r.|)*( sqrt(abs(1-((r`1/|.r.|-cn)/(1-cn))^2))) by A1,A4,A5
,A6;
     A15:g2.s=proj1.s by Lm2;
     A16:g1.s=(2 NormF).s by Lm5;
     A17:proj1.r=r`1 by PSCOMP_1:def 28;
       (2 NormF).r=|.r.| by Def1;
    hence f.x=g3.x by A3,A14,A15,A16,A17;
   end;
hence f is continuous by A3,A5,FUNCT_1:9;
end;

theorem Th62:
 for cn being Real,K1 being non empty Subset of TOP-REAL 2,
f being map of (TOP-REAL 2)|K1,R^1 st -1<cn & cn<1 &
(for p being Point of (TOP-REAL 2) st p in the carrier of (TOP-REAL 2)|K1
  holds f.p=|.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2)))
  & (for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds
        q`2>=0 & q`1/|.q.|<=cn & q<>0.REAL 2) holds
f is continuous
proof
 let cn be Real,K1 be non empty Subset of TOP-REAL 2,
f be map of (TOP-REAL 2)|K1,R^1;
 assume A1: -1<cn & cn<1 &
(for p being Point of (TOP-REAL 2) st p in the carrier of (TOP-REAL 2)|K1
  holds f.p=|.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2)))
  & (for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds
            q`2>=0 & q`1/|.q.|<=cn & q<>0.REAL 2);
  set a=cn, b=(1+cn);
  A2: 1+cn>0 by A1,Th3;
  reconsider g2=proj1|K1 as continuous map of (TOP-REAL 2)|K1,R^1 by Lm2;
  reconsider g1=(2 NormF)|K1 as continuous map of (TOP-REAL 2)|K1,R^1
                          by Lm5;
    for q being Point of TOP-REAL 2 st q in
    the carrier of (TOP-REAL 2)|K1 holds q<>0.REAL 2 by A1;
  then for q being Point of (TOP-REAL 2)|K1 holds g1.q<>0 by Lm6;
  then consider g3 being map of (TOP-REAL 2)|K1,R^1 such that
  A3: (for q being Point of
  (TOP-REAL 2)|K1,r1,r2 being real number st g2.q=r1 & g1.q=r2
  holds g3.q=r2*( sqrt(abs(1-((r1/r2-a)/b)^2))))
          & g3 is continuous by A2,Th15;
  A4:dom g3=the carrier of (TOP-REAL 2)|K1 by FUNCT_2:def 1;
  then A5:dom f=dom g3 by FUNCT_2:def 1;
    for x being set st x in dom f holds f.x=g3.x
   proof let x be set;assume
     A6:x in dom f;
     then x in K1 by A4,A5,JORDAN1:1;
     then reconsider r=x as Point of (TOP-REAL 2);
     reconsider s=x as Point of (TOP-REAL 2)|K1 by A6,FUNCT_2:def 1;
     A7: now assume |.r.|=0; then r=0.REAL 2 by TOPRNS_1:25;
      hence contradiction by A1,A4,A5,A6;
     end;
     A8: |.r.|>=0 by TOPRNS_1:26;
       |.r.|^2=(r`1)^2+(r`2)^2 by JGRAPH_3:10;
     then |.r.|^2 -(r`1)^2 =(r`2)^2 by XCMPLX_1:26;
     then (r`1)^2 -(|.r.|)^2 =-(r`2)^2 by XCMPLX_1:143;
     then A9: ((r`1) -(|.r.|))*((r`1)+|.r.|) =-(r`2)^2 by SQUARE_1:67;
       (r`2)^2>=0 by SQUARE_1:72;
     then -(r`2)^2 <=0 by REAL_1:66;
     then -(|.r.|)<=r`1 & r`1<= |.r.| by A8,A9,JGRAPH_3:4;
     then r`1/|.r.| >= (-(|.r.|))/|.r.| by A7,A8,REAL_1:73;
     then r`1/|.r.|>= -1 by A7,XCMPLX_1:198;
     then r`1/|.r.|-cn>=-1-cn by REAL_1:49;
     then A10: r`1/|.r.|-cn>=-(1+cn) by XCMPLX_1:161;
     A11: (1+cn)^2>0 by A2,SQUARE_1:74;
       cn>=r`1/(|.r.|) by A1,A4,A5,A6;
     then cn-r`1/|.r.|>=0 by SQUARE_1:12;
     then -(cn-r`1/|.r.|)<=-0 by REAL_1:50;
     then r`1/|.r.|-cn<=-0 by XCMPLX_1:143;
     then r`1/|.r.|-cn<=1+cn by A2,AXIOMS:22;
     then (r`1/|.r.|-cn)^2<=(1+cn)^2 by A10,JGRAPH_2:7;
     then (r`1/|.r.|-cn)^2/(1+cn)^2<=(1+cn)^2/(1+cn)^2 by A11,REAL_1:73;
     then (r`1/|.r.|-cn)^2/(1+cn)^2<=1 by A11,XCMPLX_1:60;
     then ((r`1/|.r.|-cn)/(1+cn))^2<=1 by SQUARE_1:69;
     then 1-((r`1/|.r.|-cn)/(1+cn))^2>=0 by SQUARE_1:12;
     then abs(1-((r`1/|.r.|-cn)/(1+cn))^2)
          =1-((r`1/|.r.|-cn)/(1+cn))^2 by ABSVALUE:def 1;
     then A12:f.r=(|.r.|)*( sqrt(abs(1-((r`1/|.r.|-cn)/(1+cn))^2))) by A1,A4,A5
,A6;
     A13:g2.s=proj1.s by Lm2;
     A14:g1.s=(2 NormF).s by Lm5;
     A15:proj1.r=r`1 by PSCOMP_1:def 28;
       (2 NormF).r=|.r.| by Def1;
    hence f.x=g3.x by A3,A12,A13,A14,A15;
   end;
hence f is continuous by A3,A5,FUNCT_1:9;
end;

theorem Th63: for cn being Real,
K0,B0 being Subset of TOP-REAL 2,f being
map of (TOP-REAL 2)|K0,(TOP-REAL 2)|B0
st -1<cn & cn<1 & f=(cn-FanMorphN)|K0 &
B0={q where q is Point of TOP-REAL 2: q`2>=0 & q<>0.REAL 2}
& K0={p: p`1/|.p.|>=cn & p`2>=0 & p<>0.REAL 2}
holds f is continuous
proof let cn be Real,K0,B0 be Subset of TOP-REAL 2,f be
map of (TOP-REAL 2)|K0,(TOP-REAL 2)|B0;
assume A1: -1<cn & cn<1 & f=(cn-FanMorphN)|K0 &
B0={q where q is Point of TOP-REAL 2: q`2>=0 & q<>0.REAL 2} &
K0={p: p`1/|.p.|>=cn & p`2>=0 & p<>0.REAL 2};
 set sn=sqrt(1-cn^2);
 set p0=|[cn,sn]|;
 A2: p0`1=cn by EUCLID:56;
 A3: p0`2=sn by EUCLID:56;
     cn^2<1^2 by A1,JGRAPH_2:8;
 then A4: 1-cn^2>0 by SQUARE_1:11,59;
 then A5: p0`2>0 by A3,SQUARE_1:93;
 A6: |.p0.|=sqrt((sn)^2+cn^2) by A2,A3,JGRAPH_3:10;
   sn^2=1-cn^2 by A4,SQUARE_1:def 4;
 then |.p0.|=1 by A6,SQUARE_1:83,XCMPLX_1:27;
 then p0`1/|.p0.|=cn by EUCLID:56;
 then A7: p0 in K0 by A1,A5,JGRAPH_2:11;
then reconsider K1=K0 as non empty Subset of TOP-REAL 2;
  A8: K0 c= B0
   proof let x be set;assume x in K0;
      then consider p8 being Point of TOP-REAL 2 such that
      A9: x=p8 & (p8`1/|.p8.|>=cn & p8`2>=0 & p8<>0.REAL 2) by A1;
    thus x in B0 by A1,A9;
   end;
A10:dom ((proj1)*((cn-FanMorphN)|K1)) c= dom ((cn-FanMorphN)|K1)
                                  by RELAT_1:44;
 dom ((cn-FanMorphN)|K1) c= dom ((proj1)*((cn-FanMorphN)|K1))
proof let x be set;assume A11:x in dom ((cn-FanMorphN)|K1);
   then A12:x in dom (cn-FanMorphN) /\ K1 by FUNCT_1:68;
   A13:((cn-FanMorphN)|K1).x=(cn-FanMorphN).x by A11,FUNCT_1:68;
   A14: dom proj1 = (the carrier of TOP-REAL 2) by FUNCT_2:def 1;
     x in dom (cn-FanMorphN) by A12,XBOOLE_0:def 3;
   then (cn-FanMorphN).x in rng (cn-FanMorphN) by FUNCT_1:12;
  hence x in dom ((proj1)*((cn-FanMorphN)|K1)) by A11,A13,A14,FUNCT_1:21;
end;
then A15:dom ((proj1)*((cn-FanMorphN)|K1))
=dom ((cn-FanMorphN)|K1) by A10,XBOOLE_0:def 10
.=dom (cn-FanMorphN) /\ K1 by FUNCT_1:68
.=(the carrier of TOP-REAL 2)/\ K1 by FUNCT_2:def 1
.=K1 by XBOOLE_1:28
.=the carrier of (TOP-REAL 2)|K1 by JORDAN1:1;
A16:rng ((proj1)*((cn-FanMorphN)|K1)) c= rng (proj1) by RELAT_1:45;
  rng proj1 c= the carrier of R^1 by JGRAPH_2:16,RELSET_1:12;
then rng ((proj1)*((cn-FanMorphN)|K1)) c= the carrier of R^1 by A16,XBOOLE_1:1
;
then (proj1)*((cn-FanMorphN)|K1) is Function of the carrier of (TOP-REAL 2)|K1
,
the carrier of R^1 by A15,FUNCT_2:4;
then reconsider g2=(proj1)*((cn-FanMorphN)|K1) as map of (TOP-REAL 2)|K1,R^1
                                ;
A17:dom ((proj2)*((cn-FanMorphN)|K1))
                     c= dom ((cn-FanMorphN)|K1) by RELAT_1:44;
   dom ((cn-FanMorphN)|K1) c= dom ((proj2)*((cn-FanMorphN)|K1))
proof let x be set;assume A18:x in dom ((cn-FanMorphN)|K1);
   then A19:x in dom (cn-FanMorphN) /\ K1 by FUNCT_1:68;
   A20:((cn-FanMorphN)|K1).x=(cn-FanMorphN).x by A18,FUNCT_1:68;
   A21: dom proj2 = (the carrier of TOP-REAL 2) by FUNCT_2:def 1;
     x in dom (cn-FanMorphN) by A19,XBOOLE_0:def 3;
   then (cn-FanMorphN).x in rng (cn-FanMorphN) by FUNCT_1:12;
  hence x in dom ((proj2)*((cn-FanMorphN)|K1)) by A18,A20,A21,FUNCT_1:21;
end;
then A22:dom ((proj2)*((cn-FanMorphN)|K1))
=dom ((cn-FanMorphN)|K1) by A17,XBOOLE_0:def 10
.=dom (cn-FanMorphN) /\ K1 by FUNCT_1:68
.=((the carrier of TOP-REAL 2))/\ K1 by FUNCT_2:def 1
.=K1 by XBOOLE_1:28
.=the carrier of (TOP-REAL 2)|K1 by JORDAN1:1;
A23:rng ((proj2)*((cn-FanMorphN)|K1)) c= rng (proj2) by RELAT_1:45;
  rng proj2 c= the carrier of R^1 by JGRAPH_2:17,RELSET_1:12;
then rng ((proj2)*((cn-FanMorphN)|K1)) c= the carrier of R^1 by A23,XBOOLE_1:1
;
then (proj2)*((cn-FanMorphN)|K1) is Function of the carrier of (TOP-REAL 2)|K1
,
the carrier of R^1 by A22,FUNCT_2:4;
then reconsider g1=(proj2)*((cn-FanMorphN)|K1) as map of (TOP-REAL 2)|K1,R^1
                                ;
A24: for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds q`2>=0 & q<>0.REAL 2
   proof let q be Point of TOP-REAL 2;
    assume A25:q in the carrier of (TOP-REAL 2)|K1;
        the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
     then consider p3 being Point of TOP-REAL 2 such that
     A26: q=p3 & (p3`1/|.p3.|>=cn & p3`2>=0 & p3<>0.REAL 2) by A1,A25;
    thus q`2>=0 & q<>0.REAL 2 by A26;
   end;
    for p being Point of TOP-REAL 2 st p in the carrier of (TOP-REAL 2)|K1
  holds g2.p=|.p.|* ((p`1/|.p.|-cn)/(1-cn))
   proof let p be Point of TOP-REAL 2;
    assume A27: p in the carrier of (TOP-REAL 2)|K1;
     A28:dom ((cn-FanMorphN)|K1)=dom (cn-FanMorphN) /\ K1 by FUNCT_1:68
     .=((the carrier of TOP-REAL 2))/\ K1 by FUNCT_2:def 1
     .=K1 by XBOOLE_1:28;
     A29:the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
     then consider p3 being Point of TOP-REAL 2 such that
     A30: p=p3 & (p3`1/|.p3.|>=cn & p3`2>=0 & p3<>0.REAL 2) by A1,A27;
     A31:(cn-FanMorphN).p
     =|[ |.p.|* ((p`1/|.p.|-cn)/(1-cn)),
        |.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2))]| by A1,A30,Th58;
      ((cn-FanMorphN)|K1).p=(cn-FanMorphN).p by A27,A29,FUNCT_1:72;
     then g2.p=(proj1).(|[ |.p.|* ((p`1/|.p.|-cn)/(1-cn)),
        |.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2))]|)
                             by A27,A28,A29,A31,FUNCT_1:23
      .=(|[ |.p.|* ((p`1/|.p.|-cn)/(1-cn)),
        |.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2))]|)`1 by PSCOMP_1:def 28
      .=|.p.|* ((p`1/|.p.|-cn)/(1-cn)) by EUCLID:56;
    hence g2.p=|.p.|* ((p`1/|.p.|-cn)/(1-cn));
   end;
then consider f2 being map of (TOP-REAL 2)|K1,R^1 such that
  A32:for p being Point of TOP-REAL 2 st p in the carrier of (TOP-REAL 2)|K1
  holds f2.p=|.p.|* ((p`1/|.p.|-cn)/(1-cn));
A33:f2 is continuous by A1,A24,A32,Th59;
    for p being Point of TOP-REAL 2 st p in the carrier of (TOP-REAL 2)|K1
  holds g1.p=|.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2))
   proof let p be Point of TOP-REAL 2;
    assume A34: p in the carrier of (TOP-REAL 2)|K1;
     A35:dom ((cn-FanMorphN)|K1)=dom (cn-FanMorphN) /\ K1 by FUNCT_1:68
     .=((the carrier of TOP-REAL 2))/\ K1 by FUNCT_2:def 1
     .=K1 by XBOOLE_1:28;
     A36:the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
     then consider p3 being Point of TOP-REAL 2 such that
     A37: p=p3 & (p3`1/|.p3.|>=cn & p3`2>=0 & p3<>0.REAL 2) by A1,A34;
     A38:(cn-FanMorphN).p=|[ |.p.|* ((p`1/|.p.|-cn)/(1-cn)),
       |.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2))]| by A1,A37,Th58;
      ((cn-FanMorphN)|K1).p=(cn-FanMorphN).p by A34,A36,FUNCT_1:72;
     then g1.p=(proj2).(|[ |.p.|* ((p`1/|.p.|-cn)/(1-cn)),
          |.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2))]|)
                             by A34,A35,A36,A38,FUNCT_1:23
        .=(|[ |.p.|* ((p`1/|.p.|-cn)/(1-cn)),
          |.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2))]|)`2 by PSCOMP_1:def 29
        .= |.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2)) by EUCLID:56;
    hence g1.p=|.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2));
   end;
then consider f1 being map of (TOP-REAL 2)|K1,R^1 such that
  A39:for p being Point of TOP-REAL 2 st p in the carrier of (TOP-REAL 2)|K1
  holds f1.p=|.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2));
    for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds q`2>=0 & q`1/|.q.|>=cn & q<>0.REAL 2
   proof let q be Point of TOP-REAL 2;
    assume A40:q in the carrier of (TOP-REAL 2)|K1;
        the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
     then consider p3 being Point of TOP-REAL 2 such that
     A41: q=p3 & (p3`1/|.p3.|>=cn & p3`2>=0 & p3<>0.REAL 2) by A1,A40;
    thus thesis by A41;
   end;
then A42:f1 is continuous by A1,A39,Th61;
  for x,y,s,r being real number st |[x,y]| in K1 &
  s=f2.(|[x,y]|) & r=f1.(|[x,y]|) holds f.(|[x,y]|)=|[s,r]|
  proof let x,y,s,r be real number;
   assume A43: |[x,y]| in K1 & s=f2.(|[x,y]|) & r=f1.(|[x,y]|);
     set p99=|[x,y]|;
     A44:the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
     consider p3 being Point of TOP-REAL 2 such that
     A45: p99=p3 & (p3`1/|.p3.|>=cn & p3`2>=0 & p3<>0.REAL 2) by A1,A43;
     A46:f1.p99=|.p99.|*( sqrt(1-((p99`1/|.p99.|-cn)/(1-cn))^2))
                          by A39,A43,A44;
      ((cn-FanMorphN)|K0).(|[x,y]|)=((cn-FanMorphN)).(|[x,y]|) by A43,FUNCT_1:
72
    .= |[ |.p99.|* ((p99`1/|.p99.|-cn)/(1-cn)),
      |.p99.|*( sqrt(1-((p99`1/|.p99.|-cn)/(1-cn))^2))]| by A1,A45,Th58
    .=|[s,r]| by A32,A43,A44,A46;
   hence f.(|[x,y]|)=|[s,r]| by A1;
  end;
hence f is continuous by A7,A8,A33,A42,JGRAPH_2:45;
end;

theorem Th64: for cn being Real,
K0,B0 being Subset of TOP-REAL 2,f being
map of (TOP-REAL 2)|K0,(TOP-REAL 2)|B0
st -1<cn & cn<1 & f=(cn-FanMorphN)|K0 &
B0={q where q is Point of TOP-REAL 2: q`2>=0 & q<>0.REAL 2}
& K0={p: p`1/|.p.|<=cn & p`2>=0 & p<>0.REAL 2}
holds f is continuous
proof let cn be Real,K0,B0 be Subset of TOP-REAL 2,f be
map of (TOP-REAL 2)|K0,(TOP-REAL 2)|B0;
assume A1: -1<cn & cn<1 & f=(cn-FanMorphN)|K0 &
B0={q where q is Point of TOP-REAL 2: q`2>=0 & q<>0.REAL 2} &
K0={p: p`1/|.p.|<=cn & p`2>=0 & p<>0.REAL 2};
 set sn=sqrt(1-cn^2);
 set p0=|[cn,sn]|;
 A2: p0`1=cn by EUCLID:56;
 A3: p0`2=sn by EUCLID:56;
     cn^2<1^2 by A1,JGRAPH_2:8;
 then A4: 1-cn^2>0 by SQUARE_1:11,59;
 then A5: p0`2>0 by A3,SQUARE_1:93;
 A6: |.p0.|=sqrt((sn)^2+cn^2) by A2,A3,JGRAPH_3:10;
   sn^2=1-cn^2 by A4,SQUARE_1:def 4;
 then |.p0.|=1 by A6,SQUARE_1:83,XCMPLX_1:27;
 then p0`1/|.p0.|=cn by EUCLID:56;
 then A7: p0 in K0 by A1,A5,JGRAPH_2:11;
 then reconsider K1=K0 as non empty Subset of TOP-REAL 2;
  A8: K0 c= B0
   proof let x be set;assume x in K0;
      then consider p8 being Point of TOP-REAL 2 such that
      A9: x=p8 & (
      p8`1/|.p8.|<=cn & p8`2>=0 & p8<>0.REAL 2) by A1;
    thus x in B0 by A1,A9;
   end;
A10:dom ((proj1)*((cn-FanMorphN)|K1)) c= dom ((cn-FanMorphN)|K1)
                                  by RELAT_1:44;
 dom ((cn-FanMorphN)|K1) c= dom ((proj1)*((cn-FanMorphN)|K1))
proof let x be set;assume A11:x in dom ((cn-FanMorphN)|K1);
   then A12:x in dom (cn-FanMorphN) /\ K1 by FUNCT_1:68;
   A13:((cn-FanMorphN)|K1).x=(cn-FanMorphN).x by A11,FUNCT_1:68;
   A14: dom proj1 = (the carrier of TOP-REAL 2) by FUNCT_2:def 1;
     x in dom (cn-FanMorphN) by A12,XBOOLE_0:def 3;
   then (cn-FanMorphN).x in rng (cn-FanMorphN) by FUNCT_1:12;
  hence x in dom ((proj1)*((cn-FanMorphN)|K1)) by A11,A13,A14,FUNCT_1:21;
end;
then A15:dom ((proj1)*((cn-FanMorphN)|K1))
=dom ((cn-FanMorphN)|K1) by A10,XBOOLE_0:def 10
.=dom (cn-FanMorphN) /\ K1 by FUNCT_1:68
.=(the carrier of TOP-REAL 2)/\ K1 by FUNCT_2:def 1
.=K1 by XBOOLE_1:28
.=the carrier of (TOP-REAL 2)|K1 by JORDAN1:1;
A16:rng ((proj1)*((cn-FanMorphN)|K1)) c= rng (proj1) by RELAT_1:45;
  rng proj1 c= the carrier of R^1 by JGRAPH_2:16,RELSET_1:12;
then rng ((proj1)*((cn-FanMorphN)|K1)) c= the carrier of R^1 by A16,XBOOLE_1:1
;
then (proj1)*((cn-FanMorphN)|K1) is Function of the carrier of (TOP-REAL 2)|K1
,
the carrier of R^1 by A15,FUNCT_2:4;
then reconsider g2=(proj1)*((cn-FanMorphN)|K1) as map of (TOP-REAL 2)|K1,R^1
                                ;
A17:dom ((proj2)*((cn-FanMorphN)|K1))
                     c= dom ((cn-FanMorphN)|K1) by RELAT_1:44;
 dom ((cn-FanMorphN)|K1) c= dom ((proj2)*((cn-FanMorphN)|K1))
proof let x be set;assume A18:x in dom ((cn-FanMorphN)|K1);
   then A19:x in dom (cn-FanMorphN) /\ K1 by FUNCT_1:68;
   A20:((cn-FanMorphN)|K1).x=(cn-FanMorphN).x by A18,FUNCT_1:68;
   A21: dom proj2 = (the carrier of TOP-REAL 2) by FUNCT_2:def 1;
     x in dom (cn-FanMorphN) by A19,XBOOLE_0:def 3;
   then (cn-FanMorphN).x in rng (cn-FanMorphN) by FUNCT_1:12;
  hence x in dom ((proj2)*((cn-FanMorphN)|K1)) by A18,A20,A21,FUNCT_1:21;
end;
then A22:dom ((proj2)*((cn-FanMorphN)|K1))
=dom ((cn-FanMorphN)|K1) by A17,XBOOLE_0:def 10
.=dom (cn-FanMorphN) /\ K1 by FUNCT_1:68
.=((the carrier of TOP-REAL 2))/\ K1 by FUNCT_2:def 1
.=K1 by XBOOLE_1:28
.=the carrier of (TOP-REAL 2)|K1 by JORDAN1:1;
A23:rng ((proj2)*((cn-FanMorphN)|K1)) c= rng (proj2) by RELAT_1:45;
  rng proj2 c= the carrier of R^1 by JGRAPH_2:17,RELSET_1:12;
then rng ((proj2)*((cn-FanMorphN)|K1)) c= the carrier of R^1 by A23,XBOOLE_1:1
;
then (proj2)*((cn-FanMorphN)|K1) is Function of the carrier of (TOP-REAL 2)|K1
,
the carrier of R^1 by A22,FUNCT_2:4;
then reconsider g1=(proj2)*((cn-FanMorphN)|K1) as map of (TOP-REAL 2)|K1,R^1
                                ;
A24: for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds q`2>=0 & q<>0.REAL 2
   proof let q be Point of TOP-REAL 2;
    assume A25:q in the carrier of (TOP-REAL 2)|K1;
        the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
     then consider p3 being Point of TOP-REAL 2 such that
     A26: q=p3 & (p3`1/|.p3.|<=cn & p3`2>=0 & p3<>0.REAL 2) by A1,A25;
    thus q`2>=0 & q<>0.REAL 2 by A26;
   end;
    for p being Point of TOP-REAL 2 st p in the carrier of (TOP-REAL 2)|K1
  holds g2.p=|.p.|* ((p`1/|.p.|-cn)/(1+cn))
   proof let p be Point of TOP-REAL 2;
    assume A27: p in the carrier of (TOP-REAL 2)|K1;
     A28:dom ((cn-FanMorphN)|K1)=dom (cn-FanMorphN) /\ K1 by FUNCT_1:68
     .=((the carrier of TOP-REAL 2))/\ K1 by FUNCT_2:def 1
     .=K1 by XBOOLE_1:28;
     A29:the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
     then consider p3 being Point of TOP-REAL 2 such that
     A30: p=p3 & (p3`1/|.p3.|<=cn & p3`2>=0 & p3<>0.REAL 2) by A1,A27;
     A31:(cn-FanMorphN).p
     =|[ |.p.|* ((p`1/|.p.|-cn)/(1+cn)),
       |.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2))]| by A1,A30,Th58;
      ((cn-FanMorphN)|K1).p=(cn-FanMorphN).p by A27,A29,FUNCT_1:72;
     then g2.p=(proj1).(|[ |.p.|* ((p`1/|.p.|-cn)/(1+cn)),
        |.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2))]|)
                             by A27,A28,A29,A31,FUNCT_1:23
      .=(|[ |.p.|* ((p`1/|.p.|-cn)/(1+cn)),
       |.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2))]|)`1 by PSCOMP_1:def 28
      .=|.p.|* ((p`1/|.p.|-cn)/(1+cn)) by EUCLID:56;
    hence g2.p=|.p.|* ((p`1/|.p.|-cn)/(1+cn));
   end;
then consider f2 being map of (TOP-REAL 2)|K1,R^1 such that
  A32:for p being Point of TOP-REAL 2 st p in the carrier of (TOP-REAL 2)|K1
  holds f2.p=|.p.|* ((p`1/|.p.|-cn)/(1+cn));
A33:f2 is continuous by A1,A24,A32,Th60;
    for p being Point of TOP-REAL 2 st p in the carrier of (TOP-REAL 2)|K1
  holds g1.p=|.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2))
   proof let p be Point of TOP-REAL 2;
    assume A34: p in the carrier of (TOP-REAL 2)|K1;
     A35:dom ((cn-FanMorphN)|K1)=dom (cn-FanMorphN) /\ K1 by FUNCT_1:68
     .=((the carrier of TOP-REAL 2))/\ K1 by FUNCT_2:def 1
     .=K1 by XBOOLE_1:28;
     A36:the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
     then consider p3 being Point of TOP-REAL 2 such that
     A37: p=p3 & (p3`1/|.p3.|<=cn & p3`2>=0 & p3<>0.REAL 2) by A1,A34;
     A38:(cn-FanMorphN).p=|[ |.p.|* ((p`1/|.p.|-cn)/(1+cn)),
       |.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2))]| by A1,A37,Th58;
      ((cn-FanMorphN)|K1).p=(cn-FanMorphN).p by A34,A36,FUNCT_1:72;
     then g1.p=(proj2).(|[ |.p.|* ((p`1/|.p.|-cn)/(1+cn)),
          |.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2))]|)
                             by A34,A35,A36,A38,FUNCT_1:23
        .=(|[ |.p.|* ((p`1/|.p.|-cn)/(1+cn)),
          |.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2))]|)`2 by PSCOMP_1:def 29
        .= |.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2)) by EUCLID:56;
    hence g1.p=|.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2));
   end;
then consider f1 being map of (TOP-REAL 2)|K1,R^1 such that
  A39:for p being Point of TOP-REAL 2 st p in the carrier of (TOP-REAL 2)|K1
  holds f1.p=|.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2));
    for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds q`2>=0 & q`1/|.q.|<=cn & q<>0.REAL 2
   proof let q be Point of TOP-REAL 2;
    assume A40:q in the carrier of (TOP-REAL 2)|K1;
        the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
     then consider p3 being Point of TOP-REAL 2 such that
     A41: q=p3 & (p3`1/|.p3.|<=cn & p3`2>=0 & p3<>0.REAL 2) by A1,A40;
    thus thesis by A41;
   end;
then A42:f1 is continuous by A1,A39,Th62;
  for x,y,s,r being real number st |[x,y]| in K1 &
  s=f2.(|[x,y]|) & r=f1.(|[x,y]|) holds f.(|[x,y]|)=|[s,r]|
  proof let x,y,s,r be real number;
   assume A43: |[x,y]| in K1 & s=f2.(|[x,y]|) & r=f1.(|[x,y]|);
     set p99=|[x,y]|;
     A44:the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
     consider p3 being Point of TOP-REAL 2 such that
     A45: p99=p3 & (p3`1/|.p3.|<=cn & p3`2>=0 & p3<>0.REAL 2) by A1,A43;
     A46:f1.p99=|.p99.|*( sqrt(1-((p99`1/|.p99.|-cn)/(1+cn))^2))
                          by A39,A43,A44;
      ((cn-FanMorphN)|K0).(|[x,y]|)=((cn-FanMorphN)).(|[x,y]|) by A43,FUNCT_1:
72
    .= |[ |.p99.|* ((p99`1/|.p99.|-cn)/(1+cn)),
      |.p99.|*( sqrt(1-((p99`1/|.p99.|-cn)/(1+cn))^2))]| by A1,A45,Th58
    .=|[s,r]| by A32,A43,A44,A46;
   hence f.(|[x,y]|)=|[s,r]| by A1;
  end;
hence f is continuous by A7,A8,A33,A42,JGRAPH_2:45;
end;

theorem Th65: for cn being Real, K03 being Subset of TOP-REAL 2
 st K03={p: p`1>=(cn)*(|.p.|) & p`2>=0}
 holds K03 is closed
proof let sn be Real, K003 be Subset of TOP-REAL 2;
 assume A1: K003={p: p`1>=(sn)*(|.p.|) & p`2>=0};
 defpred P[Point of TOP-REAL 2] means ($1`1>=sn*|.$1.|);
  reconsider K1={p7 where p7 is Point of TOP-REAL 2:P[p7]}
  as Subset of TOP-REAL 2 from TopSubset;
  A2: K1 is closed by Lm8;
  defpred Q[Point of TOP-REAL 2] means $1`2>=0;
  reconsider KX={p where p is Point of TOP-REAL 2: Q[p]}
    as Subset of TOP-REAL 2 from TopSubset;
A3: KX is closed by JORDAN6:8;
     {p : P[p] & Q[p]}
    = {p7 where p7 is Point of TOP-REAL 2:P[p7]} /\
     {p1 where p1 is Point of TOP-REAL 2: Q[p1]} from TopInter2;
  hence thesis by A1,A2,A3,TOPS_1:35;
end;

theorem Th66: for cn being Real, K03 being Subset of TOP-REAL 2
 st K03={p: p`1<=(cn)*(|.p.|) & p`2>=0}
 holds K03 is closed
proof let sn be Real, K003 be Subset of TOP-REAL 2;
 assume A1: K003={p: p`1<=(sn)*(|.p.|) & p`2>=0};
 defpred P[Point of TOP-REAL 2] means ($1`1<=sn*|.$1.|);
  reconsider K1={p7 where p7 is Point of TOP-REAL 2:P[p7]}
  as Subset of TOP-REAL 2 from TopSubset;
  A2: K1 is closed by Lm10;
  defpred Q[Point of TOP-REAL 2] means $1`2>=0;
  reconsider KX={p where p is Point of TOP-REAL 2: Q[p]}
    as Subset of TOP-REAL 2 from TopSubset;
A3: KX is closed by JORDAN6:8;
     {p : P[p] & Q[p]}
    = {p7 where p7 is Point of TOP-REAL 2:P[p7]} /\
     {p1 where p1 is Point of TOP-REAL 2: Q[p1]} from TopInter2;
  hence thesis by A1,A2,A3,TOPS_1:35;
end;

theorem Th67: for cn being Real,
K0,B0 being Subset of TOP-REAL 2,f being
map of (TOP-REAL 2)|K0,(TOP-REAL 2)|B0
st -1<cn & cn<1 & f=(cn-FanMorphN)|K0 &
B0=(the carrier of TOP-REAL 2) \ {0.REAL 2}
& K0={p: p`2>=0 & p<>0.REAL 2}
holds f is continuous
proof let cn be Real,K0,B0 be Subset of TOP-REAL 2,f be
map of (TOP-REAL 2)|K0,(TOP-REAL 2)|B0;
assume A1: -1<cn & cn<1 & f=(cn-FanMorphN)|K0 &
B0=(the carrier of TOP-REAL 2)
\ {0.REAL 2} &
K0={p: p`2>=0 & p<>0.REAL 2};
 set sn=sqrt(1-cn^2);
 set p0=|[cn,sn]|;
 A2: p0`1=cn by EUCLID:56;
 A3: p0`2=sn by EUCLID:56;
     cn^2<1 by A1,JGRAPH_2:8,SQUARE_1:59;
 then A4: 1-cn^2>0 by SQUARE_1:11;
 then A5: p0<>0.REAL 2 by A3,JGRAPH_2:11,SQUARE_1:93;
 A6: p0`2>0 by A3,A4,SQUARE_1:93;
 A7: |.p0.|=sqrt((sn)^2+cn^2) by A2,A3,JGRAPH_3:10;
   sn^2=1-cn^2 by A4,SQUARE_1:def 4;
 then |.p0.|=1 by A7,SQUARE_1:83,XCMPLX_1:27;
 then A8: p0`1/|.p0.|=cn by EUCLID:56;
   p0 in K0 by A1,A6,JGRAPH_2:11;
 then reconsider K1=K0 as non empty Subset of TOP-REAL 2;
    p0 in the carrier of TOP-REAL 2 & not p0 in {0.REAL 2}
               by A5,TARSKI:def 1;
  then reconsider D=B0 as non empty Subset of TOP-REAL 2 by A1,XBOOLE_0:def 4;
   A9: p0 in {p: p`1/|.p.|>=cn & p`2>=0 & p<>0.REAL 2} by A6,A8,JGRAPH_2:11;
   defpred P[Point of TOP-REAL 2] means
   $1`1/|.$1.|>=cn & $1`2>=0 & $1<>0.REAL 2;
   A10: {p: P[p]}
     is Subset of TOP-REAL 2 from SubsetD;
   A11: the carrier of ((TOP-REAL 2)|K1)=K1 by JORDAN1:1;
   A12: {p: p`1/|.p.|>=cn & p`2>=0 & p<>0.REAL 2} c= K1
    proof let x be set;assume x in
      {p: p`1/|.p.|>=cn & p`2>=0 & p<>0.REAL 2};
      then consider p such that
        A13:p=x & p`1/|.p.|>=cn & p`2>=0 & p<>0.REAL 2;
     thus x in K1 by A1,A13;
    end;
   then reconsider K00={p: p`1/|.p.|>=cn & p`2>=0 & p<>0.REAL 2}
      as non empty Subset of ((TOP-REAL 2)|K1)
                 by A9,A11;
   reconsider K001={p: p`1/|.p.|>=cn & p`2>=0 & p<>0.REAL 2}
            as non empty Subset of (TOP-REAL 2)
                       by A9,A10;
   defpred P[Point of TOP-REAL 2] means $1`1>=(cn)*(|.$1.|) & $1`2>=0;
     {p: P[p]} is Subset of
   TOP-REAL 2 from SubsetD;
   then reconsider K003={p: p`1>=(cn)*(|.p.|) & p`2>=0}
            as Subset of (TOP-REAL 2);
   A14: p0 in {p: p`1/|.p.|<=cn & p`2>=0 & p<>0.REAL 2} by A6,A8,JGRAPH_2:11;
   A15: {p: p`1/|.p.|<=cn & p`2>=0 & p<>0.REAL 2} c= K1
    proof let x be set;assume x in
      {p: p`1/|.p.|<=cn & p`2>=0 & p<>0.REAL 2};
      then consider p such that
        A16: p=x & p`1/|.p.|<=cn & p`2>=0 & p<>0.REAL 2;
     thus x in K1 by A1,A16;
    end;
   then reconsider K11={p: p`1/|.p.|<=cn & p`2>=0 & p<>0.REAL 2}
            as non empty Subset of ((TOP-REAL 2)|K1)
            by A11,A14;
    A17: p0 in {p: p`1/|.p.|<=cn & p`2>=0 & p<>0.REAL 2} by A6,A8,JGRAPH_2:11;
    defpred P[Point of TOP-REAL 2] means
    $1`1/|.$1.|<=cn & $1`2>=0 & $1<>0.REAL 2;
     {p: P[p]} is Subset of
   TOP-REAL 2 from SubsetD;
   then reconsider K111={p: p`1/|.p.|<=cn & p`2>=0 & p<>0.REAL 2}
            as non empty Subset of (TOP-REAL 2)
                       by A17;
    defpred P[Point of TOP-REAL 2] means $1`1<=(cn)*(|.$1.|) & $1`2>=0;
     {p: P[p]} is Subset of
   TOP-REAL 2 from SubsetD;
   then reconsider K004={p: p`1<=(cn)*(|.p.|) & p`2>=0}
            as Subset of (TOP-REAL 2);
    the carrier of (TOP-REAL 2)|B0=the carrier of (TOP-REAL 2)|D;
  then A18: dom f=the carrier of (TOP-REAL 2)|K1 by FUNCT_2:def 1
    .=K1 by JORDAN1:1;
  then A19: dom (f|K00)=K00 by A12,RELAT_1:91
  .= the carrier of ((TOP-REAL 2)|K1)|K00 by JORDAN1:1;
  A20: the carrier of (TOP-REAL 2)|D =D by JORDAN1:1;
    rng (f|K00) c= rng f by RELAT_1:99;
  then rng (f|K00) c=D by A20,XBOOLE_1:1;
  then f|K00 is Function of the carrier of ((TOP-REAL 2)|K1)|K00,
   the carrier of (TOP-REAL 2)|D by A19,A20,FUNCT_2:4;
  then reconsider f1=f|K00 as map of ((TOP-REAL 2)|K1)|K00,(TOP-REAL 2)|D
                           ;
  A21: dom f1=the carrier of ((TOP-REAL 2)|K1)|K00 by FUNCT_2:def 1
  .=K00 by JORDAN1:1;
  A22: dom (cn-FanMorphN)=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
  then A23: dom ((cn-FanMorphN)|K001)=K001 by RELAT_1:91
  .= the carrier of (TOP-REAL 2)|K001 by JORDAN1:1;
  A24: the carrier of (TOP-REAL 2)|K1 = K1 by JORDAN1:1;
    rng ((cn-FanMorphN)|K001) c= K1
   proof let y be set;assume y in rng ((cn-FanMorphN)|K001);
     then consider x being set such that
     A25: x in dom ((cn-FanMorphN)|K001) & y=((cn-FanMorphN)|K001).x
        by FUNCT_1:def 5;
     A26: dom ((cn-FanMorphN)|K001)=(dom (cn-FanMorphN))/\ K001 by FUNCT_1:68
     .=(the carrier of TOP-REAL 2)/\ K001 by FUNCT_2:def 1
     .=K001 by XBOOLE_1:28;
     then reconsider q=x as Point of TOP-REAL 2 by A25;
     A27: y=(cn-FanMorphN).q by A25,FUNCT_1:70;
     consider p2 being Point of TOP-REAL 2 such that
     A28: p2=q & p2`1/|.p2.|>=cn & p2`2>=0 & p2<>0.REAL 2 by A25,A26;
     A29: cn-FanMorphN.q= |[ |.q.|* ((q`1/|.q.|-cn)/(1-cn)),
        |.q.|*( sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))]| by A1,A28,Th58;
     set q4= |[ |.q.|* ((q`1/|.q.|-cn)/(1-cn)),
        |.q.|*( sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))]|;
     A30: |.q.|>=0 by TOPRNS_1:26;
       |.q.|<>0 by A28,TOPRNS_1:25;
     then A31: (|.q.|)^2>0 by SQUARE_1:74;
     A32: 1-cn>0 by A1,Th3;
        A33: q4`2= |.q.|*( sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))&
        q4`1= |.q.|* ((q`1/|.q.|-cn)/(1-cn)) by EUCLID:56;
        A34: (q`1/|.q.|-cn)>= 0 by A28,SQUARE_1:12;
          0<=(q`2)^2 by SQUARE_1:72;
        then 0+(q`1)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
        then (q`1)^2 <= (|.q.|)^2 by JGRAPH_3:10;
        then (q`1)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A31,REAL_1:73;
        then (q`1)^2/(|.q.|)^2 <= 1 by A31,XCMPLX_1:60;
        then ((q`1)/|.q.|)^2 <= 1 by SQUARE_1:69;
        then 1>=q`1/|.q.| by Th4;
        then A35: 1-cn>=q`1/|.q.|-cn by REAL_1:49;
        A36: -(1-cn)<= -0 by A32,REAL_1:50;
           -(1-cn)<= -( q`1/|.q.|-cn) by A35,REAL_1:50;
         then (-(1-cn))/(1-cn)<=(-( q`1/|.q.|-cn))/(1-cn)
                        by A32,REAL_1:73;
        then A37: -1<=(-( q`1/|.q.|-cn))/(1-cn) by A32,XCMPLX_1:198;
         --(1-cn)>= -(q`1/|.q.|-cn) by A34,A36,REAL_1:50;
       then (-(q`1/|.q.|-cn))/(1-cn)<=1 by A32,REAL_2:117;
       then ((-(q`1/|.q.|-cn))/(1-cn))^2<=1 by A37,JGRAPH_2:7,SQUARE_1:59;
       then A38: 1-((-(q`1/|.q.|-cn))/(1-cn))^2>=0 by SQUARE_1:12;
       then 1-(-((q`1/|.q.|-cn))/(1-cn))^2>=0 by XCMPLX_1:188;
       then A39: 1-((q`1/|.q.|-cn)/(1-cn))^2>=0 by SQUARE_1:61;
         sqrt(1-((-(q`1/|.q.|-cn))/(1-cn))^2)>=0 by A38,SQUARE_1:def 4;
       then sqrt(1-(-(q`1/|.q.|-cn))^2/(1-cn)^2)>=0 by SQUARE_1:69;
       then sqrt(1-(q`1/|.q.|-cn)^2/(1-cn)^2)>=0 by SQUARE_1:61;
       then A40: sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2)>=0 by SQUARE_1:69;
        A41: (q4`2)^2= (|.q.|)^2*( sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))^2
                         by A33,SQUARE_1:68
         .= (|.q.|)^2*(1-((q`1/|.q.|-cn)/(1-cn))^2) by A39,SQUARE_1:def 4;
        A42: (q4`1)^2= (|.q.|)^2*((q`1/|.q.|-cn)/(1-cn))^2
                         by A33,SQUARE_1:68;
          (|.q4.|)^2=(q4`1)^2+(q4`2)^2 by JGRAPH_3:10
        .=(|.q.|)^2*(1-((q`1/|.q.|-cn)/(1-cn))^2
           +((q`1/|.q.|-cn)/(1-cn))^2) by A41,A42,XCMPLX_1:8
        .=(|.q.|)^2*1 by XCMPLX_1:27 .=(|.q.|)^2;
     then q4`2>=0 & q4<>0.REAL 2 by A30,A31,A33,A40,SQUARE_1:19,60,TOPRNS_1:24
;
    hence y in K1 by A1,A27,A29;
   end;
  then (cn-FanMorphN)|K001 is Function of the carrier of (TOP-REAL 2)|K001,
   the carrier of (TOP-REAL 2)|K1 by A23,A24,FUNCT_2:4;
  then reconsider f3=(cn-FanMorphN)|K001
   as map of (TOP-REAL 2)|K001,(TOP-REAL 2)|K1 ;
  A43: dom (f|K11)=K11 by A15,A18,RELAT_1:91
  .= the carrier of ((TOP-REAL 2)|K1)|K11 by JORDAN1:1;
  A44: the carrier of (TOP-REAL 2)|D =D by JORDAN1:1;
    rng (f|K11) c= rng f by RELAT_1:99;
  then rng (f|K11) c=D by A44,XBOOLE_1:1;
  then f|K11 is Function of the carrier of ((TOP-REAL 2)|K1)|K11,
   the carrier of (TOP-REAL 2)|D by A43,A44,FUNCT_2:4;
  then reconsider f2=f|K11 as map of ((TOP-REAL 2)|K1)|K11,(TOP-REAL 2)|D
                   ;
  A45: dom f2=the carrier of ((TOP-REAL 2)|K1)|K11 by FUNCT_2:def 1
  .=K11 by JORDAN1:1;
  A46: dom ((cn-FanMorphN)|K111)=K111 by A22,RELAT_1:91
  .= the carrier of (TOP-REAL 2)|K111 by JORDAN1:1;
  A47: the carrier of (TOP-REAL 2)|K1 =K1 by JORDAN1:1;
    rng ((cn-FanMorphN)|K111) c= K1
   proof let y be set;assume y in rng ((cn-FanMorphN)|K111);
     then consider x being set such that
     A48: x in dom ((cn-FanMorphN)|K111) & y=((cn-FanMorphN)|K111).x
        by FUNCT_1:def 5;
     A49: dom ((cn-FanMorphN)|K111)=(dom (cn-FanMorphN))/\ K111 by FUNCT_1:68
     .=(the carrier of TOP-REAL 2)/\ K111 by FUNCT_2:def 1
     .=K111 by XBOOLE_1:28;
     then reconsider q=x as Point of TOP-REAL 2 by A48;
     A50: y=(cn-FanMorphN).q by A48,FUNCT_1:70;
     consider p2 being Point of TOP-REAL 2 such that
     A51: p2=q & p2`1/|.p2.|<=cn & p2`2>=0 & p2<>0.REAL 2 by A48,A49;
     A52: cn-FanMorphN.q= |[ |.q.|* ((q`1/|.q.|-cn)/(1+cn)),
        |.q.|*( sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))]| by A1,A51,Th58;
     set q4= |[ |.q.|* ((q`1/|.q.|-cn)/(1+cn)),
       |.q.|*( sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))]|;
     A53: |.q.|>=0 by TOPRNS_1:26;
       |.q.|<>0 by A51,TOPRNS_1:25;
     then A54: (|.q.|)^2>0 by SQUARE_1:74;
   A55: 1+cn>0 by A1,Th3;
        A56: q4`2= |.q.|*( sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))&
        q4`1= |.q.|* ((q`1/|.q.|-cn)/(1+cn))
                     by EUCLID:56;
        A57: (q`1/|.q.|-cn)<=0 by A51,REAL_2:106;
        A58: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
          0<=(q`2)^2 by SQUARE_1:72;
        then 0+(q`1)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
        then (q`1)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A54,A58,REAL_1:73;
        then (q`1)^2/(|.q.|)^2 <= 1 by A54,XCMPLX_1:60;
        then ((q`1)/|.q.|)^2 <= 1 by SQUARE_1:69;
        then -1<=q`1/|.q.| by Th4;
        then -1-cn<=q`1/|.q.|-cn by REAL_1:49;
        then A59: -(1+cn)<=q`1/|.q.|-cn by XCMPLX_1:161;
         (1+cn)/(1+cn)>=(q`1/|.q.|-cn)/(1+cn) by A55,A57,REAL_1:73;
        then A60: 1>=(q`1/|.q.|-cn)/(1+cn) by A55,XCMPLX_1:60;
           (-(1+cn))/(1+cn)<=(( q`1/|.q.|-cn))/(1+cn)
                        by A55,A59,REAL_1:73;
        then -1<=(( q`1/|.q.|-cn))/(1+cn)
                        by A55,XCMPLX_1:198;
       then (((q`1/|.q.|-cn))/(1+cn))^2<=1^2 by A60,JGRAPH_2:7;
       then A61: 1-((q`1/|.q.|-cn)/(1+cn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((q`1/|.q.|-cn)/(1+cn)))^2>=0 by SQUARE_1:61;
       then 1-((-(q`1/|.q.|-cn))/(1+cn))^2>=0 by XCMPLX_1:188;
       then sqrt(1-((-(q`1/|.q.|-cn))/(1+cn))^2)>=0 by SQUARE_1:def 4;
       then sqrt(1-(-(q`1/|.q.|-cn))^2/(1+cn)^2)>=0 by SQUARE_1:69;
       then sqrt(1-(q`1/|.q.|-cn)^2/(1+cn)^2)>=0 by SQUARE_1:61;
       then sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2)>=0 by SQUARE_1:69;
       then A62: q4`2>= 0 by A53,A56,SQUARE_1:19;
        A63: (q4`2)^2= (|.q.|)^2*( sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))^2
                         by A56,SQUARE_1:68
         .= (|.q.|)^2*(1-((q`1/|.q.|-cn)/(1+cn))^2) by A61,SQUARE_1:def 4;
        A64: (q4`1)^2= (|.q.|)^2*((q`1/|.q.|-cn)/(1+cn))^2
                         by A56,SQUARE_1:68;
          (|.q4.|)^2=(q4`1)^2+(q4`2)^2 by JGRAPH_3:10
        .=(|.q.|)^2*(1-((q`1/|.q.|-cn)/(1+cn))^2
           +((q`1/|.q.|-cn)/(1+cn))^2) by A63,A64,XCMPLX_1:8
        .=(|.q.|)^2*1 by XCMPLX_1:27 .=(|.q.|)^2;
        then q4<>0.REAL 2 by A54,SQUARE_1:60,TOPRNS_1:24;
    hence y in K1 by A1,A50,A52,A62;
   end;
  then (cn-FanMorphN)|K111 is Function of the carrier of (TOP-REAL 2)|K111,
   the carrier of (TOP-REAL 2)|K1 by A46,A47,FUNCT_2:4;
  then reconsider f4=(cn-FanMorphN)|K111
   as map of (TOP-REAL 2)|K111,(TOP-REAL 2)|K1 ;
  set T1= ((TOP-REAL 2)|K1)|K00,T2=((TOP-REAL 2)|K1)|K11;
  A65: [#](((TOP-REAL 2)|K1)|K00)=K00 by PRE_TOPC:def 10;
  A66: [#](((TOP-REAL 2)|K1)|K11)=K11 by PRE_TOPC:def 10;
  A67: [#]((TOP-REAL 2)|K1)=K1 by PRE_TOPC:def 10;
  A68: K00 \/ K11 c= K1
   proof let x be set;assume
     A69: x in K00 \/ K11;
       now per cases by A69,XBOOLE_0:def 2;
     case x in K00;
       then consider p8 being Point of TOP-REAL 2 such that
       A70: p8=x & p8`1/|.p8.|>=cn & p8`2>=0 & p8<>0.REAL 2;
      thus x in K1 by A1,A70;
     case x in K11;
       then consider p8 being Point of TOP-REAL 2 such that
       A71: p8=x & p8`1/|.p8.|<=cn & p8`2>=0 & p8<>0.REAL 2;
      thus x in K1 by A1,A71;
     end;
    hence x in K1;
   end;
  A72: K1 c= K00 \/ K11
   proof let x be set;assume x in K1;
     then consider p such that
     A73: p=x & (p`2>=0 & p<>0.REAL 2) by A1;
       now per cases;
     case p`1/|.p.|>=cn;
       then x in K00 by A73;
      hence x in K00 \/ K11 by XBOOLE_0:def 2;
     case p`1/|.p.|<cn;
       then x in K11 by A73;
      hence x in K00 \/ K11 by XBOOLE_0:def 2;
     end;
    hence x in K00 \/ K11;
   end;
  then A74: [#](((TOP-REAL 2)|K1)|K00) \/ [#](((TOP-REAL 2)|K1)|K11)
        =[#]((TOP-REAL 2)|K1) by A65,A66,A67,A68,XBOOLE_0:def 10;
  A75: K003 is closed by Th65;
  A76: K003 /\ K1 c= K00
   proof let x be set;assume x in K003 /\ K1;
     then A77: x in K003 & x in K1 by XBOOLE_0:def 3;
     then consider q1 being Point of TOP-REAL 2 such that
     A78: q1=x & q1`1>=(cn)*(|.q1.|) & q1`2>=0;
     consider q2 being Point of TOP-REAL 2 such that
     A79: q2=x & q2`2>=0 & q2<>0.REAL 2 by A1,A77;
     A80: |.q2.|<>0 by A79,TOPRNS_1:25;
       |.q1.|>0 by A78,A79,Lm1;
     then q1`1/|.q1.|>=(cn)*(|.q1.|)/|.q1.| by A78,REAL_1:73;
     then q1`1/|.q1.|>=(cn) by A78,A79,A80,XCMPLX_1:90;
    hence x in K00 by A78,A79;
   end;
    K00 c= K003 /\ K1
   proof let x be set;assume x in K00;
     then consider p such that
     A81: p=x & (p`1/|.p.|>=cn & p`2>=0 & p<>0.REAL 2);
     A82: |.p.|<>0 by A81,TOPRNS_1:25;
       |.p.|>=0 by TOPRNS_1:26;
     then p`1/|.p.|*|.p.|>=(cn)*(|.p.|) by A81,AXIOMS:25;
     then p`1>=(cn)*(|.p.|) by A82,XCMPLX_1:88;
     then A83: x in K003 by A81;
       x in K1 by A1,A81;
    hence x in K003 /\ K1 by A83,XBOOLE_0:def 3;
   end;
  then K00=K003 /\ [#]((TOP-REAL 2)|K1) by A67,A76,XBOOLE_0:def 10;
  then A84: K00 is closed by A75,PRE_TOPC:43;
  A85: K004 is closed by Th66;
  A86: K004 /\ K1 c= K11
   proof let x be set;assume x in K004 /\ K1;
     then A87: x in K004 & x in K1 by XBOOLE_0:def 3;
     then consider q1 being Point of TOP-REAL 2 such that
     A88: q1=x & q1`1<=(cn)*(|.q1.|) & q1`2>=0;
     consider q2 being Point of TOP-REAL 2 such that
     A89: q2=x & q2`2>=0 & q2<>0.REAL 2 by A1,A87;
     A90: |.q2.|<>0 by A89,TOPRNS_1:25;
       |.q1.|>0 by A88,A89,Lm1;
     then q1`1/|.q1.|<=(cn)*(|.q1.|)/|.q1.| by A88,REAL_1:73;
     then q1`1/|.q1.|<=(cn) by A88,A89,A90,XCMPLX_1:90;
    hence x in K11 by A88,A89;
   end;
    K11 c= K004 /\ K1
   proof let x be set;assume x in K11;
     then consider p such that
     A91: p=x & (p`1/|.p.|<=cn & p`2>=0 & p<>0.REAL 2);
     A92: |.p.|<>0 by A91,TOPRNS_1:25;
       |.p.|>=0 by TOPRNS_1:26;
     then p`1/|.p.|*|.p.|<=(cn)*(|.p.|) by A91,AXIOMS:25;
     then p`1<=(cn)*(|.p.|) by A92,XCMPLX_1:88;
     then A93: x in K004 by A91;
       x in K1 by A1,A91;
    hence x in K004 /\ K1 by A93,XBOOLE_0:def 3;
   end;
  then K11=K004 /\ [#]((TOP-REAL 2)|K1) by A67,A86,XBOOLE_0:def 10;
  then A94: K11 is closed by A85,PRE_TOPC:43;
  A95: ((TOP-REAL 2)|K1)|K00=(TOP-REAL 2)|K001 by GOBOARD9:4;
    K1 c= D
   proof let x be set;assume x in K1;
     then consider p6 being Point of TOP-REAL 2 such that
     A96: p6=x & p6`2>=0 & p6<>0.REAL 2 by A1;
       x in the carrier of TOP-REAL 2 & not x in {0.REAL 2} by A96,TARSKI:def 1
;
    hence x in D by A1,XBOOLE_0:def 4;
   end;
  then D=K1 \/ D by XBOOLE_1:12;
  then A97: (TOP-REAL 2)|K1 is SubSpace of (TOP-REAL 2)|D by TOPMETR:5;
    the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
  then A98: f1= f3 by A1,FUNCT_1:82;
    f3 is continuous by A1,Th63;
  then A99: f1 is continuous by A95,A97,A98,TOPMETR:7;
  A100: ((TOP-REAL 2)|K1)|K11=(TOP-REAL 2)|K111 by GOBOARD9:4;
    the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
  then A101: f2= f4 by A1,FUNCT_1:82;
    f4 is continuous by A1,Th64;
  then A102: f2 is continuous by A97,A100,A101,TOPMETR:7;
  A103: D<>{};
  A104: the carrier of (TOP-REAL 2)|B0=B0 by JORDAN1:1;
    for p being set st p in ([#]T1)/\([#]T2) holds f1.p = f2.p
   proof let p be set;
    assume p in ([#]T1)/\([#]T2);
     then A105: p in K00 & p in K11 by A65,A66,XBOOLE_0:def 3;
    hence f1.p=f.p by FUNCT_1:72 .=f2.p by A105,FUNCT_1:72;
   end;
  then consider h being map of (TOP-REAL 2)|K1,(TOP-REAL 2)|D
  such that
  A106: h=f1+*f2 & h is continuous by A65,A66,A74,A84,A94,A99,A102,JGRAPH_2:9;
  A107: dom h=the carrier of ((TOP-REAL 2)|K1) by FUNCT_2:def 1;
  A108: the carrier of ((TOP-REAL 2)|K1)=K0 by JORDAN1:1;
  A109: K0=the carrier of ((TOP-REAL 2)|K0) by JORDAN1:1
  .=dom f by A103,A104,FUNCT_2:def 1;
    for y being set st y in dom h holds h.y=f.y
   proof let y be set;assume
    A110: y in dom h;
      now per cases by A72,A107,A108,A110,XBOOLE_0:def 2;
    case A111: y in K00 & not y in K11;
      then y in dom f1 \/ dom f2 by A21,XBOOLE_0:def 2;
     hence h.y=f1.y by A45,A106,A111,FUNCT_4:def 1 .=f.y by A111,FUNCT_1:72;
    case A112: y in K11;
      then y in dom f1 \/ dom f2 by A45,XBOOLE_0:def 2;
     hence h.y=f2.y by A45,A106,A112,FUNCT_4:def 1 .=f.y by A112,FUNCT_1:72;
    end;
   hence h.y=f.y;
   end;
 hence thesis by A106,A107,A108,A109,FUNCT_1:9;
end;

theorem Th68: for cn being Real,
K0,B0 being Subset of TOP-REAL 2,f being
map of (TOP-REAL 2)|K0,(TOP-REAL 2)|B0
st -1<cn & cn<1 & f=(cn-FanMorphN)|K0 &
B0=(the carrier of TOP-REAL 2) \ {0.REAL 2}
& K0={p: p`2<=0 & p<>0.REAL 2}
holds f is continuous
proof let cn be Real,K0,B0 be Subset of TOP-REAL 2,f be
map of (TOP-REAL 2)|K0,(TOP-REAL 2)|B0;
assume A1:  -1<cn & cn<1 & f=(cn-FanMorphN)|K0 &
B0=(the carrier of TOP-REAL 2)
\ {0.REAL 2} &
K0={p: p`2<=0 & p<>0.REAL 2};
 set sn=sqrt(1-cn^2);
 set p0=|[cn,-sn]|;
 A2: p0`2=-sn by EUCLID:56;
     cn^2<1^2 by A1,JGRAPH_2:8;
 then 1-cn^2>0 by SQUARE_1:11,59;
 then --sn>0 by SQUARE_1:93;
 then A3: p0`2<0 by A2,REAL_1:66;
 then p0 in K0 by A1,JGRAPH_2:11;
then reconsider K1=K0 as non empty Subset of TOP-REAL 2;
   p0 in the carrier of TOP-REAL 2 & not p0 in {0.REAL 2}
               by A3,JGRAPH_2:11,TARSKI:def 1;
  then reconsider D=B0 as non empty Subset of TOP-REAL 2 by A1,XBOOLE_0:def 4;
   A4: K1 c= D
    proof let x be set;assume x in K1;
      then consider p2 being Point of TOP-REAL 2 such that
      A5: p2=x & p2`2<=0 & p2<>0.REAL 2 by A1;
        not p2 in {0.REAL 2} by A5,TARSKI:def 1;
     hence x in D by A1,A5,XBOOLE_0:def 4;
    end;
     for p being Point of (TOP-REAL 2)|K1,V being Subset of (TOP-REAL 2)|D
      st f.p in V & V is open holds
      ex W being Subset of (TOP-REAL 2)|K1 st
      p in W & W is open & f.:W c= V
    proof let p be Point of (TOP-REAL 2)|K1,V be Subset of (TOP-REAL 2)|D;
      assume A6: f.p in V & V is open;
      then consider V2 being Subset of TOP-REAL 2 such that
      A7: V2 is open & V2 /\ [#]((TOP-REAL 2)|D)=V by TOPS_2:32;
      A8: p in the carrier of (TOP-REAL 2)|K1;
      A9: the carrier of (TOP-REAL 2)|K1=[#]((TOP-REAL 2)|K1) by PRE_TOPC:12;
      A10: [#]((TOP-REAL 2)|K1)=K1 by PRE_TOPC:def 10;
        V2 /\ [#]((TOP-REAL 2)|K1) c= the carrier of (TOP-REAL 2)|K1
       proof let x be set;assume x in V2 /\ [#]((TOP-REAL 2)|K1);
         then x in [#]((TOP-REAL 2)|K1) by XBOOLE_0:def 3;
        hence x in the carrier of (TOP-REAL 2)|K1;
       end;
      then reconsider W2=V2 /\ [#]((TOP-REAL 2)|K1) as Subset of (TOP-REAL 2)|
K1
                               ;
      A11: W2 is open by A7,TOPS_2:32;
      A12: f.p=(cn-FanMorphN).p by A1,A9,A10,FUNCT_1:72;
      consider q being Point of TOP-REAL 2 such that
      A13: q=p & q`2<=0 & q <>0.REAL 2 by A1,A8,A9,A10;
        (cn-FanMorphN).q=q by A13,Th56;
      then p in V2 & p in [#]((TOP-REAL 2)|D)
                       by A6,A7,A12,A13,XBOOLE_0:def 3;
      then A14: p in W2 by A9,XBOOLE_0:def 3;
        f.:W2 c= V
       proof let y be set;assume y in f.:W2;
         then consider x being set such that
         A15: x in dom f & x in W2 & y=f.x by FUNCT_1:def 12;
           f is Function of the carrier of (TOP-REAL 2)|K1,
           the carrier of (TOP-REAL 2)|D;
         then dom f= K1 by A9,A10,FUNCT_2:def 1;
         then consider p4 being Point of TOP-REAL 2 such that
         A16: x=p4 & p4`2<=0 & p4<>0.REAL 2 by A1,A15;
         A17: f.p4=(cn-FanMorphN).p4 by A1,A9,A10,A15,A16,FUNCT_1:72
         .=p4 by A16,Th56;
         A18: p4 in V2 & p4 in [#]((TOP-REAL 2)|K1) by A15,A16,XBOOLE_0:def 3
;
         then p4 in D by A4,A10;
         then p4 in [#]((TOP-REAL 2)|D) by PRE_TOPC:def 10;
        hence y in V by A7,A15,A16,A17,A18,XBOOLE_0:def 3;
       end;
     hence ex W being Subset of (TOP-REAL 2)|K1 st
      p in W & W is open & f.:W c= V by A11,A14;
    end;
  hence f is continuous by JGRAPH_2:20;
end;

theorem Th69:
 for B0 being Subset of TOP-REAL 2,
     K0 being Subset of (TOP-REAL 2)|B0
st B0=(the carrier of TOP-REAL 2) \ {0.REAL 2}
& K0={p: p`2>=0 & p<>0.REAL 2} holds K0 is closed
proof
 defpred P[Point of TOP-REAL 2] means $1`2>=0;
 set I1 = {p: P[p] & p<>0.REAL 2};
 set J0 = (the carrier of TOP-REAL 2) \ {0.REAL 2};
 let B0 be Subset of TOP-REAL 2,K0 be Subset of (TOP-REAL 2)|B0;
 assume A1: B0=J0 & K0=I1;
  reconsider K1={p7 where p7 is Point of TOP-REAL 2:P[p7]}
    as Subset of TOP-REAL 2 from TopSubset;
  A2:K1 is closed by JORDAN6:8;
    I1 = {p7 where p7 is Point of TOP-REAL 2 : P[p7]} /\ J0
    from TopInter;
  then K0=K1 /\ [#]((TOP-REAL 2)|B0) by A1,PRE_TOPC:def 10;
hence K0 is closed by A2,PRE_TOPC:43;
end;

theorem Th70:
 for B0 being Subset of TOP-REAL 2,
     K0 being Subset of (TOP-REAL 2)|B0
st B0=(the carrier of TOP-REAL 2) \ {0.REAL 2}
& K0={p: p`2<=0 & p<>0.REAL 2} holds K0 is closed
proof
defpred P[Point of TOP-REAL 2] means $1`2<=0;
 set I1 = {p: P[p] & p<>0.REAL 2};
 set J0 = (the carrier of TOP-REAL 2) \ {0.REAL 2};
 let B0 be Subset of TOP-REAL 2,K0 be Subset of (TOP-REAL 2)|B0;
 assume A1: B0=J0 & K0=I1;
  reconsider K1={p7 where p7 is Point of TOP-REAL 2:P[p7]}
    as Subset of TOP-REAL 2 from TopSubset;
  A2:K1 is closed by JORDAN6:9;
    I1 = {p7 where p7 is Point of TOP-REAL 2 : P[p7]} /\ J0
    from TopInter;
  then K0=K1 /\ [#]((TOP-REAL 2)|B0) by A1,PRE_TOPC:def 10;
hence K0 is closed by A2,PRE_TOPC:43;
end;

theorem Th71: for cn being Real,
 B0 being Subset of TOP-REAL 2,K0 being Subset of
(TOP-REAL 2)|B0,f being map of ((TOP-REAL 2)|B0)|K0,((TOP-REAL 2)|B0)
st -1<cn & cn<1 & f=(cn-FanMorphN)|K0 &
B0=(the carrier of TOP-REAL 2) \ {0.REAL 2}
& K0={p: p`2>=0 & p<>0.REAL 2}
holds f is continuous
proof let cn be Real,
 B0 be Subset of TOP-REAL 2,K0 be Subset of
(TOP-REAL 2)|B0,f being map of ((TOP-REAL 2)|B0)|K0,((TOP-REAL 2)|B0);
assume A1:  -1<cn & cn<1 & f=(cn-FanMorphN)|K0 &
  B0=(the carrier of TOP-REAL 2) \ {0.REAL 2}
   & K0={p: p`2>=0 & p<>0.REAL 2 };
    the carrier of (TOP-REAL 2)|B0= B0 by JORDAN1:1;
  then K0 c= the carrier of TOP-REAL 2 by XBOOLE_1:1;
  then reconsider K1=K0 as Subset of TOP-REAL 2;
    K0 c= B0
   proof let x be set;assume x in K0;
      then consider p8 being Point of TOP-REAL 2 such that
      A2: x=p8 & (p8`2>=0 & p8<>0.REAL 2) by A1;
       not x in {0.REAL 2} by A2,TARSKI:def 1;
    hence x in B0 by A1,A2,XBOOLE_0:def 4;
   end;
  then ((TOP-REAL 2)|B0)|K0=(TOP-REAL 2)|K1 by JORDAN6:47;
hence thesis by A1,Th67;
end;

theorem Th72: for cn being Real,
 B0 being Subset of TOP-REAL 2,K0 being Subset of
(TOP-REAL 2)|B0,f being map of ((TOP-REAL 2)|B0)|K0,((TOP-REAL 2)|B0)
st -1<cn & cn<1 & f=(cn-FanMorphN)|K0 &
B0=(the carrier of TOP-REAL 2) \ {0.REAL 2}
& K0={p: p`2<=0 & p<>0.REAL 2}
holds f is continuous
proof let cn be Real,
 B0 be Subset of TOP-REAL 2,K0 be Subset of
(TOP-REAL 2)|B0,f be map of ((TOP-REAL 2)|B0)|K0,((TOP-REAL 2)|B0);
 assume A1:  -1<cn & cn<1 & f=(cn-FanMorphN)|K0 &
  B0=(the carrier of TOP-REAL 2) \ {0.REAL 2}
  & K0={p: p`2<=0 & p<>0.REAL 2};
    the carrier of (TOP-REAL 2)|B0= B0 by JORDAN1:1;
  then K0 c= the carrier of TOP-REAL 2 by XBOOLE_1:1;
  then reconsider K1=K0 as Subset of TOP-REAL 2;
    K0 c= B0
   proof let x be set;assume x in K0;
      then consider p8 being Point of TOP-REAL 2 such that
      A2: x=p8 & (p8`2<=0 & p8<>0.REAL 2) by A1;
       not x in {0.REAL 2} by A2,TARSKI:def 1;
    hence x in B0 by A1,A2,XBOOLE_0:def 4;
   end;
  then ((TOP-REAL 2)|B0)|K0=(TOP-REAL 2)|K1 by JORDAN6:47;
hence thesis by A1,Th68;
end;

theorem Th73: for cn being Real,p being Point of TOP-REAL 2
holds |.(cn-FanMorphN).p.|=|.p.|
proof let cn be Real,p be Point of TOP-REAL 2;
       set f=(cn-FanMorphN);
       set z=f.p;
        reconsider q=p as Point of TOP-REAL 2;
        reconsider qz=z as Point of TOP-REAL 2;
        A1: |.q.|>=0 by TOPRNS_1:26;
          now per cases;
        case A2: q`1/|.q.|>=cn & q`2>0;
          then (cn-FanMorphN).q= |[ |.q.|* ((q`1/|.q.|-cn)/(1-cn)),
          |.q.|*( sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))]| by Th56;
          then A3: qz`2= |.q.|*( sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))&
                  qz`1= |.q.|* ((q`1/|.q.|-cn)/(1-cn)) by EUCLID:56;
            |.q.|<>0 by A2,JGRAPH_2:11,TOPRNS_1:25;
          then A4: (|.q.|)^2>0 by SQUARE_1:74;
          A5: (q`1/|.q.|-cn)>=0 by A2,SQUARE_1:12;
          A6: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
            0<=(q`2)^2 by SQUARE_1:72;
          then 0+(q`1)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
          then (q`1)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A4,A6,REAL_1:73;
          then (q`1)^2/(|.q.|)^2 <= 1 by A4,XCMPLX_1:60;
          then ((q`1)/|.q.|)^2 <= 1 by SQUARE_1:69;
          then 1>=q`1/|.q.| by Th4;
          then A7: 1-cn>=q`1/|.q.|-cn by REAL_1:49;
            now per cases;
          case A8: 1-cn=0;
            A9:((q`1/|.q.|-cn)/(1-cn))=(q`1/|.q.|-cn)*(1-cn)" by XCMPLX_0:def 9
             .= (q`1/|.q.|-cn)*0 by A8,XCMPLX_0:def 7 .=0;
            then 1-((q`1/|.q.|-cn)/(1-cn))^2=1 by SQUARE_1:60;
            then (cn-FanMorphN).q= |[ |.q.|*0,|.q.|*(1)]| by A2,A9,Th56,
SQUARE_1:83
            .=|[0,(|.q.|)]|;
             then ((cn-FanMorphN).q)`2=(|.q.|) & ((cn-FanMorphN).q)`1=0
                              by EUCLID:56;
            then |.(cn-FanMorphN).p.|=sqrt(((|.q.|))^2+0) by JGRAPH_3:10,
SQUARE_1:60
            .=|.q.| by A1,SQUARE_1:89;
           hence |.(cn-FanMorphN).p.|=|.p.|;
          case A10: 1-cn<>0;
              now per cases by A10;
            case A11: 1-cn>0;
              then A12: -(1-cn)<= -0 by REAL_1:50;
                -(1-cn)<= -( q`1/|.q.|-cn) by A7,REAL_1:50;
              then (-(1-cn))/(1-cn)<=(-( q`1/|.q.|-cn))/(1-cn)
                        by A11,REAL_1:73;
              then A13: -1<=(-( q`1/|.q.|-cn))/(1-cn) by A11,XCMPLX_1:198;
               --(1-cn)>= -(q`1/|.q.|-cn) by A5,A12,REAL_1:50;
              then (-(q`1/|.q.|-cn))/(1-cn)<=1 by A11,REAL_2:117;
              then ((-(q`1/|.q.|-cn))/(1-cn))^2<=1^2 by A13,JGRAPH_2:7;
              then 1-((-(q`1/|.q.|-cn))/(1-cn))^2>=0 by SQUARE_1:12,59;
              then 1-(-((q`1/|.q.|-cn))/(1-cn))^2>=0 by XCMPLX_1:188;
              then A14: 1-((q`1/|.q.|-cn)/(1-cn))^2>=0 by SQUARE_1:61;
              A15: (qz`2)^2= (|.q.|)^2*( sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))^2
                         by A3,SQUARE_1:68
                  .= (|.q.|)^2*(1-((q`1/|.q.|-cn)/(1-cn))^2)
                            by A14,SQUARE_1:def 4;
              A16: (qz`1)^2= (|.q.|)^2*((q`1/|.q.|-cn)/(1-cn))^2
                         by A3,SQUARE_1:68;
                (|.qz.|)^2=(qz`1)^2+(qz`2)^2 by JGRAPH_3:10
               .=(|.q.|)^2*(1-((q`1/|.q.|-cn)/(1-cn))^2
                    +((q`1/|.q.|-cn)/(1-cn))^2) by A15,A16,XCMPLX_1:8
               .=(|.q.|)^2*1 by XCMPLX_1:27 .=(|.q.|)^2;
              then A17: sqrt((|.qz.|)^2)=|.q.| by A1,SQUARE_1:89;
                |.qz.|>=0 by TOPRNS_1:26;
             hence |.(cn-FanMorphN).p.|=|.p.| by A17,SQUARE_1:89;
            case A18: 1-cn<0;
              A19: q`1/|.q.|-cn>=0 by A2,SQUARE_1:12;
                0<(q`2)^2 by A2,SQUARE_1:74;
              then 0+(q`1)^2<(q`1)^2+(q`2)^2 by REAL_1:67;
              then (q`1)^2/(|.q.|)^2 < (|.q.|)^2/(|.q.|)^2 by A4,A6,REAL_1:73;
              then (q`1)^2/(|.q.|)^2 < 1 by A4,XCMPLX_1:60;
              then ((q`1)/|.q.|)^2 < 1 by SQUARE_1:69;
              then 1 > q`1/|.p.| by Th5;
              then q`1/|.q.|-cn<1-cn by REAL_1:54;
             hence |.(cn-FanMorphN).p.|=|.p.| by A18,A19,AXIOMS:22;
            end;
           hence |.(cn-FanMorphN).p.|=|.p.|;
          end;
         hence |.(cn-FanMorphN).p.|=|.p.|;
        case A20: q`1/|.q.|<cn & q`2>0;
          then A21:(cn-FanMorphN).q= |[ |.q.|* ((q`1/|.q.|-cn)/(1+cn)),
          |.q.|*( sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))]| by Th57;
          then A22: qz`2= |.q.|*( sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))&
                  qz`1= |.q.|* ((q`1/|.q.|-cn)/(1+cn)) by EUCLID:56;
            |.q.|<>0 by A20,JGRAPH_2:11,TOPRNS_1:25;
          then A23: (|.q.|)^2>0 by SQUARE_1:74;
          A24: (q`1/|.q.|-cn)<0 by A20,REAL_2:105;
          A25: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
            0<=(q`2)^2 by SQUARE_1:72;
          then 0+(q`1)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
          then (q`1)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A23,A25,REAL_1:73;
          then (q`1)^2/(|.q.|)^2 <= 1 by A23,XCMPLX_1:60;
          then ((q`1)/|.q.|)^2 <= 1 by SQUARE_1:69;
          then -1<=q`1/|.q.| by Th4;
          then A26: -1-cn<=q`1/|.q.|-cn by REAL_1:49;
            now per cases;
          case A27: 1+cn=0;
             ((q`1/|.q.|-cn)/(1+cn))=(q`1/|.q.|-cn)*(1+cn)" by XCMPLX_0:def 9
             .= (q`1/|.q.|-cn)*0 by A27,XCMPLX_0:def 7 .=0;
             then ((cn-FanMorphN).q)`2=(|.q.|) & ((cn-FanMorphN).q)`1=0
                              by A21,EUCLID:56,SQUARE_1:60,83;
            then |.(cn-FanMorphN).p.|=sqrt(((|.q.|))^2+0) by JGRAPH_3:10,
SQUARE_1:60
            .=|.q.| by A1,SQUARE_1:89;
           hence |.(cn-FanMorphN).p.|=|.p.|;
          case A28: 1+cn<>0;
              now per cases by A28;
            case A29: 1+cn>0;
                -(1+cn)<= ( q`1/|.q.|-cn) by A26,XCMPLX_1:161;
              then (-(1+cn))/(1+cn)<=(( q`1/|.q.|-cn))/(1+cn)
                        by A29,REAL_1:73;
              then A30: -1<=(( q`1/|.q.|-cn))/(1+cn)
                        by A29,XCMPLX_1:198;
                (1+cn)>= (q`1/|.q.|-cn) by A24,A29,AXIOMS:22;
              then ((q`1/|.q.|-cn))/(1+cn)<=1 by A29,REAL_2:117;
              then (((q`1/|.q.|-cn))/(1+cn))^2<=1 by A30,JGRAPH_2:7,SQUARE_1:59
;
              then A31: 1-(((q`1/|.q.|-cn))/(1+cn))^2>=0 by SQUARE_1:12;
              A32: (qz`2)^2= (|.q.|)^2*( sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))^2
                         by A22,SQUARE_1:68
                  .= (|.q.|)^2*(1-((q`1/|.q.|-cn)/(1+cn))^2)
                            by A31,SQUARE_1:def 4;
              A33: (qz`1)^2= (|.q.|)^2*((q`1/|.q.|-cn)/(1+cn))^2
                         by A22,SQUARE_1:68;
                (|.qz.|)^2=(qz`1)^2+(qz`2)^2 by JGRAPH_3:10
               .=(|.q.|)^2*(1-((q`1/|.q.|-cn)/(1+cn))^2
                    +((q`1/|.q.|-cn)/(1+cn))^2) by A32,A33,XCMPLX_1:8
               .=(|.q.|)^2*1 by XCMPLX_1:27 .=(|.q.|)^2;
              then A34: sqrt((|.qz.|)^2)=|.q.| by A1,SQUARE_1:89;
                |.qz.|>=0 by TOPRNS_1:26;
             hence |.(cn-FanMorphN).p.|=|.p.| by A34,SQUARE_1:89;
            case 1+cn<0;
              then A35: -(1+cn)>-0 by REAL_1:50;
                0<(q`2)^2 by A20,SQUARE_1:74;
              then 0+(q`1)^2<(q`1)^2+(q`2)^2 by REAL_1:67;
              then (q`1)^2/(|.q.|)^2 < (|.q.|)^2/(|.q.|)^2 by A23,A25,REAL_1:73
;
              then (q`1)^2/(|.q.|)^2 < 1 by A23,XCMPLX_1:60;
              then ((q`1)/|.q.|)^2 < 1 by SQUARE_1:69;
              then -1 < q`1/|.p.| by Th5;
              then q`1/|.q.|-cn>-1-cn by REAL_1:54;
              then q`1/|.q.|-cn>-(1+cn) by XCMPLX_1:161;
             hence |.(cn-FanMorphN).p.|=|.p.| by A24,A35,AXIOMS:22;
            end;
           hence |.(cn-FanMorphN).p.|=|.p.|;
          end;
         hence |.(cn-FanMorphN).p.|=|.p.|;
        case q`2<=0;
         hence |.(cn-FanMorphN).p.|=|.p.| by Th56;
        end;
 hence |.(cn-FanMorphN).p.|=|.p.|;
end;

theorem Th74: for cn being Real,x,K0 being set
 st -1<cn & cn<1 & x in K0 &
 K0={p: p`2>=0 & p<>0.REAL 2}
holds (cn-FanMorphN).x in K0
proof let cn be Real,x,K0 be set;
 assume A1:  -1<cn & cn<1 & x in K0 &
   K0={p: p`2>=0 & p<>0.REAL 2};
  then consider p such that
  A2: p=x & p`2>=0 & p<>0.REAL 2;
  A3:now assume |.p.|<=0;
    then |.p.|=0 by TOPRNS_1:26;
   hence contradiction by A2,TOPRNS_1:25;
  end;
  then A4: (|.p.|)^2>0 by SQUARE_1:74;
    now per cases;
  case A5: p`1/|.p.|<=cn;
   then A6: (cn-FanMorphN).p= |[ |.p.|*((p`1/|.p.|-cn)/(1+cn)),
      |.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2))]| by A1,A2,Th58;
   reconsider p9= (cn-FanMorphN).p as Point of TOP-REAL 2;
   A7: p9`2=|.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2)) by A6,EUCLID:56;
   A8: 1+cn>0 by A1,Th3;
       A9: (|.p.|)^2 =(p`1)^2+(p`2)^2 by JGRAPH_3:10;
         now per cases;
       case p`2=0;
        hence (cn-FanMorphN).x in K0 by A1,A2,Th56;
       case p`2<>0;
         then 0<(p`2)^2 by SQUARE_1:74;
         then 0+(p`1)^2<(p`1)^2+(p`2)^2 by REAL_1:67;
         then (p`1)^2/(|.p.|)^2 < (|.p.|)^2/(|.p.|)^2 by A4,A9,REAL_1:73;
         then (p`1)^2/(|.p.|)^2 < 1 by A4,XCMPLX_1:60;
         then A10: ((p`1)/|.p.|)^2 < 1 by SQUARE_1:69;
           p`1/|.p.|-cn<=0 by A5,REAL_2:106;
         then A11: (p`1/|.p.|-cn)/(1+cn)<(1+cn)/(1+cn) by A8,REAL_1:73;
           -1 < p`1/|.p.| by A10,Th5;
         then -1-cn< p`1/|.p.|-cn by REAL_1:54;
         then -(1+cn)<p`1/|.p.|-cn by XCMPLX_1:161;
         then (-1)*(1+cn)< p`1/|.p.|-cn by XCMPLX_1:180;
         then (-1)*(1+cn)/(1+cn)< (p`1/|.p.|-cn)/(1+cn)
                       by A8,REAL_1:73;
         then -1< (p`1/|.p.|-cn)/(1+cn) & (p`1/|.p.|-cn)/(1+cn)<1
                           by A8,A11,XCMPLX_1:60,90;
         then 1^2> ((p`1/|.p.|-cn)/(1+cn))^2 by JGRAPH_2:8;
         then 1-((p`1/|.p.|-cn)/(1+cn))^2>0 by SQUARE_1:11,59;
         then sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2)>0 by SQUARE_1:93;
         then |.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2))>0
                           by A3,SQUARE_1:21;
        hence (cn-FanMorphN).x in K0 by A1,A2,A7,JGRAPH_2:11;
        end;
       hence (cn-FanMorphN).x in K0;
  case A12: p`1/|.p.|>cn;
   then A13: (cn-FanMorphN).p= |[ |.p.|*((p`1/|.p.|-cn)/(1-cn)),
       |.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2))]| by A1,A2,Th58;
   reconsider p9= (cn-FanMorphN).p as Point of TOP-REAL 2;
   A14: p9`2=|.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2)) by A13,EUCLID:56;
    A15: 1-cn>0 by A1,Th3;
       A16: (|.p.|)^2 =(p`1)^2+(p`2)^2 by JGRAPH_3:10;
         now per cases;
       case p`2=0;
        hence (cn-FanMorphN).x in K0 by A1,A2,Th56;
       case p`2<>0;
         then 0<(p`2)^2 by SQUARE_1:74;
         then 0+(p`1)^2<(p`1)^2+(p`2)^2 by REAL_1:67;
         then (p`1)^2/(|.p.|)^2 < (|.p.|)^2/(|.p.|)^2 by A4,A16,REAL_1:73;
         then (p`1)^2/(|.p.|)^2 < 1 by A4,XCMPLX_1:60;
         then ((p`1)/|.p.|)^2 < 1 by SQUARE_1:69;
         then p`1/|.p.|<1 by Th5;
         then (p`1/|.p.|-cn)<1-cn by REAL_1:54;
         then A17: (p`1/|.p.|-cn)/(1-cn)<(1-cn)/(1-cn) by A15,REAL_1:73;
           -(1-cn)< -0 by A15,REAL_1:50;
         then A18: -(1-cn)<cn-cn by XCMPLX_1:14;
           p`1/|.p.|-cn>=cn-cn by A12,REAL_1:49;
         then -(1-cn)<p`1/|.p.|-cn by A18,AXIOMS:22;
         then (-1)*(1-cn)< p`1/|.p.|-cn by XCMPLX_1:180;
         then (-1)*(1-cn)/(1-cn)< (p`1/|.p.|-cn)/(1-cn)
                       by A15,REAL_1:73;
         then -1< (p`1/|.p.|-cn)/(1-cn) & (p`1/|.p.|-cn)/(1-cn)<1
                           by A15,A17,XCMPLX_1:60,90;
         then 1> ((p`1/|.p.|-cn)/(1-cn))^2 by JGRAPH_2:8,SQUARE_1:59;
         then 1-((p`1/|.p.|-cn)/(1-cn))^2>0 by SQUARE_1:11;
         then sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2)>0 by SQUARE_1:93;
         then p9`2>0 by A3,A14,SQUARE_1:21;
        hence (cn-FanMorphN).x in K0 by A1,A2,JGRAPH_2:11;
       end;
        hence (cn-FanMorphN).x in K0;
        end;
 hence (cn-FanMorphN).x in K0;
end;

theorem Th75:
  for cn being Real,x,K0 being set st -1<cn & cn<1 & x in K0 &
 K0={p: p`2<=0 & p<>0.REAL 2} holds (cn-FanMorphN).x in K0
proof let cn be Real,x,K0 be set;
 assume A1:  -1<cn & cn<1 & x in K0 &
 K0={p: p`2<=0 & p<>0.REAL 2};
 then consider p such that
 A2: p=x & p`2<=0 & p<>0.REAL 2;
 thus thesis by A1,A2,Th56;
end;

theorem Th76:for cn being Real,
 D being non empty Subset of TOP-REAL 2
st -1<cn & cn<1 & D`={0.REAL 2} holds
ex h being map of (TOP-REAL 2)|D,(TOP-REAL 2)|D
st h=(cn-FanMorphN)|D & h is continuous
proof let cn be Real,D be non empty Subset of TOP-REAL 2;
assume A1:  -1<cn & cn<1 & D`={0.REAL 2};
defpred P[Point of TOP-REAL 2] means $1`2>=0;
  reconsider B0= {0.REAL 2} as Subset of TOP-REAL 2;
  A2: D =(B0)` by A1
     .=((the carrier of TOP-REAL 2)\ {0.REAL 2}) by SUBSET_1:def 5;
    A3:{p: P[p] & p<>0.REAL 2} c= the carrier of (TOP-REAL 2)|D
        from InclSub(A2);
        (|[0,1]|)`1=0 & (|[0,1]|)`2=1 by EUCLID:56;
      then |[0,1]| in {p where p is Point of TOP-REAL 2:
        p`2>=0 & p<>0.REAL 2} by JGRAPH_2:11;
      then reconsider K0={p:p`2>=0 & p<>0.REAL 2}
          as non empty Subset of (TOP-REAL 2)|D by A3;
          defpred P[Point of TOP-REAL 2] means $1`2<=0;
    A4:{p: P[p] & p<>0.REAL 2} c= the carrier of (TOP-REAL 2)|D
        from InclSub(A2);
     set Y1=|[0,-1]|;
     A5:Y1`1=0 & Y1`2=-1 by EUCLID:56;
       0.REAL 2 <> Y1 by EUCLID:56,JGRAPH_2:11;
     then Y1 in {p where p is Point of TOP-REAL 2:
     p`2<=0 & p<>0.REAL 2} by A5;
  then reconsider K1={p: p`2<=0 & p<>0.REAL 2}
       as non empty Subset of (TOP-REAL 2)|D by A4;
   A6:the carrier of ((TOP-REAL 2)|D)=D by JORDAN1:1;
   A7:K0 c= (the carrier of TOP-REAL 2)
   proof let z be set;assume z in K0;
        then consider p8 being Point of TOP-REAL 2 such that
        A8: p8=z & p8`2>=0 & p8<>0.REAL 2;
    thus z in (the carrier of TOP-REAL 2) by A8;
   end;
  A9:dom ((cn-FanMorphN)|K0)= dom ((cn-FanMorphN)) /\ K0 by FUNCT_1:68
       .=((the carrier of TOP-REAL 2)) /\ K0 by FUNCT_2:def 1
       .=K0 by A7,XBOOLE_1:28;
  A10: K0=the carrier of ((TOP-REAL 2)|D)|K0 by JORDAN1:1;
    rng ((cn-FanMorphN)|K0) c= the carrier of ((TOP-REAL 2)|D)|K0
   proof let y be set;assume y in rng ((cn-FanMorphN)|K0);
     then consider x being set such that
     A11:x in dom ((cn-FanMorphN)|K0)
     & y=((cn-FanMorphN)|K0).x by FUNCT_1:def 5;
       x in (dom (cn-FanMorphN)) /\ K0 by A11,FUNCT_1:68;
     then A12:x in K0 by XBOOLE_0:def 3;
       K0 c= the carrier of TOP-REAL 2 by A6,XBOOLE_1:1;
     then reconsider p=x as Point of TOP-REAL 2 by A12;
      (cn-FanMorphN).p=y by A11,A12,FUNCT_1:72;
     then y in K0 by A1,A12,Th74;
    hence y in the carrier of ((TOP-REAL 2)|D)|K0 by JORDAN1:1;
   end;
  then rng ((cn-FanMorphN)|K0)c= the carrier of ((TOP-REAL 2)|D)
                            by A10,XBOOLE_1:1;
  then (cn-FanMorphN)|K0 is Function of the carrier of ((TOP-REAL 2)|D)|K0,
  the carrier of ((TOP-REAL 2)|D) by A9,A10,FUNCT_2:4;
  then reconsider f=(cn-FanMorphN)|K0 as map of ((TOP-REAL 2)|D)|K0,
  ((TOP-REAL 2)|D) ;
   A13:K1 c= (the carrier of TOP-REAL 2)
   proof let z be set;assume z in K1;
        then consider p8 being Point of TOP-REAL 2 such that
        A14: p8=z & p8`2<=0 & p8<>0.REAL 2;
    thus z in (the carrier of TOP-REAL 2) by A14;
   end;
  A15:dom ((cn-FanMorphN)|K1)= dom ((cn-FanMorphN)) /\ K1 by FUNCT_1:68
       .=((the carrier of TOP-REAL 2)) /\ K1 by FUNCT_2:def 1
       .=K1 by A13,XBOOLE_1:28;
  A16: K1=the carrier of ((TOP-REAL 2)|D)|K1 by JORDAN1:1;
    rng ((cn-FanMorphN)|K1) c= the carrier of ((TOP-REAL 2)|D)|K1
   proof let y be set;assume y in rng ((cn-FanMorphN)|K1);
     then consider x being set such that
     A17:x in dom ((cn-FanMorphN)|K1)
     & y=((cn-FanMorphN)|K1).x by FUNCT_1:def 5;
       x in (dom (cn-FanMorphN)) /\ K1 by A17,FUNCT_1:68;
     then A18:x in K1 by XBOOLE_0:def 3;
       K1 c= the carrier of TOP-REAL 2 by A6,XBOOLE_1:1;
     then reconsider p=x as Point of TOP-REAL 2 by A18;
      (cn-FanMorphN).p=y by A17,A18,FUNCT_1:72;
     then y in K1 by A1,A18,Th75;
    hence y in the carrier of ((TOP-REAL 2)|D)|K1 by JORDAN1:1;
   end;
  then rng ((cn-FanMorphN)|K1)c= the carrier of ((TOP-REAL 2)|D)
                     by A16,XBOOLE_1:1;
  then (cn-FanMorphN)|K1 is Function of the carrier of ((TOP-REAL 2)|D)|K1,
  the carrier of ((TOP-REAL 2)|D) by A15,A16,FUNCT_2:4;
  then reconsider g=(cn-FanMorphN)|K1 as map of ((TOP-REAL 2)|D)|K1,
  ((TOP-REAL 2)|D) ;
  A19:K0=[#](((TOP-REAL 2)|D)|K0) by PRE_TOPC:def 10;
  A20:K1=[#](((TOP-REAL 2)|D)|K1) by PRE_TOPC:def 10;
  A21:D= [#]((TOP-REAL 2)|D) by PRE_TOPC:def 10;
  A22:K0 \/ K1 c= D
   proof let x be set;assume A23:x in K0 \/ K1;
       now per cases by A23,XBOOLE_0:def 2;
     case x in K0;
       then consider p such that A24:p=x &
       p`2>=0 & p<>0.REAL 2;
     thus
       x in the carrier of TOP-REAL 2 & not x=0.REAL 2 by A24;
     case x in K1;
       then consider p such that A25:p=x &
       p`2<=0 & p<>0.REAL 2;
     thus
       x in the carrier of TOP-REAL 2 & not x=0.REAL 2 by A25;
     end;
     then x in the carrier of TOP-REAL 2 & not x in {0.REAL 2} by TARSKI:def 1
;
    hence x in D by A2,XBOOLE_0:def 4;
   end;
    D c= K0 \/ K1
   proof let x be set;assume A26: x in D;
     then A27:x in (the carrier of TOP-REAL 2) & not x in {0.REAL 2}
                                     by A2,XBOOLE_0:def 4;
     reconsider px=x as Point of TOP-REAL 2 by A26;
       px`2>=0 & px<>0.REAL 2 or
       px`2<=0 & px<>0.REAL 2 by A27,TARSKI:def 1;
     then x in K0 or x in K1;
    hence x in K0 \/ K1 by XBOOLE_0:def 2;
   end;
  then A28:([#](((TOP-REAL 2)|D)|K0)) \/ ([#](((TOP-REAL 2)|D)|K1))
  = [#]((TOP-REAL 2)|D) by A19,A20,A21,A22,XBOOLE_0:def 10;
  A29: f is continuous & K0 is closed by A1,A2,Th69,Th71;
  A30: g is continuous & K1 is closed by A1,A2,Th70,Th72;
  A31: for x be set st x in ([#]((((TOP-REAL 2)|D)|K0)))
        /\ ([#] ((((TOP-REAL 2)|D)|K1))) holds f.x = g.x
   proof let x be set;assume x in ([#]((((TOP-REAL 2)|D)|K0)))
        /\ ([#] ((((TOP-REAL 2)|D)|K1)));
     then A32:x in K0 & x in K1 by A19,A20,XBOOLE_0:def 3;
     then f.x=(cn-FanMorphN).x by FUNCT_1:72;
    hence f.x = g.x by A32,FUNCT_1:72;
   end;
  then consider h being map of (TOP-REAL 2)|D,(TOP-REAL 2)|D
  such that A33: h= f+*g & h is continuous
                          by A19,A20,A28,A29,A30,JGRAPH_2:9;
  A34:dom h=the carrier of ((TOP-REAL 2)|D) by FUNCT_2:def 1;
  A35:the carrier of ((TOP-REAL 2)|D)=D by JORDAN1:1;
  A36:the carrier of ((TOP-REAL 2)|D) =
((the carrier of TOP-REAL 2)\ {0.REAL 2}) by A2,JORDAN1:1;
    dom (cn-FanMorphN)=(the carrier of (TOP-REAL 2))
                    by FUNCT_2:def 1;
  then A37:dom ((cn-FanMorphN)|D)=(the carrier of (TOP-REAL 2))/\ D
                    by FUNCT_1:68
              .=the carrier of ((TOP-REAL 2)|D) by A35,XBOOLE_1:28;
  A38:dom f=K0 by A10,FUNCT_2:def 1;
  A39:K0=[#](((TOP-REAL 2)|D)|K0) by PRE_TOPC:def 10;
  A40:dom g=K1 by A16,FUNCT_2:def 1;
   K1=[#](((TOP-REAL 2)|D)|K1) by PRE_TOPC:def 10;
  then A41: f tolerates g by A31,A38,A39,A40,PARTFUN1:def 6;
    for x being set st x in dom h holds h.x=((cn-FanMorphN)|D).x
   proof let x be set;assume A42: x in dom h;
     then x in (the carrier of TOP-REAL 2) & not x in {0.REAL 2}
                     by A34,A36,XBOOLE_0:def 4;
     then A43:x <>0.REAL 2 by TARSKI:def 1;
     reconsider p=x as Point of TOP-REAL 2 by A34,A36,A42,XBOOLE_0:def 4;
     A44: x in D`` by A34,A42,JORDAN1:1;
       now per cases;
     case A45:x in K0;
       A46:(cn-FanMorphN)|D.p=(cn-FanMorphN).p by A44,FUNCT_1:72
          .=f.p by A45,FUNCT_1:72;
         h.p=(g+*f).p by A33,A41,FUNCT_4:35
       .=f.p by A38,A45,FUNCT_4:14;
      hence h.x=(cn-FanMorphN)|D.x by A46;
     case not x in K0;
       then not (p`2>=0) by A43;
       then A47:x in K1 by A43;
        (cn-FanMorphN)|D.p=(cn-FanMorphN).p by A44,FUNCT_1:72
                           .=g.p by A47,FUNCT_1:72;
      hence h.x=(cn-FanMorphN)|D.x by A33,A40,A47,FUNCT_4:14;
     end;
    hence h.x=(cn-FanMorphN)|D.x;
   end;
  then f+*g=(cn-FanMorphN)|D by A33,A34,A37,FUNCT_1:9;
 hence thesis by A19,A20,A28,A29,A30,A31,JGRAPH_2:9;
end;

theorem Th77: for cn being Real
  st -1<cn & cn<1 holds
  ex h being map of (TOP-REAL 2),(TOP-REAL 2)
  st h=(cn-FanMorphN) & h is continuous
proof let cn be Real;
 assume A1:  -1<cn & cn<1;
 reconsider f=(cn-FanMorphN) as map of (TOP-REAL 2),(TOP-REAL 2)
                            ;
  (the carrier of TOP-REAL 2) \{0.REAL 2} c=(the carrier of TOP-REAL 2)
                         by XBOOLE_1:36;
 then reconsider D=(the carrier of TOP-REAL 2) \{0.REAL 2}
   as non empty Subset of TOP-REAL 2 by JGRAPH_2:19;
 A2: f.(0.REAL 2)=0.REAL 2 by Th56,JGRAPH_2:11;
 A3: D`= {0.REAL 2} by JGRAPH_3:30;
 A4: for p being Point of (TOP-REAL 2)|D holds f.p<>f.(0.REAL 2)
  proof let p be Point of (TOP-REAL 2)|D;
    A5:the carrier of (TOP-REAL 2)|D=[#]((TOP-REAL 2)|D) by PRE_TOPC:12;
    A6:[#]((TOP-REAL 2)|D)=D by PRE_TOPC:def 10;
    then A7:p in the carrier of TOP-REAL 2 & not p in {0.REAL 2}
                         by A5,XBOOLE_0:def 4;
    reconsider q=p as Point of TOP-REAL 2 by A5,A6,XBOOLE_0:def 4;
    A8:not p=0.REAL 2 by A7,TARSKI:def 1;
      now per cases;
    case A9: q`1/|.q.|>=cn & q`2>=0;
      set q9= |[ |.q.|* ((q`1/|.q.|-cn)/(1-cn)),
        |.q.|*( sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))]|;
      A10:q9`2= |.q.|*( sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2)) &
      q9`1= |.q.|* ((q`1/|.q.|-cn)/(1-cn))
                                  by EUCLID:56;
        now assume A11: q9=0.REAL 2;
       A12: |.q.|<>0 by A8,TOPRNS_1:25;
        then sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2)= sqrt(1-0) by A10,A11,JGRAPH_2:
11,SQUARE_1:60,XCMPLX_1:6
       .=1 by SQUARE_1:83;
      hence contradiction by A11,A12,EUCLID:56,JGRAPH_2:11;
     end;
     hence f.p<>f.(0.REAL 2) by A1,A2,A8,A9,Th58;
    case A13: q`1/|.q.|<cn & q`2>=0;
      set q9=|[ |.q.|* ((q`1/|.q.|-cn)/(1+cn)),
         |.q.|*( sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))]|;
      A14:q9`2= |.q.|*( sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2)) &
          q9`1= |.q.|* ((q`1/|.q.|-cn)/(1+cn)) by EUCLID:56;
        now assume A15: q9=0.REAL 2;
       A16: |.q.|<>0 by A8,TOPRNS_1:25;
        then sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2)= sqrt(1-0) by A14,A15,JGRAPH_2:
11,SQUARE_1:60,XCMPLX_1:6
       .=1 by SQUARE_1:83;
      hence contradiction by A15,A16,EUCLID:56,JGRAPH_2:11;
     end;
     hence f.p<>f.(0.REAL 2) by A1,A2,A8,A13,Th58;
     case q`2<0; then f.p=p by Th56;
     hence f.p<>f.(0.REAL 2) by A8,Th56,JGRAPH_2:11;
    end;
   hence f.p<>f.(0.REAL 2);
  end;
 A17:ex h being map of (TOP-REAL 2)|D,(TOP-REAL 2)|D
  st h=(cn-FanMorphN)|D & h is continuous by A1,A3,Th76;
   for V being Subset of (TOP-REAL 2) st f.(0.REAL 2) in V & V is open
 ex W being Subset of (TOP-REAL 2) st 0.REAL 2 in W
 & W is open & f.:W c= V
  proof let V be Subset of (TOP-REAL 2);
    assume A18:f.(0.REAL 2) in V & V is open;
     A19:TOP-REAL 2=TopSpaceMetr(Euclid 2) by EUCLID:def 8;
     A20:the carrier of TOP-REAL 2=REAL 2 by EUCLID:25;
     A21:Euclid 2=MetrStruct(#REAL 2,Pitag_dist 2#) by EUCLID:def 7;
     then reconsider u0=0.REAL 2 as Point of Euclid 2 by EUCLID:25;
     consider r being real number such that A22:r>0 & Ball(u0,r) c= V
                                   by A2,A18,A19,TOPMETR:22;
     reconsider r as Real by XREAL_0:def 1;
     reconsider W1=Ball(u0,r) as Subset of TOP-REAL 2
                       by A20,A21;
     A23:u0 in W1 by A22,GOBOARD6:4;
     A24:W1 is open by GOBOARD6:6;
       f.:W1 c= W1
      proof let z be set;assume z in f.:W1;
        then consider y being set such that
        A25: y in dom f & y in W1 & z=f.y by FUNCT_1:def 12;
        reconsider q=y as Point of TOP-REAL 2 by A25;
        reconsider qy=q as Point of Euclid 2 by A21,EUCLID:25;
          z in rng f by A25,FUNCT_1:def 5;
        then reconsider qz=z as Point of TOP-REAL 2;
        reconsider pz=qz as Point of Euclid 2 by A21,EUCLID:25;
          dist(u0,qy)<r by A25,METRIC_1:12;
        then A26: |.(0.REAL 2) - q.|<r by JGRAPH_1:45;
        A27: |.q.|>=0 by TOPRNS_1:26;
          now per cases by JGRAPH_2:11;
        case q`2<=0;
         hence z in W1 by A25,Th56;
        case A28: q<>0.REAL 2 &
          (q`1/|.q.|>=cn & q`2>=0);
          then A29:(cn-FanMorphN).q= |[ |.q.|* ((q`1/|.q.|-cn)/(1-cn)) ,
          |.q.|*( sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))]| by A1,Th58;
            |.q.|<>0 by A28,TOPRNS_1:25;
          then A30: (|.q.|)^2>0 by SQUARE_1:74;
        A31: 1-cn>0 by A1,Th3;
        A32: qz`2= |.q.|*( sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))&
        qz`1= |.q.|* ((q`1/|.q.|-cn)/(1-cn))
                     by A25,A29,EUCLID:56;
        A33: (q`1/|.q.|-cn)>= 0 by A28,SQUARE_1:12;
        A34: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
          0<=(q`2)^2 by SQUARE_1:72;
        then 0+(q`1)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
        then (q`1)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A30,A34,REAL_1:73;
        then (q`1)^2/(|.q.|)^2 <= 1 by A30,XCMPLX_1:60;
        then ((q`1)/|.q.|)^2 <= 1 by SQUARE_1:69;
        then 1>=q`1/|.q.| by Th4;
        then A35: 1-cn>=q`1/|.q.|-cn by REAL_1:49;
        A36: -(1-cn)<= -0 by A31,REAL_1:50;
           -(1-cn)<= -( q`1/|.q.|-cn) by A35,REAL_1:50;
         then (-(1-cn))/(1-cn)<=(-( q`1/|.q.|-cn))/(1-cn)
                        by A31,REAL_1:73;
        then A37: -1<=(-( q`1/|.q.|-cn))/(1-cn)
                        by A31,XCMPLX_1:198;
         --(1-cn)>= -(q`1/|.q.|-cn) by A33,A36,REAL_1:50;
       then (-(q`1/|.q.|-cn))/(1-cn)<=1 by A31,REAL_2:117;
       then ((-(q`1/|.q.|-cn))/(1-cn))^2<=1^2 by A37,JGRAPH_2:7;
       then 1-((-(q`1/|.q.|-cn))/(1-cn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((q`1/|.q.|-cn))/(1-cn))^2>=0 by XCMPLX_1:188;
       then A38: 1-((q`1/|.q.|-cn)/(1-cn))^2>=0 by SQUARE_1:61;
        A39: (qz`2)^2= (|.q.|)^2*( sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))^2
                         by A32,SQUARE_1:68
         .= (|.q.|)^2*(1-((q`1/|.q.|-cn)/(1-cn))^2) by A38,SQUARE_1:def 4;
        A40: (qz`1)^2= (|.q.|)^2*((q`1/|.q.|-cn)/(1-cn))^2
                         by A32,SQUARE_1:68;
          (|.qz.|)^2=(qz`1)^2+(qz`2)^2 by JGRAPH_3:10
        .=(|.q.|)^2*(1-((q`1/|.q.|-cn)/(1-cn))^2
           +((q`1/|.q.|-cn)/(1-cn))^2) by A39,A40,XCMPLX_1:8
        .=(|.q.|)^2*1 by XCMPLX_1:27 .=(|.q.|)^2;
        then A41: sqrt((|.qz.|)^2)=|.q.| by A27,SQUARE_1:89;
            |.qz.|>=0 by TOPRNS_1:26;
          then A42: |.qz.|=|.q.| by A41,SQUARE_1:89;
            |.(0.REAL 2) +- q.|<r by A26,EUCLID:45;
          then |.- q.|<r by EUCLID:31;
          then |.q.|<r by TOPRNS_1:27;
          then |.- qz.|<r by A42,TOPRNS_1:27;
          then |.(0.REAL 2) +- qz.|<r by EUCLID:31;
          then |.(0.REAL 2) - qz.|<r by EUCLID:45;
          then dist(u0,pz)<r by JGRAPH_1:45;
         hence z in W1 by METRIC_1:12;
        case A43: q<>0.REAL 2 &
          q`1/|.q.|<cn & q`2>=0;
          then A44:(cn-FanMorphN).q= |[ |.q.|* ((q`1/|.q.|-cn)/(1+cn)),
          |.q.|*( sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))]| by A1,Th58;
            |.q.|<>0 by A43,TOPRNS_1:25;
          then A45: (|.q.|)^2>0 by SQUARE_1:74;
        A46: (cn-q`1/|.q.|)>=0 by A43,SQUARE_1:12;
        A47: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
        A48: 1+cn>0 by A1,Th3;
        A49: qz`2= |.q.|*( sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))&
        qz`1= |.q.|* ((q`1/|.q.|-cn)/(1+cn))
                     by A25,A44,EUCLID:56;
          0<=(q`2)^2 by SQUARE_1:72;
        then 0+(q`1)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
        then (q`1)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A45,A47,REAL_1:73;
        then (q`1)^2/(|.q.|)^2 <= 1 by A45,XCMPLX_1:60;
        then ((q`1)/|.q.|)^2 <= 1 by SQUARE_1:69;
        then -1<=q`1/|.q.| by Th4;
        then --1>=-q`1/|.q.| by REAL_1:50;
        then A50: 1+cn>=-q`1/|.q.|+cn by REAL_1:55;
          -(1+cn)<= -0 by A48,REAL_1:50;
         then -(1+cn)<= -( q`1/|.q.|-cn) by A46,XCMPLX_1:143;
         then (-(1+cn))/(1+cn)<=(-( q`1/|.q.|-cn))/(1+cn)
                        by A48,REAL_1:73;
        then A51: -1<=(-( q`1/|.q.|-cn))/(1+cn)
                        by A48,XCMPLX_1:198;
          -(1+cn)<=-(-q`1/|.q.|+cn) by A50,REAL_1:50;
        then -(1+cn)<=--(q`1/|.q.|)-cn by XCMPLX_1:161;
        then --(1+cn)>= -(q`1/|.q.|-cn) by REAL_1:50;
       then (-(q`1/|.q.|-cn))/(1+cn)<=1 by A48,REAL_2:117;
       then ((-(q`1/|.q.|-cn))/(1+cn))^2<=1^2 by A51,JGRAPH_2:7;
       then 1-((-(q`1/|.q.|-cn))/(1+cn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((q`1/|.q.|-cn))/(1+cn))^2>=0 by XCMPLX_1:188;
       then A52: 1-((q`1/|.q.|-cn)/(1+cn))^2>=0 by SQUARE_1:61;
        A53: (qz`2)^2= (|.q.|)^2*( sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))^2
                         by A49,SQUARE_1:68
         .= (|.q.|)^2*(1-((q`1/|.q.|-cn)/(1+cn))^2) by A52,SQUARE_1:def 4;
        A54: (qz`1)^2= (|.q.|)^2*((q`1/|.q.|-cn)/(1+cn))^2
                         by A49,SQUARE_1:68;
          (|.qz.|)^2=(qz`1)^2+(qz`2)^2 by JGRAPH_3:10
        .=(|.q.|)^2*(1-((q`1/|.q.|-cn)/(1+cn))^2
           +((q`1/|.q.|-cn)/(1+cn))^2) by A53,A54,XCMPLX_1:8
        .=(|.q.|)^2*1 by XCMPLX_1:27 .=(|.q.|)^2;
        then A55: sqrt((|.qz.|)^2)=|.q.| by A27,SQUARE_1:89;
            |.qz.|>=0 by TOPRNS_1:26;
          then A56: |.qz.|=|.q.| by A55,SQUARE_1:89;
            |.(0.REAL 2) +- q.|<r by A26,EUCLID:45;
          then |.- q.|<r by EUCLID:31;
          then |.q.|<r by TOPRNS_1:27;
          then |.- qz.|<r by A56,TOPRNS_1:27;
          then |.(0.REAL 2) +- qz.|<r by EUCLID:31;
          then |.(0.REAL 2) - qz.|<r by EUCLID:45;
          then dist(u0,pz)<r by JGRAPH_1:45;
         hence z in W1 by METRIC_1:12;
        end;
       hence z in W1;
      end;
     then f.:W1 c= V by A22,XBOOLE_1:1;
   hence ex W being Subset of (TOP-REAL 2) st 0.REAL 2 in W
       & W is open & f.:W c= V by A23,A24;
  end;
  then f is continuous by A2,A3,A4,A17,JGRAPH_3:13;
 hence thesis;
end;

theorem Th78: for cn being Real
  st -1<cn & cn<1 holds
  (cn-FanMorphN) is one-to-one
proof let cn be Real;
 assume A1: -1<cn & cn<1;
    for x1,x2 being set st x1 in dom (cn-FanMorphN) & x2 in dom (cn-FanMorphN)
   & (cn-FanMorphN).x1=(cn-FanMorphN).x2 holds x1=x2
  proof let x1,x2 be set;
   assume A2: x1 in dom (cn-FanMorphN) & x2 in dom (cn-FanMorphN)
    & (cn-FanMorphN).x1=(cn-FanMorphN).x2;
    then reconsider p1=x1 as Point of TOP-REAL 2 by FUNCT_2:def 1;
    reconsider p2=x2 as Point of TOP-REAL 2 by A2,FUNCT_2:def 1;
    set q=p1,p=p2;
    A3: 1-cn>0 by A1,Th3;
      now per cases by JGRAPH_2:11;
    case A4: q`2<=0;
      then A5:(cn-FanMorphN).q=q by Th56;
        now per cases by JGRAPH_2:11;
      case p`2<=0;
       hence x1=x2 by A2,A5,Th56;
      case A6:p<>0.REAL 2 &
          (p`1/|.p.|>=cn & p`2>=0);
    then A7:cn-FanMorphN.p= |[ |.p.|* ((p`1/|.p.|-cn)/(1-cn)),
           |.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2))]| by A1,Th58;
         set p4= |[ |.p.|* ((p`1/|.p.|-cn)/(1-cn)),
           |.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2))]|;
        A8: p4`2= |.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2)) &
          p4`1= |.p.|* ((p`1/|.p.|-cn)/(1-cn)) by EUCLID:56;
     A9: |.p.|>=0 by TOPRNS_1:26;
     A10: |.p.|>0 by A6,Lm1;
     then A11: (|.p.|)^2>0 by SQUARE_1:74;
        A12: (p`1/|.p.|-cn)>=0 by A6,SQUARE_1:12;
        A13: (p`1/|.p.|-cn)>= 0 by A6,SQUARE_1:12;
        A14: (p`1/|.p.|-cn)/(1-cn)>=0 by A3,A12,REAL_2:125;
        A15: (|.p.|)^2 =(p`1)^2+(p`2)^2 by JGRAPH_3:10;
          0<=(p`2)^2 by SQUARE_1:72;
        then 0+(p`1)^2<=(p`1)^2+(p`2)^2 by REAL_1:55;
        then (p`1)^2/(|.p.|)^2 <= (|.p.|)^2/(|.p.|)^2 by A11,A15,REAL_1:73;
        then (p`1)^2/(|.p.|)^2 <= 1 by A11,XCMPLX_1:60;
        then ((p`1)/|.p.|)^2 <= 1 by SQUARE_1:69;
        then 1>=p`1/|.p.| by Th4;
        then A16: 1-cn>=p`1/|.p.|-cn by REAL_1:49;
        A17: -(1-cn)<= -0 by A3,REAL_1:50;
           -(1-cn)<= -( p`1/|.p.|-cn) by A16,REAL_1:50;
         then (-(1-cn))/(1-cn)<=(-( p`1/|.p.|-cn))/(1-cn)
                        by A3,REAL_1:73;
        then A18: -1<=(-( p`1/|.p.|-cn))/(1-cn) by A3,XCMPLX_1:198;
         --(1-cn)>= -(p`1/|.p.|-cn) by A13,A17,REAL_1:50;
       then (-(p`1/|.p.|-cn))/(1-cn)<=1 by A3,REAL_2:117;
       then ((-(p`1/|.p.|-cn))/(1-cn))^2<=1^2 by A18,JGRAPH_2:7;
       then A19: 1-((-(p`1/|.p.|-cn))/(1-cn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((p`1/|.p.|-cn))/(1-cn))^2>=0 by XCMPLX_1:188;
       then A20: 1-((p`1/|.p.|-cn)/(1-cn))^2>=0 by SQUARE_1:61;
         sqrt(1-((-(p`1/|.p.|-cn))/(1-cn))^2)>=0 by A19,SQUARE_1:def 4;
       then sqrt(1-(-(p`1/|.p.|-cn))^2/(1-cn)^2)>=0 by SQUARE_1:69;
       then sqrt(1-(p`1/|.p.|-cn)^2/(1-cn)^2)>=0 by SQUARE_1:61;
       then sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2)>=0 by SQUARE_1:69;
        then q`2=0 by A2,A4,A5,A7,A8,A9,SQUARE_1:19;
        then ( sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2))=0 by A2,A5,A7,A8,A10,XCMPLX_1
:6;
        then 1-((p`1/|.p.|-cn)/(1-cn))^2=0 by A20,SQUARE_1:92;
        then 1=0+((p`1/|.p.|-cn)/(1-cn))^2 by XCMPLX_1:27;
        then 1= (p`1/|.p.|-cn)/(1-cn) by A14,SQUARE_1:83,89;
        then (1)*(1-cn)=(p`1/|.p.|-cn) by A3,XCMPLX_1:88;
        then 1-cn+cn=p`1/|.p.| by XCMPLX_1:27;
        then 1=p`1/|.p.| by XCMPLX_1:27;
        then (1)*|.p.|=p`1 by A10,XCMPLX_1:88;
        then (p`1)^2-(p`1)^2 =(p`2)^2 by A15,XCMPLX_1:26;
        then 0= (p`2)^2 by XCMPLX_1:14;
        then p`2=0 by SQUARE_1:73;
       hence x1=x2 by A2,A5,Th56;
      case A21:p<>0.REAL 2 &
        (p`1/|.p.|<cn & p`2>=0);
    then A22: cn-FanMorphN.p= |[ |.p.|* ((p`1/|.p.|-cn)/(1+cn)),
           |.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2))]| by A1,Th58;
         set p4= |[ |.p.|* ((p`1/|.p.|-cn)/(1+cn)),
           |.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2))]|;
        A23: p4`2= |.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2)) &
          p4`1= |.p.|* ((p`1/|.p.|-cn)/(1+cn)) by EUCLID:56;
     A24: |.p.|>=0 by TOPRNS_1:26;
     A25: |.p.|<>0 by A21,TOPRNS_1:25;
     then A26: (|.p.|)^2>0 by SQUARE_1:74;
     A27: 1+cn>0 by A1,Th3;
        A28: (p`1/|.p.|-cn)<=0 by A21,REAL_2:106;
        then A29: -(p`1/|.p.|-cn)>= -0 by REAL_1:50;
          --(p`1/|.p.|-cn)/(1+cn)<=0 by A27,A28,REAL_2:126;
        then A30: -((p`1/|.p.|-cn)/(1+cn))>=0 by REAL_1:66;
        A31: (|.p.|)^2 =(p`1)^2+(p`2)^2 by JGRAPH_3:10;
          0<=(p`2)^2 by SQUARE_1:72;
        then 0+(p`1)^2<=(p`1)^2+(p`2)^2 by REAL_1:55;
        then (p`1)^2/(|.p.|)^2 <= (|.p.|)^2/(|.p.|)^2 by A26,A31,REAL_1:73;
        then (p`1)^2/(|.p.|)^2 <= 1 by A26,XCMPLX_1:60;
        then ((p`1)/|.p.|)^2 <= 1 by SQUARE_1:69;
        then A32: (-((p`1)/|.p.|))^2 <= 1 by SQUARE_1:61;
          -(1+cn)<= -0 by A27,REAL_1:50;
        then (-(1+cn))/(1+cn)<=(-( p`1/|.p.|-cn))/(1+cn)
                        by A27,A29,REAL_1:73;
        then A33: -1<=(-( p`1/|.p.|-cn))/(1+cn)
                        by A27,XCMPLX_1:198;
         1>= -p`1/|.p.| by A32,Th4;
       then (1+cn)>= -p`1/|.p.|+cn by REAL_1:55;
       then (1+cn)>= -(p`1/|.p.|-cn) by XCMPLX_1:162;
       then (-(p`1/|.p.|-cn))/(1+cn)<=1 by A27,REAL_2:117;
       then ((-(p`1/|.p.|-cn))/(1+cn))^2<=1^2 by A33,JGRAPH_2:7;
       then A34: 1-((-(p`1/|.p.|-cn))/(1+cn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((p`1/|.p.|-cn))/(1+cn))^2>=0 by XCMPLX_1:188;
       then A35: 1-((p`1/|.p.|-cn)/(1+cn))^2>=0 by SQUARE_1:61;
         sqrt(1-((-(p`1/|.p.|-cn))/(1+cn))^2)>=0 by A34,SQUARE_1:def 4;
       then sqrt(1-(-(p`1/|.p.|-cn))^2/(1+cn)^2)>=0 by SQUARE_1:69;
       then sqrt(1-((p`1/|.p.|-cn))^2/(1+cn)^2)>=0 by SQUARE_1:61;
       then sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2)>=0 by SQUARE_1:69;
        then q`2=0 by A2,A4,A5,A22,A23,A24,SQUARE_1:19;
        then ( sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2))=0 by A2,A5,A22,A23,A25,
XCMPLX_1:6;
        then 1-((p`1/|.p.|-cn)/(1+cn))^2=0 by A35,SQUARE_1:92;
        then 1=0+((p`1/|.p.|-cn)/(1+cn))^2 by XCMPLX_1:27;
        then 1=sqrt((-((p`1/|.p.|-cn)/(1+cn)))^2) by SQUARE_1:61,83;
        then 1= -((p`1/|.p.|-cn)/(1+cn)) by A30,SQUARE_1:89;
        then 1= ((-(p`1/|.p.|-cn))/(1+cn)) by XCMPLX_1:188;
        then (1)*(1+cn)=-(p`1/|.p.|-cn) by A27,XCMPLX_1:88;
        then (1+cn)=-p`1/|.p.|+cn by XCMPLX_1:162;
        then 1+cn-cn=-p`1/|.p.| by XCMPLX_1:26;
        then 1=-p`1/|.p.| by XCMPLX_1:26;
        then 1=(-p`1)/|.p.| by XCMPLX_1:188;
        then (1)*|.p.|=-p`1 by A25,XCMPLX_1:88;
        then (-p`1)^2-(p`1)^2 =(p`2)^2 by A31,XCMPLX_1:26;
        then (p`1)^2-(p`1)^2 =(p`2)^2 by SQUARE_1:61;
        then 0= (p`2)^2 by XCMPLX_1:14;
        then p`2=0 by SQUARE_1:73;
       hence x1=x2 by A2,A5,Th56;
      end;
     hence x1=x2;
    case A36:
      q`1/|.q.|>=cn & q`2>=0 & q<>0.REAL 2;
        then A37:
         cn-FanMorphN.q= |[ |.q.|* ((q`1/|.q.|-cn)/(1-cn)),
           |.q.|*( sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))]| by A1,Th58;
         set q4= |[ |.q.|* ((q`1/|.q.|-cn)/(1-cn)),
           |.q.|*( sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))]|;
        A38: q4`2= |.q.|*( sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2)) &
          q4`1= |.q.|* ((q`1/|.q.|-cn)/(1-cn)) by EUCLID:56;
     A39: |.q.|>=0 by TOPRNS_1:26;
       |.q.|<>0 by A36,TOPRNS_1:25;
     then A40: (|.q.|)^2>0 by SQUARE_1:74;
        now per cases by JGRAPH_2:11;
      case A41: p`2<=0;
        then A42: (cn-FanMorphN).p=p by Th56;
        A43: |.q.|>=0 by TOPRNS_1:26;
        A44: |.q.|<>0 by A36,TOPRNS_1:25;
        then A45: (|.q.|)^2>0 by SQUARE_1:74;
        A46: 1-cn>0 by A1,Th3;
        A47: (q`1/|.q.|-cn)>=0 by A36,SQUARE_1:12;
        A48: (q`1/|.q.|-cn)>= 0 by A36,SQUARE_1:12;
        A49: (q`1/|.q.|-cn)/(1-cn)>=0 by A46,A47,REAL_2:125;
        A50: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
          0<=(q`2)^2 by SQUARE_1:72;
        then 0+(q`1)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
        then (q`1)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A45,A50,REAL_1:73;
        then (q`1)^2/(|.q.|)^2 <= 1 by A45,XCMPLX_1:60;
        then ((q`1)/|.q.|)^2 <= 1 by SQUARE_1:69;
        then 1>=q`1/|.q.| by Th4;
        then A51: 1-cn>=q`1/|.q.|-cn by REAL_1:49;
        A52: -(1-cn)<= -0 by A46,REAL_1:50;
           -(1-cn)<= -( q`1/|.q.|-cn) by A51,REAL_1:50;
         then (-(1-cn))/(1-cn)<=(-( q`1/|.q.|-cn))/(1-cn)
                        by A46,REAL_1:73;
        then A53: -1<=(-( q`1/|.q.|-cn))/(1-cn) by A46,XCMPLX_1:198;
         --(1-cn)>= -(q`1/|.q.|-cn) by A48,A52,REAL_1:50;
       then (-(q`1/|.q.|-cn))/(1-cn)<=1 by A46,REAL_2:117;
       then ((-(q`1/|.q.|-cn))/(1-cn))^2<=1^2 by A53,JGRAPH_2:7;
       then A54: 1-((-(q`1/|.q.|-cn))/(1-cn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((q`1/|.q.|-cn))/(1-cn))^2>=0 by XCMPLX_1:188;
       then A55: 1-((q`1/|.q.|-cn)/(1-cn))^2>=0 by SQUARE_1:61;
         sqrt(1-((-(q`1/|.q.|-cn))/(1-cn))^2)>=0 by A54,SQUARE_1:def 4;
       then sqrt(1-(-(q`1/|.q.|-cn))^2/(1-cn)^2)>=0 by SQUARE_1:69;
       then sqrt(1-(q`1/|.q.|-cn)^2/(1-cn)^2)>=0 by SQUARE_1:61;
       then sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2)>=0 by SQUARE_1:69;
        then p`2=0 by A2,A37,A38,A41,A42,A43,SQUARE_1:19;
        then ( sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))=0 by A2,A37,A38,A42,A44,
XCMPLX_1:6;
        then 1-((q`1/|.q.|-cn)/(1-cn))^2=0 by A55,SQUARE_1:92;
        then 1=0+((q`1/|.q.|-cn)/(1-cn))^2 by XCMPLX_1:27;
        then 1= (q`1/|.q.|-cn)/(1-cn) by A49,SQUARE_1:83,89;
        then (1)*(1-cn)=(q`1/|.q.|-cn) by A46,XCMPLX_1:88;
        then 1-cn+cn=q`1/|.q.| by XCMPLX_1:27;
        then 1=q`1/|.q.| by XCMPLX_1:27;
        then (1)*|.q.|=q`1 by A44,XCMPLX_1:88;
        then (q`1)^2-(q`1)^2 =(q`2)^2 by A50,XCMPLX_1:26;
        then 0= (q`2)^2 by XCMPLX_1:14;
        then q`2=0 by SQUARE_1:73;
       hence x1=x2 by A2,A42,Th56;
      case A56:p<>0.REAL 2 &
          (p`1/|.p.|>=cn & p`2>=0);
       then A57: cn-FanMorphN.p= |[ |.p.|* ((p`1/|.p.|-cn)/(1-cn)),
           |.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2))]| by A1,Th58;
         set p4= |[ |.p.|* ((p`1/|.p.|-cn)/(1-cn)),
           |.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2))]|;
        A58: p4`2= |.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2)) &
          p4`1= |.p.|* ((p`1/|.p.|-cn)/(1-cn)) by EUCLID:56;
        A59: |.q4.|>=0 by TOPRNS_1:26;
        A60: q4`2= |.q.|*( sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))&
        q4`1= |.q.|* ((q`1/|.q.|-cn)/(1-cn)) by EUCLID:56;
        A61: (q`1/|.q.|-cn)>= 0 by A36,SQUARE_1:12;
        A62: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
          0<=(q`2)^2 by SQUARE_1:72;
        then 0+(q`1)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
        then (q`1)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A40,A62,REAL_1:73;
        then (q`1)^2/(|.q.|)^2 <= 1 by A40,XCMPLX_1:60;
        then ((q`1)/|.q.|)^2 <= 1 by SQUARE_1:69;
        then 1>=q`1/|.q.| by Th4;
        then A63: 1-cn>=q`1/|.q.|-cn by REAL_1:49;
        A64: -(1-cn)<= -0 by A3,REAL_1:50;
           -(1-cn)<= -( q`1/|.q.|-cn) by A63,REAL_1:50;
         then (-(1-cn))/(1-cn)<=(-( q`1/|.q.|-cn))/(1-cn)
                        by A3,REAL_1:73;
        then A65: -1<=(-( q`1/|.q.|-cn))/(1-cn) by A3,XCMPLX_1:198;
         --(1-cn)>= -(q`1/|.q.|-cn) by A61,A64,REAL_1:50;
       then (-(q`1/|.q.|-cn))/(1-cn)<=1 by A3,REAL_2:117;
       then ((-(q`1/|.q.|-cn))/(1-cn))^2<=1^2 by A65,JGRAPH_2:7;
       then 1-((-(q`1/|.q.|-cn))/(1-cn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((q`1/|.q.|-cn))/(1-cn))^2>=0 by XCMPLX_1:188;
       then A66: 1-((q`1/|.q.|-cn)/(1-cn))^2>=0 by SQUARE_1:61;
        A67: (q4`2)^2= (|.q.|)^2*( sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))^2
                         by A60,SQUARE_1:68
         .= (|.q.|)^2*(1-((q`1/|.q.|-cn)/(1-cn))^2) by A66,SQUARE_1:def 4;
        A68: (q4`1)^2= (|.q.|)^2*((q`1/|.q.|-cn)/(1-cn))^2
                         by A60,SQUARE_1:68;
          (|.q4.|)^2=(q4`1)^2+(q4`2)^2 by JGRAPH_3:10
        .=(|.q.|)^2*(1-((q`1/|.q.|-cn)/(1-cn))^2
           +((q`1/|.q.|-cn)/(1-cn))^2) by A67,A68,XCMPLX_1:8
        .=(|.q.|)^2*1 by XCMPLX_1:27 .=(|.q.|)^2;
        then A69: sqrt((|.q4.|)^2)=|.q.| by A39,SQUARE_1:89;
        then A70: |.q4.|=|.q.| by A59,SQUARE_1:89;
        A71: |.p4.|>=0 by TOPRNS_1:26;
        A72: |.p.|>=0 by TOPRNS_1:26;
        A73: |.p.|<>0 by A56,TOPRNS_1:25;
        then A74: (|.p.|)^2>0 by SQUARE_1:74;
        A75: (p`1/|.p.|-cn)>= 0 by A56,SQUARE_1:12;
        A76: (|.p.|)^2 =(p`1)^2+(p`2)^2 by JGRAPH_3:10;
          0<=(p`2)^2 by SQUARE_1:72;
        then 0+(p`1)^2<=(p`1)^2+(p`2)^2 by REAL_1:55;
        then (p`1)^2/(|.p.|)^2 <= (|.p.|)^2/(|.p.|)^2 by A74,A76,REAL_1:73;
        then (p`1)^2/(|.p.|)^2 <= 1 by A74,XCMPLX_1:60;
        then ((p`1)/|.p.|)^2 <= 1 by SQUARE_1:69;
        then 1>=p`1/|.p.| by Th4;
        then A77: 1-cn>=p`1/|.p.|-cn by REAL_1:49;
        A78: -(1-cn)<= -0 by A3,REAL_1:50;
           -(1-cn)<= -( p`1/|.p.|-cn) by A77,REAL_1:50;
         then (-(1-cn))/(1-cn)<=(-( p`1/|.p.|-cn))/(1-cn)
                        by A3,REAL_1:73;
        then A79: -1<=(-( p`1/|.p.|-cn))/(1-cn) by A3,XCMPLX_1:198;
         --(1-cn)>= -(p`1/|.p.|-cn) by A75,A78,REAL_1:50;
       then (-(p`1/|.p.|-cn))/(1-cn)<=1 by A3,REAL_2:117;
       then ((-(p`1/|.p.|-cn))/(1-cn))^2<=1^2 by A79,JGRAPH_2:7;
       then 1-((-(p`1/|.p.|-cn))/(1-cn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((p`1/|.p.|-cn))/(1-cn))^2>=0 by XCMPLX_1:188;
       then A80: 1-((p`1/|.p.|-cn)/(1-cn))^2>=0 by SQUARE_1:61;
        A81: (p4`2)^2= (|.p.|)^2*( sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2))^2
                         by A58,SQUARE_1:68
         .= (|.p.|)^2*(1-((p`1/|.p.|-cn)/(1-cn))^2) by A80,SQUARE_1:def 4;
        A82: (p4`1)^2= (|.p.|)^2*((p`1/|.p.|-cn)/(1-cn))^2
                         by A58,SQUARE_1:68;
          (|.p4.|)^2=(p4`1)^2+(p4`2)^2 by JGRAPH_3:10
        .=(|.p.|)^2*(1-((p`1/|.p.|-cn)/(1-cn))^2
           +((p`1/|.p.|-cn)/(1-cn))^2) by A81,A82,XCMPLX_1:8
        .=(|.p.|)^2*1 by XCMPLX_1:27 .=(|.p.|)^2;
        then A83: sqrt((|.p4.|)^2)=|.p.| by A72,SQUARE_1:89;
        then A84: |.p4.|=|.p.| by A71,SQUARE_1:89;
        ((p`1/|.p.|-cn)/(1-cn))
        =|.q.|* ((q`1/|.q.|-cn)/(1-cn))/|.p.| by A2,A37,A38,A57,A58,A73,
XCMPLX_1:90;
        then (p`1/|.p.|-cn)/(1-cn)=(q`1/|.q.|-cn)/(1-cn) by A2,A37,A57,A69,A73,
A83,XCMPLX_1:90;
        then (p`1/|.p.|-cn)/(1-cn)*(1-cn)=q`1/|.q.|-cn by A3,XCMPLX_1:88;
        then p`1/|.p.|-cn=q`1/|.q.|-cn by A3,XCMPLX_1:88;
        then p`1/|.p.|-cn+cn=q`1/|.q.| by XCMPLX_1:27;
        then p`1/|.p.|=q`1/|.q.| by XCMPLX_1:27;
        then p`1/|.p.|*|.p.|=q`1 by A2,A37,A57,A70,A73,A84,XCMPLX_1:88;
        then A85: p`1=q`1 by A73,XCMPLX_1:88;
        A86: |.p.|^2=(p`1)^2+(p`2)^2 by JGRAPH_3:10;
          |.q.|^2=(q`1)^2+(q`2)^2 by JGRAPH_3:10;
        then (p`1)^2+(p`2)^2-(p`1)^2=(q`2)^2 by A2,A37,A57,A70,A84,A85,A86,
XCMPLX_1:26;
        then (p`2)^2=(q`2)^2 by XCMPLX_1:26;
        then p`2=sqrt((q`2)^2) by A56,SQUARE_1:89;
        then A87: p`2=q`2 by A36,SQUARE_1:89;
          p=|[p`1,p`2]| by EUCLID:57;
       hence x1=x2 by A85,A87,EUCLID:57;
      case A88: p<>0.REAL 2 &
        (p`1/|.p.|<cn & p`2>=0);
        then A89: cn-FanMorphN.p= |[ |.p.|* ((p`1/|.p.|-cn)/(1+cn)),
           |.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2))]| by A1,Th58;
         set p4= |[ |.p.|* ((p`1/|.p.|-cn)/(1+cn)),
           |.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2))]|;
        A90: p4`2= |.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2)) &
          p4`1= |.p.|* ((p`1/|.p.|-cn)/(1+cn)) by EUCLID:56;
         A91: p`1/|.p.|-cn<0 by A88,REAL_2:105;
           1+cn>0 by A1,Th3;
         then A92: ((p`1/|.p.|-cn)/(1+cn))<0 by A91,REAL_2:128;
        A93: |.p.|>=0 by TOPRNS_1:26;
          |.p.|<>0 by A88,TOPRNS_1:25;
         then A94: |.p.|* ((p`1/|.p.|-cn)/(1+cn))<0 by A92,A93,SQUARE_1:24;
         A95: q`1/|.q.|-cn>=0 by A36,SQUARE_1:12;
           1-cn>0 by A1,Th3;
         then ((q`1/|.q.|-cn)/(1-cn))>=0 by A95,REAL_2:125;
       hence x1=x2 by A2,A37,A38,A39,A89,A90,A94,REAL_2:121;
      end;
     hence x1=x2;
    case A96:
      q`1/|.q.|<cn & q`2>=0 & q<>0.REAL 2;
        then A97: cn-FanMorphN.q= |[ |.q.|* ((q`1/|.q.|-cn)/(1+cn)),
           |.q.|*( sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))]| by A1,Th58;
         set q4= |[ |.q.|* ((q`1/|.q.|-cn)/(1+cn)),
           |.q.|*( sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))]|;
        A98: q4`2= |.q.|*( sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2)) &
          q4`1= |.q.|* ((q`1/|.q.|-cn)/(1+cn)) by EUCLID:56;
     A99: |.q.|>=0 by TOPRNS_1:26;
     A100: |.q.|<>0 by A96,TOPRNS_1:25;
     then A101: (|.q.|)^2>0 by SQUARE_1:74;
        now per cases by JGRAPH_2:11;
      case A102: p`2<=0;
        then A103:(cn-FanMorphN).p=p by Th56;
        A104: 1+cn>0 by A1,Th3;
        A105: (q`1/|.q.|-cn)<=0 by A96,REAL_2:106;
        then A106: -(q`1/|.q.|-cn)>= -0 by REAL_1:50;
          --(q`1/|.q.|-cn)/(1+cn)<=0 by A104,A105,REAL_2:126;
        then A107: -((q`1/|.q.|-cn)/(1+cn))>=0 by REAL_1:66;
        A108: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
          0<=(q`2)^2 by SQUARE_1:72;
        then 0+(q`1)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
        then (q`1)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A101,A108,REAL_1:73;
        then (q`1)^2/(|.q.|)^2 <= 1 by A101,XCMPLX_1:60;
        then ((q`1)/|.q.|)^2 <= 1 by SQUARE_1:69;
        then A109: (-((q`1)/|.q.|))^2 <= 1 by SQUARE_1:61;
          -(1+cn)<= -0 by A104,REAL_1:50;
        then (-(1+cn))/(1+cn)<=(-( q`1/|.q.|-cn))/(1+cn) by A104,A106,REAL_1:73
;
        then A110: -1<=(-( q`1/|.q.|-cn))/(1+cn) by A104,XCMPLX_1:198;
         1>= -q`1/|.q.| by A109,Th4;
       then (1+cn)>= -q`1/|.q.|+cn by REAL_1:55;
       then (1+cn)>= -(q`1/|.q.|-cn) by XCMPLX_1:162;
       then (-(q`1/|.q.|-cn))/(1+cn)<=1 by A104,REAL_2:117;
       then ((-(q`1/|.q.|-cn))/(1+cn))^2<=1^2 by A110,JGRAPH_2:7;
       then A111: 1-((-(q`1/|.q.|-cn))/(1+cn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((q`1/|.q.|-cn))/(1+cn))^2>=0 by XCMPLX_1:188;
       then A112: 1-((q`1/|.q.|-cn)/(1+cn))^2>=0 by SQUARE_1:61;
         sqrt(1-((-(q`1/|.q.|-cn))/(1+cn))^2)>=0 by A111,SQUARE_1:def 4;
       then sqrt(1-(-(q`1/|.q.|-cn))^2/(1+cn)^2)>=0 by SQUARE_1:69;
       then sqrt(1-((q`1/|.q.|-cn))^2/(1+cn)^2)>=0 by SQUARE_1:61;
       then sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2)>=0 by SQUARE_1:69;
        then p`2=0 by A2,A97,A98,A99,A102,A103,SQUARE_1:19;
        then ( sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))=0 by A2,A97,A98,A100,A103,
XCMPLX_1:6;
        then 1-((q`1/|.q.|-cn)/(1+cn))^2=0 by A112,SQUARE_1:92;
        then 1=0+((q`1/|.q.|-cn)/(1+cn))^2 by XCMPLX_1:27;
        then 1=sqrt((-((q`1/|.q.|-cn)/(1+cn)))^2) by SQUARE_1:61,83;
        then 1= -((q`1/|.q.|-cn)/(1+cn)) by A107,SQUARE_1:89;
        then 1= ((-(q`1/|.q.|-cn))/(1+cn)) by XCMPLX_1:188;
        then (1)*(1+cn)=-(q`1/|.q.|-cn) by A104,XCMPLX_1:88;
        then (1+cn)=-q`1/|.q.|+cn by XCMPLX_1:162;
        then 1+cn-cn=-q`1/|.q.| by XCMPLX_1:26;
        then 1=-q`1/|.q.| by XCMPLX_1:26;
        then 1=(-q`1)/|.q.| by XCMPLX_1:188;
        then (1)*|.q.|=-q`1 by A100,XCMPLX_1:88;
        then (-q`1)^2-(q`1)^2 =(q`2)^2 by A108,XCMPLX_1:26;
        then (q`1)^2-(q`1)^2 =(q`2)^2 by SQUARE_1:61;
        then 0= (q`2)^2 by XCMPLX_1:14;
        then q`2=0 by SQUARE_1:73;
       hence x1=x2 by A2,A103,Th56;
      case A113:p<>0.REAL 2 &
          (p`1/|.p.|>=cn & p`2>=0);
        then A114: cn-FanMorphN.p= |[ |.p.|* ((p`1/|.p.|-cn)/(1-cn)),
           |.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2))]| by A1,Th58;
         set p4= |[ |.p.|* ((p`1/|.p.|-cn)/(1-cn)),
           |.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2))]|;
        A115: p4`2= |.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2)) &
          p4`1= |.p.|* ((p`1/|.p.|-cn)/(1-cn)) by EUCLID:56;
         A116: q`1/|.q.|-cn<0 by A96,REAL_2:105;
           1+cn>0 by A1,Th3;
         then A117: ((q`1/|.q.|-cn)/(1+cn))<0 by A116,REAL_2:128;
        A118: |.q.|>=0 by TOPRNS_1:26;
          |.q.|<>0 by A96,TOPRNS_1:25;
         then A119: |.q.|* ((q`1/|.q.|-cn)/(1+cn))<0 by A117,A118,SQUARE_1:24;
         A120: p`1/|.p.|-cn>=0 by A113,SQUARE_1:12;
           1-cn>0 by A1,Th3;
         then A121: ((p`1/|.p.|-cn)/(1-cn))>=0 by A120,REAL_2:125;
           |.p.|>=0 by TOPRNS_1:26;
       hence x1=x2 by A2,A97,A98,A114,A115,A119,A121,REAL_2:121;
      case A122:p<>0.REAL 2 &
          (p`1/|.p.|<cn & p`2>=0);
        then A123:
         cn-FanMorphN.p= |[ |.p.|* ((p`1/|.p.|-cn)/(1+cn)),
           |.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2))]| by A1,Th58;
         set p4= |[ |.p.|* ((p`1/|.p.|-cn)/(1+cn)),
           |.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2))]|;
        A124: p4`2= |.p.|*( sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2)) &
          p4`1= |.p.|* ((p`1/|.p.|-cn)/(1+cn)) by EUCLID:56;
        A125: |.q4.|>=0 by TOPRNS_1:26;
        A126: q4`2= |.q.|*( sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))&
        q4`1= |.q.|* ((q`1/|.q.|-cn)/(1+cn)) by EUCLID:56;
          (q`1/|.q.|-cn)<=0 by A96,REAL_2:106;
        then A127: -(q`1/|.q.|-cn)>= -0 by REAL_1:50;
        A128: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
          0<=(q`2)^2 by SQUARE_1:72;
        then 0+(q`1)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
        then (q`1)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A101,A128,REAL_1:73;
        then (q`1)^2/(|.q.|)^2 <= 1 by A101,XCMPLX_1:60;
        then ((q`1)/|.q.|)^2 <= 1 by SQUARE_1:69;
        then -1<=q`1/|.q.| by Th4;
        then A129: -1-cn<=q`1/|.q.|-cn by REAL_1:49;
        A130: 1+cn>0 by A1,Th3;
        then -(1+cn)<= -0 by REAL_1:50;
        then (-(1+cn))/(1+cn)<=(-( q`1/|.q.|-cn))/(1+cn)
                        by A127,A130,REAL_1:73;
        then A131: -1<=(-( q`1/|.q.|-cn))/(1+cn) by A130,XCMPLX_1:198;
          -(1+cn)<=q`1/|.q.|-cn by A129,XCMPLX_1:161;
        then --(1+cn)>= -(q`1/|.q.|-cn) by REAL_1:50;
       then (-(q`1/|.q.|-cn))/(1+cn)<=1 by A130,REAL_2:117;
       then ((-(q`1/|.q.|-cn))/(1+cn))^2<=1^2 by A131,JGRAPH_2:7;
       then 1-((-(q`1/|.q.|-cn))/(1+cn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((q`1/|.q.|-cn))/(1+cn))^2>=0 by XCMPLX_1:188;
       then A132: 1-((q`1/|.q.|-cn)/(1+cn))^2>=0 by SQUARE_1:61;
        A133: (q4`2)^2= (|.q.|)^2*( sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))^2
                         by A126,SQUARE_1:68
         .= (|.q.|)^2*(1-((q`1/|.q.|-cn)/(1+cn))^2) by A132,SQUARE_1:def 4;
        A134: (q4`1)^2= (|.q.|)^2*((q`1/|.q.|-cn)/(1+cn))^2
                         by A126,SQUARE_1:68;
          (|.q4.|)^2=(q4`1)^2+(q4`2)^2 by JGRAPH_3:10
        .=(|.q.|)^2*(1-((q`1/|.q.|-cn)/(1+cn))^2
           +((q`1/|.q.|-cn)/(1+cn))^2) by A133,A134,XCMPLX_1:8
        .=(|.q.|)^2*1 by XCMPLX_1:27 .=(|.q.|)^2;
        then A135: sqrt((|.q4.|)^2)=|.q.| by A99,SQUARE_1:89;
        then A136: |.q4.|=|.q.| by A125,SQUARE_1:89;
        A137: |.p4.|>=0 by TOPRNS_1:26;
        A138: |.p.|>=0 by TOPRNS_1:26;
        A139: |.p.|<>0 by A122,TOPRNS_1:25;
        then A140: (|.p.|)^2>0 by SQUARE_1:74;
          (p`1/|.p.|-cn)<=0 by A122,REAL_2:106;
        then A141: -(p`1/|.p.|-cn)>= -0 by REAL_1:50;
        A142: (|.p.|)^2 =(p`1)^2+(p`2)^2 by JGRAPH_3:10;
          0<=(p`2)^2 by SQUARE_1:72;
        then 0+(p`1)^2<=(p`1)^2+(p`2)^2 by REAL_1:55;
        then (p`1)^2/(|.p.|)^2 <= (|.p.|)^2/(|.p.|)^2 by A140,A142,REAL_1:73;
        then (p`1)^2/(|.p.|)^2 <= 1 by A140,XCMPLX_1:60;
        then ((p`1)/|.p.|)^2 <= 1 by SQUARE_1:69;
        then -1<=p`1/|.p.| by Th4;
        then A143: -1-cn<=p`1/|.p.|-cn by REAL_1:49;
          -(1+cn)<= -0 by A130,REAL_1:50;
        then (-(1+cn))/(1+cn)<=(-( p`1/|.p.|-cn))/(1+cn) by A130,A141,REAL_1:73
;
        then A144: -1<=(-( p`1/|.p.|-cn))/(1+cn) by A130,XCMPLX_1:198;
          -(1+cn)<=p`1/|.p.|-cn by A143,XCMPLX_1:161;
        then --(1+cn)>= -(p`1/|.p.|-cn) by REAL_1:50;
       then (-(p`1/|.p.|-cn))/(1+cn)<=1 by A130,REAL_2:117;
       then ((-(p`1/|.p.|-cn))/(1+cn))^2<=1^2 by A144,JGRAPH_2:7;
       then 1-((-(p`1/|.p.|-cn))/(1+cn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((p`1/|.p.|-cn))/(1+cn))^2>=0 by XCMPLX_1:188;
       then A145: 1-((p`1/|.p.|-cn)/(1+cn))^2>=0 by SQUARE_1:61;
        A146: (p4`2)^2= (|.p.|)^2*( sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2))^2
                         by A124,SQUARE_1:68
         .= (|.p.|)^2*(1-((p`1/|.p.|-cn)/(1+cn))^2)
                      by A145,SQUARE_1:def 4;
        A147: (p4`1)^2= (|.p.|)^2*((p`1/|.p.|-cn)/(1+cn))^2
                         by A124,SQUARE_1:68;
          (|.p4.|)^2=(p4`1)^2+(p4`2)^2 by JGRAPH_3:10
        .=(|.p.|)^2*(1-((p`1/|.p.|-cn)/(1+cn))^2
           +((p`1/|.p.|-cn)/(1+cn))^2) by A146,A147,XCMPLX_1:8
        .=(|.p.|)^2*1 by XCMPLX_1:27 .=(|.p.|)^2;
        then A148: sqrt((|.p4.|)^2)=|.p.| by A138,SQUARE_1:89;
        then A149: |.p4.|=|.p.| by A137,SQUARE_1:89;
          ((p`1/|.p.|-cn)/(1+cn))=|.q.|* ((q`1/|.q.|-cn)/(1+cn))/|.p.|
                         by A2,A97,A98,A123,A124,A139,XCMPLX_1:90;
        then (p`1/|.p.|-cn)/(1+cn)=(q`1/|.q.|-cn)/(1+cn)
                              by A2,A97,A123,A135,A139,A148,XCMPLX_1:90
;
        then (p`1/|.p.|-cn)/(1+cn)*(1+cn)=q`1/|.q.|-cn
                            by A130,XCMPLX_1:88;
        then p`1/|.p.|-cn=q`1/|.q.|-cn by A130,XCMPLX_1:88;
        then p`1/|.p.|-cn+cn=q`1/|.q.| by XCMPLX_1:27;
        then p`1/|.p.|=q`1/|.q.| by XCMPLX_1:27;
        then p`1/|.p.|*|.p.|=q`1 by A2,A97,A123,A136,A139,A149,XCMPLX_1:88;
        then A150: p`1=q`1 by A139,XCMPLX_1:88;
        A151: |.p.|^2=(p`1)^2+(p`2)^2 by JGRAPH_3:10;
          |.q.|^2=(q`1)^2+(q`2)^2 by JGRAPH_3:10;
        then (p`2)^2+(p`1)^2-(p`1)^2=(q`2)^2 by A2,A97,A123,A136,A149,A150,A151
,XCMPLX_1:26;
        then (p`2)^2=(q`2)^2 by XCMPLX_1:26;
        then p`2=sqrt((q`2)^2) by A122,SQUARE_1:89;
        then A152: p`2=q`2 by A96,SQUARE_1:89;
          p=|[p`1,p`2]| by EUCLID:57;
       hence x1=x2 by A150,A152,EUCLID:57;
      end;
     hence x1=x2;
    end;
   hence x1=x2;
  end;
 hence thesis by FUNCT_1:def 8;
end;

theorem Th79: for cn being Real
  st -1<cn & cn<1 holds
  (cn-FanMorphN) is map of TOP-REAL 2,TOP-REAL 2 &
 rng (cn-FanMorphN) = the carrier of TOP-REAL 2
proof let cn be Real;
  assume A1:  -1<cn & cn<1;
thus (cn-FanMorphN) is map of TOP-REAL 2,TOP-REAL 2 ;
  A2:for f being map of TOP-REAL 2,TOP-REAL 2 st f=(cn-FanMorphN) holds
    rng (cn-FanMorphN)=the carrier of TOP-REAL 2
  proof let f be map of TOP-REAL 2,TOP-REAL 2;
    assume A3:f=(cn-FanMorphN);
    A4:dom f=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
      the carrier of TOP-REAL 2 c= rng f
    proof let y be set;assume y in the carrier of TOP-REAL 2;
     then reconsider p2=y as Point of TOP-REAL 2;
     set q=p2;
       now per cases by JGRAPH_2:11;
     case q`2<=0;
       then q in dom (cn-FanMorphN) & y=(cn-FanMorphN).q by A3,A4,Th56;
      hence ex x being set st x in dom (cn-FanMorphN) & y=(cn-FanMorphN).x;
     case A5: q`1/|.q.|>=0 & q`2>=0 & q<>0.REAL 2;
          set px=|[ |.q.|*(q`1/|.q.|*(1-cn)+cn),
            (|.q.|)*sqrt(1-(q`1/|.q.|*(1-cn)+cn)^2)]|;
        A6: px`2 = (|.q.|)*sqrt(1-(q`1/|.q.|*(1-cn)+cn)^2) &
         px`1 = |.q.|*(q`1/|.q.|*(1-cn)+cn) by EUCLID:56;
        then A7: |.px.|^2=((|.q.|)*sqrt(1-(q`1/|.q.|*(1-cn)+cn)^2))^2
        +(|.q.|*(q`1/|.q.|*(1-cn)+cn))^2 by JGRAPH_3:10
        .=(|.q.|)^2*(sqrt(1-(q`1/|.q.|*(1-cn)+cn)^2))^2
        +(|.q.|*(q`1/|.q.|*(1-cn)+cn))^2 by SQUARE_1:68
        .=(|.q.|)^2*(sqrt(1-(q`1/|.q.|*(1-cn)+cn)^2))^2
        +(|.q.|)^2*((q`1/|.q.|*(1-cn)+cn))^2 by SQUARE_1:68;
       A8: |.q.|<>0 by A5,TOPRNS_1:25;
         1-cn>=0 by A1,Th3;
       then A9: q`1/|.q.|*(1-cn)>=0 by A5,REAL_2:121;
       then A10: (q`1/|.q.|*(1-cn)+cn)>=0+cn by REAL_1:55;
       A11: |.q.|>=0 by TOPRNS_1:26;
       A12: |.q.|^2>0 by A8,SQUARE_1:74;
         --(1+cn)>0 by A1,Th3;
       then -(1+cn)<0 by REAL_1:66;
       then -1-cn <0 by XCMPLX_1:161;
       then -1-cn<= q`1/|.q.|*(1-cn) by A9,AXIOMS:22;
       then -1-cn+cn<= q`1/|.q.|*(1-cn)+cn by REAL_1:55;
       then A13: -1<= (q`1/|.q.|*(1-cn)+cn) by XCMPLX_1:27;
       A14: 1-cn>0 by A1,Th3;
       A15: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
         0<=(q`2)^2 by SQUARE_1:72;
       then 0+(q`1)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
       then (q`1)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A12,A15,REAL_1:73;
       then (q`1)^2/(|.q.|)^2 <= 1 by A12,XCMPLX_1:60;
       then ((q`1)/|.q.|)^2 <= 1 by SQUARE_1:69;
       then q`1/|.q.|<=1 by Th4;
       then q`1/|.q.|*(1-cn) <=(1)*(1-cn) by A14,AXIOMS:25;
       then q`1/|.q.|*(1-cn)+cn-cn <=1-cn by XCMPLX_1:26;
       then (q`1/|.q.|*(1-cn)+cn) <=1 by REAL_1:54;
       then 1^2>=(q`1/|.q.|*(1-cn)+cn)^2 by A13,JGRAPH_2:7;
       then A16: 1-(q`1/|.q.|*(1-cn)+cn)^2>=0 by SQUARE_1:12,59;
       then A17: |.px.|^2=(|.q.|)^2*(1-(q`1/|.q.|*(1-cn)+cn)^2)
        +(|.q.|)^2*((q`1/|.q.|*(1-cn)+cn))^2 by A7,SQUARE_1:def 4
        .= (|.q.|)^2*((1-(q`1/|.q.|*(1-cn)+cn)^2)
        +((q`1/|.q.|*(1-cn)+cn))^2) by XCMPLX_1:8
        .= (|.q.|)^2*(1) by XCMPLX_1:27
        .= (|.q.|)^2;
          |.px.|>=0 by TOPRNS_1:26;
        then A18: |.px.|=sqrt(|.q.|^2) by A17,SQUARE_1:89
        .=|.q.| by A11,SQUARE_1:89;
        then A19: |.px.|<>0 by A5,TOPRNS_1:25;
         sqrt(1-(q`1/|.q.|*(1-cn)+cn)^2)>=0 by A16,SQUARE_1:def 4;
        then px`1/|.px.| >=cn & px`2>=0 & px<>0.REAL 2
                                by A6,A8,A10,A11,A18,REAL_2:121,TOPRNS_1:24,
XCMPLX_1:90;
    then A20: (cn-FanMorphN).px
       =|[ |.px.|* ((px`1/|.px.|-cn)/(1-cn)),
         |.px.|*( sqrt(1-((px`1/|.px.|-cn)/(1-cn))^2))]|
                            by A1,Th58;
        A21: |.px.|* ((px`1/|.px.|-cn)/(1-cn))
        =|.q.|* (( ((q`1/|.q.|*(1-cn)+cn))-cn)/(1-cn)) by A6,A8,A18,XCMPLX_1:90
        .=|.q.|* (( q`1/|.q.|*(1-cn) )/(1-cn)) by XCMPLX_1:26
        .=|.q.|* ( q`1/|.q.|) by A14,XCMPLX_1:90
        .= q`1 by A8,XCMPLX_1:88;
        then A22: |.px.|*( sqrt(1-((px`1/|.px.|-cn)/(1-cn))^2))
           = |.px.|*( sqrt(1-(q`1/|.px.|)^2)) by A19,XCMPLX_1:90
           .= |.px.|*( sqrt(1-(q`1)^2/(|.px.|)^2)) by SQUARE_1:69
           .= |.px.|*( sqrt( (|.px.|)^2/(|.px.|)^2-(q`1)^2/(|.px.|)^2))
                            by A12,A17,XCMPLX_1:60
           .= |.px.|*( sqrt( ((|.px.|)^2-(q`1)^2)/(|.px.|)^2))
                            by XCMPLX_1:121
           .= |.px.|*( sqrt( ((q`1)^2+(q`2)^2-(q`1)^2)/(|.px.|)^2))
                            by A17,JGRAPH_3:10
           .= |.px.|*( sqrt( ((q`2)^2)/(|.px.|)^2)) by XCMPLX_1:26
           .= |.px.|*( sqrt((q`2/|.q.|)^2)) by A18,SQUARE_1:69;
           q`2/|.q.|>=0 by A5,A11,REAL_2:125;
         then |.px.|*( sqrt((q`2/|.q.|)^2))=|.q.|*(q`2/|.q.|) by A18,SQUARE_1:
89
         .=q`2 by A8,XCMPLX_1:88;
        then A23:q=(cn-FanMorphN).px by A20,A21,A22,EUCLID:57;
          dom (cn-FanMorphN)=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
      hence ex x being set st x in dom (cn-FanMorphN) & y=(cn-FanMorphN).x
                               by A23;
     case A24: q`1/|.q.|<0 & q`2>=0 & q<>0.REAL 2;
          set px=|[ |.q.|*(q`1/|.q.|*(1+cn)+cn),
            (|.q.|)*sqrt(1-(q`1/|.q.|*(1+cn)+cn)^2)]|;
        A25: px`2 = (|.q.|)*sqrt(1-(q`1/|.q.|*(1+cn)+cn)^2) &
         px`1 = |.q.|*(q`1/|.q.|*(1+cn)+cn) by EUCLID:56;
        then A26: |.px.|^2=((|.q.|)*sqrt(1-(q`1/|.q.|*(1+cn)+cn)^2))^2
        +(|.q.|*(q`1/|.q.|*(1+cn)+cn))^2 by JGRAPH_3:10
        .=(|.q.|)^2*(sqrt(1-(q`1/|.q.|*(1+cn)+cn)^2))^2
        +(|.q.|*(q`1/|.q.|*(1+cn)+cn))^2 by SQUARE_1:68
        .=(|.q.|)^2*(sqrt(1-(q`1/|.q.|*(1+cn)+cn)^2))^2
        +(|.q.|)^2*((q`1/|.q.|*(1+cn)+cn))^2 by SQUARE_1:68;
       A27: |.q.|<>0 by A24,TOPRNS_1:25;
         1+cn>=0 by A1,Th3;
       then A28: q`1/|.q.|*(1+cn)<=0 by A24,REAL_2:123;
       then A29: (q`1/|.q.|*(1+cn)+cn)<=0+cn by REAL_1:55;
       A30: |.q.|>=0 by TOPRNS_1:26;
       A31: |.q.|^2>0 by A27,SQUARE_1:74;
         (1-cn)>0 by A1,Th3;
       then 1-cn>= q`1/|.q.|*(1+cn) by A28,AXIOMS:22;
       then 1-cn+cn>= q`1/|.q.|*(1+cn)+cn by REAL_1:55;
       then A32: 1>= (q`1/|.q.|*(1+cn)+cn) by XCMPLX_1:27;
       A33: 1+cn>0 by A1,Th3;
       A34: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
         0<=(q`2)^2 by SQUARE_1:72;
       then 0+(q`1)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
       then (q`1)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A31,A34,REAL_1:73;
       then (q`1)^2/(|.q.|)^2 <= 1 by A31,XCMPLX_1:60;
       then ((q`1)/|.q.|)^2 <= 1 by SQUARE_1:69;
       then q`1/|.q.|>=-1 by Th4;
       then q`1/|.q.|*(1+cn) >=(-1)*(1+cn) by A33,AXIOMS:25;
       then q`1/|.q.|*(1+cn) >=-(1+cn) by XCMPLX_1:180;
       then q`1/|.q.|*(1+cn) >=-1-cn by XCMPLX_1:161;
       then q`1/|.q.|*(1+cn)+cn-cn >=-1-cn by XCMPLX_1:26;
       then (q`1/|.q.|*(1+cn)+cn) >=-1 by REAL_1:54;
       then 1^2>=(q`1/|.q.|*(1+cn)+cn)^2 by A32,JGRAPH_2:7;
       then A35: 1-(q`1/|.q.|*(1+cn)+cn)^2>=0 by SQUARE_1:12,59;
       then A36: |.px.|^2=(|.q.|)^2*(1-(q`1/|.q.|*(1+cn)+cn)^2)
        +(|.q.|)^2*((q`1/|.q.|*(1+cn)+cn))^2 by A26,SQUARE_1:def 4
        .= (|.q.|)^2*((1-(q`1/|.q.|*(1+cn)+cn)^2)
        +((q`1/|.q.|*(1+cn)+cn))^2) by XCMPLX_1:8
        .= (|.q.|)^2*(1) by XCMPLX_1:27
        .= (|.q.|)^2;
          |.px.|>=0 by TOPRNS_1:26;
        then A37: |.px.|=sqrt(|.q.|^2) by A36,SQUARE_1:89
        .=|.q.| by A30,SQUARE_1:89;
        then A38: |.px.|<>0 by A24,TOPRNS_1:25;
         sqrt(1-(q`1/|.q.|*(1+cn)+cn)^2)>=0 by A35,SQUARE_1:def 4;
        then px`1/|.px.| <=cn & px`2>=0 & px<>0.REAL 2
                           by A25,A27,A29,A30,A37,REAL_2:121,TOPRNS_1:24,
XCMPLX_1:90;
    then A39: (cn-FanMorphN).px
       =|[ |.px.|* ((px`1/|.px.|-cn)/(1+cn)),
         |.px.|*( sqrt(1-((px`1/|.px.|-cn)/(1+cn))^2))]|
                            by A1,Th58;
        A40: |.px.|* ((px`1/|.px.|-cn)/(1+cn))
        =|.q.|* (( ((q`1/|.q.|*(1+cn)+cn))-cn)/(1+cn)) by A25,A27,A37,XCMPLX_1:
90
        .=|.q.|* (( q`1/|.q.|*(1+cn) )/(1+cn)) by XCMPLX_1:26
        .=|.q.|* ( q`1/|.q.|) by A33,XCMPLX_1:90
        .= q`1 by A27,XCMPLX_1:88;
        then A41: |.px.|*( sqrt(1-((px`1/|.px.|-cn)/(1+cn))^2))
           = |.px.|*( sqrt(1-(q`1/|.px.|)^2)) by A38,XCMPLX_1:90
           .= |.px.|*( sqrt(1-(q`1)^2/(|.px.|)^2)) by SQUARE_1:69
           .= |.px.|*( sqrt( (|.px.|)^2/(|.px.|)^2-(q`1)^2/(|.px.|)^2))
                            by A31,A36,XCMPLX_1:60
           .= |.px.|*( sqrt( ((|.px.|)^2-(q`1)^2)/(|.px.|)^2))
                            by XCMPLX_1:121
           .= |.px.|*( sqrt( ((q`1)^2+(q`2)^2-(q`1)^2)/(|.px.|)^2))
                            by A36,JGRAPH_3:10
           .= |.px.|*( sqrt( ((q`2)^2)/(|.px.|)^2))
                            by XCMPLX_1:26
           .= |.px.|*( sqrt((q`2/|.q.|)^2)) by A37,SQUARE_1:69;
           q`2/|.q.|>=0 by A24,A30,REAL_2:125;
         then |.px.|*( sqrt((q`2/|.q.|)^2))=|.q.|*(q`2/|.q.|) by A37,SQUARE_1:
89
         .=q`2 by A27,XCMPLX_1:88;
        then A42:q=(cn-FanMorphN).px by A39,A40,A41,EUCLID:57;
          dom (cn-FanMorphN)=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
      hence ex x being set st x in dom (cn-FanMorphN) & y=(cn-FanMorphN).x
                               by A42;
     end;
    hence y in rng f by A3,FUNCT_1:def 5;
   end;
   hence rng (cn-FanMorphN)=the carrier of TOP-REAL 2 by A3,XBOOLE_0:def 10;
  end;
 thus thesis by A2;
end;

theorem Th80: for cn being Real,p2 being Point of TOP-REAL 2
  st -1<cn & cn<1
  ex K being non empty compact Subset of TOP-REAL 2 st
  K = (cn-FanMorphN).:K &
 (ex V2 being Subset of TOP-REAL 2 st
  p2 in V2 & V2 is open & V2 c= K & (cn-FanMorphN).p2 in V2)
 proof let cn be Real,p2 be Point of TOP-REAL 2;
  assume A1: -1<cn & cn<1;
    A2: the carrier of TOP-REAL 2=REAL 2 by EUCLID:25;
    A3: Euclid 2=MetrStruct(#REAL 2,Pitag_dist 2#) by EUCLID:def 7;
    then reconsider O=0.REAL 2 as Point of Euclid 2 by EUCLID:25;
    A4: TOP-REAL 2=TopSpaceMetr(Euclid 2) by EUCLID:def 8;
    reconsider V0=Ball(O,|.p2.|+1) as Subset of TOP-REAL 2
                              by A2,A3;
    A5: |. p2.|<|.p2.|+1 by REAL_1:69;
      0<=|.p2.| by TOPRNS_1:26;
    then A6: O in V0 by A5,GOBOARD6:4;
    A7: V0 c= cl_Ball(O,|.p2.|+1) by METRIC_1:15;
    reconsider K01=cl_Ball(O,|.p2.|+1) as Subset of TOP-REAL 2
                            by A2,A3;
      K01 is compact by Th22;
    then reconsider K0=cl_Ball(O,|.p2.|+1)
          as non empty compact Subset of TOP-REAL 2 by A6,A7;
    reconsider u2=p2 as Point of Euclid 2 by A3,EUCLID:25;
    A8: (cn-FanMorphN).:K0 c= K0
     proof let y be set;assume y in (cn-FanMorphN).:K0;
       then consider x being set such that
       A9: x in dom (cn-FanMorphN) & x in K0 & y=(cn-FanMorphN).x
                     by FUNCT_1:def 12;
       reconsider q=x as Point of TOP-REAL 2 by A9;
       reconsider uq=q as Point of Euclid 2 by A3,EUCLID:25;
       A10: y in rng (cn-FanMorphN) by A9,FUNCT_1:def 5;
       then reconsider u=y as Point of Euclid 2 by A3,EUCLID:25;
       reconsider q4=y as Point of TOP-REAL 2 by A10;
       A11: |.q4.|=|.q.| by A9,Th73;
         dist(O,uq)<= |.p2.|+1 by A9,METRIC_1:13;
       then |.0.REAL 2 - q.|<= |.p2.|+1 by JGRAPH_1:45;
       then |.0.REAL 2 +- q.|<= |.p2.|+1 by EUCLID:45;
       then |. -q.|<= |.p2.|+1 by EUCLID:31;
       then |.q.|<= |.p2.|+1 by TOPRNS_1:27;
       then |. -q4.|<= |.p2.|+1 by A11,TOPRNS_1:27;
       then |.0.REAL 2 +- q4.|<= |.p2.|+1 by EUCLID:31;
       then |.0.REAL 2 - q4.|<= |.p2.|+1 by EUCLID:45;
       then dist(O,u)<= |.p2.|+1 by JGRAPH_1:45;
      hence y in K0 by METRIC_1:13;
     end;
      K0 c= (cn-FanMorphN).:K0
     proof let y be set;assume A12: y in K0;
       then reconsider u=y as Point of Euclid 2;
       reconsider q4=y as Point of TOP-REAL 2 by A12;
         the carrier of TOP-REAL 2 c= rng (cn-FanMorphN) by A1,Th79;
       then q4 in rng (cn-FanMorphN) by TARSKI:def 3;
       then consider x being set such that
       A13: x in dom (cn-FanMorphN) & y=(cn-FanMorphN).x by FUNCT_1:def 5;
       reconsider q=x as Point of TOP-REAL 2 by A13,FUNCT_2:def 1;
       reconsider uq=q as Point of Euclid 2 by A3,EUCLID:25;
       A14: |.q4.|=|.q.| by A13,Th73;
        q = uq & q4 = u & O=0.REAL 2;
      then q in K0 by A12,A14,Lm11;
      hence y in (cn-FanMorphN).:K0 by A13,FUNCT_1:def 12;
     end;
    then A15: K0= (cn-FanMorphN).:K0 by A8,XBOOLE_0:def 10;
      |. -p2.|<|.p2.|+1 by A5,TOPRNS_1:27;
    then |.0.REAL 2 +- p2.|<|.p2.|+1 by EUCLID:31;
    then |.0.REAL 2 - p2.|<|.p2.|+1 by EUCLID:45;
    then dist(O,u2)<|.p2.|+1 by JGRAPH_1:45;
    then A16: p2 in V0 by METRIC_1:12;
    A17: V0 is open by A4,TOPMETR:21;
     set q3= (cn-FanMorphN).p2;
     reconsider u3=q3 as Point of Euclid 2 by A3,EUCLID:25;
       |.q3.|=|.p2.| by Th73;
    then |. -q3.|<|.p2.|+1 by A5,TOPRNS_1:27;
    then |.0.REAL 2 +- q3.|<|.p2.|+1 by EUCLID:31;
    then |.0.REAL 2 - q3.|<|.p2.|+1 by EUCLID:45;
     then dist(O,u3)< |.p2.|+1 by JGRAPH_1:45;
    then (cn-FanMorphN).p2 in V0 by METRIC_1:12;
   hence thesis by A7,A15,A16,A17;
 end;

theorem for cn being Real st -1<cn & cn<1
 ex f being map of TOP-REAL 2,TOP-REAL 2 st f=(cn-FanMorphN)
   & f is_homeomorphism
proof let cn be Real;assume A1: -1<cn & cn<1;
  then A2: (cn-FanMorphN) is map of TOP-REAL 2,TOP-REAL 2 &
  rng (cn-FanMorphN) = the carrier of TOP-REAL 2 by Th79;
  reconsider f=(cn-FanMorphN) as map of TOP-REAL 2,TOP-REAL 2;
  consider h being map of (TOP-REAL 2),(TOP-REAL 2)
  such that A3: h=(cn-FanMorphN) & h is continuous by A1,Th77;
  A4: f is one-to-one by A1,Th78;
  A5: rng f=[#](TOP-REAL 2) by A2,PRE_TOPC:12;
   for p2 being Point of TOP-REAL 2
  ex K being non empty compact Subset of TOP-REAL 2 st K = f.:K &
 (ex V2 being Subset of TOP-REAL 2 st
  p2 in V2 & V2 is open & V2 c= K & f.p2 in V2) by A1,Th80;
  then f is_homeomorphism by A3,A4,A5,Th8;
 hence thesis;
end;

theorem Th82:
 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-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;
  A2: |.q.|>=0 by TOPRNS_1:26;
  let p be Point of TOP-REAL 2;
    assume p=(cn-FanMorphN).q;
    then A3: p=|[ |.q.|* ((q`1/|.q.|-cn)/(1-cn)),
         |.q.|*( sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))]| by A1,Th56;
      set qz=p;
      A4: |.q.|<>0 by A1,JGRAPH_2:11,TOPRNS_1:25;
      then A5: (|.q.|)^2>0 by SQUARE_1:74;
     A6: 1-cn>0 by A1,Th3;
        A7: qz`2= |.q.|*( sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))&
        qz`1= |.q.|* ((q`1/|.q.|-cn)/(1-cn)) by A3,EUCLID:56;
        A8: (q`1/|.q.|-cn)>= 0 by A1,SQUARE_1:12;
        A9: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
          0<(q`2)^2 by A1,SQUARE_1:74;
        then 0+(q`1)^2<(q`1)^2+(q`2)^2 by REAL_1:67;
        then (q`1)^2/(|.q.|)^2 < (|.q.|)^2/(|.q.|)^2 by A5,A9,REAL_1:73;
        then (q`1)^2/(|.q.|)^2 < 1 by A5,XCMPLX_1:60;
        then ((q`1)/|.q.|)^2 < 1 by SQUARE_1:69;
        then 1>q`1/|.q.| by Th5;
        then A10: 1-cn>q`1/|.q.|-cn by REAL_1:54;
        A11: -(1-cn)< -0 by A6,REAL_1:50;
           -(1-cn)< -( q`1/|.q.|-cn) by A10,REAL_1:50;
         then (-(1-cn))/(1-cn)<(-( q`1/|.q.|-cn))/(1-cn)
                        by A6,REAL_1:73;
        then A12: -1<(-( q`1/|.q.|-cn))/(1-cn) by A6,XCMPLX_1:198;
         --(1-cn)> -(q`1/|.q.|-cn) by A8,A11,REAL_1:50;
       then (-(q`1/|.q.|-cn))/(1-cn)<1 by A6,REAL_2:118;
       then ((-(q`1/|.q.|-cn))/(1-cn))^2<1^2 by A12,JGRAPH_2:8;
       then 1-((-(q`1/|.q.|-cn))/(1-cn))^2>0 by SQUARE_1:11,59;
       then sqrt(1-((-(q`1/|.q.|-cn))/(1-cn))^2)>0 by SQUARE_1:93;
       then sqrt(1-(-(q`1/|.q.|-cn))^2/(1-cn)^2)> 0 by SQUARE_1:69;
       then sqrt(1-(q`1/|.q.|-cn)^2/(1-cn)^2)> 0 by SQUARE_1:61;
       then A13: sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2)> 0 by SQUARE_1:69;
         ((q`1/|.q.|-cn)/(1-cn))>=0 by A6,A8,SQUARE_1:27;
   hence p`2>0 & p`1>=0 by A2,A4,A7,A13,SQUARE_1:19,21;
 end;

theorem Th83:
 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-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;
  A2: |.q.|>=0 by TOPRNS_1:26;
  let p be Point of TOP-REAL 2;
    assume p=(cn-FanMorphN).q;
    then A3: p=|[ |.q.|* ((q`1/|.q.|-cn)/(1+cn)),
      |.q.|*( sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))]| by A1,Th57;
      set qz=p;
          A4: |.q.|<>0 by A1,JGRAPH_2:11,TOPRNS_1:25;
      then A5: (|.q.|)^2>0 by SQUARE_1:74;
      A6: (q`1/|.q.|-cn)< 0 by A1,REAL_2:105;
     A7: 1+cn>0 by A1,Th3;
        A8: qz`2= |.q.|*( sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))&
        qz`1= |.q.|* ((q`1/|.q.|-cn)/(1+cn)) by A3,EUCLID:56;
        A9: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
          0<(q`2)^2 by A1,SQUARE_1:74;
        then 0+(q`1)^2<(q`1)^2+(q`2)^2 by REAL_1:67;
        then (q`1)^2/(|.q.|)^2 < (|.q.|)^2/(|.q.|)^2 by A5,A9,REAL_1:73;
        then (q`1)^2/(|.q.|)^2 < 1 by A5,XCMPLX_1:60;
        then ((q`1)/|.q.|)^2 < 1 by SQUARE_1:69;
        then -1<q`1/|.q.| by Th5;
        then -1-cn<q`1/|.q.|-cn by REAL_1:54;
        then -(1+cn)<q`1/|.q.|-cn by XCMPLX_1:161;
        then A10: --(1+cn)> -(q`1/|.q.|-cn) by REAL_1:50;
        A11: -( q`1/|.q.|-cn)>0 by A6,REAL_1:66;
          -(1+cn)< -0 by A7,REAL_1:50;
         then -(1+cn)< -( q`1/|.q.|-cn) by A11,AXIOMS:22;
         then (-(1+cn))/(1+cn)<(-( q`1/|.q.|-cn))/(1+cn) by A7,REAL_1:73;
        then A12: -1<(-( q`1/|.q.|-cn))/(1+cn) by A7,XCMPLX_1:198;
         (-(q`1/|.q.|-cn))/(1+cn)<1 by A7,A10,REAL_2:118;
       then ((-(q`1/|.q.|-cn))/(1+cn))^2<1^2 by A12,JGRAPH_2:8;
       then 1-((-(q`1/|.q.|-cn))/(1+cn))^2>0 by SQUARE_1:11,59;
       then sqrt(1-((-(q`1/|.q.|-cn))/(1+cn))^2)>0 by SQUARE_1:93;
       then sqrt(1-(-(q`1/|.q.|-cn))^2/(1+cn)^2)> 0 by SQUARE_1:69;
       then sqrt(1-(q`1/|.q.|-cn)^2/(1+cn)^2)> 0 by SQUARE_1:61;
       then A13: sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2)> 0 by SQUARE_1:69;
         ((q`1/|.q.|-cn)/(1+cn))<0 by A6,A7,REAL_2:128;
   hence p`2>0 & p`1<0 by A2,A4,A8,A13,SQUARE_1:21,24;
 end;

theorem Th84:
 for cn being Real,q1,q2 being Point of TOP-REAL 2 st -1<cn & cn<1 & q1`2>0
 & q1`1/|.q1.|>=cn & q2`2>0 & q2`1/|.q2.|>=cn & 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
   & q1`1/|.q1.|>=cn & q2`2>0 & q2`1/|.q2.|>=cn & q1`1/|.q1.|<q2`1/|.q2.|;
  let p1,p2 be Point of TOP-REAL 2;
   assume A2: p1=(cn-FanMorphN).q1 & p2=(cn-FanMorphN).q2;
    then p1=|[ |.q1.|* ((q1`1/|.q1.|-cn)/(1-cn)),
       |.q1.|*( sqrt(1-((q1`1/|.q1.|-cn)/(1-cn))^2))]| by A1,Th56;
    then A3: p1`2= |.q1.|*( sqrt(1-((q1`1/|.q1.|-cn)/(1-cn))^2))&
        p1`1= |.q1.|* ((q1`1/|.q1.|-cn)/(1-cn)) by EUCLID:56;
      p2=|[ |.q2.|* ((q2`1/|.q2.|-cn)/(1-cn)),
      |.q2.|*( sqrt(1-((q2`1/|.q2.|-cn)/(1-cn))^2))]| by A1,A2,Th56;
    then A4: p2`2= |.q2.|*( sqrt(1-((q2`1/|.q2.|-cn)/(1-cn))^2))&
        p2`1= |.q2.|* ((q2`1/|.q2.|-cn)/(1-cn))
                     by EUCLID:56;
    A5: |.q1.|>0 & |.q2.|>0 by A1,Lm1,JGRAPH_2:11;
    A6: |.p1.|=|.q1.| & |.p2.|=|.q2.| by A2,Th73;
    then A7: p1`1/|.p1.|= ((q1`1/|.q1.|-cn)/(1-cn)) by A3,A5,XCMPLX_1:90
;
    A8: p2`1/|.p2.|= ((q2`1/|.q2.|-cn)/(1-cn)) by A4,A5,A6,XCMPLX_1:90
;
    A9: q1`1/|.q1.|-cn< q2`1/|.q2.|-cn by A1,REAL_1:54;
      1-cn>0 by A1,Th3;
   hence thesis by A7,A8,A9,REAL_1:73;
 end;

theorem Th85:
 for cn being Real,q1,q2 being Point of TOP-REAL 2 st -1<cn & cn<1 & q1`2>0
 & q1`1/|.q1.|<cn & q2`2>0 & q2`1/|.q2.|<cn & 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
   & q1`1/|.q1.|<cn & q2`2>0 & q2`1/|.q2.|<cn & q1`1/|.q1.|<q2`1/|.q2.|;
  let p1,p2 be Point of TOP-REAL 2;
   assume A2: p1=(cn-FanMorphN).q1 & p2=(cn-FanMorphN).q2;
    then p1=|[ |.q1.|* ((q1`1/|.q1.|-cn)/(1+cn)),
      |.q1.|*( sqrt(1-((q1`1/|.q1.|-cn)/(1+cn))^2))]| by A1,Th57;
    then A3: p1`2= |.q1.|*( sqrt(1-((q1`1/|.q1.|-cn)/(1+cn))^2))&
        p1`1= |.q1.|* ((q1`1/|.q1.|-cn)/(1+cn))
                     by EUCLID:56;
      p2=|[ |.q2.|* ((q2`1/|.q2.|-cn)/(1+cn)),
      |.q2.|*( sqrt(1-((q2`1/|.q2.|-cn)/(1+cn))^2))]| by A1,A2,Th57;
    then A4: p2`2= |.q2.|*( sqrt(1-((q2`1/|.q2.|-cn)/(1+cn))^2))&
        p2`1= |.q2.|* ((q2`1/|.q2.|-cn)/(1+cn))
                     by EUCLID:56;
    A5: |.q1.|>0 & |.q2.|>0 by A1,Lm1,JGRAPH_2:11;
    A6: |.p1.|=|.q1.| & |.p2.|=|.q2.| by A2,Th73;
    then A7: p1`1/|.p1.|= ((q1`1/|.q1.|-cn)/(1+cn)) by A3,A5,XCMPLX_1:90
;
    A8: p2`1/|.p2.|= ((q2`1/|.q2.|-cn)/(1+cn)) by A4,A5,A6,XCMPLX_1:90
;
    A9: q1`1/|.q1.|-cn< q2`1/|.q2.|-cn by A1,REAL_1:54;
      1+cn>0 by A1,Th3;
   hence p1`1/|.p1.|<p2`1/|.p2.| by A7,A8,A9,REAL_1:73;
 end;

theorem
   for cn being Real,q1,q2 being Point of TOP-REAL 2 st -1<cn & cn<1 & q1`2>0
 & q2`2>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`1/|.q1.|<q2`1/|.q2.|;
  let p1,p2 be Point of TOP-REAL 2;
   assume A2: p1=(cn-FanMorphN).q1 & p2=(cn-FanMorphN).q2;
      now per cases;
    case q1`1/|.q1.|>=cn & q2`1/|.q2.|>=cn;
     hence p1`1/|.p1.|<p2`1/|.p2.| by A1,A2,Th84;
    case q1`1/|.q1.|>=cn & q2`1/|.q2.|<cn;
     hence p1`1/|.p1.|<p2`1/|.p2.| by A1,AXIOMS:22;
    case A3: q1`1/|.q1.|<cn & q2`1/|.q2.|>=cn;
       then A4: p1`2>0 & p1`1<0 by A1,A2,Th83;
       then A5: |.p1.|>0 by Lm1,JGRAPH_2:11;
       A6: p2`2>0 & p2`1>=0 by A1,A2,A3,Th82;
       then |.p2.|>0 by Lm1,JGRAPH_2:11;
       then p2`1/|.p2.|>=0 by A6,REAL_2:125;
     hence p1`1/|.p1.|<p2`1/|.p2.| by A4,A5,REAL_2:128;
    case q1`1/|.q1.|<cn & q2`1/|.q2.|<cn;
     hence p1`1/|.p1.|<p2`1/|.p2.| by A1,A2,Th85;
    end;
   hence p1`1/|.p1.|<p2`1/|.p2.|;
 end;

theorem
   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-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;
  A2: |.q.|>=0 by TOPRNS_1:26;
  let p be Point of TOP-REAL 2;
   assume p=(cn-FanMorphN).q;
    then A3: p=|[ |.q.|* ((q`1/|.q.|-cn)/(1-cn)),
      |.q.|*( sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))]| by A1,Th56;
     A4: |.q.|<>0 by A1,JGRAPH_2:11,TOPRNS_1:25;
     A5: (q`1/|.q.|-cn)=0 by A1,XCMPLX_1:14;
     A6: p`2= |.q.|*( sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))&
        p`1= |.q.|* ((q`1/|.q.|-cn)/(1-cn)) by A3,EUCLID:56;
         ((-(q`1/|.q.|-cn))/(1-cn))^2<1^2 by A5,JGRAPH_2:8,XCMPLX_1:198;
       then 1-((-(q`1/|.q.|-cn))/(1-cn))^2>0 by SQUARE_1:11,59;
       then sqrt(1-((-(q`1/|.q.|-cn))/(1-cn))^2)>0 by SQUARE_1:93;
       then sqrt(1-(-(q`1/|.q.|-cn))^2/(1-cn)^2)> 0 by SQUARE_1:69;
       then sqrt(1-(q`1/|.q.|-cn)^2/(1-cn)^2)> 0 by SQUARE_1:61;
       then sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2)> 0 by SQUARE_1:69;
   hence thesis by A2,A4,A5,A6,SQUARE_1:21;
 end;

theorem
   for cn being real number holds 0.REAL 2=(cn-FanMorphN).(0.REAL 2)
   by Th56,JGRAPH_2:11;

begin :: Fan Morphism for East

definition let s be real number, q be Point of TOP-REAL 2;
  func FanE(s,q) -> Point of TOP-REAL 2 equals
  :Def6: |.q.|*|[sqrt(1-((q`2/|.q.|-s)/(1-s))^2),
          (q`2/|.q.|-s)/(1-s)]| if q`2/|.q.|>=s & q`1>0,
       |.q.|*|[sqrt(1-((q`2/|.q.|-s)/(1+s))^2),
          (q`2/|.q.|-s)/(1+s)]| if q`2/|.q.|<s & q`1>0
   otherwise q;
 correctness;
end;

definition let s be real number;
 func s-FanMorphE -> Function of the carrier of TOP-REAL 2,
 the carrier of TOP-REAL 2 means
  :Def7: for q being Point of TOP-REAL 2 holds
       it.q=FanE(s,q);
 existence
  proof
    deffunc F(Point of TOP-REAL 2)=FanE(s,$1);
    thus ex IT being Function of the carrier of TOP-REAL 2,
 the carrier of TOP-REAL 2 st
  for q being Point of TOP-REAL 2 holds
       IT.q=F(q)from LambdaD;
  end;
 uniqueness
  proof
    deffunc F(Point of TOP-REAL 2)=FanE(s,$1);
    thus for a,b being Function of the carrier of TOP-REAL 2,
 the carrier of TOP-REAL 2 st
  (for q being Point of TOP-REAL 2 holds a.q=F(q)) &
  (for q being Point of TOP-REAL 2 holds b.q=F(q)) holds a=b
    from FuncDefUniq;
  end;
end;

theorem Th89: for sn being real number holds
  (q`2/|.q.|>=sn & q`1>0 implies
    sn-FanMorphE.q= |[ |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1-sn))]|)&
  (q`1<=0 implies sn-FanMorphE.q=q)
proof let sn be real number;
  A1: sn-FanMorphE.q=FanE(sn,q) by Def7;
   hereby assume q`2/|.q.|>=sn & q`1>0;
    then FanE(sn,q)= |.q.|*|[sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2),
          (q`2/|.q.|-sn)/(1-sn)]| by Def6
          .= |[ |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1-sn))]| by EUCLID:62;
     hence sn-FanMorphE.q= |[ |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1-sn))]| by Def7;
    end;
   assume q`1<=0; hence thesis by A1,Def6;
end;

theorem Th90: for sn being Real holds
  (q`2/|.q.|<=sn & q`1>0 implies
    sn-FanMorphE.q= |[ |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)),
           |.q.|*((q`2/|.q.|-sn)/(1+sn))]|)
proof let sn be Real;
   assume A1: q`2/|.q.|<=sn & q`1>0;
     now per cases by A1,REAL_1:def 5;
   case q`2/|.q.|<sn;
    then FanE(sn,q)= |.q.|*|[sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2),
          (q`2/|.q.|-sn)/(1+sn)]| by A1,Def6
          .= |[ |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1+sn))]| by EUCLID:62;
    hence sn-FanMorphE.q= |[ |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)),
           |.q.|*((q`2/|.q.|-sn)/(1+sn))]| by Def7;
   case A2: q`2/|.q.|=sn;
    then A3: q`2/|.q.|-sn=0 by XCMPLX_1:14;
    then (q`2/|.q.|-sn)/(1-sn)=0;
    hence sn-FanMorphE.q= |[ |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)),
           |.q.|*((q`2/|.q.|-sn)/(1+sn))]| by A1,A2,A3,Th89;
   end;
   hence thesis;
end;

theorem Th91: for sn being Real st -1<sn & sn<1 holds
  (q`2/|.q.|>=sn & q`1>=0 & q<>0.REAL 2 implies
    sn-FanMorphE.q= |[ |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1-sn))]|) &
  (q`2/|.q.|<=sn & q`1>=0 & q<>0.REAL 2 implies
    sn-FanMorphE.q= |[ |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)),
           |.q.|*((q`2/|.q.|-sn)/(1+sn))]|)
proof let sn be Real;
 assume A1: -1<sn & sn<1;
     now per cases;
   case A2: q`2/|.q.|>=sn & q`1>=0 & q<>0.REAL 2;
      now per cases;
    case A3: q`1>0;
      then FanE(sn,q)= |.q.|*|[sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2),
          (q`2/|.q.|-sn)/(1-sn)]| by A2,Def6
          .= |[ |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1-sn))]| by EUCLID:62;
     hence thesis by A3,Def7,Th90;
    case A4: q`1<=0;
     then A5: sn-FanMorphE.q=q by Th89;
     A6: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
     A7: |.q.|<>0 by A2,TOPRNS_1:25;
     then A8: |.q.|^2>0 by SQUARE_1:74;
     A9: q`1=0 by A2,A4;
     A10: |.q.|>=0 by TOPRNS_1:26;
     then A11: sqrt((|.q.|)^2)=|.q.| by SQUARE_1:89;
     A12: 1-sn>0 by A1,Th3;
       (q`2)^2/|.q.|^2=1^2 by A6,A8,A9,SQUARE_1:59,60,XCMPLX_1:60;
      then ((q`2)/|.q.|)^2=1^2 by SQUARE_1:69;
      then A13: sqrt(((q`2)/|.q.|)^2)=1 by SQUARE_1:89;
       now assume q`2<0;
       then q`2/|.q.|<0 by A7,A10,REAL_2:128;
       then -((q`2)/|.q.|)=1 by A13,SQUARE_1:90;
      hence contradiction by A1,A2;
     end;
     then A14: |.q.|=q`2 by A6,A9,A11,SQUARE_1:60,89;
     then 1=q`2/|.q.| by A7,XCMPLX_1:60;
     then (q`2/|.q.|-sn)/(1-sn)=1 by A12,XCMPLX_1:60;
     hence thesis by A1,A5,A7,A9,A14,EUCLID:57,SQUARE_1:59,82,XCMPLX_1:60;
    end;
    hence thesis;
   case A15: q`2/|.q.|<=sn & q`1>=0 & q<>0.REAL 2;
      now per cases;
    case q`1>0;
     hence thesis by Th89,Th90;
    case A16: q`1<=0;
     A17: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
     A18: |.q.|<>0 by A15,TOPRNS_1:25;
     A19: q`1=0 by A15,A16;
      A20: |.q.|>=0 by TOPRNS_1:26;
     A21: 1+sn>0 by A1,Th3;
       1-sn>0 by A1,Th3;
     then 1-sn+sn>0+sn by REAL_1:53;
     then 1>sn by XCMPLX_1:27;
     then 1>q`2/|.q.| by A15,AXIOMS:22;
     then 1 *(|.q.|)>q`2/|.q.|*(|.q.|) by A18,A20,REAL_1:70;
     then A22: (|.q.|)>q`2 by A18,XCMPLX_1:88;
     then A23: |.q.|=-q`2 by A17,A19,JGRAPH_3:1,SQUARE_1:60;
     A24: q`2= -(|.q.|) by A17,A19,A22,JGRAPH_3:1,SQUARE_1:60;
     then -1=q`2/|.q.| by A18,XCMPLX_1:198;
     then (q`2/|.q.|-sn)/(1+sn)
     =(-(1+sn))/(1+sn) by XCMPLX_1:161 .=-1 by A21,XCMPLX_1:198;
     then |[ |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1+sn))]|
      =|[ |.q.|*(sqrt(1-1)),|.q.|*(-1)]| by SQUARE_1:59,61
     .=|[ |.q.|*(0),-(|.q.|)]| by SQUARE_1:82,XCMPLX_1:180
     .=q by A19,A23,EUCLID:57;
     hence thesis by A1,A16,A18,A24,Th89,XCMPLX_1:198;
    end;
     hence thesis;
   case q`1<0 or q=0.REAL 2;
    hence thesis;
   end;
  hence thesis;
end;

theorem Th92:
 for sn being Real,K1 being non empty Subset of TOP-REAL 2,
f being map of (TOP-REAL 2)|K1,R^1 st
  -1<sn & sn<1 &
(for p being Point of (TOP-REAL 2) st p in the carrier of (TOP-REAL 2)|K1
  holds f.p=|.p.|* ((p`2/|.p.|-sn)/(1-sn)))
  & (for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds q`1>=0 & q<>0.REAL 2) holds
f is continuous
proof
 let sn be Real,K1 be non empty Subset of TOP-REAL 2,
f be map of (TOP-REAL 2)|K1,R^1;
 assume A1: -1<sn & sn<1 &
(for p being Point of (TOP-REAL 2) st p in the carrier of (TOP-REAL 2)|K1
  holds f.p=|.p.|* ((p`2/|.p.|-sn)/(1-sn)))
  & (for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds q`1>=0 & q<>0.REAL 2);
  set a=sn, b=(1-sn);
  A2: b>0 by A1,Th3;
  reconsider g2=proj2|K1 as continuous map of (TOP-REAL 2)|K1,R^1 by Lm3;
  reconsider g1=(2 NormF)|K1 as continuous map of (TOP-REAL 2)|K1,R^1
                          by Lm5;
    for q being Point of TOP-REAL 2 st q in
    the carrier of (TOP-REAL 2)|K1 holds q<>0.REAL 2 by A1;
  then for q being Point of (TOP-REAL 2)|K1 holds g1.q<>0 by Lm6;
  then consider g3 being map of (TOP-REAL 2)|K1,R^1 such that
  A3: (for q being Point of
  (TOP-REAL 2)|K1,r1,r2 being Real st g2.q=r1 & g1.q=r2
  holds g3.q=r2*((r1/r2-a)/b)) & g3 is continuous by A2,Th10;
  A4:dom g3=the carrier of (TOP-REAL 2)|K1 by FUNCT_2:def 1;
  then A5:dom f=dom g3 by FUNCT_2:def 1;
    for x being set st x in dom f holds f.x=g3.x
   proof let x be set;assume
     A6: x in dom f;
     then x in K1 by A4,A5,JORDAN1:1;
     then reconsider r=x as Point of (TOP-REAL 2);
     reconsider s=x as Point of (TOP-REAL 2)|K1 by A6,FUNCT_2:def 1;
     A7:f.r=(|.r.|)* ((r`2/|.r.|-sn)/(1-sn)) by A1,A4,A5,A6;
     A8:g2.s=proj2.s by Lm3;
     A9:g1.s=(2 NormF).s by Lm5;
     A10:proj2.r=r`2 by PSCOMP_1:def 29;
       (2 NormF).r=|.r.| by Def1;
    hence f.x=g3.x by A3,A7,A8,A9,A10;
   end;
hence f is continuous by A3,A5,FUNCT_1:9;
end;

theorem Th93:
 for sn being Real,K1 being non empty Subset of TOP-REAL 2,
f being map of (TOP-REAL 2)|K1,R^1 st
  -1<sn & sn<1 &
(for p being Point of (TOP-REAL 2) st p in the carrier of (TOP-REAL 2)|K1
  holds f.p=|.p.|* ((p`2/|.p.|-sn)/(1+sn)))
  & (for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds q`1>=0 & q<>0.REAL 2) holds
f is continuous
proof
 let sn be Real,K1 be non empty Subset of TOP-REAL 2,
f be map of (TOP-REAL 2)|K1,R^1;
 assume A1: -1<sn & sn<1 &
(for p being Point of (TOP-REAL 2) st p in the carrier of (TOP-REAL 2)|K1
  holds f.p=|.p.|* ((p`2/|.p.|-sn)/(1+sn)))
  & (for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds q`1>=0 & q<>0.REAL 2);
  set a=sn, b=(1+sn);
  A2: 1+sn>0 by A1,Th3;
  reconsider g2=proj2|K1 as continuous map of (TOP-REAL 2)|K1,R^1 by Lm3;
  reconsider g1=(2 NormF)|K1 as continuous map of (TOP-REAL 2)|K1,R^1
                          by Lm5;
    for q being Point of TOP-REAL 2 st q in
    the carrier of (TOP-REAL 2)|K1 holds q<>0.REAL 2 by A1;
  then for q being Point of (TOP-REAL 2)|K1 holds g1.q<>0 by Lm6;
  then consider g3 being map of (TOP-REAL 2)|K1,R^1 such that
  A3: (for q being Point of
  (TOP-REAL 2)|K1,r1,r2 being Real st g2.q=r1 & g1.q=r2
  holds g3.q=r2*((r1/r2-a)/b)) & g3 is continuous by A2,Th10;
  A4:dom g3=the carrier of (TOP-REAL 2)|K1 by FUNCT_2:def 1;
  then A5:dom f=dom g3 by FUNCT_2:def 1;
    for x being set st x in dom f holds f.x=g3.x
   proof let x be set;assume
     A6: x in dom f; then A7: x in dom g3 by A4,FUNCT_2:def 1;
     then x in K1 by A4,JORDAN1:1;
     then reconsider r=x as Point of (TOP-REAL 2);
     reconsider s=x as Point of (TOP-REAL 2)|K1 by A6,FUNCT_2:def 1;
     A8:f.r=(|.r.|)* ((r`2/|.r.|-sn)/(1+sn)) by A1,A4,A7;
     A9:g2.s=proj2.s by Lm3;
     A10:g1.s=(2 NormF).s by Lm5;
     A11:proj2.r=r`2 by PSCOMP_1:def 29;
       (2 NormF).r=|.r.| by Def1;
    hence f.x=g3.x by A3,A8,A9,A10,A11;
   end;
hence f is continuous by A3,A5,FUNCT_1:9;
end;

theorem Th94:
 for sn being Real,K1 being non empty Subset of TOP-REAL 2,
f being map of (TOP-REAL 2)|K1,R^1 st -1<sn & sn<1 &
(for p being Point of (TOP-REAL 2) st p in the carrier of (TOP-REAL 2)|K1
  holds f.p=|.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)))
  & (for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds
        q`1>=0 & q`2/|.q.|>=sn & q<>0.REAL 2) holds
f is continuous
proof
 let sn be Real,K1 be non empty Subset of TOP-REAL 2,
f be map of (TOP-REAL 2)|K1,R^1;
 assume A1: -1<sn & sn<1 &
(for p being Point of (TOP-REAL 2) st p in the carrier of (TOP-REAL 2)|K1
  holds f.p=|.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)))
  & (for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds
            q`1>=0 & q`2/|.q.|>=sn & q<>0.REAL 2);
  set a=sn, b=(1-sn);
  A2: b>0 by A1,Th3;
  reconsider g2=proj2|K1 as continuous map of (TOP-REAL 2)|K1,R^1 by Lm3;
  reconsider g1=(2 NormF)|K1 as continuous map of (TOP-REAL 2)|K1,R^1
                          by Lm5;
    for q being Point of TOP-REAL 2 st q in
    the carrier of (TOP-REAL 2)|K1 holds q<>0.REAL 2 by A1;
  then for q being Point of (TOP-REAL 2)|K1 holds g1.q<>0 by Lm6;
  then consider g3 being map of (TOP-REAL 2)|K1,R^1 such that
  A3: (for q being Point of
  (TOP-REAL 2)|K1,r1,r2 being real number st g2.q=r1 & g1.q=r2
  holds g3.q=r2*(sqrt(abs(1-((r1/r2-a)/b)^2))))
          & g3 is continuous by A2,Th15;
  A4:dom g3=the carrier of (TOP-REAL 2)|K1 by FUNCT_2:def 1;
  then A5:dom f=dom g3 by FUNCT_2:def 1;
    for x being set st x in dom f holds f.x=g3.x
   proof let x be set;assume
     A6: x in dom f;
     then x in K1 by A4,A5,JORDAN1:1;
     then reconsider r=x as Point of (TOP-REAL 2);
     reconsider s=x as Point of (TOP-REAL 2)|K1 by A6,FUNCT_2:def 1;
     A7: now assume |.r.|=0; then r=0.REAL 2 by TOPRNS_1:25;
      hence contradiction by A1,A4,A5,A6;
     end;
     A8: |.r.|>=0 by TOPRNS_1:26;
       |.r.|^2=(r`1)^2+(r`2)^2 by JGRAPH_3:10;
     then |.r.|^2 -(r`2)^2 =(r`1)^2 by XCMPLX_1:26;
     then (r`2)^2 -(|.r.|)^2 =-(r`1)^2 by XCMPLX_1:143;
     then A9: ((r`2) -(|.r.|))*((r`2)+|.r.|) =-(r`1)^2 by SQUARE_1:67;
       (r`1)^2>=0 by SQUARE_1:72;
     then -(r`1)^2 <=0 by REAL_1:66;
     then -(|.r.|)<=r`2 & r`2<= |.r.| by A8,A9,JGRAPH_3:4;
     then r`2/|.r.| <= |.r.|/|.r.| by A7,A8,REAL_1:73;
     then r`2/|.r.|<=1 by A7,XCMPLX_1:60;
     then A10: r`2/|.r.|-sn<=(1-sn) by REAL_1:49;
     A11: now assume (1-sn)^2=0;
       then 1-sn+sn=0+sn by SQUARE_1:73;
      hence contradiction by A1,XCMPLX_1:27;
     end;
     A12: 1-sn>0 by A1,Th3;
      A13: (1-sn)^2>=0 by SQUARE_1:72;
       sn<=r`2/(|.r.|) by A1,A4,A5,A6;
     then sn-r`2/|.r.|<=0 by REAL_2:106;
     then -(sn- r`2/|.r.|)>=-(1-sn) by A12,REAL_1:50;
     then -(1-sn)<= r`2/|.r.|-sn by XCMPLX_1:143;
     then (r`2/|.r.|-sn)^2<=(1-sn)^2 by A10,JGRAPH_2:7;
     then (r`2/|.r.|-sn)^2/(1-sn)^2<=(1-sn)^2/(1-sn)^2 by A11,A13,REAL_1:73;
     then (r`2/|.r.|-sn)^2/(1-sn)^2<=1 by A11,XCMPLX_1:60;
     then ((r`2/|.r.|-sn)/(1-sn))^2<=1 by SQUARE_1:69;
     then 1-((r`2/|.r.|-sn)/(1-sn))^2>=0 by SQUARE_1:12;
     then abs(1-((r`2/|.r.|-sn)/(1-sn))^2)
          =1-((r`2/|.r.|-sn)/(1-sn))^2 by ABSVALUE:def 1;
     then A14:f.r=(|.r.|)*(sqrt(abs(1-((r`2/|.r.|-sn)/(1-sn))^2))) by A1,A4,A5,
A6;
     A15:g2.s=proj2.s by Lm3;
     A16:g1.s=(2 NormF).s by Lm5;
     A17:proj2.r=r`2 by PSCOMP_1:def 29;
       (2 NormF).r=|.r.| by Def1;
    hence f.x=g3.x by A3,A14,A15,A16,A17;
   end;
hence f is continuous by A3,A5,FUNCT_1:9;
end;

theorem Th95:
 for sn being Real,K1 being non empty Subset of TOP-REAL 2,
f being map of (TOP-REAL 2)|K1,R^1 st
  -1<sn & sn<1 &
(for p being Point of (TOP-REAL 2) st p in the carrier of (TOP-REAL 2)|K1
  holds f.p=|.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)))
  & (for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds
        q`1>=0 & q`2/|.q.|<=sn & q<>0.REAL 2) holds
f is continuous
proof
 let sn be Real,K1 be non empty Subset of TOP-REAL 2,
f be map of (TOP-REAL 2)|K1,R^1;
 assume A1:-1<sn & sn<1 &
(for p being Point of (TOP-REAL 2) st p in the carrier of (TOP-REAL 2)|K1
  holds f.p=|.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)))
  & (for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds
            q`1>=0 & q`2/|.q.|<=sn & q<>0.REAL 2);
  set a=sn, b=(1+sn);
  A2: 1+sn>0 by A1,Th3;
  reconsider g2=proj2|K1 as continuous map of (TOP-REAL 2)|K1,R^1 by Lm3;
  reconsider g1=(2 NormF)|K1 as continuous map of (TOP-REAL 2)|K1,R^1
                          by Lm5;
    for q being Point of TOP-REAL 2 st q in
    the carrier of (TOP-REAL 2)|K1 holds q<>0.REAL 2 by A1;
  then for q being Point of (TOP-REAL 2)|K1 holds g1.q<>0 by Lm6;
  then consider g3 being map of (TOP-REAL 2)|K1,R^1 such that
  A3: (for q being Point of
  (TOP-REAL 2)|K1,r1,r2 being real number st g2.q=r1 & g1.q=r2
  holds g3.q=r2*(sqrt(abs(1-((r1/r2-a)/b)^2))))
          & g3 is continuous by A2,Th15;
  A4:dom g3=the carrier of (TOP-REAL 2)|K1 by FUNCT_2:def 1;
  then A5:dom f=dom g3 by FUNCT_2:def 1;
    for x being set st x in dom f holds f.x=g3.x
   proof let x be set;assume
     A6:x in dom f;
     then x in K1 by A4,A5,JORDAN1:1;
     then reconsider r=x as Point of (TOP-REAL 2);
     reconsider s=x as Point of (TOP-REAL 2)|K1 by A6,FUNCT_2:def 1;
     A7: now assume |.r.|=0; then r=0.REAL 2 by TOPRNS_1:25;
      hence contradiction by A1,A4,A5,A6;
     end;
     A8: |.r.|>=0 by TOPRNS_1:26;
       |.r.|^2=(r`1)^2+(r`2)^2 by JGRAPH_3:10;
     then |.r.|^2 -(r`2)^2 =(r`1)^2 by XCMPLX_1:26;
     then (r`2)^2 -(|.r.|)^2 =-(r`1)^2 by XCMPLX_1:143;
     then A9: ((r`2) -(|.r.|))*((r`2)+|.r.|) =-(r`1)^2 by SQUARE_1:67;
       (r`1)^2>=0 by SQUARE_1:72;
     then -(r`1)^2 <=0 by REAL_1:66;
     then -(|.r.|)<=r`2 & r`2<= |.r.| by A8,A9,JGRAPH_3:4;
     then r`2/|.r.| >= (-(|.r.|))/|.r.| by A7,A8,REAL_1:73;
     then r`2/|.r.|>= -1 by A7,XCMPLX_1:198;
     then r`2/|.r.|-sn>=-1-sn by REAL_1:49;
     then A10: r`2/|.r.|-sn>=-(1+sn) by XCMPLX_1:161;
     A11: (1+sn)^2>0 by A2,SQUARE_1:74;
       sn>=r`2/(|.r.|) by A1,A4,A5,A6;
     then sn-r`2/|.r.|>=0 by SQUARE_1:12;
     then -(sn-r`2/|.r.|)<=-0 by REAL_1:50;
     then r`2/|.r.|-sn<=-0 by XCMPLX_1:143;
     then r`2/|.r.|-sn<=1+sn by A2,AXIOMS:22;
     then (r`2/|.r.|-sn)^2<=(1+sn)^2 by A10,JGRAPH_2:7;
     then (r`2/|.r.|-sn)^2/(1+sn)^2<=(1+sn)^2/(1+sn)^2 by A11,REAL_1:73;
     then (r`2/|.r.|-sn)^2/(1+sn)^2<=1 by A11,XCMPLX_1:60;
     then ((r`2/|.r.|-sn)/(1+sn))^2<=1 by SQUARE_1:69;
     then 1-((r`2/|.r.|-sn)/(1+sn))^2>=0 by SQUARE_1:12;
     then abs(1-((r`2/|.r.|-sn)/(1+sn))^2)
          =1-((r`2/|.r.|-sn)/(1+sn))^2 by ABSVALUE:def 1;
     then A12:f.r=(|.r.|)*(sqrt(abs(1-((r`2/|.r.|-sn)/(1+sn))^2))) by A1,A4,A5,
A6;
     A13:g2.s=proj2.s by Lm3;
     A14:g1.s=(2 NormF).s by Lm5;
     A15:proj2.r=r`2 by PSCOMP_1:def 29;
       (2 NormF).r=|.r.| by Def1;
    hence f.x=g3.x by A3,A12,A13,A14,A15;
   end;
hence f is continuous by A3,A5,FUNCT_1:9;
end;

theorem Th96: for sn being Real,
K0,B0 being Subset of TOP-REAL 2,f being
map of (TOP-REAL 2)|K0,(TOP-REAL 2)|B0
st -1<sn & sn<1 & f=(sn-FanMorphE)|K0 &
B0={q where q is Point of TOP-REAL 2: q`1>=0 & q<>0.REAL 2}
& K0={p: p`2/|.p.|>=sn & p`1>=0 & p<>0.REAL 2}
holds f is continuous
proof let sn be Real,K0,B0 be Subset of TOP-REAL 2,f be
map of (TOP-REAL 2)|K0,(TOP-REAL 2)|B0;
assume A1: -1<sn & sn<1 & f=(sn-FanMorphE)|K0 &
B0={q where q is Point of TOP-REAL 2: q`1>=0 & q<>0.REAL 2} &
K0={p: p`2/|.p.|>=sn & p`1>=0 & p<>0.REAL 2};
 set cn=sqrt(1-sn^2);
 set p0=|[cn,sn]|;
 A2: p0`2=sn by EUCLID:56;
 A3: p0`1=cn by EUCLID:56;
     sn^2<1^2 by A1,JGRAPH_2:8;
 then A4: 1-sn^2>0 by SQUARE_1:11,59;
 then A5: --cn>0 by SQUARE_1:93;
 A6: |.p0.|=sqrt((cn)^2+sn^2) by A2,A3,JGRAPH_3:10;
   cn^2=1-sn^2 by A4,SQUARE_1:def 4;
 then |.p0.|=1 by A6,SQUARE_1:83,XCMPLX_1:27;
 then p0`2/|.p0.|=sn by EUCLID:56;
 then A7: p0 in K0 by A1,A3,A5,JGRAPH_2:11;
then reconsider K1=K0 as non empty Subset of TOP-REAL 2;
  A8: K0 c= B0
   proof let x be set;assume x in K0;
      then consider p8 being Point of TOP-REAL 2 such that
      A9: x=p8 & (p8`2/|.p8.|>=sn & p8`1>=0 & p8<>0.REAL 2) by A1;
    thus x in B0 by A1,A9;
   end;
A10:dom ((proj2)*((sn-FanMorphE)|K1)) c= dom ((sn-FanMorphE)|K1)
                                  by RELAT_1:44;
 dom ((sn-FanMorphE)|K1) c= dom ((proj2)*((sn-FanMorphE)|K1))
proof let x be set;assume A11:x in dom ((sn-FanMorphE)|K1);
   then A12:x in dom (sn-FanMorphE) /\ K1 by FUNCT_1:68;
   A13:((sn-FanMorphE)|K1).x=(sn-FanMorphE).x by A11,FUNCT_1:68;
   A14: dom proj2 = (the carrier of TOP-REAL 2) by FUNCT_2:def 1;
     x in dom (sn-FanMorphE) by A12,XBOOLE_0:def 3;
   then (sn-FanMorphE).x in rng (sn-FanMorphE) by FUNCT_1:12;
  hence x in dom ((proj2)*((sn-FanMorphE)|K1)) by A11,A13,A14,FUNCT_1:21;
end;
then A15:dom ((proj2)*((sn-FanMorphE)|K1))
=dom ((sn-FanMorphE)|K1) by A10,XBOOLE_0:def 10
.=dom (sn-FanMorphE) /\ K1 by FUNCT_1:68
.=(the carrier of TOP-REAL 2)/\ K1 by FUNCT_2:def 1
.=K1 by XBOOLE_1:28
.=the carrier of (TOP-REAL 2)|K1 by JORDAN1:1;
A16:rng ((proj2)*((sn-FanMorphE)|K1)) c= rng (proj2) by RELAT_1:45;
  rng proj2 c= the carrier of R^1 by JGRAPH_2:17,RELSET_1:12;
then rng ((proj2)*((sn-FanMorphE)|K1)) c= the carrier of R^1 by A16,XBOOLE_1:1
;
then (proj2)*((sn-FanMorphE)|K1) is Function of the carrier of (TOP-REAL 2)|K1
,
the carrier of R^1 by A15,FUNCT_2:4;
then reconsider g2=(proj2)*((sn-FanMorphE)|K1) as map of (TOP-REAL 2)|K1,R^1
                                ;
A17:dom ((proj1)*((sn-FanMorphE)|K1))
                     c= dom ((sn-FanMorphE)|K1) by RELAT_1:44;
 dom ((sn-FanMorphE)|K1) c= dom ((proj1)*((sn-FanMorphE)|K1))
proof let x be set;assume A18:x in dom ((sn-FanMorphE)|K1);
   then A19:x in dom (sn-FanMorphE) /\ K1 by FUNCT_1:68;
   A20:((sn-FanMorphE)|K1).x=(sn-FanMorphE).x by A18,FUNCT_1:68;
   A21: dom proj1 = (the carrier of TOP-REAL 2) by FUNCT_2:def 1;
     x in dom (sn-FanMorphE) by A19,XBOOLE_0:def 3;
   then (sn-FanMorphE).x in rng (sn-FanMorphE) by FUNCT_1:12;
  hence x in dom ((proj1)*((sn-FanMorphE)|K1)) by A18,A20,A21,FUNCT_1:21;
end;
then A22:dom ((proj1)*((sn-FanMorphE)|K1))
=dom ((sn-FanMorphE)|K1) by A17,XBOOLE_0:def 10
.=dom (sn-FanMorphE) /\ K1 by FUNCT_1:68
.=((the carrier of TOP-REAL 2))/\ K1 by FUNCT_2:def 1
.=K1 by XBOOLE_1:28
.=the carrier of (TOP-REAL 2)|K1 by JORDAN1:1;
A23:rng ((proj1)*((sn-FanMorphE)|K1)) c= rng (proj1) by RELAT_1:45;
  rng proj1 c= the carrier of R^1 by JGRAPH_2:16,RELSET_1:12;
then rng ((proj1)*((sn-FanMorphE)|K1)) c= the carrier of R^1 by A23,XBOOLE_1:1
;
then (proj1)*((sn-FanMorphE)|K1) is Function of the carrier of (TOP-REAL 2)|K1
,
the carrier of R^1 by A22,FUNCT_2:4;
then reconsider g1=(proj1)*((sn-FanMorphE)|K1) as map of (TOP-REAL 2)|K1,R^1
                                ;
A24: for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds q`1>=0 & q<>0.REAL 2
   proof let q be Point of TOP-REAL 2;
    assume A25:q in the carrier of (TOP-REAL 2)|K1;
        the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
     then consider p3 being Point of TOP-REAL 2 such that
     A26: q=p3 & (p3`2/|.p3.|>=sn & p3`1>=0 & p3<>0.REAL 2) by A1,A25;
    thus q`1>=0 & q<>0.REAL 2 by A26;
   end;
    for p being Point of TOP-REAL 2 st p in the carrier of (TOP-REAL 2)|K1
  holds g2.p=|.p.|* ((p`2/|.p.|-sn)/(1-sn))
   proof let p be Point of TOP-REAL 2;
    assume A27: p in the carrier of (TOP-REAL 2)|K1;
     A28:dom ((sn-FanMorphE)|K1)=dom (sn-FanMorphE) /\ K1 by FUNCT_1:68
     .=((the carrier of TOP-REAL 2))/\ K1 by FUNCT_2:def 1
     .=K1 by XBOOLE_1:28;
     A29:the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
     then consider p3 being Point of TOP-REAL 2 such that
     A30: p=p3 & (p3`2/|.p3.|>=sn & p3`1>=0 & p3<>0.REAL 2) by A1,A27;
     A31:(sn-FanMorphE).p
     =|[|.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)),
      |.p.|* ((p`2/|.p.|-sn)/(1-sn))]| by A1,A30,Th91;
      ((sn-FanMorphE)|K1).p=(sn-FanMorphE).p by A27,A29,FUNCT_1:72;
     then g2.p=(proj2).(|[|.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)),
            |.p.|* ((p`2/|.p.|-sn)/(1-sn))]|)
                             by A27,A28,A29,A31,FUNCT_1:23
      .=(|[|.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)),
        |.p.|* ((p`2/|.p.|-sn)/(1-sn))]|)`2 by PSCOMP_1:def 29
      .=|.p.|* ((p`2/|.p.|-sn)/(1-sn)) by EUCLID:56;
    hence g2.p=|.p.|* ((p`2/|.p.|-sn)/(1-sn));
   end;
then consider f2 being map of (TOP-REAL 2)|K1,R^1 such that
  A32:for p being Point of TOP-REAL 2 st p in the carrier of (TOP-REAL 2)|K1
  holds f2.p=|.p.|* ((p`2/|.p.|-sn)/(1-sn));
A33:f2 is continuous by A1,A24,A32,Th92;
    for p being Point of TOP-REAL 2 st p in the carrier of (TOP-REAL 2)|K1
  holds g1.p=|.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2))
   proof let p be Point of TOP-REAL 2;
    assume A34: p in the carrier of (TOP-REAL 2)|K1;
     A35:dom ((sn-FanMorphE)|K1)=dom (sn-FanMorphE) /\ K1 by FUNCT_1:68
     .=((the carrier of TOP-REAL 2))/\ K1 by FUNCT_2:def 1
     .=K1 by XBOOLE_1:28;
     A36:the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
     then consider p3 being Point of TOP-REAL 2 such that
     A37: p=p3 & (p3`2/|.p3.|>=sn & p3`1>=0 & p3<>0.REAL 2) by A1,A34;
     A38:(sn-FanMorphE).p=|[|.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)),
      |.p.|* ((p`2/|.p.|-sn)/(1-sn))]| by A1,A37,Th91;
      ((sn-FanMorphE)|K1).p=(sn-FanMorphE).p by A34,A36,FUNCT_1:72;
     then g1.p=(proj1).(|[|.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)),
          |.p.|* ((p`2/|.p.|-sn)/(1-sn))]|)
                             by A34,A35,A36,A38,FUNCT_1:23
        .=(|[|.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)),
          |.p.|* ((p`2/|.p.|-sn)/(1-sn)) ]|)`1 by PSCOMP_1:def 28
        .= |.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)) by EUCLID:56;
    hence g1.p=|.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2));
   end;
then consider f1 being map of (TOP-REAL 2)|K1,R^1 such that
  A39:for p being Point of TOP-REAL 2 st p in the carrier of (TOP-REAL 2)|K1
  holds f1.p=|.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2));
    for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds q`1>=0 & q`2/|.q.|>=sn & q<>0.REAL 2
   proof let q be Point of TOP-REAL 2;
    assume A40:q in the carrier of (TOP-REAL 2)|K1;
        the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
     then consider p3 being Point of TOP-REAL 2 such that
     A41: q=p3 & (p3`2/|.p3.|>=sn & p3`1>=0 & p3<>0.REAL 2) by A1,A40;
    thus thesis by A41;
   end;
then A42:f1 is continuous by A1,A39,Th94;
  for x,y,r,s being real number st |[x,y]| in K1 &
  r=f1.(|[x,y]|) & s=f2.(|[x,y]|) holds f.(|[x,y]|)=|[r,s]|
  proof let x,y,r,s be real number;
   assume A43: |[x,y]| in K1 & r=f1.(|[x,y]|) & s=f2.(|[x,y]|);
     set p99=|[x,y]|;
     A44:the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
     consider p3 being Point of TOP-REAL 2 such that
     A45: p99=p3 & (p3`2/|.p3.|>=sn & p3`1>=0 & p3<>0.REAL 2) by A1,A43;
     A46:f1.p99=|.p99.|*(sqrt(1-((p99`2/|.p99.|-sn)/(1-sn))^2))
                          by A39,A43,A44;
      ((sn-FanMorphE)|K0).(|[x,y]|)=((sn-FanMorphE)).(|[x,y]|) by A43,FUNCT_1:
72
    .= |[|.p99.|*(sqrt(1-((p99`2/|.p99.|-sn)/(1-sn))^2)),
      |.p99.|* ((p99`2/|.p99.|-sn)/(1-sn))]| by A1,A45,Th91
    .=|[r,s]| by A32,A43,A44,A46;
   hence f.(|[x,y]|)=|[r,s]| by A1;
  end;
hence f is continuous by A7,A8,A33,A42,JGRAPH_2:45;
end;

theorem Th97: for sn being Real,
K0,B0 being Subset of TOP-REAL 2,f being
map of (TOP-REAL 2)|K0,(TOP-REAL 2)|B0
st -1<sn & sn<1 & f=(sn-FanMorphE)|K0 &
B0={q where q is Point of TOP-REAL 2: q`1>=0 & q<>0.REAL 2}
& K0={p: p`2/|.p.|<=sn & p`1>=0 & p<>0.REAL 2}
holds f is continuous
proof let sn be Real,K0,B0 be Subset of TOP-REAL 2,f be
map of (TOP-REAL 2)|K0,(TOP-REAL 2)|B0;
assume A1: -1<sn & sn<1 & f=(sn-FanMorphE)|K0 &
B0={q where q is Point of TOP-REAL 2: q`1>=0 & q<>0.REAL 2} &
K0={p: p`2/|.p.|<=sn & p`1>=0 & p<>0.REAL 2};
 set cn=sqrt(1-sn^2);
 set p0=|[cn,sn]|;
 A2: p0`2=sn by EUCLID:56;
 A3: p0`1=cn by EUCLID:56;
     sn^2<1^2 by A1,JGRAPH_2:8;
 then A4: 1-sn^2>0 by SQUARE_1:11,59;
 then A5: --cn>0 by SQUARE_1:93;
 A6: |.p0.|=sqrt((cn)^2+sn^2) by A2,A3,JGRAPH_3:10;
   cn^2=1-sn^2 by A4,SQUARE_1:def 4;
 then |.p0.|=1 by A6,SQUARE_1:83,XCMPLX_1:27;
 then p0`2/|.p0.|=sn by EUCLID:56;
 then A7: p0 in K0 by A1,A3,A5,JGRAPH_2:11;
 then reconsider K1=K0 as non empty Subset of TOP-REAL 2;
  A8: K0 c= B0
   proof let x be set;assume x in K0;
      then consider p8 being Point of TOP-REAL 2 such that
      A9: x=p8 & (
      p8`2/|.p8.|<=sn & p8`1>=0 & p8<>0.REAL 2) by A1;
    thus x in B0 by A1,A9;
   end;
A10:dom ((proj2)*((sn-FanMorphE)|K1)) c= dom ((sn-FanMorphE)|K1)
                                  by RELAT_1:44;
 dom ((sn-FanMorphE)|K1) c= dom ((proj2)*((sn-FanMorphE)|K1))
proof let x be set;assume A11:x in dom ((sn-FanMorphE)|K1);
   then A12:x in dom (sn-FanMorphE) /\ K1 by FUNCT_1:68;
   A13:((sn-FanMorphE)|K1).x=(sn-FanMorphE).x by A11,FUNCT_1:68;
   A14: dom proj2 = (the carrier of TOP-REAL 2) by FUNCT_2:def 1;
     x in dom (sn-FanMorphE) by A12,XBOOLE_0:def 3;
   then (sn-FanMorphE).x in rng (sn-FanMorphE) by FUNCT_1:12;
  hence x in dom ((proj2)*((sn-FanMorphE)|K1)) by A11,A13,A14,FUNCT_1:21;
end;
then A15:dom ((proj2)*((sn-FanMorphE)|K1))
=dom ((sn-FanMorphE)|K1) by A10,XBOOLE_0:def 10
.=dom (sn-FanMorphE) /\ K1 by FUNCT_1:68
.=(the carrier of TOP-REAL 2)/\ K1 by FUNCT_2:def 1
.=K1 by XBOOLE_1:28
.=the carrier of (TOP-REAL 2)|K1 by JORDAN1:1;
A16:rng ((proj2)*((sn-FanMorphE)|K1)) c= rng (proj2) by RELAT_1:45;
  rng proj2 c= the carrier of R^1 by JGRAPH_2:17,RELSET_1:12;
then rng ((proj2)*((sn-FanMorphE)|K1)) c= the carrier of R^1 by A16,XBOOLE_1:1
;
then (proj2)*((sn-FanMorphE)|K1) is Function of the carrier of (TOP-REAL 2)|K1
,
the carrier of R^1 by A15,FUNCT_2:4;
then reconsider g2=(proj2)*((sn-FanMorphE)|K1) as map of (TOP-REAL 2)|K1,R^1
                                ;
A17:dom ((proj1)*((sn-FanMorphE)|K1))
                     c= dom ((sn-FanMorphE)|K1) by RELAT_1:44;
 dom ((sn-FanMorphE)|K1) c= dom ((proj1)*((sn-FanMorphE)|K1))
proof let x be set;assume A18:x in dom ((sn-FanMorphE)|K1);
   then A19:x in dom (sn-FanMorphE) /\ K1 by FUNCT_1:68;
   A20:((sn-FanMorphE)|K1).x=(sn-FanMorphE).x by A18,FUNCT_1:68;
   A21: dom proj1 = (the carrier of TOP-REAL 2) by FUNCT_2:def 1;
     x in dom (sn-FanMorphE) by A19,XBOOLE_0:def 3;
   then (sn-FanMorphE).x in rng (sn-FanMorphE) by FUNCT_1:12;
  hence x in dom ((proj1)*((sn-FanMorphE)|K1)) by A18,A20,A21,FUNCT_1:21;
end;
then A22:dom ((proj1)*((sn-FanMorphE)|K1))
=dom ((sn-FanMorphE)|K1) by A17,XBOOLE_0:def 10
.=dom (sn-FanMorphE) /\ K1 by FUNCT_1:68
.=((the carrier of TOP-REAL 2))/\ K1 by FUNCT_2:def 1
.=K1 by XBOOLE_1:28
.=the carrier of (TOP-REAL 2)|K1 by JORDAN1:1;
A23:rng ((proj1)*((sn-FanMorphE)|K1)) c= rng (proj1) by RELAT_1:45;
  rng proj1 c= the carrier of R^1 by JGRAPH_2:16,RELSET_1:12;
then rng ((proj1)*((sn-FanMorphE)|K1)) c= the carrier of R^1 by A23,XBOOLE_1:1
;
then (proj1)*((sn-FanMorphE)|K1) is Function of the carrier of (TOP-REAL 2)|K1
,
the carrier of R^1 by A22,FUNCT_2:4;
then reconsider g1=(proj1)*((sn-FanMorphE)|K1) as map of (TOP-REAL 2)|K1,R^1
                                ;
A24: for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds q`1>=0 & q<>0.REAL 2
   proof let q be Point of TOP-REAL 2;
    assume A25:q in the carrier of (TOP-REAL 2)|K1;
        the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
     then consider p3 being Point of TOP-REAL 2 such that
     A26: q=p3 & (p3`2/|.p3.|<=sn & p3`1>=0 & p3<>0.REAL 2) by A1,A25;
    thus q`1>=0 & q<>0.REAL 2 by A26;
   end;
    for p being Point of TOP-REAL 2 st p in the carrier of (TOP-REAL 2)|K1
  holds g2.p=|.p.|* ((p`2/|.p.|-sn)/(1+sn))
   proof let p be Point of TOP-REAL 2;
    assume A27: p in the carrier of (TOP-REAL 2)|K1;
     A28:dom ((sn-FanMorphE)|K1)=dom (sn-FanMorphE) /\ K1 by FUNCT_1:68
     .=((the carrier of TOP-REAL 2))/\ K1 by FUNCT_2:def 1
     .=K1 by XBOOLE_1:28;
     A29:the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
     then consider p3 being Point of TOP-REAL 2 such that
     A30: p=p3 & (p3`2/|.p3.|<=sn & p3`1>=0 & p3<>0.REAL 2) by A1,A27;
     A31:(sn-FanMorphE).p
     =|[|.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)),
      |.p.|* ((p`2/|.p.|-sn)/(1+sn))]| by A1,A30,Th91;
      ((sn-FanMorphE)|K1).p=(sn-FanMorphE).p by A27,A29,FUNCT_1:72;
     then g2.p=(proj2).(|[|.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)),
            |.p.|* ((p`2/|.p.|-sn)/(1+sn))]|)
                             by A27,A28,A29,A31,FUNCT_1:23
      .=(|[|.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)),
        |.p.|* ((p`2/|.p.|-sn)/(1+sn))]|)`2 by PSCOMP_1:def 29
      .=|.p.|* ((p`2/|.p.|-sn)/(1+sn)) by EUCLID:56;
    hence g2.p=|.p.|* ((p`2/|.p.|-sn)/(1+sn));
   end;
then consider f2 being map of (TOP-REAL 2)|K1,R^1 such that
  A32:for p being Point of TOP-REAL 2 st p in the carrier of (TOP-REAL 2)|K1
  holds f2.p=|.p.|* ((p`2/|.p.|-sn)/(1+sn));
A33:f2 is continuous by A1,A24,A32,Th93;
    for p being Point of TOP-REAL 2 st p in the carrier of (TOP-REAL 2)|K1
  holds g1.p=|.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2))
   proof let p be Point of TOP-REAL 2;
    assume A34: p in the carrier of (TOP-REAL 2)|K1;
     A35:dom ((sn-FanMorphE)|K1)=dom (sn-FanMorphE) /\ K1 by FUNCT_1:68
     .=((the carrier of TOP-REAL 2))/\ K1 by FUNCT_2:def 1
     .=K1 by XBOOLE_1:28;
     A36:the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
     then consider p3 being Point of TOP-REAL 2 such that
     A37: p=p3 & (p3`2/|.p3.|<=sn & p3`1>=0 & p3<>0.REAL 2) by A1,A34;
     A38:(sn-FanMorphE).p=|[|.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)),
      |.p.|* ((p`2/|.p.|-sn)/(1+sn))]| by A1,A37,Th91;
      ((sn-FanMorphE)|K1).p=(sn-FanMorphE).p by A34,A36,FUNCT_1:72;
     then g1.p=(proj1).(|[|.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)),
          |.p.|* ((p`2/|.p.|-sn)/(1+sn))]|)
                             by A34,A35,A36,A38,FUNCT_1:23
        .=(|[|.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)),
          |.p.|* ((p`2/|.p.|-sn)/(1+sn)) ]|)`1 by PSCOMP_1:def 28
        .= |.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)) by EUCLID:56;
    hence g1.p=|.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2));
   end;
then consider f1 being map of (TOP-REAL 2)|K1,R^1 such that
  A39:for p being Point of TOP-REAL 2 st p in the carrier of (TOP-REAL 2)|K1
  holds f1.p=|.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2));
    for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds q`1>=0 & q`2/|.q.|<=sn & q<>0.REAL 2
   proof let q be Point of TOP-REAL 2;
    assume A40:q in the carrier of (TOP-REAL 2)|K1;
        the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
     then consider p3 being Point of TOP-REAL 2 such that
     A41: q=p3 & (p3`2/|.p3.|<=sn & p3`1>=0 & p3<>0.REAL 2) by A1,A40;
    thus thesis by A41;
   end;
then A42:f1 is continuous by A1,A39,Th95;
  for x,y,r,s being real number st |[x,y]| in K1 &
  r=f1.(|[x,y]|) & s=f2.(|[x,y]|) holds f.(|[x,y]|)=|[r,s]|
  proof let x,y,r,s be real number;
   assume A43: |[x,y]| in K1 & r=f1.(|[x,y]|) & s=f2.(|[x,y]|);
     set p99=|[x,y]|;
     A44:the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
     consider p3 being Point of TOP-REAL 2 such that
     A45: p99=p3 & (p3`2/|.p3.|<=sn & p3`1>=0 & p3<>0.REAL 2) by A1,A43;
     A46:f1.p99=|.p99.|*(sqrt(1-((p99`2/|.p99.|-sn)/(1+sn))^2))
                          by A39,A43,A44;
      ((sn-FanMorphE)|K0).(|[x,y]|)=((sn-FanMorphE)).(|[x,y]|) by A43,FUNCT_1:
72
    .= |[|.p99.|*(sqrt(1-((p99`2/|.p99.|-sn)/(1+sn))^2)),
      |.p99.|* ((p99`2/|.p99.|-sn)/(1+sn))]| by A1,A45,Th91
    .=|[r,s]| by A32,A43,A44,A46;
   hence f.(|[x,y]|)=|[r,s]| by A1;
  end;
hence f is continuous by A7,A8,A33,A42,JGRAPH_2:45;
end;

theorem Th98: for sn being Real, K03 being Subset of TOP-REAL 2
 st K03={p: p`2>=(sn)*(|.p.|) & p`1>=0}
 holds K03 is closed
proof let sn be Real, K003 be Subset of TOP-REAL 2;
 assume A1:K003={p: p`2>=(sn)*(|.p.|) & p`1>=0};
 defpred P[Point of TOP-REAL 2] means ($1`2>=sn*|.$1.|);
  reconsider K1={p7 where p7 is Point of TOP-REAL 2:P[p7]}
  as Subset of TOP-REAL 2 from TopSubset;
  A2: K1 is closed by Lm7;
  defpred Q[Point of TOP-REAL 2] means $1`1>=0;
  reconsider KX={p where p is Point of TOP-REAL 2: Q[p]}
    as Subset of TOP-REAL 2 from TopSubset;
A3: KX is closed by JORDAN6:5;
     {p : P[p] & Q[p]}
    = {p7 where p7 is Point of TOP-REAL 2:P[p7]} /\
     {p1 where p1 is Point of TOP-REAL 2: Q[p1]} from TopInter2;
  hence thesis by A1,A2,A3,TOPS_1:35;
end;

theorem Th99: for sn being Real, K03 being Subset of TOP-REAL 2
 st K03={p: p`2<=(sn)*(|.p.|) & p`1>=0}
 holds K03 is closed
proof let sn be Real, K003 be Subset of TOP-REAL 2;
 assume A1: K003={p: p`2<=(sn)*(|.p.|) & p`1>=0};
 defpred P[Point of TOP-REAL 2] means ($1`2<=sn*|.$1.|);
  reconsider K1={p7 where p7 is Point of TOP-REAL 2:P[p7]}
  as Subset of TOP-REAL 2 from TopSubset;
  A2: K1 is closed by Lm9;
  defpred Q[Point of TOP-REAL 2] means $1`1>=0;
  reconsider KX={p where p is Point of TOP-REAL 2: Q[p]}
    as Subset of TOP-REAL 2 from TopSubset;
A3: KX is closed by JORDAN6:5;
     {p : P[p] & Q[p]}
    = {p7 where p7 is Point of TOP-REAL 2:P[p7]} /\
     {p1 where p1 is Point of TOP-REAL 2: Q[p1]} from TopInter2;
  hence thesis by A1,A2,A3,TOPS_1:35;
end;

theorem Th100: for sn being Real,
K0,B0 being Subset of TOP-REAL 2,f being
map of (TOP-REAL 2)|K0,(TOP-REAL 2)|B0
st -1<sn & sn<1 & f=(sn-FanMorphE)|K0 &
B0=(the carrier of TOP-REAL 2) \ {0.REAL 2}
& K0={p: p`1>=0 & p<>0.REAL 2}
holds f is continuous
proof let sn be Real,K0,B0 be Subset of TOP-REAL 2,f be
map of (TOP-REAL 2)|K0,(TOP-REAL 2)|B0;
assume A1: -1<sn & sn<1 & f=(sn-FanMorphE)|K0 &
B0=(the carrier of TOP-REAL 2)
\ {0.REAL 2} &
K0={p: p`1>=0 & p<>0.REAL 2};
 set cn=sqrt(1-sn^2);
 set p0=|[cn,sn]|;
 A2: p0`2=sn by EUCLID:56;
 A3: p0`1=cn by EUCLID:56;
     sn^2<1 by A1,JGRAPH_2:8,SQUARE_1:59;
 then A4: 1-sn^2>0 by SQUARE_1:11;
 then A5: p0<>0.REAL 2 by A3,JGRAPH_2:11,SQUARE_1:93;
 A6: --cn>0 by A4,SQUARE_1:93;
 A7: p0`1>0 by A3,A4,SQUARE_1:93;
 A8: |.p0.|=sqrt((cn)^2+sn^2) by A2,A3,JGRAPH_3:10;
   cn^2=1-sn^2 by A4,SQUARE_1:def 4;
 then |.p0.|=1 by A8,SQUARE_1:83,XCMPLX_1:27;
 then A9: p0`2/|.p0.|=sn by EUCLID:56;
   p0 in K0 by A1,A7,JGRAPH_2:11;
then reconsider K1=K0 as non empty Subset of TOP-REAL 2;
    p0 in the carrier of TOP-REAL 2 & not p0 in {0.REAL 2}
               by A5,TARSKI:def 1;
  then reconsider D=B0 as non empty Subset of TOP-REAL 2 by A1,XBOOLE_0:def 4;
   A10: p0 in {p: p`2/|.p.|>=sn & p`1>=0 & p<>0.REAL 2} by A7,A9,JGRAPH_2:11;
   defpred P[Point of TOP-REAL 2] means
   $1`2/|.$1.|>=sn & $1`1>=0 & $1<>0.REAL 2;
   A11: {p: P[p]}
     is Subset of TOP-REAL 2 from SubsetD;
   A12: the carrier of ((TOP-REAL 2)|K1)=K1 by JORDAN1:1;
   A13: {p: p`2/|.p.|>=sn & p`1>=0 & p<>0.REAL 2} c= K1
    proof let x be set;assume x in
      {p: p`2/|.p.|>=sn & p`1>=0 & p<>0.REAL 2};
      then consider p such that
        A14:p=x & p`2/|.p.|>=sn & p`1>=0 & p<>0.REAL 2;
     thus x in K1 by A1,A14;
    end;
   then reconsider K00={p: p`2/|.p.|>=sn & p`1>=0 & p<>0.REAL 2}
      as non empty Subset of ((TOP-REAL 2)|K1)
                 by A10,A12;
   reconsider K001={p: p`2/|.p.|>=sn & p`1>=0 & p<>0.REAL 2}
            as non empty Subset of (TOP-REAL 2)
                       by A10,A11;
   defpred P[Point of TOP-REAL 2] means $1`2>=(sn)*(|.$1.|) & $1`1>=0;
     {p: P[p]} is Subset of
   TOP-REAL 2 from SubsetD;
   then reconsider K003={p: p`2>=(sn)*(|.p.|) & p`1>=0}
            as Subset of (TOP-REAL 2);
   A15: p0 in {p: p`2/|.p.|<=sn & p`1>=0 & p<>0.REAL 2} by A3,A6,A9,JGRAPH_2:11
;
   A16: {p: p`2/|.p.|<=sn & p`1>=0 & p<>0.REAL 2} c= K1
    proof let x be set;assume x in
      {p: p`2/|.p.|<=sn & p`1>=0 & p<>0.REAL 2};
      then consider p such that
        A17: p=x & p`2/|.p.|<=sn & p`1>=0 & p<>0.REAL 2;
     thus x in K1 by A1,A17;
    end;
   then reconsider K11={p: p`2/|.p.|<=sn & p`1>=0 & p<>0.REAL 2}
            as non empty Subset of ((TOP-REAL 2)|K1)
            by A12,A15;
    A18: p0 in {p: p`2/|.p.|<=sn & p`1>=0 & p<>0.REAL 2} by A3,A6,A9,JGRAPH_2:
11;
defpred P[Point of TOP-REAL 2] means
$1`2/|.$1.|<=sn & $1`1>=0 & $1<>0.REAL 2;
     {p: P[p]} is Subset of
   TOP-REAL 2 from SubsetD;
   then reconsider K111={p: p`2/|.p.|<=sn & p`1>=0 & p<>0.REAL 2}
            as non empty Subset of (TOP-REAL 2)
                       by A18;
    defpred P[Point of TOP-REAL 2] means
    $1`2<=(sn)*(|.$1.|) & $1`1>=0;
     {p: P[p]} is Subset of
   TOP-REAL 2 from SubsetD;
   then reconsider K004={p: p`2<=(sn)*(|.p.|) & p`1>=0}
            as Subset of (TOP-REAL 2);
    the carrier of (TOP-REAL 2)|B0=the carrier of (TOP-REAL 2)|D;
  then A19: dom f=the carrier of (TOP-REAL 2)|K1 by FUNCT_2:def 1
    .=K1 by JORDAN1:1;
  then A20: dom (f|K00)=K00 by A13,RELAT_1:91
  .= the carrier of ((TOP-REAL 2)|K1)|K00 by JORDAN1:1;
  A21: the carrier of (TOP-REAL 2)|D=D by JORDAN1:1;
    rng (f|K00) c= rng f by RELAT_1:99;
  then rng (f|K00) c=D by A21,XBOOLE_1:1;
  then f|K00 is Function of the carrier of ((TOP-REAL 2)|K1)|K00,
   the carrier of (TOP-REAL 2)|D by A20,A21,FUNCT_2:4;
  then reconsider f1=f|K00 as map of ((TOP-REAL 2)|K1)|K00,(TOP-REAL 2)|D
                           ;
  A22: dom f1=the carrier of ((TOP-REAL 2)|K1)|K00 by FUNCT_2:def 1
  .=K00 by JORDAN1:1;
  A23: dom (sn-FanMorphE)=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
  then A24: dom ((sn-FanMorphE)|K001)=K001 by RELAT_1:91
  .= the carrier of (TOP-REAL 2)|K001 by JORDAN1:1;
  A25: the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
    rng ((sn-FanMorphE)|K001) c= K1
   proof let y be set;assume y in rng ((sn-FanMorphE)|K001);
     then consider x being set such that
     A26: x in dom ((sn-FanMorphE)|K001) & y=((sn-FanMorphE)|K001).x
        by FUNCT_1:def 5;
     A27: dom ((sn-FanMorphE)|K001)=(dom (sn-FanMorphE))/\ K001 by FUNCT_1:68
     .=(the carrier of TOP-REAL 2)/\ K001 by FUNCT_2:def 1
     .=K001 by XBOOLE_1:28;
     then reconsider q=x as Point of TOP-REAL 2 by A26;
     A28: y=(sn-FanMorphE).q by A26,FUNCT_1:70;
     consider p2 being Point of TOP-REAL 2 such that
     A29: p2=q & p2`2/|.p2.|>=sn & p2`1>=0 & p2<>0.REAL 2 by A26,A27;
     A30: sn-FanMorphE.q= |[ |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1-sn))]| by A1,A29,Th91;
     set q4= |[ |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1-sn))]|;
     A31: |.q.|>=0 by TOPRNS_1:26;
       |.q.|<>0 by A29,TOPRNS_1:25;
     then A32: (|.q.|)^2>0 by SQUARE_1:74;
     A33: 1-sn>0 by A1,Th3;
        A34: q4`1= |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2))&
        q4`2= |.q.|* ((q`2/|.q.|-sn)/(1-sn))
                     by EUCLID:56;
        A35: (q`2/|.q.|-sn)>= 0 by A29,SQUARE_1:12;
          0<=(q`1)^2 by SQUARE_1:72;
        then 0+(q`2)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
        then (q`2)^2 <= (|.q.|)^2 by JGRAPH_3:10;
        then (q`2)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A32,REAL_1:73;
        then (q`2)^2/(|.q.|)^2 <= 1 by A32,XCMPLX_1:60;
        then ((q`2)/|.q.|)^2 <= 1 by SQUARE_1:69;
        then 1>=q`2/|.q.| by Th4;
        then A36: 1-sn>=q`2/|.q.|-sn by REAL_1:49;
        A37: -(1-sn)<= -0 by A33,REAL_1:50;
           -(1-sn)<= -( q`2/|.q.|-sn) by A36,REAL_1:50;
         then (-(1-sn))/(1-sn)<=(-( q`2/|.q.|-sn))/(1-sn)
                        by A33,REAL_1:73;
        then A38: -1<=(-( q`2/|.q.|-sn))/(1-sn)
                        by A33,XCMPLX_1:198;
         --(1-sn)>= -(q`2/|.q.|-sn) by A35,A37,REAL_1:50;
       then (-(q`2/|.q.|-sn))/(1-sn)<=1 by A33,REAL_2:117;
       then ((-(q`2/|.q.|-sn))/(1-sn))^2<=1 by A38,JGRAPH_2:7,SQUARE_1:59;
       then A39: 1-((-(q`2/|.q.|-sn))/(1-sn))^2>=0 by SQUARE_1:12;
       then 1-(-((q`2/|.q.|-sn))/(1-sn))^2>=0 by XCMPLX_1:188;
       then A40: 1-((q`2/|.q.|-sn)/(1-sn))^2>=0 by SQUARE_1:61;
         sqrt(1-((-(q`2/|.q.|-sn))/(1-sn))^2)>=0 by A39,SQUARE_1:def 4;
       then sqrt(1-(-(q`2/|.q.|-sn))^2/(1-sn)^2)>=0 by SQUARE_1:69;
       then sqrt(1-(q`2/|.q.|-sn)^2/(1-sn)^2)>=0 by SQUARE_1:61;
       then A41: sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)>=0 by SQUARE_1:69;
        A42: (q4`1)^2= (|.q.|)^2*(sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2))^2
                         by A34,SQUARE_1:68
         .= (|.q.|)^2*(1-((q`2/|.q.|-sn)/(1-sn))^2) by A40,SQUARE_1:def 4;
        A43: (q4`2)^2= (|.q.|)^2*((q`2/|.q.|-sn)/(1-sn))^2
                         by A34,SQUARE_1:68;
          (|.q4.|)^2=(q4`1)^2+(q4`2)^2 by JGRAPH_3:10
        .=(|.q.|)^2*(1-((q`2/|.q.|-sn)/(1-sn))^2
           +((q`2/|.q.|-sn)/(1-sn))^2) by A42,A43,XCMPLX_1:8
        .=(|.q.|)^2*1 by XCMPLX_1:27 .=(|.q.|)^2;
     then q4`1>=0 & q4<>0.REAL 2 by A31,A32,A34,A41,SQUARE_1:19,60,TOPRNS_1:24
;
    hence y in K1 by A1,A28,A30;
   end;
  then (sn-FanMorphE)|K001 is Function of the carrier of (TOP-REAL 2)|K001,
   the carrier of (TOP-REAL 2)|K1 by A24,A25,FUNCT_2:4;
  then reconsider f3=(sn-FanMorphE)|K001
   as map of (TOP-REAL 2)|K001,(TOP-REAL 2)|K1 ;
  A44: dom (f|K11)=K11 by A16,A19,RELAT_1:91
  .= the carrier of ((TOP-REAL 2)|K1)|K11 by JORDAN1:1;
  A45: the carrier of (TOP-REAL 2)|D =D by JORDAN1:1;
    rng (f|K11) c= rng f by RELAT_1:99;
  then rng (f|K11) c=D by A45,XBOOLE_1:1;
  then f|K11 is Function of the carrier of ((TOP-REAL 2)|K1)|K11,
   the carrier of (TOP-REAL 2)|D by A44,A45,FUNCT_2:4;
  then reconsider f2=f|K11 as map of ((TOP-REAL 2)|K1)|K11,(TOP-REAL 2)|D
                   ;
  A46: dom f2=the carrier of ((TOP-REAL 2)|K1)|K11 by FUNCT_2:def 1
  .=K11 by JORDAN1:1;
  A47: dom ((sn-FanMorphE)|K111)=K111 by A23,RELAT_1:91
  .= the carrier of (TOP-REAL 2)|K111 by JORDAN1:1;
  A48: the carrier of (TOP-REAL 2)|K1 =K1 by JORDAN1:1;
    rng ((sn-FanMorphE)|K111) c= K1
   proof let y be set;assume y in rng ((sn-FanMorphE)|K111);
     then consider x being set such that
     A49: x in dom ((sn-FanMorphE)|K111) & y=((sn-FanMorphE)|K111).x
        by FUNCT_1:def 5;
     A50: dom ((sn-FanMorphE)|K111)=(dom (sn-FanMorphE))/\ K111 by FUNCT_1:68
     .=(the carrier of TOP-REAL 2)/\ K111 by FUNCT_2:def 1
     .=K111 by XBOOLE_1:28;
     then reconsider q=x as Point of TOP-REAL 2 by A49;
     A51: y=(sn-FanMorphE).q by A49,FUNCT_1:70;
     consider p2 being Point of TOP-REAL 2 such that
     A52: p2=q & p2`2/|.p2.|<=sn & p2`1>=0 & p2<>0.REAL 2 by A49,A50;
     A53: sn-FanMorphE.q= |[ |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1+sn))]| by A1,A52,Th91;
     set q4= |[ |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1+sn))]|;
     A54: |.q.|>=0 by TOPRNS_1:26;
       |.q.|<>0 by A52,TOPRNS_1:25;
     then A55: (|.q.|)^2>0 by SQUARE_1:74;
   A56: 1+sn>0 by A1,Th3;
        A57: q4`1= |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2))&
        q4`2= |.q.|* ((q`2/|.q.|-sn)/(1+sn))
                     by EUCLID:56;
        A58: (q`2/|.q.|-sn)<=0 by A52,REAL_2:106;
        A59: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
          0<=(q`1)^2 by SQUARE_1:72;
        then 0+(q`2)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
        then (q`2)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A55,A59,REAL_1:73;
        then (q`2)^2/(|.q.|)^2 <= 1 by A55,XCMPLX_1:60;
        then ((q`2)/|.q.|)^2 <= 1 by SQUARE_1:69;
        then -1<=q`2/|.q.| by Th4;
        then -1-sn<=q`2/|.q.|-sn by REAL_1:49;
        then A60: -(1+sn)<=q`2/|.q.|-sn by XCMPLX_1:161;
         (1+sn)/(1+sn)>=(q`2/|.q.|-sn)/(1+sn) by A56,A58,REAL_1:73;
        then A61: 1>=(q`2/|.q.|-sn)/(1+sn) by A56,XCMPLX_1:60;
           (-(1+sn))/(1+sn)<=(( q`2/|.q.|-sn))/(1+sn)
                        by A56,A60,REAL_1:73;
        then -1<=(( q`2/|.q.|-sn))/(1+sn)
                        by A56,XCMPLX_1:198;
       then (((q`2/|.q.|-sn))/(1+sn))^2<=1^2 by A61,JGRAPH_2:7;
       then A62: 1-((q`2/|.q.|-sn)/(1+sn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((q`2/|.q.|-sn)/(1+sn)))^2>=0 by SQUARE_1:61;
       then 1-((-(q`2/|.q.|-sn))/(1+sn))^2>=0 by XCMPLX_1:188;
       then sqrt(1-((-(q`2/|.q.|-sn))/(1+sn))^2)>=0 by SQUARE_1:def 4;
       then sqrt(1-(-(q`2/|.q.|-sn))^2/(1+sn)^2)>=0 by SQUARE_1:69;
       then sqrt(1-(q`2/|.q.|-sn)^2/(1+sn)^2)>=0 by SQUARE_1:61;
       then sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)>=0 by SQUARE_1:69;
       then A63: q4`1>= 0 by A54,A57,SQUARE_1:19;
       A64: (q4`1)^2= (|.q.|)^2*(sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2))^2
                         by A57,SQUARE_1:68
         .= (|.q.|)^2*(1-((q`2/|.q.|-sn)/(1+sn))^2) by A62,SQUARE_1:def 4;
        A65: (q4`2)^2= (|.q.|)^2*((q`2/|.q.|-sn)/(1+sn))^2
                         by A57,SQUARE_1:68;
          (|.q4.|)^2=(q4`1)^2+(q4`2)^2 by JGRAPH_3:10
        .=(|.q.|)^2*(1-((q`2/|.q.|-sn)/(1+sn))^2
           +((q`2/|.q.|-sn)/(1+sn))^2) by A64,A65,XCMPLX_1:8
        .=(|.q.|)^2*1 by XCMPLX_1:27 .=(|.q.|)^2;
        then q4<>0.REAL 2 by A55,SQUARE_1:60,TOPRNS_1:24;
    hence y in K1 by A1,A51,A53,A63;
   end;
  then (sn-FanMorphE)|K111 is Function of the carrier of (TOP-REAL 2)|K111,
   the carrier of (TOP-REAL 2)|K1 by A47,A48,FUNCT_2:4;
  then reconsider f4=(sn-FanMorphE)|K111
   as map of (TOP-REAL 2)|K111,(TOP-REAL 2)|K1 ;
  set T1= ((TOP-REAL 2)|K1)|K00,T2=((TOP-REAL 2)|K1)|K11;
  A66: [#](((TOP-REAL 2)|K1)|K00)=K00 by PRE_TOPC:def 10;
  A67: [#](((TOP-REAL 2)|K1)|K11)=K11 by PRE_TOPC:def 10;
  A68: [#]((TOP-REAL 2)|K1)=K1 by PRE_TOPC:def 10;
  A69: K00 \/ K11 c= K1
   proof let x be set;assume
     A70: x in K00 \/ K11;
       now per cases by A70,XBOOLE_0:def 2;
     case x in K00;
       then consider p8 being Point of TOP-REAL 2 such that
       A71: p8=x & p8`2/|.p8.|>=sn & p8`1>=0 & p8<>0.REAL 2;
      thus x in K1 by A1,A71;
     case x in K11;
       then consider p8 being Point of TOP-REAL 2 such that
       A72: p8=x & p8`2/|.p8.|<=sn & p8`1>=0 & p8<>0.REAL 2;
      thus x in K1 by A1,A72;
     end;
    hence x in K1;
   end;
  A73: K1 c= K00 \/ K11
   proof let x be set;assume x in K1;
     then consider p such that
     A74: p=x & (p`1>=0 & p<>0.REAL 2) by A1;
       now per cases;
     case p`2/|.p.|>=sn;
       then x in K00 by A74;
      hence x in K00 \/ K11 by XBOOLE_0:def 2;
     case p`2/|.p.|<sn;
       then x in K11 by A74;
      hence x in K00 \/ K11 by XBOOLE_0:def 2;
     end;
    hence x in K00 \/ K11;
   end;
  then A75: [#](((TOP-REAL 2)|K1)|K00) \/ [#](((TOP-REAL 2)|K1)|K11)
        =[#]((TOP-REAL 2)|K1) by A66,A67,A68,A69,XBOOLE_0:def 10;
  A76: K003 is closed by Th98;
  A77: K003 /\ K1 c= K00
   proof let x be set;assume x in K003 /\ K1;
     then A78: x in K003 & x in K1 by XBOOLE_0:def 3;
     then consider q1 being Point of TOP-REAL 2 such that
     A79: q1=x & q1`2>=(sn)*(|.q1.|) & q1`1>=0;
     consider q2 being Point of TOP-REAL 2 such that
     A80: q2=x & q2`1>=0 & q2<>0.REAL 2 by A1,A78;
     A81: |.q2.|<>0 by A80,TOPRNS_1:25;
       |.q1.|>=0 by TOPRNS_1:26;
     then q1`2/|.q1.|>=(sn)*(|.q1.|)/|.q1.| by A79,A80,A81,REAL_1:73;
     then q1`2/|.q1.|>=(sn) by A79,A80,A81,XCMPLX_1:90;
    hence x in K00 by A79,A80;
   end;
    K00 c= K003 /\ K1
   proof let x be set;assume x in K00;
     then consider p such that
     A82: p=x & (p`2/|.p.|>=sn & p`1>=0 & p<>0.REAL 2);
     A83: |.p.|<>0 by A82,TOPRNS_1:25;
       |.p.|>=0 by TOPRNS_1:26;
     then p`2/|.p.|*|.p.|>=(sn)*(|.p.|) by A82,AXIOMS:25;
     then p`2>=(sn)*(|.p.|) by A83,XCMPLX_1:88;
     then A84: x in K003 by A82;
       x in K1 by A1,A82;
    hence x in K003 /\ K1 by A84,XBOOLE_0:def 3;
   end;
  then K00=K003 /\ [#]((TOP-REAL 2)|K1) by A68,A77,XBOOLE_0:def 10;
  then A85: K00 is closed by A76,PRE_TOPC:43;
  A86: K004 is closed by Th99;
  A87: K004 /\ K1 c= K11
   proof let x be set;assume x in K004 /\ K1;
     then A88: x in K004 & x in K1 by XBOOLE_0:def 3;
     then consider q1 being Point of TOP-REAL 2 such that
     A89: q1=x & q1`2<=(sn)*(|.q1.|) & q1`1>=0;
     consider q2 being Point of TOP-REAL 2 such that
     A90: q2=x & q2`1>=0 & q2<>0.REAL 2 by A1,A88;
     A91: |.q2.|<>0 by A90,TOPRNS_1:25;
       |.q1.|>=0 by TOPRNS_1:26;
     then q1`2/|.q1.|<=(sn)*(|.q1.|)/|.q1.| by A89,A90,A91,REAL_1:73;
     then q1`2/|.q1.|<=(sn) by A89,A90,A91,XCMPLX_1:90;
    hence x in K11 by A89,A90;
   end;
    K11 c= K004 /\ K1
   proof let x be set;assume x in K11;
     then consider p such that
     A92: p=x & (p`2/|.p.|<=sn & p`1>=0 & p<>0.REAL 2);
     A93: |.p.|<>0 by A92,TOPRNS_1:25;
       |.p.|>=0 by TOPRNS_1:26;
     then p`2/|.p.|*|.p.|<=(sn)*(|.p.|) by A92,AXIOMS:25;
     then p`2<=(sn)*(|.p.|) by A93,XCMPLX_1:88;
     then A94: x in K004 by A92;
       x in K1 by A1,A92;
    hence x in K004 /\ K1 by A94,XBOOLE_0:def 3;
   end;
  then K11=K004 /\ [#]((TOP-REAL 2)|K1) by A68,A87,XBOOLE_0:def 10;
  then A95: K11 is closed by A86,PRE_TOPC:43;
  A96: ((TOP-REAL 2)|K1)|K00=(TOP-REAL 2)|K001 by GOBOARD9:4;
    K1 c= D
   proof let x be set;assume x in K1;
     then consider p6 being Point of TOP-REAL 2 such that
     A97: p6=x & p6`1>=0 & p6<>0.REAL 2 by A1;
       x in the carrier of TOP-REAL 2 & not x in {0.REAL 2} by A97,TARSKI:def 1
;
    hence x in D by A1,XBOOLE_0:def 4;
   end;
  then D=K1 \/ D by XBOOLE_1:12;
  then A98: (TOP-REAL 2)|K1 is SubSpace of (TOP-REAL 2)|D by TOPMETR:5;
    the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
  then A99: f1= f3 by A1,FUNCT_1:82;
    f3 is continuous by A1,Th96;
  then A100: f1 is continuous by A96,A98,A99,TOPMETR:7;
  A101: ((TOP-REAL 2)|K1)|K11=(TOP-REAL 2)|K111 by GOBOARD9:4;
    the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
  then A102: f2= f4 by A1,FUNCT_1:82;
    f4 is continuous by A1,Th97;
  then A103: f2 is continuous by A98,A101,A102,TOPMETR:7;
  A104: D<>{};
  A105: the carrier of (TOP-REAL 2)|B0=B0 by JORDAN1:1;
    for p being set st p in ([#]T1)/\([#]T2) holds f1.p = f2.p
   proof let p be set;
    assume p in ([#]T1)/\([#]T2);
     then A106: p in K00 & p in K11 by A66,A67,XBOOLE_0:def 3;
    hence f1.p=f.p by FUNCT_1:72 .=f2.p by A106,FUNCT_1:72;
   end;
  then consider h being map of (TOP-REAL 2)|K1,(TOP-REAL 2)|D
  such that
  A107: h=f1+*f2 & h is continuous by A66,A67,A75,A85,A95,A100,A103,JGRAPH_2:9;
  A108: dom h=the carrier of ((TOP-REAL 2)|K1) by FUNCT_2:def 1;
  A109: the carrier of ((TOP-REAL 2)|K1)=K0 by JORDAN1:1;
  A110: K0=the carrier of ((TOP-REAL 2)|K0) by JORDAN1:1
  .=dom f by A104,A105,FUNCT_2:def 1;
    for y being set st y in dom h holds h.y=f.y
   proof let y be set;assume A111: y in dom h;
      now per cases by A73,A108,A109,A111,XBOOLE_0:def 2;
    case A112: y in K00 & not y in K11;
      then y in dom f1 \/ dom f2 by A22,XBOOLE_0:def 2;
     hence h.y=f1.y by A46,A107,A112,FUNCT_4:def 1 .=f.y by A112,FUNCT_1:72;
    case A113: y in K11;
      then y in dom f1 \/ dom f2 by A46,XBOOLE_0:def 2;
     hence h.y=f2.y by A46,A107,A113,FUNCT_4:def 1 .=f.y by A113,FUNCT_1:72;
    end;
   hence h.y=f.y;
   end;
 hence thesis by A107,A108,A109,A110,FUNCT_1:9;
end;

theorem Th101: for sn being Real,
K0,B0 being Subset of TOP-REAL 2,f being
map of (TOP-REAL 2)|K0,(TOP-REAL 2)|B0
st -1<sn & sn<1 & f=(sn-FanMorphE)|K0 &
B0=(the carrier of TOP-REAL 2) \ {0.REAL 2}
& K0={p: p`1<=0 & p<>0.REAL 2}
holds f is continuous
proof let sn be Real,K0,B0 be Subset of TOP-REAL 2,f be
map of (TOP-REAL 2)|K0,(TOP-REAL 2)|B0;
assume A1:  -1<sn & sn<1 & f=(sn-FanMorphE)|K0 &
B0=(the carrier of TOP-REAL 2)
\ {0.REAL 2} &
K0={p: p`1<=0 & p<>0.REAL 2};
 set cn=sqrt(1-sn^2);
 set p0=|[-cn,-sn]|;
 A2: p0`1=-cn by EUCLID:56;
     sn^2<1^2 by A1,JGRAPH_2:8;
 then 1-sn^2>0 by SQUARE_1:11,59;
 then --cn>0 by SQUARE_1:93;
 then A3: -cn<0 by REAL_1:66;
 then p0 in K0 by A1,A2,JGRAPH_2:11;
then reconsider K1=K0 as non empty Subset of TOP-REAL 2;
   p0 in the carrier of TOP-REAL 2 & not p0 in {0.REAL 2}
               by A2,A3,JGRAPH_2:11,TARSKI:def 1;
  then reconsider D=B0 as non empty Subset of TOP-REAL 2 by A1,XBOOLE_0:def 4;
   A4: K1 c= D
    proof let x be set;assume x in K1;
      then consider p2 being Point of TOP-REAL 2 such that
      A5: p2=x & p2`1<=0 & p2<>0.REAL 2 by A1;
        not p2 in {0.REAL 2} by A5,TARSKI:def 1;
     hence x in D by A1,A5,XBOOLE_0:def 4;
    end;
     for p being Point of (TOP-REAL 2)|K1,V being Subset of (TOP-REAL 2)|D
      st f.p in V & V is open holds
      ex W being Subset of (TOP-REAL 2)|K1 st
      p in W & W is open & f.:W c= V
    proof let p be Point of (TOP-REAL 2)|K1,V be Subset of (TOP-REAL 2)|D;
      assume A6: f.p in V & V is open;
      then consider V2 being Subset of TOP-REAL 2 such that
      A7: V2 is open & V2 /\ [#]((TOP-REAL 2)|D)=V by TOPS_2:32;
      A8: p in the carrier of (TOP-REAL 2)|K1;
      A9: the carrier of (TOP-REAL 2)|K1=[#]((TOP-REAL 2)|K1) by PRE_TOPC:12;
      A10: [#]((TOP-REAL 2)|K1)=K1 by PRE_TOPC:def 10;
        V2 /\ [#]((TOP-REAL 2)|K1) c= the carrier of (TOP-REAL 2)|K1
       proof let x be set;assume x in V2 /\ [#]((TOP-REAL 2)|K1);
         then x in [#]((TOP-REAL 2)|K1) by XBOOLE_0:def 3;
        hence x in the carrier of (TOP-REAL 2)|K1;
       end;
      then reconsider W2=V2 /\ [#]((TOP-REAL 2)|K1) as Subset of (TOP-REAL 2)|
K1
                               ;
      A11: W2 is open by A7,TOPS_2:32;
      A12: f.p=(sn-FanMorphE).p by A1,A9,A10,FUNCT_1:72;
      consider q being Point of TOP-REAL 2 such that
      A13: q=p & q`1<=0 & q <>0.REAL 2 by A1,A8,A9,A10;
        (sn-FanMorphE).q=q by A13,Th89;
      then p in V2 & p in [#]((TOP-REAL 2)|D)
                       by A6,A7,A12,A13,XBOOLE_0:def 3;
      then A14: p in W2 by A9,XBOOLE_0:def 3;
        f.:W2 c= V
       proof let y be set;assume y in f.:W2;
         then consider x being set such that
         A15: x in dom f & x in W2 & y=f.x by FUNCT_1:def 12;
           f is Function of the carrier of (TOP-REAL 2)|K1,
           the carrier of (TOP-REAL 2)|D;
         then dom f= K1 by A9,A10,FUNCT_2:def 1;
         then consider p4 being Point of TOP-REAL 2 such that
         A16: x=p4 & p4`1<=0 & p4<>0.REAL 2 by A1,A15;
         A17: f.p4=(sn-FanMorphE).p4 by A1,A9,A10,A15,A16,FUNCT_1:72
         .=p4 by A16,Th89;
         A18: p4 in V2 & p4 in [#]((TOP-REAL 2)|K1) by A15,A16,XBOOLE_0:def 3
;
         then p4 in D by A4,A10;
         then p4 in [#]((TOP-REAL 2)|D) by PRE_TOPC:def 10;
        hence y in V by A7,A15,A16,A17,A18,XBOOLE_0:def 3;
       end;
     hence ex W being Subset of (TOP-REAL 2)|K1 st
      p in W & W is open & f.:W c= V by A11,A14;
    end;
  hence f is continuous by JGRAPH_2:20;
end;

theorem Th102: for sn being Real,
 B0 being Subset of TOP-REAL 2,K0 being Subset of
(TOP-REAL 2)|B0,f being map of ((TOP-REAL 2)|B0)|K0,((TOP-REAL 2)|B0)
st -1<sn & sn<1 & f=(sn-FanMorphE)|K0 &
B0=(the carrier of TOP-REAL 2) \ {0.REAL 2}
& K0={p: p`1>=0 & p<>0.REAL 2}
holds f is continuous
proof let sn be Real,
 B0 be Subset of TOP-REAL 2,K0 be Subset of
(TOP-REAL 2)|B0,f being map of ((TOP-REAL 2)|B0)|K0,((TOP-REAL 2)|B0);
assume A1:  -1<sn & sn<1 & f=(sn-FanMorphE)|K0 &
  B0=(the carrier of TOP-REAL 2) \ {0.REAL 2}
   & K0={p: p`1>=0 & p<>0.REAL 2 };
    the carrier of (TOP-REAL 2)|B0= B0 by JORDAN1:1;
  then K0 c= the carrier of TOP-REAL 2 by XBOOLE_1:1;
  then reconsider K1=K0 as Subset of TOP-REAL 2;
    K0 c= B0
   proof let x be set;assume x in K0;
      then consider p8 being Point of TOP-REAL 2 such that
      A2: x=p8 & (p8`1>=0 & p8<>0.REAL 2) by A1;
       not x in {0.REAL 2} by A2,TARSKI:def 1;
    hence x in B0 by A1,A2,XBOOLE_0:def 4;
   end;
  then ((TOP-REAL 2)|B0)|K0=(TOP-REAL 2)|K1 by JORDAN6:47;
hence f is continuous by A1,Th100;
end;

theorem Th103: for sn being Real,
 B0 being Subset of TOP-REAL 2,K0 being Subset of
(TOP-REAL 2)|B0,f being map of ((TOP-REAL 2)|B0)|K0,((TOP-REAL 2)|B0)
st -1<sn & sn<1 & f=(sn-FanMorphE)|K0 &
B0=(the carrier of TOP-REAL 2) \ {0.REAL 2}
& K0={p: p`1<=0 & p<>0.REAL 2}
holds f is continuous
proof let sn be Real,
 B0 be Subset of TOP-REAL 2,K0 be Subset of
(TOP-REAL 2)|B0,f be map of ((TOP-REAL 2)|B0)|K0,((TOP-REAL 2)|B0);
 assume A1:  -1<sn & sn<1 & f=(sn-FanMorphE)|K0 &
  B0=(the carrier of TOP-REAL 2) \ {0.REAL 2}
  & K0={p: p`1<=0 & p<>0.REAL 2};
    the carrier of (TOP-REAL 2)|B0= B0 by JORDAN1:1;
  then K0 c= the carrier of TOP-REAL 2 by XBOOLE_1:1;
  then reconsider K1=K0 as Subset of TOP-REAL 2;
    K0 c= B0
   proof let x be set;assume x in K0;
      then consider p8 being Point of TOP-REAL 2 such that
      A2: x=p8 & (p8`1<=0 & p8<>0.REAL 2) by A1;
       not x in {0.REAL 2} by A2,TARSKI:def 1;
    hence x in B0 by A1,A2,XBOOLE_0:def 4;
   end;
  then ((TOP-REAL 2)|B0)|K0=(TOP-REAL 2)|K1 by JORDAN6:47;
hence f is continuous by A1,Th101;
end;

theorem Th104: for sn being Real,p being Point of TOP-REAL 2
holds |.(sn-FanMorphE).p.|=|.p.|
proof let sn be Real,p be Point of TOP-REAL 2;
       set f=(sn-FanMorphE);
       set z=f.p;
        reconsider q=p as Point of TOP-REAL 2;
        reconsider qz=z as Point of TOP-REAL 2;
        A1: |.q.|>=0 by TOPRNS_1:26;
          now per cases;
        case A2: q`2/|.q.|>=sn & q`1>0;
          then (sn-FanMorphE).q= |[
          |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1-sn))]| by Th89;
          then A3: qz`1= |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2))&
                  qz`2= |.q.|* ((q`2/|.q.|-sn)/(1-sn))
                                by EUCLID:56;
            |.q.|<>0 by A2,JGRAPH_2:11,TOPRNS_1:25;
          then A4: (|.q.|)^2>0 by SQUARE_1:74;
          A5: (q`2/|.q.|-sn)>=0 by A2,SQUARE_1:12;
          A6: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
            0<=(q`1)^2 by SQUARE_1:72;
          then 0+(q`2)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
          then (q`2)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A4,A6,REAL_1:73;
          then (q`2)^2/(|.q.|)^2 <= 1 by A4,XCMPLX_1:60;
          then ((q`2)/|.q.|)^2 <= 1 by SQUARE_1:69;
          then 1>=q`2/|.q.| by Th4;
          then A7: 1-sn>=q`2/|.q.|-sn by REAL_1:49;
            now per cases;
          case A8: 1-sn=0;
            A9:((q`2/|.q.|-sn)/(1-sn))=(q`2/|.q.|-sn)*(1-sn)" by XCMPLX_0:def 9
             .= (q`2/|.q.|-sn)*0 by A8,XCMPLX_0:def 7 .=0;
            then 1-((q`2/|.q.|-sn)/(1-sn))^2=1 by SQUARE_1:60;
             then (sn-FanMorphE).q= |[|.q.|*(1),|.q.|*0]| by A2,A9,Th89,
SQUARE_1:83
            .=|[(|.q.|),0]|;
             then ((sn-FanMorphE).q)`1=(|.q.|) & ((sn-FanMorphE).q)`2=0
                              by EUCLID:56;
            then |.(sn-FanMorphE).p.|=sqrt(((|.q.|))^2+0) by JGRAPH_3:10,
SQUARE_1:60
            .=|.q.| by A1,SQUARE_1:89;
           hence |.(sn-FanMorphE).p.|=|.p.|;
          case A10: 1-sn<>0;
              now per cases by A10;
            case A11: 1-sn>0;
              then A12: -(1-sn)<= -0 by REAL_1:50;
                -(1-sn)<= -( q`2/|.q.|-sn) by A7,REAL_1:50;
              then (-(1-sn))/(1-sn)<=(-( q`2/|.q.|-sn))/(1-sn)
                        by A11,REAL_1:73;
              then A13: -1<=(-( q`2/|.q.|-sn))/(1-sn)
                        by A11,XCMPLX_1:198;
               --(1-sn)>= -(q`2/|.q.|-sn) by A5,A12,REAL_1:50;
              then (-(q`2/|.q.|-sn))/(1-sn)<=1 by A11,REAL_2:117;
              then ((-(q`2/|.q.|-sn))/(1-sn))^2<=1^2 by A13,JGRAPH_2:7;
              then 1-((-(q`2/|.q.|-sn))/(1-sn))^2>=0 by SQUARE_1:12,59;
              then 1-(-((q`2/|.q.|-sn))/(1-sn))^2>=0 by XCMPLX_1:188;
              then A14: 1-((q`2/|.q.|-sn)/(1-sn))^2>=0 by SQUARE_1:61;
              A15: (qz`1)^2= (|.q.|)^2*(sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2))^2
                         by A3,SQUARE_1:68
                  .= (|.q.|)^2*(1-((q`2/|.q.|-sn)/(1-sn))^2)
                            by A14,SQUARE_1:def 4;
              A16: (qz`2)^2= (|.q.|)^2*((q`2/|.q.|-sn)/(1-sn))^2
                         by A3,SQUARE_1:68;
                (|.qz.|)^2=(qz`1)^2+(qz`2)^2 by JGRAPH_3:10
               .=(|.q.|)^2*(1-((q`2/|.q.|-sn)/(1-sn))^2
                    +((q`2/|.q.|-sn)/(1-sn))^2) by A15,A16,XCMPLX_1:8
               .=(|.q.|)^2*1 by XCMPLX_1:27 .=(|.q.|)^2;
              then A17: sqrt((|.qz.|)^2)=|.q.| by A1,SQUARE_1:89;
                |.qz.|>=0 by TOPRNS_1:26;
             hence |.(sn-FanMorphE).p.|=|.p.| by A17,SQUARE_1:89;
            case A18: 1-sn<0;
              A19: q`2/|.q.|-sn>=0 by A2,SQUARE_1:12;
                0<(q`1)^2 by A2,SQUARE_1:74;
              then 0+(q`2)^2<(q`1)^2+(q`2)^2 by REAL_1:67;
              then (q`2)^2/(|.q.|)^2 < (|.q.|)^2/(|.q.|)^2 by A4,A6,REAL_1:73;
              then (q`2)^2/(|.q.|)^2 < 1 by A4,XCMPLX_1:60;
              then ((q`2)/|.q.|)^2 < 1 by SQUARE_1:69;
              then 1 > q`2/|.p.| by Th5;
              then q`2/|.q.|-sn<1-sn by REAL_1:54;
             hence |.(sn-FanMorphE).p.|=|.p.| by A18,A19,AXIOMS:22;
            end;
           hence |.(sn-FanMorphE).p.|=|.p.|;
          end;
         hence |.(sn-FanMorphE).p.|=|.p.|;
        case A20: q`2/|.q.|<sn & q`1>0;
          then A21:(sn-FanMorphE).q= |[
          |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1+sn))]| by Th90;
          then A22: qz`1= |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2))&
                  qz`2= |.q.|* ((q`2/|.q.|-sn)/(1+sn))
                                by EUCLID:56;
            |.q.|<>0 by A20,JGRAPH_2:11,TOPRNS_1:25;
          then A23: (|.q.|)^2>0 by SQUARE_1:74;
          A24: (q`2/|.q.|-sn)<0 by A20,REAL_2:105;
          A25: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
            0<=(q`1)^2 by SQUARE_1:72;
          then 0+(q`2)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
          then (q`2)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A23,A25,REAL_1:73;
          then (q`2)^2/(|.q.|)^2 <= 1 by A23,XCMPLX_1:60;
          then ((q`2)/|.q.|)^2 <= 1 by SQUARE_1:69;
          then -1<=q`2/|.q.| by Th4;
          then A26: -1-sn<=q`2/|.q.|-sn by REAL_1:49;
            now per cases;
          case A27: 1+sn=0;
             ((q`2/|.q.|-sn)/(1+sn))=(q`2/|.q.|-sn)*(1+sn)" by XCMPLX_0:def 9
             .= (q`2/|.q.|-sn)*0 by A27,XCMPLX_0:def 7 .=0;
             then ((sn-FanMorphE).q)`1=(|.q.|) & ((sn-FanMorphE).q)`2=0
                              by A21,EUCLID:56,SQUARE_1:60,83;
            then |.(sn-FanMorphE).p.|=sqrt(((|.q.|))^2+0) by JGRAPH_3:10,
SQUARE_1:60
            .=|.q.| by A1,SQUARE_1:89;
           hence |.(sn-FanMorphE).p.|=|.p.|;
          case A28: 1+sn<>0;
              now per cases by A28;
            case A29: 1+sn>0;
                -(1+sn)<= ( q`2/|.q.|-sn) by A26,XCMPLX_1:161;
              then (-(1+sn))/(1+sn)<=(( q`2/|.q.|-sn))/(1+sn)
                        by A29,REAL_1:73;
              then A30: -1<=(( q`2/|.q.|-sn))/(1+sn)
                        by A29,XCMPLX_1:198;
                (1+sn)>= (q`2/|.q.|-sn) by A24,A29,AXIOMS:22;
              then ((q`2/|.q.|-sn))/(1+sn)<=1 by A29,REAL_2:117;
              then (((q`2/|.q.|-sn))/(1+sn))^2<=1 by A30,JGRAPH_2:7,SQUARE_1:59
;
              then A31: 1-(((q`2/|.q.|-sn))/(1+sn))^2>=0 by SQUARE_1:12;
              A32: (qz`1)^2= (|.q.|)^2*(sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2))^2
                         by A22,SQUARE_1:68
                  .= (|.q.|)^2*(1-((q`2/|.q.|-sn)/(1+sn))^2)
                            by A31,SQUARE_1:def 4;
              A33: (qz`2)^2= (|.q.|)^2*((q`2/|.q.|-sn)/(1+sn))^2
                         by A22,SQUARE_1:68;
                (|.qz.|)^2=(qz`1)^2+(qz`2)^2 by JGRAPH_3:10
               .=(|.q.|)^2*(1-((q`2/|.q.|-sn)/(1+sn))^2
                    +((q`2/|.q.|-sn)/(1+sn))^2) by A32,A33,XCMPLX_1:8
               .=(|.q.|)^2*1 by XCMPLX_1:27 .=(|.q.|)^2;
              then A34: sqrt((|.qz.|)^2)=|.q.| by A1,SQUARE_1:89;
                |.qz.|>=0 by TOPRNS_1:26;
             hence |.(sn-FanMorphE).p.|=|.p.| by A34,SQUARE_1:89;
            case 1+sn<0;
              then A35: -(1+sn)>-0 by REAL_1:50;
                0<(q`1)^2 by A20,SQUARE_1:74;
              then 0+(q`2)^2<(q`1)^2+(q`2)^2 by REAL_1:67;
              then (q`2)^2/(|.q.|)^2 < (|.q.|)^2/(|.q.|)^2 by A23,A25,REAL_1:73
;
              then (q`2)^2/(|.q.|)^2 < 1 by A23,XCMPLX_1:60;
              then ((q`2)/|.q.|)^2 < 1 by SQUARE_1:69;
              then -1 < q`2/|.p.| by Th5;
              then q`2/|.q.|-sn>-1-sn by REAL_1:54;
              then q`2/|.q.|-sn>-(1+sn) by XCMPLX_1:161;
             hence |.(sn-FanMorphE).p.|=|.p.| by A24,A35,AXIOMS:22;
            end;
           hence |.(sn-FanMorphE).p.|=|.p.|;
          end;
         hence |.(sn-FanMorphE).p.|=|.p.|;
        case q`1<=0;
         hence |.(sn-FanMorphE).p.|=|.p.| by Th89;
        end;
 hence |.(sn-FanMorphE).p.|=|.p.|;
end;

theorem Th105: for sn being Real,x,K0 being set
 st -1<sn & sn<1 & x in K0 &
 K0={p: p`1>=0 & p<>0.REAL 2}
holds (sn-FanMorphE).x in K0
proof let sn be Real,x,K0 be set;
 assume A1:  -1<sn & sn<1 & x in K0 &
   K0={p: p`1>=0 & p<>0.REAL 2};
  then consider p such that
  A2: p=x & p`1>=0 & p<>0.REAL 2;
  A3:now assume |.p.|<=0;
    then |.p.|=0 by TOPRNS_1:26;
   hence contradiction by A2,TOPRNS_1:25;
  end;
  then A4: (|.p.|)^2>0 by SQUARE_1:74;
    now per cases;
  case A5: p`2/|.p.|<=sn;
   then A6: (sn-FanMorphE).p= |[ |.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)),
           |.p.|*((p`2/|.p.|-sn)/(1+sn))]| by A1,A2,Th91;
   reconsider p9= (sn-FanMorphE).p as Point of TOP-REAL 2;
   A7: p9`1=|.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)) by A6,EUCLID:56;
   A8: 1+sn>0 by A1,Th3;
       A9: (|.p.|)^2 =(p`1)^2+(p`2)^2 by JGRAPH_3:10;
         now per cases;
       case p`1=0;
        hence (sn-FanMorphE).x in K0 by A1,A2,Th89;
       case p`1<>0;
         then 0<(p`1)^2 by SQUARE_1:74;
         then 0+(p`2)^2<(p`1)^2+(p`2)^2 by REAL_1:67;
         then (p`2)^2/(|.p.|)^2 < (|.p.|)^2/(|.p.|)^2 by A4,A9,REAL_1:73;
         then (p`2)^2/(|.p.|)^2 < 1 by A4,XCMPLX_1:60;
         then A10: ((p`2)/|.p.|)^2 < 1 by SQUARE_1:69;
           p`2/|.p.|-sn<=0 by A5,REAL_2:106;
         then A11: (p`2/|.p.|-sn)/(1+sn)<(1+sn)/(1+sn) by A8,REAL_1:73;
           -1 < p`2/|.p.| by A10,Th5;
         then -1-sn< p`2/|.p.|-sn by REAL_1:54;
         then -(1+sn)<p`2/|.p.|-sn by XCMPLX_1:161;
         then (-1)*(1+sn)< p`2/|.p.|-sn by XCMPLX_1:180;
         then (-1)*(1+sn)/(1+sn)< (p`2/|.p.|-sn)/(1+sn)
                       by A8,REAL_1:73;
         then -1< (p`2/|.p.|-sn)/(1+sn) & (p`2/|.p.|-sn)/(1+sn)<1
                           by A8,A11,XCMPLX_1:60,90;
         then 1^2> ((p`2/|.p.|-sn)/(1+sn))^2 by JGRAPH_2:8;
         then 1-((p`2/|.p.|-sn)/(1+sn))^2>0 by SQUARE_1:11,59;
         then sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)>0 by SQUARE_1:93;
         then |.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2))>0
                           by A3,SQUARE_1:21;
        hence (sn-FanMorphE).x in K0 by A1,A2,A7,JGRAPH_2:11;
        end;
       hence (sn-FanMorphE).x in K0;
  case A12: p`2/|.p.|>sn;
   then A13: (sn-FanMorphE).p= |[ |.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)),
           |.p.|*((p`2/|.p.|-sn)/(1-sn))]| by A1,A2,Th91;
   reconsider p9= (sn-FanMorphE).p as Point of TOP-REAL 2;
   A14: p9`1=|.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)) by A13,EUCLID:56;
    A15: 1-sn>0 by A1,Th3;
       A16: (|.p.|)^2 =(p`1)^2+(p`2)^2 by JGRAPH_3:10;
         now per cases;
       case p`1=0;
        hence (sn-FanMorphE).x in K0 by A1,A2,Th89;
       case p`1<>0;
         then 0<(p`1)^2 by SQUARE_1:74;
         then 0+(p`2)^2<(p`1)^2+(p`2)^2 by REAL_1:67;
         then (p`2)^2/(|.p.|)^2 < (|.p.|)^2/(|.p.|)^2 by A4,A16,REAL_1:73;
         then (p`2)^2/(|.p.|)^2 < 1 by A4,XCMPLX_1:60;
         then ((p`2)/|.p.|)^2 < 1 by SQUARE_1:69;
         then p`2/|.p.|<1 by Th5;
         then (p`2/|.p.|-sn)<1-sn by REAL_1:54;
         then A17: (p`2/|.p.|-sn)/(1-sn)<(1-sn)/(1-sn) by A15,REAL_1:73;
           -(1-sn)< -0 by A15,REAL_1:50;
         then A18: -(1-sn)<sn-sn by XCMPLX_1:14;
           p`2/|.p.|-sn>=sn-sn by A12,REAL_1:49;
         then -(1-sn)<p`2/|.p.|-sn by A18,AXIOMS:22;
         then (-1)*(1-sn)< p`2/|.p.|-sn by XCMPLX_1:180;
         then (-1)*(1-sn)/(1-sn)< (p`2/|.p.|-sn)/(1-sn)
                       by A15,REAL_1:73;
         then -1< (p`2/|.p.|-sn)/(1-sn) & (p`2/|.p.|-sn)/(1-sn)<1
                           by A15,A17,XCMPLX_1:60,90;
         then 1> ((p`2/|.p.|-sn)/(1-sn))^2 by JGRAPH_2:8,SQUARE_1:59;
         then 1-((p`2/|.p.|-sn)/(1-sn))^2>0 by SQUARE_1:11;
         then sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)>0 by SQUARE_1:93;
         then p9`1>0 by A3,A14,SQUARE_1:21;
        hence (sn-FanMorphE).x in K0 by A1,A2,JGRAPH_2:11;
       end;
        hence (sn-FanMorphE).x in K0;
        end;
 hence (sn-FanMorphE).x in K0;
end;

theorem Th106: for sn being Real,x,K0 being set
 st -1<sn & sn<1 & x in K0 &
 K0={p: p`1<=0 & p<>0.REAL 2}
holds (sn-FanMorphE).x in K0
proof let sn be Real,x,K0 be set;
 assume A1:  -1<sn & sn<1 & x in K0 &
 K0={p: p`1<=0 & p<>0.REAL 2};
 then consider p such that
 A2: p=x & p`1<=0 & p<>0.REAL 2;
 thus (sn-FanMorphE).x in K0 by A1,A2,Th89;
end;

theorem Th107:for sn being Real,
 D being non empty Subset of TOP-REAL 2
st -1<sn & sn<1 & D`={0.REAL 2} holds
ex h being map of (TOP-REAL 2)|D,(TOP-REAL 2)|D
st h=(sn-FanMorphE)|D & h is continuous
proof let sn be Real,D be non empty Subset of TOP-REAL 2;
assume A1:  -1<sn & sn<1 & D`={0.REAL 2};
  reconsider B0= {0.REAL 2} as Subset of TOP-REAL 2;
  A2: D =(B0)` by A1
     .=((the carrier of TOP-REAL 2)\ {0.REAL 2}) by SUBSET_1:def 5;
     defpred P[Point of TOP-REAL 2] means $1`1>=0;
    A3:{p: P[p] & p<>0.REAL 2} c= the carrier of (TOP-REAL 2)|D
         from InclSub(A2);
        (|[0,1]|)`1=0 & (|[0,1]|)`2=1 by EUCLID:56;
      then |[0,1]| in {p where p is Point of TOP-REAL 2:
        p`1>=0 & p<>0.REAL 2} by JGRAPH_2:11;
      then reconsider K0={p:p`1>=0 & p<>0.REAL 2}
          as non empty Subset of (TOP-REAL 2)|D by A3;
          defpred P[Point of TOP-REAL 2] means $1`1<=0;
    A4:{p: P[p] & p<>0.REAL 2} c= the carrier of (TOP-REAL 2)|D
           from InclSub(A2);
     set Y1=|[0,1]|;
      Y1`1=0 & Y1`2=1 by EUCLID:56;
     then Y1 in {p where p is Point of TOP-REAL 2:
     p`1<=0 & p<>0.REAL 2} by JGRAPH_2:11;
  then reconsider K1={p: p`1<=0 & p<>0.REAL 2}
       as non empty Subset of (TOP-REAL 2)|D by A4;
   A5:the carrier of ((TOP-REAL 2)|D)=D by JORDAN1:1;
   A6:K0 c= (the carrier of TOP-REAL 2)
   proof let z be set;assume z in K0;
        then consider p8 being Point of TOP-REAL 2 such that
        A7: p8=z & p8`1>=0 & p8<>0.REAL 2;
    thus z in (the carrier of TOP-REAL 2) by A7;
   end;
  A8:dom ((sn-FanMorphE)|K0)= dom ((sn-FanMorphE)) /\ K0 by FUNCT_1:68
       .=((the carrier of TOP-REAL 2)) /\ K0 by FUNCT_2:def 1
       .=K0 by A6,XBOOLE_1:28;
  A9: K0=the carrier of ((TOP-REAL 2)|D)|K0 by JORDAN1:1;
    rng ((sn-FanMorphE)|K0) c= the carrier of ((TOP-REAL 2)|D)|K0
   proof let y be set;assume y in rng ((sn-FanMorphE)|K0);
     then consider x being set such that
     A10:x in dom ((sn-FanMorphE)|K0)
     & y=((sn-FanMorphE)|K0).x by FUNCT_1:def 5;
       x in (dom (sn-FanMorphE)) /\ K0 by A10,FUNCT_1:68;
     then A11:x in K0 by XBOOLE_0:def 3;
       K0 c= the carrier of TOP-REAL 2 by A5,XBOOLE_1:1;
     then reconsider p=x as Point of TOP-REAL 2 by A11;
      (sn-FanMorphE).p=y by A10,A11,FUNCT_1:72;
     then y in K0 by A1,A11,Th105;
    hence y in the carrier of ((TOP-REAL 2)|D)|K0 by JORDAN1:1;
   end;
  then rng ((sn-FanMorphE)|K0)c= the carrier of ((TOP-REAL 2)|D)
                            by A9,XBOOLE_1:1;
  then (sn-FanMorphE)|K0 is Function of the carrier of ((TOP-REAL 2)|D)|K0,
  the carrier of ((TOP-REAL 2)|D) by A8,A9,FUNCT_2:4;
  then reconsider f=(sn-FanMorphE)|K0 as map of ((TOP-REAL 2)|D)|K0,
  ((TOP-REAL 2)|D) ;
   A12:K1 c= (the carrier of TOP-REAL 2)
   proof let z be set;assume z in K1;
        then consider p8 being Point of TOP-REAL 2 such that
        A13: p8=z & p8`1<=0 & p8<>0.REAL 2;
    thus z in (the carrier of TOP-REAL 2) by A13;
   end;
  A14:dom ((sn-FanMorphE)|K1)= dom ((sn-FanMorphE)) /\ K1 by FUNCT_1:68
       .=((the carrier of TOP-REAL 2)) /\ K1 by FUNCT_2:def 1
       .=K1 by A12,XBOOLE_1:28;
  A15: K1=the carrier of ((TOP-REAL 2)|D)|K1 by JORDAN1:1;
    rng ((sn-FanMorphE)|K1) c= the carrier of ((TOP-REAL 2)|D)|K1
   proof let y be set;assume y in rng ((sn-FanMorphE)|K1);
     then consider x being set such that
     A16:x in dom ((sn-FanMorphE)|K1)
     & y=((sn-FanMorphE)|K1).x by FUNCT_1:def 5;
       x in (dom (sn-FanMorphE)) /\ K1 by A16,FUNCT_1:68;
     then A17:x in K1 by XBOOLE_0:def 3;
       K1 c= the carrier of TOP-REAL 2 by A5,XBOOLE_1:1;
     then reconsider p=x as Point of TOP-REAL 2 by A17;
      (sn-FanMorphE).p=y by A16,A17,FUNCT_1:72;
     then y in K1 by A1,A17,Th106;
    hence y in the carrier of ((TOP-REAL 2)|D)|K1 by JORDAN1:1;
   end;
  then rng ((sn-FanMorphE)|K1)c= the carrier of ((TOP-REAL 2)|D)
                     by A15,XBOOLE_1:1;
  then (sn-FanMorphE)|K1 is Function of the carrier of ((TOP-REAL 2)|D)|K1,
  the carrier of ((TOP-REAL 2)|D) by A14,A15,FUNCT_2:4;
  then reconsider g=(sn-FanMorphE)|K1 as map of ((TOP-REAL 2)|D)|K1,
  ((TOP-REAL 2)|D) ;
  A18:K0=[#](((TOP-REAL 2)|D)|K0) by PRE_TOPC:def 10;
  A19:K1=[#](((TOP-REAL 2)|D)|K1) by PRE_TOPC:def 10;
  A20:D= [#]((TOP-REAL 2)|D) by PRE_TOPC:def 10;
  A21:K0 \/ K1 c= D
   proof let x be set;assume A22:x in K0 \/ K1;
       now per cases by A22,XBOOLE_0:def 2;
     case x in K0;
       then consider p such that A23:p=x &
       p`1>=0 & p<>0.REAL 2;
     thus x in the carrier of TOP-REAL 2 & not x=0.REAL 2 by A23;
     case x in K1;
       then consider p such that A24:p=x & p`1<=0 & p<>0.REAL 2;
     thus x in the carrier of TOP-REAL 2 & not x=0.REAL 2 by A24;
     end;
     then x in the carrier of TOP-REAL 2 & not x in {0.REAL 2} by TARSKI:def 1
;
    hence x in D by A2,XBOOLE_0:def 4;
   end;
    D c= K0 \/ K1
   proof let x be set;assume A25: x in D;
     then A26:x in (the carrier of TOP-REAL 2) & not x in {0.REAL 2}
                                     by A2,XBOOLE_0:def 4;
     reconsider px=x as Point of TOP-REAL 2 by A25;
       px`1>=0 & px<>0.REAL 2 or
       px`1<=0 & px<>0.REAL 2 by A26,TARSKI:def 1;
     then x in K0 or x in K1;
    hence x in K0 \/ K1 by XBOOLE_0:def 2;
   end;
  then A27:([#](((TOP-REAL 2)|D)|K0)) \/ ([#](((TOP-REAL 2)|D)|K1))
  = [#]((TOP-REAL 2)|D) by A18,A19,A20,A21,XBOOLE_0:def 10;
  A28: f is continuous & K0 is closed by A1,A2,Th38,Th102;
  A29: g is continuous & K1 is closed by A1,A2,Th36,Th103;
  A30: for x be set st x in ([#]((((TOP-REAL 2)|D)|K0)))
        /\ ([#] ((((TOP-REAL 2)|D)|K1))) holds f.x = g.x
   proof let x be set;assume x in ([#]((((TOP-REAL 2)|D)|K0)))
        /\ ([#] ((((TOP-REAL 2)|D)|K1)));
     then A31:x in K0 & x in K1 by A18,A19,XBOOLE_0:def 3;
     then f.x=(sn-FanMorphE).x by FUNCT_1:72;
    hence f.x = g.x by A31,FUNCT_1:72;
   end;
  then consider h being map of (TOP-REAL 2)|D,(TOP-REAL 2)|D
  such that A32: h= f+*g & h is continuous
                          by A18,A19,A27,A28,A29,JGRAPH_2:9;
  A33:dom h=the carrier of ((TOP-REAL 2)|D) by FUNCT_2:def 1;
  A34:the carrier of ((TOP-REAL 2)|D)=D by JORDAN1:1;
  A35:the carrier of ((TOP-REAL 2)|D)=
((the carrier of TOP-REAL 2)\ {0.REAL 2}) by A2,JORDAN1:1;
    dom (sn-FanMorphE)=(the carrier of (TOP-REAL 2))
                    by FUNCT_2:def 1;
  then A36:dom ((sn-FanMorphE)|D)=(the carrier of (TOP-REAL 2))/\ D
                    by FUNCT_1:68
              .=the carrier of ((TOP-REAL 2)|D) by A34,XBOOLE_1:28;
  A37:dom f=K0 by A9,FUNCT_2:def 1;
  A38:K0=[#](((TOP-REAL 2)|D)|K0) by PRE_TOPC:def 10;
  A39:dom g=K1 by A15,FUNCT_2:def 1;
   K1=[#](((TOP-REAL 2)|D)|K1) by PRE_TOPC:def 10;
  then A40: f tolerates g by A30,A37,A38,A39,PARTFUN1:def 6;
    for x being set st x in dom h holds h.x=((sn-FanMorphE)|D).x
   proof let x be set;assume A41: x in dom h;
     then x in (the carrier of TOP-REAL 2) & not x in {0.REAL 2}
                     by A33,A35,XBOOLE_0:def 4;
     then A42:x <>0.REAL 2 by TARSKI:def 1;
     reconsider p=x as Point of TOP-REAL 2 by A33,A35,A41,XBOOLE_0:def 4;
     A43: x in D`` by A33,A41,JORDAN1:1;
       now per cases;
     case A44:x in K0;
       A45:(sn-FanMorphE)|D.p=(sn-FanMorphE).p by A43,FUNCT_1:72
          .=f.p by A44,FUNCT_1:72;
         h.p=(g+*f).p by A32,A40,FUNCT_4:35
       .=f.p by A37,A44,FUNCT_4:14;
      hence h.x=(sn-FanMorphE)|D.x by A45;
     case not x in K0;
       then not (p`1>=0) by A42;
       then A46:x in K1 by A42;
        (sn-FanMorphE)|D.p=(sn-FanMorphE).p by A43,FUNCT_1:72
                           .=g.p by A46,FUNCT_1:72;
      hence h.x=(sn-FanMorphE)|D.x by A32,A39,A46,FUNCT_4:14;
     end;
    hence h.x=(sn-FanMorphE)|D.x;
   end;
  then f+*g=(sn-FanMorphE)|D by A32,A33,A36,FUNCT_1:9;
 hence thesis by A18,A19,A27,A28,A29,A30,JGRAPH_2:9;
end;

theorem Th108: for sn being Real
  st -1<sn & sn<1 holds
  ex h being map of (TOP-REAL 2),(TOP-REAL 2)
  st h=(sn-FanMorphE) & h is continuous
proof let sn be Real;
 assume A1:  -1<sn & sn<1;
 reconsider f=(sn-FanMorphE) as map of (TOP-REAL 2),(TOP-REAL 2)
                            ;
  (the carrier of TOP-REAL 2) \{0.REAL 2} c=(the carrier of TOP-REAL 2)
                         by XBOOLE_1:36;
 then reconsider D=(the carrier of TOP-REAL 2) \{0.REAL 2}
   as non empty Subset of TOP-REAL 2 by JGRAPH_2:19;
 A2: f.(0.REAL 2)=0.REAL 2 by Th89,JGRAPH_2:11;
 A3: D`= {0.REAL 2} by JGRAPH_3:30;
 A4: for p being Point of (TOP-REAL 2)|D holds f.p<>f.(0.REAL 2)
  proof let p be Point of (TOP-REAL 2)|D;
    A5:the carrier of (TOP-REAL 2)|D=[#]((TOP-REAL 2)|D) by PRE_TOPC:12;
    A6:[#]((TOP-REAL 2)|D)=D by PRE_TOPC:def 10;
    then A7:p in the carrier of TOP-REAL 2 & not p in {0.REAL 2}
                         by A5,XBOOLE_0:def 4;
    reconsider q=p as Point of TOP-REAL 2 by A5,A6,XBOOLE_0:def 4;
    A8:not p=0.REAL 2 by A7,TARSKI:def 1;
      now per cases;
    case A9: q`2/|.q.|>=sn & q`1>=0;
      set q9= |[ |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1-sn))]|;
      A10:q9`1= |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)) &
      q9`2= |.q.|* ((q`2/|.q.|-sn)/(1-sn))
                                  by EUCLID:56;
        now assume A11: q9=0.REAL 2;
       A12: |.q.|<>0 by A8,TOPRNS_1:25;
       then sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)=sqrt(1-0) by A10,A11,JGRAPH_2:11,
SQUARE_1:60,XCMPLX_1:6
       .=1 by SQUARE_1:83;
      hence contradiction by A11,A12,EUCLID:56,JGRAPH_2:11;
     end;
     hence f.p<>f.(0.REAL 2) by A1,A2,A8,A9,Th91;
    case A13: q`2/|.q.|<sn & q`1>=0;
      set q9=|[ |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)),
              |.q.|* ((q`2/|.q.|-sn)/(1+sn))]|;
      A14:q9`1= |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)) &
          q9`2= |.q.|* ((q`2/|.q.|-sn)/(1+sn)) by EUCLID:56;
        now assume A15: q9=0.REAL 2;
       A16: |.q.|<>0 by A8,TOPRNS_1:25;
       then sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)=sqrt(1-0) by A14,A15,JGRAPH_2:11,
SQUARE_1:60,XCMPLX_1:6
       .=1 by SQUARE_1:83;
      hence contradiction by A15,A16,EUCLID:56,JGRAPH_2:11;
     end;
     hence f.p<>f.(0.REAL 2) by A1,A2,A8,A13,Th91;
     case q`1<0; then f.p=p by Th89;
     hence f.p<>f.(0.REAL 2) by A8,Th89,JGRAPH_2:11;
    end;
   hence f.p<>f.(0.REAL 2);
  end;
 A17:ex h being map of (TOP-REAL 2)|D,(TOP-REAL 2)|D
  st h=(sn-FanMorphE)|D & h is continuous by A1,A3,Th107;
   for V being Subset of (TOP-REAL 2) st f.(0.REAL 2) in V & V is open
 ex W being Subset of (TOP-REAL 2) st 0.REAL 2 in W
 & W is open & f.:W c= V
  proof let V be Subset of (TOP-REAL 2);
    assume A18:f.(0.REAL 2) in V & V is open;
     A19:TOP-REAL 2=TopSpaceMetr(Euclid 2) by EUCLID:def 8;
     A20:the carrier of TOP-REAL 2=REAL 2 by EUCLID:25;
     A21:Euclid 2=MetrStruct(#REAL 2,Pitag_dist 2#) by EUCLID:def 7;
     then reconsider u0=0.REAL 2 as Point of Euclid 2 by EUCLID:25;
     consider r being real number such that A22:r>0 & Ball(u0,r) c= V
                                   by A2,A18,A19,TOPMETR:22;
     reconsider r as Real by XREAL_0:def 1;
     reconsider W1=Ball(u0,r) as Subset of TOP-REAL 2
                       by A20,A21;
     A23:u0 in W1 by A22,GOBOARD6:4;
     A24:W1 is open by GOBOARD6:6;
       f.:W1 c= W1
      proof let z be set;assume z in f.:W1;
        then consider y being set such that
        A25: y in dom f & y in W1 & z=f.y by FUNCT_1:def 12;
        reconsider q=y as Point of TOP-REAL 2 by A25;
        reconsider qy=q as Point of Euclid 2 by A21,EUCLID:25;
          z in rng f by A25,FUNCT_1:def 5;
        then reconsider qz=z as Point of TOP-REAL 2;
        reconsider pz=qz as Point of Euclid 2 by A21,EUCLID:25;
          dist(u0,qy)<r by A25,METRIC_1:12;
        then A26: |.(0.REAL 2) - q.|<r by JGRAPH_1:45;
        A27: |.q.|>=0 by TOPRNS_1:26;
          now per cases by JGRAPH_2:11;
        case q`1<=0;
         hence z in W1 by A25,Th89;
        case A28: q<>0.REAL 2 &
          (q`2/|.q.|>=sn & q`1>=0);
          then A29:(sn-FanMorphE).q= |[
          |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1-sn))]| by A1,Th91;
            |.q.|<>0 by A28,TOPRNS_1:25;
          then A30: (|.q.|)^2>0 by SQUARE_1:74;
        A31: 1-sn>0 by A1,Th3;
        A32: qz`1= |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2))&
        qz`2= |.q.|* ((q`2/|.q.|-sn)/(1-sn)) by A25,A29,EUCLID:56;
        A33: (q`2/|.q.|-sn)>= 0 by A28,SQUARE_1:12;
        A34: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
          0<=(q`1)^2 by SQUARE_1:72;
        then 0+(q`2)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
        then (q`2)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A30,A34,REAL_1:73;
        then (q`2)^2/(|.q.|)^2 <= 1 by A30,XCMPLX_1:60;
        then ((q`2)/|.q.|)^2 <= 1 by SQUARE_1:69;
        then 1>=q`2/|.q.| by Th4;
        then A35: 1-sn>=q`2/|.q.|-sn by REAL_1:49;
        A36: -(1-sn)<= -0 by A31,REAL_1:50;
           -(1-sn)<= -( q`2/|.q.|-sn) by A35,REAL_1:50;
         then (-(1-sn))/(1-sn)<=(-( q`2/|.q.|-sn))/(1-sn) by A31,REAL_1:73;
        then A37: -1<=(-( q`2/|.q.|-sn))/(1-sn) by A31,XCMPLX_1:198;
         --(1-sn)>= -(q`2/|.q.|-sn) by A33,A36,REAL_1:50;
       then (-(q`2/|.q.|-sn))/(1-sn)<=1 by A31,REAL_2:117;
       then ((-(q`2/|.q.|-sn))/(1-sn))^2<=1^2 by A37,JGRAPH_2:7;
       then 1-((-(q`2/|.q.|-sn))/(1-sn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((q`2/|.q.|-sn))/(1-sn))^2>=0 by XCMPLX_1:188;
       then A38: 1-((q`2/|.q.|-sn)/(1-sn))^2>=0 by SQUARE_1:61;
        A39: (qz`1)^2= (|.q.|)^2*(sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2))^2
                         by A32,SQUARE_1:68
         .= (|.q.|)^2*(1-((q`2/|.q.|-sn)/(1-sn))^2) by A38,SQUARE_1:def 4;
        A40: (qz`2)^2= (|.q.|)^2*((q`2/|.q.|-sn)/(1-sn))^2
                         by A32,SQUARE_1:68;
          (|.qz.|)^2=(qz`1)^2+(qz`2)^2 by JGRAPH_3:10
        .=(|.q.|)^2*(1-((q`2/|.q.|-sn)/(1-sn))^2
           +((q`2/|.q.|-sn)/(1-sn))^2) by A39,A40,XCMPLX_1:8
        .=(|.q.|)^2*1 by XCMPLX_1:27 .=(|.q.|)^2;
        then A41: sqrt((|.qz.|)^2)=|.q.| by A27,SQUARE_1:89;
            |.qz.|>=0 by TOPRNS_1:26;
          then A42: |.qz.|=|.q.| by A41,SQUARE_1:89;
            |.(0.REAL 2) +- q.|<r by A26,EUCLID:45;
          then |.- q.|<r by EUCLID:31;
          then |.q.|<r by TOPRNS_1:27;
          then |.- qz.|<r by A42,TOPRNS_1:27;
          then |.(0.REAL 2) +- qz.|<r by EUCLID:31;
          then |.(0.REAL 2) - qz.|<r by EUCLID:45;
          then dist(u0,pz)<r by JGRAPH_1:45;
         hence z in W1 by METRIC_1:12;
        case A43: q<>0.REAL 2 & q`2/|.q.|<sn & q`1>=0;
          then A44:(sn-FanMorphE).q= |[|.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)
),
          |.q.|* ((q`2/|.q.|-sn)/(1+sn))]| by A1,Th91;
            |.q.|<>0 by A43,TOPRNS_1:25;
          then A45: (|.q.|)^2>0 by SQUARE_1:74;
        A46: (sn-q`2/|.q.|)>=0 by A43,SQUARE_1:12;
        A47: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
        A48: 1+sn>0 by A1,Th3;
        A49: qz`1= |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2))&
        qz`2= |.q.|* ((q`2/|.q.|-sn)/(1+sn))
                     by A25,A44,EUCLID:56;
          0<=(q`1)^2 by SQUARE_1:72;
        then 0+(q`2)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
        then (q`2)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A45,A47,REAL_1:73;
        then (q`2)^2/(|.q.|)^2 <= 1 by A45,XCMPLX_1:60;
        then ((q`2)/|.q.|)^2 <= 1 by SQUARE_1:69;
        then -1<=q`2/|.q.| by Th4;
        then --1>=-q`2/|.q.| by REAL_1:50;
        then A50: 1+sn>=-q`2/|.q.|+sn by REAL_1:55;
          -(1+sn)<= -0 by A48,REAL_1:50;
         then -(1+sn)<= -( q`2/|.q.|-sn) by A46,XCMPLX_1:143;
         then (-(1+sn))/(1+sn)<=(-( q`2/|.q.|-sn))/(1+sn) by A48,REAL_1:73;
        then A51: -1<=(-( q`2/|.q.|-sn))/(1+sn) by A48,XCMPLX_1:198;
          -(1+sn)<=-(-q`2/|.q.|+sn) by A50,REAL_1:50;
        then -(1+sn)<=--(q`2/|.q.|)-sn by XCMPLX_1:161;
        then --(1+sn)>= -(q`2/|.q.|-sn) by REAL_1:50;
       then (-(q`2/|.q.|-sn))/(1+sn)<=1 by A48,REAL_2:117;
       then ((-(q`2/|.q.|-sn))/(1+sn))^2<=1^2 by A51,JGRAPH_2:7;
       then 1-((-(q`2/|.q.|-sn))/(1+sn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((q`2/|.q.|-sn))/(1+sn))^2>=0 by XCMPLX_1:188;
       then A52: 1-((q`2/|.q.|-sn)/(1+sn))^2>=0 by SQUARE_1:61;
        A53: (qz`1)^2= (|.q.|)^2*(sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2))^2
                         by A49,SQUARE_1:68
         .= (|.q.|)^2*(1-((q`2/|.q.|-sn)/(1+sn))^2) by A52,SQUARE_1:def 4;
        A54: (qz`2)^2= (|.q.|)^2*((q`2/|.q.|-sn)/(1+sn))^2
                         by A49,SQUARE_1:68;
          (|.qz.|)^2=(qz`1)^2+(qz`2)^2 by JGRAPH_3:10
        .=(|.q.|)^2*(1-((q`2/|.q.|-sn)/(1+sn))^2
           +((q`2/|.q.|-sn)/(1+sn))^2) by A53,A54,XCMPLX_1:8
        .=(|.q.|)^2*1 by XCMPLX_1:27 .=(|.q.|)^2;
        then A55: sqrt((|.qz.|)^2)=|.q.| by A27,SQUARE_1:89;
            |.qz.|>=0 by TOPRNS_1:26;
          then A56: |.qz.|=|.q.| by A55,SQUARE_1:89;
            |.(0.REAL 2) +- q.|<r by A26,EUCLID:45;
          then |.- q.|<r by EUCLID:31;
          then |.q.|<r by TOPRNS_1:27;
          then |.- qz.|<r by A56,TOPRNS_1:27;
          then |.(0.REAL 2) +- qz.|<r by EUCLID:31;
          then |.(0.REAL 2) - qz.|<r by EUCLID:45;
          then dist(u0,pz)<r by JGRAPH_1:45;
         hence z in W1 by METRIC_1:12;
        end;
       hence z in W1;
      end;
     then f.:W1 c= V by A22,XBOOLE_1:1;
   hence ex W being Subset of (TOP-REAL 2) st 0.REAL 2 in W
       & W is open & f.:W c= V by A23,A24;
  end;
  then f is continuous by A2,A3,A4,A17,JGRAPH_3:13;
 hence thesis;
end;

theorem Th109: for sn being Real
  st -1<sn & sn<1 holds
  (sn-FanMorphE) is one-to-one
proof let sn be Real;
 assume A1:  -1<sn & sn<1;
    for x1,x2 being set st x1 in dom (sn-FanMorphE) & x2 in dom (sn-FanMorphE)
   & (sn-FanMorphE).x1=(sn-FanMorphE).x2 holds x1=x2
  proof let x1,x2 be set;
   assume A2: x1 in dom (sn-FanMorphE) & x2 in dom (sn-FanMorphE)
    & (sn-FanMorphE).x1=(sn-FanMorphE).x2;
    then reconsider p1=x1 as Point of TOP-REAL 2 by FUNCT_2:def 1;
    reconsider p2=x2 as Point of TOP-REAL 2 by A2,FUNCT_2:def 1;
    set q=p1,p=p2;
    A3: 1-sn>0 by A1,Th3;
      now per cases by JGRAPH_2:11;
    case A4: q`1<=0;
      then A5:(sn-FanMorphE).q=q by Th89;
        now per cases by JGRAPH_2:11;
      case p`1<=0;
       hence x1=x2 by A2,A5,Th89;
      case A6:p<>0.REAL 2 &
          (p`2/|.p.|>=sn & p`1>=0);
        then A7:
         sn-FanMorphE.p= |[ |.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)),
          |.p.|* ((p`2/|.p.|-sn)/(1-sn))]| by A1,Th91;
         set p4= |[ |.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)),
          |.p.|* ((p`2/|.p.|-sn)/(1-sn))]|;
        A8: p4`1= |.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)) &
          p4`2= |.p.|* ((p`2/|.p.|-sn)/(1-sn)) by EUCLID:56;
     A9: |.p.|>=0 by TOPRNS_1:26;
     A10: |.p.|<>0 by A6,TOPRNS_1:25;
     then A11: (|.p.|)^2>0 by SQUARE_1:74;
        A12: (p`2/|.p.|-sn)>=0 by A6,SQUARE_1:12;
        A13: (p`2/|.p.|-sn)>= 0 by A6,SQUARE_1:12;
        A14: (p`2/|.p.|-sn)/(1-sn)>=0 by A3,A12,REAL_2:125;
        A15: (|.p.|)^2 =(p`1)^2+(p`2)^2 by JGRAPH_3:10;
          0<=(p`1)^2 by SQUARE_1:72;
        then 0+(p`2)^2<=(p`1)^2+(p`2)^2 by REAL_1:55;
        then (p`2)^2/(|.p.|)^2 <= (|.p.|)^2/(|.p.|)^2 by A11,A15,REAL_1:73;
        then (p`2)^2/(|.p.|)^2 <= 1 by A11,XCMPLX_1:60;
        then ((p`2)/|.p.|)^2 <= 1 by SQUARE_1:69;
        then 1>=p`2/|.p.| by Th4;
        then A16: 1-sn>=p`2/|.p.|-sn by REAL_1:49;
        A17: -(1-sn)<= -0 by A3,REAL_1:50;
           -(1-sn)<= -( p`2/|.p.|-sn) by A16,REAL_1:50;
         then (-(1-sn))/(1-sn)<=(-( p`2/|.p.|-sn))/(1-sn) by A3,REAL_1:73;
        then A18: -1<=(-( p`2/|.p.|-sn))/(1-sn) by A3,XCMPLX_1:198;
         --(1-sn)>= -(p`2/|.p.|-sn) by A13,A17,REAL_1:50;
       then (-(p`2/|.p.|-sn))/(1-sn)<=1 by A3,REAL_2:117;
       then ((-(p`2/|.p.|-sn))/(1-sn))^2<=1^2 by A18,JGRAPH_2:7;
       then A19: 1-((-(p`2/|.p.|-sn))/(1-sn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((p`2/|.p.|-sn))/(1-sn))^2>=0 by XCMPLX_1:188;
       then A20: 1-((p`2/|.p.|-sn)/(1-sn))^2>=0 by SQUARE_1:61;
         sqrt(1-((-(p`2/|.p.|-sn))/(1-sn))^2)>=0 by A19,SQUARE_1:def 4;
       then sqrt(1-(-(p`2/|.p.|-sn))^2/(1-sn)^2)>=0 by SQUARE_1:69;
       then sqrt(1-(p`2/|.p.|-sn)^2/(1-sn)^2)>=0 by SQUARE_1:61;
       then sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)>=0 by SQUARE_1:69;
        then q`1=0 by A2,A4,A5,A7,A8,A9,SQUARE_1:19;
        then (sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2))=0 by A2,A5,A7,A8,A10,XCMPLX_1:
6;
        then 1-((p`2/|.p.|-sn)/(1-sn))^2=0 by A20,SQUARE_1:92;
        then 1=0+((p`2/|.p.|-sn)/(1-sn))^2 by XCMPLX_1:27;
        then 1= (p`2/|.p.|-sn)/(1-sn) by A14,SQUARE_1:83,89;
        then (1)*(1-sn)=(p`2/|.p.|-sn) by A3,XCMPLX_1:88;
        then 1-sn+sn=p`2/|.p.| by XCMPLX_1:27;
        then 1=p`2/|.p.| by XCMPLX_1:27;
        then (1)*|.p.|=p`2 by A10,XCMPLX_1:88;
        then (p`2)^2-(p`2)^2 =(p`1)^2 by A15,XCMPLX_1:26;
        then 0= (p`1)^2 by XCMPLX_1:14;
        then p`1=0 by SQUARE_1:73;
       hence x1=x2 by A2,A5,Th89;
      case A21:p<>0.REAL 2 &
        (p`2/|.p.|<sn & p`1>=0);
        then A22:
         sn-FanMorphE.p= |[ |.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)),
          |.p.|* ((p`2/|.p.|-sn)/(1+sn))]| by A1,Th91;
         set p4= |[ |.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)),
          |.p.|* ((p`2/|.p.|-sn)/(1+sn))]|;
        A23: p4`1= |.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)) &
          p4`2= |.p.|* ((p`2/|.p.|-sn)/(1+sn)) by EUCLID:56;
     A24: |.p.|>=0 by TOPRNS_1:26;
     A25: |.p.|<>0 by A21,TOPRNS_1:25;
     then A26: (|.p.|)^2>0 by SQUARE_1:74;
     A27: 1+sn>0 by A1,Th3;
        A28: (p`2/|.p.|-sn)<=0 by A21,REAL_2:106;
        then A29: -(p`2/|.p.|-sn)>= -0 by REAL_1:50;
          --(p`2/|.p.|-sn)/(1+sn)<=0 by A27,A28,REAL_2:126;
        then A30: -((p`2/|.p.|-sn)/(1+sn))>=0 by REAL_1:66;
        A31: (|.p.|)^2 =(p`1)^2+(p`2)^2 by JGRAPH_3:10;
          0<=(p`1)^2 by SQUARE_1:72;
        then 0+(p`2)^2<=(p`1)^2+(p`2)^2 by REAL_1:55;
        then (p`2)^2/(|.p.|)^2 <= (|.p.|)^2/(|.p.|)^2 by A26,A31,REAL_1:73;
        then (p`2)^2/(|.p.|)^2 <= 1 by A26,XCMPLX_1:60;
        then ((p`2)/|.p.|)^2 <= 1 by SQUARE_1:69;
        then A32: (-((p`2)/|.p.|))^2 <= 1 by SQUARE_1:61;
          -(1+sn)<= -0 by A27,REAL_1:50;
        then (-(1+sn))/(1+sn)<=(-( p`2/|.p.|-sn))/(1+sn) by A27,A29,REAL_1:73;
        then A33: -1<=(-( p`2/|.p.|-sn))/(1+sn) by A27,XCMPLX_1:198;
         1>= -p`2/|.p.| by A32,Th4;
       then (1+sn)>= -p`2/|.p.|+sn by REAL_1:55;
       then (1+sn)>= -(p`2/|.p.|-sn) by XCMPLX_1:162;
       then (-(p`2/|.p.|-sn))/(1+sn)<=1 by A27,REAL_2:117;
       then ((-(p`2/|.p.|-sn))/(1+sn))^2<=1^2 by A33,JGRAPH_2:7;
       then A34: 1-((-(p`2/|.p.|-sn))/(1+sn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((p`2/|.p.|-sn))/(1+sn))^2>=0 by XCMPLX_1:188;
       then A35: 1-((p`2/|.p.|-sn)/(1+sn))^2>=0 by SQUARE_1:61;
         sqrt(1-((-(p`2/|.p.|-sn))/(1+sn))^2)>=0 by A34,SQUARE_1:def 4;
       then sqrt(1-(-(p`2/|.p.|-sn))^2/(1+sn)^2)>=0 by SQUARE_1:69;
       then sqrt(1-((p`2/|.p.|-sn))^2/(1+sn)^2)>=0 by SQUARE_1:61;
       then sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)>=0 by SQUARE_1:69;
        then q`1=0 by A2,A4,A5,A22,A23,A24,SQUARE_1:19;
        then (sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2))=0 by A2,A5,A22,A23,A25,
XCMPLX_1:6;
        then 1-((p`2/|.p.|-sn)/(1+sn))^2=0 by A35,SQUARE_1:92;
        then 1=0+((p`2/|.p.|-sn)/(1+sn))^2 by XCMPLX_1:27;
        then 1=sqrt((-((p`2/|.p.|-sn)/(1+sn)))^2) by SQUARE_1:61,83;
        then 1= -((p`2/|.p.|-sn)/(1+sn)) by A30,SQUARE_1:89;
        then 1= ((-(p`2/|.p.|-sn))/(1+sn)) by XCMPLX_1:188;
        then (1)*(1+sn)=-(p`2/|.p.|-sn) by A27,XCMPLX_1:88;
        then (1+sn)=-p`2/|.p.|+sn by XCMPLX_1:162;
        then 1+sn-sn=-p`2/|.p.| by XCMPLX_1:26;
        then 1=-p`2/|.p.| by XCMPLX_1:26;
        then 1=(-p`2)/|.p.| by XCMPLX_1:188;
        then (1)*|.p.|=-p`2 by A25,XCMPLX_1:88;
        then (-p`2)^2-(p`2)^2 =(p`1)^2 by A31,XCMPLX_1:26;
        then (p`2)^2-(p`2)^2 =(p`1)^2 by SQUARE_1:61;
        then 0= (p`1)^2 by XCMPLX_1:14;
        then p`1=0 by SQUARE_1:73;
       hence x1=x2 by A2,A5,Th89;
      end;
     hence x1=x2;
    case A36:
      q`2/|.q.|>=sn & q`1>=0 & q<>0.REAL 2;
        then A37:
         sn-FanMorphE.q= |[ |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1-sn))]| by A1,Th91;
         set q4= |[ |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1-sn))]|;
        A38: q4`1= |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)) &
          q4`2= |.q.|* ((q`2/|.q.|-sn)/(1-sn)) by EUCLID:56;
     A39: |.q.|>=0 by TOPRNS_1:26;
       |.q.|<>0 by A36,TOPRNS_1:25;
     then A40: (|.q.|)^2>0 by SQUARE_1:74;
        now per cases by JGRAPH_2:11;
      case A41: p`1<=0;
        then A42: (sn-FanMorphE).p=p by Th89;
        A43: |.q.|>=0 by TOPRNS_1:26;
        A44: |.q.|<>0 by A36,TOPRNS_1:25;
        then A45: (|.q.|)^2>0 by SQUARE_1:74;
        A46: 1-sn>0 by A1,Th3;
        A47: (q`2/|.q.|-sn)>=0 by A36,SQUARE_1:12;
        A48: (q`2/|.q.|-sn)>= 0 by A36,SQUARE_1:12;
        A49: (q`2/|.q.|-sn)/(1-sn)>=0 by A46,A47,REAL_2:125;
        A50: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
          0<=(q`1)^2 by SQUARE_1:72;
        then 0+(q`2)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
        then (q`2)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A45,A50,REAL_1:73;
        then (q`2)^2/(|.q.|)^2 <= 1 by A45,XCMPLX_1:60;
        then ((q`2)/|.q.|)^2 <= 1 by SQUARE_1:69;
        then 1>=q`2/|.q.| by Th4;
        then A51: 1-sn>=q`2/|.q.|-sn by REAL_1:49;
        A52: -(1-sn)<= -0 by A46,REAL_1:50;
           -(1-sn)<= -( q`2/|.q.|-sn) by A51,REAL_1:50;
         then (-(1-sn))/(1-sn)<=(-( q`2/|.q.|-sn))/(1-sn)
                        by A46,REAL_1:73;
        then A53: -1<=(-( q`2/|.q.|-sn))/(1-sn) by A46,XCMPLX_1:198;
         --(1-sn)>= -(q`2/|.q.|-sn) by A48,A52,REAL_1:50;
       then (-(q`2/|.q.|-sn))/(1-sn)<=1 by A46,REAL_2:117;
       then ((-(q`2/|.q.|-sn))/(1-sn))^2<=1^2 by A53,JGRAPH_2:7;
       then A54: 1-((-(q`2/|.q.|-sn))/(1-sn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((q`2/|.q.|-sn))/(1-sn))^2>=0 by XCMPLX_1:188;
       then A55: 1-((q`2/|.q.|-sn)/(1-sn))^2>=0 by SQUARE_1:61;
         sqrt(1-((-(q`2/|.q.|-sn))/(1-sn))^2)>=0 by A54,SQUARE_1:def 4;
       then sqrt(1-(-(q`2/|.q.|-sn))^2/(1-sn)^2)>=0 by SQUARE_1:69;
       then sqrt(1-(q`2/|.q.|-sn)^2/(1-sn)^2)>=0 by SQUARE_1:61;
       then sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)>=0 by SQUARE_1:69;
        then p`1=0 by A2,A37,A38,A41,A42,A43,SQUARE_1:19;
        then (sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2))=0 by A2,A37,A38,A42,A44,
XCMPLX_1:6;
        then 1-((q`2/|.q.|-sn)/(1-sn))^2=0 by A55,SQUARE_1:92;
        then 1=0+((q`2/|.q.|-sn)/(1-sn))^2 by XCMPLX_1:27;
        then 1= (q`2/|.q.|-sn)/(1-sn) by A49,SQUARE_1:83,89;
        then (1)*(1-sn)=(q`2/|.q.|-sn) by A46,XCMPLX_1:88;
        then 1-sn+sn=q`2/|.q.| by XCMPLX_1:27;
        then 1=q`2/|.q.| by XCMPLX_1:27;
        then (1)*|.q.|=q`2 by A44,XCMPLX_1:88;
        then (q`2)^2-(q`2)^2 =(q`1)^2 by A50,XCMPLX_1:26;
        then 0= (q`1)^2 by XCMPLX_1:14;
        then q`1=0 by SQUARE_1:73;
       hence x1=x2 by A2,A42,Th89;
      case A56:p<>0.REAL 2 &
          (p`2/|.p.|>=sn & p`1>=0);
        then A57:
         sn-FanMorphE.p= |[ |.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)),
          |.p.|* ((p`2/|.p.|-sn)/(1-sn))]| by A1,Th91;
         set p4= |[ |.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)),
          |.p.|* ((p`2/|.p.|-sn)/(1-sn))]|;
        A58: p4`1= |.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)) &
          p4`2= |.p.|* ((p`2/|.p.|-sn)/(1-sn)) by EUCLID:56;
        A59: |.q4.|>=0 by TOPRNS_1:26;
        A60: q4`1= |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2))&
        q4`2= |.q.|* ((q`2/|.q.|-sn)/(1-sn)) by EUCLID:56;
        A61: (q`2/|.q.|-sn)>= 0 by A36,SQUARE_1:12;
        A62: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
          0<=(q`1)^2 by SQUARE_1:72;
        then 0+(q`2)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
        then (q`2)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A40,A62,REAL_1:73;
        then (q`2)^2/(|.q.|)^2 <= 1 by A40,XCMPLX_1:60;
        then ((q`2)/|.q.|)^2 <= 1 by SQUARE_1:69;
        then 1>=q`2/|.q.| by Th4;
        then A63: 1-sn>=q`2/|.q.|-sn by REAL_1:49;
        A64: -(1-sn)<= -0 by A3,REAL_1:50;
           -(1-sn)<= -( q`2/|.q.|-sn) by A63,REAL_1:50;
         then (-(1-sn))/(1-sn)<=(-( q`2/|.q.|-sn))/(1-sn) by A3,REAL_1:73;
        then A65: -1<=(-( q`2/|.q.|-sn))/(1-sn) by A3,XCMPLX_1:198;
         --(1-sn)>= -(q`2/|.q.|-sn) by A61,A64,REAL_1:50;
       then (-(q`2/|.q.|-sn))/(1-sn)<=1 by A3,REAL_2:117;
       then ((-(q`2/|.q.|-sn))/(1-sn))^2<=1^2 by A65,JGRAPH_2:7;
       then 1-((-(q`2/|.q.|-sn))/(1-sn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((q`2/|.q.|-sn))/(1-sn))^2>=0 by XCMPLX_1:188;
       then A66: 1-((q`2/|.q.|-sn)/(1-sn))^2>=0 by SQUARE_1:61;
        A67: (q4`1)^2= (|.q.|)^2*(sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2))^2
                         by A60,SQUARE_1:68
         .= (|.q.|)^2*(1-((q`2/|.q.|-sn)/(1-sn))^2) by A66,SQUARE_1:def 4;
        A68: (q4`2)^2= (|.q.|)^2*((q`2/|.q.|-sn)/(1-sn))^2
                         by A60,SQUARE_1:68;
          (|.q4.|)^2=(q4`1)^2+(q4`2)^2 by JGRAPH_3:10
        .=(|.q.|)^2*(1-((q`2/|.q.|-sn)/(1-sn))^2
           +((q`2/|.q.|-sn)/(1-sn))^2) by A67,A68,XCMPLX_1:8
        .=(|.q.|)^2*1 by XCMPLX_1:27 .=(|.q.|)^2;
        then A69: sqrt((|.q4.|)^2)=|.q.| by A39,SQUARE_1:89;
        then A70: |.q4.|=|.q.| by A59,SQUARE_1:89;
        A71: |.p4.|>=0 by TOPRNS_1:26;
        A72: |.p.|>=0 by TOPRNS_1:26;
        A73: |.p.|<>0 by A56,TOPRNS_1:25;
        then A74: (|.p.|)^2>0 by SQUARE_1:74;
        A75: (p`2/|.p.|-sn)>= 0 by A56,SQUARE_1:12;
        A76: (|.p.|)^2 =(p`1)^2+(p`2)^2 by JGRAPH_3:10;
          0<=(p`1)^2 by SQUARE_1:72;
        then 0+(p`2)^2<=(p`1)^2+(p`2)^2 by REAL_1:55;
        then (p`2)^2/(|.p.|)^2 <= (|.p.|)^2/(|.p.|)^2 by A74,A76,REAL_1:73;
        then (p`2)^2/(|.p.|)^2 <= 1 by A74,XCMPLX_1:60;
        then ((p`2)/|.p.|)^2 <= 1 by SQUARE_1:69;
        then 1>=p`2/|.p.| by Th4;
        then A77: 1-sn>=p`2/|.p.|-sn by REAL_1:49;
        A78: -(1-sn)<= -0 by A3,REAL_1:50;
           -(1-sn)<= -( p`2/|.p.|-sn) by A77,REAL_1:50;
         then (-(1-sn))/(1-sn)<=(-( p`2/|.p.|-sn))/(1-sn) by A3,REAL_1:73;
        then A79: -1<=(-( p`2/|.p.|-sn))/(1-sn) by A3,XCMPLX_1:198;
         --(1-sn)>= -(p`2/|.p.|-sn) by A75,A78,REAL_1:50;
       then (-(p`2/|.p.|-sn))/(1-sn)<=1 by A3,REAL_2:117;
       then ((-(p`2/|.p.|-sn))/(1-sn))^2<=1^2 by A79,JGRAPH_2:7;
       then 1-((-(p`2/|.p.|-sn))/(1-sn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((p`2/|.p.|-sn))/(1-sn))^2>=0 by XCMPLX_1:188;
       then A80: 1-((p`2/|.p.|-sn)/(1-sn))^2>=0 by SQUARE_1:61;
        A81: (p4`1)^2= (|.p.|)^2*(sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2))^2
                         by A58,SQUARE_1:68
         .= (|.p.|)^2*(1-((p`2/|.p.|-sn)/(1-sn))^2) by A80,SQUARE_1:def 4;
        A82: (p4`2)^2= (|.p.|)^2*((p`2/|.p.|-sn)/(1-sn))^2
                         by A58,SQUARE_1:68;
          (|.p4.|)^2=(p4`1)^2+(p4`2)^2 by JGRAPH_3:10
        .=(|.p.|)^2*(1-((p`2/|.p.|-sn)/(1-sn))^2
           +((p`2/|.p.|-sn)/(1-sn))^2) by A81,A82,XCMPLX_1:8
        .=(|.p.|)^2*1 by XCMPLX_1:27 .=(|.p.|)^2;
        then A83: sqrt((|.p4.|)^2)=|.p.| by A72,SQUARE_1:89;
        then A84: |.p4.|=|.p.| by A71,SQUARE_1:89;
        ((p`2/|.p.|-sn)/(1-sn))
        =|.q.|* ((q`2/|.q.|-sn)/(1-sn))/|.p.| by A2,A37,A38,A57,A58,A73,
XCMPLX_1:90;
        then (p`2/|.p.|-sn)/(1-sn)=(q`2/|.q.|-sn)/(1-sn)
                              by A2,A37,A57,A69,A73,A83,XCMPLX_1:90;
        then (p`2/|.p.|-sn)/(1-sn)*(1-sn)=q`2/|.q.|-sn by A3,XCMPLX_1:88;
        then p`2/|.p.|-sn=q`2/|.q.|-sn by A3,XCMPLX_1:88;
        then p`2/|.p.|-sn+sn=q`2/|.q.| by XCMPLX_1:27;
        then p`2/|.p.|=q`2/|.q.| by XCMPLX_1:27;
        then p`2/|.p.|*|.p.|=q`2 by A2,A37,A57,A70,A73,A84,XCMPLX_1:88;
        then A85: p`2=q`2 by A73,XCMPLX_1:88;
        A86: |.p.|^2=(p`1)^2+(p`2)^2 by JGRAPH_3:10;
          |.q.|^2=(q`1)^2+(q`2)^2 by JGRAPH_3:10;
        then (p`1)^2+(p`2)^2-(p`2)^2=(q`1)^2 by A2,A37,A57,A70,A84,A85,A86,
XCMPLX_1:26;
        then (p`1)^2=(q`1)^2 by XCMPLX_1:26;
        then p`1=sqrt((q`1)^2) by A56,SQUARE_1:89;
        then A87: p`1=q`1 by A36,SQUARE_1:89;
          p=|[p`1,p`2]| by EUCLID:57;
       hence x1=x2 by A85,A87,EUCLID:57;
      case A88: p<>0.REAL 2 &
        (p`2/|.p.|<sn & p`1>=0);
        then A89:
         sn-FanMorphE.p= |[ |.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)),
          |.p.|* ((p`2/|.p.|-sn)/(1+sn))]| by A1,Th91;
         set p4= |[ |.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)),
          |.p.|* ((p`2/|.p.|-sn)/(1+sn))]|;
        A90: p4`1= |.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)) &
          p4`2= |.p.|* ((p`2/|.p.|-sn)/(1+sn)) by EUCLID:56;
         A91: p`2/|.p.|-sn<0 by A88,REAL_2:105;
           1+sn>0 by A1,Th3;
         then A92: ((p`2/|.p.|-sn)/(1+sn))<0 by A91,REAL_2:128;
        A93: |.p.|>=0 by TOPRNS_1:26;
          |.p.|<>0 by A88,TOPRNS_1:25;
         then A94: |.p.|* ((p`2/|.p.|-sn)/(1+sn))<0 by A92,A93,SQUARE_1:24;
         A95: q`2/|.q.|-sn>=0 by A36,SQUARE_1:12;
           1-sn>0 by A1,Th3;
         then ((q`2/|.q.|-sn)/(1-sn))>=0 by A95,REAL_2:125;
       hence x1=x2 by A2,A37,A38,A39,A89,A90,A94,REAL_2:121;
      end;
     hence x1=x2;
    case A96:
      q`2/|.q.|<sn & q`1>=0 & q<>0.REAL 2;
        then A97:
         sn-FanMorphE.q= |[ |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1+sn))]| by A1,Th91;
         set q4= |[ |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1+sn))]|;
        A98: q4`1= |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)) &
          q4`2= |.q.|* ((q`2/|.q.|-sn)/(1+sn)) by EUCLID:56;
     A99: |.q.|>=0 by TOPRNS_1:26;
     A100: |.q.|<>0 by A96,TOPRNS_1:25;
     then A101: (|.q.|)^2>0 by SQUARE_1:74;
        now per cases by JGRAPH_2:11;
      case A102: p`1<=0;
        then A103:(sn-FanMorphE).p=p by Th89;
        A104: 1+sn>0 by A1,Th3;
        A105: (q`2/|.q.|-sn)<=0 by A96,REAL_2:106;
        then A106: -(q`2/|.q.|-sn)>= -0 by REAL_1:50;
          --(q`2/|.q.|-sn)/(1+sn)<=0 by A104,A105,REAL_2:126;
        then A107: -((q`2/|.q.|-sn)/(1+sn))>=0 by REAL_1:66;
        A108: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
          0<=(q`1)^2 by SQUARE_1:72;
        then 0+(q`2)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
        then (q`2)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A101,A108,REAL_1:73;
        then (q`2)^2/(|.q.|)^2 <= 1 by A101,XCMPLX_1:60;
        then ((q`2)/|.q.|)^2 <= 1 by SQUARE_1:69;
        then A109: (-((q`2)/|.q.|))^2 <= 1 by SQUARE_1:61;
          -(1+sn)<= -0 by A104,REAL_1:50;
        then (-(1+sn))/(1+sn)<=(-( q`2/|.q.|-sn))/(1+sn) by A104,A106,REAL_1:73
;
        then A110: -1<=(-( q`2/|.q.|-sn))/(1+sn) by A104,XCMPLX_1:198;
         1>= -q`2/|.q.| by A109,Th4;
       then (1+sn)>= -q`2/|.q.|+sn by REAL_1:55;
       then (1+sn)>= -(q`2/|.q.|-sn) by XCMPLX_1:162;
       then (-(q`2/|.q.|-sn))/(1+sn)<=1 by A104,REAL_2:117;
       then ((-(q`2/|.q.|-sn))/(1+sn))^2<=1^2 by A110,JGRAPH_2:7;
       then A111: 1-((-(q`2/|.q.|-sn))/(1+sn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((q`2/|.q.|-sn))/(1+sn))^2>=0 by XCMPLX_1:188;
       then A112: 1-((q`2/|.q.|-sn)/(1+sn))^2>=0 by SQUARE_1:61;
         sqrt(1-((-(q`2/|.q.|-sn))/(1+sn))^2)>=0 by A111,SQUARE_1:def 4;
       then sqrt(1-(-(q`2/|.q.|-sn))^2/(1+sn)^2)>=0 by SQUARE_1:69;
       then sqrt(1-((q`2/|.q.|-sn))^2/(1+sn)^2)>=0 by SQUARE_1:61;
       then sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)>=0 by SQUARE_1:69;
        then p`1=0 by A2,A97,A98,A99,A102,A103,SQUARE_1:19;
        then (sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2))=0 by A2,A97,A98,A100,A103,
XCMPLX_1:6;
        then 1-((q`2/|.q.|-sn)/(1+sn))^2=0 by A112,SQUARE_1:92;
        then 1=0+((q`2/|.q.|-sn)/(1+sn))^2 by XCMPLX_1:27;
        then 1=sqrt((-((q`2/|.q.|-sn)/(1+sn)))^2) by SQUARE_1:61,83;
        then 1= -((q`2/|.q.|-sn)/(1+sn)) by A107,SQUARE_1:89;
        then 1= ((-(q`2/|.q.|-sn))/(1+sn)) by XCMPLX_1:188;
        then (1)*(1+sn)=-(q`2/|.q.|-sn) by A104,XCMPLX_1:88;
        then (1+sn)=-q`2/|.q.|+sn by XCMPLX_1:162;
        then 1+sn-sn=-q`2/|.q.| by XCMPLX_1:26;
        then 1=-q`2/|.q.| by XCMPLX_1:26;
        then 1=(-q`2)/|.q.| by XCMPLX_1:188;
        then (1)*|.q.|=-q`2 by A100,XCMPLX_1:88;
        then (-q`2)^2-(q`2)^2 =(q`1)^2 by A108,XCMPLX_1:26;
        then (q`2)^2-(q`2)^2 =(q`1)^2 by SQUARE_1:61;
        then 0= (q`1)^2 by XCMPLX_1:14;
        then q`1=0 by SQUARE_1:73;
       hence x1=x2 by A2,A103,Th89;
      case A113:p<>0.REAL 2 &
          (p`2/|.p.|>=sn & p`1>=0);
        then A114:
         sn-FanMorphE.p= |[ |.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)),
          |.p.|* ((p`2/|.p.|-sn)/(1-sn))]| by A1,Th91;
         set p4= |[ |.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)),
          |.p.|* ((p`2/|.p.|-sn)/(1-sn))]|;
        A115: p4`1= |.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1-sn))^2)) &
          p4`2= |.p.|* ((p`2/|.p.|-sn)/(1-sn)) by EUCLID:56;
         A116: q`2/|.q.|-sn<0 by A96,REAL_2:105;
           1+sn>0 by A1,Th3;
         then A117: ((q`2/|.q.|-sn)/(1+sn))<0 by A116,REAL_2:128;
        A118: |.q.|>=0 by TOPRNS_1:26;
          |.q.|<>0 by A96,TOPRNS_1:25;
         then A119: |.q.|* ((q`2/|.q.|-sn)/(1+sn))<0 by A117,A118,SQUARE_1:24;
         A120: p`2/|.p.|-sn>=0 by A113,SQUARE_1:12;
           1-sn>0 by A1,Th3;
         then A121: ((p`2/|.p.|-sn)/(1-sn))>=0 by A120,REAL_2:125;
           |.p.|>=0 by TOPRNS_1:26;
       hence x1=x2 by A2,A97,A98,A114,A115,A119,A121,REAL_2:121;
      case A122:p<>0.REAL 2 &
          (p`2/|.p.|<sn & p`1>=0);
        then A123:
         sn-FanMorphE.p= |[ |.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)),
          |.p.|* ((p`2/|.p.|-sn)/(1+sn))]| by A1,Th91;
         set p4= |[ |.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)),
          |.p.|* ((p`2/|.p.|-sn)/(1+sn))]|;
        A124: p4`1= |.p.|*(sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2)) &
          p4`2= |.p.|* ((p`2/|.p.|-sn)/(1+sn)) by EUCLID:56;
        A125: |.q4.|>=0 by TOPRNS_1:26;
        A126: q4`1= |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2))&
        q4`2= |.q.|* ((q`2/|.q.|-sn)/(1+sn)) by EUCLID:56;
          (q`2/|.q.|-sn)<=0 by A96,REAL_2:106;
        then A127: -(q`2/|.q.|-sn)>= -0 by REAL_1:50;
        A128: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
          0<=(q`1)^2 by SQUARE_1:72;
        then 0+(q`2)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
        then (q`2)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A101,A128,REAL_1:73;
        then (q`2)^2/(|.q.|)^2 <= 1 by A101,XCMPLX_1:60;
        then ((q`2)/|.q.|)^2 <= 1 by SQUARE_1:69;
        then -1<=q`2/|.q.| by Th4;
        then A129: -1-sn<=q`2/|.q.|-sn by REAL_1:49;
        A130: 1+sn>0 by A1,Th3;
        then -(1+sn)<= -0 by REAL_1:50;
        then (-(1+sn))/(1+sn)<=(-( q`2/|.q.|-sn))/(1+sn) by A127,A130,REAL_1:73
;
        then A131: -1<=(-( q`2/|.q.|-sn))/(1+sn) by A130,XCMPLX_1:198;
          -(1+sn)<=q`2/|.q.|-sn by A129,XCMPLX_1:161;
        then --(1+sn)>= -(q`2/|.q.|-sn) by REAL_1:50;
       then (-(q`2/|.q.|-sn))/(1+sn)<=1 by A130,REAL_2:117;
       then ((-(q`2/|.q.|-sn))/(1+sn))^2<=1^2 by A131,JGRAPH_2:7;
       then 1-((-(q`2/|.q.|-sn))/(1+sn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((q`2/|.q.|-sn))/(1+sn))^2>=0 by XCMPLX_1:188;
       then A132: 1-((q`2/|.q.|-sn)/(1+sn))^2>=0 by SQUARE_1:61;
        A133: (q4`1)^2= (|.q.|)^2*(sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2))^2
                         by A126,SQUARE_1:68
         .= (|.q.|)^2*(1-((q`2/|.q.|-sn)/(1+sn))^2) by A132,SQUARE_1:def 4;
        A134: (q4`2)^2= (|.q.|)^2*((q`2/|.q.|-sn)/(1+sn))^2
                         by A126,SQUARE_1:68;
          (|.q4.|)^2=(q4`1)^2+(q4`2)^2 by JGRAPH_3:10
        .=(|.q.|)^2*(1-((q`2/|.q.|-sn)/(1+sn))^2
           +((q`2/|.q.|-sn)/(1+sn))^2) by A133,A134,XCMPLX_1:8
        .=(|.q.|)^2*1 by XCMPLX_1:27 .=(|.q.|)^2;
        then A135: sqrt((|.q4.|)^2)=|.q.| by A99,SQUARE_1:89;
        then A136: |.q4.|=|.q.| by A125,SQUARE_1:89;
        A137: |.p4.|>=0 by TOPRNS_1:26;
        A138: |.p.|>=0 by TOPRNS_1:26;
        A139: |.p.|<>0 by A122,TOPRNS_1:25;
        then A140: (|.p.|)^2>0 by SQUARE_1:74;
          (p`2/|.p.|-sn)<=0 by A122,REAL_2:106;
        then A141: -(p`2/|.p.|-sn)>= -0 by REAL_1:50;
        A142: (|.p.|)^2 =(p`1)^2+(p`2)^2 by JGRAPH_3:10;
          0<=(p`1)^2 by SQUARE_1:72;
        then 0+(p`2)^2<=(p`1)^2+(p`2)^2 by REAL_1:55;
        then (p`2)^2/(|.p.|)^2 <= (|.p.|)^2/(|.p.|)^2 by A140,A142,REAL_1:73;
        then (p`2)^2/(|.p.|)^2 <= 1 by A140,XCMPLX_1:60;
        then ((p`2)/|.p.|)^2 <= 1 by SQUARE_1:69;
        then -1<=p`2/|.p.| by Th4;
        then A143: -1-sn<=p`2/|.p.|-sn by REAL_1:49;
          -(1+sn)<= -0 by A130,REAL_1:50;
        then (-(1+sn))/(1+sn)<=(-( p`2/|.p.|-sn))/(1+sn) by A130,A141,REAL_1:73
;
        then A144: -1<=(-( p`2/|.p.|-sn))/(1+sn) by A130,XCMPLX_1:198;
          -(1+sn)<=p`2/|.p.|-sn by A143,XCMPLX_1:161;
        then --(1+sn)>= -(p`2/|.p.|-sn) by REAL_1:50;
       then (-(p`2/|.p.|-sn))/(1+sn)<=1 by A130,REAL_2:117;
       then ((-(p`2/|.p.|-sn))/(1+sn))^2<=1^2 by A144,JGRAPH_2:7;
       then 1-((-(p`2/|.p.|-sn))/(1+sn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((p`2/|.p.|-sn))/(1+sn))^2>=0 by XCMPLX_1:188;
       then A145: 1-((p`2/|.p.|-sn)/(1+sn))^2>=0 by SQUARE_1:61;
        A146: (p4`1)^2= (|.p.|)^2*(sqrt(1-((p`2/|.p.|-sn)/(1+sn))^2))^2
                         by A124,SQUARE_1:68
         .= (|.p.|)^2*(1-((p`2/|.p.|-sn)/(1+sn))^2)
                      by A145,SQUARE_1:def 4;
        A147: (p4`2)^2= (|.p.|)^2*((p`2/|.p.|-sn)/(1+sn))^2
                         by A124,SQUARE_1:68;
          (|.p4.|)^2=(p4`1)^2+(p4`2)^2 by JGRAPH_3:10
        .=(|.p.|)^2*(1-((p`2/|.p.|-sn)/(1+sn))^2
           +((p`2/|.p.|-sn)/(1+sn))^2) by A146,A147,XCMPLX_1:8
        .=(|.p.|)^2*1 by XCMPLX_1:27 .=(|.p.|)^2;
        then A148: sqrt((|.p4.|)^2)=|.p.| by A138,SQUARE_1:89;
        then A149: |.p4.|=|.p.| by A137,SQUARE_1:89;
        ((p`2/|.p.|-sn)/(1+sn))
        =|.q.|* ((q`2/|.q.|-sn)/(1+sn))/|.p.| by A2,A97,A98,A123,A124,A139,
XCMPLX_1:90;
        then (p`2/|.p.|-sn)/(1+sn)=(q`2/|.q.|-sn)/(1+sn)
                              by A2,A97,A123,A135,A139,A148,XCMPLX_1:90
;
        then (p`2/|.p.|-sn)/(1+sn)*(1+sn)=q`2/|.q.|-sn by A130,XCMPLX_1:88;
        then p`2/|.p.|-sn=q`2/|.q.|-sn by A130,XCMPLX_1:88;
        then p`2/|.p.|-sn+sn=q`2/|.q.| by XCMPLX_1:27;
        then p`2/|.p.|=q`2/|.q.| by XCMPLX_1:27;
        then p`2/|.p.|*|.p.|=q`2 by A2,A97,A123,A136,A139,A149,XCMPLX_1:88;
        then A150: p`2=q`2 by A139,XCMPLX_1:88;
        A151: |.p.|^2=(p`1)^2+(p`2)^2 by JGRAPH_3:10;
          |.q.|^2=(q`1)^2+(q`2)^2 by JGRAPH_3:10;
        then (p`1)^2+(p`2)^2-(p`2)^2=(q`1)^2 by A2,A97,A123,A136,A149,A150,A151
,XCMPLX_1:26;
        then (p`1)^2=(q`1)^2 by XCMPLX_1:26;
        then p`1=sqrt((q`1)^2) by A122,SQUARE_1:89;
        then A152: p`1=q`1 by A96,SQUARE_1:89;
          p=|[p`1,p`2]| by EUCLID:57;
       hence x1=x2 by A150,A152,EUCLID:57;
      end;
     hence x1=x2;
    end;
   hence x1=x2;
  end;
 hence thesis by FUNCT_1:def 8;
end;

theorem Th110: for sn being Real
  st -1<sn & sn<1 holds
  (sn-FanMorphE) is map of TOP-REAL 2,TOP-REAL 2 &
 rng (sn-FanMorphE) = the carrier of TOP-REAL 2
proof let sn be Real;
  assume A1:  -1<sn & sn<1;
thus (sn-FanMorphE) is map of TOP-REAL 2,TOP-REAL 2 ;
  A2:for f being map of TOP-REAL 2,TOP-REAL 2 st f=(sn-FanMorphE) holds
    rng (sn-FanMorphE)=the carrier of TOP-REAL 2
  proof let f be map of TOP-REAL 2,TOP-REAL 2;
    assume A3:f=(sn-FanMorphE);
    A4:dom f=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
      the carrier of TOP-REAL 2 c= rng f
    proof let y be set;assume y in the carrier of TOP-REAL 2;
     then reconsider p2=y as Point of TOP-REAL 2;
     set q=p2;
       now per cases by JGRAPH_2:11;
     case q`1<=0;
       then q in dom (sn-FanMorphE) & y=(sn-FanMorphE).q by A3,A4,Th89;
      hence ex x being set st x in dom (sn-FanMorphE) & y=(sn-FanMorphE).x;
     case A5: q`2/|.q.|>=0 & q`1>=0 & q<>0.REAL 2;
          set px=|[ (|.q.|)*sqrt(1-(q`2/|.q.|*(1-sn)+sn)^2),
          |.q.|*(q`2/|.q.|*(1-sn)+sn)]|;
        A6: px`1 = (|.q.|)*sqrt(1-(q`2/|.q.|*(1-sn)+sn)^2) &
         px`2 = |.q.|*(q`2/|.q.|*(1-sn)+sn) by EUCLID:56;
        then A7: |.px.|^2=((|.q.|)*sqrt(1-(q`2/|.q.|*(1-sn)+sn)^2))^2
        +(|.q.|*(q`2/|.q.|*(1-sn)+sn))^2 by JGRAPH_3:10
        .=(|.q.|)^2*(sqrt(1-(q`2/|.q.|*(1-sn)+sn)^2))^2
        +(|.q.|*(q`2/|.q.|*(1-sn)+sn))^2 by SQUARE_1:68
        .=(|.q.|)^2*(sqrt(1-(q`2/|.q.|*(1-sn)+sn)^2))^2
        +(|.q.|)^2*((q`2/|.q.|*(1-sn)+sn))^2 by SQUARE_1:68;
       A8: |.q.|<>0 by A5,TOPRNS_1:25;
         1-sn>=0 by A1,Th3;
       then A9: q`2/|.q.|*(1-sn)>=0 by A5,REAL_2:121;
       then A10: (q`2/|.q.|*(1-sn)+sn)>=0+sn by REAL_1:55;
       A11: |.q.|>=0 by TOPRNS_1:26;
       A12: |.q.|^2>0 by A8,SQUARE_1:74;
         --(1+sn)>0 by A1,Th3;
       then -(1+sn)<0 by REAL_1:66;
       then -1-sn <0 by XCMPLX_1:161;
       then -1-sn<= q`2/|.q.|*(1-sn) by A9,AXIOMS:22;
       then -1-sn+sn<= q`2/|.q.|*(1-sn)+sn by REAL_1:55;
       then A13: -1<= (q`2/|.q.|*(1-sn)+sn) by XCMPLX_1:27;
       A14: 1-sn>0 by A1,Th3;
       A15: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
         0<=(q`1)^2 by SQUARE_1:72;
       then 0+(q`2)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
       then (q`2)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A12,A15,REAL_1:73;
       then (q`2)^2/(|.q.|)^2 <= 1 by A12,XCMPLX_1:60;
       then ((q`2)/|.q.|)^2 <= 1 by SQUARE_1:69;
       then q`2/|.q.|<=1 by Th4;
       then q`2/|.q.|*(1-sn) <=(1)*(1-sn) by A14,AXIOMS:25;
       then q`2/|.q.|*(1-sn)+sn-sn <=1-sn by XCMPLX_1:26;
       then (q`2/|.q.|*(1-sn)+sn) <=1 by REAL_1:54;
       then 1^2>=(q`2/|.q.|*(1-sn)+sn)^2 by A13,JGRAPH_2:7;
       then A16: 1-(q`2/|.q.|*(1-sn)+sn)^2>=0 by SQUARE_1:12,59;
       then A17: |.px.|^2=(|.q.|)^2*(1-(q`2/|.q.|*(1-sn)+sn)^2)
        +(|.q.|)^2*((q`2/|.q.|*(1-sn)+sn))^2 by A7,SQUARE_1:def 4
        .= (|.q.|)^2*((1-(q`2/|.q.|*(1-sn)+sn)^2)
        +((q`2/|.q.|*(1-sn)+sn))^2) by XCMPLX_1:8
        .= (|.q.|)^2*(1) by XCMPLX_1:27
        .= (|.q.|)^2;
          |.px.|>=0 by TOPRNS_1:26;
        then A18: |.px.|=sqrt(|.q.|^2) by A17,SQUARE_1:89
        .=|.q.| by A11,SQUARE_1:89;
        then A19: |.px.|<>0 by A5,TOPRNS_1:25;
         sqrt(1-(q`2/|.q.|*(1-sn)+sn)^2)>=0 by A16,SQUARE_1:def 4;
        then px`2/|.px.| >=sn & px`1>=0 & px<>0.REAL 2
                                by A6,A8,A10,A11,A18,REAL_2:121,TOPRNS_1:24,
XCMPLX_1:90;
    then A20: (sn-FanMorphE).px
       =|[ |.px.|*(sqrt(1-((px`2/|.px.|-sn)/(1-sn))^2)),
               |.px.|* ((px`2/|.px.|-sn)/(1-sn))]|
                            by A1,Th91;
        A21: |.px.|* ((px`2/|.px.|-sn)/(1-sn))
        =|.q.|* (( ((q`2/|.q.|*(1-sn)+sn))-sn)/(1-sn)) by A6,A8,A18,XCMPLX_1:90
        .=|.q.|* (( q`2/|.q.|*(1-sn) )/(1-sn)) by XCMPLX_1:26
        .=|.q.|* ( q`2/|.q.|) by A14,XCMPLX_1:90
        .= q`2 by A8,XCMPLX_1:88;
        then A22: |.px.|*(sqrt(1-((px`2/|.px.|-sn)/(1-sn))^2))
           = |.px.|*(sqrt(1-(q`2/|.px.|)^2)) by A19,XCMPLX_1:90
           .= |.px.|*(sqrt(1-(q`2)^2/(|.px.|)^2)) by SQUARE_1:69
           .= |.px.|*(sqrt( (|.px.|)^2/(|.px.|)^2-(q`2)^2/(|.px.|)^2))
                            by A12,A17,XCMPLX_1:60
           .= |.px.|*(sqrt( ((|.px.|)^2-(q`2)^2)/(|.px.|)^2))
                            by XCMPLX_1:121
           .= |.px.|*(sqrt( ((q`1)^2+(q`2)^2-(q`2)^2)/(|.px.|)^2))
                            by A17,JGRAPH_3:10
           .= |.px.|*(sqrt( ((q`1)^2)/(|.px.|)^2)) by XCMPLX_1:26
           .= |.px.|*(sqrt((q`1/|.q.|)^2)) by A18,SQUARE_1:69;
           q`1/|.q.|>=0 by A5,A11,REAL_2:125;
         then |.px.|*(sqrt((q`1/|.q.|)^2))=|.q.|*(q`1/|.q.|) by A18,SQUARE_1:89
         .=q`1 by A8,XCMPLX_1:88;
        then A23:q=(sn-FanMorphE).px by A20,A21,A22,EUCLID:57;
          dom (sn-FanMorphE)=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
      hence ex x being set st x in dom (sn-FanMorphE) & y=(sn-FanMorphE).x
                               by A23;
     case A24: q`2/|.q.|<0 & q`1>=0 & q<>0.REAL 2;
          set px=|[ (|.q.|)*sqrt(1-(q`2/|.q.|*(1+sn)+sn)^2),
                   |.q.|*(q`2/|.q.|*(1+sn)+sn)]|;
        A25: px`1 = (|.q.|)*sqrt(1-(q`2/|.q.|*(1+sn)+sn)^2) &
         px`2 = |.q.|*(q`2/|.q.|*(1+sn)+sn) by EUCLID:56;
        then A26: |.px.|^2=((|.q.|)*sqrt(1-(q`2/|.q.|*(1+sn)+sn)^2))^2
        +(|.q.|*(q`2/|.q.|*(1+sn)+sn))^2 by JGRAPH_3:10
        .=(|.q.|)^2*(sqrt(1-(q`2/|.q.|*(1+sn)+sn)^2))^2
        +(|.q.|*(q`2/|.q.|*(1+sn)+sn))^2 by SQUARE_1:68
        .=(|.q.|)^2*(sqrt(1-(q`2/|.q.|*(1+sn)+sn)^2))^2
        +(|.q.|)^2*((q`2/|.q.|*(1+sn)+sn))^2 by SQUARE_1:68;
       A27: |.q.|<>0 by A24,TOPRNS_1:25;
         1+sn>=0 by A1,Th3;
       then A28: q`2/|.q.|*(1+sn)<=0 by A24,REAL_2:123;
       then A29: (q`2/|.q.|*(1+sn)+sn)<=0+sn by REAL_1:55;
       A30: |.q.|>=0 by TOPRNS_1:26;
       A31: |.q.|^2>0 by A27,SQUARE_1:74;
         (1-sn)>0 by A1,Th3;
       then 1-sn>= q`2/|.q.|*(1+sn) by A28,AXIOMS:22;
       then 1-sn+sn>= q`2/|.q.|*(1+sn)+sn by REAL_1:55;
       then A32: 1>= (q`2/|.q.|*(1+sn)+sn) by XCMPLX_1:27;
       A33: 1+sn>0 by A1,Th3;
       A34: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
         0<=(q`1)^2 by SQUARE_1:72;
       then 0+(q`2)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
       then (q`2)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A31,A34,REAL_1:73;
       then (q`2)^2/(|.q.|)^2 <= 1 by A31,XCMPLX_1:60;
       then ((q`2)/|.q.|)^2 <= 1 by SQUARE_1:69;
       then q`2/|.q.|>=-1 by Th4;
       then q`2/|.q.|*(1+sn) >=(-1)*(1+sn) by A33,AXIOMS:25;
       then q`2/|.q.|*(1+sn) >=-(1+sn) by XCMPLX_1:180;
       then q`2/|.q.|*(1+sn) >=-1-sn by XCMPLX_1:161;
       then q`2/|.q.|*(1+sn)+sn-sn >=-1-sn by XCMPLX_1:26;
       then (q`2/|.q.|*(1+sn)+sn) >=-1 by REAL_1:54;
       then 1^2>=(q`2/|.q.|*(1+sn)+sn)^2 by A32,JGRAPH_2:7;
       then A35: 1-(q`2/|.q.|*(1+sn)+sn)^2>=0 by SQUARE_1:12,59;
       then A36: |.px.|^2=(|.q.|)^2*(1-(q`2/|.q.|*(1+sn)+sn)^2)
        +(|.q.|)^2*((q`2/|.q.|*(1+sn)+sn))^2 by A26,SQUARE_1:def 4
        .= (|.q.|)^2*((1-(q`2/|.q.|*(1+sn)+sn)^2)
        +((q`2/|.q.|*(1+sn)+sn))^2) by XCMPLX_1:8
        .= (|.q.|)^2*(1) by XCMPLX_1:27
        .= (|.q.|)^2;
          |.px.|>=0 by TOPRNS_1:26;
        then A37: |.px.|=sqrt(|.q.|^2) by A36,SQUARE_1:89
        .=|.q.| by A30,SQUARE_1:89;
        then A38: |.px.|<>0 by A24,TOPRNS_1:25;
         sqrt(1-(q`2/|.q.|*(1+sn)+sn)^2)>=0 by A35,SQUARE_1:def 4;
        then px`2/|.px.| <=sn & px`1>=0 & px<>0.REAL 2
                           by A25,A27,A29,A30,A37,REAL_2:121,TOPRNS_1:24,
XCMPLX_1:90;
    then A39: (sn-FanMorphE).px
       =|[ |.px.|*(sqrt(1-((px`2/|.px.|-sn)/(1+sn))^2)),
               |.px.|* ((px`2/|.px.|-sn)/(1+sn))]| by A1,Th91;
        A40: |.px.|* ((px`2/|.px.|-sn)/(1+sn))
        =|.q.|* (( ((q`2/|.q.|*(1+sn)+sn))-sn)/(1+sn)) by A25,A27,A37,XCMPLX_1:
90
        .=|.q.|* (( q`2/|.q.|*(1+sn) )/(1+sn)) by XCMPLX_1:26
        .=|.q.|* ( q`2/|.q.|) by A33,XCMPLX_1:90
        .= q`2 by A27,XCMPLX_1:88;
        then A41: |.px.|*(sqrt(1-((px`2/|.px.|-sn)/(1+sn))^2))
           = |.px.|*(sqrt(1-(q`2/|.px.|)^2)) by A38,XCMPLX_1:90
           .= |.px.|*(sqrt(1-(q`2)^2/(|.px.|)^2)) by SQUARE_1:69
           .= |.px.|*(sqrt( (|.px.|)^2/(|.px.|)^2-(q`2)^2/(|.px.|)^2))
                            by A31,A36,XCMPLX_1:60
           .= |.px.|*(sqrt( ((|.px.|)^2-(q`2)^2)/(|.px.|)^2))
                            by XCMPLX_1:121
           .= |.px.|*(sqrt( ((q`1)^2+(q`2)^2-(q`2)^2)/(|.px.|)^2))
                            by A36,JGRAPH_3:10
           .= |.px.|*(sqrt( ((q`1)^2)/(|.px.|)^2)) by XCMPLX_1:26
           .= |.px.|*(sqrt((q`1/|.q.|)^2)) by A37,SQUARE_1:69;
           q`1/|.q.|>=0 by A24,A30,REAL_2:125;
         then |.px.|*(sqrt((q`1/|.q.|)^2))
         =|.q.|*(q`1/|.q.|) by A37,SQUARE_1:89
         .=q`1 by A27,XCMPLX_1:88;
        then A42:q=(sn-FanMorphE).px by A39,A40,A41,EUCLID:57;
          dom (sn-FanMorphE)=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
      hence ex x being set st x in dom (sn-FanMorphE) & y=(sn-FanMorphE).x
                               by A42;
     end;
    hence y in rng f by A3,FUNCT_1:def 5;
   end;
   hence rng (sn-FanMorphE)=the carrier of TOP-REAL 2
                 by A3,XBOOLE_0:def 10;
  end;
 thus rng (sn-FanMorphE)=the carrier of TOP-REAL 2 by A2;
end;

theorem Th111: for sn being Real,p2 being Point of TOP-REAL 2
  st -1<sn & sn<1
  ex K being non empty compact Subset of TOP-REAL 2 st
  K = (sn-FanMorphE).:K &
 (ex V2 being Subset of TOP-REAL 2 st
  p2 in V2 & V2 is open & V2 c= K & (sn-FanMorphE).p2 in V2)
 proof let sn be Real,p2 be Point of TOP-REAL 2;
  assume A1: -1<sn & sn<1;
    A2: the carrier of TOP-REAL 2=REAL 2 by EUCLID:25;
    A3: Euclid 2=MetrStruct(#REAL 2,Pitag_dist 2#) by EUCLID:def 7;
    then reconsider O=0.REAL 2 as Point of Euclid 2 by EUCLID:25;
    A4: TOP-REAL 2=TopSpaceMetr(Euclid 2) by EUCLID:def 8;
    reconsider V0=Ball(O,|.p2.|+1) as Subset of TOP-REAL 2
                              by A2,A3;
    A5: |. p2.|<|.p2.|+1 by REAL_1:69;
      0<=|.p2.| by TOPRNS_1:26;
    then A6: O in V0 by A5,GOBOARD6:4;
    A7: V0 c= cl_Ball(O,|.p2.|+1) by METRIC_1:15;
    reconsider K01=cl_Ball(O,|.p2.|+1) as Subset of TOP-REAL 2
                            by A2,A3;
      K01 is compact by Th22;
    then reconsider K0=cl_Ball(O,|.p2.|+1)
          as non empty compact Subset of TOP-REAL 2 by A6,A7;
    reconsider u2=p2 as Point of Euclid 2 by A3,EUCLID:25;
    A8: (sn-FanMorphE).:K0 c= K0
     proof let y be set;assume y in (sn-FanMorphE).:K0;
       then consider x being set such that
       A9: x in dom (sn-FanMorphE) & x in K0 & y=(sn-FanMorphE).x
                     by FUNCT_1:def 12;
       reconsider q=x as Point of TOP-REAL 2 by A9;
       reconsider uq=q as Point of Euclid 2 by A3,EUCLID:25;
       A10: y in rng (sn-FanMorphE) by A9,FUNCT_1:def 5;
       then reconsider u=y as Point of Euclid 2 by A3,EUCLID:25;
       reconsider q4=y as Point of TOP-REAL 2 by A10;
       A11: |.q4.|=|.q.| by A9,Th104;
         dist(O,uq)<= |.p2.|+1 by A9,METRIC_1:13;
       then |.0.REAL 2 - q.|<= |.p2.|+1 by JGRAPH_1:45;
       then |.0.REAL 2 +- q.|<= |.p2.|+1 by EUCLID:45;
       then |. -q.|<= |.p2.|+1 by EUCLID:31;
       then |.q.|<= |.p2.|+1 by TOPRNS_1:27;
       then |. -q4.|<= |.p2.|+1 by A11,TOPRNS_1:27;
       then |.0.REAL 2 +- q4.|<= |.p2.|+1 by EUCLID:31;
       then |.0.REAL 2 - q4.|<= |.p2.|+1 by EUCLID:45;
       then dist(O,u)<= |.p2.|+1 by JGRAPH_1:45;
      hence y in K0 by METRIC_1:13;
     end;
      K0 c= (sn-FanMorphE).:K0
     proof let y be set;assume A12: y in K0;
       then reconsider u=y as Point of Euclid 2;
       reconsider q4=y as Point of TOP-REAL 2 by A12;
         the carrier of TOP-REAL 2 c= rng (sn-FanMorphE) by A1,Th110;
       then q4 in rng (sn-FanMorphE) by TARSKI:def 3;
       then consider x being set such that
       A13: x in dom (sn-FanMorphE) & y=(sn-FanMorphE).x by FUNCT_1:def 5;
       reconsider q=x as Point of TOP-REAL 2 by A13,FUNCT_2:def 1;
       reconsider uq=q as Point of Euclid 2 by A3,EUCLID:25;
       A14: |.q4.|=|.q.| by A13,Th104;
         q = uq & q4 = u & O=0.REAL 2;
       then q in K0 by A12,A14,Lm11;
      hence y in (sn-FanMorphE).:K0 by A13,FUNCT_1:def 12;
     end;
    then A15: K0= (sn-FanMorphE).:K0 by A8,XBOOLE_0:def 10;
      |. -p2.|<|.p2.|+1 by A5,TOPRNS_1:27;
    then |.0.REAL 2 +- p2.|<|.p2.|+1 by EUCLID:31;
    then |.0.REAL 2 - p2.|<|.p2.|+1 by EUCLID:45;
    then dist(O,u2)<|.p2.|+1 by JGRAPH_1:45;
    then A16: p2 in V0 by METRIC_1:12;
    A17: V0 is open by A4,TOPMETR:21;
     set q3= (sn-FanMorphE).p2;
     reconsider u3=q3 as Point of Euclid 2 by A3,EUCLID:25;
       |.q3.|=|.p2.| by Th104;
    then |. -q3.|<|.p2.|+1 by A5,TOPRNS_1:27;
    then |.0.REAL 2 +- q3.|<|.p2.|+1 by EUCLID:31;
    then |.0.REAL 2 - q3.|<|.p2.|+1 by EUCLID:45;
     then dist(O,u3)< |.p2.|+1 by JGRAPH_1:45;
    then (sn-FanMorphE).p2 in V0 by METRIC_1:12;
   hence thesis by A7,A15,A16,A17;
 end;

theorem for sn being Real st -1<sn & sn<1
 ex f being map of TOP-REAL 2,TOP-REAL 2 st f=(sn-FanMorphE)
   & f is_homeomorphism
proof let sn be Real;assume A1:  -1<sn & sn<1;
  then A2: (sn-FanMorphE) is map of TOP-REAL 2,TOP-REAL 2 &
  rng (sn-FanMorphE) = the carrier of TOP-REAL 2 by Th110;
  reconsider f=(sn-FanMorphE) as map of TOP-REAL 2,TOP-REAL 2;
  consider h being map of (TOP-REAL 2),(TOP-REAL 2)
  such that A3: h=(sn-FanMorphE) & h is continuous by A1,Th108;
  A4: f is one-to-one by A1,Th109;
  A5: rng f=[#](TOP-REAL 2) by A2,PRE_TOPC:12;
   for p2 being Point of TOP-REAL 2
  ex K being non empty compact Subset of TOP-REAL 2 st K = f.:K &
 (ex V2 being Subset of TOP-REAL 2 st
  p2 in V2 & V2 is open & V2 c= K & f.p2 in V2) by A1,Th111;
  then f is_homeomorphism by A3,A4,A5,Th8;
 hence thesis;
end;

theorem Th113:
 for sn being Real,q being Point of TOP-REAL 2 st -1<sn & sn<1 & q`1>0
 & q`2/|.q.|>=sn 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;
  A2: |.q.|>=0 by TOPRNS_1:26;
  let p be Point of TOP-REAL 2;
    assume p=(sn-FanMorphE).q;
    then A3: p=|[ |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1-sn))]| by A1,Th89;
      set qz=p;
      A4: |.q.|<>0 by A1,JGRAPH_2:11,TOPRNS_1:25;
      then A5: (|.q.|)^2>0 by SQUARE_1:74;
     A6: 1-sn>0 by A1,Th3;
        A7: qz`1= |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2))&
        qz`2= |.q.|* ((q`2/|.q.|-sn)/(1-sn)) by A3,EUCLID:56;
        A8: (q`2/|.q.|-sn)>= 0 by A1,SQUARE_1:12;
        A9: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
          0<(q`1)^2 by A1,SQUARE_1:74;
        then 0+(q`2)^2<(q`1)^2+(q`2)^2 by REAL_1:67;
        then (q`2)^2/(|.q.|)^2 < (|.q.|)^2/(|.q.|)^2 by A5,A9,REAL_1:73;
        then (q`2)^2/(|.q.|)^2 < 1 by A5,XCMPLX_1:60;
        then ((q`2)/|.q.|)^2 < 1 by SQUARE_1:69;
        then 1>q`2/|.q.| by Th5;
        then A10: 1-sn>q`2/|.q.|-sn by REAL_1:54;
        A11: -(1-sn)< -0 by A6,REAL_1:50;
           -(1-sn)< -( q`2/|.q.|-sn) by A10,REAL_1:50;
         then (-(1-sn))/(1-sn)<(-( q`2/|.q.|-sn))/(1-sn)
                        by A6,REAL_1:73;
        then A12: -1<(-( q`2/|.q.|-sn))/(1-sn) by A6,XCMPLX_1:198;
         --(1-sn)> -(q`2/|.q.|-sn) by A8,A11,REAL_1:50;
       then (-(q`2/|.q.|-sn))/(1-sn)<1 by A6,REAL_2:118;
       then ((-(q`2/|.q.|-sn))/(1-sn))^2<1^2 by A12,JGRAPH_2:8;
       then 1-((-(q`2/|.q.|-sn))/(1-sn))^2>0 by SQUARE_1:11,59;
       then sqrt(1-((-(q`2/|.q.|-sn))/(1-sn))^2)>0 by SQUARE_1:93;
       then sqrt(1-(-(q`2/|.q.|-sn))^2/(1-sn)^2)> 0 by SQUARE_1:69;
       then sqrt(1-(q`2/|.q.|-sn)^2/(1-sn)^2)> 0 by SQUARE_1:61;
       then A13: sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)> 0 by SQUARE_1:69;
         ((q`2/|.q.|-sn)/(1-sn))>=0 by A6,A8,SQUARE_1:27;
   hence p`1>0 & p`2>=0 by A2,A4,A7,A13,SQUARE_1:19,21;
 end;

theorem Th114:
 for sn being Real,q being Point of TOP-REAL 2 st -1<sn & sn<1 & q`1>0
 & q`2/|.q.|<sn 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;
  A2: |.q.|>=0 by TOPRNS_1:26;
  let p be Point of TOP-REAL 2;
    assume p=(sn-FanMorphE).q;
    then A3: p=|[ |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1+sn))]| by A1,Th90;
      set qz=p;
          A4: |.q.|<>0 by A1,JGRAPH_2:11,TOPRNS_1:25;
      then A5: (|.q.|)^2>0 by SQUARE_1:74;
      A6: (q`2/|.q.|-sn)< 0 by A1,REAL_2:105;
     A7: 1+sn>0 by A1,Th3;
        A8: qz`1= |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2))&
        qz`2= |.q.|* ((q`2/|.q.|-sn)/(1+sn)) by A3,EUCLID:56;
        A9: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
          0<(q`1)^2 by A1,SQUARE_1:74;
        then 0+(q`2)^2<(q`1)^2+(q`2)^2 by REAL_1:67;
        then (q`2)^2/(|.q.|)^2 < (|.q.|)^2/(|.q.|)^2 by A5,A9,REAL_1:73;
        then (q`2)^2/(|.q.|)^2 < 1 by A5,XCMPLX_1:60;
        then ((q`2)/|.q.|)^2 < 1 by SQUARE_1:69;
        then -1<q`2/|.q.| by Th5;
        then -1-sn<q`2/|.q.|-sn by REAL_1:54;
        then -(1+sn)<q`2/|.q.|-sn by XCMPLX_1:161;
        then A10: --(1+sn)> -(q`2/|.q.|-sn) by REAL_1:50;
        A11: -( q`2/|.q.|-sn)>0 by A6,REAL_1:66;
          -(1+sn)< -0 by A7,REAL_1:50;
         then -(1+sn)< -( q`2/|.q.|-sn) by A11,AXIOMS:22;
         then (-(1+sn))/(1+sn)<(-( q`2/|.q.|-sn))/(1+sn) by A7,REAL_1:73;
        then A12: -1<(-( q`2/|.q.|-sn))/(1+sn) by A7,XCMPLX_1:198;
         (-(q`2/|.q.|-sn))/(1+sn)<1 by A7,A10,REAL_2:118;
       then ((-(q`2/|.q.|-sn))/(1+sn))^2<1^2 by A12,JGRAPH_2:8;
       then 1-((-(q`2/|.q.|-sn))/(1+sn))^2>0 by SQUARE_1:11,59;
       then sqrt(1-((-(q`2/|.q.|-sn))/(1+sn))^2)>0 by SQUARE_1:93;
       then sqrt(1-(-(q`2/|.q.|-sn))^2/(1+sn)^2)> 0 by SQUARE_1:69;
       then sqrt(1-(q`2/|.q.|-sn)^2/(1+sn)^2)> 0 by SQUARE_1:61;
       then A13: sqrt(1-((q`2/|.q.|-sn)/(1+sn))^2)> 0 by SQUARE_1:69;
         ((q`2/|.q.|-sn)/(1+sn))<0 by A6,A7,REAL_2:128;
   hence p`1>0 & p`2<0 by A2,A4,A8,A13,SQUARE_1:21,24;
 end;

theorem Th115:
 for sn being Real,q1,q2 being Point of TOP-REAL 2 st -1<sn & sn<1 & q1`1>0
 & q1`2/|.q1.|>=sn & q2`1>0 & q2`2/|.q2.|>=sn & 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
   & q1`2/|.q1.|>=sn & q2`1>0 & q2`2/|.q2.|>=sn & q1`2/|.q1.|<q2`2/|.q2.|;
  let p1,p2 be Point of TOP-REAL 2;
   assume A2: p1=(sn-FanMorphE).q1 & p2=(sn-FanMorphE).q2;
    then p1=|[ |.q1.|*(sqrt(1-((q1`2/|.q1.|-sn)/(1-sn))^2)),
          |.q1.|* ((q1`2/|.q1.|-sn)/(1-sn))]| by A1,Th89;
    then A3: p1`1= |.q1.|*(sqrt(1-((q1`2/|.q1.|-sn)/(1-sn))^2))&
        p1`2= |.q1.|* ((q1`2/|.q1.|-sn)/(1-sn)) by EUCLID:56;
      p2=|[ |.q2.|*(sqrt(1-((q2`2/|.q2.|-sn)/(1-sn))^2)),
          |.q2.|* ((q2`2/|.q2.|-sn)/(1-sn))]| by A1,A2,Th89;
    then A4: p2`1= |.q2.|*(sqrt(1-((q2`2/|.q2.|-sn)/(1-sn))^2))&
        p2`2= |.q2.|* ((q2`2/|.q2.|-sn)/(1-sn)) by EUCLID:56;
    A5: |.q1.|>0 & |.q2.|>0 by A1,Lm1,JGRAPH_2:11;
    A6: |.p1.|=|.q1.| & |.p2.|=|.q2.| by A2,Th104;
    then A7: p1`2/|.p1.|= ((q1`2/|.q1.|-sn)/(1-sn)) by A3,A5,XCMPLX_1:90
;
    A8: p2`2/|.p2.|= ((q2`2/|.q2.|-sn)/(1-sn)) by A4,A5,A6,XCMPLX_1:90
;
    A9: q1`2/|.q1.|-sn< q2`2/|.q2.|-sn by A1,REAL_1:54;
      1-sn>0 by A1,Th3;
   hence thesis by A7,A8,A9,REAL_1:73;
 end;

theorem Th116:
 for sn being Real,q1,q2 being Point of TOP-REAL 2 st -1<sn & sn<1 & q1`1>0
 & q1`2/|.q1.|<sn & q2`1>0 & q2`2/|.q2.|<sn & 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
   & q1`2/|.q1.|<sn & q2`1>0 & q2`2/|.q2.|<sn & q1`2/|.q1.|<q2`2/|.q2.|;
  let p1,p2 be Point of TOP-REAL 2;
   assume A2: p1=(sn-FanMorphE).q1 & p2=(sn-FanMorphE).q2;
    then p1=|[ |.q1.|*(sqrt(1-((q1`2/|.q1.|-sn)/(1+sn))^2)),
          |.q1.|* ((q1`2/|.q1.|-sn)/(1+sn))]| by A1,Th90;
    then A3: p1`1= |.q1.|*(sqrt(1-((q1`2/|.q1.|-sn)/(1+sn))^2))&
        p1`2= |.q1.|* ((q1`2/|.q1.|-sn)/(1+sn)) by EUCLID:56;
      p2=|[ |.q2.|*(sqrt(1-((q2`2/|.q2.|-sn)/(1+sn))^2)),
          |.q2.|* ((q2`2/|.q2.|-sn)/(1+sn))]| by A1,A2,Th90;
    then A4: p2`1= |.q2.|*(sqrt(1-((q2`2/|.q2.|-sn)/(1+sn))^2))&
        p2`2= |.q2.|* ((q2`2/|.q2.|-sn)/(1+sn)) by EUCLID:56;
    A5: |.q1.|>0 & |.q2.|>0 by A1,Lm1,JGRAPH_2:11;
    A6: |.p1.|=|.q1.| & |.p2.|=|.q2.| by A2,Th104;
    then A7: p1`2/|.p1.|= (q1`2/|.q1.|-sn)/(1+sn) by A3,A5,XCMPLX_1:90
;
    A8: p2`2/|.p2.|= (q2`2/|.q2.|-sn)/(1+sn) by A4,A5,A6,XCMPLX_1:90;
    A9: q1`2/|.q1.|-sn< q2`2/|.q2.|-sn by A1,REAL_1:54;
      1+sn>0 by A1,Th3;
   hence p1`2/|.p1.|<p2`2/|.p2.| by A7,A8,A9,REAL_1:73;
 end;

theorem
   for sn being Real,q1,q2 being Point of TOP-REAL 2 st -1<sn & sn<1 & q1`1>0
 & q2`1>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`2/|.q1.|<q2`2/|.q2.|;
  let p1,p2 be Point of TOP-REAL 2;
   assume A2: p1=(sn-FanMorphE).q1 & p2=(sn-FanMorphE).q2;
      now per cases;
    case q1`2/|.q1.|>=sn & q2`2/|.q2.|>=sn;
     hence p1`2/|.p1.|<p2`2/|.p2.| by A1,A2,Th115;
    case q1`2/|.q1.|>=sn & q2`2/|.q2.|<sn;
     hence p1`2/|.p1.|<p2`2/|.p2.| by A1,AXIOMS:22;
    case A3: q1`2/|.q1.|<sn & q2`2/|.q2.|>=sn;
       then A4: p1`1>0 & p1`2<0 by A1,A2,Th114;
       then A5: |.p1.|>0 by Lm1,JGRAPH_2:11;
       A6: p2`1>0 & p2`2>=0 by A1,A2,A3,Th113;
       then |.p2.|>0 by Lm1,JGRAPH_2:11;
       then p2`2/|.p2.|>=0 by A6,REAL_2:125;
     hence p1`2/|.p1.|<p2`2/|.p2.| by A4,A5,REAL_2:128;
    case q1`2/|.q1.|<sn & q2`2/|.q2.|<sn;
     hence p1`2/|.p1.|<p2`2/|.p2.| by A1,A2,Th116;
    end;
   hence p1`2/|.p1.|<p2`2/|.p2.|;
 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 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;
  A2: |.q.|>=0 by TOPRNS_1:26;
  let p be Point of TOP-REAL 2;
   assume p=(sn-FanMorphE).q;
    then A3: p=|[ |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)),
          |.q.|* ((q`2/|.q.|-sn)/(1-sn))]| by A1,Th89;
     A4: |.q.|<>0 by A1,JGRAPH_2:11,TOPRNS_1:25;
     A5: (q`2/|.q.|-sn)=0 by A1,XCMPLX_1:14;
     A6: p`1= |.q.|*(sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2))&
        p`2= |.q.|* ((q`2/|.q.|-sn)/(1-sn)) by A3,EUCLID:56;
         ((-(q`2/|.q.|-sn))/(1-sn))^2<1^2 by A5,JGRAPH_2:8,XCMPLX_1:198;
       then 1-((-(q`2/|.q.|-sn))/(1-sn))^2>0 by SQUARE_1:11,59;
       then sqrt(1-((-(q`2/|.q.|-sn))/(1-sn))^2)>0 by SQUARE_1:93;
       then sqrt(1-(-(q`2/|.q.|-sn))^2/(1-sn)^2)> 0 by SQUARE_1:69;
       then sqrt(1-(q`2/|.q.|-sn)^2/(1-sn)^2)> 0 by SQUARE_1:61;
       then sqrt(1-((q`2/|.q.|-sn)/(1-sn))^2)> 0 by SQUARE_1:69;
   hence thesis by A2,A4,A5,A6,SQUARE_1:21;
 end;

theorem
   for sn being real number holds 0.REAL 2=(sn-FanMorphE).(0.REAL 2)
   by Th89,JGRAPH_2:11;

begin :: Fan Morphism for South

definition let s be real number, q be Point of TOP-REAL 2;
  func FanS(s,q) -> Point of TOP-REAL 2 equals
  :Def8: |.q.|*|[(q`1/|.q.|-s)/(1-s),
     -sqrt(1-((q`1/|.q.|-s)/(1-s))^2)]| if q`1/|.q.|>=s & q`2<0,
       |.q.|*|[(q`1/|.q.|-s)/(1+s),
     -sqrt(1-((q`1/|.q.|-s)/(1+s))^2)]| if q`1/|.q.|<s & q`2<0
   otherwise q;
 correctness;
end;

definition let c be real number;
 func c-FanMorphS -> Function of the carrier of TOP-REAL 2,
 the carrier of TOP-REAL 2 means
  :Def9: for q being Point of TOP-REAL 2 holds
       it.q=FanS(c,q);
 existence
  proof
    deffunc F(Point of TOP-REAL 2)=FanS(c,$1);
    thus ex IT being Function of the carrier of TOP-REAL 2,
 the carrier of TOP-REAL 2 st
  for q being Point of TOP-REAL 2 holds
       IT.q=F(q)from LambdaD;
  end;
 uniqueness
  proof
    deffunc F(Point of TOP-REAL 2)=FanS(c,$1);
    thus for a,b being Function of the carrier of TOP-REAL 2,
 the carrier of TOP-REAL 2 st
  (for q being Point of TOP-REAL 2 holds a.q=F(q)) &
  (for q being Point of TOP-REAL 2 holds b.q=F(q)) holds
  a=b from FuncDefUniq;
  end;
end;

theorem Th120: for cn being real number holds
  (q`1/|.q.|>=cn & q`2<0 implies
    cn-FanMorphS.q= |[ |.q.|* ((q`1/|.q.|-cn)/(1-cn)),
    |.q.|*(-sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))]|)&
  (q`2>=0 implies cn-FanMorphS.q=q)
proof let cn be real number;
  A1: cn-FanMorphS.q=FanS(cn,q) by Def9;
   hereby assume q`1/|.q.|>=cn & q`2<0;
    then FanS(cn,q)= |.q.|*|[(q`1/|.q.|-cn)/(1-cn),
         -sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2)]| by Def8
          .= |[ |.q.|* ((q`1/|.q.|-cn)/(1-cn)),
          |.q.|*(-sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))]| by EUCLID:62;
     hence cn-FanMorphS.q= |[ |.q.|* ((q`1/|.q.|-cn)/(1-cn)),
    |.q.|*(-sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))]| by Def9;
   end;
   assume q`2>=0; hence thesis by A1,Def8;
end;

theorem Th121: for cn being Real holds
  (q`1/|.q.|<=cn & q`2<0 implies
    cn-FanMorphS.q= |[ |.q.|*((q`1/|.q.|-cn)/(1+cn)),
       |.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))]|)
proof let cn be Real;
   assume A1: q`1/|.q.|<=cn & q`2<0;
     now per cases by A1,REAL_1:def 5;
   case q`1/|.q.|<cn;
    then FanS(cn,q)= |.q.|*|[(q`1/|.q.|-cn)/(1+cn),
        -sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2)]| by A1,Def8
          .= |[ |.q.|* ((q`1/|.q.|-cn)/(1+cn)),
           |.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))]| by EUCLID:62;
    hence cn-FanMorphS.q= |[ |.q.|*((q`1/|.q.|-cn)/(1+cn)),
      |.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))]| by Def9;
   case A2: q`1/|.q.|=cn;
    then A3: q`1/|.q.|-cn=0 by XCMPLX_1:14;
    then (q`1/|.q.|-cn)/(1-cn)=0;
    hence cn-FanMorphS.q= |[ |.q.|*((q`1/|.q.|-cn)/(1+cn)),
        |.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))]| by A1,A2,A3,Th120;
   end;
   hence thesis;
end;

theorem Th122: for cn being Real st -1<cn & cn<1 holds
  (q`1/|.q.|>=cn & q`2<=0 & q<>0.REAL 2 implies
    cn-FanMorphS.q= |[ |.q.|* ((q`1/|.q.|-cn)/(1-cn)),
        |.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))]|) &
  (q`1/|.q.|<=cn & q`2<=0 & q<>0.REAL 2 implies
    cn-FanMorphS.q= |[ |.q.|*((q`1/|.q.|-cn)/(1+cn)),
        |.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))]|)
proof let cn be Real;
 assume A1: -1<cn & cn<1;
     now per cases;
   case A2: q`1/|.q.|>=cn & q`2<=0 & q<>0.REAL 2;
      now per cases;
    case A3: q`2<0;
      then FanS(cn,q)= |.q.|*|[(q`1/|.q.|-cn)/(1-cn),
          -sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2)]| by A2,Def8
          .= |[|.q.|* ((q`1/|.q.|-cn)/(1-cn)),
          |.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))]| by EUCLID:62;
     hence thesis by A3,Def9,Th121;
    case A4: q`2>=0;
     then A5: cn-FanMorphS.q=q by Th120;
     A6: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
     A7: |.q.|<>0 by A2,TOPRNS_1:25;
     then A8: |.q.|^2>0 by SQUARE_1:74;
     A9: q`2=0 by A2,A4;
     A10: |.q.|>=0 by TOPRNS_1:26;
     then A11: sqrt((|.q.|)^2)=|.q.| by SQUARE_1:89;
     A12: 1-cn>0 by A1,Th3;
       (q`1)^2/|.q.|^2=1^2 by A6,A8,A9,SQUARE_1:59,60,XCMPLX_1:60;
      then ((q`1)/|.q.|)^2=1^2 by SQUARE_1:69;
      then A13: sqrt(((q`1)/|.q.|)^2)=1 by SQUARE_1:89;
       now assume q`1<0;
       then q`1/|.q.|<0 by A7,A10,REAL_2:128;
       then -((q`1)/|.q.|)=1 by A13,SQUARE_1:90;
      hence contradiction by A1,A2;
     end;
     then A14: |.q.|=q`1 by A6,A9,A11,SQUARE_1:60,89;
     then 1=q`1/|.q.| by A7,XCMPLX_1:60;
     then (q`1/|.q.|-cn)/(1-cn)=1 by A12,XCMPLX_1:60;
     hence thesis by A1,A5,A7,A9,A14,EUCLID:57,SQUARE_1:59,82,XCMPLX_1:60;
    end;
    hence thesis;
   case A15: q`1/|.q.|<=cn & q`2<=0 & q<>0.REAL 2;
      now per cases;
    case q`2<0;
     hence thesis by Th120,Th121;
    case A16: q`2>=0;
     A17: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
     A18: |.q.|<>0 by A15,TOPRNS_1:25;
     A19: q`2=0 by A15,A16;
      A20: |.q.|>=0 by TOPRNS_1:26;
     A21: 1+cn>0 by A1,Th3;
       1-cn>0 by A1,Th3;
     then 1-cn+cn>0+cn by REAL_1:53;
     then 1>cn by XCMPLX_1:27;
     then 1>q`1/|.q.| by A15,AXIOMS:22;
     then (1)*(|.q.|)>q`1/|.q.|*(|.q.|) by A18,A20,REAL_1:70;
     then A22: (|.q.|)>q`1 by A18,XCMPLX_1:88;
     then A23: |.q.|=-q`1 by A17,A19,JGRAPH_3:1,SQUARE_1:60;
     A24: q`1= -(|.q.|) by A17,A19,A22,JGRAPH_3:1,SQUARE_1:60;
     then -1=q`1/|.q.| by A18,XCMPLX_1:198;
     then (q`1/|.q.|-cn)/(1+cn)
     =(-(1+cn))/(1+cn) by XCMPLX_1:161 .=-1 by A21,XCMPLX_1:198;
     then |[ |.q.|* ((q`1/|.q.|-cn)/(1+cn)),
       |.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))]|
      =|[ |.q.|*(-1), |.q.|*( -sqrt(1-1))]| by SQUARE_1:59,61
     .=|[ -(|.q.|), |.q.|*(0)]| by SQUARE_1:82,XCMPLX_1:180
     .=q by A19,A23,EUCLID:57;
     hence thesis by A1,A16,A18,A24,Th120,XCMPLX_1:198;
    end;
     hence thesis;
   case q`2>0 or q=0.REAL 2;
    hence thesis;
   end;
  hence thesis;
end;

theorem Th123:
 for cn being Real,K1 being non empty Subset of TOP-REAL 2,
f being map of (TOP-REAL 2)|K1,R^1 st -1<cn & cn<1 &
(for p being Point of (TOP-REAL 2) st p in the carrier of (TOP-REAL 2)|K1
  holds f.p=|.p.|* ((p`1/|.p.|-cn)/(1-cn)))
  & (for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds q`2<=0 & q<>0.REAL 2) holds
f is continuous
proof
 let cn be Real,K1 be non empty Subset of TOP-REAL 2,
f be map of (TOP-REAL 2)|K1,R^1;
 assume A1: -1<cn & cn<1 &
(for p being Point of (TOP-REAL 2) st p in the carrier of (TOP-REAL 2)|K1
  holds f.p=|.p.|* ((p`1/|.p.|-cn)/(1-cn)))
  & (for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds q`2<=0 & q<>0.REAL 2);
  set a=cn, b=(1-cn);
  A2: b>0 by A1,Th3;
  reconsider g2=proj1|K1 as continuous map of (TOP-REAL 2)|K1,R^1 by Lm2;
  reconsider g1=(2 NormF)|K1 as continuous map of (TOP-REAL 2)|K1,R^1
                          by Lm5;
    for q being Point of TOP-REAL 2 st q in
    the carrier of (TOP-REAL 2)|K1 holds q<>0.REAL 2 by A1;
  then for q being Point of (TOP-REAL 2)|K1 holds g1.q<>0 by Lm6;
  then consider g3 being map of (TOP-REAL 2)|K1,R^1 such that
  A3: (for q being Point of
  (TOP-REAL 2)|K1,r1,r2 being Real st g2.q=r1 & g1.q=r2
  holds g3.q=r2*((r1/r2-a)/b)) & g3 is continuous by A2,Th10;
  A4:dom g3=the carrier of (TOP-REAL 2)|K1 by FUNCT_2:def 1;
  then A5:dom f=dom g3 by FUNCT_2:def 1;
    for x being set st x in dom f holds f.x=g3.x
   proof let x be set;assume
     A6: x in dom f;
     then x in K1 by A4,A5,JORDAN1:1;
     then reconsider r=x as Point of (TOP-REAL 2);
     reconsider s=x as Point of (TOP-REAL 2)|K1 by A6,FUNCT_2:def 1;
     A7:f.r=(|.r.|)* ((r`1/|.r.|-cn)/(1-cn)) by A1,A4,A5,A6;
     A8:g2.s=proj1.s by Lm2;
     A9:g1.s=(2 NormF).s by Lm5;
     A10:proj1.r=r`1 by PSCOMP_1:def 28;
       (2 NormF).r=|.r.| by Def1;
    hence f.x=g3.x by A3,A7,A8,A9,A10;
   end;
hence f is continuous by A3,A5,FUNCT_1:9;
end;

theorem Th124:
 for cn being Real,K1 being non empty Subset of TOP-REAL 2,
f being map of (TOP-REAL 2)|K1,R^1 st -1<cn & cn<1 &
(for p being Point of (TOP-REAL 2) st p in the carrier of (TOP-REAL 2)|K1
  holds f.p=|.p.|* ((p`1/|.p.|-cn)/(1+cn)))
  & (for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds q`2<=0 & q<>0.REAL 2) holds
f is continuous
proof
 let cn be Real,K1 be non empty Subset of TOP-REAL 2,
f be map of (TOP-REAL 2)|K1,R^1;
 assume A1: -1<cn & cn<1 &
(for p being Point of (TOP-REAL 2) st p in the carrier of (TOP-REAL 2)|K1
  holds f.p=|.p.|* ((p`1/|.p.|-cn)/(1+cn)))
  & (for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds q`2<=0 & q<>0.REAL 2);
  set a=cn, b=(1+cn);
  A2: 1+cn>0 by A1,Th3;
  reconsider g2=proj1|K1 as continuous map of (TOP-REAL 2)|K1,R^1 by Lm2;
  reconsider g1=(2 NormF)|K1 as continuous map of (TOP-REAL 2)|K1,R^1
                          by Lm5;
    for q being Point of TOP-REAL 2 st q in
    the carrier of (TOP-REAL 2)|K1 holds q<>0.REAL 2 by A1;
  then for q being Point of (TOP-REAL 2)|K1 holds g1.q<>0 by Lm6;
  then consider g3 being map of (TOP-REAL 2)|K1,R^1 such that
  A3: (for q being Point of
  (TOP-REAL 2)|K1,r1,r2 being Real st g2.q=r1 & g1.q=r2
  holds g3.q=r2*((r1/r2-a)/b)) & g3 is continuous by A2,Th10;
  A4:dom g3=the carrier of (TOP-REAL 2)|K1 by FUNCT_2:def 1;
  then A5:dom f=dom g3 by FUNCT_2:def 1;
    for x being set st x in dom f holds f.x=g3.x
   proof let x be set;assume
     A6: x in dom f; then A7: x in dom g3 by A4,FUNCT_2:def 1;
     then x in K1 by A4,JORDAN1:1;
     then reconsider r=x as Point of (TOP-REAL 2);
     reconsider s=x as Point of (TOP-REAL 2)|K1 by A6,FUNCT_2:def 1;
     A8:f.r=(|.r.|)* ((r`1/|.r.|-cn)/(1+cn)) by A1,A4,A7;
     A9:g2.s=proj1.s by Lm2;
     A10:g1.s=(2 NormF).s by Lm5;
     A11:proj1.r=r`1 by PSCOMP_1:def 28;
       (2 NormF).r=|.r.| by Def1;
    hence f.x=g3.x by A3,A8,A9,A10,A11;
   end;
hence f is continuous by A3,A5,FUNCT_1:9;
end;

theorem Th125:
 for cn being Real,K1 being non empty Subset of TOP-REAL 2,
f being map of (TOP-REAL 2)|K1,R^1 st
  -1<cn & cn<1 &
(for p being Point of (TOP-REAL 2) st p in the carrier of (TOP-REAL 2)|K1
  holds f.p=|.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2)))
  & (for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds
        q`2<=0 & q`1/|.q.|>=cn & q<>0.REAL 2) holds
f is continuous
proof
 let cn be Real,K1 be non empty Subset of TOP-REAL 2,
f be map of (TOP-REAL 2)|K1,R^1;
 assume A1:
  -1<cn & cn<1 &
(for p being Point of (TOP-REAL 2) st p in the carrier of (TOP-REAL 2)|K1
  holds f.p=|.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2)))
  & (for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds
            q`2<=0 & q`1/|.q.|>=cn & q<>0.REAL 2);
  set a=cn, b=(1-cn);
  A2: b>0 by A1,Th3;
  reconsider g2=proj1|K1 as continuous map of (TOP-REAL 2)|K1,R^1 by Lm2;
  reconsider g1=(2 NormF)|K1 as continuous map of (TOP-REAL 2)|K1,R^1
                          by Lm5;
    for q being Point of TOP-REAL 2 st q in
    the carrier of (TOP-REAL 2)|K1 holds q<>0.REAL 2 by A1;
  then for q being Point of (TOP-REAL 2)|K1 holds g1.q<>0 by Lm6;
  then consider g3 being map of (TOP-REAL 2)|K1,R^1 such that
  A3: (for q being Point of
  (TOP-REAL 2)|K1,r1,r2 being real number st g2.q=r1 & g1.q=r2
  holds g3.q=r2*( -sqrt(abs(1-((r1/r2-a)/b)^2))))
          & g3 is continuous by A2,Th14;
  A4:dom g3=the carrier of (TOP-REAL 2)|K1 by FUNCT_2:def 1;
  then A5:dom f=dom g3 by FUNCT_2:def 1;
    for x being set st x in dom f holds f.x=g3.x
   proof let x be set;assume
     A6: x in dom f;
     then x in K1 by A4,A5,JORDAN1:1;
     then reconsider r=x as Point of (TOP-REAL 2);
     reconsider s=x as Point of (TOP-REAL 2)|K1 by A6,FUNCT_2:def 1;
     A7: now assume |.r.|=0; then r=0.REAL 2 by TOPRNS_1:25;
      hence contradiction by A1,A4,A5,A6;
     end;
     A8: |.r.|>=0 by TOPRNS_1:26;
       |.r.|^2=(r`1)^2+(r`2)^2 by JGRAPH_3:10;
     then |.r.|^2 -(r`1)^2 =(r`2)^2 by XCMPLX_1:26;
     then (r`1)^2 -(|.r.|)^2 =-(r`2)^2 by XCMPLX_1:143;
     then A9: ((r`1) -(|.r.|))*((r`1)+|.r.|) =-(r`2)^2 by SQUARE_1:67;
       (r`2)^2>=0 by SQUARE_1:72;
     then -(r`2)^2 <=0 by REAL_1:66;
     then -(|.r.|)<=r`1 & r`1<= |.r.| by A8,A9,JGRAPH_3:4;
     then r`1/|.r.| <= |.r.|/|.r.| by A7,A8,REAL_1:73;
     then r`1/|.r.|<=1 by A7,XCMPLX_1:60;
     then A10: r`1/|.r.|-cn<=(1-cn) by REAL_1:49;
     A11: now assume (1-cn)^2=0;
       then 1-cn+cn=0+cn by SQUARE_1:73;
      hence contradiction by A1,XCMPLX_1:27;
     end;
     A12: 1-cn>0 by A1,Th3;
      A13: (1-cn)^2>=0 by SQUARE_1:72;
       cn<=r`1/(|.r.|) by A1,A4,A5,A6;
     then cn-r`1/|.r.|<=0 by REAL_2:106;
     then -(cn- r`1/|.r.|)>=-(1-cn) by A12,REAL_1:50;
     then -(1-cn)<= r`1/|.r.|-cn by XCMPLX_1:143;
     then (r`1/|.r.|-cn)^2<=(1-cn)^2 by A10,JGRAPH_2:7;
     then (r`1/|.r.|-cn)^2/(1-cn)^2<=(1-cn)^2/(1-cn)^2
                         by A11,A13,REAL_1:73;
     then (r`1/|.r.|-cn)^2/(1-cn)^2<=1 by A11,XCMPLX_1:60;
     then ((r`1/|.r.|-cn)/(1-cn))^2<=1 by SQUARE_1:69;
     then 1-((r`1/|.r.|-cn)/(1-cn))^2>=0 by SQUARE_1:12;
     then abs(1-((r`1/|.r.|-cn)/(1-cn))^2)
          =1-((r`1/|.r.|-cn)/(1-cn))^2 by ABSVALUE:def 1;
     then A14:f.r=(|.r.|)*( -sqrt(abs(1-((r`1/|.r.|-cn)/(1-cn))^2)))
                             by A1,A4,A5,A6;
     A15:g2.s=proj1.s by Lm2;
     A16:g1.s=(2 NormF).s by Lm5;
     A17:proj1.r=r`1 by PSCOMP_1:def 28;
       (2 NormF).r=|.r.| by Def1;
    hence f.x=g3.x by A3,A14,A15,A16,A17;
   end;
hence f is continuous by A3,A5,FUNCT_1:9;
end;

theorem Th126:
 for cn being Real,K1 being non empty Subset of TOP-REAL 2,
f being map of (TOP-REAL 2)|K1,R^1 st
  -1<cn & cn<1 &
(for p being Point of (TOP-REAL 2) st p in the carrier of (TOP-REAL 2)|K1
  holds f.p=|.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2)))
  & (for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds
        q`2<=0 & q`1/|.q.|<=cn & q<>0.REAL 2) holds
f is continuous
proof
 let cn be Real,K1 be non empty Subset of TOP-REAL 2,
f be map of (TOP-REAL 2)|K1,R^1;
 assume A1:
  -1<cn & cn<1 &
(for p being Point of (TOP-REAL 2) st p in the carrier of (TOP-REAL 2)|K1
  holds f.p=|.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2)))
  & (for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds
            q`2<=0 & q`1/|.q.|<=cn & q<>0.REAL 2);
  set a=cn, b=(1+cn);
  A2: 1+cn>0 by A1,Th3;
  reconsider g2=proj1|K1 as continuous map of (TOP-REAL 2)|K1,R^1 by Lm2;
  reconsider g1=(2 NormF)|K1 as continuous map of (TOP-REAL 2)|K1,R^1
                          by Lm5;
    for q being Point of TOP-REAL 2 st q in
    the carrier of (TOP-REAL 2)|K1 holds q<>0.REAL 2 by A1;
  then for q being Point of (TOP-REAL 2)|K1 holds g1.q<>0 by Lm6;
  then consider g3 being map of (TOP-REAL 2)|K1,R^1 such that
  A3: (for q being Point of
  (TOP-REAL 2)|K1,r1,r2 being real number st g2.q=r1 & g1.q=r2
  holds g3.q=r2*( -sqrt(abs(1-((r1/r2-a)/b)^2))))
          & g3 is continuous by A2,Th14;
  A4:dom g3=the carrier of (TOP-REAL 2)|K1 by FUNCT_2:def 1;
  then A5:dom f=dom g3 by FUNCT_2:def 1;
    for x being set st x in dom f holds f.x=g3.x
   proof let x be set;assume
     A6:x in dom f;
     then x in K1 by A4,A5,JORDAN1:1;
     then reconsider r=x as Point of (TOP-REAL 2);
     reconsider s=x as Point of (TOP-REAL 2)|K1 by A6,FUNCT_2:def 1;
     A7: now assume |.r.|=0; then r=0.REAL 2 by TOPRNS_1:25;
      hence contradiction by A1,A4,A5,A6;
     end;
     A8: |.r.|>=0 by TOPRNS_1:26;
       |.r.|^2=(r`1)^2+(r`2)^2 by JGRAPH_3:10;
     then |.r.|^2 -(r`1)^2 =(r`2)^2 by XCMPLX_1:26;
     then (r`1)^2 -(|.r.|)^2 =-(r`2)^2 by XCMPLX_1:143;
     then A9: ((r`1) -(|.r.|))*((r`1)+|.r.|) =-(r`2)^2 by SQUARE_1:67;
       (r`2)^2>=0 by SQUARE_1:72;
     then -(r`2)^2 <=0 by REAL_1:66;
     then -(|.r.|)<=r`1 & r`1<= |.r.| by A8,A9,JGRAPH_3:4;
     then r`1/|.r.| >= (-(|.r.|))/|.r.| by A7,A8,REAL_1:73;
     then r`1/|.r.|>= -1 by A7,XCMPLX_1:198;
     then r`1/|.r.|-cn>=-1-cn by REAL_1:49;
     then A10: r`1/|.r.|-cn>=-(1+cn) by XCMPLX_1:161;
     A11: (1+cn)^2>0 by A2,SQUARE_1:74;
       cn>=r`1/(|.r.|) by A1,A4,A5,A6;
     then cn-r`1/|.r.|>=0 by SQUARE_1:12;
     then -(cn-r`1/|.r.|)<=-0 by REAL_1:50;
     then r`1/|.r.|-cn<=-0 by XCMPLX_1:143;
     then r`1/|.r.|-cn<=1+cn by A2,AXIOMS:22;
     then (r`1/|.r.|-cn)^2<=(1+cn)^2 by A10,JGRAPH_2:7;
     then (r`1/|.r.|-cn)^2/(1+cn)^2<=(1+cn)^2/(1+cn)^2 by A11,REAL_1:73;
     then (r`1/|.r.|-cn)^2/(1+cn)^2<=1 by A11,XCMPLX_1:60;
     then ((r`1/|.r.|-cn)/(1+cn))^2<=1 by SQUARE_1:69;
     then 1-((r`1/|.r.|-cn)/(1+cn))^2>=0 by SQUARE_1:12;
     then abs(1-((r`1/|.r.|-cn)/(1+cn))^2)
          =1-((r`1/|.r.|-cn)/(1+cn))^2 by ABSVALUE:def 1;
     then A12:f.r=(|.r.|)*( -sqrt(abs(1-((r`1/|.r.|-cn)/(1+cn))^2))) by A1,A4,
A5,A6;
     A13:g2.s=proj1.s by Lm2;
     A14:g1.s=(2 NormF).s by Lm5;
     A15:proj1.r=r`1 by PSCOMP_1:def 28;
       (2 NormF).r=|.r.| by Def1;
    hence f.x=g3.x by A3,A12,A13,A14,A15;
   end;
hence f is continuous by A3,A5,FUNCT_1:9;
end;

theorem Th127: for cn being Real,
K0,B0 being Subset of TOP-REAL 2,f being
map of (TOP-REAL 2)|K0,(TOP-REAL 2)|B0
st -1<cn & cn<1 & f=(cn-FanMorphS)|K0 &
B0={q where q is Point of TOP-REAL 2: q`2<=0 & q<>0.REAL 2}
& K0={p: p`1/|.p.|>=cn & p`2<=0 & p<>0.REAL 2}
holds f is continuous
proof let cn be Real,K0,B0 be Subset of TOP-REAL 2,f be
map of (TOP-REAL 2)|K0,(TOP-REAL 2)|B0;
assume A1: -1<cn & cn<1 & f=(cn-FanMorphS)|K0 &
B0={q where q is Point of TOP-REAL 2: q`2<=0 & q<>0.REAL 2} &
K0={p: p`1/|.p.|>=cn & p`2<=0 & p<>0.REAL 2};
 set sn=-sqrt(1-cn^2);
 set p0=|[cn,sn]|;
 A2: p0`1=cn by EUCLID:56;
 A3: p0`2=sn by EUCLID:56;
     cn^2<1^2 by A1,JGRAPH_2:8;
 then A4: 1-cn^2>0 by SQUARE_1:11,59;
 A5: now assume p0=0.REAL 2;
   then --sn=-0 by EUCLID:56,JGRAPH_2:11;
  hence contradiction by A4,SQUARE_1:93;
 end;
   -sn>0 by A4,SQUARE_1:93;
 then A6: sn<0 by REAL_1:66;
 A7: |.p0.|=sqrt((sn)^2+cn^2) by A2,A3,JGRAPH_3:10;
   (-sn)^2=1-cn^2 by A4,SQUARE_1:def 4;
 then (sn)^2=1-cn^2 by SQUARE_1:61;
 then |.p0.|=1 by A7,SQUARE_1:83,XCMPLX_1:27;
 then p0`1/|.p0.|=cn by EUCLID:56;
 then A8: p0 in K0 by A1,A3,A5,A6;
then reconsider K1=K0 as non empty Subset of TOP-REAL 2;
  A9: K0 c= B0
   proof let x be set;assume x in K0;
      then consider p8 being Point of TOP-REAL 2 such that
      A10: x=p8 & (
      p8`1/|.p8.|>=cn & p8`2<=0 & p8<>0.REAL 2) by A1;
    thus x in B0 by A1,A10;
   end;
A11:dom ((proj1)*((cn-FanMorphS)|K1)) c= dom ((cn-FanMorphS)|K1)
                                  by RELAT_1:44;
 dom ((cn-FanMorphS)|K1) c= dom ((proj1)*((cn-FanMorphS)|K1))
proof let x be set;assume A12:x in dom ((cn-FanMorphS)|K1);
   then A13:x in dom (cn-FanMorphS) /\ K1 by FUNCT_1:68;
   A14:((cn-FanMorphS)|K1).x=(cn-FanMorphS).x by A12,FUNCT_1:68;
   A15: dom proj1 = (the carrier of TOP-REAL 2) by FUNCT_2:def 1;
     x in dom (cn-FanMorphS) by A13,XBOOLE_0:def 3;
   then (cn-FanMorphS).x in rng (cn-FanMorphS) by FUNCT_1:12;
  hence x in dom ((proj1)*((cn-FanMorphS)|K1)) by A12,A14,A15,FUNCT_1:21;
end;
then A16:dom ((proj1)*((cn-FanMorphS)|K1))
=dom ((cn-FanMorphS)|K1) by A11,XBOOLE_0:def 10
.=dom (cn-FanMorphS) /\ K1 by FUNCT_1:68
.=(the carrier of TOP-REAL 2)/\ K1 by FUNCT_2:def 1
.=K1 by XBOOLE_1:28
.=the carrier of (TOP-REAL 2)|K1 by JORDAN1:1;
A17:rng ((proj1)*((cn-FanMorphS)|K1)) c= rng (proj1) by RELAT_1:45;
  rng proj1 c= the carrier of R^1 by JGRAPH_2:16,RELSET_1:12;
then rng ((proj1)*((cn-FanMorphS)|K1)) c= the carrier of R^1 by A17,XBOOLE_1:1
;
then (proj1)*((cn-FanMorphS)|K1) is Function of the carrier of (TOP-REAL 2)|K1
,
the carrier of R^1 by A16,FUNCT_2:4;
then reconsider g2=(proj1)*((cn-FanMorphS)|K1) as map of (TOP-REAL 2)|K1,R^1
                                ;
A18:dom ((proj2)*((cn-FanMorphS)|K1))
                     c= dom ((cn-FanMorphS)|K1) by RELAT_1:44;
   dom ((cn-FanMorphS)|K1) c= dom ((proj2)*((cn-FanMorphS)|K1))
proof let x be set;assume A19:x in dom ((cn-FanMorphS)|K1);
   then A20:x in dom (cn-FanMorphS) /\ K1 by FUNCT_1:68;
   A21:((cn-FanMorphS)|K1).x=(cn-FanMorphS).x by A19,FUNCT_1:68;
   A22: dom proj2 = (the carrier of TOP-REAL 2) by FUNCT_2:def 1;
     x in dom (cn-FanMorphS) by A20,XBOOLE_0:def 3;
   then (cn-FanMorphS).x in rng (cn-FanMorphS) by FUNCT_1:12;
  hence x in dom ((proj2)*((cn-FanMorphS)|K1)) by A19,A21,A22,FUNCT_1:21;
end;
then A23:dom ((proj2)*((cn-FanMorphS)|K1))
=dom ((cn-FanMorphS)|K1) by A18,XBOOLE_0:def 10
.=dom (cn-FanMorphS) /\ K1 by FUNCT_1:68
.=((the carrier of TOP-REAL 2))/\ K1 by FUNCT_2:def 1
.=K1 by XBOOLE_1:28
.=the carrier of (TOP-REAL 2)|K1 by JORDAN1:1;
A24:rng ((proj2)*((cn-FanMorphS)|K1)) c= rng (proj2) by RELAT_1:45;
  rng proj2 c= the carrier of R^1 by JGRAPH_2:17,RELSET_1:12;
then rng ((proj2)*((cn-FanMorphS)|K1)) c= the carrier of R^1 by A24,XBOOLE_1:1
;
then (proj2)*((cn-FanMorphS)|K1) is Function of the carrier of (TOP-REAL 2)|K1
,
the carrier of R^1 by A23,FUNCT_2:4;
then reconsider g1=(proj2)*((cn-FanMorphS)|K1) as map of (TOP-REAL 2)|K1,R^1
                                ;
A25: for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds q`2<=0 & q<>0.REAL 2
   proof let q be Point of TOP-REAL 2;
    assume A26:q in the carrier of (TOP-REAL 2)|K1;
        the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
     then consider p3 being Point of TOP-REAL 2 such that
     A27: q=p3 & (p3`1/|.p3.|>=cn & p3`2<=0 & p3<>0.REAL 2) by A1,A26;
    thus q`2<=0 & q<>0.REAL 2 by A27;
   end;
    for p being Point of TOP-REAL 2 st p in the carrier of (TOP-REAL 2)|K1
  holds g2.p=|.p.|* ((p`1/|.p.|-cn)/(1-cn))
   proof let p be Point of TOP-REAL 2;
    assume A28: p in the carrier of (TOP-REAL 2)|K1;
     A29:dom ((cn-FanMorphS)|K1)=dom (cn-FanMorphS) /\ K1 by FUNCT_1:68
     .=((the carrier of TOP-REAL 2))/\ K1 by FUNCT_2:def 1
     .=K1 by XBOOLE_1:28;
     A30:the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
     then consider p3 being Point of TOP-REAL 2 such that
     A31: p=p3 & (p3`1/|.p3.|>=cn & p3`2<=0 & p3<>0.REAL 2) by A1,A28;
     A32:(cn-FanMorphS).p
     =|[ |.p.|* ((p`1/|.p.|-cn)/(1-cn)),
        |.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2))]| by A1,A31,Th122;
      ((cn-FanMorphS)|K1).p=(cn-FanMorphS).p by A28,A30,FUNCT_1:72;
     then g2.p=(proj1).(|[ |.p.|* ((p`1/|.p.|-cn)/(1-cn)),
        |.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2))]|)
                             by A28,A29,A30,A32,FUNCT_1:23
      .=(|[ |.p.|* ((p`1/|.p.|-cn)/(1-cn)),
        |.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2))]|)`1 by PSCOMP_1:def 28
      .=|.p.|* ((p`1/|.p.|-cn)/(1-cn)) by EUCLID:56;
    hence g2.p=|.p.|* ((p`1/|.p.|-cn)/(1-cn));
   end;
then consider f2 being map of (TOP-REAL 2)|K1,R^1 such that
  A33:for p being Point of TOP-REAL 2 st p in the carrier of (TOP-REAL 2)|K1
  holds f2.p=|.p.|* ((p`1/|.p.|-cn)/(1-cn));
A34:f2 is continuous by A1,A25,A33,Th123;
    for p being Point of TOP-REAL 2 st p in the carrier of (TOP-REAL 2)|K1
  holds g1.p=|.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2))
   proof let p be Point of TOP-REAL 2;
    assume A35: p in the carrier of (TOP-REAL 2)|K1;
     A36:dom ((cn-FanMorphS)|K1)=dom (cn-FanMorphS) /\ K1 by FUNCT_1:68
     .=((the carrier of TOP-REAL 2))/\ K1 by FUNCT_2:def 1
       .=K1 by XBOOLE_1:28;
     A37:the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
     then consider p3 being Point of TOP-REAL 2 such that
     A38: p=p3 & (p3`1/|.p3.|>=cn & p3`2<=0 & p3<>0.REAL 2) by A1,A35;
     A39:(cn-FanMorphS).p=|[ |.p.|* ((p`1/|.p.|-cn)/(1-cn)),
       |.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2))]| by A1,A38,Th122;
      ((cn-FanMorphS)|K1).p=(cn-FanMorphS).p by A35,A37,FUNCT_1:72;
     then g1.p=(proj2).(|[ |.p.|* ((p`1/|.p.|-cn)/(1-cn)),
          |.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2))]|)
                             by A35,A36,A37,A39,FUNCT_1:23
        .=(|[ |.p.|* ((p`1/|.p.|-cn)/(1-cn)),
          |.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2))]|)`2 by PSCOMP_1:def 29
        .= |.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2)) by EUCLID:56;
    hence g1.p=|.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2));
   end;
then consider f1 being map of (TOP-REAL 2)|K1,R^1 such that
  A40:for p being Point of TOP-REAL 2 st p in the carrier of (TOP-REAL 2)|K1
  holds f1.p=|.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2));
    for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds q`2<=0 & q`1/|.q.|>=cn & q<>0.REAL 2
   proof let q be Point of TOP-REAL 2;
    assume A41:q in the carrier of (TOP-REAL 2)|K1;
        the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
     then consider p3 being Point of TOP-REAL 2 such that
     A42: q=p3 & (p3`1/|.p3.|>=cn & p3`2<=0 & p3<>0.REAL 2) by A1,A41;
    thus thesis by A42;
   end;
then A43:f1 is continuous by A1,A40,Th125;
  for x,y,s,r being real number st |[x,y]| in K1 &
  s=f2.(|[x,y]|) & r=f1.(|[x,y]|) holds f.(|[x,y]|)=|[s,r]|
  proof let x,y,s,r be real number;
   assume A44: |[x,y]| in K1 & s=f2.(|[x,y]|) & r=f1.(|[x,y]|);
     set p99=|[x,y]|;
     A45:the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
     consider p3 being Point of TOP-REAL 2 such that
     A46: p99=p3 & (p3`1/|.p3.|>=cn & p3`2<=0 & p3<>0.REAL 2) by A1,A44;
     A47:f1.p99=|.p99.|*( -sqrt(1-((p99`1/|.p99.|-cn)/(1-cn))^2))
                          by A40,A44,A45;
      ((cn-FanMorphS)|K0).(|[x,y]|)=((cn-FanMorphS)).(|[x,y]|) by A44,FUNCT_1:
72
    .= |[ |.p99.|* ((p99`1/|.p99.|-cn)/(1-cn)),
      |.p99.|*( -sqrt(1-((p99`1/|.p99.|-cn)/(1-cn))^2))]| by A1,A46,Th122
    .=|[s,r]| by A33,A44,A45,A47;
   hence f.(|[x,y]|)=|[s,r]| by A1;
  end;
hence f is continuous by A8,A9,A34,A43,JGRAPH_2:45;
end;

theorem Th128: for cn being Real,
K0,B0 being Subset of TOP-REAL 2,f being
map of (TOP-REAL 2)|K0,(TOP-REAL 2)|B0
st -1<cn & cn<1 & f=(cn-FanMorphS)|K0 &
B0={q where q is Point of TOP-REAL 2: q`2<=0 & q<>0.REAL 2}
& K0={p: p`1/|.p.|<=cn & p`2<=0 & p<>0.REAL 2}
holds f is continuous
proof let cn be Real,K0,B0 be Subset of TOP-REAL 2,f be
map of (TOP-REAL 2)|K0,(TOP-REAL 2)|B0;
assume A1: -1<cn & cn<1 & f=(cn-FanMorphS)|K0 &
B0={q where q is Point of TOP-REAL 2: q`2<=0 & q<>0.REAL 2} &
K0={p: p`1/|.p.|<=cn & p`2<=0 & p<>0.REAL 2};
 set sn=-sqrt(1-cn^2);
 set p0=|[cn,sn]|;
 A2: p0`1=cn by EUCLID:56;
 A3: p0`2=sn by EUCLID:56;
     cn^2<1^2 by A1,JGRAPH_2:8;
 then A4: 1-cn^2>0 by SQUARE_1:11,59;
 A5: now assume p0=0.REAL 2;
   then --sn=-0 by EUCLID:56,JGRAPH_2:11;
  hence contradiction by A4,SQUARE_1:93;
 end;
   -sn>0 by A4,SQUARE_1:93;
 then A6: sn<0 by REAL_1:66;
 A7: |.p0.|=sqrt((sn)^2+cn^2) by A2,A3,JGRAPH_3:10;
   (-sn)^2=1-cn^2 by A4,SQUARE_1:def 4;
 then sn^2=1-cn^2 by SQUARE_1:61;
 then |.p0.|=1 by A7,SQUARE_1:83,XCMPLX_1:27;
 then p0`1/|.p0.|=cn by EUCLID:56;
 then A8: p0 in K0 by A1,A3,A5,A6;
 then reconsider K1=K0 as non empty Subset of TOP-REAL 2;
  A9: K0 c= B0
   proof let x be set;assume x in K0;
      then consider p8 being Point of TOP-REAL 2 such that
      A10: x=p8 & (p8`1/|.p8.|<=cn & p8`2<=0 & p8<>0.REAL 2) by A1;
    thus x in B0 by A1,A10;
   end;
A11:dom ((proj1)*((cn-FanMorphS)|K1)) c= dom ((cn-FanMorphS)|K1)
                                  by RELAT_1:44;
 dom ((cn-FanMorphS)|K1) c= dom ((proj1)*((cn-FanMorphS)|K1))
proof let x be set;assume A12:x in dom ((cn-FanMorphS)|K1);
   then A13:x in dom (cn-FanMorphS) /\ K1 by FUNCT_1:68;
   A14:((cn-FanMorphS)|K1).x=(cn-FanMorphS).x by A12,FUNCT_1:68;
   A15: dom proj1 = (the carrier of TOP-REAL 2) by FUNCT_2:def 1;
     x in dom (cn-FanMorphS) by A13,XBOOLE_0:def 3;
   then (cn-FanMorphS).x in rng (cn-FanMorphS) by FUNCT_1:12;
  hence x in dom ((proj1)*((cn-FanMorphS)|K1)) by A12,A14,A15,FUNCT_1:21;
end;
then A16:dom ((proj1)*((cn-FanMorphS)|K1))
=dom ((cn-FanMorphS)|K1) by A11,XBOOLE_0:def 10
.=dom (cn-FanMorphS) /\ K1 by FUNCT_1:68
.=(the carrier of TOP-REAL 2)/\ K1 by FUNCT_2:def 1
.=K1 by XBOOLE_1:28
.=the carrier of (TOP-REAL 2)|K1 by JORDAN1:1;
A17:rng ((proj1)*((cn-FanMorphS)|K1)) c= rng (proj1) by RELAT_1:45;
  rng proj1 c= the carrier of R^1 by JGRAPH_2:16,RELSET_1:12;
then rng ((proj1)*((cn-FanMorphS)|K1)) c= the carrier of R^1 by A17,XBOOLE_1:1
;
then (proj1)*((cn-FanMorphS)|K1) is Function of the carrier of (TOP-REAL 2)|K1
,
the carrier of R^1 by A16,FUNCT_2:4;
then reconsider g2=(proj1)*((cn-FanMorphS)|K1) as map of (TOP-REAL 2)|K1,R^1
                                ;
A18:dom ((proj2)*((cn-FanMorphS)|K1))
                     c= dom ((cn-FanMorphS)|K1) by RELAT_1:44;
 dom ((cn-FanMorphS)|K1) c= dom ((proj2)*((cn-FanMorphS)|K1))
proof let x be set;assume A19:x in dom ((cn-FanMorphS)|K1);
   then A20:x in dom (cn-FanMorphS) /\ K1 by FUNCT_1:68;
   A21:((cn-FanMorphS)|K1).x=(cn-FanMorphS).x by A19,FUNCT_1:68;
   A22: dom proj2 = (the carrier of TOP-REAL 2) by FUNCT_2:def 1;
     x in dom (cn-FanMorphS) by A20,XBOOLE_0:def 3;
   then (cn-FanMorphS).x in rng (cn-FanMorphS) by FUNCT_1:12;
  hence x in dom ((proj2)*((cn-FanMorphS)|K1)) by A19,A21,A22,FUNCT_1:21;
end;
then A23:dom ((proj2)*((cn-FanMorphS)|K1))
=dom ((cn-FanMorphS)|K1) by A18,XBOOLE_0:def 10
.=dom (cn-FanMorphS) /\ K1 by FUNCT_1:68
.=((the carrier of TOP-REAL 2))/\ K1 by FUNCT_2:def 1
.=K1 by XBOOLE_1:28
.=the carrier of (TOP-REAL 2)|K1 by JORDAN1:1;
A24:rng ((proj2)*((cn-FanMorphS)|K1)) c= rng (proj2) by RELAT_1:45;
  rng proj2 c= the carrier of R^1 by JGRAPH_2:17,RELSET_1:12;
then rng ((proj2)*((cn-FanMorphS)|K1)) c= the carrier of R^1 by A24,XBOOLE_1:1
;
then (proj2)*((cn-FanMorphS)|K1) is Function of the carrier of (TOP-REAL 2)|K1
,
the carrier of R^1 by A23,FUNCT_2:4;
then reconsider g1=(proj2)*((cn-FanMorphS)|K1) as map of (TOP-REAL 2)|K1,R^1
                                ;
A25: for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds q`2<=0 & q<>0.REAL 2
   proof let q be Point of TOP-REAL 2;
    assume A26:q in the carrier of (TOP-REAL 2)|K1;
        the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
     then consider p3 being Point of TOP-REAL 2 such that
     A27: q=p3 & (p3`1/|.p3.|<=cn & p3`2<=0 & p3<>0.REAL 2) by A1,A26;
    thus q`2<=0 & q<>0.REAL 2 by A27;
   end;
    for p being Point of TOP-REAL 2 st p in the carrier of (TOP-REAL 2)|K1
  holds g2.p=|.p.|* ((p`1/|.p.|-cn)/(1+cn))
   proof let p be Point of TOP-REAL 2;
    assume A28: p in the carrier of (TOP-REAL 2)|K1;
     A29:dom ((cn-FanMorphS)|K1)=dom (cn-FanMorphS) /\ K1 by FUNCT_1:68
     .=((the carrier of TOP-REAL 2))/\ K1 by FUNCT_2:def 1
     .=K1 by XBOOLE_1:28;
     A30:the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
     then consider p3 being Point of TOP-REAL 2 such that
     A31: p=p3 & (p3`1/|.p3.|<=cn & p3`2<=0 & p3<>0.REAL 2) by A1,A28;
     A32:(cn-FanMorphS).p
     =|[ |.p.|* ((p`1/|.p.|-cn)/(1+cn)),
       |.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2))]| by A1,A31,Th122;
      ((cn-FanMorphS)|K1).p=(cn-FanMorphS).p by A28,A30,FUNCT_1:72;
     then g2.p=(proj1).(|[ |.p.|* ((p`1/|.p.|-cn)/(1+cn)),
        |.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2))]|)
                             by A28,A29,A30,A32,FUNCT_1:23
      .=(|[ |.p.|* ((p`1/|.p.|-cn)/(1+cn)),
       |.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2))]|)`1 by PSCOMP_1:def 28
      .=|.p.|* ((p`1/|.p.|-cn)/(1+cn)) by EUCLID:56;
    hence g2.p=|.p.|* ((p`1/|.p.|-cn)/(1+cn));
   end;
then consider f2 being map of (TOP-REAL 2)|K1,R^1 such that
  A33:for p being Point of TOP-REAL 2 st p in the carrier of (TOP-REAL 2)|K1
  holds f2.p=|.p.|* ((p`1/|.p.|-cn)/(1+cn));
A34:f2 is continuous by A1,A25,A33,Th124;
    for p being Point of TOP-REAL 2 st p in the carrier of (TOP-REAL 2)|K1
  holds g1.p=|.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2))
   proof let p be Point of TOP-REAL 2;
    assume A35: p in the carrier of (TOP-REAL 2)|K1;
     A36:dom ((cn-FanMorphS)|K1)=dom (cn-FanMorphS) /\ K1 by FUNCT_1:68
     .=((the carrier of TOP-REAL 2))/\ K1 by FUNCT_2:def 1
     .=K1 by XBOOLE_1:28;
     A37:the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
     then consider p3 being Point of TOP-REAL 2 such that
     A38: p=p3 & (p3`1/|.p3.|<=cn & p3`2<=0 & p3<>0.REAL 2) by A1,A35;
     A39:(cn-FanMorphS).p=|[ |.p.|* ((p`1/|.p.|-cn)/(1+cn)),
       |.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2))]| by A1,A38,Th122;
      ((cn-FanMorphS)|K1).p=(cn-FanMorphS).p by A35,A37,FUNCT_1:72;
     then g1.p=(proj2).(|[ |.p.|* ((p`1/|.p.|-cn)/(1+cn)),
          |.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2))]|)
                             by A35,A36,A37,A39,FUNCT_1:23
        .=(|[ |.p.|* ((p`1/|.p.|-cn)/(1+cn)),
          |.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2))]|)`2 by PSCOMP_1:def 29
        .= |.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2)) by EUCLID:56;
    hence g1.p=|.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2));
   end;
then consider f1 being map of (TOP-REAL 2)|K1,R^1 such that
  A40:for p being Point of TOP-REAL 2 st p in the carrier of (TOP-REAL 2)|K1
  holds f1.p=|.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2));
    for q being Point of TOP-REAL 2 st q in
  the carrier of (TOP-REAL 2)|K1 holds q`2<=0 & q`1/|.q.|<=cn & q<>0.REAL 2
   proof let q be Point of TOP-REAL 2;
    assume A41:q in the carrier of (TOP-REAL 2)|K1;
        the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
     then consider p3 being Point of TOP-REAL 2 such that
     A42: q=p3 & (p3`1/|.p3.|<=cn & p3`2<=0 & p3<>0.REAL 2) by A1,A41;
    thus thesis by A42;
   end;
then A43:f1 is continuous by A1,A40,Th126;
  for x,y,s,r being real number st |[x,y]| in K1 &
  s=f2.(|[x,y]|) & r=f1.(|[x,y]|) holds f.(|[x,y]|)=|[s,r]|
  proof let x,y,s,r be real number;
   assume A44: |[x,y]| in K1 & s=f2.(|[x,y]|) & r=f1.(|[x,y]|);
     set p99=|[x,y]|;
     A45:the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
     consider p3 being Point of TOP-REAL 2 such that
     A46: p99=p3 & (p3`1/|.p3.|<=cn & p3`2<=0 & p3<>0.REAL 2) by A1,A44;
     A47:f1.p99=|.p99.|*( -sqrt(1-((p99`1/|.p99.|-cn)/(1+cn))^2))
                          by A40,A44,A45;
      ((cn-FanMorphS)|K0).(|[x,y]|)=((cn-FanMorphS)).(|[x,y]|) by A44,FUNCT_1:
72
    .= |[ |.p99.|* ((p99`1/|.p99.|-cn)/(1+cn)),
      |.p99.|*( -sqrt(1-((p99`1/|.p99.|-cn)/(1+cn))^2))]| by A1,A46,Th122
    .=|[s,r]| by A33,A44,A45,A47;
   hence f.(|[x,y]|)=|[s,r]| by A1;
  end;
hence f is continuous by A8,A9,A34,A43,JGRAPH_2:45;
end;

theorem Th129: for cn being Real, K03 being Subset of TOP-REAL 2
 st K03={p: p`1>=(cn)*(|.p.|) & p`2<=0}
 holds K03 is closed
proof let sn be Real, K003 be Subset of TOP-REAL 2;
 assume A1: K003={p: p`1>=(sn)*(|.p.|) & p`2<=0};
 defpred P[Point of TOP-REAL 2] means
 ($1`1>=sn*|.$1.|);
  reconsider K1={p7 where p7 is Point of TOP-REAL 2:P[p7]}
  as Subset of TOP-REAL 2 from TopSubset;
  A2: K1 is closed by Lm8;
  defpred Q[Point of TOP-REAL 2] means $1`2<=0;
  reconsider KX={p where p is Point of TOP-REAL 2: Q[p]}
    as Subset of TOP-REAL 2 from TopSubset;
A3: KX is closed by JORDAN6:9;
     {p : P[p] & Q[p]}
    = {p7 where p7 is Point of TOP-REAL 2:P[p7]} /\
     {p1 where p1 is Point of TOP-REAL 2: Q[p1]} from TopInter2;
  hence thesis by A1,A2,A3,TOPS_1:35;
end;

theorem Th130: for cn being Real, K03 being Subset of TOP-REAL 2
 st K03={p: p`1<=(cn)*(|.p.|) & p`2<=0}
 holds K03 is closed
proof let sn be Real, K003 be Subset of TOP-REAL 2;
 assume A1: K003={p: p`1<=(sn)*(|.p.|) & p`2<=0};
 defpred P[Point of TOP-REAL 2] means ($1`1<=sn*|.$1.|);
  reconsider K1={p7 where p7 is Point of TOP-REAL 2:P[p7]}
  as Subset of TOP-REAL 2 from TopSubset;
  A2: K1 is closed by Lm10;
  defpred Q[Point of TOP-REAL 2] means $1`2<=0;
  reconsider KX={p where p is Point of TOP-REAL 2: Q[p]}
    as Subset of TOP-REAL 2 from TopSubset;
A3: KX is closed by JORDAN6:9;
     {p : P[p] & Q[p]}
    = {p7 where p7 is Point of TOP-REAL 2:P[p7]} /\
     {p1 where p1 is Point of TOP-REAL 2: Q[p1]} from TopInter2;
  hence thesis by A1,A2,A3,TOPS_1:35;
end;

theorem Th131: for cn being Real,
K0,B0 being Subset of TOP-REAL 2,f being
map of (TOP-REAL 2)|K0,(TOP-REAL 2)|B0
st -1<cn & cn<1 & f=(cn-FanMorphS)|K0 &
B0=(the carrier of TOP-REAL 2) \ {0.REAL 2}
& K0={p: p`2<=0 & p<>0.REAL 2}
holds f is continuous
proof let cn be Real,K0,B0 be Subset of TOP-REAL 2,f be
map of (TOP-REAL 2)|K0,(TOP-REAL 2)|B0;
assume A1: -1<cn & cn<1 & f=(cn-FanMorphS)|K0 &
B0=(the carrier of TOP-REAL 2) \ {0.REAL 2} & K0={p: p`2<=0 & p<>0.REAL 2};
 set sn=-sqrt(1-cn^2);
 set p0=|[cn,sn]|;
 A2: p0`1=cn by EUCLID:56;
 A3: p0`2=sn by EUCLID:56;
     cn^2<1 by A1,JGRAPH_2:8,SQUARE_1:59;
 then A4: 1-cn^2>0 by SQUARE_1:11;
 A5: now assume p0=0.REAL 2;
   then --sn=-0 by EUCLID:56,JGRAPH_2:11;
  hence contradiction by A4,SQUARE_1:93;
 end;
   -sn>0 by A4,SQUARE_1:93;
 then A6: sn<0 by REAL_1:66;
 A7: |.p0.|=sqrt((sn)^2+cn^2) by A2,A3,JGRAPH_3:10;
   (-sn)^2=1-cn^2 by A4,SQUARE_1:def 4;
 then sn^2=1-cn^2 by SQUARE_1:61;
 then |.p0.|=1 by A7,SQUARE_1:83,XCMPLX_1:27;
 then A8: p0`1/|.p0.|=cn by EUCLID:56;
   p0 in K0 by A1,A3,A5,A6;
then reconsider K1=K0 as non empty Subset of TOP-REAL 2;
    p0 in the carrier of TOP-REAL 2 & not p0 in {0.REAL 2}
               by A5,TARSKI:def 1;
  then reconsider D=B0 as non empty Subset of TOP-REAL 2 by A1,XBOOLE_0:def 4;
   A9: p0 in {p: p`1/|.p.|>=cn & p`2<=0 & p<>0.REAL 2} by A3,A5,A6,A8;
   defpred P[Point of TOP-REAL 2] means
   $1`1/|.$1.|>=cn & $1`2<=0 & $1<>0.REAL 2;
   A10: {p: P[p]}
     is Subset of TOP-REAL 2 from SubsetD;
   A11: the carrier of ((TOP-REAL 2)|K1)=K1 by JORDAN1:1;
   A12: {p: p`1/|.p.|>=cn & p`2<=0 & p<>0.REAL 2} c= K1
    proof let x be set;assume x in
      {p: p`1/|.p.|>=cn & p`2<=0 & p<>0.REAL 2};
      then consider p such that
        A13:p=x & p`1/|.p.|>=cn & p`2<=0 & p<>0.REAL 2;
     thus x in K1 by A1,A13;
    end;
   then reconsider K00={p: p`1/|.p.|>=cn & p`2<=0 & p<>0.REAL 2}
      as non empty Subset of ((TOP-REAL 2)|K1)
                 by A9,A11;
   reconsider K001={p: p`1/|.p.|>=cn & p`2<=0 & p<>0.REAL 2}
            as non empty Subset of (TOP-REAL 2)
                       by A9,A10;
   defpred P[Point of TOP-REAL 2] means $1`1>=(cn)*(|.$1.|) & $1`2<=0;
     {p: P[p]} is Subset of
   TOP-REAL 2 from SubsetD;
   then reconsider K003={p: p`1>=(cn)*(|.p.|) & p`2<=0}
            as Subset of (TOP-REAL 2);
   A14: p0 in {p: p`1/|.p.|<=cn & p`2<=0 & p<>0.REAL 2} by A3,A5,A6,A8;
   A15: {p: p`1/|.p.|<=cn & p`2<=0 & p<>0.REAL 2} c= K1
    proof let x be set;assume x in
      {p: p`1/|.p.|<=cn & p`2<=0 & p<>0.REAL 2};
      then consider p such that
        A16: p=x & p`1/|.p.|<=cn & p`2<=0 & p<>0.REAL 2;
     thus x in K1 by A1,A16;
    end;
   then reconsider K11={p: p`1/|.p.|<=cn & p`2<=0 & p<>0.REAL 2}
            as non empty Subset of ((TOP-REAL 2)|K1)
            by A11,A14;
    A17: p0 in {p: p`1/|.p.|<=cn & p`2<=0 & p<>0.REAL 2} by A3,A5,A6,A8;
    defpred P[Point of TOP-REAL 2] means
    $1`1/|.$1.|<=cn & $1`2<=0 & $1<>0.REAL 2;
     {p: P[p]} is Subset of
   TOP-REAL 2 from SubsetD;
   then reconsider K111={p: p`1/|.p.|<=cn & p`2<=0 & p<>0.REAL 2}
            as non empty Subset of TOP-REAL 2 by A17;
   defpred P[Point of TOP-REAL 2] means $1`1<=(cn)*(|.$1.|) & $1`2<=0;
     {p: P[p]} is Subset of
   TOP-REAL 2 from SubsetD;
   then reconsider K004={p: p`1<=(cn)*(|.p.|) & p`2<=0}
            as Subset of (TOP-REAL 2);
    the carrier of (TOP-REAL 2)|B0=the carrier of (TOP-REAL 2)|D;
  then A18: dom f=the carrier of (TOP-REAL 2)|K1 by FUNCT_2:def 1
    .=K1 by JORDAN1:1;
  then A19: dom (f|K00)=K00 by A12,RELAT_1:91
  .= the carrier of ((TOP-REAL 2)|K1)|K00 by JORDAN1:1;
  A20: the carrier of (TOP-REAL 2)|D = D by JORDAN1:1;
    rng (f|K00) c= rng f by RELAT_1:99;
  then rng (f|K00) c=D by A20,XBOOLE_1:1;
  then f|K00 is Function of the carrier of ((TOP-REAL 2)|K1)|K00,
   the carrier of (TOP-REAL 2)|D by A19,A20,FUNCT_2:4;
  then reconsider f1=f|K00 as map of ((TOP-REAL 2)|K1)|K00,(TOP-REAL 2)|D
                           ;
  A21: dom f1=the carrier of ((TOP-REAL 2)|K1)|K00 by FUNCT_2:def 1
  .=K00 by JORDAN1:1;
  A22: dom (cn-FanMorphS)=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
  then A23: dom ((cn-FanMorphS)|K001)=K001 by RELAT_1:91
  .= the carrier of (TOP-REAL 2)|K001 by JORDAN1:1;
  A24: the carrier of (TOP-REAL 2)|K1 =K1 by JORDAN1:1;
    rng ((cn-FanMorphS)|K001) c= K1
   proof let y be set;assume y in rng ((cn-FanMorphS)|K001);
     then consider x being set such that
     A25: x in dom ((cn-FanMorphS)|K001) & y=((cn-FanMorphS)|K001).x
        by FUNCT_1:def 5;
     A26: dom ((cn-FanMorphS)|K001)=(dom (cn-FanMorphS))/\ K001 by FUNCT_1:68
     .=(the carrier of TOP-REAL 2)/\ K001 by FUNCT_2:def 1
     .=K001 by XBOOLE_1:28;
     then reconsider q=x as Point of TOP-REAL 2 by A25;
     A27: y=(cn-FanMorphS).q by A25,FUNCT_1:70;
     consider p2 being Point of TOP-REAL 2 such that
     A28: p2=q & p2`1/|.p2.|>=cn & p2`2<=0 & p2<>0.REAL 2 by A25,A26;
     A29: cn-FanMorphS.q= |[ |.q.|* ((q`1/|.q.|-cn)/(1-cn)),
        |.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))]| by A1,A28,Th122;
     set q4= |[ |.q.|* ((q`1/|.q.|-cn)/(1-cn)),
        |.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))]|;
     A30: |.q.|>=0 by TOPRNS_1:26;
       |.q.|<>0 by A28,TOPRNS_1:25;
     then A31: (|.q.|)^2>0 by SQUARE_1:74;
     A32: 1-cn>0 by A1,Th3;
        A33: q4`2= |.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))&
        q4`1= |.q.|* ((q`1/|.q.|-cn)/(1-cn)) by EUCLID:56;
        A34: (q`1/|.q.|-cn)>= 0 by A28,SQUARE_1:12;
          0<=(q`2)^2 by SQUARE_1:72;
        then 0+(q`1)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
        then (q`1)^2 <= (|.q.|)^2 by JGRAPH_3:10;
        then (q`1)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A31,REAL_1:73;
        then (q`1)^2/(|.q.|)^2 <= 1 by A31,XCMPLX_1:60;
        then ((q`1)/|.q.|)^2 <= 1 by SQUARE_1:69;
        then 1>=q`1/|.q.| by Th4;
        then A35: 1-cn>=q`1/|.q.|-cn by REAL_1:49;
        A36: -(1-cn)<= -0 by A32,REAL_1:50;
           -(1-cn)<= -( q`1/|.q.|-cn) by A35,REAL_1:50;
         then (-(1-cn))/(1-cn)<=(-( q`1/|.q.|-cn))/(1-cn)
                        by A32,REAL_1:73;
        then A37: -1<=(-( q`1/|.q.|-cn))/(1-cn) by A32,XCMPLX_1:198;
         --(1-cn)>= -(q`1/|.q.|-cn) by A34,A36,REAL_1:50;
       then (-(q`1/|.q.|-cn))/(1-cn)<=1 by A32,REAL_2:117;
       then ((-(q`1/|.q.|-cn))/(1-cn))^2<=1 by A37,JGRAPH_2:7,SQUARE_1:59;
       then A38: 1-((-(q`1/|.q.|-cn))/(1-cn))^2>=0 by SQUARE_1:12;
       then 1-(-((q`1/|.q.|-cn))/(1-cn))^2>=0 by XCMPLX_1:188;
       then A39: 1-((q`1/|.q.|-cn)/(1-cn))^2>=0 by SQUARE_1:61;
         sqrt(1-((-(q`1/|.q.|-cn))/(1-cn))^2)>=0 by A38,SQUARE_1:def 4;
       then sqrt(1-(-(q`1/|.q.|-cn))^2/(1-cn)^2)>=0 by SQUARE_1:69;
       then sqrt(1-(q`1/|.q.|-cn)^2/(1-cn)^2)>=0 by SQUARE_1:61;
       then sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2)>=0 by SQUARE_1:69;
       then A40: -sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2)<=0 by REAL_1:66;
        A41: (q4`2)^2= (|.q.|)^2*( -sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))^2
                         by A33,SQUARE_1:68
         .= (|.q.|)^2*(sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))^2 by SQUARE_1:61
         .= (|.q.|)^2*(1-((q`1/|.q.|-cn)/(1-cn))^2) by A39,SQUARE_1:def 4;
        A42: (q4`1)^2= (|.q.|)^2*((q`1/|.q.|-cn)/(1-cn))^2
                         by A33,SQUARE_1:68;
          (|.q4.|)^2=(q4`1)^2+(q4`2)^2 by JGRAPH_3:10
        .=(|.q.|)^2*(1-((q`1/|.q.|-cn)/(1-cn))^2
           +((q`1/|.q.|-cn)/(1-cn))^2) by A41,A42,XCMPLX_1:8
        .=(|.q.|)^2*1 by XCMPLX_1:27 .=(|.q.|)^2;
     then q4`2<=0 & q4<>0.REAL 2 by A30,A31,A33,A40,SQUARE_1:23,60,TOPRNS_1:24
;
    hence y in K1 by A1,A27,A29;
   end;
  then (cn-FanMorphS)|K001 is Function of the carrier of (TOP-REAL 2)|K001,
   the carrier of (TOP-REAL 2)|K1 by A23,A24,FUNCT_2:4;
  then reconsider f3=(cn-FanMorphS)|K001
   as map of (TOP-REAL 2)|K001,(TOP-REAL 2)|K1 ;
  A43: dom (f|K11)=K11 by A15,A18,RELAT_1:91
  .= the carrier of ((TOP-REAL 2)|K1)|K11 by JORDAN1:1;
  A44: the carrier of (TOP-REAL 2)|D = D by JORDAN1:1;
    rng (f|K11) c= rng f by RELAT_1:99;
  then rng (f|K11) c=D by A44,XBOOLE_1:1;
  then f|K11 is Function of the carrier of ((TOP-REAL 2)|K1)|K11,
   the carrier of (TOP-REAL 2)|D by A43,A44,FUNCT_2:4;
  then reconsider f2=f|K11 as map of ((TOP-REAL 2)|K1)|K11,(TOP-REAL 2)|D
                   ;
  A45: dom f2=the carrier of ((TOP-REAL 2)|K1)|K11 by FUNCT_2:def 1
  .=K11 by JORDAN1:1;
  A46: dom ((cn-FanMorphS)|K111)=K111 by A22,RELAT_1:91
  .= the carrier of (TOP-REAL 2)|K111 by JORDAN1:1;
  A47: the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
    rng ((cn-FanMorphS)|K111) c= K1
   proof let y be set;assume y in rng ((cn-FanMorphS)|K111);
     then consider x being set such that
     A48: x in dom ((cn-FanMorphS)|K111) & y=((cn-FanMorphS)|K111).x
        by FUNCT_1:def 5;
     A49: dom ((cn-FanMorphS)|K111)=(dom (cn-FanMorphS))/\ K111 by FUNCT_1:68
     .=(the carrier of TOP-REAL 2)/\ K111 by FUNCT_2:def 1
     .=K111 by XBOOLE_1:28;
     then reconsider q=x as Point of TOP-REAL 2 by A48;
     A50: y=(cn-FanMorphS).q by A48,FUNCT_1:70;
     consider p2 being Point of TOP-REAL 2 such that
     A51: p2=q & p2`1/|.p2.|<=cn & p2`2<=0 & p2<>0.REAL 2 by A48,A49;
     A52: cn-FanMorphS.q= |[ |.q.|* ((q`1/|.q.|-cn)/(1+cn)),
        |.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))]| by A1,A51,Th122;
     set q4= |[ |.q.|* ((q`1/|.q.|-cn)/(1+cn)),
       |.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))]|;
     A53: |.q.|>=0 by TOPRNS_1:26;
       |.q.|<>0 by A51,TOPRNS_1:25;
     then A54: (|.q.|)^2>0 by SQUARE_1:74;
   A55: 1+cn>0 by A1,Th3;
        A56: q4`2= |.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))&
        q4`1= |.q.|* ((q`1/|.q.|-cn)/(1+cn)) by EUCLID:56;
        A57: (q`1/|.q.|-cn)<=0 by A51,REAL_2:106;
        A58: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
          0<=(q`2)^2 by SQUARE_1:72;
        then 0+(q`1)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
        then (q`1)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A54,A58,REAL_1:73;
        then (q`1)^2/(|.q.|)^2 <= 1 by A54,XCMPLX_1:60;
        then ((q`1)/|.q.|)^2 <= 1 by SQUARE_1:69;
        then -1<=q`1/|.q.| by Th4;
        then -1-cn<=q`1/|.q.|-cn by REAL_1:49;
        then A59: -(1+cn)<=q`1/|.q.|-cn by XCMPLX_1:161;
         (1+cn)/(1+cn)>=(q`1/|.q.|-cn)/(1+cn) by A55,A57,REAL_1:73;
        then A60: 1>=(q`1/|.q.|-cn)/(1+cn) by A55,XCMPLX_1:60;
           (-(1+cn))/(1+cn)<=(( q`1/|.q.|-cn))/(1+cn)
                        by A55,A59,REAL_1:73;
        then -1<=(( q`1/|.q.|-cn))/(1+cn) by A55,XCMPLX_1:198;
       then (((q`1/|.q.|-cn))/(1+cn))^2<=1^2 by A60,JGRAPH_2:7;
       then A61: 1-((q`1/|.q.|-cn)/(1+cn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((q`1/|.q.|-cn)/(1+cn)))^2>=0 by SQUARE_1:61;
       then 1-((-(q`1/|.q.|-cn))/(1+cn))^2>=0 by XCMPLX_1:188;
       then sqrt(1-((-(q`1/|.q.|-cn))/(1+cn))^2)>=0 by SQUARE_1:def 4;
       then sqrt(1-(-(q`1/|.q.|-cn))^2/(1+cn)^2)>=0 by SQUARE_1:69;
       then sqrt(1-(q`1/|.q.|-cn)^2/(1+cn)^2)>=0 by SQUARE_1:61;
       then sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2)>=0 by SQUARE_1:69;
       then -sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2)<=0 by REAL_1:66;
       then A62: q4`2<= 0 by A53,A56,SQUARE_1:23;
        A63: (q4`2)^2= (|.q.|)^2*( -sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))^2
                         by A56,SQUARE_1:68
         .= (|.q.|)^2*(sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))^2 by SQUARE_1:61
         .= (|.q.|)^2*(1-((q`1/|.q.|-cn)/(1+cn))^2) by A61,SQUARE_1:def 4;
        A64: (q4`1)^2= (|.q.|)^2*((q`1/|.q.|-cn)/(1+cn))^2
                         by A56,SQUARE_1:68;
          (|.q4.|)^2=(q4`1)^2+(q4`2)^2 by JGRAPH_3:10
        .=(|.q.|)^2*(1-((q`1/|.q.|-cn)/(1+cn))^2
           +((q`1/|.q.|-cn)/(1+cn))^2) by A63,A64,XCMPLX_1:8
        .=(|.q.|)^2*1 by XCMPLX_1:27 .=(|.q.|)^2;
        then q4<>0.REAL 2 by A54,SQUARE_1:60,TOPRNS_1:24;
    hence y in K1 by A1,A50,A52,A62;
   end;
  then (cn-FanMorphS)|K111 is Function of the carrier of (TOP-REAL 2)|K111,
   the carrier of (TOP-REAL 2)|K1 by A46,A47,FUNCT_2:4;
  then reconsider f4=(cn-FanMorphS)|K111
   as map of (TOP-REAL 2)|K111,(TOP-REAL 2)|K1 ;
  set T1= ((TOP-REAL 2)|K1)|K00,T2=((TOP-REAL 2)|K1)|K11;
  A65: [#](((TOP-REAL 2)|K1)|K00)=K00 by PRE_TOPC:def 10;
  A66: [#](((TOP-REAL 2)|K1)|K11)=K11 by PRE_TOPC:def 10;
  A67: [#]((TOP-REAL 2)|K1)=K1 by PRE_TOPC:def 10;
  A68: K00 \/ K11 c= K1
   proof let x be set;assume
     A69: x in K00 \/ K11;
       now per cases by A69,XBOOLE_0:def 2;
     case x in K00;
       then consider p8 being Point of TOP-REAL 2 such that
       A70: p8=x & p8`1/|.p8.|>=cn & p8`2<=0 & p8<>0.REAL 2;
      thus x in K1 by A1,A70;
     case x in K11;
       then consider p8 being Point of TOP-REAL 2 such that
       A71: p8=x & p8`1/|.p8.|<=cn & p8`2<=0 & p8<>0.REAL 2;
      thus x in K1 by A1,A71;
     end;
    hence x in K1;
   end;
  A72: K1 c= K00 \/ K11
   proof let x be set;assume x in K1;
     then consider p such that
     A73: p=x & (p`2<=0 & p<>0.REAL 2) by A1;
       now per cases;
     case p`1/|.p.|>=cn;
       then x in K00 by A73;
      hence x in K00 \/ K11 by XBOOLE_0:def 2;
     case p`1/|.p.|<cn;
       then x in K11 by A73;
      hence x in K00 \/ K11 by XBOOLE_0:def 2;
     end;
    hence x in K00 \/ K11;
   end;
  then A74: [#](((TOP-REAL 2)|K1)|K00) \/ [#](((TOP-REAL 2)|K1)|K11)
        =[#]((TOP-REAL 2)|K1) by A65,A66,A67,A68,XBOOLE_0:def 10;
  A75: K003 is closed by Th129;
  A76: K003 /\ K1 c= K00
   proof let x be set;assume x in K003 /\ K1;
     then A77: x in K003 & x in K1 by XBOOLE_0:def 3;
     then consider q1 being Point of TOP-REAL 2 such that
     A78: q1=x & q1`1>=(cn)*(|.q1.|) & q1`2<=0;
     consider q2 being Point of TOP-REAL 2 such that
     A79: q2=x & q2`2<=0 & q2<>0.REAL 2 by A1,A77;
     A80: |.q2.|<>0 by A79,TOPRNS_1:25;
       |.q1.|>=0 by TOPRNS_1:26;
     then q1`1/|.q1.|>=(cn)*(|.q1.|)/|.q1.| by A78,A79,A80,REAL_1:73;
     then q1`1/|.q1.|>=(cn) by A78,A79,A80,XCMPLX_1:90;
    hence x in K00 by A78,A79;
   end;
    K00 c= K003 /\ K1
   proof let x be set;assume x in K00;
     then consider p such that
     A81: p=x & (p`1/|.p.|>=cn & p`2<=0 & p<>0.REAL 2);
     A82: |.p.|<>0 by A81,TOPRNS_1:25;
       |.p.|>=0 by TOPRNS_1:26;
     then p`1/|.p.|*|.p.|>=(cn)*(|.p.|) by A81,AXIOMS:25;
     then p`1>=(cn)*(|.p.|) by A82,XCMPLX_1:88;
     then A83: x in K003 by A81;
       x in K1 by A1,A81;
    hence x in K003 /\ K1 by A83,XBOOLE_0:def 3;
   end;
  then K00=K003 /\ [#]((TOP-REAL 2)|K1) by A67,A76,XBOOLE_0:def 10;
  then A84: K00 is closed by A75,PRE_TOPC:43;
  A85: K004 is closed by Th130;
  A86: K004 /\ K1 c= K11
   proof let x be set;assume x in K004 /\ K1;
     then A87: x in K004 & x in K1 by XBOOLE_0:def 3;
     then consider q1 being Point of TOP-REAL 2 such that
     A88: q1=x & q1`1<=(cn)*(|.q1.|) & q1`2<=0;
     consider q2 being Point of TOP-REAL 2 such that
     A89: q2=x & q2`2<=0 & q2<>0.REAL 2 by A1,A87;
     A90: |.q2.|<>0 by A89,TOPRNS_1:25;
       |.q1.|>=0 by TOPRNS_1:26;
     then q1`1/|.q1.|<=(cn)*(|.q1.|)/|.q1.| by A88,A89,A90,REAL_1:73;
     then q1`1/|.q1.|<=(cn) by A88,A89,A90,XCMPLX_1:90;
    hence x in K11 by A88,A89;
   end;
    K11 c= K004 /\ K1
   proof let x be set;assume x in K11;
     then consider p such that
     A91: p=x & (p`1/|.p.|<=cn & p`2<=0 & p<>0.REAL 2);
     A92: |.p.|<>0 by A91,TOPRNS_1:25;
       |.p.|>=0 by TOPRNS_1:26;
     then p`1/|.p.|*|.p.|<=(cn)*(|.p.|) by A91,AXIOMS:25;
     then p`1<=(cn)*(|.p.|) by A92,XCMPLX_1:88;
     then A93: x in K004 by A91;
       x in K1 by A1,A91;
    hence x in K004 /\ K1 by A93,XBOOLE_0:def 3;
   end;
  then K11=K004 /\ [#]((TOP-REAL 2)|K1) by A67,A86,XBOOLE_0:def 10;
  then A94: K11 is closed by A85,PRE_TOPC:43;
  A95: ((TOP-REAL 2)|K1)|K00=(TOP-REAL 2)|K001 by GOBOARD9:4;
    K1 c= D
   proof let x be set;assume x in K1;
     then consider p6 being Point of TOP-REAL 2 such that
     A96: p6=x & p6`2<=0 & p6<>0.REAL 2 by A1;
       x in the carrier of TOP-REAL 2 & not x in {0.REAL 2} by A96,TARSKI:def 1
;
    hence x in D by A1,XBOOLE_0:def 4;
   end;
  then D=K1 \/ D by XBOOLE_1:12;
  then A97: (TOP-REAL 2)|K1 is SubSpace of (TOP-REAL 2)|D by TOPMETR:5;
    the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
  then A98: f1= f3 by A1,FUNCT_1:82;
    f3 is continuous by A1,Th127;
  then A99: f1 is continuous by A95,A97,A98,TOPMETR:7;
  A100: ((TOP-REAL 2)|K1)|K11=(TOP-REAL 2)|K111 by GOBOARD9:4;
    the carrier of (TOP-REAL 2)|K1=K1 by JORDAN1:1;
  then A101: f2= f4 by A1,FUNCT_1:82;
    f4 is continuous by A1,Th128;
  then A102: f2 is continuous by A97,A100,A101,TOPMETR:7;
  A103: D<>{};
  A104: the carrier of (TOP-REAL 2)|B0=B0 by JORDAN1:1;
    for p being set st p in ([#]T1)/\([#]T2) holds f1.p = f2.p
   proof let p be set;
    assume p in ([#]T1)/\([#]T2);
     then A105: p in K00 & p in K11 by A65,A66,XBOOLE_0:def 3;
    hence f1.p=f.p by FUNCT_1:72 .=f2.p by A105,FUNCT_1:72;
   end;
  then consider h being map of (TOP-REAL 2)|K1,(TOP-REAL 2)|D
  such that
  A106: h=f1+*f2 & h is continuous by A65,A66,A74,A84,A94,A99,A102,JGRAPH_2:9;
  A107: dom h=the carrier of ((TOP-REAL 2)|K1) by FUNCT_2:def 1;
  A108: the carrier of ((TOP-REAL 2)|K1)=K0 by JORDAN1:1;
  A109: K0=the carrier of ((TOP-REAL 2)|K0) by JORDAN1:1
  .=dom f by A103,A104,FUNCT_2:def 1;
    for y being set st y in dom h holds h.y=f.y
   proof let y be set;assume
    A110: y in dom h;
      now per cases by A72,A107,A108,A110,XBOOLE_0:def 2;
    case A111: y in K00 & not y in K11;
      then y in dom f1 \/ dom f2 by A21,XBOOLE_0:def 2;
     hence h.y=f1.y by A45,A106,A111,FUNCT_4:def 1 .=f.y by A111,FUNCT_1:72;
    case A112: y in K11;
      then y in dom f1 \/ dom f2 by A45,XBOOLE_0:def 2;
     hence h.y=f2.y by A45,A106,A112,FUNCT_4:def 1 .=f.y by A112,FUNCT_1:72;
    end;
   hence h.y=f.y;
   end;
 hence thesis by A106,A107,A108,A109,FUNCT_1:9;
end;

theorem Th132: for cn being Real,
K0,B0 being Subset of TOP-REAL 2,f being
map of (TOP-REAL 2)|K0,(TOP-REAL 2)|B0
st -1<cn & cn<1 & f=(cn-FanMorphS)|K0 &
B0=(the carrier of TOP-REAL 2) \ {0.REAL 2}
& K0={p: p`2>=0 & p<>0.REAL 2}
holds f is continuous
proof let cn be Real,K0,B0 be Subset of TOP-REAL 2,f be
map of (TOP-REAL 2)|K0,(TOP-REAL 2)|B0;
assume A1:  -1<cn & cn<1 & f=(cn-FanMorphS)|K0 &
B0=(the carrier of TOP-REAL 2)
\ {0.REAL 2} &
K0={p: p`2>=0 & p<>0.REAL 2};
 set sn=sqrt(1-cn^2);
 set p0=|[cn,sn]|;
 A2: p0`2=sn by EUCLID:56;
     cn^2<1^2 by A1,JGRAPH_2:8;
 then A3: 1-cn^2>0 by SQUARE_1:11,59;
 then A4: sn>0 by SQUARE_1:93;
 A5: p0`2>0 by A2,A3,SQUARE_1:93;
   p0 in K0 by A1,A2,A4,JGRAPH_2:11;
then reconsider K1=K0 as non empty Subset of TOP-REAL 2;
   p0 in the carrier of TOP-REAL 2 & not p0 in {0.REAL 2}
               by A5,JGRAPH_2:11,TARSKI:def 1;
  then reconsider D=B0 as non empty Subset of TOP-REAL 2 by A1,XBOOLE_0:def 4;
   A6: K1 c= D
    proof let x be set;assume x in K1;
      then consider p2 being Point of TOP-REAL 2 such that
      A7: p2=x & p2`2>=0 & p2<>0.REAL 2 by A1;
        not p2 in {0.REAL 2} by A7,TARSKI:def 1;
     hence x in D by A1,A7,XBOOLE_0:def 4;
    end;
     for p being Point of (TOP-REAL 2)|K1,V being Subset of (TOP-REAL 2)|D
      st f.p in V & V is open holds
      ex W being Subset of (TOP-REAL 2)|K1 st
      p in W & W is open & f.:W c= V
    proof let p be Point of (TOP-REAL 2)|K1,V be Subset of (TOP-REAL 2)|D;
      assume A8: f.p in V & V is open;
      then consider V2 being Subset of TOP-REAL 2 such that
      A9: V2 is open & V2 /\ [#]((TOP-REAL 2)|D)=V by TOPS_2:32;
      A10: p in the carrier of (TOP-REAL 2)|K1;
      A11: the carrier of (TOP-REAL 2)|K1=[#]((TOP-REAL 2)|K1) by PRE_TOPC:12;
      A12: [#]((TOP-REAL 2)|K1)=K1 by PRE_TOPC:def 10;
        V2 /\ [#]((TOP-REAL 2)|K1) c= the carrier of (TOP-REAL 2)|K1
       proof let x be set;assume x in V2 /\ [#]((TOP-REAL 2)|K1);
         then x in [#]((TOP-REAL 2)|K1) by XBOOLE_0:def 3;
        hence x in the carrier of (TOP-REAL 2)|K1;
       end;
      then reconsider W2=V2 /\ [#]((TOP-REAL 2)|K1) as Subset of (TOP-REAL 2)|
K1
                               ;
      A13: W2 is open by A9,TOPS_2:32;
      A14: f.p=(cn-FanMorphS).p by A1,A11,A12,FUNCT_1:72;
      consider q being Point of TOP-REAL 2 such that
      A15: q=p & q`2>=0 & q <>0.REAL 2 by A1,A10,A11,A12;
        (cn-FanMorphS).q=q by A15,Th120;
      then p in V2 & p in [#]((TOP-REAL 2)|D)
                       by A8,A9,A14,A15,XBOOLE_0:def 3;
      then A16: p in W2 by A11,XBOOLE_0:def 3;
        f.:W2 c= V
       proof let y be set;assume y in f.:W2;
         then consider x being set such that
         A17: x in dom f & x in W2 & y=f.x by FUNCT_1:def 12;
           f is Function of the carrier of (TOP-REAL 2)|K1,
           the carrier of (TOP-REAL 2)|D;
         then dom f= K1 by A11,A12,FUNCT_2:def 1;
         then consider p4 being Point of TOP-REAL 2 such that
         A18: x=p4 & p4`2>=0 & p4<>0.REAL 2 by A1,A17;
         A19: f.p4=(cn-FanMorphS).p4 by A1,A11,A12,A17,A18,FUNCT_1:72
         .=p4 by A18,Th120;
         A20: p4 in V2 & p4 in [#]((TOP-REAL 2)|K1) by A17,A18,XBOOLE_0:def 3
;
         then p4 in D by A6,A12;
         then p4 in [#]((TOP-REAL 2)|D) by PRE_TOPC:def 10;
        hence y in V by A9,A17,A18,A19,A20,XBOOLE_0:def 3;
       end;
     hence ex W being Subset of (TOP-REAL 2)|K1 st
      p in W & W is open & f.:W c= V by A13,A16;
    end;
  hence f is continuous by JGRAPH_2:20;
end;

theorem Th133: for cn being Real,
 B0 being Subset of TOP-REAL 2,K0 being Subset of
(TOP-REAL 2)|B0,f being map of ((TOP-REAL 2)|B0)|K0,((TOP-REAL 2)|B0)
st -1<cn & cn<1 & f=(cn-FanMorphS)|K0 &
B0=(the carrier of TOP-REAL 2) \ {0.REAL 2}
& K0={p: p`2<=0 & p<>0.REAL 2}
holds f is continuous
proof let cn be Real,
 B0 be Subset of TOP-REAL 2,K0 be Subset of
(TOP-REAL 2)|B0,f being map of ((TOP-REAL 2)|B0)|K0,((TOP-REAL 2)|B0);
assume A1:  -1<cn & cn<1 & f=(cn-FanMorphS)|K0 &
  B0=(the carrier of TOP-REAL 2) \ {0.REAL 2}
   & K0={p: p`2<=0 & p<>0.REAL 2 };
    the carrier of (TOP-REAL 2)|B0= B0 by JORDAN1:1;
  then K0 c= the carrier of TOP-REAL 2 by XBOOLE_1:1;
  then reconsider K1=K0 as Subset of TOP-REAL 2;
    K0 c= B0
   proof let x be set;assume x in K0;
      then consider p8 being Point of TOP-REAL 2 such that
      A2: x=p8 & (p8`2<=0 & p8<>0.REAL 2) by A1;
       not x in {0.REAL 2} by A2,TARSKI:def 1;
    hence x in B0 by A1,A2,XBOOLE_0:def 4;
   end;
  then ((TOP-REAL 2)|B0)|K0=(TOP-REAL 2)|K1 by JORDAN6:47;
hence f is continuous by A1,Th131;
end;

theorem Th134: for cn being Real,
 B0 being Subset of TOP-REAL 2,K0 being Subset of
(TOP-REAL 2)|B0,f being map of ((TOP-REAL 2)|B0)|K0,((TOP-REAL 2)|B0)
st -1<cn & cn<1 & f=(cn-FanMorphS)|K0 &
B0=(the carrier of TOP-REAL 2) \ {0.REAL 2}
& K0={p: p`2>=0 & p<>0.REAL 2}
holds f is continuous
proof let cn be Real,
 B0 be Subset of TOP-REAL 2,K0 be Subset of
(TOP-REAL 2)|B0,f be map of ((TOP-REAL 2)|B0)|K0,((TOP-REAL 2)|B0);
 assume A1:  -1<cn & cn<1 & f=(cn-FanMorphS)|K0 &
  B0=(the carrier of TOP-REAL 2) \ {0.REAL 2}
  & K0={p: p`2>=0 & p<>0.REAL 2};
    the carrier of (TOP-REAL 2)|B0=B0 by JORDAN1:1;
  then K0 c= the carrier of TOP-REAL 2 by XBOOLE_1:1;
  then reconsider K1=K0 as Subset of TOP-REAL 2;
    K0 c= B0
   proof let x be set;assume x in K0;
      then consider p8 being Point of TOP-REAL 2 such that
      A2: x=p8 & (p8`2>=0 & p8<>0.REAL 2) by A1;
       not x in {0.REAL 2} by A2,TARSKI:def 1;
    hence x in B0 by A1,A2,XBOOLE_0:def 4;
   end;
  then ((TOP-REAL 2)|B0)|K0=(TOP-REAL 2)|K1 by JORDAN6:47;
hence f is continuous by A1,Th132;
end;

theorem Th135: for cn being Real,p being Point of TOP-REAL 2
holds |.(cn-FanMorphS).p.|=|.p.|
proof let cn be Real,p be Point of TOP-REAL 2;
       set f=(cn-FanMorphS);
       set z=f.p;
        set q=p;
        reconsider qz=z as Point of TOP-REAL 2;
        A1: |.q.|>=0 by TOPRNS_1:26;
          now per cases;
        case A2: q`1/|.q.|>=cn & q`2<0;
          then (cn-FanMorphS).q= |[ |.q.|* ((q`1/|.q.|-cn)/(1-cn)),
          |.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))]| by Th120;
          then A3: qz`2= |.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))&
                  qz`1= |.q.|* ((q`1/|.q.|-cn)/(1-cn)) by EUCLID:56;
            |.q.|<>0 by A2,JGRAPH_2:11,TOPRNS_1:25;
          then A4: (|.q.|)^2>0 by SQUARE_1:74;
          A5: (q`1/|.q.|-cn)>=0 by A2,SQUARE_1:12;
          A6: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
            0<=(q`2)^2 by SQUARE_1:72;
          then 0+(q`1)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
          then (q`1)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A4,A6,REAL_1:73;
          then (q`1)^2/(|.q.|)^2 <= 1 by A4,XCMPLX_1:60;
          then ((q`1)/|.q.|)^2 <= 1 by SQUARE_1:69;
          then 1>=q`1/|.q.| by Th4;
          then A7: 1-cn>=q`1/|.q.|-cn by REAL_1:49;
            now per cases;
          case A8: 1-cn=0;
            A9:((q`1/|.q.|-cn)/(1-cn))=(q`1/|.q.|-cn)*(1-cn)" by XCMPLX_0:def 9
             .= (q`1/|.q.|-cn)*0 by A8,XCMPLX_0:def 7 .=0;
            then 1-((q`1/|.q.|-cn)/(1-cn))^2=1 by SQUARE_1:60;
             then (cn-FanMorphS).q= |[ |.q.|*0,|.q.|*(-1)]| by A2,A9,Th120,
SQUARE_1:83
            .=|[0,-(|.q.|)]| by XCMPLX_1:180;
             then ((cn-FanMorphS).q)`2=(- |.q.|) & ((cn-FanMorphS).q)`1=0
                              by EUCLID:56;
            then |.(cn-FanMorphS).p.|=sqrt((-(|.q.|))^2+0) by JGRAPH_3:10,
SQUARE_1:60
            .=sqrt(((|.q.|))^2) by SQUARE_1:61
            .=|.q.| by A1,SQUARE_1:89;
           hence |.(cn-FanMorphS).p.|=|.p.|;
          case A10: 1-cn<>0;
              now per cases by A10;
            case A11: 1-cn>0;
              then A12: -(1-cn)<= -0 by REAL_1:50;
                -(1-cn)<= -( q`1/|.q.|-cn) by A7,REAL_1:50;
              then (-(1-cn))/(1-cn)<=(-( q`1/|.q.|-cn))/(1-cn)
                        by A11,REAL_1:73;
              then A13: -1<=(-( q`1/|.q.|-cn))/(1-cn) by A11,XCMPLX_1:198;
               --(1-cn)>= -(q`1/|.q.|-cn) by A5,A12,REAL_1:50;
              then (-(q`1/|.q.|-cn))/(1-cn)<=1 by A11,REAL_2:117;
              then ((-(q`1/|.q.|-cn))/(1-cn))^2<=1^2 by A13,JGRAPH_2:7;
              then 1-((-(q`1/|.q.|-cn))/(1-cn))^2>=0 by SQUARE_1:12,59;
              then 1-(-((q`1/|.q.|-cn))/(1-cn))^2>=0 by XCMPLX_1:188;
              then A14: 1-((q`1/|.q.|-cn)/(1-cn))^2>=0 by SQUARE_1:61;
              A15: (qz`2)^2= (|.q.|)^2*( -sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))^2
                         by A3,SQUARE_1:68
                  .= (|.q.|)^2*(sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))^2
                            by SQUARE_1:61
                  .= (|.q.|)^2*(1-((q`1/|.q.|-cn)/(1-cn))^2)
                            by A14,SQUARE_1:def 4;
              A16: (qz`1)^2= (|.q.|)^2*((q`1/|.q.|-cn)/(1-cn))^2
                         by A3,SQUARE_1:68;
                (|.qz.|)^2=(qz`1)^2+(qz`2)^2 by JGRAPH_3:10
               .=(|.q.|)^2*(1-((q`1/|.q.|-cn)/(1-cn))^2
                    +((q`1/|.q.|-cn)/(1-cn))^2) by A15,A16,XCMPLX_1:8
               .=(|.q.|)^2*1 by XCMPLX_1:27 .=(|.q.|)^2;

              then A17: sqrt((|.qz.|)^2)=|.q.| by A1,SQUARE_1:89;
                |.qz.|>=0 by TOPRNS_1:26;
             hence |.(cn-FanMorphS).p.|=|.p.| by A17,SQUARE_1:89;
            case A18: 1-cn<0;
              A19: q`1/|.q.|-cn>=0 by A2,SQUARE_1:12;
                0<(q`2)^2 by A2,SQUARE_1:74;
              then 0+(q`1)^2<(q`1)^2+(q`2)^2 by REAL_1:67;
              then (q`1)^2/(|.q.|)^2 < (|.q.|)^2/(|.q.|)^2 by A4,A6,REAL_1:73;
              then (q`1)^2/(|.q.|)^2 < 1 by A4,XCMPLX_1:60;
              then ((q`1)/|.q.|)^2 < 1 by SQUARE_1:69;
              then 1 > q`1/|.p.| by Th5;
              then q`1/|.q.|-cn<1-cn by REAL_1:54;
             hence |.(cn-FanMorphS).p.|=|.p.| by A18,A19,AXIOMS:22;
            end;
           hence |.(cn-FanMorphS).p.|=|.p.|;
          end;
         hence |.(cn-FanMorphS).p.|=|.p.|;
        case A20: q`1/|.q.|<cn & q`2<0;
          then A21:(cn-FanMorphS).q= |[ |.q.|* ((q`1/|.q.|-cn)/(1+cn)),
          |.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))]| by Th121;
          then A22: qz`2= |.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))&
                  qz`1= |.q.|* ((q`1/|.q.|-cn)/(1+cn))
                                by EUCLID:56;
            |.q.|<>0 by A20,JGRAPH_2:11,TOPRNS_1:25;
          then A23: (|.q.|)^2>0 by SQUARE_1:74;
          A24: (q`1/|.q.|-cn)<0 by A20,REAL_2:105;
          A25: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
            0<=(q`2)^2 by SQUARE_1:72;
          then 0+(q`1)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
          then (q`1)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A23,A25,REAL_1:73;
          then (q`1)^2/(|.q.|)^2 <= 1 by A23,XCMPLX_1:60;
          then ((q`1)/|.q.|)^2 <= 1 by SQUARE_1:69;
          then -1<=q`1/|.q.| by Th4;
          then A26: -1-cn<=q`1/|.q.|-cn by REAL_1:49;
            now per cases;
          case A27: 1+cn=0;
             ((q`1/|.q.|-cn)/(1+cn))=(q`1/|.q.|-cn)*(1+cn)" by XCMPLX_0:def 9
             .= (q`1/|.q.|-cn)*0 by A27,XCMPLX_0:def 7 .=0;
             then (cn-FanMorphS).q=|[0,-(|.q.|)]| by A21,SQUARE_1:60,83,
XCMPLX_1:180;
             then ((cn-FanMorphS).q)`2=-(|.q.|) & ((cn-FanMorphS).q)`1=0
                              by EUCLID:56;
            then |.(cn-FanMorphS).p.|=sqrt((-(|.q.|))^2+0) by JGRAPH_3:10,
SQUARE_1:60
            .=sqrt(((|.q.|))^2) by SQUARE_1:61
            .=|.q.| by A1,SQUARE_1:89;
           hence |.(cn-FanMorphS).p.|=|.p.|;
          case A28: 1+cn<>0;
              now per cases by A28;
            case A29: 1+cn>0;
                -(1+cn)<= ( q`1/|.q.|-cn) by A26,XCMPLX_1:161;
              then (-(1+cn))/(1+cn)<=(( q`1/|.q.|-cn))/(1+cn)
                        by A29,REAL_1:73;
              then A30: -1<=(( q`1/|.q.|-cn))/(1+cn)
                        by A29,XCMPLX_1:198;
                (1+cn)>= (q`1/|.q.|-cn) by A24,A29,AXIOMS:22;
              then ((q`1/|.q.|-cn))/(1+cn)<=1 by A29,REAL_2:117;
              then (((q`1/|.q.|-cn))/(1+cn))^2<=1 by A30,JGRAPH_2:7,SQUARE_1:59
;
              then A31: 1-(((q`1/|.q.|-cn))/(1+cn))^2>=0 by SQUARE_1:12;
              A32: (qz`2)^2= (|.q.|)^2*( -sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))^2
                         by A22,SQUARE_1:68
                  .= (|.q.|)^2*(sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))^2
                         by SQUARE_1:61
                  .= (|.q.|)^2*(1-((q`1/|.q.|-cn)/(1+cn))^2)
                            by A31,SQUARE_1:def 4;
              A33: (qz`1)^2= (|.q.|)^2*((q`1/|.q.|-cn)/(1+cn))^2
                         by A22,SQUARE_1:68;
                (|.qz.|)^2=(qz`1)^2+(qz`2)^2 by JGRAPH_3:10
               .=(|.q.|)^2*(1-((q`1/|.q.|-cn)/(1+cn))^2
                    +((q`1/|.q.|-cn)/(1+cn))^2) by A32,A33,XCMPLX_1:8
               .=(|.q.|)^2*1 by XCMPLX_1:27 .=(|.q.|)^2;
              then A34: sqrt((|.qz.|)^2)=|.q.| by A1,SQUARE_1:89;
                |.qz.|>=0 by TOPRNS_1:26;
             hence |.(cn-FanMorphS).p.|=|.p.| by A34,SQUARE_1:89;
            case 1+cn<0;
              then A35: -(1+cn)>-0 by REAL_1:50;
                0<(q`2)^2 by A20,SQUARE_1:74;
              then 0+(q`1)^2<(q`1)^2+(q`2)^2 by REAL_1:67;
              then (q`1)^2/(|.q.|)^2 < (|.q.|)^2/(|.q.|)^2 by A23,A25,REAL_1:73
;
              then (q`1)^2/(|.q.|)^2 < 1 by A23,XCMPLX_1:60;
              then ((q`1)/|.q.|)^2 < 1 by SQUARE_1:69;
              then -1 < q`1/|.p.| by Th5;
              then q`1/|.q.|-cn>-1-cn by REAL_1:54;
              then q`1/|.q.|-cn>-(1+cn) by XCMPLX_1:161;
             hence |.(cn-FanMorphS).p.|=|.p.| by A24,A35,AXIOMS:22;
            end;
           hence |.(cn-FanMorphS).p.|=|.p.|;
          end;
         hence |.(cn-FanMorphS).p.|=|.p.|;
        case q`2>=0;
         hence |.(cn-FanMorphS).p.|=|.p.| by Th120;
        end;
 hence thesis;
end;

theorem Th136: for cn being Real,x,K0 being set
 st -1<cn & cn<1 & x in K0 &
 K0={p: p`2<=0 & p<>0.REAL 2}
holds (cn-FanMorphS).x in K0
proof let cn be Real,x,K0 be set;
 assume A1:  -1<cn & cn<1 & x in K0 &
   K0={p: p`2<=0 & p<>0.REAL 2};
  then consider p such that
  A2: p=x & p`2<=0 & p<>0.REAL 2;
  A3:now assume |.p.|<=0;
    then |.p.|=0 by TOPRNS_1:26;
   hence contradiction by A2,TOPRNS_1:25;
  end;
  then A4: (|.p.|)^2>0 by SQUARE_1:74;
    now per cases;
  case A5: p`1/|.p.|<=cn;
   then A6: (cn-FanMorphS).p= |[ |.p.|*((p`1/|.p.|-cn)/(1+cn)),
      |.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2))]| by A1,A2,Th122;
   reconsider p9= (cn-FanMorphS).p as Point of TOP-REAL 2;
   A7: p9`2=|.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2)) by A6,EUCLID:56;
   A8: 1+cn>0 by A1,Th3;
       A9: (|.p.|)^2 =(p`1)^2+(p`2)^2 by JGRAPH_3:10;
         now per cases;
       case p`2=0;
        hence (cn-FanMorphS).x in K0 by A1,A2,Th120;
       case p`2<>0;
         then 0<(p`2)^2 by SQUARE_1:74;
         then 0+(p`1)^2<(p`1)^2+(p`2)^2 by REAL_1:67;
         then (p`1)^2/(|.p.|)^2 < (|.p.|)^2/(|.p.|)^2 by A4,A9,REAL_1:73;
         then (p`1)^2/(|.p.|)^2 < 1 by A4,XCMPLX_1:60;
         then A10: ((p`1)/|.p.|)^2 < 1 by SQUARE_1:69;
           p`1/|.p.|-cn<=0 by A5,REAL_2:106;
         then A11: (p`1/|.p.|-cn)/(1+cn)<(1+cn)/(1+cn) by A8,REAL_1:73;
           -1 < p`1/|.p.| by A10,Th5;
         then -1-cn< p`1/|.p.|-cn by REAL_1:54;
         then -(1+cn)<p`1/|.p.|-cn by XCMPLX_1:161;
         then (-1)*(1+cn)< p`1/|.p.|-cn by XCMPLX_1:180;
         then (-1)*(1+cn)/(1+cn)< (p`1/|.p.|-cn)/(1+cn)
                       by A8,REAL_1:73;
         then -1< (p`1/|.p.|-cn)/(1+cn) & (p`1/|.p.|-cn)/(1+cn)<1
                           by A8,A11,XCMPLX_1:60,90;
         then 1^2> ((p`1/|.p.|-cn)/(1+cn))^2 by JGRAPH_2:8;
         then 1-((p`1/|.p.|-cn)/(1+cn))^2>0 by SQUARE_1:11,59;
         then --sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2)>0 by SQUARE_1:93;
         then -sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2)<0 by REAL_1:66;
         then |.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2))<0
                           by A3,SQUARE_1:24;
        hence (cn-FanMorphS).x in K0 by A1,A2,A7,JGRAPH_2:11;
        end;
       hence (cn-FanMorphS).x in K0;
  case A12: p`1/|.p.|>cn;
   then A13: (cn-FanMorphS).p= |[ |.p.|*((p`1/|.p.|-cn)/(1-cn)),
       |.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2))]| by A1,A2,Th122;
   reconsider p9= (cn-FanMorphS).p as Point of TOP-REAL 2;
   A14: p9`2=|.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2)) by A13,EUCLID:56;
    A15: 1-cn>0 by A1,Th3;
       A16: (|.p.|)^2 =(p`1)^2+(p`2)^2 by JGRAPH_3:10;
         now per cases;
       case p`2=0;
        hence (cn-FanMorphS).x in K0 by A1,A2,Th120;
       case p`2<>0;
         then 0<(p`2)^2 by SQUARE_1:74;
         then 0+(p`1)^2<(p`1)^2+(p`2)^2 by REAL_1:67;
         then (p`1)^2/(|.p.|)^2 < (|.p.|)^2/(|.p.|)^2 by A4,A16,REAL_1:73;
         then (p`1)^2/(|.p.|)^2 < 1 by A4,XCMPLX_1:60;
         then ((p`1)/|.p.|)^2 < 1 by SQUARE_1:69;
         then p`1/|.p.|<1 by Th5;
         then (p`1/|.p.|-cn)<1-cn by REAL_1:54;
         then A17: (p`1/|.p.|-cn)/(1-cn)<(1-cn)/(1-cn) by A15,REAL_1:73;
           -(1-cn)< -0 by A15,REAL_1:50;
         then A18: -(1-cn)<cn-cn by XCMPLX_1:14;
           p`1/|.p.|-cn>=cn-cn by A12,REAL_1:49;
         then -(1-cn)<p`1/|.p.|-cn by A18,AXIOMS:22;
         then (-1)*(1-cn)< p`1/|.p.|-cn by XCMPLX_1:180;
         then (-1)*(1-cn)/(1-cn)< (p`1/|.p.|-cn)/(1-cn)
                       by A15,REAL_1:73;
         then -1< (p`1/|.p.|-cn)/(1-cn) & (p`1/|.p.|-cn)/(1-cn)<1
                           by A15,A17,XCMPLX_1:60,90;
         then 1> ((p`1/|.p.|-cn)/(1-cn))^2 by JGRAPH_2:8,SQUARE_1:59;
         then 1-((p`1/|.p.|-cn)/(1-cn))^2>0 by SQUARE_1:11;
         then --sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2)>0 by SQUARE_1:93;
         then -sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2)<0 by REAL_1:66;
         then p9`2<0 by A3,A14,SQUARE_1:24;
        hence (cn-FanMorphS).x in K0 by A1,A2,JGRAPH_2:11;
       end;
        hence (cn-FanMorphS).x in K0;
        end;
 hence thesis;
end;

theorem Th137: for cn being Real,x,K0 being set
 st -1<cn & cn<1 & x in K0 &
 K0={p: p`2>=0 & p<>0.REAL 2}
holds (cn-FanMorphS).x in K0
proof let cn be Real,x,K0 be set;
 assume A1:  -1<cn & cn<1 & x in K0 &
 K0={p: p`2>=0 & p<>0.REAL 2};
 then consider p such that
 A2: p=x & p`2>=0 & p<>0.REAL 2;
 thus (cn-FanMorphS).x in K0 by A1,A2,Th120;
end;

theorem Th138:for cn being Real,
 D being non empty Subset of TOP-REAL 2
st -1<cn & cn<1 & D`={0.REAL 2} holds
ex h being map of (TOP-REAL 2)|D,(TOP-REAL 2)|D
st h=(cn-FanMorphS)|D & h is continuous
proof let cn be Real,D be non empty Subset of TOP-REAL 2;
assume A1:  -1<cn & cn<1 & D`={0.REAL 2};
  reconsider B0= {0.REAL 2} as Subset of TOP-REAL 2;
  A2: D =(B0)` by A1
     .=((the carrier of TOP-REAL 2)\ {0.REAL 2}) by SUBSET_1:def 5;
     defpred P[Point of TOP-REAL 2] means $1`2<=0;
    A3:{p: P[p] & p<>0.REAL 2} c= the carrier of (TOP-REAL 2)|D
           from InclSub(A2);
      A4: (|[0,-1]|)`1=0 & (|[0,-1]|)`2=-1 by EUCLID:56;
        |[0,-1]|<>0.REAL 2 by EUCLID:56,JGRAPH_2:11;
      then |[0,-1]| in {p where p is Point of TOP-REAL 2:
        p`2<=0 & p<>0.REAL 2} by A4;
      then reconsider K0={p:p`2<=0 & p<>0.REAL 2}
          as non empty Subset of (TOP-REAL 2)|D by A3;
          defpred P[Point of TOP-REAL 2] means $1`2>=0;
    A5:{p: P[p] & p<>0.REAL 2} c= the carrier of (TOP-REAL 2)|D
           from InclSub(A2);
     set Y1=|[0,1]|;
      Y1`1=0 & Y1`2=1 by EUCLID:56;
     then Y1 in {p where p is Point of TOP-REAL 2:
     p`2>=0 & p<>0.REAL 2} by JGRAPH_2:11;
  then reconsider K1={p: p`2>=0 & p<>0.REAL 2}
       as non empty Subset of (TOP-REAL 2)|D by A5;
   A6:the carrier of ((TOP-REAL 2)|D) =D by JORDAN1:1;
   A7:K0 c= (the carrier of TOP-REAL 2)
   proof let z be set;assume z in K0;
        then consider p8 being Point of TOP-REAL 2 such that
        A8: p8=z & p8`2<=0 & p8<>0.REAL 2;
    thus z in (the carrier of TOP-REAL 2) by A8;
   end;
  A9:dom ((cn-FanMorphS)|K0)= dom ((cn-FanMorphS)) /\ K0 by FUNCT_1:68
       .=((the carrier of TOP-REAL 2)) /\ K0 by FUNCT_2:def 1
       .=K0 by A7,XBOOLE_1:28;
  A10: K0=the carrier of ((TOP-REAL 2)|D)|K0 by JORDAN1:1;
    rng ((cn-FanMorphS)|K0) c= the carrier of ((TOP-REAL 2)|D)|K0
   proof let y be set;assume y in rng ((cn-FanMorphS)|K0);
     then consider x being set such that
     A11:x in dom ((cn-FanMorphS)|K0)
     & y=((cn-FanMorphS)|K0).x by FUNCT_1:def 5;
       x in (dom (cn-FanMorphS)) /\ K0 by A11,FUNCT_1:68;
     then A12:x in K0 by XBOOLE_0:def 3;
       K0 c= the carrier of TOP-REAL 2 by A6,XBOOLE_1:1;
     then reconsider p=x as Point of TOP-REAL 2 by A12;
      (cn-FanMorphS).p=y by A11,A12,FUNCT_1:72;
     then y in K0 by A1,A12,Th136;
    hence y in the carrier of ((TOP-REAL 2)|D)|K0 by JORDAN1:1;
   end;
  then rng ((cn-FanMorphS)|K0)c= the carrier of ((TOP-REAL 2)|D)
                            by A10,XBOOLE_1:1;
  then (cn-FanMorphS)|K0 is Function of the carrier of ((TOP-REAL 2)|D)|K0,
  the carrier of ((TOP-REAL 2)|D) by A9,A10,FUNCT_2:4;
  then reconsider f=(cn-FanMorphS)|K0 as map of ((TOP-REAL 2)|D)|K0,
  ((TOP-REAL 2)|D) ;
   A13:K1 c= (the carrier of TOP-REAL 2)
   proof let z be set;assume z in K1;
        then consider p8 being Point of TOP-REAL 2 such that
        A14: p8=z & p8`2>=0 & p8<>0.REAL 2;
    thus z in (the carrier of TOP-REAL 2) by A14;
   end;
  A15:dom ((cn-FanMorphS)|K1)= dom ((cn-FanMorphS)) /\ K1 by FUNCT_1:68
       .=((the carrier of TOP-REAL 2)) /\ K1 by FUNCT_2:def 1
       .=K1 by A13,XBOOLE_1:28;
  A16: K1=the carrier of ((TOP-REAL 2)|D)|K1 by JORDAN1:1;
    rng ((cn-FanMorphS)|K1) c= the carrier of ((TOP-REAL 2)|D)|K1
   proof let y be set;assume y in rng ((cn-FanMorphS)|K1);
     then consider x being set such that
     A17:x in dom ((cn-FanMorphS)|K1)
     & y=((cn-FanMorphS)|K1).x by FUNCT_1:def 5;
       x in (dom (cn-FanMorphS)) /\ K1 by A17,FUNCT_1:68;
     then A18:x in K1 by XBOOLE_0:def 3;
       K1 c= the carrier of TOP-REAL 2 by A6,XBOOLE_1:1;
     then reconsider p=x as Point of TOP-REAL 2 by A18;
      (cn-FanMorphS).p=y by A17,A18,FUNCT_1:72;
     then y in K1 by A1,A18,Th137;
    hence y in the carrier of ((TOP-REAL 2)|D)|K1 by JORDAN1:1;
   end;
  then rng ((cn-FanMorphS)|K1)c= the carrier of ((TOP-REAL 2)|D)
                     by A16,XBOOLE_1:1;
  then (cn-FanMorphS)|K1 is Function of the carrier of ((TOP-REAL 2)|D)|K1,
  the carrier of ((TOP-REAL 2)|D) by A15,A16,FUNCT_2:4;
  then reconsider g=(cn-FanMorphS)|K1 as map of ((TOP-REAL 2)|D)|K1,
  ((TOP-REAL 2)|D) ;
  A19:K0=[#](((TOP-REAL 2)|D)|K0) by PRE_TOPC:def 10;
  A20:K1=[#](((TOP-REAL 2)|D)|K1) by PRE_TOPC:def 10;
  A21:D= [#]((TOP-REAL 2)|D) by PRE_TOPC:def 10;
  A22:K0 \/ K1 c= D
   proof let x be set;assume A23:x in K0 \/ K1;
       now per cases by A23,XBOOLE_0:def 2;
     case x in K0;
       then consider p such that A24:p=x & p`2<=0 & p<>0.REAL 2;
     thus x in the carrier of TOP-REAL 2 & not x=0.REAL 2 by A24;
     case x in K1;
       then consider p such that A25:p=x & p`2>=0 & p<>0.REAL 2;
     thus x in the carrier of TOP-REAL 2 & not x=0.REAL 2 by A25;
     end;
     then x in the carrier of TOP-REAL 2 & not x in {0.REAL 2} by TARSKI:def 1
;
    hence x in D by A2,XBOOLE_0:def 4;
   end;
    D c= K0 \/ K1
   proof let x be set;assume A26: x in D;
     then A27:x in (the carrier of TOP-REAL 2) & not x in {0.REAL 2}
                                     by A2,XBOOLE_0:def 4;
     reconsider px=x as Point of TOP-REAL 2 by A26;
       px`2>=0 & px<>0.REAL 2 or
       px`2<=0 & px<>0.REAL 2 by A27,TARSKI:def 1;
     then x in K1 or x in K0;
    hence x in K0 \/ K1 by XBOOLE_0:def 2;
   end;
  then A28:([#](((TOP-REAL 2)|D)|K0)) \/ ([#](((TOP-REAL 2)|D)|K1))
  = [#]((TOP-REAL 2)|D) by A19,A20,A21,A22,XBOOLE_0:def 10;
  A29: f is continuous & K0 is closed by A1,A2,Th70,Th133;
  A30: g is continuous & K1 is closed by A1,A2,Th69,Th134;
  A31: for x be set st x in ([#]((((TOP-REAL 2)|D)|K0)))
        /\ ([#] ((((TOP-REAL 2)|D)|K1))) holds f.x = g.x
   proof let x be set;assume x in ([#]((((TOP-REAL 2)|D)|K0)))
        /\ ([#] ((((TOP-REAL 2)|D)|K1)));
     then A32:x in K0 & x in K1 by A19,A20,XBOOLE_0:def 3;
     then f.x=(cn-FanMorphS).x by FUNCT_1:72;
    hence f.x = g.x by A32,FUNCT_1:72;
   end;
  then consider h being map of (TOP-REAL 2)|D,(TOP-REAL 2)|D
  such that A33: h= f+*g & h is continuous
                          by A19,A20,A28,A29,A30,JGRAPH_2:9;
  A34:dom h=the carrier of ((TOP-REAL 2)|D) by FUNCT_2:def 1;
  A35:the carrier of ((TOP-REAL 2)|D)=D by JORDAN1:1;
  A36:the carrier of ((TOP-REAL 2)|D)
       =((the carrier of TOP-REAL 2)\ {0.REAL 2}) by A2,JORDAN1:1;
    dom (cn-FanMorphS)=(the carrier of (TOP-REAL 2))
                    by FUNCT_2:def 1;
  then A37:dom ((cn-FanMorphS)|D)=(the carrier of (TOP-REAL 2))/\ D
                    by FUNCT_1:68
              .=the carrier of ((TOP-REAL 2)|D) by A35,XBOOLE_1:28;
  A38:dom f=K0 by A10,FUNCT_2:def 1;
  A39:K0=[#](((TOP-REAL 2)|D)|K0) by PRE_TOPC:def 10;
  A40:dom g=K1 by A16,FUNCT_2:def 1;
   K1=[#](((TOP-REAL 2)|D)|K1) by PRE_TOPC:def 10;
  then A41: f tolerates g by A31,A38,A39,A40,PARTFUN1:def 6;
    for x being set st x in dom h holds h.x=((cn-FanMorphS)|D).x
   proof let x be set;assume A42: x in dom h;
     then x in (the carrier of TOP-REAL 2) & not x in {0.REAL 2}
                     by A2,A34,A35,XBOOLE_0:def 4;
     then A43:x <>0.REAL 2 by TARSKI:def 1;
     reconsider p=x as Point of TOP-REAL 2 by A34,A36,A42,XBOOLE_0:def 4;
       now per cases;
     case A44:x in K0;
       A45:(cn-FanMorphS)|D.p=(cn-FanMorphS).p by A34,A35,A42,FUNCT_1:72
          .=f.p by A44,FUNCT_1:72;
         h.p=(g+*f).p by A33,A41,FUNCT_4:35
       .=f.p by A38,A44,FUNCT_4:14;
      hence h.x=(cn-FanMorphS)|D.x by A45;
     case not x in K0;
       then not (p`2<=0) by A43;
       then A46:x in K1 by A43;
        (cn-FanMorphS)|D.p=(cn-FanMorphS).p by A34,A35,A42,FUNCT_1:72
                           .=g.p by A46,FUNCT_1:72;
      hence h.x=(cn-FanMorphS)|D.x by A33,A40,A46,FUNCT_4:14;
     end;
    hence h.x=(cn-FanMorphS)|D.x;
   end;
  then f+*g=(cn-FanMorphS)|D by A33,A34,A37,FUNCT_1:9;
 hence thesis by A19,A20,A28,A29,A30,A31,JGRAPH_2:9;
end;

theorem Th139: for cn being Real
  st -1<cn & cn<1 holds
  ex h being map of (TOP-REAL 2),(TOP-REAL 2)
  st h=(cn-FanMorphS) & h is continuous
proof let cn be Real;
 assume A1:  -1<cn & cn<1;
 reconsider f=(cn-FanMorphS) as map of (TOP-REAL 2),(TOP-REAL 2)
                            ;
  (the carrier of TOP-REAL 2) \{0.REAL 2} c=(the carrier of TOP-REAL 2)
                         by XBOOLE_1:36;
 then reconsider D=(the carrier of TOP-REAL 2) \{0.REAL 2}
   as non empty Subset of TOP-REAL 2 by JGRAPH_2:19;
 A2: f.(0.REAL 2)=0.REAL 2 by Th120,JGRAPH_2:11;
 A3: D`= {0.REAL 2} by JGRAPH_3:30;
 A4: for p being Point of (TOP-REAL 2)|D holds f.p<>f.(0.REAL 2)
  proof let p be Point of (TOP-REAL 2)|D;
    A5:the carrier of (TOP-REAL 2)|D=[#]((TOP-REAL 2)|D) by PRE_TOPC:12;
    A6:[#]((TOP-REAL 2)|D)=D by PRE_TOPC:def 10;
    then A7:p in the carrier of TOP-REAL 2 & not p in {0.REAL 2}
                         by A5,XBOOLE_0:def 4;
    reconsider q=p as Point of TOP-REAL 2 by A5,A6,XBOOLE_0:def 4;
    A8:not p=0.REAL 2 by A7,TARSKI:def 1;
      now per cases;
    case A9: q`1/|.q.|>=cn & q`2<=0;
      set q9= |[ |.q.|* ((q`1/|.q.|-cn)/(1-cn)),
        |.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))]|;
      A10:q9`2= |.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2)) &
      q9`1= |.q.|* ((q`1/|.q.|-cn)/(1-cn))
                                  by EUCLID:56;
        now assume A11: q9=0.REAL 2;
       A12: |.q.|<>0 by A8,TOPRNS_1:25;
        then -sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2)= -sqrt(1-0) by A10,A11,JGRAPH_2
:11,SQUARE_1:60,XCMPLX_1:6
       .=-1 by SQUARE_1:83;
      hence contradiction by A10,A11,A12,JGRAPH_2:11,XCMPLX_1:6;
     end;
     hence f.p<>f.(0.REAL 2) by A1,A2,A8,A9,Th122;
    case A13: q`1/|.q.|<cn & q`2<=0;
      set q9=|[ |.q.|* ((q`1/|.q.|-cn)/(1+cn)),
         |.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))]|;
      A14:q9`2= |.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2)) &
          q9`1= |.q.|* ((q`1/|.q.|-cn)/(1+cn)) by EUCLID:56;
        now assume A15: q9=0.REAL 2;
       A16: |.q.|<>0 by A8,TOPRNS_1:25;
        then -sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2)= -sqrt(1-0) by A14,A15,JGRAPH_2
:11,SQUARE_1:60,XCMPLX_1:6
       .=-1 by SQUARE_1:83;
      hence contradiction by A14,A15,A16,JGRAPH_2:11,XCMPLX_1:6;
     end;
     hence f.p<>f.(0.REAL 2) by A1,A2,A8,A13,Th122;
     case q`2>0; then f.p=p by Th120;
     hence f.p<>f.(0.REAL 2) by A8,Th120,JGRAPH_2:11;
    end;
   hence f.p<>f.(0.REAL 2);
  end;
 A17:ex h being map of (TOP-REAL 2)|D,(TOP-REAL 2)|D
  st h=(cn-FanMorphS)|D & h is continuous by A1,A3,Th138;
   for V being Subset of (TOP-REAL 2) st f.(0.REAL 2) in V & V is open
 ex W being Subset of (TOP-REAL 2) st 0.REAL 2 in W
 & W is open & f.:W c= V
  proof let V be Subset of (TOP-REAL 2);
    assume A18:f.(0.REAL 2) in V & V is open;
     A19:TOP-REAL 2=TopSpaceMetr(Euclid 2) by EUCLID:def 8;
     A20:the carrier of TOP-REAL 2=REAL 2 by EUCLID:25;
     A21:Euclid 2=MetrStruct(#REAL 2,Pitag_dist 2#) by EUCLID:def 7;
     then reconsider u0=0.REAL 2 as Point of Euclid 2 by EUCLID:25;
     consider r being real number such that A22:r>0 & Ball(u0,r) c= V
                                   by A2,A18,A19,TOPMETR:22;
     reconsider r as Real by XREAL_0:def 1;
     reconsider W1=Ball(u0,r) as Subset of TOP-REAL 2
                       by A20,A21;
     A23:u0 in W1 by A22,GOBOARD6:4;
     A24:W1 is open by GOBOARD6:6;
       f.:W1 c= W1
      proof let z be set;assume z in f.:W1;
        then consider y being set such that
        A25: y in dom f & y in W1 & z=f.y by FUNCT_1:def 12;
        reconsider q=y as Point of TOP-REAL 2 by A25;
        reconsider qy=q as Point of Euclid 2 by A21,EUCLID:25;
          z in rng f by A25,FUNCT_1:def 5;
        then reconsider qz=z as Point of TOP-REAL 2;
        reconsider pz=qz as Point of Euclid 2 by A21,EUCLID:25;
          dist(u0,qy)<r by A25,METRIC_1:12;
        then A26: |.(0.REAL 2) - q.|<r by JGRAPH_1:45;
        A27: |.q.|>=0 by TOPRNS_1:26;
          now per cases by JGRAPH_2:11;
        case q`2>=0;
         hence z in W1 by A25,Th120;
        case A28: q<>0.REAL 2 &
          (q`1/|.q.|>=cn & q`2<=0);
          then A29:(cn-FanMorphS).q= |[ |.q.|* ((q`1/|.q.|-cn)/(1-cn)) ,
          |.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))]| by A1,Th122;
            |.q.|<>0 by A28,TOPRNS_1:25;
          then A30: (|.q.|)^2>0 by SQUARE_1:74;
        A31: 1-cn>0 by A1,Th3;
        A32: qz`2= |.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))&
        qz`1= |.q.|* ((q`1/|.q.|-cn)/(1-cn)) by A25,A29,EUCLID:56;
        A33: (q`1/|.q.|-cn)>= 0 by A28,SQUARE_1:12;
        A34: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
          0<=(q`2)^2 by SQUARE_1:72;
        then 0+(q`1)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
        then (q`1)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A30,A34,REAL_1:73;
        then (q`1)^2/(|.q.|)^2 <= 1 by A30,XCMPLX_1:60;
        then ((q`1)/|.q.|)^2 <= 1 by SQUARE_1:69;
        then 1>=q`1/|.q.| by Th4;
        then A35: 1-cn>=q`1/|.q.|-cn by REAL_1:49;
        A36: -(1-cn)<= -0 by A31,REAL_1:50;
           -(1-cn)<= -( q`1/|.q.|-cn) by A35,REAL_1:50;
         then (-(1-cn))/(1-cn)<=(-( q`1/|.q.|-cn))/(1-cn)
                        by A31,REAL_1:73;
        then A37: -1<=(-( q`1/|.q.|-cn))/(1-cn) by A31,XCMPLX_1:198;
         --(1-cn)>= -(q`1/|.q.|-cn) by A33,A36,REAL_1:50;
       then (-(q`1/|.q.|-cn))/(1-cn)<=1 by A31,REAL_2:117;
       then ((-(q`1/|.q.|-cn))/(1-cn))^2<=1^2 by A37,JGRAPH_2:7;
       then 1-((-(q`1/|.q.|-cn))/(1-cn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((q`1/|.q.|-cn))/(1-cn))^2>=0 by XCMPLX_1:188;
       then A38: 1-((q`1/|.q.|-cn)/(1-cn))^2>=0 by SQUARE_1:61;
        A39: (qz`2)^2= (|.q.|)^2*( -sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))^2
                         by A32,SQUARE_1:68
         .= (|.q.|)^2*(sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))^2 by SQUARE_1:61
         .= (|.q.|)^2*(1-((q`1/|.q.|-cn)/(1-cn))^2) by A38,SQUARE_1:def 4;
        A40: (qz`1)^2= (|.q.|)^2*((q`1/|.q.|-cn)/(1-cn))^2
                         by A32,SQUARE_1:68;
          (|.qz.|)^2=(qz`1)^2+(qz`2)^2 by JGRAPH_3:10
        .=(|.q.|)^2*(1-((q`1/|.q.|-cn)/(1-cn))^2
           +((q`1/|.q.|-cn)/(1-cn))^2) by A39,A40,XCMPLX_1:8
        .=(|.q.|)^2*1 by XCMPLX_1:27 .=(|.q.|)^2;
        then A41: sqrt((|.qz.|)^2)=|.q.| by A27,SQUARE_1:89;
            |.qz.|>=0 by TOPRNS_1:26;
          then A42: |.qz.|=|.q.| by A41,SQUARE_1:89;
            |.(0.REAL 2) +- q.|<r by A26,EUCLID:45;
          then |.- q.|<r by EUCLID:31;
          then |.q.|<r by TOPRNS_1:27;
          then |.- qz.|<r by A42,TOPRNS_1:27;
          then |.(0.REAL 2) +- qz.|<r by EUCLID:31;
          then |.(0.REAL 2) - qz.|<r by EUCLID:45;
          then dist(u0,pz)<r by JGRAPH_1:45;
         hence z in W1 by METRIC_1:12;
        case A43: q<>0.REAL 2 &
          q`1/|.q.|<cn & q`2<=0;
          then A44:(cn-FanMorphS).q= |[ |.q.|* ((q`1/|.q.|-cn)/(1+cn)),
          |.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))]| by A1,Th122;
            |.q.|<>0 by A43,TOPRNS_1:25;
          then A45: (|.q.|)^2>0 by SQUARE_1:74;
        A46: (cn-q`1/|.q.|)>=0 by A43,SQUARE_1:12;
        A47: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
        A48: 1+cn>0 by A1,Th3;
        A49: qz`2= |.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))&
        qz`1= |.q.|* ((q`1/|.q.|-cn)/(1+cn)) by A25,A44,EUCLID:56;
          0<=(q`2)^2 by SQUARE_1:72;
        then 0+(q`1)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
        then (q`1)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A45,A47,REAL_1:73;
        then (q`1)^2/(|.q.|)^2 <= 1 by A45,XCMPLX_1:60;
        then ((q`1)/|.q.|)^2 <= 1 by SQUARE_1:69;
        then -1<=q`1/|.q.| by Th4;
        then --1>=-q`1/|.q.| by REAL_1:50;
        then A50: 1+cn>=-q`1/|.q.|+cn by REAL_1:55;
          -(1+cn)<= -0 by A48,REAL_1:50;
         then -(1+cn)<= -( q`1/|.q.|-cn) by A46,XCMPLX_1:143;
         then (-(1+cn))/(1+cn)<=(-( q`1/|.q.|-cn))/(1+cn) by A48,REAL_1:73;
        then A51: -1<=(-( q`1/|.q.|-cn))/(1+cn) by A48,XCMPLX_1:198;
          -(1+cn)<=-(-q`1/|.q.|+cn) by A50,REAL_1:50;
        then -(1+cn)<=--(q`1/|.q.|)-cn by XCMPLX_1:161;
        then --(1+cn)>= -(q`1/|.q.|-cn) by REAL_1:50;
       then (-(q`1/|.q.|-cn))/(1+cn)<=1 by A48,REAL_2:117;
       then ((-(q`1/|.q.|-cn))/(1+cn))^2<=1^2 by A51,JGRAPH_2:7;
       then 1-((-(q`1/|.q.|-cn))/(1+cn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((q`1/|.q.|-cn))/(1+cn))^2>=0 by XCMPLX_1:188;
       then A52: 1-((q`1/|.q.|-cn)/(1+cn))^2>=0 by SQUARE_1:61;
        A53: (qz`2)^2= (|.q.|)^2*( -sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))^2
                         by A49,SQUARE_1:68
         .= (|.q.|)^2*(sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))^2 by SQUARE_1:61
         .= (|.q.|)^2*(1-((q`1/|.q.|-cn)/(1+cn))^2) by A52,SQUARE_1:def 4;
        A54: (qz`1)^2= (|.q.|)^2*((q`1/|.q.|-cn)/(1+cn))^2
                         by A49,SQUARE_1:68;
          (|.qz.|)^2=(qz`1)^2+(qz`2)^2 by JGRAPH_3:10
        .=(|.q.|)^2*(1-((q`1/|.q.|-cn)/(1+cn))^2
           +((q`1/|.q.|-cn)/(1+cn))^2) by A53,A54,XCMPLX_1:8
        .=(|.q.|)^2*1 by XCMPLX_1:27 .=(|.q.|)^2;
        then A55: sqrt((|.qz.|)^2)=|.q.| by A27,SQUARE_1:89;
            |.qz.|>=0 by TOPRNS_1:26;
          then A56: |.qz.|=|.q.| by A55,SQUARE_1:89;
            |.(0.REAL 2) +- q.|<r by A26,EUCLID:45;
          then |.- q.|<r by EUCLID:31;
          then |.q.|<r by TOPRNS_1:27;
          then |.- qz.|<r by A56,TOPRNS_1:27;
          then |.(0.REAL 2) +- qz.|<r by EUCLID:31;
          then |.(0.REAL 2) - qz.|<r by EUCLID:45;
          then dist(u0,pz)<r by JGRAPH_1:45;
         hence z in W1 by METRIC_1:12;
        end;
       hence z in W1;
      end;
     then f.:W1 c= V by A22,XBOOLE_1:1;
   hence ex W being Subset of (TOP-REAL 2) st 0.REAL 2 in W
       & W is open & f.:W c= V by A23,A24;
  end;
  then f is continuous by A2,A3,A4,A17,JGRAPH_3:13;
 hence thesis;
end;

theorem Th140: for cn being Real
  st -1<cn & cn<1 holds cn-FanMorphS is one-to-one
proof let cn be Real;
 assume A1: -1<cn & cn<1;
    for x1,x2 being set st x1 in dom (cn-FanMorphS) & x2 in dom (cn-FanMorphS)
   & (cn-FanMorphS).x1=(cn-FanMorphS).x2 holds x1=x2
  proof let x1,x2 be set;
   assume A2: x1 in dom (cn-FanMorphS) & x2 in dom (cn-FanMorphS)
    & (cn-FanMorphS).x1=(cn-FanMorphS).x2;
    then reconsider p1=x1 as Point of TOP-REAL 2 by FUNCT_2:def 1;
    reconsider p2=x2 as Point of TOP-REAL 2 by A2,FUNCT_2:def 1;
    set q=p1,p=p2;
    A3: 1-cn>0 by A1,Th3;
      now per cases by JGRAPH_2:11;
    case A4: q`2>=0;
      then A5:(cn-FanMorphS).q=q by Th120;
        now per cases by JGRAPH_2:11;
      case p`2>=0;
       hence x1=x2 by A2,A5,Th120;
      case A6:p<>0.REAL 2 &
          (p`1/|.p.|>=cn & p`2<=0);
        then A7:cn-FanMorphS.p= |[ |.p.|* ((p`1/|.p.|-cn)/(1-cn)),
           |.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2))]| by A1,Th122;
         set p4= |[ |.p.|* ((p`1/|.p.|-cn)/(1-cn)),
           |.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2))]|;
        A8: p4`2= |.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2)) &
          p4`1= |.p.|* ((p`1/|.p.|-cn)/(1-cn)) by EUCLID:56;
        A9: q`2=|.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2)) &
          q`1=|.p.|* ((p`1/|.p.|-cn)/(1-cn)) by A2,A5,A7,EUCLID:56;
     A10: |.p.|>=0 by TOPRNS_1:26;
     A11: |.p.|<>0 by A6,TOPRNS_1:25;
     then A12: (|.p.|)^2>0 by SQUARE_1:74;
        A13: (p`1/|.p.|-cn)>=0 by A6,SQUARE_1:12;
        A14: (p`1/|.p.|-cn)>= 0 by A6,SQUARE_1:12;
        A15: (p`1/|.p.|-cn)/(1-cn)>=0 by A3,A13,REAL_2:125;
        A16: (|.p.|)^2 =(p`1)^2+(p`2)^2 by JGRAPH_3:10;
          0<=(p`2)^2 by SQUARE_1:72;
        then 0+(p`1)^2<=(p`1)^2+(p`2)^2 by REAL_1:55;
        then (p`1)^2/(|.p.|)^2 <= (|.p.|)^2/(|.p.|)^2 by A12,A16,REAL_1:73;
        then (p`1)^2/(|.p.|)^2 <= 1 by A12,XCMPLX_1:60;
        then ((p`1)/|.p.|)^2 <= 1 by SQUARE_1:69;
        then 1>=p`1/|.p.| by Th4;
        then A17: 1-cn>=p`1/|.p.|-cn by REAL_1:49;
        A18: -(1-cn)<= -0 by A3,REAL_1:50;
           -(1-cn)<= -( p`1/|.p.|-cn) by A17,REAL_1:50;
         then (-(1-cn))/(1-cn)<=(-( p`1/|.p.|-cn))/(1-cn)
                        by A3,REAL_1:73;
        then A19: -1<=(-( p`1/|.p.|-cn))/(1-cn) by A3,XCMPLX_1:198;
         --(1-cn)>= -(p`1/|.p.|-cn) by A14,A18,REAL_1:50;
       then (-(p`1/|.p.|-cn))/(1-cn)<=1 by A3,REAL_2:117;
       then ((-(p`1/|.p.|-cn))/(1-cn))^2<=1^2 by A19,JGRAPH_2:7;
       then A20: 1-((-(p`1/|.p.|-cn))/(1-cn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((p`1/|.p.|-cn))/(1-cn))^2>=0 by XCMPLX_1:188;
       then A21: 1-((p`1/|.p.|-cn)/(1-cn))^2>=0 by SQUARE_1:61;
         sqrt(1-((-(p`1/|.p.|-cn))/(1-cn))^2)>=0 by A20,SQUARE_1:def 4;
       then sqrt(1-(-(p`1/|.p.|-cn))^2/(1-cn)^2)>=0 by SQUARE_1:69;
       then sqrt(1-(p`1/|.p.|-cn)^2/(1-cn)^2)>=0 by SQUARE_1:61;
       then sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2)>=0 by SQUARE_1:69;
       then -sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2)<=0 by REAL_1:66;
        then q`2=0 by A2,A4,A5,A7,A8,A10,SQUARE_1:23;
        then -sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2)=-0 by A9,A11,XCMPLX_1:6;
        then 1-((p`1/|.p.|-cn)/(1-cn))^2=0 by A21,SQUARE_1:92;
        then 1=0+((p`1/|.p.|-cn)/(1-cn))^2 by XCMPLX_1:27;
        then 1= (p`1/|.p.|-cn)/(1-cn) by A15,SQUARE_1:83,89;
        then (1)*(1-cn)=(p`1/|.p.|-cn) by A3,XCMPLX_1:88;
        then 1-cn+cn=p`1/|.p.| by XCMPLX_1:27;
        then 1=p`1/|.p.| by XCMPLX_1:27;
        then (1)*|.p.|=p`1 by A11,XCMPLX_1:88;
        then (p`1)^2-(p`1)^2 =(p`2)^2 by A16,XCMPLX_1:26;
        then 0= (p`2)^2 by XCMPLX_1:14;
        then p`2=0 by SQUARE_1:73;
       hence x1=x2 by A2,A5,Th120;
      case A22:p<>0.REAL 2 &
        (p`1/|.p.|<cn & p`2<=0);
        then A23:
         cn-FanMorphS.p= |[ |.p.|* ((p`1/|.p.|-cn)/(1+cn)),
           |.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2))]| by A1,Th122;
         set p4= |[ |.p.|* ((p`1/|.p.|-cn)/(1+cn)),
           |.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2))]|;
        A24: p4`2= |.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2)) &
          p4`1= |.p.|* ((p`1/|.p.|-cn)/(1+cn)) by EUCLID:56;
        A25: q`2=|.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2)) &
          q`1=|.p.|* ((p`1/|.p.|-cn)/(1+cn)) by A2,A5,A23,EUCLID:56;
     A26: |.p.|>=0 by TOPRNS_1:26;
     A27: |.p.|<>0 by A22,TOPRNS_1:25;
     then A28: (|.p.|)^2>0 by SQUARE_1:74;
     A29: 1+cn>0 by A1,Th3;
        A30: (p`1/|.p.|-cn)<=0 by A22,REAL_2:106;
        then A31: -(p`1/|.p.|-cn)>= -0 by REAL_1:50;
          --(p`1/|.p.|-cn)/(1+cn)<=0 by A29,A30,REAL_2:126;
        then A32: -((p`1/|.p.|-cn)/(1+cn))>=0 by REAL_1:66;
        A33: (|.p.|)^2 =(p`1)^2+(p`2)^2 by JGRAPH_3:10;
          0<=(p`2)^2 by SQUARE_1:72;
        then 0+(p`1)^2<=(p`1)^2+(p`2)^2 by REAL_1:55;
        then (p`1)^2/(|.p.|)^2 <= (|.p.|)^2/(|.p.|)^2 by A28,A33,REAL_1:73;
        then (p`1)^2/(|.p.|)^2 <= 1 by A28,XCMPLX_1:60;
        then ((p`1)/|.p.|)^2 <= 1 by SQUARE_1:69;
        then A34: (-((p`1)/|.p.|))^2 <= 1 by SQUARE_1:61;
          -(1+cn)<= -0 by A29,REAL_1:50;
        then (-(1+cn))/(1+cn)<=(-( p`1/|.p.|-cn))/(1+cn) by A29,A31,REAL_1:73;
        then A35: -1<=(-( p`1/|.p.|-cn))/(1+cn) by A29,XCMPLX_1:198;
         1>= -p`1/|.p.| by A34,Th4;
       then (1+cn)>= -p`1/|.p.|+cn by REAL_1:55;
       then (1+cn)>= -(p`1/|.p.|-cn) by XCMPLX_1:162;
       then (-(p`1/|.p.|-cn))/(1+cn)<=1 by A29,REAL_2:117;
       then ((-(p`1/|.p.|-cn))/(1+cn))^2<=1^2 by A35,JGRAPH_2:7;
       then A36: 1-((-(p`1/|.p.|-cn))/(1+cn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((p`1/|.p.|-cn))/(1+cn))^2>=0 by XCMPLX_1:188;
       then A37: 1-((p`1/|.p.|-cn)/(1+cn))^2>=0 by SQUARE_1:61;
         sqrt(1-((-(p`1/|.p.|-cn))/(1+cn))^2)>=0 by A36,SQUARE_1:def 4;
       then sqrt(1-(-(p`1/|.p.|-cn))^2/(1+cn)^2)>=0 by SQUARE_1:69;
       then sqrt(1-((p`1/|.p.|-cn))^2/(1+cn)^2)>=0 by SQUARE_1:61;
       then sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2)>=0 by SQUARE_1:69;
       then -sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2)<=0 by REAL_1:66;
        then q`2=0 by A2,A4,A5,A23,A24,A26,SQUARE_1:23;
        then -sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2)=-0 by A25,A27,XCMPLX_1:6;
        then 1-((p`1/|.p.|-cn)/(1+cn))^2=0 by A37,SQUARE_1:92;
        then 1=0+((p`1/|.p.|-cn)/(1+cn))^2 by XCMPLX_1:27;
        then 1=(-((p`1/|.p.|-cn)/(1+cn)))^2 by SQUARE_1:61;
        then 1= -((p`1/|.p.|-cn)/(1+cn)) by A32,SQUARE_1:83,89;
        then 1= ((-(p`1/|.p.|-cn))/(1+cn)) by XCMPLX_1:188;
        then (1)*(1+cn)=-(p`1/|.p.|-cn) by A29,XCMPLX_1:88;
        then (1+cn)=-p`1/|.p.|+cn by XCMPLX_1:162;
        then 1+cn-cn=-p`1/|.p.| by XCMPLX_1:26;
        then 1=-p`1/|.p.| by XCMPLX_1:26;
        then 1=(-p`1)/|.p.| by XCMPLX_1:188;
        then (1)*|.p.|=-p`1 by A27,XCMPLX_1:88;
        then (-p`1)^2-(p`1)^2 =(p`2)^2 by A33,XCMPLX_1:26;
        then (p`1)^2-(p`1)^2 =(p`2)^2 by SQUARE_1:61;
        then 0= (p`2)^2 by XCMPLX_1:14;
        then p`2=0 by SQUARE_1:73;
       hence x1=x2 by A2,A5,Th120;
      end;
     hence x1=x2;
    case A38:
      q`1/|.q.|>=cn & q`2<=0 & q<>0.REAL 2;
        then A39:
         cn-FanMorphS.q= |[ |.q.|* ((q`1/|.q.|-cn)/(1-cn)),
           |.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))]| by A1,Th122;
         set q4= |[ |.q.|* ((q`1/|.q.|-cn)/(1-cn)),
           |.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))]|;
        A40: q4`2= |.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2)) &
          q4`1= |.q.|* ((q`1/|.q.|-cn)/(1-cn)) by EUCLID:56;
     A41: |.q.|>=0 by TOPRNS_1:26;
       |.q.|<>0 by A38,TOPRNS_1:25;
     then A42: (|.q.|)^2>0 by SQUARE_1:74;
        now per cases by JGRAPH_2:11;
      case A43: p`2>=0;
        then A44: (cn-FanMorphS).p=p by Th120;
        then A45: p`2=|.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2)) &
        p`1=|.q.|* ((q`1/|.q.|-cn)/(1-cn)) by A2,A39,EUCLID:56;
        A46: |.q.|>=0 by TOPRNS_1:26;
        A47: |.q.|<>0 by A38,TOPRNS_1:25;
        then A48: (|.q.|)^2>0 by SQUARE_1:74;
        A49: 1-cn>0 by A1,Th3;
        A50: (q`1/|.q.|-cn)>=0 by A38,SQUARE_1:12;
        A51: (q`1/|.q.|-cn)>= 0 by A38,SQUARE_1:12;
        A52: (q`1/|.q.|-cn)/(1-cn)>=0 by A49,A50,REAL_2:125;
        A53: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
          0<=(q`2)^2 by SQUARE_1:72;
        then 0+(q`1)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
        then (q`1)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A48,A53,REAL_1:73;
        then (q`1)^2/(|.q.|)^2 <= 1 by A48,XCMPLX_1:60;
        then ((q`1)/|.q.|)^2 <= 1 by SQUARE_1:69;
        then 1>=q`1/|.q.| by Th4;
        then A54: 1-cn>=q`1/|.q.|-cn by REAL_1:49;
        A55: -(1-cn)<= -0 by A49,REAL_1:50;
           -(1-cn)<= -( q`1/|.q.|-cn) by A54,REAL_1:50;
         then (-(1-cn))/(1-cn)<=(-( q`1/|.q.|-cn))/(1-cn)
                        by A49,REAL_1:73;
        then A56: -1<=(-( q`1/|.q.|-cn))/(1-cn) by A49,XCMPLX_1:198;
         --(1-cn)>= -(q`1/|.q.|-cn) by A51,A55,REAL_1:50;
       then (-(q`1/|.q.|-cn))/(1-cn)<=1 by A49,REAL_2:117;
       then ((-(q`1/|.q.|-cn))/(1-cn))^2<=1^2 by A56,JGRAPH_2:7;
       then A57: 1-((-(q`1/|.q.|-cn))/(1-cn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((q`1/|.q.|-cn))/(1-cn))^2>=0 by XCMPLX_1:188;
       then A58: 1-((q`1/|.q.|-cn)/(1-cn))^2>=0 by SQUARE_1:61;
         sqrt(1-((-(q`1/|.q.|-cn))/(1-cn))^2)>=0 by A57,SQUARE_1:def 4;
       then sqrt(1-(-(q`1/|.q.|-cn))^2/(1-cn)^2)>=0 by SQUARE_1:69;
       then sqrt(1-(q`1/|.q.|-cn)^2/(1-cn)^2)>=0 by SQUARE_1:61;
       then sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2)>=0 by SQUARE_1:69;
       then -sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2)<=0 by REAL_1:66;
        then p`2=0 by A2,A39,A40,A43,A44,A46,SQUARE_1:23;
        then -sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2)=-0 by A45,A47,XCMPLX_1:6;
        then 1-((q`1/|.q.|-cn)/(1-cn))^2=0 by A58,SQUARE_1:92;
        then 1=0+((q`1/|.q.|-cn)/(1-cn))^2 by XCMPLX_1:27;
        then 1= (q`1/|.q.|-cn)/(1-cn) by A52,SQUARE_1:83,89;
        then (1)*(1-cn)=(q`1/|.q.|-cn) by A49,XCMPLX_1:88;
        then 1-cn+cn=q`1/|.q.| by XCMPLX_1:27;
        then 1=q`1/|.q.| by XCMPLX_1:27;
        then (1)*|.q.|=q`1 by A47,XCMPLX_1:88;
        then (q`1)^2-(q`1)^2 =(q`2)^2 by A53,XCMPLX_1:26;
        then 0= (q`2)^2 by XCMPLX_1:14;
        then q`2=0 by SQUARE_1:73;
       hence x1=x2 by A2,A44,Th120;
      case A59:p<>0.REAL 2 & (p`1/|.p.|>=cn & p`2<=0);
        then A60: cn-FanMorphS.p= |[ |.p.|* ((p`1/|.p.|-cn)/(1-cn)),
           |.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2))]| by A1,Th122;
         set p4= |[ |.p.|* ((p`1/|.p.|-cn)/(1-cn)),
           |.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2))]|;
        A61: p4`2= |.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2)) &
          p4`1= |.p.|* ((p`1/|.p.|-cn)/(1-cn)) by EUCLID:56;
        A62: |.q4.|>=0 by TOPRNS_1:26;
        A63: q4`2= |.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))&
        q4`1= |.q.|* ((q`1/|.q.|-cn)/(1-cn))
                     by EUCLID:56;
        A64: (q`1/|.q.|-cn)>= 0 by A38,SQUARE_1:12;
        A65: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
          0<=(q`2)^2 by SQUARE_1:72;
        then 0+(q`1)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
        then (q`1)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A42,A65,REAL_1:73;
        then (q`1)^2/(|.q.|)^2 <= 1 by A42,XCMPLX_1:60;
        then ((q`1)/|.q.|)^2 <= 1 by SQUARE_1:69;
        then 1>=q`1/|.q.| by Th4;
        then A66: 1-cn>=q`1/|.q.|-cn by REAL_1:49;
        A67: -(1-cn)<= -0 by A3,REAL_1:50;
           -(1-cn)<= -( q`1/|.q.|-cn) by A66,REAL_1:50;
         then (-(1-cn))/(1-cn)<=(-( q`1/|.q.|-cn))/(1-cn)
                        by A3,REAL_1:73;
        then A68: -1<=(-( q`1/|.q.|-cn))/(1-cn)
                        by A3,XCMPLX_1:198;
         --(1-cn)>= -(q`1/|.q.|-cn) by A64,A67,REAL_1:50;
       then (-(q`1/|.q.|-cn))/(1-cn)<=1 by A3,REAL_2:117;
       then ((-(q`1/|.q.|-cn))/(1-cn))^2<=1^2 by A68,JGRAPH_2:7;
       then 1-((-(q`1/|.q.|-cn))/(1-cn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((q`1/|.q.|-cn))/(1-cn))^2>=0 by XCMPLX_1:188;
       then A69: 1-((q`1/|.q.|-cn)/(1-cn))^2>=0 by SQUARE_1:61;
        A70: (q4`2)^2= (|.q.|)^2*( -sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))^2
                         by A63,SQUARE_1:68
         .= (|.q.|)^2*(sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))^2 by SQUARE_1:61
         .= (|.q.|)^2*(1-((q`1/|.q.|-cn)/(1-cn))^2) by A69,SQUARE_1:def 4;
        A71: (q4`1)^2= (|.q.|)^2*((q`1/|.q.|-cn)/(1-cn))^2
                         by A63,SQUARE_1:68;
          (|.q4.|)^2=(q4`1)^2+(q4`2)^2 by JGRAPH_3:10
        .=(|.q.|)^2*(1-((q`1/|.q.|-cn)/(1-cn))^2
           +((q`1/|.q.|-cn)/(1-cn))^2) by A70,A71,XCMPLX_1:8
        .=(|.q.|)^2*1 by XCMPLX_1:27 .=(|.q.|)^2;
        then A72: sqrt((|.q4.|)^2)=|.q.| by A41,SQUARE_1:89;
        then A73: |.q4.|=|.q.| by A62,SQUARE_1:89;
        A74: |.p4.|>=0 by TOPRNS_1:26;
        A75: |.p.|>=0 by TOPRNS_1:26;
        A76: |.p.|<>0 by A59,TOPRNS_1:25;
        then A77: (|.p.|)^2>0 by SQUARE_1:74;
        A78: (p`1/|.p.|-cn)>= 0 by A59,SQUARE_1:12;
        A79: (|.p.|)^2 =(p`1)^2+(p`2)^2 by JGRAPH_3:10;
          0<=(p`2)^2 by SQUARE_1:72;
        then 0+(p`1)^2<=(p`1)^2+(p`2)^2 by REAL_1:55;
        then (p`1)^2/(|.p.|)^2 <= (|.p.|)^2/(|.p.|)^2 by A77,A79,REAL_1:73;
        then (p`1)^2/(|.p.|)^2 <= 1 by A77,XCMPLX_1:60;
        then ((p`1)/|.p.|)^2 <= 1 by SQUARE_1:69;
        then 1>=p`1/|.p.| by Th4;
        then A80: 1-cn>=p`1/|.p.|-cn by REAL_1:49;
        A81: -(1-cn)<= -0 by A3,REAL_1:50;
           -(1-cn)<= -( p`1/|.p.|-cn) by A80,REAL_1:50;
         then (-(1-cn))/(1-cn)<=(-( p`1/|.p.|-cn))/(1-cn)
                        by A3,REAL_1:73;
        then A82: -1<=(-( p`1/|.p.|-cn))/(1-cn)
                        by A3,XCMPLX_1:198;
         --(1-cn)>= -(p`1/|.p.|-cn) by A78,A81,REAL_1:50;
       then (-(p`1/|.p.|-cn))/(1-cn)<=1 by A3,REAL_2:117;
       then ((-(p`1/|.p.|-cn))/(1-cn))^2<=1^2 by A82,JGRAPH_2:7;
       then 1-((-(p`1/|.p.|-cn))/(1-cn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((p`1/|.p.|-cn))/(1-cn))^2>=0 by XCMPLX_1:188;
       then A83: 1-((p`1/|.p.|-cn)/(1-cn))^2>=0 by SQUARE_1:61;
        A84: (p4`2)^2= (|.p.|)^2*( -sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2))^2
                         by A61,SQUARE_1:68
         .= (|.p.|)^2*(sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2))^2 by SQUARE_1:61
         .= (|.p.|)^2*(1-((p`1/|.p.|-cn)/(1-cn))^2) by A83,SQUARE_1:def 4;
        A85: (p4`1)^2= (|.p.|)^2*((p`1/|.p.|-cn)/(1-cn))^2
                         by A61,SQUARE_1:68;
          (|.p4.|)^2=(p4`1)^2+(p4`2)^2 by JGRAPH_3:10
        .=(|.p.|)^2*(1-((p`1/|.p.|-cn)/(1-cn))^2
           +((p`1/|.p.|-cn)/(1-cn))^2) by A84,A85,XCMPLX_1:8
        .=(|.p.|)^2*1 by XCMPLX_1:27 .=(|.p.|)^2;
        then A86: sqrt((|.p4.|)^2)=|.p.| by A75,SQUARE_1:89;
        then A87: |.p4.|=|.p.| by A74,SQUARE_1:89;
        ((p`1/|.p.|-cn)/(1-cn))
        =|.q.|* ((q`1/|.q.|-cn)/(1-cn))/|.p.| by A2,A39,A40,A60,A61,A76,
XCMPLX_1:90;
        then (p`1/|.p.|-cn)/(1-cn)=(q`1/|.q.|-cn)/(1-cn)
                              by A2,A39,A60,A72,A76,A86,XCMPLX_1:90;
        then (p`1/|.p.|-cn)/(1-cn)*(1-cn)=q`1/|.q.|-cn by A3,XCMPLX_1:88;
        then p`1/|.p.|-cn=q`1/|.q.|-cn by A3,XCMPLX_1:88;
        then p`1/|.p.|-cn+cn=q`1/|.q.| by XCMPLX_1:27;
        then p`1/|.p.|=q`1/|.q.| by XCMPLX_1:27;
        then p`1/|.p.|*|.p.|=q`1 by A2,A39,A60,A73,A76,A87,XCMPLX_1:88;
        then A88: p`1=q`1 by A76,XCMPLX_1:88;
        A89: |.p.|^2=(p`1)^2+(p`2)^2 by JGRAPH_3:10;
          |.q.|^2=(q`1)^2+(q`2)^2 by JGRAPH_3:10;
        then (p`1)^2+(p`2)^2-(p`1)^2=(q`2)^2 by A2,A39,A60,A73,A87,A88,A89,
XCMPLX_1:26;
        then (p`2)^2=(q`2)^2 by XCMPLX_1:26;
        then (-p`2)^2=(q`2)^2 by SQUARE_1:61;
        then A90: (-p`2)^2=(-q`2)^2 by SQUARE_1:61;
          --q`2<=0 by A38;
        then A91: -q`2>=0 by REAL_1:66;
          --p`2<=0 by A59; then -p`2>=0 by REAL_1:66;
        then -p`2=sqrt((-q`2)^2) by A90,SQUARE_1:89;
        then A92: --p`2=--q`2 by A91,SQUARE_1:89;
          p=|[p`1,p`2]| by EUCLID:57;
       hence x1=x2 by A88,A92,EUCLID:57;
      case A93: p<>0.REAL 2 &
        (p`1/|.p.|<cn & p`2<=0);
        then A94:
         cn-FanMorphS.p= |[ |.p.|* ((p`1/|.p.|-cn)/(1+cn)),
           |.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2))]| by A1,Th122;
         set p4= |[ |.p.|* ((p`1/|.p.|-cn)/(1+cn)),
           |.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2))]|;
        A95: p4`2= |.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2)) &
          p4`1= |.p.|* ((p`1/|.p.|-cn)/(1+cn)) by EUCLID:56;
         A96: p`1/|.p.|-cn<0 by A93,REAL_2:105;
           1+cn>0 by A1,Th3;
         then A97: ((p`1/|.p.|-cn)/(1+cn))<0 by A96,REAL_2:128;
        A98: |.p.|>=0 by TOPRNS_1:26;
          |.p.|<>0 by A93,TOPRNS_1:25;
         then A99: |.p.|* ((p`1/|.p.|-cn)/(1+cn))<0 by A97,A98,SQUARE_1:24;
         A100: q`1/|.q.|-cn>=0 by A38,SQUARE_1:12;
           1-cn>0 by A1,Th3;
         then ((q`1/|.q.|-cn)/(1-cn))>=0 by A100,REAL_2:125;
       hence x1=x2 by A2,A39,A40,A41,A94,A95,A99,REAL_2:121;
      end;
     hence x1=x2;
    case A101:
      q`1/|.q.|<cn & q`2<=0 & q<>0.REAL 2;
        then A102:
         cn-FanMorphS.q= |[ |.q.|* ((q`1/|.q.|-cn)/(1+cn)),
           |.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))]| by A1,Th122;
         set q4= |[ |.q.|* ((q`1/|.q.|-cn)/(1+cn)),
           |.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))]|;
        A103: q4`2= |.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2)) &
          q4`1= |.q.|* ((q`1/|.q.|-cn)/(1+cn)) by EUCLID:56;
     A104: |.q.|>=0 by TOPRNS_1:26;
     A105: |.q.|<>0 by A101,TOPRNS_1:25;
     then A106: (|.q.|)^2>0 by SQUARE_1:74;
        now per cases by JGRAPH_2:11;
      case A107: p`2>=0;
        then A108:(cn-FanMorphS).p=p by Th120;
        then A109: p`2=|.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2)) &
          p`1=|.q.|* ((q`1/|.q.|-cn)/(1+cn)) by A2,A102,EUCLID:56;
        A110: 1+cn>0 by A1,Th3;
        A111: (q`1/|.q.|-cn)<=0 by A101,REAL_2:106;
        then A112: -(q`1/|.q.|-cn)>= -0 by REAL_1:50;
          --(q`1/|.q.|-cn)/(1+cn)<=0 by A110,A111,REAL_2:126;
        then A113: -((q`1/|.q.|-cn)/(1+cn))>=0 by REAL_1:66;
        A114: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
          0<=(q`2)^2 by SQUARE_1:72;
        then 0+(q`1)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
        then (q`1)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A106,A114,REAL_1:73;
        then (q`1)^2/(|.q.|)^2 <= 1 by A106,XCMPLX_1:60;
        then ((q`1)/|.q.|)^2 <= 1 by SQUARE_1:69;
        then A115: (-((q`1)/|.q.|))^2 <= 1 by SQUARE_1:61;
          -(1+cn)<= -0 by A110,REAL_1:50;
        then (-(1+cn))/(1+cn)<=(-( q`1/|.q.|-cn))/(1+cn)
                        by A110,A112,REAL_1:73;
        then A116: -1<=(-( q`1/|.q.|-cn))/(1+cn) by A110,XCMPLX_1:198;
         1>= -q`1/|.q.| by A115,Th4;
       then (1+cn)>= -q`1/|.q.|+cn by REAL_1:55;
       then (1+cn)>= -(q`1/|.q.|-cn) by XCMPLX_1:162;
       then (-(q`1/|.q.|-cn))/(1+cn)<=1 by A110,REAL_2:117;
       then ((-(q`1/|.q.|-cn))/(1+cn))^2<=1^2 by A116,JGRAPH_2:7;
       then A117: 1-((-(q`1/|.q.|-cn))/(1+cn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((q`1/|.q.|-cn))/(1+cn))^2>=0 by XCMPLX_1:188;
       then A118: 1-((q`1/|.q.|-cn)/(1+cn))^2>=0 by SQUARE_1:61;
         sqrt(1-((-(q`1/|.q.|-cn))/(1+cn))^2)>=0 by A117,SQUARE_1:def 4;
       then sqrt(1-(-(q`1/|.q.|-cn))^2/(1+cn)^2)>=0 by SQUARE_1:69;
       then sqrt(1-((q`1/|.q.|-cn))^2/(1+cn)^2)>=0 by SQUARE_1:61;
       then sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2)>=0 by SQUARE_1:69;
       then -sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2)<=0 by REAL_1:66;
        then p`2=0 by A2,A102,A103,A104,A107,A108,SQUARE_1:23;
        then -sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2)=-0 by A105,A109,XCMPLX_1:6;
        then 1-((q`1/|.q.|-cn)/(1+cn))^2=0 by A118,SQUARE_1:92;
        then 1=0+((q`1/|.q.|-cn)/(1+cn))^2 by XCMPLX_1:27;
        then 1=(-((q`1/|.q.|-cn)/(1+cn)))^2 by SQUARE_1:61;
        then 1= -((q`1/|.q.|-cn)/(1+cn)) by A113,SQUARE_1:83,89;
        then 1= ((-(q`1/|.q.|-cn))/(1+cn)) by XCMPLX_1:188;
        then (1)*(1+cn)=-(q`1/|.q.|-cn) by A110,XCMPLX_1:88;
        then (1+cn)=-q`1/|.q.|+cn by XCMPLX_1:162;
        then 1+cn-cn=-q`1/|.q.| by XCMPLX_1:26;
        then 1=-q`1/|.q.| by XCMPLX_1:26;
        then 1=(-q`1)/|.q.| by XCMPLX_1:188;
        then (1)*|.q.|=-q`1 by A105,XCMPLX_1:88;
        then (-q`1)^2-(q`1)^2 =(q`2)^2 by A114,XCMPLX_1:26;
        then (q`1)^2-(q`1)^2 =(q`2)^2 by SQUARE_1:61;
        then 0= (q`2)^2 by XCMPLX_1:14;
        then q`2=0 by SQUARE_1:73;
       hence x1=x2 by A2,A108,Th120;
      case A119:p<>0.REAL 2 & (p`1/|.p.|>=cn & p`2<=0);
        then A120: cn-FanMorphS.p= |[ |.p.|* ((p`1/|.p.|-cn)/(1-cn)),
           |.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2))]| by A1,Th122;
         set p4= |[ |.p.|* ((p`1/|.p.|-cn)/(1-cn)),
           |.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2))]|;
        A121: p4`2= |.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1-cn))^2)) &
          p4`1= |.p.|* ((p`1/|.p.|-cn)/(1-cn)) by EUCLID:56;
         A122: q`1/|.q.|-cn<0 by A101,REAL_2:105;
           1+cn>0 by A1,Th3;
         then A123: ((q`1/|.q.|-cn)/(1+cn))<0 by A122,REAL_2:128;
        A124: |.q.|>=0 by TOPRNS_1:26;
          |.q.|<>0 by A101,TOPRNS_1:25;
         then A125: |.q.|* ((q`1/|.q.|-cn)/(1+cn))<0 by A123,A124,SQUARE_1:24;
         A126: p`1/|.p.|-cn>=0 by A119,SQUARE_1:12;
           1-cn>0 by A1,Th3;
         then A127: ((p`1/|.p.|-cn)/(1-cn))>=0 by A126,REAL_2:125;
           |.p.|>=0 by TOPRNS_1:26;
       hence x1=x2 by A2,A102,A103,A120,A121,A125,A127,REAL_2:121;
      case A128:p<>0.REAL 2 &
          (p`1/|.p.|<cn & p`2<=0);
        then A129:
         cn-FanMorphS.p= |[ |.p.|* ((p`1/|.p.|-cn)/(1+cn)),
           |.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2))]| by A1,Th122;
         set p4= |[ |.p.|* ((p`1/|.p.|-cn)/(1+cn)),
           |.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2))]|;
        A130: p4`2= |.p.|*( -sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2)) &
          p4`1= |.p.|* ((p`1/|.p.|-cn)/(1+cn)) by EUCLID:56;
        A131: |.q4.|>=0 by TOPRNS_1:26;
        A132: q4`2= |.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))&
        q4`1= |.q.|* ((q`1/|.q.|-cn)/(1+cn)) by EUCLID:56;
          (q`1/|.q.|-cn)<=0 by A101,REAL_2:106;
        then A133: -(q`1/|.q.|-cn)>= -0 by REAL_1:50;
        A134: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
          0<=(q`2)^2 by SQUARE_1:72;
        then 0+(q`1)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
        then (q`1)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A106,A134,REAL_1:73;
        then (q`1)^2/(|.q.|)^2 <= 1 by A106,XCMPLX_1:60;
        then ((q`1)/|.q.|)^2 <= 1 by SQUARE_1:69;
        then -1<=q`1/|.q.| by Th4;
        then A135: -1-cn<=q`1/|.q.|-cn by REAL_1:49;
        A136: 1+cn>0 by A1,Th3;
        then -(1+cn)<= -0 by REAL_1:50;
        then (-(1+cn))/(1+cn)<=(-( q`1/|.q.|-cn))/(1+cn)
                        by A133,A136,REAL_1:73;
        then A137: -1<=(-( q`1/|.q.|-cn))/(1+cn) by A136,XCMPLX_1:198;
          -(1+cn)<=q`1/|.q.|-cn by A135,XCMPLX_1:161;
        then --(1+cn)>= -(q`1/|.q.|-cn) by REAL_1:50;
       then (-(q`1/|.q.|-cn))/(1+cn)<=1 by A136,REAL_2:117;
       then ((-(q`1/|.q.|-cn))/(1+cn))^2<=1^2 by A137,JGRAPH_2:7;
       then 1-((-(q`1/|.q.|-cn))/(1+cn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((q`1/|.q.|-cn))/(1+cn))^2>=0 by XCMPLX_1:188;
       then A138: 1-((q`1/|.q.|-cn)/(1+cn))^2>=0 by SQUARE_1:61;
        A139: (q4`2)^2= (|.q.|)^2*( -sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))^2
                         by A132,SQUARE_1:68
         .= (|.q.|)^2*(sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))^2 by SQUARE_1:61
         .= (|.q.|)^2*(1-((q`1/|.q.|-cn)/(1+cn))^2) by A138,SQUARE_1:def 4;
        A140: (q4`1)^2= (|.q.|)^2*((q`1/|.q.|-cn)/(1+cn))^2
                         by A132,SQUARE_1:68;
          (|.q4.|)^2=(q4`1)^2+(q4`2)^2 by JGRAPH_3:10
        .=(|.q.|)^2*(1-((q`1/|.q.|-cn)/(1+cn))^2
           +((q`1/|.q.|-cn)/(1+cn))^2) by A139,A140,XCMPLX_1:8
        .=(|.q.|)^2*1 by XCMPLX_1:27 .=(|.q.|)^2;
        then A141: sqrt((|.q4.|)^2)=|.q.| by A104,SQUARE_1:89;
        then A142: |.q4.|=|.q.| by A131,SQUARE_1:89;
        A143: |.p4.|>=0 by TOPRNS_1:26;
        A144: |.p.|>=0 by TOPRNS_1:26;
        A145: |.p.|<>0 by A128,TOPRNS_1:25;
        then A146: (|.p.|)^2>0 by SQUARE_1:74;
          (p`1/|.p.|-cn)<=0 by A128,REAL_2:106;
        then A147: -(p`1/|.p.|-cn)>= -0 by REAL_1:50;
        A148: (|.p.|)^2 =(p`1)^2+(p`2)^2 by JGRAPH_3:10;
          0<=(p`2)^2 by SQUARE_1:72;
        then 0+(p`1)^2<=(p`1)^2+(p`2)^2 by REAL_1:55;
        then (p`1)^2/(|.p.|)^2 <= (|.p.|)^2/(|.p.|)^2 by A146,A148,REAL_1:73;
        then (p`1)^2/(|.p.|)^2 <= 1 by A146,XCMPLX_1:60;
        then ((p`1)/|.p.|)^2 <= 1 by SQUARE_1:69;
        then -1<=p`1/|.p.| by Th4;
        then A149: -1-cn<=p`1/|.p.|-cn by REAL_1:49;
          -(1+cn)<= -0 by A136,REAL_1:50;
        then (-(1+cn))/(1+cn)<=(-( p`1/|.p.|-cn))/(1+cn)
                        by A136,A147,REAL_1:73;
        then A150: -1<=(-( p`1/|.p.|-cn))/(1+cn) by A136,XCMPLX_1:198;
          -(1+cn)<=p`1/|.p.|-cn by A149,XCMPLX_1:161;
        then --(1+cn)>= -(p`1/|.p.|-cn) by REAL_1:50;
       then (-(p`1/|.p.|-cn))/(1+cn)<=1 by A136,REAL_2:117;
       then ((-(p`1/|.p.|-cn))/(1+cn))^2<=1^2 by A150,JGRAPH_2:7;
       then 1-((-(p`1/|.p.|-cn))/(1+cn))^2>=0 by SQUARE_1:12,59;
       then 1-(-((p`1/|.p.|-cn))/(1+cn))^2>=0 by XCMPLX_1:188;
       then A151: 1-((p`1/|.p.|-cn)/(1+cn))^2>=0 by SQUARE_1:61;
        A152: (p4`2)^2= (|.p.|)^2*( -sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2))^2
                         by A130,SQUARE_1:68
         .= (|.p.|)^2*(sqrt(1-((p`1/|.p.|-cn)/(1+cn))^2))^2 by SQUARE_1:61
         .= (|.p.|)^2*(1-((p`1/|.p.|-cn)/(1+cn))^2)
                      by A151,SQUARE_1:def 4;
        A153: (p4`1)^2= (|.p.|)^2*((p`1/|.p.|-cn)/(1+cn))^2
                         by A130,SQUARE_1:68;
          (|.p4.|)^2=(p4`1)^2+(p4`2)^2 by JGRAPH_3:10
        .=(|.p.|)^2*(1-((p`1/|.p.|-cn)/(1+cn))^2
           +((p`1/|.p.|-cn)/(1+cn))^2) by A152,A153,XCMPLX_1:8
        .=(|.p.|)^2*1 by XCMPLX_1:27 .=(|.p.|)^2;
        then A154: sqrt((|.p4.|)^2)=|.p.| by A144,SQUARE_1:89;
        then A155: |.p4.|=|.p.| by A143,SQUARE_1:89;
        ((p`1/|.p.|-cn)/(1+cn))
        =|.q.|* ((q`1/|.q.|-cn)/(1+cn))/|.p.| by A2,A102,A103,A129,A130,A145,
XCMPLX_1:90;
        then (p`1/|.p.|-cn)/(1+cn)=(q`1/|.q.|-cn)/(1+cn)
                              by A2,A102,A129,A141,A145,A154,XCMPLX_1:90
;
        then (p`1/|.p.|-cn)/(1+cn)*(1+cn)=q`1/|.q.|-cn by A136,XCMPLX_1:88;
        then p`1/|.p.|-cn=q`1/|.q.|-cn by A136,XCMPLX_1:88;
        then p`1/|.p.|-cn+cn=q`1/|.q.| by XCMPLX_1:27;
        then p`1/|.p.|=q`1/|.q.| by XCMPLX_1:27;
        then p`1/|.p.|*|.p.|=q`1 by A2,A102,A129,A142,A145,A155,XCMPLX_1:88;
        then A156: p`1=q`1 by A145,XCMPLX_1:88;
        A157: |.p.|^2=(p`1)^2+(p`2)^2 by JGRAPH_3:10;
          |.q.|^2=(q`1)^2+(q`2)^2 by JGRAPH_3:10;
        then (p`2)^2+(p`1)^2-(p`1)^2=(q`2)^2 by A2,A102,A129,A142,A155,A156,
A157,XCMPLX_1:26;
        then (p`2)^2=(q`2)^2 by XCMPLX_1:26;
        then (-p`2)^2=(q`2)^2 by SQUARE_1:61;
        then A158: (-p`2)^2=(-q`2)^2 by SQUARE_1:61;
          --q`2<=0 by A101;
        then A159: -q`2>=0 by REAL_1:66;
          --p`2<=0 by A128; then -p`2>=0 by REAL_1:66;
        then -p`2=sqrt((-q`2)^2) by A158,SQUARE_1:89;
        then A160: --p`2=--q`2 by A159,SQUARE_1:89;
          p=|[p`1,p`2]| by EUCLID:57;
       hence x1=x2 by A156,A160,EUCLID:57;
      end;
     hence x1=x2;
    end;
   hence x1=x2;
  end;
 hence thesis by FUNCT_1:def 8;
end;

theorem Th141: for cn being Real
  st -1<cn & cn<1 holds
  (cn-FanMorphS) is map of TOP-REAL 2,TOP-REAL 2 &
 rng (cn-FanMorphS) = the carrier of TOP-REAL 2
proof let cn be Real;
  assume A1:  -1<cn & cn<1;
thus (cn-FanMorphS) is map of TOP-REAL 2,TOP-REAL 2 ;
  A2:for f being map of TOP-REAL 2,TOP-REAL 2 st f=(cn-FanMorphS) holds
    rng (cn-FanMorphS)=the carrier of TOP-REAL 2
  proof let f be map of TOP-REAL 2,TOP-REAL 2;
    assume A3:f=(cn-FanMorphS);
    A4:dom f=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
      the carrier of TOP-REAL 2 c= rng f
    proof let y be set;assume y in the carrier of TOP-REAL 2;
     then reconsider p2=y as Point of TOP-REAL 2;
     set q=p2;
       now per cases by JGRAPH_2:11;
     case q`2>=0;
       then q in dom (cn-FanMorphS) & y=(cn-FanMorphS).q by A3,A4,Th120;
      hence ex x being set st x in dom (cn-FanMorphS) & y=(cn-FanMorphS).x;
     case A5: q`1/|.q.|>=0 & q`2<=0 & q<>0.REAL 2;
          set px=|[ |.q.|*(q`1/|.q.|*(1-cn)+cn),
            -((|.q.|)*sqrt(1-(q`1/|.q.|*(1-cn)+cn)^2))]|;
        A6: px`2 = -(|.q.|)*sqrt(1-(q`1/|.q.|*(1-cn)+cn)^2) &
         px`1 = |.q.|*(q`1/|.q.|*(1-cn)+cn) by EUCLID:56;
        then A7: |.px.|^2=(-((|.q.|)*sqrt(1-(q`1/|.q.|*(1-cn)+cn)^2)))^2
        +(|.q.|*(q`1/|.q.|*(1-cn)+cn))^2 by JGRAPH_3:10
        .=((- |.q.|)*sqrt(1-(q`1/|.q.|*(1-cn)+cn)^2))^2
        +(|.q.|*(q`1/|.q.|*(1-cn)+cn))^2 by XCMPLX_1:175
        .=(- |.q.|)^2*(sqrt(1-(q`1/|.q.|*(1-cn)+cn)^2))^2
        +(|.q.|*(q`1/|.q.|*(1-cn)+cn))^2 by SQUARE_1:68
        .=(|.q.|)^2*(sqrt(1-(q`1/|.q.|*(1-cn)+cn)^2))^2
        +(|.q.|*(q`1/|.q.|*(1-cn)+cn))^2 by SQUARE_1:61
        .=(|.q.|)^2*(sqrt(1-(q`1/|.q.|*(1-cn)+cn)^2))^2
        +(|.q.|)^2*((q`1/|.q.|*(1-cn)+cn))^2 by SQUARE_1:68;
       A8: |.q.|<>0 by A5,TOPRNS_1:25;
         1-cn>=0 by A1,Th3;
       then A9: q`1/|.q.|*(1-cn)>=0 by A5,REAL_2:121;
       then A10: (q`1/|.q.|*(1-cn)+cn)>=0+cn by REAL_1:55;
       A11: |.q.|>=0 by TOPRNS_1:26;
       A12: |.q.|^2>0 by A8,SQUARE_1:74;
         --(1+cn)>0 by A1,Th3;
       then -(1+cn)<0 by REAL_1:66;
       then -1-cn <0 by XCMPLX_1:161;
       then -1-cn<= q`1/|.q.|*(1-cn) by A9,AXIOMS:22;
       then -1-cn+cn<= q`1/|.q.|*(1-cn)+cn by REAL_1:55;
       then A13: -1<= (q`1/|.q.|*(1-cn)+cn) by XCMPLX_1:27;
       A14: 1-cn>0 by A1,Th3;
       A15: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
         0<=(q`2)^2 by SQUARE_1:72;
       then 0+(q`1)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
       then (q`1)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A12,A15,REAL_1:73;
       then (q`1)^2/(|.q.|)^2 <= 1 by A12,XCMPLX_1:60;
       then ((q`1)/|.q.|)^2 <= 1 by SQUARE_1:69;
       then q`1/|.q.|<=1 by Th4;
       then q`1/|.q.|*(1-cn) <=(1)*(1-cn) by A14,AXIOMS:25;
       then q`1/|.q.|*(1-cn)+cn-cn <=1-cn by XCMPLX_1:26;
       then (q`1/|.q.|*(1-cn)+cn) <=1 by REAL_1:54;
       then 1^2>=(q`1/|.q.|*(1-cn)+cn)^2 by A13,JGRAPH_2:7;
       then A16: 1-(q`1/|.q.|*(1-cn)+cn)^2>=0 by SQUARE_1:12,59;
       then A17: |.px.|^2=(|.q.|)^2*(1-(q`1/|.q.|*(1-cn)+cn)^2)
        +(|.q.|)^2*((q`1/|.q.|*(1-cn)+cn))^2 by A7,SQUARE_1:def 4
        .= (|.q.|)^2*((1-(q`1/|.q.|*(1-cn)+cn)^2)
        +((q`1/|.q.|*(1-cn)+cn))^2) by XCMPLX_1:8
        .= (|.q.|)^2*(1) by XCMPLX_1:27
        .= (|.q.|)^2;
          |.px.|>=0 by TOPRNS_1:26;
        then A18: |.px.|=sqrt(|.q.|^2) by A17,SQUARE_1:89
        .=|.q.| by A11,SQUARE_1:89;
        then A19: |.px.|<>0 by A5,TOPRNS_1:25;
         sqrt(1-(q`1/|.q.|*(1-cn)+cn)^2)>=0 by A16,SQUARE_1:def 4;
       then (q`1/|.q.|*(1-cn)+cn)>=cn &
       (|.q.|)*sqrt(1-(q`1/|.q.|*(1-cn)+cn)^2)>=0 by A10,A11,REAL_2:121;
        then px`1/|.px.| >=cn & px`2<=0 & px<>0.REAL 2
                                by A6,A8,A18,REAL_1:66,TOPRNS_1:24,XCMPLX_1:90;
    then A20: (cn-FanMorphS).px
       =|[ |.px.|* ((px`1/|.px.|-cn)/(1-cn)),
         |.px.|*( -sqrt(1-((px`1/|.px.|-cn)/(1-cn))^2))]| by A1,Th122;
        A21: |.px.|* ((px`1/|.px.|-cn)/(1-cn))
        =|.q.|* (( ((q`1/|.q.|*(1-cn)+cn))-cn)/(1-cn)) by A6,A8,A18,XCMPLX_1:90
        .=|.q.|* (( q`1/|.q.|*(1-cn) )/(1-cn)) by XCMPLX_1:26
        .=|.q.|* ( q`1/|.q.|) by A14,XCMPLX_1:90
        .= q`1 by A8,XCMPLX_1:88;
        then A22: |.px.|*( -sqrt(1-((px`1/|.px.|-cn)/(1-cn))^2))
           = |.px.|*( -sqrt(1-(q`1/|.px.|)^2)) by A19,XCMPLX_1:90
           .= |.px.|*( -sqrt(1-(q`1)^2/(|.px.|)^2)) by SQUARE_1:69
           .= |.px.|*( -sqrt( (|.px.|)^2/(|.px.|)^2-(q`1)^2/(|.px.|)^2))
                            by A12,A17,XCMPLX_1:60
           .= |.px.|*( -sqrt( ((|.px.|)^2-(q`1)^2)/(|.px.|)^2))
                            by XCMPLX_1:121
           .= |.px.|*( -sqrt( ((q`1)^2+(q`2)^2-(q`1)^2)/(|.px.|)^2))
                            by A17,JGRAPH_3:10
           .= |.px.|*( -sqrt( ((q`2)^2)/(|.px.|)^2)) by XCMPLX_1:26
           .= |.px.|*( -sqrt((q`2/|.q.|)^2)) by A18,SQUARE_1:69;
           --(q`2/|.q.|)<=0 by A5,A8,A11,REAL_2:126;
         then -(q`2/|.q.|)>=0 by REAL_1:66;
         then |.px.|*( sqrt((-(q`2/|.q.|))^2))=|.q.|*(-(q`2/|.q.|)) by A18,
SQUARE_1:89
         .= (-q`2)/|.q.|*|.q.| by XCMPLX_1:188
         .=-q`2 by A8,XCMPLX_1:88;
         then - |.px.|*( sqrt((-(q`2/|.q.|))^2))=q`2;
          then |.px.|*(- sqrt((-(q`2/|.q.|))^2))=q`2 by XCMPLX_1:175;
        then |.px.|*(- sqrt(((q`2/|.q.|))^2))=q`2 by SQUARE_1:61;
        then A23:q=(cn-FanMorphS).px by A20,A21,A22,EUCLID:57;
          dom (cn-FanMorphS)=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
      hence ex x being set st x in dom (cn-FanMorphS) & y=(cn-FanMorphS).x
                               by A23;
     case A24: q`1/|.q.|<0 & q`2<=0 & q<>0.REAL 2;
          set px=|[ |.q.|*(q`1/|.q.|*(1+cn)+cn),
            -(|.q.|)*(sqrt(1-(q`1/|.q.|*(1+cn)+cn)^2))]|;
        A25: px`2 = -(|.q.|)*sqrt(1-(q`1/|.q.|*(1+cn)+cn)^2) &
         px`1 = |.q.|*(q`1/|.q.|*(1+cn)+cn) by EUCLID:56;
        then A26: |.px.|^2=(-(|.q.|)*sqrt(1-(q`1/|.q.|*(1+cn)+cn)^2))^2
        +(|.q.|*(q`1/|.q.|*(1+cn)+cn))^2 by JGRAPH_3:10
        .=((|.q.|)*sqrt(1-(q`1/|.q.|*(1+cn)+cn)^2))^2
        +(|.q.|*(q`1/|.q.|*(1+cn)+cn))^2 by SQUARE_1:61
        .=(|.q.|)^2*(sqrt(1-(q`1/|.q.|*(1+cn)+cn)^2))^2
        +(|.q.|*(q`1/|.q.|*(1+cn)+cn))^2 by SQUARE_1:68
        .=(|.q.|)^2*(sqrt(1-(q`1/|.q.|*(1+cn)+cn)^2))^2
        +(|.q.|)^2*((q`1/|.q.|*(1+cn)+cn))^2 by SQUARE_1:68;
       A27: |.q.|<>0 by A24,TOPRNS_1:25;
         1+cn>=0 by A1,Th3;
       then A28: q`1/|.q.|*(1+cn)<=0 by A24,REAL_2:123;
       then A29: (q`1/|.q.|*(1+cn)+cn)<=0+cn by REAL_1:55;
       A30: |.q.|>=0 by TOPRNS_1:26;
       A31: |.q.|^2>0 by A27,SQUARE_1:74;
         (1-cn)>0 by A1,Th3;
       then 1-cn>= q`1/|.q.|*(1+cn) by A28,AXIOMS:22;
       then 1-cn+cn>= q`1/|.q.|*(1+cn)+cn by REAL_1:55;
       then A32: 1>= (q`1/|.q.|*(1+cn)+cn) by XCMPLX_1:27;
       A33: 1+cn>0 by A1,Th3;
       A34: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
         0<=(q`2)^2 by SQUARE_1:72;
       then 0+(q`1)^2<=(q`1)^2+(q`2)^2 by REAL_1:55;
       then (q`1)^2/(|.q.|)^2 <= (|.q.|)^2/(|.q.|)^2 by A31,A34,REAL_1:73;
       then (q`1)^2/(|.q.|)^2 <= 1 by A31,XCMPLX_1:60;
       then ((q`1)/|.q.|)^2 <= 1 by SQUARE_1:69;
       then q`1/|.q.|>=-1 by Th4;
       then q`1/|.q.|*(1+cn) >=(-1)*(1+cn) by A33,AXIOMS:25;
       then q`1/|.q.|*(1+cn) >=-(1+cn) by XCMPLX_1:180;
       then q`1/|.q.|*(1+cn) >=-1-cn by XCMPLX_1:161;
       then q`1/|.q.|*(1+cn)+cn-cn >=-1-cn by XCMPLX_1:26;
       then (q`1/|.q.|*(1+cn)+cn) >=-1 by REAL_1:54;
       then 1^2>=(q`1/|.q.|*(1+cn)+cn)^2 by A32,JGRAPH_2:7;
       then A35: 1-(q`1/|.q.|*(1+cn)+cn)^2>=0 by SQUARE_1:12,59;
       then A36: |.px.|^2=(|.q.|)^2*(1-(q`1/|.q.|*(1+cn)+cn)^2)
        +(|.q.|)^2*((q`1/|.q.|*(1+cn)+cn))^2 by A26,SQUARE_1:def 4
        .= (|.q.|)^2*((1-(q`1/|.q.|*(1+cn)+cn)^2)
        +((q`1/|.q.|*(1+cn)+cn))^2) by XCMPLX_1:8
        .= (|.q.|)^2*(1) by XCMPLX_1:27
        .= (|.q.|)^2;
          |.px.|>=0 by TOPRNS_1:26;
        then A37: |.px.|=sqrt(|.q.|^2) by A36,SQUARE_1:89
        .=|.q.| by A30,SQUARE_1:89;
        then A38: |.px.|<>0 by A24,TOPRNS_1:25;
          sqrt(1-(q`1/|.q.|*(1+cn)+cn)^2)>=0 by A35,SQUARE_1:def 4;
        then -sqrt(1-(q`1/|.q.|*(1+cn)+cn)^2)<=0 by REAL_1:66;
       then (q`1/|.q.|*(1+cn)+cn)<=cn &
       (|.q.|)*(-sqrt(1-(q`1/|.q.|*(1+cn)+cn)^2))<=0 by A29,A30,REAL_2:123;
        then px`1/|.px.| <=cn & px`2<=0 & px<>0.REAL 2
                           by A25,A37,A38,TOPRNS_1:24,XCMPLX_1:90,175;
    then A39: (cn-FanMorphS).px
       =|[ |.px.|* ((px`1/|.px.|-cn)/(1+cn)),
         |.px.|*( -sqrt(1-((px`1/|.px.|-cn)/(1+cn))^2))]| by A1,Th122;
        A40: |.px.|* ((px`1/|.px.|-cn)/(1+cn))
        =|.q.|* (( ((q`1/|.q.|*(1+cn)+cn))-cn)/(1+cn)) by A25,A27,A37,XCMPLX_1:
90
        .=|.q.|* (( q`1/|.q.|*(1+cn) )/(1+cn)) by XCMPLX_1:26
        .=|.q.|* ( q`1/|.q.|) by A33,XCMPLX_1:90
        .= q`1 by A27,XCMPLX_1:88;
        then A41: |.px.|*( -sqrt(1-((px`1/|.px.|-cn)/(1+cn))^2))
           = |.px.|*( -sqrt(1-(q`1/|.px.|)^2)) by A38,XCMPLX_1:90
           .= |.px.|*( -sqrt(1-(q`1)^2/(|.px.|)^2)) by SQUARE_1:69
           .= |.px.|*( -sqrt( (|.px.|)^2/(|.px.|)^2-(q`1)^2/(|.px.|)^2))
                            by A31,A36,XCMPLX_1:60
           .= |.px.|*( -sqrt( ((|.px.|)^2-(q`1)^2)/(|.px.|)^2))
                            by XCMPLX_1:121
           .= |.px.|*( -sqrt( ((q`1)^2+(q`2)^2-(q`1)^2)/(|.px.|)^2))
                            by A36,JGRAPH_3:10
           .= |.px.|*( -sqrt( ((q`2)^2)/(|.px.|)^2)) by XCMPLX_1:26
           .= |.px.|*( -sqrt((q`2/|.q.|)^2)) by A37,SQUARE_1:69;
           --q`2/|.q.|<=0 by A24,A27,A30,REAL_2:126;
         then A42: -q`2/|.q.|>=0 by REAL_1:66;
           |.px.|*( -sqrt((q`2/|.q.|)^2))
          = |.px.|*( -sqrt((-q`2/|.q.|)^2)) by SQUARE_1:61
         .=|.px.|*(--(q`2/|.q.|)) by A42,SQUARE_1:89
         .=q`2 by A27,A37,XCMPLX_1:88;
        then A43:q=(cn-FanMorphS).px by A39,A40,A41,EUCLID:57;
          dom (cn-FanMorphS)=the carrier of TOP-REAL 2 by FUNCT_2:def 1;
      hence ex x being set st x in dom (cn-FanMorphS) & y=(cn-FanMorphS).x
                               by A43;
     end;
    hence y in rng f by A3,FUNCT_1:def 5;
   end;
   hence rng (cn-FanMorphS)=the carrier of TOP-REAL 2
                 by A3,XBOOLE_0:def 10;
  end;
 thus thesis by A2;
end;

theorem Th142: for cn being Real,p2 being Point of TOP-REAL 2
  st -1<cn & cn<1
  ex K being non empty compact Subset of TOP-REAL 2 st
  K = (cn-FanMorphS).:K &
 (ex V2 being Subset of TOP-REAL 2 st
  p2 in V2 & V2 is open & V2 c= K & (cn-FanMorphS).p2 in V2)
 proof let cn be Real,p2 be Point of TOP-REAL 2;
  assume A1: -1<cn & cn<1;
    A2: the carrier of TOP-REAL 2=REAL 2 by EUCLID:25;
    A3: Euclid 2=MetrStruct(#REAL 2,Pitag_dist 2#) by EUCLID:def 7;
    then reconsider O=0.REAL 2 as Point of Euclid 2 by EUCLID:25;
    A4: TOP-REAL 2=TopSpaceMetr(Euclid 2) by EUCLID:def 8;
    reconsider V0=Ball(O,|.p2.|+1) as Subset of TOP-REAL 2
                              by A2,A3;
    A5: |. p2.|<|.p2.|+1 by REAL_1:69;
      0<=|.p2.| by TOPRNS_1:26;
    then A6: O in V0 by A5,GOBOARD6:4;
    A7: V0 c= cl_Ball(O,|.p2.|+1) by METRIC_1:15;
    reconsider K01=cl_Ball(O,|.p2.|+1) as Subset of TOP-REAL 2
                            by A2,A3;
      K01 is compact by Th22;
    then reconsider K0=cl_Ball(O,|.p2.|+1)
          as non empty compact Subset of TOP-REAL 2 by A6,A7;
    reconsider u2=p2 as Point of Euclid 2 by A3,EUCLID:25;
    A8: (cn-FanMorphS).:K0 c= K0
     proof let y be set;assume y in (cn-FanMorphS).:K0;
       then consider x being set such that
       A9: x in dom (cn-FanMorphS) & x in K0 & y=(cn-FanMorphS).x
                     by FUNCT_1:def 12;
       reconsider q=x as Point of TOP-REAL 2 by A9;
       reconsider uq=q as Point of Euclid 2 by A3,EUCLID:25;
       A10: y in rng (cn-FanMorphS) by A9,FUNCT_1:def 5;
       then reconsider u=y as Point of Euclid 2 by A3,EUCLID:25;
       reconsider q4=y as Point of TOP-REAL 2 by A10;
       A11: |.q4.|=|.q.| by A9,Th135;
         dist(O,uq)<= |.p2.|+1 by A9,METRIC_1:13;
       then |.0.REAL 2 - q.|<= |.p2.|+1 by JGRAPH_1:45;
       then |.0.REAL 2 +- q.|<= |.p2.|+1 by EUCLID:45;
       then |. -q.|<= |.p2.|+1 by EUCLID:31;
       then |.q.|<= |.p2.|+1 by TOPRNS_1:27;
       then |. -q4.|<= |.p2.|+1 by A11,TOPRNS_1:27;
       then |.0.REAL 2 +- q4.|<= |.p2.|+1 by EUCLID:31;
       then |.0.REAL 2 - q4.|<= |.p2.|+1 by EUCLID:45;
       then dist(O,u)<= |.p2.|+1 by JGRAPH_1:45;
      hence y in K0 by METRIC_1:13;
     end;
      K0 c= (cn-FanMorphS).:K0
     proof let y be set;assume A12: y in K0;
       then reconsider u=y as Point of Euclid 2;
       reconsider q4=y as Point of TOP-REAL 2 by A12;
         the carrier of TOP-REAL 2 c= rng (cn-FanMorphS) by A1,Th141;
       then q4 in rng (cn-FanMorphS) by TARSKI:def 3;
       then consider x being set such that
       A13: x in dom (cn-FanMorphS) & y=(cn-FanMorphS).x by FUNCT_1:def 5;
       reconsider q=x as Point of TOP-REAL 2 by A13,FUNCT_2:def 1;
       reconsider uq=q as Point of Euclid 2 by A3,EUCLID:25;
       A14: |.q4.|=|.q.| by A13,Th135;
         q = uq & q4 = u & O=0.REAL 2;
       then q in K0 by A12,A14,Lm11;
      hence y in (cn-FanMorphS).:K0 by A13,FUNCT_1:def 12;
     end;
    then A15: K0= (cn-FanMorphS).:K0 by A8,XBOOLE_0:def 10;
      |. -p2.|<|.p2.|+1 by A5,TOPRNS_1:27;
    then |.0.REAL 2 +- p2.|<|.p2.|+1 by EUCLID:31;
    then |.0.REAL 2 - p2.|<|.p2.|+1 by EUCLID:45;
    then dist(O,u2)<|.p2.|+1 by JGRAPH_1:45;
    then A16: p2 in V0 by METRIC_1:12;
    A17: V0 is open by A4,TOPMETR:21;
     set q3= (cn-FanMorphS).p2;
     reconsider u3=q3 as Point of Euclid 2 by A3,EUCLID:25;
       |.q3.|=|.p2.| by Th135;
    then |. -q3.|<|.p2.|+1 by A5,TOPRNS_1:27;
    then |.0.REAL 2 +- q3.|<|.p2.|+1 by EUCLID:31;
    then |.0.REAL 2 - q3.|<|.p2.|+1 by EUCLID:45;
     then dist(O,u3)< |.p2.|+1 by JGRAPH_1:45;
    then (cn-FanMorphS).p2 in V0 by METRIC_1:12;
   hence thesis by A7,A15,A16,A17;
 end;

theorem for cn being Real st -1<cn & cn<1
 ex f being map of TOP-REAL 2,TOP-REAL 2 st f=cn-FanMorphS
   & f is_homeomorphism
proof let cn be Real;assume A1:  -1<cn & cn<1;
  then A2: (cn-FanMorphS) is map of TOP-REAL 2,TOP-REAL 2 &
  rng (cn-FanMorphS) = the carrier of TOP-REAL 2 by Th141;
  reconsider f=(cn-FanMorphS) as map of TOP-REAL 2,TOP-REAL 2;
  consider h being map of (TOP-REAL 2),(TOP-REAL 2)
  such that A3: h=(cn-FanMorphS) & h is continuous by A1,Th139;
  A4: f is one-to-one by A1,Th140;
  A5: rng f=[#](TOP-REAL 2) by A2,PRE_TOPC:12;
   for p2 being Point of TOP-REAL 2
  ex K being non empty compact Subset of TOP-REAL 2 st K = f.:K &
 (ex V2 being Subset of TOP-REAL 2 st
  p2 in V2 & V2 is open & V2 c= K & f.p2 in V2) by A1,Th142;
  then f is_homeomorphism by A3,A4,A5,Th8;
 hence thesis;
end;

theorem Th144:
 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;
  A2: |.q.|>=0 by TOPRNS_1:26;
  let p be Point of TOP-REAL 2;
    assume p=(cn-FanMorphS).q;
    then A3: p=|[ |.q.|* ((q`1/|.q.|-cn)/(1-cn)),
         |.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))]| by A1,Th120;
      set qz=p;
      A4: |.q.|<>0 by A1,JGRAPH_2:11,TOPRNS_1:25;
      then A5: (|.q.|)^2>0 by SQUARE_1:74;
     A6: 1-cn>0 by A1,Th3;
        A7: qz`2= |.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))&
        qz`1= |.q.|* ((q`1/|.q.|-cn)/(1-cn)) by A3,EUCLID:56;
        A8: (q`1/|.q.|-cn)>= 0 by A1,SQUARE_1:12;
        A9: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
          0<(q`2)^2 by A1,SQUARE_1:74;
        then 0+(q`1)^2<(q`1)^2+(q`2)^2 by REAL_1:67;
        then (q`1)^2/(|.q.|)^2 < (|.q.|)^2/(|.q.|)^2 by A5,A9,REAL_1:73;
        then (q`1)^2/(|.q.|)^2 < 1 by A5,XCMPLX_1:60;
        then ((q`1)/|.q.|)^2 < 1 by SQUARE_1:69;
        then 1>q`1/|.q.| by Th5;
        then A10: 1-cn>q`1/|.q.|-cn by REAL_1:54;
        A11: -(1-cn)< -0 by A6,REAL_1:50;
           -(1-cn)< -( q`1/|.q.|-cn) by A10,REAL_1:50;
         then (-(1-cn))/(1-cn)<(-( q`1/|.q.|-cn))/(1-cn) by A6,REAL_1:73;
        then A12: -1<(-( q`1/|.q.|-cn))/(1-cn) by A6,XCMPLX_1:198;
         --(1-cn)> -(q`1/|.q.|-cn) by A8,A11,REAL_1:50;
       then (-(q`1/|.q.|-cn))/(1-cn)<1 by A6,REAL_2:118;
       then ((-(q`1/|.q.|-cn))/(1-cn))^2<1^2 by A12,JGRAPH_2:8;
       then A13: -sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2)< 0 by Lm12;
         ((q`1/|.q.|-cn)/(1-cn))>=0 by A6,A8,SQUARE_1:27;
   hence thesis by A2,A4,A7,A13,SQUARE_1:19,24;
 end;

theorem Th145:
 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;
  A2: |.q.|>=0 by TOPRNS_1:26;
  let p be Point of TOP-REAL 2;
    assume p=(cn-FanMorphS).q;
    then A3: p=|[ |.q.|* ((q`1/|.q.|-cn)/(1+cn)),
      |.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))]| by A1,Th121;
      set qz=p;
          A4: |.q.|<>0 by A1,JGRAPH_2:11,TOPRNS_1:25;
      then A5: (|.q.|)^2>0 by SQUARE_1:74;
      A6: (q`1/|.q.|-cn)< 0 by A1,REAL_2:105;
     A7: 1+cn>0 by A1,Th3;
        A8: qz`2= |.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2))&
        qz`1= |.q.|* ((q`1/|.q.|-cn)/(1+cn)) by A3,EUCLID:56;
        A9: (|.q.|)^2 =(q`1)^2+(q`2)^2 by JGRAPH_3:10;
          0<(q`2)^2 by A1,SQUARE_1:74;
        then 0+(q`1)^2<(q`1)^2+(q`2)^2 by REAL_1:67;
        then (q`1)^2/(|.q.|)^2 < (|.q.|)^2/(|.q.|)^2 by A5,A9,REAL_1:73;
        then (q`1)^2/(|.q.|)^2 < 1 by A5,XCMPLX_1:60;
        then ((q`1)/|.q.|)^2 < 1 by SQUARE_1:69;
        then -1<q`1/|.q.| by Th5;
        then -1-cn<q`1/|.q.|-cn by REAL_1:54;
        then -(1+cn)<q`1/|.q.|-cn by XCMPLX_1:161;
        then A10: --(1+cn)> -(q`1/|.q.|-cn) by REAL_1:50;
        A11: -( q`1/|.q.|-cn)>0 by A6,REAL_1:66;
          -(1+cn)< -0 by A7,REAL_1:50;
         then -(1+cn)< -( q`1/|.q.|-cn) by A11,AXIOMS:22;
         then (-(1+cn))/(1+cn)<(-( q`1/|.q.|-cn))/(1+cn) by A7,REAL_1:73;
        then A12: -1<(-( q`1/|.q.|-cn))/(1+cn) by A7,XCMPLX_1:198;
         (-(q`1/|.q.|-cn))/(1+cn)<1 by A7,A10,REAL_2:118;
       then ((-(q`1/|.q.|-cn))/(1+cn))^2<1^2 by A12,JGRAPH_2:8;
       then 1-((-(q`1/|.q.|-cn))/(1+cn))^2>0 by SQUARE_1:11,59;
       then sqrt(1-((-(q`1/|.q.|-cn))/(1+cn))^2)>0 by SQUARE_1:93;
       then sqrt(1-(-(q`1/|.q.|-cn))^2/(1+cn)^2)> 0 by SQUARE_1:69;
       then sqrt(1-(q`1/|.q.|-cn)^2/(1+cn)^2)> 0 by SQUARE_1:61;
       then --sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2)> 0 by SQUARE_1:69;
       then A13: -sqrt(1-((q`1/|.q.|-cn)/(1+cn))^2)< 0 by REAL_1:66;
         ((q`1/|.q.|-cn)/(1+cn))<0 by A6,A7,REAL_2:128;
   hence thesis by A2,A4,A8,A13,SQUARE_1:24;
 end;

theorem Th146:
 for cn being Real,q1,q2 being Point of TOP-REAL 2 st -1<cn & cn<1 & q1`2<0
 & q1`1/|.q1.|>=cn & q2`2<0 & q2`1/|.q2.|>=cn & 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
   & q1`1/|.q1.|>=cn & q2`2<0 & q2`1/|.q2.|>=cn & q1`1/|.q1.|<q2`1/|.q2.|;
  let p1,p2 be Point of TOP-REAL 2;
   assume A2: p1=(cn-FanMorphS).q1 & p2=(cn-FanMorphS).q2;
    then p1=|[ |.q1.|* ((q1`1/|.q1.|-cn)/(1-cn)),
       |.q1.|*( -sqrt(1-((q1`1/|.q1.|-cn)/(1-cn))^2))]| by A1,Th120;
    then A3: p1`2= |.q1.|*( -sqrt(1-((q1`1/|.q1.|-cn)/(1-cn))^2))&
        p1`1= |.q1.|* ((q1`1/|.q1.|-cn)/(1-cn)) by EUCLID:56;
      p2=|[ |.q2.|* ((q2`1/|.q2.|-cn)/(1-cn)),
      |.q2.|*( -sqrt(1-((q2`1/|.q2.|-cn)/(1-cn))^2))]| by A1,A2,Th120;
    then A4: p2`2= |.q2.|*( -sqrt(1-((q2`1/|.q2.|-cn)/(1-cn))^2))&
        p2`1= |.q2.|* ((q2`1/|.q2.|-cn)/(1-cn)) by EUCLID:56;
    A5: |.q1.|>0 & |.q2.|>0 by A1,Lm1,JGRAPH_2:11;
    A6: |.p1.|=|.q1.| & |.p2.|=|.q2.| by A2,Th135;
    then A7: p1`1/|.p1.|= ((q1`1/|.q1.|-cn)/(1-cn)) by A3,A5,XCMPLX_1:90
;
    A8: p2`1/|.p2.|= ((q2`1/|.q2.|-cn)/(1-cn)) by A4,A5,A6,XCMPLX_1:90
;
    A9: q1`1/|.q1.|-cn< q2`1/|.q2.|-cn by A1,REAL_1:54;
      1-cn>0 by A1,Th3;
   hence thesis by A7,A8,A9,REAL_1:73;
 end;

theorem Th147:
 for cn being Real,q1,q2 being Point of TOP-REAL 2 st -1<cn & cn<1 & q1`2<0
 & q1`1/|.q1.|<cn & q2`2<0 & q2`1/|.q2.|<cn & 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
   & q1`1/|.q1.|<cn & q2`2<0 & q2`1/|.q2.|<cn & q1`1/|.q1.|<q2`1/|.q2.|;
  let p1,p2 be Point of TOP-REAL 2;
   assume A2: p1=(cn-FanMorphS).q1 & p2=(cn-FanMorphS).q2;
    then p1=|[ |.q1.|* ((q1`1/|.q1.|-cn)/(1+cn)),
      |.q1.|*( -sqrt(1-((q1`1/|.q1.|-cn)/(1+cn))^2))]| by A1,Th121;
    then A3: p1`2= |.q1.|*( -sqrt(1-((q1`1/|.q1.|-cn)/(1+cn))^2))&
        p1`1= |.q1.|* ((q1`1/|.q1.|-cn)/(1+cn)) by EUCLID:56;
      p2=|[ |.q2.|* ((q2`1/|.q2.|-cn)/(1+cn)),
      |.q2.|*( -sqrt(1-((q2`1/|.q2.|-cn)/(1+cn))^2))]| by A1,A2,Th121;
    then A4: p2`2= |.q2.|*( -sqrt(1-((q2`1/|.q2.|-cn)/(1+cn))^2))&
        p2`1= |.q2.|* ((q2`1/|.q2.|-cn)/(1+cn)) by EUCLID:56;
    A5: |.q1.|>0 & |.q2.|>0 by A1,Lm1,JGRAPH_2:11;
    A6: |.p1.|=|.q1.| & |.p2.|=|.q2.| by A2,Th135;
    then A7: p1`1/|.p1.|= (q1`1/|.q1.|-cn)/(1+cn) by A3,A5,XCMPLX_1:90
;
    A8: p2`1/|.p2.|= (q2`1/|.q2.|-cn)/(1+cn) by A4,A5,A6,XCMPLX_1:90;
    A9: q1`1/|.q1.|-cn< q2`1/|.q2.|-cn by A1,REAL_1:54;
      1+cn>0 by A1,Th3;
   hence p1`1/|.p1.|<p2`1/|.p2.| by A7,A8,A9,REAL_1:73;
 end;

theorem
   for cn being Real,q1,q2 being Point of TOP-REAL 2 st -1<cn & cn<1 & q1`2<0
 & q2`2<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`1/|.q1.|<q2`1/|.q2.|;
  let p1,p2 be Point of TOP-REAL 2;
   assume A2: p1=(cn-FanMorphS).q1 & p2=(cn-FanMorphS).q2;
      now per cases;
    case q1`1/|.q1.|>=cn & q2`1/|.q2.|>=cn;
     hence p1`1/|.p1.|<p2`1/|.p2.| by A1,A2,Th146;
    case q1`1/|.q1.|>=cn & q2`1/|.q2.|<cn;
     hence p1`1/|.p1.|<p2`1/|.p2.| by A1,AXIOMS:22;
    case A3: q1`1/|.q1.|<cn & q2`1/|.q2.|>=cn;
       then A4: p1`2<0 & p1`1<0 by A1,A2,Th145;
       then A5: |.p1.|>0 by Lm1,JGRAPH_2:11;
       A6: p2`2<0 & p2`1>=0 by A1,A2,A3,Th144;
       then |.p2.|>0 by Lm1,JGRAPH_2:11;
       then p2`1/|.p2.|>=0 by A6,REAL_2:125;
     hence p1`1/|.p1.|<p2`1/|.p2.| by A4,A5,REAL_2:128;
    case q1`1/|.q1.|<cn & q2`1/|.q2.|<cn;
     hence p1`1/|.p1.|<p2`1/|.p2.| by A1,A2,Th147;
    end;
   hence p1`1/|.p1.|<p2`1/|.p2.|;
 end;

theorem
   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;
  A2: |.q.|>=0 by TOPRNS_1:26;
  let p be Point of TOP-REAL 2;
   assume p=(cn-FanMorphS).q;
    then A3: p=|[ |.q.|* ((q`1/|.q.|-cn)/(1-cn)),
      |.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))]| by A1,Th120;
     A4: |.q.|<>0 by A1,JGRAPH_2:11,TOPRNS_1:25;
     A5: (q`1/|.q.|-cn)=0 by A1,XCMPLX_1:14;
     A6: p`2= |.q.|*( -sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2))&
        p`1= |.q.|* ((q`1/|.q.|-cn)/(1-cn)) by A3,EUCLID:56;
         ((-(q`1/|.q.|-cn))/(1-cn))^2<1^2 by A5,JGRAPH_2:8,XCMPLX_1:198;
       then -sqrt(1-((q`1/|.q.|-cn)/(1-cn))^2)< 0 by Lm12;
   hence thesis by A2,A4,A5,A6,SQUARE_1:24;
 end;

theorem
   for cn being real number holds 0.REAL 2=(cn-FanMorphS).(0.REAL 2)
   by Th120,JGRAPH_2:11;

Back to top