The Mizar article:

Bounded Domains and Unbounded Domains

by
Yatsuka Nakamura,
Andrzej Trybulec, and
Czeslaw Bylinski

Received January 7, 1999

Copyright (c) 1999 Association of Mizar Users

MML identifier: JORDAN2C
[ MML identifier index ]


environ

 vocabulary ABSVALUE, ARYTM_1, FINSEQ_1, RELAT_1, FUNCT_1, ORDINAL2, PRE_TOPC,
      EUCLID, COMPLEX1, SQUARE_1, RLVECT_1, RVSUM_1, TOPREAL1, RELAT_2,
      BORSUK_1, TOPS_2, SUBSET_1, RCOMP_1, BOOLE, LATTICES, CONNSP_1, TARSKI,
      CONNSP_3, SETFAM_1, GRAPH_1, ARYTM_3, TREAL_1, SEQ_1, FUNCT_4, TOPMETR,
      COMPTS_1, FINSEQ_2, METRIC_1, PCOMPS_1, WEIERSTR, SEQ_4, SEQ_2, VECTSP_1,
      FUNCOP_1, PARTFUN1, JORDAN3, TBSP_1, CONNSP_2, SPPOL_1, GOBOARD2,
      SPRECT_1, TREES_1, PSCOMP_1, GOBOARD1, MCART_1, CARD_1, MATRIX_1,
      GOBOARD9, SEQM_3, GOBOARD5, TOPS_1, JORDAN1, SPRECT_2, JORDAN2C,
      FINSEQ_4, CARD_3, ARYTM;
 notation TARSKI, XBOOLE_0, ZFMISC_1, SUBSET_1, SETFAM_1, RELAT_1, FUNCT_1,
      FUNCT_2, ORDINAL1, NUMBERS, XCMPLX_0, XREAL_0, REAL_1, NAT_1, CARD_1,
      PRE_TOPC, TOPS_1, TOPS_2, COMPTS_1, METRIC_1, STRUCT_0, PCOMPS_1,
      CONNSP_1, CONNSP_2, TBSP_1, CONNSP_3, TOPRNS_1, TOPMETR, RCOMP_1,
      FINSEQ_1, FINSEQ_2, FINSEQ_4, SQUARE_1, ABSVALUE, BORSUK_2, WEIERSTR,
      SEQ_4, JORDAN3, BINARITH, FUNCOP_1, FUNCT_3, VECTSP_1, TREAL_1, FUNCT_4,
      RVSUM_1, EUCLID, SPPOL_1, PSCOMP_1, SPRECT_1, SPRECT_2, TOPREAL1,
      MATRIX_1, GOBOARD1, JORDAN1, GOBOARD2, GOBOARD5, GOBOARD9, JORDAN2B;
 constructors WEIERSTR, REAL_1, TOPS_1, TOPS_2, COMPTS_1, REALSET1, CONNSP_1,
      TBSP_1, CONNSP_3, FINSEQ_4, JORDAN1, RCOMP_1, GOBOARD2, GOBOARD9,
      FINSEQOP, SQUARE_1, ABSVALUE, BORSUK_2, JORDAN3, TREAL_1, FUNCT_4,
      SPPOL_1, SPRECT_1, TOPREAL2, SPRECT_2, PSCOMP_1, BINARITH, JORDAN2B,
      PARTFUN1, TOPRNS_1, DOMAIN_1, MEMBERED, XCMPLX_0;
 clusters SUBSET_1, RELSET_1, STRUCT_0, PRE_TOPC, METRIC_1, PCOMPS_1, SPRECT_1,
      TOPMETR, TOPREAL1, BORSUK_1, FUNCT_1, FUNCOP_1, FINSEQ_1, XREAL_0,
      FINSET_1, SPRECT_3, EUCLID, GOBOARD1, GOBOARD2, ARYTM_3, MEMBERED,
      ZFMISC_1;
 requirements REAL, NUMERALS, BOOLE, SUBSET, ARITHM;
 definitions TARSKI, XBOOLE_0, TBSP_1, JORDAN1;
 theorems PRE_TOPC, CONNSP_1, EUCLID, TBSP_1, TOPREAL3, REAL_1, REAL_2, AXIOMS,
      NAT_1, JGRAPH_1, JORDAN1, TOPS_2, FUNCT_2, BORSUK_1, TOPMETR, TOPREAL1,
      FINSEQ_2, FINSEQOP, ABSVALUE, RVSUM_1, SQUARE_1, BORSUK_2, TOPREAL5,
      WEIERSTR, SPPOL_1, TARSKI, SEQ_4, FUNCT_1, METRIC_1, SUBSET_1, FUNCOP_1,
      ZFMISC_1, FINSEQ_1, FINSEQ_4, JORDAN3, BINARITH, RELAT_1, FUNCT_3,
      VECTSP_1, BINOP_1, RCOMP_1, FUNCT_4, HEINE, TOPMETR2, TREAL_1, PCOMPS_1,
      CONNSP_3, JORDAN6, COMPTS_1, TSEP_1, CONNSP_2, TOPS_1, JORDAN5D,
      JORDAN2B, GOBOARD1, GOBOARD2, GOBOARD5, GOBOARD6, UNIFORM1, FINSEQ_6,
      MATRIX_1, GOBOARD9, GOBRD12, SPRECT_1, CARD_1, CARD_2, ENUMSET1,
      PSCOMP_1, SPRECT_2, SPRECT_3, SEQ_2, RELSET_1, SCMFSA_7, SETFAM_1,
      XBOOLE_0, XBOOLE_1, XREAL_0, TOPRNS_1, XCMPLX_0, XCMPLX_1;
 schemes FUNCT_1, NAT_1;

begin ::Definitions of Bounded Domain and Unbounded Domain

reserve m,n,i,i2,j for Nat,
        r,r1,r2,s for Real,
        x,y,z,y1,y2 for set;

theorem Th1: r<=0 implies abs(r)=-r
proof assume
A1:r<=0;
  per cases by A1;
  suppose r<0;
  hence abs(r)=-r by ABSVALUE:def 1;
  suppose r=0;
  hence abs(r)=-r by ABSVALUE:7;
end;

theorem Th2:for n,m st n<=m & m<=n+2 holds m=n or m=n+1 or m=n+2
 proof let n,m;assume A1: n<=m & m<=n+2;
   per cases;
   suppose m<=n+1;
    hence m=n or m=n+1 or m=n+2 by A1,NAT_1:27;
   suppose A2:m>n+1;
       m<=n+(1+1) by A1;
     then m<=n+1+1 by XCMPLX_1:1;
     then m=n+1 or m=n+1+1 by A2,NAT_1:27;
     then m=n+1 or m=n+(1+1) by XCMPLX_1:1;
    hence thesis;
 end;

theorem Th3:for n,m st n<=m & m<=n+3 holds m=n or m=n+1 or m=n+2 or m=n+3
 proof let n,m;assume A1:n<=m & m<=n+3;
   per cases;
   suppose m<=n+2;
    hence m=n or m=n+1 or m=n+2 or m=n+3 by A1,Th2;
   suppose A2:m>n+2;
       m<=n+(2+1) by A1;
     then m<=n+2+1 by XCMPLX_1:1;
     then m=n+2 or m=n+2+1 by A2,NAT_1:27;
     then m=n+2 or m=n+(2+1) by XCMPLX_1:1;
    hence thesis;
 end;

theorem Th4:for n,m st n<=m & m<=n+4 holds
m=n or m=n+1 or m=n+2 or m=n+3 or m=n+4
proof let n,m;assume A1:n<=m & m<=n+4;
   per cases;
   suppose m<=n+3;
    hence m=n or m=n+1 or m=n+2 or m=n+3 or m=n+4 by A1,Th3;
   suppose A2:m>n+3;
       m<=n+(3+1) by A1;
     then m<=n+3+1 by XCMPLX_1:1;
     then m=n+3 or m=n+3+1 by A2,NAT_1:27;
     then m=n+3 or m=n+(3+1) by XCMPLX_1:1;
    hence thesis;
end;

theorem Th5: for a,b being real number st a>=0 & b>=0 holds a+b>=0
proof let a,b be real number;assume a>=0 & b>=0;
  then a+b>=0+0 by REAL_1:55;
 hence a+b>=0;
end;

theorem Th6: for a,b being real number st a>0 & b>=0 holds a+b>0
proof let a,b be real number;assume a>0 & b>=0;
  then a+b>0+0 by REAL_1:67;
 hence a+b>0;
end;

theorem Th7: for f being FinSequence st rng f={x,y} & len f=2
holds f.1=x & f.2=y or f.1=y & f.2=x
proof let f be FinSequence;assume A1:rng f={x,y} & len f=2;
  then 1 in Seg len f by FINSEQ_1:3;
  then 1 in dom f by FINSEQ_1:def 3;
then A2:f.1 in rng f by FUNCT_1:def 5;
    2 in Seg len f by A1,FINSEQ_1:3;
  then 2 in dom f by FINSEQ_1:def 3;
then A3: f.2 in rng f by FUNCT_1:def 5;
  A4:now assume A5:f.1=x & f.2=x;
      y in rng f by A1,TARSKI:def 2;
    then consider z such that A6:z in dom f & y=f.z by FUNCT_1:def 5;
    A7:z in Seg len f by A6,FINSEQ_1:def 3;
    reconsider nz=z as Nat by A6;
A8: 1<=nz & nz<=len f by A7,FINSEQ_1:3;
    per cases by A1,A8,NAT_1:27;
    suppose nz=1; hence f.1=x & f.2=y by A5,A6;
    suppose nz=1+1; hence f.1=x & f.2=y by A5,A6;
  end;
    now assume A9:f.1=y & f.2=y;
      x in rng f by A1,TARSKI:def 2;
    then consider z such that A10:z in dom f & x=f.z by FUNCT_1:def 5;
    A11:z in Seg len f by A10,FINSEQ_1:def 3;
    reconsider nz=z as Nat by A10;
A12:1<=nz & nz<=len f by A11,FINSEQ_1:3;
    per cases by A1,A12,NAT_1:27;
    suppose nz=1; hence f.1=x & f.2=y by A9,A10;
    suppose nz=1+1; hence f.1=x & f.2=y by A9,A10;
  end;
 hence thesis by A1,A2,A3,A4,TARSKI:def 2;
end;

theorem Th8: for f being increasing FinSequence of REAL
st rng f={r,s} & len f=2 & r<=s
holds f.1=r & f.2=s
proof let f be increasing FinSequence of REAL;
  assume A1: rng f={r,s} & len f=2 & r<=s;
     now assume A2:f.1=s & f.2=r;
      1 in Seg len f by A1,FINSEQ_1:3;
    then A3:1 in dom f by FINSEQ_1:def 3;
      2 in Seg len f by A1,FINSEQ_1:3;
    then 2 in dom f by FINSEQ_1:def 3;
    hence f.1=r & f.2=s by A1,A2,A3,GOBOARD1:def 1;
   end;
 hence f.1=r & f.2=s by A1,Th7;
end;

reserve p,p1,p2,p3,q,q1,q2,q3,q4 for Point of TOP-REAL n;

theorem Th9: p1+p2-p3=p1-p3+p2
proof
 thus p1+p2-p3=p1+p2+-p3 by EUCLID:45
 .=p1+-p3+p2 by EUCLID:30 .=p1-p3+p2 by EUCLID:45;
end;

theorem abs(|.q.|)=|.q.|
 proof
     |.q.|>=0 by TOPRNS_1:26;
  hence thesis by ABSVALUE:def 1;
 end;

theorem Th11: abs(|.q1.|- |.q2.|)<=|.q1-q2.|
 proof
   per cases;
   suppose |.q1.|>=|.q2.|;
     then |.q1.|- |.q2.|>=0 by SQUARE_1:12;
     then |.q1.|- |.q2.|=abs(|.q1.|- |.q2.|) by ABSVALUE:def 1;
    hence thesis by TOPRNS_1:33;
   suppose |.q1.|<|.q2.|;
     then A1: |.q2.|- |.q1.|>0 by SQUARE_1:11;
      |.q2.|- |.q1.|<= |.q2-q1.| by TOPRNS_1:33;
    then abs(|.q2.|- |.q1.|)<= |.q2-q1.| by A1,ABSVALUE:def 1;
    then abs(|.q2.|- |.q1.|)<= |.q1-q2.| by TOPRNS_1:28;
    hence thesis by UNIFORM1:13;
 end;

theorem Th12: |.|[r]|.|=abs(r)
proof
    set p=|[r]|;
    reconsider w=|[r]| as Element of REAL 1 by EUCLID:25;
    A1: |.p.|=|.w.| by JGRAPH_1:def 5;
    A2: |.w.| = sqrt Sum sqr w by EUCLID:def 5;
      0 <= Sum sqr w by RVSUM_1:116;
    then A3: (|.p.|)^2=Sum sqr w by A1,A2,SQUARE_1:def 4;
    A4: |.p.|>=0 by TOPRNS_1:26;
      w=<*r*> by JORDAN2B:def 2;
    then sqr w=<*r^2*> by RVSUM_1:81;
    then (|.p.|)^2 = r^2 by A3,RVSUM_1:103;
    then (|.p.|) =sqrt (r^2) by A4,SQUARE_1:89
           .=abs r by SQUARE_1:91;
 hence thesis;
end;

theorem Th13: q-0.REAL n=q & (0.REAL n)-q = -q
proof
   thus q-0.REAL n=q-(q-q) by EUCLID:46 .=q-q+q by EUCLID:51 .=q by EUCLID:52;
  thus (0.REAL n)-q =-q--q -q by EUCLID:46.= -q+--q-q by EUCLID:45
  .=-q+q-q by EUCLID:39
  .=-q by EUCLID:52;
end;

theorem Th14:
for P being Subset of TOP-REAL n st P is convex
 holds P is connected
proof let P be Subset of TOP-REAL n; assume that
A1:for w3,w4 being Point of TOP-REAL n
         st w3 in P & w4 in P holds LSeg(w3,w4) c= P;
      for w1,w2 being Point of TOP-REAL n st w1 in P & w2 in P & w1<>w2
       ex h being map of I[01],(TOP-REAL n)|P st h is continuous
                        & w1=h.0 & w2=h.1
   proof
    let w1,w2 be Point of TOP-REAL n;
    assume A2:w1 in P & w2 in P & w1<>w2;
    then A3: LSeg(w1,w2) c= P by A1;
      LSeg(w1,w2) is_an_arc_of w1,w2 by A2,TOPREAL1:15;
    then consider f being map of I[01], (TOP-REAL n)|LSeg(w1,w2) such that
A4: f is_homeomorphism & f.0 = w1 & f.1 = w2 by TOPREAL1:def 2;
    A5:f is continuous by A4,TOPS_2:def 5;
    A6: rng f = [#]((TOP-REAL n)|LSeg(w1,w2)) by A4,TOPS_2:def 5;
    then A7: rng f c= P by A3,PRE_TOPC:def 10;
    then [#]((TOP-REAL n)|LSeg(w1,w2)) c= [#]((TOP-REAL n)|P)
              by A6,PRE_TOPC:def 10;
    then [#]((TOP-REAL n)|LSeg(w1,w2))
          c=the carrier of ((TOP-REAL n)|P) by PRE_TOPC:12;
    then the carrier of ((TOP-REAL n)|LSeg(w1,w2))
       c= the carrier of ((TOP-REAL n)|P) by PRE_TOPC:12;
    then A8:(TOP-REAL n)|LSeg(w1,w2)
           is SubSpace of (TOP-REAL n)|P by TOPMETR:4;
      dom f=[.0 qua Real,1 qua Real.] by BORSUK_1:83,FUNCT_2:def 1;
    then reconsider g=f as Function of [.0 qua Real,1 qua Real.],P
                                  by A7,FUNCT_2:4;
       the carrier of (TOP-REAL n)|P = P by JORDAN1:1;
    then reconsider gt=g as map of I[01],(TOP-REAL n)|P
      by BORSUK_1:83;
      gt is continuous by A5,A8,TOPMETR:7;
    hence thesis by A4;
   end;
   hence thesis by JORDAN1:5;
end;

theorem Th15:for G being non empty TopSpace,
P being Subset of G,A being Subset of G,
Q being Subset of G|A st P=Q & P is connected
holds Q is connected
proof let G be non empty TopSpace,P be Subset of G,
 A be Subset of G,
  Q be Subset of G|A;
 assume that
A1: P=Q and
A2: P is connected;
  A3:G|P is connected by A2,CONNSP_1:def 3;
    Q c= the carrier of G|A;
  then Q c= A by JORDAN1:1;
  then G|P=(G|A)|Q by A1,JORDAN6:47;
 hence Q is connected by A3,CONNSP_1:def 3;
end;

definition let n;let A be Subset of TOP-REAL n;
  canceled;
attr A is Bounded means
:Def2:ex C being Subset of Euclid n st C=A & C is bounded;
correctness;
end;

theorem Th16:for A,B being Subset of TOP-REAL n st B is Bounded &
A c= B holds A is Bounded
proof let A,B be Subset of TOP-REAL n;
 assume A1:B is Bounded & A c= B;
  then consider C being Subset of Euclid n
   such that A2:C=B & C is bounded by Def2;
    A is Subset of Euclid n by A1,A2,XBOOLE_1:1;
  then reconsider C2=A as Subset of Euclid n;
    C2 is bounded by A1,A2,TBSP_1:21;
 hence A is Bounded by Def2;
end;

definition let n;let A be Subset of TOP-REAL n;
  let B be Subset of TOP-REAL n;
pred B is_inside_component_of A means
:Def3: B is_a_component_of A` & B is Bounded;
end;

definition let M be non empty MetrStruct;
cluster bounded Subset of M;
 existence proof take {}M, 1; thus thesis; end;
end;

theorem Th17: for A being Subset of TOP-REAL n,
 B being Subset of TOP-REAL n holds
   B is_inside_component_of A iff
  ex C being Subset of ((TOP-REAL n)|(A`)) st
 C=B & C is_a_component_of ((TOP-REAL n)|(A`)) &
 C is bounded Subset of Euclid n
proof let A be Subset of TOP-REAL n,
  B be Subset of TOP-REAL n;
  A1: B is_a_component_of A` iff
  ex C being Subset of (TOP-REAL n)|(A`) st
   C=B & C is_a_component_of (TOP-REAL n)|(A`) by CONNSP_1:def 6;
A2:B is_inside_component_of A iff
    B is_a_component_of A` & B is Bounded by Def3;
   thus B is_inside_component_of A implies
   (ex C being Subset of ((TOP-REAL n)|(A`)) st
   C=B & C is_a_component_of ((TOP-REAL n)|(A`)) &
    C is bounded Subset of Euclid n)
   proof assume A3: B is_inside_component_of A;
     then consider C being Subset of ((TOP-REAL n)|(A`)) such that
     A4:C=B & C is_a_component_of ((TOP-REAL n)|(A`)) by A1,Def3;
     consider D1 being Subset of Euclid n such that
      A5: D1=B & D1 is bounded by A2,A3,Def2;
    thus thesis by A4,A5;
   end;
     given C being Subset of ((TOP-REAL n)|(A`)) such that
     A6:C=B & C is_a_component_of ((TOP-REAL n)|(A`)) &
        C is bounded Subset of Euclid n;
     A7:B is Bounded by A6,Def2;
       B is_a_component_of A` by A6,CONNSP_1:def 6;
    hence thesis by A7,Def3;
end;

definition let n;let A be Subset of TOP-REAL n;
   let B be Subset of TOP-REAL n;
pred B is_outside_component_of A means
:Def4: B is_a_component_of A` & B is not Bounded;
end;

theorem Th18: for A being Subset of TOP-REAL n,
 B being Subset of TOP-REAL n holds
   B is_outside_component_of A iff
  ex C being Subset of ((TOP-REAL n)|(A`)) st
 C=B & C is_a_component_of ((TOP-REAL n)|(A`)) &
  C is not bounded Subset of Euclid n
proof let A be Subset of TOP-REAL n,
 B be Subset of TOP-REAL n;
  A1: B is_a_component_of A` iff
  ex C being Subset of (TOP-REAL n)|(A`) st
   C=B & C is_a_component_of (TOP-REAL n)|(A`) by CONNSP_1:def 6;
A2: B is_outside_component_of A iff
     B is_a_component_of A` & not B is Bounded by Def4;
  thus B is_outside_component_of A implies
   (ex C being Subset of ((TOP-REAL n)|(A`)) st
   C=B & C is_a_component_of ((TOP-REAL n)|(A`)) &
   C is not bounded Subset of Euclid n)
   proof assume A3: B is_outside_component_of A;
     then consider C being Subset of (TOP-REAL n)|(A`) such that
     A4:C=B & C is_a_component_of (TOP-REAL n)|(A`) by A1,Def4;
       B is Subset of Euclid n by TOPREAL3:13;
     then reconsider D2=B as Subset of Euclid n;
        now assume for D being Subset of Euclid n st D=C holds D is bounded;
        then D2 is bounded by A4;
       hence contradiction by A2,A3,Def2;
      end;
    hence thesis by A4;
   end;
     given C being Subset of ((TOP-REAL n)|(A`)) such that
     A5:C=B & C is_a_component_of ((TOP-REAL n)|(A`)) &
       C is not bounded Subset of Euclid n;
       for D4 being Subset of Euclid n
      st D4=B holds not D4 is bounded by A5;
     then A6:not B is Bounded by Def2;
       B is_a_component_of A` by A5,CONNSP_1:def 6;
    hence thesis by A6,Def4;
end;

theorem for A,B being Subset of TOP-REAL n st
B is_inside_component_of A holds B c= A`
proof let A,B be Subset of TOP-REAL n;
  assume B is_inside_component_of A;
   then B is_a_component_of A` by Def3;
  hence B c= A` by SPRECT_1:7;
end;

theorem for A,B being Subset of TOP-REAL n st
B is_outside_component_of A holds B c= A`
proof let A,B be Subset of TOP-REAL n;
  assume B is_outside_component_of A;
  then B is_a_component_of A` by Def4;
  hence B c= A` by SPRECT_1:7;
end;

definition let n;let A be Subset of TOP-REAL n;
func BDD A -> Subset of TOP-REAL n equals
:Def5: union{B where B is Subset of TOP-REAL n: B is_inside_component_of A};
correctness
 proof
    union{B where B is Subset of TOP-REAL n:
      B is_inside_component_of A} c= the carrier of TOP-REAL n
   proof let x;assume x in union{B where B is Subset of TOP-REAL n:
      B is_inside_component_of A};
     then consider y being set such that
     A1: x in y & y in
     {B where B is Subset of TOP-REAL n:
      B is_inside_component_of A} by TARSKI:def 4;
      consider B being Subset of TOP-REAL n such that
      A2:y=B & B is_inside_component_of A by A1;
     thus x in the carrier of TOP-REAL n by A1,A2;
   end;
  hence thesis;
 end;
end;

definition let n;let A be Subset of TOP-REAL n;
func UBD A -> Subset of TOP-REAL n equals
:Def6: union{B where B is Subset of TOP-REAL n: B is_outside_component_of A};
correctness
 proof
    union{B where B is Subset of TOP-REAL n:
      B is_outside_component_of A} c= the carrier of TOP-REAL n
   proof let x;assume x in
     union{B where B is Subset of TOP-REAL n:
      B is_outside_component_of A};
     then consider y being set such that
     A1: x in y & y in
     {B where B is Subset of TOP-REAL n:
      B is_outside_component_of A} by TARSKI:def 4;
      consider B being Subset of TOP-REAL n such that
      A2:y=B & B is_outside_component_of A by A1;
   thus x in the carrier of TOP-REAL n by A1,A2;
   end;
  hence thesis;
 end;
end;

theorem Th21:  [#](TOP-REAL n) is convex
proof
   let w1,w2 be Point of TOP-REAL n;
       LSeg(w1,w2) c= the carrier of TOP-REAL n;
    hence thesis by PRE_TOPC:12;
end;

theorem Th22:  [#](TOP-REAL n) is connected
proof [#](TOP-REAL n) is convex by Th21;
 hence [#](TOP-REAL n) is connected by Th14;
end;

definition let n;
 cluster [#](TOP-REAL n) -> connected;
coherence by Th22;
end;

theorem Th23:  [#](TOP-REAL n) is_a_component_of TOP-REAL n
proof
  set A=[#](TOP-REAL n);
     for B being Subset of TOP-REAL n st
           B is connected holds A c= B implies A = B
    proof let B be Subset of TOP-REAL n;
     assume B is connected;
      thus A c= B implies A = B
      proof assume A1:A c= B;
          B c= the carrier of TOP-REAL n;
        then B c= [#] (TOP-REAL n) by PRE_TOPC:12;
       hence A=B by A1,XBOOLE_0:def 10;
      end;
    end;
 hence thesis by CONNSP_1:def 5;
end;

theorem Th24:for A being Subset of TOP-REAL n holds
BDD A is a_union_of_components of (TOP-REAL n)|A`
proof let A be Subset of TOP-REAL n;
    {B where B is Subset of TOP-REAL n:
    B is_inside_component_of A} c= bool (the carrier of ((TOP-REAL n)|A`))
    proof let x; assume x in {B where B is Subset of TOP-REAL n:
     B is_inside_component_of A};
     then consider B being Subset of TOP-REAL n such that
     A1: x=B & B is_inside_component_of A;
     consider C being Subset of ((TOP-REAL n)|(A`)) such that
     A2: C=B & C is_a_component_of ((TOP-REAL n)|(A`)) &
      C is bounded Subset of Euclid n by A1,Th17;
     thus x in bool (the carrier of ((TOP-REAL n)|A`)) by A1,A2;
    end;
  then reconsider F0={B where B is Subset of TOP-REAL n:
    B is_inside_component_of A} as Subset-Family of
      the carrier of ((TOP-REAL n)|A`) by SETFAM_1:def 7;
  reconsider F0 as Subset-Family of (TOP-REAL n)|A`;
  A3: BDD A=union F0 by Def5;
     for B0 being Subset of ((TOP-REAL n)|A`) st B0 in F0
    holds B0 is_a_component_of ((TOP-REAL n)|A`)
    proof let B0 be Subset of ((TOP-REAL n)|A`);
      assume B0 in F0;
      then consider B being Subset of TOP-REAL n such that
      A4:B=B0 & B is_inside_component_of A;
      consider
        C being Subset of ((TOP-REAL n)|(A`)) such that
      A5: C=B & C is_a_component_of ((TOP-REAL n)|(A`)) &
        C is bounded Subset of Euclid n by A4,Th17;
     thus B0 is_a_component_of ((TOP-REAL n)|A`) by A4,A5;
    end;
 hence thesis by A3,CONNSP_3:def 2;
end;

theorem Th25:for A being Subset of TOP-REAL n holds
UBD A is a_union_of_components of (TOP-REAL n)|A`
proof let A be Subset of TOP-REAL n;
    {B where B is Subset of TOP-REAL n:
    B is_outside_component_of A} c= bool (the carrier of ((TOP-REAL n)|A`))
    proof let x;assume x in {B where B is Subset of TOP-REAL n:
     B is_outside_component_of A};
     then consider B being Subset of TOP-REAL n such that
     A1: x=B & B is_outside_component_of A;
       ex C being Subset of ((TOP-REAL n)|(A`)) st
     C=B & C is_a_component_of ((TOP-REAL n)|(A`)) &
     C is not bounded Subset of Euclid n by A1,Th18;
     hence thesis by A1;
    end;
  then reconsider F0={B where B is Subset of TOP-REAL n:
    B is_outside_component_of A} as Subset-Family of
      the carrier of ((TOP-REAL n)|A`) by SETFAM_1:def 7;
  reconsider F0 as Subset-Family of ((TOP-REAL n)|A`);
  A2: UBD A=union F0 by Def6;
     for B0 being Subset of ((TOP-REAL n)|A`) st B0 in F0
    holds B0 is_a_component_of ((TOP-REAL n)|A`)
    proof let B0 be Subset of ((TOP-REAL n)|A`);
      assume B0 in F0;
      then consider B being Subset of TOP-REAL n such that
      A3:B=B0 & B is_outside_component_of A;
        ex C being Subset of ((TOP-REAL n)|(A`)) st
      C=B & C is_a_component_of ((TOP-REAL n)|(A`)) &
      C is not bounded Subset of Euclid n by A3,Th18;
     hence thesis by A3;
    end;
 hence thesis by A2,CONNSP_3:def 2;
end;

theorem Th26: for A being Subset of TOP-REAL n,
 B being Subset of TOP-REAL n st B is_inside_component_of A
holds B c= BDD A
proof let A be Subset of TOP-REAL n,
 B be Subset of TOP-REAL n;
 assume A1:B is_inside_component_of A;
 let x;assume A2:x in B;
     B in {B2 where B2 is Subset of TOP-REAL n:
        B2 is_inside_component_of A} by A1;
   then x in union{B2 where B2 is Subset of TOP-REAL n:
        B2 is_inside_component_of A} by A2,TARSKI:def 4;
  hence x in BDD A by Def5;
end;

theorem Th27: for A being Subset of TOP-REAL n,
 B being Subset of TOP-REAL n st B is_outside_component_of A
holds B c= UBD A
proof let A be Subset of TOP-REAL n,
  B be Subset of TOP-REAL n;
 assume A1:B is_outside_component_of A;
  let x;assume A2:x in B;
     B in {B2 where B2 is Subset of TOP-REAL n:
        B2 is_outside_component_of A} by A1;
   then x in union{B2 where B2 is Subset of TOP-REAL n:
        B2 is_outside_component_of A} by A2,TARSKI:def 4;
  hence x in UBD A by Def6;
end;

theorem Th28:for A being Subset of TOP-REAL n holds
 BDD A misses UBD A
proof let A be Subset of TOP-REAL n;
 assume A1:(BDD A) /\ (UBD A) <>{};
   consider x being Element of (BDD A) /\ (UBD A);
   A2:x in BDD A & x in UBD A by A1,XBOOLE_0:def 3;
   then x in union{B where B is Subset of TOP-REAL n:
           B is_inside_component_of A} by Def5;
   then consider y being set such that
   A3:x in y & y in {B where B is Subset of TOP-REAL n:
           B is_inside_component_of A} by TARSKI:def 4;
   consider B being Subset of TOP-REAL n such that
   A4:y=B & B is_inside_component_of A by A3;
   consider C being Subset of ((TOP-REAL n)|(A`)) such that
   A5: C=B & C is_a_component_of ((TOP-REAL n)|(A`)) &
     C is bounded Subset of Euclid n by A4,Th17;
     x in union{B2 where B2 is Subset of TOP-REAL n:
           B2 is_outside_component_of A} by A2,Def6;
   then consider y2 being set such that
   A6:x in y2 & y2 in {B2 where B2 is Subset of TOP-REAL n:
           B2 is_outside_component_of A} by TARSKI:def 4;
   consider B2 being Subset of TOP-REAL n such that
   A7:y2=B2 & B2 is_outside_component_of A by A6;
   consider C2 being Subset of ((TOP-REAL n)|(A`)) such that
   A8: C2=B2 & C2 is_a_component_of ((TOP-REAL n)|(A`)) &
   C2 is not bounded Subset of Euclid n by A7,Th18;
     C /\ C2<>{}((TOP-REAL n)|(A`)) by A3,A4,A5,A6,A7,A8,XBOOLE_0:def 3;
   then C meets C2 by XBOOLE_0:def 7;
  hence contradiction by A5,A8,CONNSP_1:37;
end;

theorem Th29:for A being Subset of TOP-REAL n holds
BDD A c= A`
proof let A be Subset of TOP-REAL n;
  reconsider D=BDD A as Subset of (TOP-REAL n)|A` by Th24;
    D c= the carrier of ((TOP-REAL n)|A`);
 hence thesis by JORDAN1:1;
end;

theorem Th30:for A being Subset of TOP-REAL n holds
UBD A c= A`
proof let A be Subset of TOP-REAL n;
  reconsider D=UBD A as Subset of (TOP-REAL n)|A` by Th25;
    D c= the carrier of ((TOP-REAL n)|A`);
 hence thesis by JORDAN1:1;
end;

theorem Th31:for A being Subset of TOP-REAL n holds
(BDD A) \/ (UBD A) = A`
proof let A be Subset of TOP-REAL n;
 A1:(BDD A) c= A` by Th29;
   (UBD A) c= A` by Th30;
 then A2:(BDD A) \/ (UBD A) c= A` by A1,XBOOLE_1:8;
   A` c= (BDD A) \/ (UBD A)
 proof let z be set;assume A3:z in A`;
   then reconsider p=z as Element of A`;
   reconsider q=p as Point of (TOP-REAL n)|A` by JORDAN1:1;
   reconsider B=A` as non empty Subset of TOP-REAL n by A3;
A4:(TOP-REAL n)|B is non empty;
   then A5:skl q is_a_component_of (TOP-REAL n)|A` by CONNSP_1:43;
     skl q is Subset of [#]((TOP-REAL n)|A`) by PRE_TOPC:12;
   then skl q is Subset of A` by PRE_TOPC:def 10;
   then skl q is Subset of TOP-REAL n by XBOOLE_1:1;
   then reconsider G=skl q as Subset of TOP-REAL n;
   A6:q in G by A4,CONNSP_1:40;
   A7:G is_a_component_of A` by A5,CONNSP_1:def 6;
   per cases;
   suppose G is Bounded;
     then G is_inside_component_of A by A7,Def3;
     then G c= BDD A by Th26;
    hence z in (BDD A) \/ (UBD A) by A6,XBOOLE_0:def 2;
   suppose not G is Bounded;
     then G is_outside_component_of A by A7,Def4;
     then G c= UBD A by Th27;
    hence z in (BDD A) \/ (UBD A) by A6,XBOOLE_0:def 2;
 end;
 hence thesis by A2,XBOOLE_0:def 10;
end;

reserve u for Point of Euclid n;

theorem Th32:
 for G being non empty TopSpace,
   w1,w2,w3 being Point of G,
   h1,h2 being map of I[01],G
   st h1 is continuous & w1=h1.0 & w2=h1.1 &
  h2 is continuous & w2=h2.0 & w3=h2.1 holds
  ex h3 being map of I[01],G st h3 is continuous & w1=h3.0 & w3=h3.1 &
  rng h3 c= (rng h1) \/ (rng h2)
proof
   let G be non empty TopSpace,
   w1,w2,w3 be Point of G,
   h1,h2 be map of I[01],G;
    assume A1:h1 is continuous & w1=h1.0 & w2=h1.1 &
     h2 is continuous & w2=h2.0 & w3=h2.1;
     then reconsider g1=h1 as Path of w1,w2 by BORSUK_2:def 1;
     reconsider g2=h2 as Path of w2,w3 by A1,BORSUK_2:def 1;
    set P1=g1,P2=g2,p1=w1,p3=w3;
    ex P0 being Path of p1,p3 st P0 is continuous & P0.0=p1 & P0.1=p3 &
   for t being Point of I[01], t' being Real st t = t' holds
    ( 0 <= t' & t' <= 1/2 implies P0.t = P1.(2*t') ) &
    ( 1/2 <= t' & t' <= 1 implies P0.t = P2.(2*t'-1) )
  proof ::This proof is almost a copy of BORSUK_2:def 4(proof of Existence)
   set e1 = P[01](0, 1/2, (#)(0,1), (0,1)(#));
   set e2 = P[01](1/2, 1, (#)(0,1), (0,1)(#));
   set E1 = P1 * e1;
   set E2 = P2 * e2;
   set f = E1 +* E2;
A2:dom e1 = the carrier of Closed-Interval-TSpace(0,1/2) by FUNCT_2:def 1
         .= [.0,1/2.] by TOPMETR:25;
A3:dom e2 = the carrier of Closed-Interval-TSpace(1/2,1) by FUNCT_2:def 1
         .= [.1/2,1 qua Real.] by TOPMETR:25;
A4: dom P1 = the carrier of I[01] & dom P2 = the carrier of I[01]
       by FUNCT_2:def 1;
then A5:rng e1 c= dom P1 by TOPMETR:27;
     rng e2 c= the carrier of Closed-Interval-TSpace(0,1);
then A6: dom E2 = dom e2 by A4,RELAT_1:46,TOPMETR:27;
A7: dom f = dom E1 \/ dom E2 by FUNCT_4:def 1
        .= [.0,1/2.] \/ [.1/2,1 qua Real.] by A2,A3,A5,A6,RELAT_1:46
        .= the carrier of I[01] by BORSUK_1:83,TREAL_1:2;
A8: for t' being Real st 0 <= t' & t' <= 1/2 holds E1.t' = P1.(2*t')
   proof
    let t' be Real such that
A9:  0 <= t' & t' <= 1/2;
      dom e1 = the carrier of Closed-Interval-TSpace(0,1/2) by FUNCT_2:def 1;
     then dom e1 = [.0, 1/2.] by TOPMETR:25
          .= {r : 0 <= r & r <= 1/2 } by RCOMP_1:def 1;
then A10:  t' in dom e1 by A9;
    then reconsider s = t' as Point of Closed-Interval-TSpace (0, 1/2)
      by FUNCT_2:def 1;
    reconsider r1 = (#)(0,1), r2 = (0,1)(#) as Real
      by BORSUK_1:def 17,def 18,TREAL_1:8;
      e1.s = ((r2 - r1)/(1/2 - 0))*t' + ((1/2)*r1 - 0 * r2)/(1/2 - 0)
      by TREAL_1:14
        .= 2*t' by BORSUK_1:def 17,def 18,TREAL_1:8;
    hence thesis by A10,FUNCT_1:23;
   end;
     not 0 in { r : 1/2 <= r & r <= 1 }
   proof
    assume 0 in { r : 1/2 <= r & r <= 1 };
     then ex rr being Real st rr = 0 & 1/2 <= rr & rr <= 1;
    hence thesis;
   end;
   then not 0 in dom E2 by A3,A6,RCOMP_1:def 1;
then A11:f.0 = E1.0 by FUNCT_4:12
      .= P1.(2*0) by A8
      .= p1 by A1;
     rng E1 c= rng P1 & rng E2 c= rng P2 by RELAT_1:45;
then A12:rng E1 c= the carrier of G &
rng E2 c= the carrier of G by XBOOLE_1:1;
A13: rng f c= rng E1 \/ rng E2 by FUNCT_4:18;
     rng E1 \/ rng E2 c= (the carrier of G) \/ the carrier of G
      by A12,XBOOLE_1:13;
   then rng f c= the carrier of G by A13,XBOOLE_1:1;
   then f is Function of the carrier of I[01], the carrier of G
      by A7,FUNCT_2:def 1,RELSET_1:11;
   then reconsider f as map of I[01], G ;
   reconsider T1 = Closed-Interval-TSpace (0, 1/2),
              T2 = Closed-Interval-TSpace (1/2, 1) as SubSpace of I[01]
                by TOPMETR:27,TREAL_1:6;
A14:e1 is continuous by TREAL_1:15;
A15:e2 is continuous by TREAL_1:15;
     E1 is Function of the carrier of Closed-Interval-TSpace (0, 1/2),
            the carrier of G by FUNCT_2:19,TOPMETR:27;
   then reconsider ff = E1 as map of T1, G ;
A16:E2 is Function of the carrier of Closed-Interval-TSpace (1/2, 1),
            the carrier of G by FUNCT_2:19,TOPMETR:27;
   then reconsider gg = E2 as map of T2, G ;
     1/2 in { r : 0 <= r & r <= 1 };
   then reconsider pol = 1/2 as Point of I[01] by BORSUK_1:83,RCOMP_1:def 1;
A17:ff is continuous & gg is continuous by A1,A14,A15,TOPMETR:27,TOPS_2:58;
A18:[#] T1 = the carrier of T1 by PRE_TOPC:12
       .= [.0,1/2.] by TOPMETR:25;
A19:[#] T2 = the carrier of T2 by PRE_TOPC:12
       .= [.1/2,1 qua Real.] by TOPMETR:25;
then A20: ([#] T1) \/ ([#] T2) = [.0,1 qua Real.] by A18,TREAL_1:2
                  .= [#] I[01] by BORSUK_1:83,PRE_TOPC:12;
A21: for t' being Real st 1/2 <= t' & t' <= 1 holds E2.t' = P2.(2*t'-1)
   proof
    let t' be Real such that
A22:      1/2 <= t' & t' <= 1;
      dom e2 = the carrier of Closed-Interval-TSpace(1/2,1) by FUNCT_2:def 1;
     then dom e2 = [.1/2,1 qua Real.] by TOPMETR:25
          .= {r : 1/2 <= r & r <= 1 } by RCOMP_1:def 1;
then A23:  t' in dom e2 by A22;
    then reconsider s = t' as Point of Closed-Interval-TSpace (1/2,1)
      by FUNCT_2:def 1;
    reconsider r1 = (#)(0,1), r2 = (0,1)(#) as Real
      by BORSUK_1:def 17,def 18,TREAL_1:8;
      e2.s = ((r2 - r1)/(1 - 1/2))*t' + (1 * r1 - (1/2)*r2)/(1 - 1/2)
      by TREAL_1:14
        .= 2*t' + (- 1) by BORSUK_1:def 17,def 18,TREAL_1:8
        .= 2*t' - 1 by XCMPLX_0:def 8;
    hence thesis by A23,FUNCT_1:23;
   end;
A24:ff.(1/2) = P2.(2*(1/2)-1) by A1,A8
           .= gg.pol by A21;
A25:([#] T1) /\ ([#] T2) = {pol} by A18,A19,TOPMETR2:1;
     R^1 is_T2 by PCOMPS_1:38,TOPMETR:def 7;
   then T1 is compact & T2 is compact & I[01] is_T2 & ff.pol = gg.pol
     by A24,HEINE:11,TOPMETR:3;
  then consider h being map of I[01], G such that
A26:   h = ff+*gg & h is continuous by A17,A20,A25,TOPMETR2:4;
     1 in { r : 1/2 <= r & r <= 1 };
   then 1 in dom E2 by A3,A6,RCOMP_1:def 1;
 then A27:  f.1 = E2.1 by FUNCT_4:14
      .= P2.(2*1-1) by A21
      .= p3 by A1;
   then reconsider f as Path of p1, p3 by A11,A26,BORSUK_2:def 1;
  for t being Point of I[01], t' being Real st t = t' holds
     ( 0 <= t' & t' <= 1/2 implies f.t = P1.(2*t') ) &
     ( 1/2 <= t' & t' <= 1 implies f.t = P2.(2*t'-1) )
   proof
    let t be Point of I[01], t' be Real;
    assume A28: t = t';
    thus 0 <= t' & t' <= 1/2 implies f.t = P1.(2*t')
    proof
     assume A29: 0 <= t' & t' <= 1/2;
     then t' in { r : 0 <= r & r <= 1/2 };
then A30:  t' in [.0,1/2.] by RCOMP_1:def 1;
     per cases;
     suppose A31: t' <> 1/2;
       not t' in dom E2
     proof
      assume t' in dom E2;
      then t' in [.0,1/2.] /\ [.1/2,1 qua Real.] by A3,A6,A30,XBOOLE_0:def 3;
      then t' in {1/2} by TOPMETR2:1;
      hence thesis by A31,TARSKI:def 1;
     end;
     then f.t = E1.t by A28,FUNCT_4:12
        .= P1.(2*t') by A8,A28,A29;
     hence thesis;
     suppose A32: t' = 1/2;
       1/2 in { r : 1/2 <= r & r <= 1 };
     then 1/2 in [.1/2, 1 qua Real.] by RCOMP_1:def 1;
     then 1/2 in the carrier of Closed-Interval-TSpace(1/2,1)
         by TOPMETR:25;
     then t in dom E2 by A16,A28,A32,FUNCT_2:def 1;
     then f.t = E2.(1/2) by A28,A32,FUNCT_4:14
        .= P1.(2*t') by A8,A24,A32;
     hence thesis;
    end;
    thus 1/2 <= t' & t' <= 1 implies f.t = P2.(2*t'-1)
    proof
     assume A33: 1/2 <= t' & t' <= 1;
     then t' in { r : 1/2 <= r & r <= 1 };
      then t' in [.1/2,1 qua Real.] by RCOMP_1:def 1;
      then f.t = E2.t by A3,A6,A28,FUNCT_4:14
        .= P2.(2*t'-1) by A21,A28,A33;
     hence thesis;
    end;
   end;
   hence thesis by A11,A26,A27;
  end;
  then consider P0 being Path of p1,p3 such that
  A34: P0 is continuous & P0.0=p1 & P0.1=p3 &
   for t being Point of I[01], t' being Real st t = t' holds
    ( 0 <= t' & t' <= 1/2 implies P0.t = P1.(2*t') ) &
    ( 1/2 <= t' & t' <= 1 implies P0.t = P2.(2*t'-1) );
    rng P0 c= (rng P1) \/ (rng P2)
  proof let x be set;assume x in rng P0;
    then consider z being set such that
    A35:z in dom P0 & x=P0.z by FUNCT_1:def 5;
A36: dom P0=the carrier of I[01] by FUNCT_2:def 1;
    then reconsider r=z as Real by A35,BORSUK_1:83;
    A37:0<=r & r<=1 by A35,A36,BORSUK_1:83,TOPREAL5:1;
    A38:dom g1=the carrier of I[01] by FUNCT_2:def 1;
    A39:dom g2=the carrier of I[01] by FUNCT_2:def 1;
    per cases;
    suppose A40:r<=1/2;
      then A41:P0.z=P1.(2*r) by A34,A35,A36,A37;
      A42:0<=2*r by A37,REAL_2:121;
       2*r <= 2*(1/2) by A40,AXIOMS:25;
      then 2*r in the carrier of I[01] by A42,BORSUK_1:83,TOPREAL5:1;
      then P0.z in rng g1 by A38,A41,FUNCT_1:def 5;
     hence x in (rng P1) \/ (rng P2) by A35,XBOOLE_0:def 2;
    suppose A43:r>1/2;
      then A44:P0.z=P2.(2*r-1) by A34,A35,A36,A37;
        2*(1/2)=1;
      then 0+1<=2*r by A43,AXIOMS:25;
      then A45:0<=2*r-1 by REAL_1:84;
        2*r<=2*1 by A37,AXIOMS:25;
      then 2*r<=1+1;
      then 2*r-1<=1 by REAL_1:86;
      then 2*r-1 in the carrier of I[01] by A45,BORSUK_1:83,TOPREAL5:1;
      then P0.z in rng g2 by A39,A44,FUNCT_1:def 5;
     hence x in (rng P1) \/ (rng P2) by A35,XBOOLE_0:def 2;
  end;
  hence thesis by A34;
end;

theorem Th33: for P being Subset of TOP-REAL n
 st P=REAL n holds P is connected
 proof let P be Subset of TOP-REAL n;
  assume A1:P=(REAL n);
    for p1,p2 being Point of TOP-REAL n
  st p1 in P & p2 in P holds LSeg(p1,p2) c= P
  proof let p1,p2 be Point of TOP-REAL n;
    assume p1 in P & p2 in P;
      the carrier of TOP-REAL n=REAL n by EUCLID:25;
   hence LSeg(p1,p2) c= P by A1;
  end;
  then P is convex by JORDAN1:def 1;
  hence P is connected by Th14;
 end;

definition let n;
   func 1*n -> FinSequence of REAL equals :Def7:
   n |-> (1 qua Real);
coherence by FINSEQ_2:77;
end;

definition let n;
  redefine func 1*n -> Element of REAL n;
    coherence
     proof
      A1:n-tuples_on REAL = REAL n & 1*n = n |-> (1 qua Real)
                              by Def7,EUCLID:def 1;
      reconsider f=1*n as FinSequence;
        len f = len (n|->(1 qua Real)) by Def7 .=n by FINSEQ_2:69;
      hence thesis by A1,FINSEQ_2:110;
     end;
end;

definition let n;
  func 1.REAL n -> Point of TOP-REAL n equals
   :Def8:  1*n;
  coherence by EUCLID:25;
end;

theorem abs 1*n = n |-> (1 qua Real)
 proof
  reconsider f= (n |-> (1 qua Real)) as FinSequence of REAL by FINSEQ_2:77;
  thus abs 1*n = abs f by Def7
          .= absreal*(n |-> (1 qua Real)) by EUCLID:def 3
          .= n |-> absreal.(1 qua Real) by FINSEQOP:17
          .= n |-> abs(1 qua Real) by EUCLID:def 2
          .= n |-> 1 by ABSVALUE:def 1;
 end;

theorem Th35: |.1*n.| = sqrt n
 proof
  reconsider f= (n |-> (1 qua Real)) as FinSequence of REAL by FINSEQ_2:77;
   thus |.1*n .| = sqrt Sum sqr 1*n by EUCLID:def 5
               .= sqrt Sum sqr f by Def7
              .= sqrt Sum f by RVSUM_1:82,SQUARE_1:59
             .= sqrt (n*1) by RVSUM_1:110
            .= sqrt n;
 end;

theorem Th36: 1.REAL 1 = <* 1 qua Real *>
 proof
  reconsider f= (1 |-> (1 qua Real)) as FinSequence of REAL by FINSEQ_2:77;
   thus 1.REAL 1=1*1 by Def8 .=f by Def7 .=<* 1 qua Real *> by FINSEQ_2:73;
 end;

theorem Th37: |. (1.REAL n) .| = sqrt n
 proof A1:1.REAL n=1*n by Def8;
     |. 1*n .|=sqrt n by Th35;
  hence thesis by A1,JGRAPH_1:def 5;
 end;

theorem Th38:
1<=n implies 1<=|. (1.REAL n) .|
 proof assume A1: 1<=n;
    |.1.REAL n.|=sqrt n by Th37;
  hence thesis by A1,SQUARE_1:83,94;
 end;

theorem Th39: for W being Subset of Euclid n
 st n>=1 & W=REAL n holds W is not bounded
 proof let W be Subset of Euclid n;
   assume A1:n>=1 & W=(REAL n);
   assume W is bounded;
     then consider r being Real such that
     A2:0<r & for x,y being Point of Euclid n st
     x in W & y in W holds dist(x,y)<=r by TBSP_1:def 9;
     reconsider x0=(r+1)*(1.REAL n) as Point of Euclid n by TOPREAL3:13;
       (r+1)*(1.REAL n) in the carrier of TOP-REAL n;
     then A3:x0 in W by A1,EUCLID:25;
     reconsider y0=0.REAL n as Point of Euclid n by TOPREAL3:13;
       0.REAL n in the carrier of TOP-REAL n;
     then y0 in W by A1,EUCLID:25;
     then dist(x0,y0)<=r by A2,A3;
     then |.(r+1)*(1.REAL n) -0.REAL n.|<=r by JGRAPH_1:45;
     then |.(r+1)*(1.REAL n).|<=r by Th13;
     then abs(r+1)*|.(1.REAL n).|<=r by TOPRNS_1:8;
     then A4:abs(r+1)*(sqrt n)<=r by Th37;
       r+1>r by REAL_1:69;
     then A5:r+1>0 by A2,AXIOMS:22;
     then A6:(r+1)*(sqrt n)<=r by A4,ABSVALUE:def 1;
       (sqrt 1)<=(sqrt n) by A1,SQUARE_1:94;
     then (r+1)*1<=(r+1)*(sqrt n) by A5,AXIOMS:25,SQUARE_1:83;
     then (r+1)*1<=r by A6,AXIOMS:22;
     then (r+1)-r<=r-r by REAL_1:49;
     then 1<=r-r by XCMPLX_1:26;
     then 1<=0 by XCMPLX_1:14;
    hence contradiction;
 end;

theorem Th40: for A being Subset of TOP-REAL n holds
 A is Bounded iff ex r being Real st for q being Point of TOP-REAL n st
 q in A holds |.q.|<r
proof let A be Subset of TOP-REAL n;
 hereby assume A is Bounded;
   then consider C being Subset of Euclid n such that
    A1:C=A & C is bounded by Def2;
  per cases;
  suppose A2:C<>{};
    consider x0 being Element of C;
         x0 in C by A2;
    then reconsider x0 as Point of Euclid n;
    consider r being Real such that
    A3: 0<r & for x,y being Point of (Euclid n) st x in C & y in C holds
                               dist(x,y) <= r by A1,TBSP_1:def 9;
    reconsider o=0.REAL n as Point of Euclid n by TOPREAL3:13;
    set R0=r+dist(o,x0)+1;
     for q being Point of TOP-REAL n st q in A holds |.q.|<R0
    proof let q1 be Point of TOP-REAL n;
     assume A4:q1 in A;
      reconsider z=q1 as Point of Euclid n by TOPREAL3:13;
        |.q1-(0.REAL n).|=dist(o,z) by JGRAPH_1:45;
      then A5: |.q1.|=dist(o,z) by Th13;
      A6:dist(x0,z)<=r by A1,A3,A4;
      A7:dist(o,z)<=dist(o,x0)+dist(x0,z) by METRIC_1:4;
        dist(o,x0)+dist(x0,z)<=dist(o,x0)+r by A6,AXIOMS:24;
then A8:   dist(o,z)<=dist(o,x0)+r by A7,AXIOMS:22;
        r+dist(o,x0)<r+dist(o,x0)+1 by REAL_1:69;
      hence |.q1.|<R0 by A5,A8,AXIOMS:22;
    end;
  hence ex r2 being Real st for q being Point of TOP-REAL n st
 q in A holds |.q.|<r2;
 suppose C={};
  then for q being Point of TOP-REAL n st q in A holds |.q.|<1 by A1;
  hence ex r2 being Real st for q being Point of TOP-REAL n st
 q in A holds |.q.|<r2;
end;
 given r being Real such that
 A9: for q being Point of TOP-REAL n st q in A holds |.q.|<r;
    A is Subset of Euclid n by TOPREAL3:13;
  then reconsider C=A as Subset of Euclid n;
    now per cases;
  suppose A10:C<>{};
    consider x0 being Element of C;
      x0 in C by A10;
    then reconsider x0 as Point of Euclid n;
    reconsider q0=x0 as Point of TOP-REAL n by TOPREAL3:13;
     |.q0.|<r by A9,A10;
    then A11:0<r by TOPRNS_1:26;
    reconsider o=0.REAL n as Point of Euclid n by TOPREAL3:13;
    set R0=r+r;
    A12:0<R0 by A11,Th6;
      for x,y being Point of (Euclid n) st x in C & y in C holds dist(x,y) <=
R0
    proof let x,y be Point of (Euclid n);
      assume A13:x in C & y in C;
      then reconsider q1=x as Point of TOP-REAL n;
      reconsider q2=y as Point of TOP-REAL n by A13;
      A14:dist(x,y)<=dist(x,o)+dist(o,y) by METRIC_1:4;
        dist(x,o)=|.q1-0.REAL n.| by JGRAPH_1:45
      .=|.q1.| by Th13;
      then A15:dist(x,o) <r by A9,A13;
        dist(o,y)=|.q2-(0.REAL n) .| by JGRAPH_1:45
      .=|.q2.| by Th13;
      then dist(o,y) <r by A9,A13;
      then dist(x,o)+dist(o,y)<=r+r by A15,REAL_1:55;
     hence dist(x,y) <= R0 by A14,AXIOMS:22;
    end;
   hence C is bounded by A12,TBSP_1:def 9;
   suppose C={};
   then C = {}Euclid n;
   hence C is bounded by TBSP_1:14;
   end;
 hence A is Bounded by Def2;
end;

theorem Th41: n>=1 implies not [#](TOP-REAL n) is Bounded
proof assume A1:n>=1;
 assume [#](TOP-REAL n) is Bounded;
   then consider C being Subset of Euclid n
    such that A2:C=[#](TOP-REAL n) & C is bounded by Def2;
     C=the carrier of TOP-REAL n by A2,PRE_TOPC:12;
   then C=REAL n by EUCLID:25;
  hence contradiction by A1,A2,Th39;
end;

theorem Th42: n>=1 implies UBD {}(TOP-REAL n)=REAL n
proof assume A1:n>=1;
    UBD {}(TOP-REAL n) c= the carrier of TOP-REAL n;
 hence UBD {}(TOP-REAL n) c= REAL n by EUCLID:25;
  let x be set;assume x in REAL n;
    then A2:x in the carrier of TOP-REAL n by EUCLID:25;
    set A={}(TOP-REAL n);
      A`=[#](TOP-REAL n) by PRE_TOPC:27;
    then A3:A`=[#](TOP-REAL n);
    A4:(TOP-REAL n)| [#](TOP-REAL n)=TOP-REAL n by TSEP_1:3;
    A5: [#]((TOP-REAL n)|A`) = [#](TOP-REAL n) by A3,TSEP_1:3;
    A6:[#]((TOP-REAL n)|A`) is_a_component_of (TOP-REAL n)|A` by A3,A4,Th23;
      now assume A7:for D being Subset of Euclid n
      st D=[#]((TOP-REAL n)|A`) holds D is bounded;
        [#]((TOP-REAL n)|A`)=the carrier of TOP-REAL n by A3,A4,PRE_TOPC:12;
      then [#]((TOP-REAL n)|A`) c= the carrier of Euclid n by TOPREAL3:13;
     then reconsider D1=[#]((TOP-REAL n)|A`) as Subset of Euclid n
     ;
        D1 is bounded by A7;
      then [#](TOP-REAL n) is Bounded by A3,A4,Def2;
     hence contradiction by A1,Th41;
    end;
    then x in [#](TOP-REAL n) &
    [#](TOP-REAL n) is_outside_component_of {}(TOP-REAL n)
     by A2,A5,A6,Th18,PRE_TOPC:12;
    then x in [#](TOP-REAL n) & [#](TOP-REAL n) in
      {B2 where B2 is Subset of TOP-REAL n:
         B2 is_outside_component_of {}(TOP-REAL n)};
    then x in union{B2 where B2 is Subset of TOP-REAL n:
         B2 is_outside_component_of {}(TOP-REAL n)} by TARSKI:def 4;
   hence x in UBD {}(TOP-REAL n) by Def6;
end;

theorem Th43:for w1,w2,w3 being Point of TOP-REAL n,
 P being non empty Subset of TOP-REAL n,
 h1,h2 being map of I[01],(TOP-REAL n)|P st
 h1 is continuous & w1=h1.0 & w2=h1.1 &
 h2 is continuous & w2=h2.0 & w3=h2.1
 holds ex h3 being map of I[01],(TOP-REAL n)|P st
 h3 is continuous & w1=h3.0 & w3=h3.1
 proof let w1,w2,w3 be Point of TOP-REAL n,
   P be non empty Subset of TOP-REAL n,
   h1,h2 be map of I[01],(TOP-REAL n)|P;
  assume A1:h1 is continuous & w1=h1.0 & w2=h1.1 &
    h2 is continuous & w2=h2.0 & w3=h2.1;
A2:  0 in [.0 qua Real,1 qua Real.] by TOPREAL5:1;
      1 in [.0 qua Real,1 qua Real.] by TOPREAL5:1;
    then reconsider p1=w1,p2=w2,p3=w3 as Point of (TOP-REAL n)|P
     by A1,A2,BORSUK_1:83,FUNCT_2:7;
    reconsider P1=h1 as Path of p1,p2 by A1,BORSUK_2:def 1;
    reconsider P2=h2 as Path of p2,p3 by A1,BORSUK_2:def 1;
    ex P0 being Path of p1,p3 st P0 is continuous & P0.0=p1 & P0.1=p3 &
   for t being Point of I[01], t' being Real st t = t' holds
    ( 0 <= t' & t' <= 1/2 implies P0.t = P1.(2*t') ) &
    ( 1/2 <= t' & t' <= 1 implies P0.t = P2.(2*t'-1) )
  proof ::This proof is almost a copy of BORSUK_2:def 4(proof of Existence)
   set e1 = P[01](0, 1/2, (#)(0,1), (0,1)(#));
   set e2 = P[01](1/2, 1, (#)(0,1), (0,1)(#));
   set E1 = P1 * e1;
   set E2 = P2 * e2;
   set f = E1 +* E2;
A3:dom e1 = the carrier of Closed-Interval-TSpace(0,1/2) by FUNCT_2:def 1
         .= [.0,1/2.] by TOPMETR:25;
A4:dom e2 = the carrier of Closed-Interval-TSpace(1/2,1) by FUNCT_2:def 1
         .= [.1/2,1 qua Real.] by TOPMETR:25;
A5: dom P1 = the carrier of I[01] & dom P2 = the carrier of I[01]
       by FUNCT_2:def 1;
then A6:rng e1 c= dom P1 by TOPMETR:27;
     rng e2 c= the carrier of Closed-Interval-TSpace(0,1);
then A7: dom E2 = dom e2 by A5,RELAT_1:46,TOPMETR:27;
A8: dom f = dom E1 \/ dom E2 by FUNCT_4:def 1
        .= [.0,1/2.] \/ [.1/2,1 qua Real.] by A3,A4,A6,A7,RELAT_1:46
        .= the carrier of I[01] by BORSUK_1:83,TREAL_1:2;
A9: for t' being Real st 0 <= t' & t' <= 1/2 holds E1.t' = P1.(2*t')
   proof
    let t' be Real such that
A10:0 <= t' & t' <= 1/2;
      dom e1 = the carrier of Closed-Interval-TSpace(0,1/2) by FUNCT_2:def 1;
     then dom e1 = [.0, 1/2.] by TOPMETR:25
          .= {r : 0 <= r & r <= 1/2 } by RCOMP_1:def 1;
then A11:  t' in dom e1 by A10;
    then reconsider s = t' as Point of Closed-Interval-TSpace (0, 1/2)
      by FUNCT_2:def 1;
    reconsider r1 = (#)(0,1), r2 = (0,1)(#) as Real
      by BORSUK_1:def 17,def 18,TREAL_1:8;
      e1.s = ((r2 - r1)/(1/2 - 0))*t' + ((1/2)*r1 - 0 * r2)/(1/2 - 0)
      by TREAL_1:14
        .= 2*t' by BORSUK_1:def 17,def 18,TREAL_1:8;
    hence thesis by A11,FUNCT_1:23;
   end;
     not 0 in { r : 1/2 <= r & r <= 1 }
   proof
    assume 0 in { r : 1/2 <= r & r <= 1 };
     then ex rr being Real st rr = 0 & 1/2 <= rr & rr <= 1;
    hence thesis;
   end;
   then not 0 in dom E2 by A4,A7,RCOMP_1:def 1;
then A12:f.0 = E1.0 by FUNCT_4:12
      .= P1.(2*0) by A9
      .= p1 by A1;
     rng E1 c= rng P1 & rng E2 c= rng P2 by RELAT_1:45;
then A13:rng E1 c= the carrier of ((TOP-REAL n)|P) &
rng E2 c= the carrier of ((TOP-REAL n)|P) by XBOOLE_1:1;
A14: rng f c= rng E1 \/ rng E2 by FUNCT_4:18;
     rng E1 \/ rng E2 c= (the carrier of ((TOP-REAL n)|P))
   \/ the carrier of ((TOP-REAL n)|P)
      by A13,XBOOLE_1:13;
   then rng f c= the carrier of ((TOP-REAL n)|P) by A14,XBOOLE_1:1;
   then f is Function of the carrier of I[01], the carrier of ((TOP-REAL n)|P)
      by A8,FUNCT_2:def 1,RELSET_1:11;
   then reconsider f as map of I[01], ((TOP-REAL n)|P) ;
   reconsider T1 = Closed-Interval-TSpace (0, 1/2),
              T2 = Closed-Interval-TSpace (1/2, 1) as SubSpace of I[01]
                by TOPMETR:27,TREAL_1:6;
A15:e1 is continuous by TREAL_1:15;
A16:e2 is continuous by TREAL_1:15;
     E1 is Function of the carrier of Closed-Interval-TSpace (0, 1/2),
            the carrier of ((TOP-REAL n)|P) by FUNCT_2:19,TOPMETR:27;
   then reconsider ff = E1 as map of T1, ((TOP-REAL n)|P) ;
A17:E2 is Function of the carrier of Closed-Interval-TSpace (1/2, 1),
            the carrier of ((TOP-REAL n)|P) by FUNCT_2:19,TOPMETR:27;
   then reconsider gg = E2 as map of T2, ((TOP-REAL n)|P) ;
     1/2 in { r : 0 <= r & r <= 1 };
   then reconsider pol = 1/2 as Point of I[01] by BORSUK_1:83,RCOMP_1:def 1;
A18:ff is continuous & gg is continuous by A1,A15,A16,TOPMETR:27,TOPS_2:58;
A19:[#] T1 = the carrier of T1 by PRE_TOPC:12
       .= [.0,1/2.] by TOPMETR:25;
A20:[#] T2 = the carrier of T2 by PRE_TOPC:12
       .= [.1/2,1 qua Real.] by TOPMETR:25;
then A21: ([#] T1) \/ ([#] T2) = [.0,1 qua Real.] by A19,TREAL_1:2
                  .= [#] I[01] by BORSUK_1:83,PRE_TOPC:12;
A22: for t' being Real st 1/2 <= t' & t' <= 1 holds E2.t' = P2.(2*t'-1)
   proof
    let t' be Real such that
A23:      1/2 <= t' & t' <= 1;
      dom e2 = the carrier of Closed-Interval-TSpace(1/2,1) by FUNCT_2:def 1;
     then dom e2 = [.1/2,1 qua Real.] by TOPMETR:25
          .= {r : 1/2 <= r & r <= 1 } by RCOMP_1:def 1;
then A24:  t' in dom e2 by A23;
    then reconsider s = t' as Point of Closed-Interval-TSpace (1/2,1)
      by FUNCT_2:def 1;
    reconsider r1 = (#)(0,1), r2 = (0,1)(#) as Real
      by BORSUK_1:def 17,def 18,TREAL_1:8;
      e2.s = ((r2 - r1)/(1 - 1/2))*t' + (1 * r1 - (1/2)*r2)/(1 - 1/2)
      by TREAL_1:14
        .= 2*t' + (- 1) by BORSUK_1:def 17,def 18,TREAL_1:8
        .= 2*t' - 1 by XCMPLX_0:def 8;
    hence thesis by A24,FUNCT_1:23;
   end;
A25:ff.(1/2) = P2.(2*(1/2)-1) by A1,A9
           .= gg.pol by A22;
A26:([#] T1) /\ ([#] T2) = {pol} by A19,A20,TOPMETR2:1;
     R^1 is_T2 by PCOMPS_1:38,TOPMETR:def 7;
   then T1 is compact & T2 is compact & I[01] is_T2 & ff.pol = gg.pol
     by A25,HEINE:11,TOPMETR:3;
  then consider h being map of I[01], ((TOP-REAL n)|P) such that
A27:   h = ff+*gg & h is continuous by A18,A21,A26,TOPMETR2:4;
     1 in { r : 1/2 <= r & r <= 1 };
   then 1 in dom E2 by A4,A7,RCOMP_1:def 1;
 then A28:  f.1 = E2.1 by FUNCT_4:14
      .= P2.(2*1-1) by A22
      .= p3 by A1;
   then reconsider f as Path of p1, p3 by A12,A27,BORSUK_2:def 1;
  for t being Point of I[01], t' being Real st t = t' holds
     ( 0 <= t' & t' <= 1/2 implies f.t = P1.(2*t') ) &
     ( 1/2 <= t' & t' <= 1 implies f.t = P2.(2*t'-1) )
   proof
    let t be Point of I[01], t' be Real;
    assume A29: t = t';
    thus 0 <= t' & t' <= 1/2 implies f.t = P1.(2*t')
    proof
     assume A30: 0 <= t' & t' <= 1/2;
     then t' in { r : 0 <= r & r <= 1/2 };
then A31:  t' in [.0,1/2.] by RCOMP_1:def 1;
     per cases;
     suppose A32: t' <> 1/2;
       not t' in dom E2
     proof
      assume t' in dom E2;
      then t' in [.0,1/2.] /\ [.1/2,1 qua Real.] by A4,A7,A31,XBOOLE_0:def 3;
      then t' in {1/2} by TOPMETR2:1;
      hence thesis by A32,TARSKI:def 1;
     end;
     then f.t = E1.t by A29,FUNCT_4:12
        .= P1.(2*t') by A9,A29,A30;
     hence thesis;
     suppose A33: t' = 1/2;
       1/2 in { r : 1/2 <= r & r <= 1 };
     then 1/2 in [.1/2, 1 qua Real.] by RCOMP_1:def 1;
     then 1/2 in the carrier of Closed-Interval-TSpace(1/2,1)
         by TOPMETR:25;
     then t in dom E2 by A17,A29,A33,FUNCT_2:def 1;
     then f.t = E2.(1/2) by A29,A33,FUNCT_4:14
        .= P1.(2*t') by A9,A25,A33;
     hence thesis;
    end;
    thus 1/2 <= t' & t' <= 1 implies f.t = P2.(2*t'-1)
    proof
     assume A34: 1/2 <= t' & t' <= 1;
     then t' in { r : 1/2 <= r & r <= 1 };
      then t' in [.1/2,1 qua Real.] by RCOMP_1:def 1;
      then f.t = E2.t by A4,A7,A29,FUNCT_4:14
        .= P2.(2*t'-1) by A22,A29,A34;
     hence thesis;
    end;
   end;
   hence thesis by A12,A27,A28;
  end;
  hence thesis;
 end;

theorem Th44: for P being Subset of TOP-REAL n,
 w1,w2,w3 being Point of TOP-REAL n
   st w1 in P & w2 in P & w3 in P & LSeg(w1,w2) c= P & LSeg(w2,w3) c= P
   ex h being map of I[01],(TOP-REAL n)|P st h is continuous
                        & w1=h.0 & w3=h.1
proof let P be Subset of TOP-REAL n,
 w1,w2,w3 be Point of TOP-REAL n;
  assume A1:w1 in P & w2 in P & w3 in P & LSeg(w1,w2) c= P
         & LSeg(w2,w3) c= P;
   then reconsider Y = P as non empty Subset of TOP-REAL n;
   per cases;
   suppose A2:w1<>w2;
    then LSeg(w1,w2) is_an_arc_of w1,w2 by TOPREAL1:15;
    then consider f being map of I[01], (TOP-REAL n)|LSeg(w1,w2)
            such that A3: f is_homeomorphism & f.0 = w1 & f.1 = w2 by TOPREAL1:
def 2;
    A4:f is continuous by A3,TOPS_2:def 5;
    A5: rng f = [#]((TOP-REAL n)|LSeg(w1,w2)) by A3,TOPS_2:def 5;
    then A6: rng f c= P by A1,PRE_TOPC:def 10;
    then [#]((TOP-REAL n)|LSeg(w1,w2)) c= [#]((TOP-REAL n)|P)
              by A5,PRE_TOPC:def 10;
    then [#]((TOP-REAL n)|LSeg(w1,w2))
          c=the carrier of ((TOP-REAL n)|P) by PRE_TOPC:12;
    then the carrier of ((TOP-REAL n)|LSeg(w1,w2))
       c= the carrier of ((TOP-REAL n)|P) by PRE_TOPC:12;
    then A7:(TOP-REAL n)|LSeg(w1,w2)
           is SubSpace of (TOP-REAL n)|P by TOPMETR:4;
      dom f= ([.0 qua Real ,1 qua Real.]) by BORSUK_1:83,FUNCT_2:def 1;
    then reconsider g=f as Function of ([.0 qua Real,1 qua Real.]),P
                                   by A6,FUNCT_2:4;
       the carrier of (TOP-REAL n)|P = P by JORDAN1:1;
    then reconsider gt=g as map of I[01],(TOP-REAL n)|Y
      by BORSUK_1:83;
      now per cases;
    suppose w2<>w3;
    then LSeg(w2,w3) is_an_arc_of w2,w3 by TOPREAL1:15;
    then consider f2 being map of I[01], (TOP-REAL n)|LSeg(w2,w3)
            such that A8: f2 is_homeomorphism & f2.0 = w2 & f2.1 = w3 by
TOPREAL1:def 2;
    A9:f2 is continuous by A8,TOPS_2:def 5;
    A10: rng f2 = [#]((TOP-REAL n)|LSeg(w2,w3)) by A8,TOPS_2:def 5;
    then A11: rng f2 c= P by A1,PRE_TOPC:def 10;
    then [#]((TOP-REAL n)|LSeg(w2,w3)) c= [#]((TOP-REAL n)|P)
              by A10,PRE_TOPC:def 10;
    then [#]((TOP-REAL n)|LSeg(w2,w3))
          c=the carrier of ((TOP-REAL n)|P) by PRE_TOPC:12;
    then the carrier of ((TOP-REAL n)|LSeg(w2,w3))
       c= the carrier of ((TOP-REAL n)|P) by PRE_TOPC:12;
    then A12:(TOP-REAL n)|LSeg(w2,w3)
           is SubSpace of (TOP-REAL n)|P by TOPMETR:4;
      dom f2=[.0 qua Real,1 qua Real.] by BORSUK_1:83,FUNCT_2:def 1;
    then reconsider g2=f2 as Function of ([.0 qua Real,1 qua Real.]),P
                              by A11,FUNCT_2:4;
       the carrier of (TOP-REAL n)|P = P by JORDAN1:1;
    then reconsider gt2=g2 as map of I[01],(TOP-REAL n)|Y
      by BORSUK_1:83;
    A13:gt2 is continuous by A9,A12,TOPMETR:7;
      [#]((TOP-REAL n)|P)=P by PRE_TOPC:def 10;
    then reconsider w1'=w1,w2'=w2,w3'=w3 as Point of (TOP-REAL n)|P by A1;
      gt is continuous & w1'=gt.0 & w2'=gt.1 by A3,A4,A7,TOPMETR:7;
   then ex h being map of I[01],(TOP-REAL n)|Y st h is continuous
       & w1'=h.0 & w3'=h.1 & rng h c= (rng gt) \/ (rng gt2)
                        by A8,A13,Th32;
  hence ex h being map of I[01],(TOP-REAL n)|P st h is continuous
                        & w1=h.0 & w3=h.1;
  suppose A14:w2=w3;
     then LSeg(w1,w3) is_an_arc_of w1,w3 by A2,TOPREAL1:15;
    then consider f being map of I[01], (TOP-REAL n)|LSeg(w1,w3)
            such that A15: f is_homeomorphism & f.0 = w1 & f.1 = w3 by TOPREAL1
:def 2;
    A16:f is continuous by A15,TOPS_2:def 5;
    A17: rng f = [#]((TOP-REAL n)|LSeg(w1,w3)) by A15,TOPS_2:def 5;
    then A18: rng f c= P by A1,A14,PRE_TOPC:def 10;
    then [#]((TOP-REAL n)|LSeg(w1,w3)) c= [#]((TOP-REAL n)|P)
              by A17,PRE_TOPC:def 10;
    then [#]((TOP-REAL n)|LSeg(w1,w3))
          c=the carrier of ((TOP-REAL n)|P) by PRE_TOPC:12;
    then the carrier of ((TOP-REAL n)|LSeg(w1,w3))
       c= the carrier of ((TOP-REAL n)|P) by PRE_TOPC:12;
    then A19:(TOP-REAL n)|LSeg(w1,w3)
           is SubSpace of (TOP-REAL n)|P by TOPMETR:4;
      dom f=[.0 qua Real ,1 qua Real.] by BORSUK_1:83,FUNCT_2:def 1;
    then reconsider g=f as Function of ([.0 qua Real,1 qua Real.]),P
                         by A18,FUNCT_2:4;
       the carrier of (TOP-REAL n)|P = P by JORDAN1:1;
    then reconsider gt=g as map of I[01],(TOP-REAL n)|Y
      by BORSUK_1:83;
      gt is continuous by A16,A19,TOPMETR:7;
  hence ex h being map of I[01],(TOP-REAL n)|P st h is continuous
                        & w1=h.0 & w3=h.1 by A15;
   end;
  hence ex h being map of I[01],(TOP-REAL n)|P st h is continuous
                        & w1=h.0 & w3=h.1;
 suppose A20:w1=w2;
     now per cases;
   case w2<>w3;
     then LSeg(w1,w3) is_an_arc_of w1,w3 by A20,TOPREAL1:15;
    then consider f being map of I[01], (TOP-REAL n)|LSeg(w1,w3)
            such that A21: f is_homeomorphism & f.0 = w1 & f.1 = w3 by TOPREAL1
:def 2;
    A22:f is continuous by A21,TOPS_2:def 5;
    A23: rng f = [#]((TOP-REAL n)|LSeg(w1,w3)) by A21,TOPS_2:def 5;
    then A24: rng f c= P by A1,A20,PRE_TOPC:def 10;
    then [#]((TOP-REAL n)|LSeg(w1,w3)) c= [#]((TOP-REAL n)|P)
              by A23,PRE_TOPC:def 10;
    then [#]((TOP-REAL n)|LSeg(w1,w3))
          c=the carrier of ((TOP-REAL n)|P) by PRE_TOPC:12;
    then the carrier of ((TOP-REAL n)|LSeg(w1,w3))
       c= the carrier of ((TOP-REAL n)|P) by PRE_TOPC:12;
    then A25:(TOP-REAL n)|LSeg(w1,w3)
           is SubSpace of (TOP-REAL n)|P by TOPMETR:4;
      dom f=[.0 qua Real,1 qua Real.] by BORSUK_1:83,FUNCT_2:def 1;
    then reconsider g=f as Function of [.0 qua Real,1 qua Real.],P
                                 by A24,FUNCT_2:4;
       the carrier of (TOP-REAL n)|P = P by JORDAN1:1;
    then reconsider gt=g as map of I[01],(TOP-REAL n)|Y
      by BORSUK_1:83;
      gt is continuous by A22,A25,TOPMETR:7;
  hence ex h being map of I[01],(TOP-REAL n)|P st h is continuous
                        & w1=h.0 & w3=h.1 by A21;
   case A26:w2=w3;
      [#]((TOP-REAL n)|P)=P by PRE_TOPC:def 10;
    then reconsider w1'=w1,w3'=w3 as Point of (TOP-REAL n)|P by A1;
      ex f be map of I[01], (TOP-REAL n)|Y st
    f is continuous & f.0 = w1' & f.1 = w3' by A20,A26,BORSUK_2:4;
    hence ex h being map of I[01],(TOP-REAL n)|P st h is continuous
                        & w1=h.0 & w3=h.1;
   end;
  hence thesis;
end;

theorem Th45: for P being Subset of TOP-REAL n,
 w1,w2,w3,w4 being Point of TOP-REAL n
   st w1 in P & w2 in P & w3 in P & w4 in P &
   LSeg(w1,w2) c= P & LSeg(w2,w3) c= P &
   LSeg(w3,w4) c= P
   ex h being map of I[01],(TOP-REAL n)|P st h is continuous
                        & w1=h.0 & w4=h.1
proof let P be Subset of TOP-REAL n,
 w1,w2,w3,w4 be Point of TOP-REAL n;
 assume A1: w1 in P & w2 in P & w3 in P & w4 in P &
   LSeg(w1,w2) c= P & LSeg(w2,w3) c= P &
   LSeg(w3,w4) c= P;
   then consider h2 being map of I[01],(TOP-REAL n)|P such that
   A2: h2 is continuous & w1=h2.0 & w3=h2.1 by Th44;
   reconsider Y = P as non empty Subset of TOP-REAL n by A1;
   per cases;
   suppose w3<>w4;
    then LSeg(w3,w4) is_an_arc_of w3,w4 by TOPREAL1:15;
    then consider f being map of I[01], (TOP-REAL n)|LSeg(w3,w4)
            such that A3: f is_homeomorphism & f.0 = w3 & f.1 = w4 by TOPREAL1:
def 2;
    A4:f is continuous by A3,TOPS_2:def 5;
    A5: rng f = [#]((TOP-REAL n)|LSeg(w3,w4)) by A3,TOPS_2:def 5;
    then A6: rng f c= P by A1,PRE_TOPC:def 10;
    then [#]((TOP-REAL n)|LSeg(w3,w4)) c= [#]((TOP-REAL n)|P)
              by A5,PRE_TOPC:def 10;
    then [#]((TOP-REAL n)|LSeg(w3,w4))
          c=the carrier of ((TOP-REAL n)|P) by PRE_TOPC:12;
    then the carrier of ((TOP-REAL n)|LSeg(w3,w4))
       c= the carrier of ((TOP-REAL n)|P) by PRE_TOPC:12;
    then A7:(TOP-REAL n)|LSeg(w3,w4)
           is SubSpace of (TOP-REAL n)|P by TOPMETR:4;
      dom f=[.0 qua Real,1 qua Real.] by BORSUK_1:83,FUNCT_2:def 1;
    then reconsider g=f as Function of [.0 qua Real,1 qua Real.],P
                             by A6,FUNCT_2:4;
       the carrier of (TOP-REAL n)|P = P by JORDAN1:1;
    then reconsider gt=g as map of I[01],(TOP-REAL n)|Y
      by BORSUK_1:83;
    A8:gt is continuous by A4,A7,TOPMETR:7;
      [#]((TOP-REAL n)|P)=P by PRE_TOPC:def 10;
    then reconsider w1'=w1,w3'=w3,w4'=w4 as Point of (TOP-REAL n)|P by A1;
      h2 is continuous & w1'=h2.0 & w3'=h2.1 by A2;
   then ex h being map of I[01],(TOP-REAL n)|Y st h is continuous
          & w1'=h.0 & w4'=h.1 & rng h c= (rng h2) \/ (rng gt) by A3,A8,Th32;
   hence thesis;
   suppose w3=w4;
   hence thesis by A2;
end;

theorem Th46: for P being Subset of TOP-REAL n,
 w1,w2,w3,w4,w5,w6,w7 being Point of TOP-REAL n
   st w1 in P & w2 in P & w3 in P & w4 in P & w5 in P & w6 in P & w7 in P &
   LSeg(w1,w2) c= P & LSeg(w2,w3) c= P &
   LSeg(w3,w4) c= P & LSeg(w4,w5) c= P & LSeg(w5,w6) c= P
   & LSeg(w6,w7) c= P
   ex h being map of I[01],(TOP-REAL n)|P st h is continuous
                        & w1=h.0 & w7=h.1
proof let P be Subset of TOP-REAL n,
 w1,w2,w3,w4,w5,w6,w7 be Point of TOP-REAL n;
 assume A1: w1 in P & w2 in P & w3 in P &
 w4 in P & w5 in P & w6 in P & w7 in P &
   LSeg(w1,w2) c= P & LSeg(w2,w3) c= P &
   LSeg(w3,w4) c= P & LSeg(w4,w5) c= P & LSeg(w5,w6) c= P
   & LSeg(w6,w7) c= P;
then A2: ex h2 being map of I[01],(TOP-REAL n)|P st
    h2 is continuous & w1=h2.0 & w4=h2.1 by Th45;
     ex h4 being map of I[01],(TOP-REAL n)|P st
    h4 is continuous & w4=h4.0 & w7=h4.1 by A1,Th45;
  hence thesis by A1,A2,Th43;
end;

reserve s2 for Real;

theorem Th47: for w1,w2 being Point of TOP-REAL n st
not (ex r being Real st w1=r*w2 or w2=r*w1)
holds not (0.REAL n) in LSeg(w1,w2)
proof let w1,w2 be Point of TOP-REAL n;
 assume A1:not (ex r being Real st w1=r*w2 or w2=r*w1);
  assume (0.REAL n) in LSeg(w1,w2);
   then (0.REAL n) in { (1-s)*w1 + s*w2 : 0 <= s & s <= 1 } by TOPREAL1:def 4;
    then consider s being Real such that
    A2:(0.REAL n)=(1-s)*w1 + s*w2 &( 0 <= s & s <= 1 );
      (0.REAL n)-s*w2=(1-s)*w1 by A2,EUCLID:52;
    then -s*w2=(1-s)*w1 by Th13;
    then A3:(-s)*w2=(1-s)*w1 by EUCLID:44;
    per cases;
    suppose A4: -s<>0;
        ((-s)"*(-s))*w2=(-s)"*((1-s)*w1) by A3,EUCLID:34;
      then ((-s)"*(-s))*w2=((-s)"*(1-s))*w1 by EUCLID:34;
      then (1)*w2=((-s)"*(1-s))*w1 by A4,XCMPLX_0:def 7;
      then w2=((-s)"*(1-s))*w1 by EUCLID:33;
     hence contradiction by A1;
    suppose -s=0; then 1+-s=1+0; then 1-s=1 by XCMPLX_0:def 8;
      then (-s)*w2=w1 by A3,EUCLID:33;
     hence contradiction by A1;
end;

theorem Th48: for w1,w2 being Point of TOP-REAL n,P being Subset of
 TopSpaceMetr(Euclid n) st P=LSeg(w1,w2)&
 not (0.REAL n) in LSeg(w1,w2) holds ex w0 being Point of TOP-REAL n st
 w0 in LSeg(w1,w2) & |.w0.|>0 & |.w0.|=(dist_min(P)).(0.REAL n)
 proof let w1,w2 be Point of TOP-REAL n,P be Subset of
 TopSpaceMetr(Euclid n);
 assume A1:P=LSeg(w1,w2)& not 0.REAL n in LSeg(w1,w2);
  set M=Euclid n;
  reconsider P0=P as Subset of TopSpaceMetr(M);
  reconsider P1=P as Subset of TOP-REAL n by EUCLID:def 8;
  reconsider o=0.REAL n as Point of M by TOPREAL3:13;
  A2:TOP-REAL n=TopSpaceMetr(M) by EUCLID:def 8;
  reconsider o2=0.REAL n as Point of TopSpaceMetr(M) by EUCLID:def 8;
  A3:P0 is compact by A1,A2,SPPOL_1:28;
  reconsider Q={0.REAL n} as Subset of TopSpaceMetr(M) by EUCLID:def 8;
    0.REAL n is Point of TopSpaceMetr(M) by EUCLID:def 8;
  then Q is compact by BORSUK_1:41;
   then consider x1,x2 being Point of M such that
   A4:x1 in P0 & x2 in Q &
   dist(x1,x2) = min_dist_min(P0,Q) by A1,A3,WEIERSTR:36;
A5:  for x being set holds x in (dist_min(P0)).:(Q) iff x=(dist_min(P0)).o
   proof let x be set;
     hereby assume x in (dist_min(P0)).:(Q);
      then consider y being set such that
      A6: y in dom(dist_min(P0)) &
      y in Q & x=(dist_min(P0)).y by FUNCT_1:def 12;
    thus x=(dist_min(P0)).o by A6,TARSKI:def 1;
    end;
    assume A7:x=(dist_min(P0)).o;
      A8:o in Q by TARSKI:def 1;
        o2 in the carrier of TopSpaceMetr(M);
      then o in dom (dist_min(P0)) by FUNCT_2:def 1;
    hence thesis by A7,A8,FUNCT_1:def 12;
   end;
  A9:[#] ((dist_min(P0)).:(Q))=(dist_min(P0)).:(Q) by WEIERSTR:def 3;
  A10:  (dist_min(P0)).:(Q)={(dist_min(P0)).o} by A5,TARSKI:def 1;
   lower_bound([#] ((dist_min(P0)).:(Q)))=lower_bound((dist_min(P0)).:(Q))
                                 by WEIERSTR:def 5;
  then lower_bound((dist_min(P0)).:(Q))=(dist_min(P0)).o
                       by A9,A10,SEQ_4:22;
  then A11:dist(x1,x2)=(dist_min(P)).(0.REAL n) by A4,WEIERSTR:def 9;
  A12:x2=0.REAL n by A4,TARSKI:def 1;
    x1 in P1 by A4;
  then reconsider w01=x1 as Point of TOP-REAL n;
  A13: |.w01.|=|.w01-0.REAL n.| by Th13 .=dist(x1,x2) by A12,JGRAPH_1:45;
  A14: |.w01.|>=0 by TOPRNS_1:26;
    |.w01.| <> 0 by A1,A4,TOPRNS_1:25;
 hence thesis by A1,A4,A11,A13,A14;
 end;

theorem Th49: for a being Real,
Q being Subset of TOP-REAL n,
w1,w4 being Point of TOP-REAL n
 st Q={q : (|.q.|) > a } & w1 in Q & w4 in Q &
 not (ex r being Real st w1=r*w4 or w4=r*w1)
 holds ex w2,w3 being Point of TOP-REAL n st w2 in Q & w3 in Q &
 LSeg(w1,w2) c=Q & LSeg(w2,w3) c= Q & LSeg(w3,w4) c= Q
 proof let a be Real,
   Q be Subset of TOP-REAL n,
   w1,w4 be Point of TOP-REAL n;
   assume A1: Q={q : (|.q.|) > a } & w1 in Q & w4 in Q &
    not (ex r being Real st w1=r*w4 or w4=r*w1);
    then A2:not (0.REAL n) in LSeg(w1,w4) by Th47;
    reconsider P=LSeg(w1,w4) as Subset of TopSpaceMetr(Euclid n)
                                    by EUCLID:def 8;
    consider w0 being Point of TOP-REAL n such that
     A3:w0 in LSeg(w1,w4) & |.w0.|>0 & |.w0.|=(dist_min(P)).(0.REAL n)
                    by A2,Th48;
     set l'=(a+1)/|.w0.|;
     set w2= l'*w1,w3=l'*w4;
     A4:LSeg(w2,w3)c=Q
     proof let x be set;assume x in LSeg(w2,w3);
       then x in { (1-r)*w2 + r*w3 : 0 <= r & r <= 1 } by TOPREAL1:def 4;
       then consider r such that
       A5: x=(1-r)*w2 + r*w3 &( 0 <= r & r <= 1);
       reconsider w5=(1-r)*w1 + r*w4 as Point of TOP-REAL n;
       reconsider w5'=w5 as Point of Euclid n by TOPREAL3:13;
         w5 in {(1-r1)*w1 + r1*w4:0 <= r1 & r1 <= 1} by A5;
       then A6:w5 in LSeg(w1,w4) by TOPREAL1:def 4;
       A7:w5' in the carrier of TOP-REAL n;
       reconsider o=0.REAL n as Point of Euclid n by TOPREAL3:13;
       reconsider P=LSeg(w1,w4) as Subset of TopSpaceMetr(Euclid n)
                                              by EUCLID:def 8;
    (dist(o)).:(P) c= REAL
  proof let x be set;assume x in ((dist(o)).:(P));
    then consider z being set such that
    A8: z in dom (dist(o)) & z in P & x=(dist(o)).z by FUNCT_1:def 12;
    reconsider z2=z as Point of Euclid n by A8,TOPREAL3:13;
      (dist(o)).z2 =dist(z2,o) by WEIERSTR:def 6;
   hence x in REAL by A8;
  end;
  then reconsider F=((dist(o)).:(P)) as Subset of REAL;
  A9:w5' in the carrier of TopSpaceMetr(Euclid n) by A7,EUCLID:def 8;
    for r being real number st r in ((dist(o)).:(P)) holds 0<=r
   proof let r be real number;assume r in ((dist(o)).:(P));
     then consider x being set
     such that A10:x in dom (dist(o)) &
     x in P & r=(dist(o)).x by FUNCT_1:def 12;
     reconsider w0=x as Point of Euclid n by A10,TOPREAL3:13;
       r=dist(w0,o) by A10,WEIERSTR:def 6;
    hence 0<=r by METRIC_1:5;
   end;
  then A11:F is bounded_below by SEQ_4:def 2;
    w5' in dom (dist(o)) & dist(w5',o)=(dist(o)).w5'
   by A9,FUNCT_2:def 1,WEIERSTR:def 6;
  then dist(w5',o) in ((dist(o)).:(P)) by A6,FUNCT_1:def 12;
  then lower_bound F <=dist(w5',o) by A11,SEQ_4:def 5;
       then dist(w5',o)>= lower_bound([#]((dist(o)).:(P))) by WEIERSTR:def 3;
       then dist(w5',o)>= lower_bound((dist(o)).:(P)) by WEIERSTR:def 5;
       then dist(w5',o)>=|.w0.| by A3,WEIERSTR:def 8;
       then |.w5-0.REAL n.| >=|.w0.| by JGRAPH_1:45;
       then A12: |.w5.| >=|.w0.| by Th13;
       A13:abs(a+1)>=0 by ABSVALUE:5;
A14:   |.w0.| >= 0 by TOPRNS_1:26;
       A15:abs l' = abs(a+1)/abs |.w0.| by ABSVALUE:16
                .=abs(a+1)/|.w0.| by A14,ABSVALUE:def 1;
         |.w5.|/|.w0.|>=1 by A3,A12,REAL_2:143;
       then abs(a+1)*(|.w5.|/|.w0.|)>=abs(a+1)*1 by A13,REAL_2:197;
       then abs(a+1)*(|.w0.|"*|.w5.|)>=abs(a+1) by XCMPLX_0:def 9;
       then abs(a+1)*|.w0.|"*|.w5.|>=abs(a+1) by XCMPLX_1:4;
       then A16:abs(a+1)/|.w0.|*|.w5.|>=abs(a+1) by XCMPLX_0:def 9;
       A17:a+1>a by REAL_1:69;
         abs(a+1)>=a+1 by ABSVALUE:11;
       then abs(a+1)>a by A17,AXIOMS:22;
       then abs(a+1)/|.w0.|*|.w5.|>a by A16,AXIOMS:22;
       then |.l'*((1-r)*w1 + r*w4).|>a by A15,TOPRNS_1:8;
       then |.l'*((1-r)*w1) + l'*(r*w4).|>a by EUCLID:36;
       then |.l'*((1-r)*w1) + (l'*r)*w4.|>a by EUCLID:34;
       then |.(l'*(1-r))*w1 + (l'*r)*w4.|>a by EUCLID:34;
       then |.((1-r)*l')*w1 + r*(l'*w4).|>a by EUCLID:34;
       then |.(1-r)*w2 + r*w3.|>a by EUCLID:34;
      hence x in Q by A1,A5;
     end;
     A18:w2 in LSeg(w2,w3) & w3 in LSeg(w2,w3) by TOPREAL1:6;
     then A19:w2 in Q by A4;
     A20:w3 in Q by A4,A18;
     A21: LSeg(w1,w2) c=Q
     proof let x be set;assume x in LSeg(w1,w2);
       then x in { (1-r)*w1 + r*w2 : 0 <= r & r <= 1 } by TOPREAL1:def 4;
       then consider r such that
       A22: x=(1-r)*w1 + r*w2 &( 0 <= r & r <= 1);
        now per cases;
      case A23:a>=0; a+1>a by REAL_1:69;
          then (a+1)/|.w0.|>0 by A3,A23,REAL_2:127;
         then A24: r*l'>=0 by A22,REAL_2:121;
       reconsider w5=(1-0)*w1 + 0 * w4 as Point of TOP-REAL n;
       A25:(1-0)*w1+0 * w4=(1-0)*w1+0.REAL n by EUCLID:33
        .=(1-0)*w1 by EUCLID:31
       .=w1 by EUCLID:33;
       reconsider w5'=w5 as Point of Euclid n by TOPREAL3:13;
         w5 in {(1-r1)*w1 + r1*w4:0 <= r1 & r1 <= 1};
       then A26:w5 in LSeg(w1,w4) by TOPREAL1:def 4;
       A27:w5' in the carrier of TOP-REAL n;
       reconsider o=0.REAL n as Point of Euclid n by TOPREAL3:13;
       reconsider P=LSeg(w1,w4) as Subset of TopSpaceMetr(Euclid n)
                                              by EUCLID:def 8;
    ((dist(o)).:(P)) c= REAL
  proof let x be set;assume x in ((dist(o)).:(P));
    then consider z being set such that
    A28: z in dom (dist(o)) & z in P & x=(dist(o)).z by FUNCT_1:def 12;
    reconsider z2=z as Point of Euclid n by A28,TOPREAL3:13;
      (dist(o)).z2 =dist(z2,o) by WEIERSTR:def 6;
   hence x in REAL by A28;
  end;
  then reconsider F=((dist(o)).:(P)) as Subset of REAL;
  A29:w5' in the carrier of TopSpaceMetr(Euclid n) by A27,EUCLID:def 8;
    for r being real number st r in ((dist(o)).:(P)) holds 0<=r
   proof let r be real number;assume r in ((dist(o)).:(P));
     then consider x being set
     such that A30:x in dom (dist(o)) &
     x in P & r=(dist(o)).x by FUNCT_1:def 12;
     reconsider w0=x as Point of Euclid n by A30,TOPREAL3:13;
       r=dist(w0,o) by A30,WEIERSTR:def 6;
    hence 0<=r by METRIC_1:5;
   end;
  then A31:F is bounded_below by SEQ_4:def 2;
    w5' in dom (dist(o)) & dist(w5',o)=(dist(o)).w5'
   by A29,FUNCT_2:def 1,WEIERSTR:def 6;
  then dist(w5',o) in ((dist(o)).:(P)) by A26,FUNCT_1:def 12;
  then lower_bound F <=dist(w5',o) by A31,SEQ_4:def 5;
       then dist(w5',o)>= lower_bound([#]((dist(o)).:(P))) by WEIERSTR:def 3;
       then dist(w5',o)>= lower_bound((dist(o)).:(P)) by WEIERSTR:def 5;
       then dist(w5',o)>=|.w0.| by A3,WEIERSTR:def 8;
       then |.w5-0.REAL n.| >=|.w0.| by JGRAPH_1:45;
then A32:   |.w5.| >=|.w0.| by Th13;
           r*l'*|.w0.|
          =r*(a+1)/|.w0.|*|.w0.| by XCMPLX_1:75
         .=r*(a+1) by A3,XCMPLX_1:88;
         then A33:r*l'*|.w1.|>= r*(a+1) by A24,A25,A32,AXIOMS:25;
         consider q1 being Point of TOP-REAL n such that
         A34:q1=w1 & |.q1.| > a by A1;
         consider q2 being Point of TOP-REAL n such that
         A35:q2=w2 & |.q2.| > a by A1,A19;
         A36:1-r>=0 by A22,SQUARE_1:12;
  A37:  a+r>=a+0 by A22,AXIOMS:24;
          now per cases;
        case 1-r>0;
         then A38:(1-r)*|.w1.|>(1-r)*a by A34,REAL_1:70;
           (1-r)+ r*l'>=0 by A24,A36,Th5;
         then abs((1-r)+ r*l')*|.w1.|=((1-r)+ r*l')*|.w1.| by ABSVALUE:def 1
         .= (1-r)*|.w1.|+r*l'*|.w1.| by XCMPLX_1:8;
         then A39:abs((1-r)+ r*l')*|.w1.|>r*(a+1)+(1-r)*a by A33,A38,REAL_1:67;
           r*(a+1)+(1-r)*a=r*a +r*1 +(1-r)*a by XCMPLX_1:8
         .=r*a +(1-r)*a+r*1 by XCMPLX_1:1
         .=(r+(1-r))*a+r*1 by XCMPLX_1:8
         .=1 * a+r*1 by XCMPLX_1:27
         .=a+r;
         then abs((1-r)+ r*l')*|.w1.|>a by A37,A39,AXIOMS:22;
         then |.((1-r)+ r*l')*w1.|>a by TOPRNS_1:8;
         then |.(1-r)*w1 + r*l'*w1.|>a by EUCLID:37;
       hence |.(1-r)*w1 + r*w2.|>a by EUCLID:34;
        case 1-r<=0; then 1-r+r<=0+r by AXIOMS:24;
         then 1<=0+r by XCMPLX_1:27;
         then r=1 by A22,AXIOMS:21;
         then (1-r)*w1+r*w2=0.REAL n +1 * w2 by EUCLID:33
         .=0.REAL n +w2 by EUCLID:33
         .=w2 by EUCLID:31;
       hence |.(1-r)*w1 + r*w2.|>a by A35;
        end;
       hence |.(1-r)*w1 + r*w2.|>a;
       case A40:a<0;
         |.(1-r)*w1 + r*w2.|>=0 by TOPRNS_1:26;
        hence |.(1-r)*w1 + r*w2.|>a by A40;
       end;
      hence x in Q by A1,A22;
     end;
       LSeg(w4,w3) c=Q
     proof let x be set;assume x in LSeg(w4,w3);
       then x in { (1-r)*w4 + r*w3 : 0 <= r & r <= 1 } by TOPREAL1:def 4;
       then consider r such that
       A41: x=(1-r)*w4 + r*w3 &( 0 <= r & r <= 1);
        now per cases;
      case A42:a>=0;
            a<a+1 by REAL_1:69;
          then (a+1)/|.w0.|>0 by A3,A42,REAL_2:127;
         then A43: r*l'>=0 by A41,REAL_2:121;
       reconsider w5=(1-0)*w4 + 0 * w1 as Point of TOP-REAL n;
       A44:(1-0)*w4+0 * w1=(1-0)*w4+0.REAL n by EUCLID:33
       .=(1-0)*w4 by EUCLID:31
       .=w4 by EUCLID:33;
       reconsider w5'=w5 as Point of Euclid n by TOPREAL3:13;
         w5 in {(1-r1)*w4 + r1 * w1:0 <= r1 & r1 <= 1};
       then A45:w5 in LSeg(w1,w4) by TOPREAL1:def 4;
       A46:w5' in the carrier of TOP-REAL n;
       reconsider o=0.REAL n as Point of Euclid n by TOPREAL3:13;
       reconsider P=LSeg(w4,w1) as Subset of TopSpaceMetr(Euclid n)
                                              by EUCLID:def 8;
    (dist(o)).:(P) c= REAL
  proof let x be set;assume x in ((dist(o)).:(P));
    then consider z being set such that
    A47: z in dom (dist(o)) & z in P & x=(dist(o)).z by FUNCT_1:def 12;
    reconsider z2=z as Point of Euclid n by A47,TOPREAL3:13;
      (dist(o)).z2 =dist(z2,o) by WEIERSTR:def 6;
   hence x in REAL by A47;
  end;
  then reconsider F=((dist(o)).:(P)) as Subset of REAL;
  A48:w5' in the carrier of TopSpaceMetr(Euclid n) by A46,EUCLID:def 8;
    for r being real number st r in ((dist(o)).:(P)) holds 0<=r
   proof let r be real number;assume r in ((dist(o)).:(P));
     then consider x being set
     such that A49:x in dom (dist(o)) &
     x in P & r=(dist(o)).x by FUNCT_1:def 12;
     reconsider w0=x as Point of Euclid n by A49,TOPREAL3:13;
       r=dist(w0,o) by A49,WEIERSTR:def 6;
    hence 0<=r by METRIC_1:5;
   end;
  then A50:F is bounded_below by SEQ_4:def 2;
    w5' in dom (dist(o)) & dist(w5',o)=(dist(o)).w5'
   by A48,FUNCT_2:def 1,WEIERSTR:def 6;
  then dist(w5',o) in ((dist(o)).:(P)) by A45,FUNCT_1:def 12;
  then lower_bound F <=dist(w5',o) by A50,SEQ_4:def 5;
       then dist(w5',o)>= lower_bound([#]((dist(o)).:(P))) by WEIERSTR:def 3;
       then dist(w5',o)>= lower_bound((dist(o)).:(P)) by WEIERSTR:def 5;
       then dist(w5',o)>=|.w0.| by A3,WEIERSTR:def 8;
       then |.w5-0.REAL n.| >=|.w0.| by JGRAPH_1:45;
then A51:    |.w5.| >=|.w0.| by Th13;
           r*l'*|.w0.|
         =r*(a+1)/|.w0.|*|.w0.| by XCMPLX_1:75
         .=r*(a+1) by A3,XCMPLX_1:88;
         then A52:r*l'*|.w4.|>= r*(a+1) by A43,A44,A51,AXIOMS:25;
         consider q1 being Point of TOP-REAL n such that
         A53:q1=w4 & |.q1.| > a by A1;
         A54:1-r>=0 by A41,SQUARE_1:12;
A55:    a+r>=a+0 by A41,AXIOMS:24;
          now per cases;
        case 1-r>0;
         then A56:(1-r)*|.w4.|>(1-r)*a by A53,REAL_1:70;
           (1-r)+ r*l'>=0 by A43,A54,Th5;
         then abs((1-r)+ r*l')*|.w4.|=((1-r)+ r*l')*|.w4.| by ABSVALUE:def 1
         .= (1-r)*|.w4.|+r*l'*|.w4.| by XCMPLX_1:8;
         then A57:abs((1-r)+r*l')*|.w4.|>r*(a+1)+(1-r)*a by A52,A56,REAL_1:67;
           r*(a+1)+(1-r)*a=r*a +r*1 +(1-r)*a by XCMPLX_1:8
         .=r*a +(1-r)*a+r*1 by XCMPLX_1:1
         .=(r+(1-r))*a+r*1 by XCMPLX_1:8
         .=1 * a+r*1 by XCMPLX_1:27
         .=a+r;
         then abs((1-r)+r*l')*|.w4.|>a by A55,A57,AXIOMS:22;
         then |.((1-r)+ r*l')*w4.|>a by TOPRNS_1:8;
         then |.(1-r)*w4 + r*l'*w4.|>a by EUCLID:37;
       hence |.(1-r)*w4 + r*w3.|>a by EUCLID:34;
        case 1-r<=0; then 1-r+r<=0+r by AXIOMS:24;
         then 1<=0+r by XCMPLX_1:27;
         then r=1 by A41,AXIOMS:21;
         then A58:(1-r)*w4+r*w3=0.REAL n +1 * w3 by EUCLID:33
         .=0.REAL n +w3 by EUCLID:33
         .=w3 by EUCLID:31;
         consider q3 being Point of TOP-REAL n such that
         A59:q3=w3 & |.q3.| > a by A1,A20;
       thus |.(1-r)*w4 + r*w3.|>a by A58,A59;
        end;
       hence |.(1-r)*w4 + r*w3.|>a;
       case A60:a<0;
         |.(1-r)*w4 + r*w3.|>=0 by TOPRNS_1:26;
        hence |.(1-r)*w4 + r*w3.|>a by A60;
       end;
      hence x in Q by A1,A41;
     end;
  hence thesis by A4,A18,A21;
 end;

theorem Th50: for a being Real,
Q being Subset of TOP-REAL n,
w1,w4 being Point of TOP-REAL n
 st Q=(REAL n)\ {q : (|.q.|) < a } & w1 in Q & w4 in Q &
 not (ex r being Real st w1=r*w4 or w4=r*w1)
 holds ex w2,w3 being Point of TOP-REAL n st w2 in Q & w3 in Q &
 LSeg(w1,w2) c=Q & LSeg(w2,w3) c= Q & LSeg(w3,w4) c= Q
 proof let a be Real,
   Q be Subset of TOP-REAL n,
   w1,w4 be Point of TOP-REAL n;
   assume A1: Q=(REAL n)\ {q : (|.q.|) < a } & w1 in Q & w4 in Q &
    not (ex r being Real st w1=r*w4 or w4=r*w1);
    then A2:not (0.REAL n) in LSeg(w1,w4) by Th47;
    reconsider P=LSeg(w1,w4) as Subset of TopSpaceMetr(Euclid n)
                                    by EUCLID:def 8;
    consider w0 being Point of TOP-REAL n such that
     A3:w0 in LSeg(w1,w4) & |.w0.|>0 & |.w0.|=(dist_min(P)).(0.REAL n)
                    by A2,Th48;
     set l'=a/|.w0.|;
     set w2= l'*w1,w3=l'*w4;
       A4:(REAL n)\ {q : (|.q.|) < a } = {q1 : (|.q1.|) >= a }
        proof thus (REAL n)\ {q : (|.q.|) < a } c= {q1 : (|.q1.|) >= a }
         proof let z be set;assume z in (REAL n)\ {q : (|.q.|) < a };
           then A5:z in REAL n & not z in {q : (|.q.|) < a } by XBOOLE_0:def 4;
           then reconsider q2=z as Point of TOP-REAL n by EUCLID:25;
              |.q2.| >= a by A5;
          hence z in {q1 : (|.q1.|) >= a };
         end;
          let z be set;assume z in {q1 : (|.q1.|) >= a };
            then consider q1 such that A6:z=q1 & (|.q1.|) >= a;
 A7:        q1 in the carrier of TOP-REAL n;
              for q st q=z holds (|.q.|) >= a by A6;
            then z in REAL n & not z in {q : (|.q.|) < a } by A6,A7,EUCLID:25;
           hence thesis by XBOOLE_0:def 4;
        end;
     A8:LSeg(w2,w3)c=Q
     proof let x be set;assume x in LSeg(w2,w3);
       then x in {(1-r)*w2 + r*w3 : 0 <= r & r <= 1 } by TOPREAL1:def 4;
       then consider r such that
       A9: x=(1-r)*w2 + r*w3 &( 0 <= r & r <= 1);
       reconsider w5=(1-r)*w1 + r*w4 as Point of TOP-REAL n;
       reconsider w5'=w5 as Point of Euclid n by TOPREAL3:13;
         w5 in {(1-r1)*w1 + r1*w4:0 <= r1 & r1 <= 1} by A9;
       then A10:w5 in LSeg(w1,w4) by TOPREAL1:def 4;
       A11:w5' in the carrier of TOP-REAL n;
       reconsider o=0.REAL n as Point of Euclid n by TOPREAL3:13;
       reconsider P=LSeg(w1,w4) as Subset of TopSpaceMetr(Euclid n)
                                              by EUCLID:def 8;
    ((dist(o)).:(P)) c= REAL
  proof let x be set;assume x in ((dist(o)).:(P));
    then consider z being set such that
    A12: z in dom (dist(o)) & z in P & x=(dist(o)).z by FUNCT_1:def 12;
    reconsider z2=z as Point of Euclid n by A12,TOPREAL3:13;
      (dist(o)).z2 =dist(z2,o) by WEIERSTR:def 6;
   hence x in REAL by A12;
  end;
  then reconsider F=(dist(o)).:(P) as Subset of REAL;
  A13:w5' in the carrier of TopSpaceMetr(Euclid n) by A11,EUCLID:def 8;
    for r being real number st r in ((dist(o)).:(P)) holds 0<=r
   proof let r be real number;assume r in ((dist(o)).:(P));
     then consider x being set
     such that A14:x in dom (dist(o)) &
     x in P & r=(dist(o)).x by FUNCT_1:def 12;
     reconsider w0=x as Point of Euclid n by A14,TOPREAL3:13;
       r=dist(w0,o) by A14,WEIERSTR:def 6;
    hence 0<=r by METRIC_1:5;
   end;
  then A15:F is bounded_below by SEQ_4:def 2;
    w5' in dom (dist(o)) & dist(w5',o)=(dist(o)).w5'
  by A13,FUNCT_2:def 1,WEIERSTR:def 6;
  then dist(w5',o) in ((dist(o)).:(P)) by A10,FUNCT_1:def 12;
  then lower_bound F <=dist(w5',o) by A15,SEQ_4:def 5;
       then dist(w5',o)>= lower_bound([#]((dist(o)).:(P))) by WEIERSTR:def 3;
       then dist(w5',o)>= lower_bound((dist(o)).:(P)) by WEIERSTR:def 5;
       then dist(w5',o)>=|.w0.| by A3,WEIERSTR:def 8;
       then |.w5-0.REAL n.| >=|.w0.| by JGRAPH_1:45;
       then A16: |.w5.| >=|.w0.| by Th13;
       A17:abs(a)>=0 by ABSVALUE:5;
A18:   |.w0.| >= 0 by TOPRNS_1:26;
       A19:abs l'=abs(a)/(abs|.w0.|) by ABSVALUE:16
                .=abs(a)/|.w0.| by A18,ABSVALUE:def 1;
         |.w5.|/|.w0.|>=1 by A3,A16,REAL_2:143;
       then abs(a)*(|.w5.|/|.w0.|)>=abs(a)*1 by A17,REAL_2:197;
       then abs(a)*(|.w5.|*|.w0.|")>=abs(a) by XCMPLX_0:def 9;
       then abs(a)*|.w0.|"*|.w5.|>=abs(a) by XCMPLX_1:4;
       then A20:abs(a)/|.w0.|*|.w5.|>=abs(a) by XCMPLX_0:def 9;
         abs(a)>=a by ABSVALUE:11;
       then abs(a)/|.w0.|*|.w5.|>=a by A20,AXIOMS:22;
       then |.l'*((1-r)*w1 + r*w4).|>=a by A19,TOPRNS_1:8;
       then |.l'*((1-r)*w1) + l'*(r*w4).|>=a by EUCLID:36;
       then |.l'*((1-r)*w1) + (l'*r)*w4.|>=a by EUCLID:34;
       then |.(l'*(1-r))*w1 + (l'*r)*w4.|>=a by EUCLID:34;
       then |.((1-r)*l')*w1 + r*(l'*w4).|>=a by EUCLID:34;
       then |.(1-r)*w2 + r*w3.|>=a by EUCLID:34;
      hence x in Q by A1,A4,A9;
     end;
     A21:w2 in LSeg(w2,w3) & w3 in LSeg(w2,w3) by TOPREAL1:6;
     A22: LSeg(w1,w2) c=Q
     proof let x be set;assume x in LSeg(w1,w2);
       then x in { (1-r)*w1 + r*w2 : 0 <= r & r <= 1 } by TOPREAL1:def 4;
       then consider r such that
       A23: x=(1-r)*w1 + r*w2 &( 0 <= r & r <= 1);
        now per cases;
      case a>0;
          then a/|.w0.|>0 by A3,REAL_2:127;
         then A24: r*l'>=0 by A23,REAL_2:121;
       reconsider w5=(1-0)*w1 + 0 * w4 as Point of TOP-REAL n;
       A25:(1-0)*w1+0 * w4=(1-0)*w1+0.REAL n
         by EUCLID:33.=(1-0)*w1 by EUCLID:31
       .=w1 by EUCLID:33;
       reconsider w5'=w5 as Point of Euclid n by TOPREAL3:13;
         w5 in {(1-r1)*w1 + r1 * w4:0 <= r1 & r1 <= 1};
       then A26:w5 in LSeg(w1,w4) by TOPREAL1:def 4;
       A27:w5' in the carrier of TOP-REAL n;
       reconsider o=0.REAL n as Point of Euclid n by TOPREAL3:13;
       reconsider P=LSeg(w1,w4) as Subset of TopSpaceMetr(Euclid n)
                                              by EUCLID:def 8;
    ((dist(o)).:(P)) c= REAL
  proof let x be set;assume x in ((dist(o)).:(P));
    then consider z being set such that
    A28: z in dom (dist(o)) & z in P & x=(dist(o)).z by FUNCT_1:def 12;
    reconsider z2=z as Point of Euclid n by A28,TOPREAL3:13;
      (dist(o)).z2 =dist(z2,o) by WEIERSTR:def 6;
   hence x in REAL by A28;
  end;
  then reconsider F=((dist(o)).:(P)) as Subset of REAL;
  A29:w5' in the carrier of TopSpaceMetr(Euclid n) by A27,EUCLID:def 8;
    for r being real number st r in ((dist(o)).:(P)) holds 0<=r
   proof let r be real number;assume r in ((dist(o)).:(P));
     then consider x being set
     such that A30:x in dom (dist(o)) &
     x in P & r=(dist(o)).x by FUNCT_1:def 12;
     reconsider w0=x as Point of Euclid n by A30,TOPREAL3:13;
       r=dist(w0,o) by A30,WEIERSTR:def 6;
    hence 0<=r by METRIC_1:5;
   end;
  then A31:F is bounded_below by SEQ_4:def 2;
    w5' in dom (dist(o)) & dist(w5',o)=(dist(o)).w5'
   by A29,FUNCT_2:def 1,WEIERSTR:def 6;
  then dist(w5',o) in ((dist(o)).:(P)) by A26,FUNCT_1:def 12;
  then lower_bound F <=dist(w5',o) by A31,SEQ_4:def 5;
       then dist(w5',o)>= lower_bound([#]((dist(o)).:(P))) by WEIERSTR:def 3;
       then dist(w5',o)>= lower_bound((dist(o)).:(P)) by WEIERSTR:def 5;
       then dist(w5',o)>=|.w0.| by A3,WEIERSTR:def 8;
       then |.w5-0.REAL n.| >=|.w0.| by JGRAPH_1:45;
  then A32: |.w5.| >=|.w0.| by Th13;
           r*l'*|.w0.|
          =r*a/|.w0.|*|.w0.| by XCMPLX_1:75
         .=r*a by A3,XCMPLX_1:88;
         then A33:r*l'*|.w1.|>= r*a by A24,A25,A32,AXIOMS:25;
         consider q1 being Point of TOP-REAL n such that
         A34:q1=w1 & |.q1.| >= a by A1,A4;
         A35:1-r>=0 by A23,SQUARE_1:12;
         then A36:(1-r)*|.w1.|>=(1-r)*a by A34,AXIOMS:25;
           (1-r)+ r*l'>=0 by A24,A35,Th5;
         then abs((1-r)+r*l')*|.w1.|=((1-r)+ r*l')*|.w1.| by ABSVALUE:def 1
         .= (1-r)*|.w1.|+r*l'*|.w1.| by XCMPLX_1:8;
         then A37:abs((1-r)+r*l')*|.w1.|>=r*a+(1-r)*a by A33,A36,REAL_1:55;
           r*a+(1-r)*a
         =(r+(1-r))*a by XCMPLX_1:8
         .=1 * a by XCMPLX_1:27
         .=a;
         then |.((1-r)+ r*l')*w1.|>=a by A37,TOPRNS_1:8;
         then |.(1-r)*w1 + r*l'*w1.|>=a by EUCLID:37;
       hence |.(1-r)*w1 + r*w2.|>=a by EUCLID:34;
       case a<=0;
        hence |.(1-r)*w1 + r*w2.|>=a by TOPRNS_1:26;
       end;
      hence x in Q by A1,A4,A23;
     end;
       LSeg(w4,w3) c=Q
     proof let x be set;assume x in LSeg(w4,w3);
       then x in { (1-r)*w4 + r*w3 : 0 <= r & r <= 1 } by TOPREAL1:def 4;
       then consider r such that
       A38: x=(1-r)*w4 + r*w3 &( 0 <= r & r <= 1);
        now per cases;
      case a>0;
          then a/|.w0.|>0 by A3,REAL_2:127;
         then A39: r*l'>=0 by A38,REAL_2:121;
       reconsider w5=(1-0)*w4 + 0 * w1 as Point of TOP-REAL n;
       A40:(1-0)*w4+0 * w1=(1-0)*w4+0.REAL n by EUCLID:33
         .=(1-0)*w4 by EUCLID:31
       .=w4 by EUCLID:33;
       reconsider w5'=w5 as Point of Euclid n by TOPREAL3:13;
         w5 in {(1-r1)*w4 + r1 * w1:0 <= r1 & r1 <= 1};
       then A41:w5 in LSeg(w1,w4) by TOPREAL1:def 4;
       A42:w5' in the carrier of TOP-REAL n;
       reconsider o=0.REAL n as Point of Euclid n by TOPREAL3:13;
       reconsider P=LSeg(w4,w1) as Subset of TopSpaceMetr(Euclid n)
                                              by EUCLID:def 8;
    ((dist(o)).:(P)) c= REAL
  proof let x be set;assume x in ((dist(o)).:(P));
    then consider z being set such that
    A43: z in dom (dist(o)) & z in P & x=(dist(o)).z by FUNCT_1:def 12;
    reconsider z2=z as Point of Euclid n by A43,TOPREAL3:13;
      (dist(o)).z2 =dist(z2,o) by WEIERSTR:def 6;
   hence x in REAL by A43;
  end;
  then reconsider F=((dist(o)).:(P)) as Subset of REAL;
  A44:w5' in the carrier of TopSpaceMetr(Euclid n) by A42,EUCLID:def 8;
    for r being real number st r in ((dist(o)).:(P)) holds 0<=r
   proof let r be real number;assume r in ((dist(o)).:(P));
     then consider x being set
     such that A45:x in dom (dist(o)) &
     x in P & r=(dist(o)).x by FUNCT_1:def 12;
     reconsider w0=x as Point of Euclid n by A45,TOPREAL3:13;
       r=dist(w0,o) by A45,WEIERSTR:def 6;
    hence 0<=r by METRIC_1:5;
   end;
  then A46:F is bounded_below by SEQ_4:def 2;
    w5' in dom (dist(o)) & dist(w5',o)=(dist(o)).w5'
   by A44,FUNCT_2:def 1,WEIERSTR:def 6;
  then dist(w5',o) in ((dist(o)).:(P)) by A41,FUNCT_1:def 12;
  then lower_bound F <=dist(w5',o) by A46,SEQ_4:def 5;
       then dist(w5',o)>= lower_bound([#]((dist(o)).:(P))) by WEIERSTR:def 3;
       then dist(w5',o)>= lower_bound((dist(o)).:(P)) by WEIERSTR:def 5;
       then dist(w5',o)>=|.w0.| by A3,WEIERSTR:def 8;
       then |.w5-0.REAL n.| >=|.w0.| by JGRAPH_1:45;
then A47:     |.w5.| >=|.w0.| by Th13;
           r*l'*|.w0.|
          =r*a/|.w0.|*|.w0.| by XCMPLX_1:75
         .=r*a by A3,XCMPLX_1:88;
         then A48:r*l'*|.w4.|>= r*a by A39,A40,A47,AXIOMS:25;
         consider q1 being Point of TOP-REAL n such that
         A49:q1=w4 & |.q1.| >= a by A1,A4;
         A50:1-r>=0 by A38,SQUARE_1:12;
         then A51:(1-r)*|.w4.|>=(1-r)*a by A49,AXIOMS:25;
           (1-r)+ r*l'>=0 by A39,A50,Th5;
         then abs((1-r)+r*l')*|.w4.|=((1-r)+ r*l')*|.w4.| by ABSVALUE:def 1
         .= (1-r)*|.w4.|+r*l'*|.w4.| by XCMPLX_1:8;
         then A52:abs((1-r)+r*l')*|.w4.|>=r*a+(1-r)*a by A48,A51,REAL_1:55;
           r*a+(1-r)*a
         =(r+(1-r))*a by XCMPLX_1:8
         .=1 * a by XCMPLX_1:27
         .=a;
         then |.((1-r)+ r*l')*w4.|>=a by A52,TOPRNS_1:8;
         then |.(1-r)*w4 + r*l'*w4.|>=a by EUCLID:37;
       hence |.(1-r)*w4 + r*w3.|>=a by EUCLID:34;
       case a<=0;
        hence |.(1-r)*w4 + r*w3.|>=a by TOPRNS_1:26;
       end;
      hence x in Q by A1,A4,A38;
     end;
  hence thesis by A8,A21,A22;
 end;

canceled;

theorem Th52:for f being FinSequence of REAL holds f is Element of REAL (len f)
 & f is Point of TOP-REAL (len f)
proof let f be FinSequence of REAL;
    f is Element of REAL* by FINSEQ_1:def 11;
 then f in { s where s is Element of REAL*: len s = len f };
 then f in ((len f)-tuples_on REAL) by FINSEQ_2:def 4;
  then f is Element of REAL (len f) by EUCLID:def 1;
 hence thesis by EUCLID:25;
end;

theorem Th53:for x being Element of REAL n,f,g being FinSequence of REAL,
r being Real st f=x & g=r*x holds len f=len g &
for i being Nat st 1<=i & i<=len f holds g/.i=r*f/.i
proof let x be Element of REAL n,f,g be FinSequence of REAL,r be Real;
  assume A1: f=x & g=r*x;
   then A2:len f=n by EUCLID:2;
   A3:len g=n by A1,EUCLID:2;
     g=(r multreal)*x by A1,RVSUM_1:def 7;
   then A4:g=(multreal[;](r,id REAL))*x by RVSUM_1:def 3;
 set h1= (dom (id REAL)) --> r;
 reconsider h2= (id REAL) as Function;
  A5:g= (multreal * (<:h1, h2:>))*x by A4,FUNCOP_1:def 5;
  A6: dom (<:h1, h2:>)=dom (h1) /\ dom ((id REAL)) &
  for u being set st u in dom (<:h1,h2:>) holds (<:h1, h2:>).u = [h1.u,h2.u]
   by FUNCT_3:def 8;
    for i being Nat st 1<=i & i<=len f holds g/.i=r*f/.i
   proof let i be Nat;assume
     A7:1<=i & i<=len f; then i in Seg len f by FINSEQ_1:3;
     then i in dom g by A2,A3,FINSEQ_1:def 3;
     then A8:g.i=(multreal * (<:h1, h2:>)).(x.i) by A5,FUNCT_1:22;
     A9:dom h2=REAL by FUNCT_1:34;
     A10:dom (<:h1, h2:>)=dom h1 /\ REAL by A6,FUNCT_1:34;
     A11:f.i=f/.i by A7,FINSEQ_4:24;
A12: dom h1=dom (id REAL) by FUNCOP_1:19 .=REAL by FUNCT_1:34;
     then A13:(<:h1, h2:>).(x.i)=[h1.(x.i),h2.(x.i)] by A10,FUNCT_3:def 8;
     A14:h1.(x.i) = r by A9,FUNCOP_1:13;
     A15:h2.(x.i)=x.i by FUNCT_1:34;
       (multreal * (<:h1, h2:>)).(x.i)=multreal.((<:h1, h2:>).(x.i))
                                          by A10,A12,FUNCT_1:23;
     then g.i=multreal.(r,f.i) by A1,A8,A13,A14,A15,BINOP_1:def 1;
     then g.i=r*(f/.i) by A11,VECTSP_1:def 14;
    hence g/.i=r*f/.i by A2,A3,A7,FINSEQ_4:24;
   end;
  hence thesis by A1,A2,EUCLID:2;
end;

theorem Th54: for x being Element of REAL n,f being FinSequence
 st x<> 0*n & x=f holds ex i being Nat st 1<=i & i<=n & f.i<>0
proof let x be Element of REAL n,f be FinSequence;
  assume A1:x<> 0*n & x=f;
   then A2:len f=n by EUCLID:2;
  assume A3: not ex i being Nat st 1<=i & i<=n & f.i<>0;
      for z being set holds z in f iff
      ex x,y st x in Seg n & y in {0} & z = [x,y]
      proof let z be set;
       hereby assume A4:z in f;
         then consider x0,y0 being set such that
         A5: z = [x0,y0] by RELAT_1:def 1;
         A6:x0 in dom f by A4,A5,RELAT_1:def 4;
         then A7:x0 in Seg len f by FINSEQ_1:def 3;
         reconsider n1=x0 as Nat by A6;
           1<=n1 & n1<=len f by A7,FINSEQ_1:3;
         then A8:f.n1=0 by A2,A3;
           y0=f.x0 by A4,A5,FUNCT_1:8;
         then y0 in {0} by A8,TARSKI:def 1;
        hence ex x,y st x in Seg n & y in {0} & z = [x,y] by A2,A5,A7;
       end;
         given x,y such that
         A9: x in Seg n & y in {0} & z = [x,y];
        A10:y=0 by A9,TARSKI:def 1;
        reconsider n1=x as Nat by A9;
          1<=n1 & n1<=n by A9,FINSEQ_1:3;
        then x in dom f & y=f.x by A2,A3,A9,A10,FINSEQ_1:def 3;
       hence z in f by A9,FUNCT_1:8;
      end;
    then f=[:Seg n,{0}:] by ZFMISC_1:def 2;
    then f=(Seg n -->0) by FUNCOP_1:def 2;
    then x=(n |-> (0 qua Real)) by A1,FINSEQ_2:def 2;
   hence contradiction by A1,EUCLID:def 4;
end;

theorem Th55: for x being Element of REAL n
 st n>=2 & x<> 0*n holds ex y being Element of REAL n st
 not ex r being Real st y=r*x or x=r*y
proof let x be Element of REAL n;
 assume A1:n>=2 & x<> 0*n;
  reconsider f=x as FinSequence of REAL;
  consider i2 being Nat such that
  A2: 1<=i2 & i2<=n & f.i2<>0 by A1,Th54;
  A3:len f=n by EUCLID:2;
  then A4:1<=len f by A1,AXIOMS:22;
  per cases;
  suppose A5:i2>1;
    reconsider g=(<*((f/.1)+1)*>)^mid(f,2,len f) as FinSequence of REAL;
    A6:len g=len (<*(f/.1+1)*>) + len (mid(f,2,len f)) by FINSEQ_1:35;
    A7:len (<*(f/.1+1)*>)=1 by FINSEQ_1:56;
    A8:len (mid(f,2,len f))=len f-'2+1 by A1,A3,A4,JORDAN3:27
    .=len f-2+1 by A1,A3,SCMFSA_7:3;
    then A9:len g=1+(len f-2+1) by A6,FINSEQ_1:56
    .=len f-2+(1+1) by XCMPLX_1:1
    .=len f by XCMPLX_1:27;
    then reconsider y2=g as Element of REAL n by A3,Th52;
      now given r being Real such that A10:y2=r*x or x=r*y2;
      per cases by A10;
      suppose A11:y2=r*x;
        A12:g/.i2=g.i2 by A2,A3,A9,FINSEQ_4:24;
        A13:f/.i2=f.i2 by A2,A3,FINSEQ_4:24;
        A14:1<=len f by A1,A3,AXIOMS:22;
        A15:i2-'1=i2-1 by A2,SCMFSA_7:3;
        A16:1<=i2-'1 by A5,JORDAN3:12;
          i2<=len f-(1+1)+(1+1) by A2,A3,XCMPLX_1:27;
        then i2<=len f-(1+1)+1+1 by XCMPLX_1:1;
        then i2-1<=len f-(1+1)+1+1-1 by REAL_1:49;
        then A17:i2-'1<=len (mid(f,2,len f)) by A8,A15,XCMPLX_1:26;
        A18:i2-'1+2-'1=i2-'1+(1+1)-'1
               .=i2-'1+1+1-'1 by XCMPLX_1:1
               .=i2-'1+1 by BINARITH:39 .=i2-1+1 by A2,SCMFSA_7:3
               .=i2 by XCMPLX_1:27;
          i2<=len f-2+(1+1) by A2,A3,XCMPLX_1:27;
        then 1+1<=i2 & i2<=1+len (mid(f,2,len f)) by A5,A8,NAT_1:38,XCMPLX_1:1
;
        then g.i2= (mid(f,2,len f)).(i2-1) by A7,FINSEQ_1:36
         .= (mid(f,2,len f)).(i2-'1) by A2,SCMFSA_7:3
         .=f.i2 by A1,A3,A14,A16,A17,A18,JORDAN3:27;
        then 1 * f/.i2=r*f/.i2 by A2,A3,A11,A12,A13,Th53;
        then A19:1=r by A2,A13,XCMPLX_1:5;
        A20:g/.1=r*f/.1 by A11,A14,Th53;
          g/.1=g.1 by A9,A14,FINSEQ_4:24;
        then f/.1+1=1 * f/.1 by A19,A20,FINSEQ_1:58;
        then f/.1+1-f/.1=0 by XCMPLX_1:14; hence contradiction by XCMPLX_1:26
;
      suppose A21:x=r*y2;
        A22:g/.i2=g.i2 by A2,A3,A9,FINSEQ_4:24;
        A23:f/.i2=f.i2 by A2,A3,FINSEQ_4:24;
        A24:1<=len f by A1,A3,AXIOMS:22;
        A25:i2-'1=i2-1 by A2,SCMFSA_7:3;
        A26:1<=i2-'1 by A5,JORDAN3:12;
          i2<=len f-(1+1)+(1+1) by A2,A3,XCMPLX_1:27;
        then i2<=len f-(1+1)+1+1 by XCMPLX_1:1;
        then i2-1<=len f-(1+1)+1+1-1 by REAL_1:49;
        then A27:i2-'1<=len (mid(f,2,len f)) by A8,A25,XCMPLX_1:26;
        A28:i2-'1+2-'1=i2-'1+(1+1)-'1
               .=i2-'1+1+1-'1 by XCMPLX_1:1
               .=i2-'1+1 by BINARITH:39 .=i2-1+1 by A2,SCMFSA_7:3
               .=i2 by XCMPLX_1:27;
          i2<=len f-2+(1+1) by A2,A3,XCMPLX_1:27;
        then 1+1<=i2 & i2<=1+len (mid(f,2,len f)) by A5,A8,NAT_1:38,XCMPLX_1:1
;
        then g.i2= (mid(f,2,len f)).(i2-1) by A7,FINSEQ_1:36
         .= (mid(f,2,len f)).(i2-'1) by A2,SCMFSA_7:3
         .=f.i2 by A1,A3,A24,A26,A27,A28,JORDAN3:27;
        then 1 * f/.i2=r*f/.i2 by A2,A3,A9,A21,A22,A23,Th53;
        then A29:1=r by A2,A23,XCMPLX_1:5;
        A30:f/.1=r*g/.1 by A9,A21,A24,Th53;
          g/.1=g.1 by A9,A24,FINSEQ_4:24;
        then f/.1+1=1 * f/.1 by A29,A30,FINSEQ_1:58;
        then f/.1+1-f/.1=0 by XCMPLX_1:14; hence contradiction by XCMPLX_1:26
;
    end;
   hence ex y being Element of REAL n st (not ex r being Real st
     y=r*x or x=r*y);
  suppose i2<=1;
    then A31:i2=1 by A2,AXIOMS:21;
    reconsider g=mid(f,1,len f-'1)^<*(f/.len f+1)*> as FinSequence of REAL;
    A32:len (<*(f/.len f+1)*>)=1 by FINSEQ_1:56;
    A33:len f-'1=len f-1 by A4,SCMFSA_7:3;
    A34:len f-'1<=len f by JORDAN3:7;
A35:1+1-1<=len f-1 by A1,A3,REAL_1:49;
    then A36:len f-'1-'1+1=len f-1-1+1 by A33,SCMFSA_7:3
    .=len f-(1+1)+1 by XCMPLX_1:36;
    then A37:len (mid(f,1,len f-'1))=len f-2+1
      by A4,A33,A34,A35,JORDAN3:27;
    A38:len (mid(f,1,len f-'1)) =len f-(1+1)+1 by A4,A33,A34,A35,A36,JORDAN3:27
    .=len f-1-1+1 by XCMPLX_1:36
    .=len f-1 by XCMPLX_1:27;
    A39:len g=(len f-2+1)+1 by A32,A37,FINSEQ_1:35
    .=len f-2+(1+1) by XCMPLX_1:1
    .=len f by XCMPLX_1:27;
    then reconsider y2=g as Element of REAL n by A3,Th52;
      now given r being Real such that A40:y2=r*x or x=r*y2;
      per cases by A40;
      suppose A41:y2=r*x;
        A42:g/.i2=g.i2 by A2,A3,A39,FINSEQ_4:24;
        A43:f/.i2=f.i2 by A2,A3,FINSEQ_4:24;
        A44:1<=len f by A1,A3,AXIOMS:22;
          g.i2= (mid(f,1,len f-'1)).i2 by A31,A35,A38,JORDAN3:17
        .=f.i2 by A31,A33,A34,A35,JORDAN3:32;
        then 1 * f/.i2=r*f/.i2 by A2,A3,A41,A42,A43,Th53;
        then A45:1=r by A2,A43,XCMPLX_1:5;
        A46:g/.len f=r*f/.len f by A41,A44,Th53;
        A47:g/.len f=g.len f by A39,A44,FINSEQ_4:24;
          g.len f= g.(len f -1+1) by XCMPLX_1:27
        .=f/.len f+1 by A38,JORDAN3:16;
        then f/.len f+1-f/.len f=0 by A45,A46,A47,XCMPLX_1:14;
       hence contradiction by XCMPLX_1:26;
      suppose A48:x=r*y2;
        A49:g/.i2=g.i2 by A2,A3,A39,FINSEQ_4:24;
        A50:f/.i2=f.i2 by A2,A3,FINSEQ_4:24;
        A51:1<=len f by A1,A3,AXIOMS:22;
          g.i2= (mid(f,1,len f-'1)).i2 by A31,A35,A38,JORDAN3:17
        .=f.(i2) by A31,A33,A34,A35,JORDAN3:32;
        then 1 * f/.i2=r*f/.i2 by A2,A3,A39,A48,A49,A50,Th53;
        then A52:1=r by A2,A50,XCMPLX_1:5;
        A53:f/.len f=r*g/.len f by A39,A48,A51,Th53;
        A54:g/.len f=g.len f by A39,A51,FINSEQ_4:24;
          g.len f=g.(len f-1+1) by XCMPLX_1:27
        .=f/.len f+1 by A38,JORDAN3:16;
        then f/.len f+1-f/.len f=0 by A52,A53,A54,XCMPLX_1:14;
     hence contradiction by XCMPLX_1:26;
    end;
   hence thesis;
end;

theorem Th56: for a being Real,
Q being Subset of TOP-REAL n,
w1,w7 being Point of TOP-REAL n
 st n>=2 & Q={q : (|.q.|) > a } & w1 in Q & w7 in Q &
 (ex r being Real st w1=r*w7 or w7=r*w1)
 holds ex w2,w3,w4,w5,w6 being Point of TOP-REAL n st
 w2 in Q & w3 in Q & w4 in Q &
 w5 in Q & w6 in Q &
 LSeg(w1,w2) c=Q & LSeg(w2,w3) c= Q & LSeg(w3,w4) c= Q & LSeg(w4,w5) c= Q &
 LSeg(w5,w6) c=Q & LSeg(w6,w7) c= Q
proof let a be Real,
  Q be Subset of TOP-REAL n,
  w1,w7 be Point of TOP-REAL n;
  assume A1: n>=2 & Q={q : (|.q.|) > a } & w1 in Q & w7 in Q &
   (ex r being Real st w1=r*w7 or w7=r*w1);
   then consider r8 being Real such that
   A2: w1=r8*w7 or w7=r8*w1;
   reconsider y1=w1 as Element of REAL n by EUCLID:25;
   per cases;
   suppose A3:a>=0; then A4:a+1>0 by REAL_1:69;
     now assume A5:w1=0.REAL n;
       ex q st q=w1 & (|.q.|)>a by A1;
    hence contradiction by A3,A5,TOPRNS_1:24;
   end; then y1<>0*n by EUCLID:def 9;
   then consider y being Element of REAL n such that
   A6:  (not ex r being Real st y=r*y1 or y1=r*y) by A1,Th55;
   set y4=((a+1)/|.y.|)*y;
   reconsider w4=y4 as Point of TOP-REAL n by EUCLID:25;
   A7: |.y.|>=0 by EUCLID:12;
A8: now assume |.y.|=0;
    then A9:y=0*n by EUCLID:11;
      0 *y1=0 *w1 by EUCLID:def 11 .=0.REAL n by EUCLID:33
    .=0*n by EUCLID:def 9;
   hence contradiction by A6,A9;
   end;
   then A10: (a+1)/|.y.|>0 by A4,A7,REAL_2:127;
     |.w4.|=|.y4.| by JGRAPH_1:def 5 .=abs((a+1)/|.y.|)*|.y.| by EUCLID:14
   .= (a+1)/|.y.|*|.y.| by A10,ABSVALUE:def 1 .=a+1 by A8,XCMPLX_1:88;
 then A11: |.w4.|>a by REAL_1:69;
   then A12:w4 in Q by A1;
   A13:now given r being Real such that
      A14: w1=r*w4 or w4=r*w1;
      reconsider y'=y,y1'=y1 as Element of n-tuples_on REAL by EUCLID:def 1;
        y1=r*(((a+1)/|.y.|)*y) or ((a+1)/|.y.|)*y=r*y1 by A14,EUCLID:def 11;
      then y1=(r*((a+1)/|.y.|))*y or
      ((a+1)/|.y.|)"*((a+1)/|.y.|)*y'=((a+1)/|.y.|)"*(r*y1) by RVSUM_1:71;
      then y1=(r*((a+1)/|.y.|))*y or
      ((a+1)/|.y.|)"*((a+1)/|.y.|)*y=((a+1)/|.y.|)"*r*y1'
                                       by RVSUM_1:71;
      then y1=(r*((a+1)/|.y.|))*y' or 1 *y=((a+1)/|.y.|)"*r*y1
                                       by A10,XCMPLX_0:def 7;
then A15:   y1=(r*((a+1)/|.y.|))*y or y'=((a+1)/|.y.|)"*r*y1 by RVSUM_1:74;
      per cases by A15;
      suppose y1=(r*((a+1)/|.y.|))*y; hence contradiction by A6;
      suppose y=((a+1)/|.y.|)"*r*y1; hence contradiction by A6;
   end;
   then consider w2,w3 being Point of TOP-REAL n such that
   A16: w2 in Q & w3 in Q &
     LSeg(w1,w2) c=Q & LSeg(w2,w3) c= Q & LSeg(w3,w4) c= Q by A1,A12,Th49;
     now given r1 being Real such that
     A17: w4=r1*w7 or w7=r1*w4;
     A18:now assume r1=0;
  then A19:w4=0.REAL n or w7=0.REAL n by A17,EUCLID:33;
         ex q7 being Point of TOP-REAL n st q7=w7 & |.q7.|>a by A1;
      hence contradiction by A3,A11,A19,TOPRNS_1:24;
     end;
      per cases by A2;
      suppose A20: w1=r8*w7;
          now per cases by A17;
        case w4=r1*w7;
          then r1"*w4=r1"*r1*w7 by EUCLID:34;
          then r1"*w4=1 *w7 by A18,XCMPLX_0:def 7;
          then r1"*w4=w7 by EUCLID:33;
          then w1=r8*r1"*w4 by A20,EUCLID:34;
         hence contradiction by A13;
        case w7=r1*w4;
          then r1"*w7=r1"*r1*w4 by EUCLID:34;
          then r1"*w7=1 *w4 by A18,XCMPLX_0:def 7;
          then r1"*w7=w4 by EUCLID:33;
          then r1""*w4=r1""*r1"*w7 by EUCLID:34;
          then r1""*w4=1 *w7 by A18,XCMPLX_0:def 7;
          then r1""*w4=w7 by EUCLID:33;
          then w1=r8*r1""*w4 by A20,EUCLID:34;
         hence contradiction by A13;
        end;
       hence contradiction;
      suppose A21: w7=r8*w1;
        then A22:r8"*w7=r8"*r8*w1 by EUCLID:34;
          now assume r8=0;
then A23:     w7=0.REAL n by A21,EUCLID:33;
           ex q7 being Point of TOP-REAL n st q7=w7 & |.q7.|>a by A1;
         hence contradiction by A3,A23,TOPRNS_1:24;
        end;
        then r8"*w7=1 *w1 by A22,XCMPLX_0:def 7;
        then A24:r8"*w7=w1 by EUCLID:33;
          now per cases by A17;
        case w4=r1*w7;
          then r1"*w4=r1"*r1*w7 by EUCLID:34;
          then r1"*w4=1 *w7 by A18,XCMPLX_0:def 7;
          then r1"*w4=w7 by EUCLID:33;
          then w1=r8"*r1"*w4 by A24,EUCLID:34;
         hence contradiction by A13;
        case w7=r1*w4;
        then r1"*w7=r1"*r1*w4 by EUCLID:34;
          then r1"*w7=1 *w4 by A18,XCMPLX_0:def 7;
          then r1"*w7=w4 by EUCLID:33;
          then r1""*w4=r1""*r1"*w7 by EUCLID:34;
          then r1""*w4=1 *w7 by A18,XCMPLX_0:def 7;
          then r1""*w4=w7 by EUCLID:33;
          then w1=r8"*r1""*w4 by A24,EUCLID:34;
         hence contradiction by A13;
        end;
       hence contradiction;
   end;
   then consider w2',w3' being Point of TOP-REAL n such that
   A25: w2' in Q & w3' in Q &
     LSeg(w4,w2') c=Q & LSeg(w2',w3') c= Q & LSeg(w3',w7) c= Q
                             by A1,A12,Th49;
 thus thesis by A12,A16,A25;
   suppose A26:a<0;
     A27:the carrier of TOP-REAL n=REAL n by EUCLID:25;
       REAL n c= Q
     proof let x be set;assume x in REAL n;
       then reconsider w=x as Point of TOP-REAL n by EUCLID:25;
         |.w.|>=0 by TOPRNS_1:26;
      hence x in Q by A1,A26;
     end;
     then A28:Q=the carrier of TOP-REAL n by A27,XBOOLE_0:def 10;
     set w2=0.REAL n;
     A29:LSeg(w1,w2) c=Q by A28;
      LSeg(w2,w7) c=Q by A28;
 hence ex w2,w3,w4,w5,w6 being Point of TOP-REAL n st
 w2 in Q & w3 in Q & w4 in Q &
 w5 in Q & w6 in Q &
 LSeg(w1,w2) c=Q & LSeg(w2,w3) c= Q & LSeg(w3,w4) c= Q & LSeg(w4,w5) c= Q &
 LSeg(w5,w6) c=Q & LSeg(w6,w7) c= Q by A28,A29;
end;

theorem Th57: for a being Real,
Q being Subset of TOP-REAL n,
w1,w7 being Point of TOP-REAL n
 st n>=2 & Q=(REAL n)\ {q : (|.q.|) < a } & w1 in Q & w7 in Q &
 (ex r being Real st w1=r*w7 or w7=r*w1)
 holds ex w2,w3,w4,w5,w6 being Point of TOP-REAL n st
 w2 in Q & w3 in Q & w4 in Q &
 w5 in Q & w6 in Q &
 LSeg(w1,w2) c=Q & LSeg(w2,w3) c= Q & LSeg(w3,w4) c= Q & LSeg(w4,w5) c= Q &
 LSeg(w5,w6) c=Q & LSeg(w6,w7) c= Q
proof let a be Real,
  Q be Subset of TOP-REAL n,
  w1,w7 be Point of TOP-REAL n;
  assume A1: n>=2 & Q=(REAL n)\ {q : (|.q.|) < a } & w1 in Q & w7 in Q &
   (ex r being Real st w1=r*w7 or w7=r*w1);
   then consider r8 being Real such that
   A2: w1=r8*w7 or w7=r8*w1;
   reconsider y1=w1 as Element of REAL n by EUCLID:25;
   per cases;
   suppose A3:a>0;
     now assume w1=0.REAL n;
     then |.w1.|=0 by TOPRNS_1:24;
     then w1 in {q : (|.q.|) < a } by A3;
    hence contradiction by A1,XBOOLE_0:def 4;
   end; then y1<>0*n by EUCLID:def 9;
   then consider y being Element of REAL n such that
   A4:  (not ex r being Real st y=r*y1 or y1=r*y) by A1,Th55;
   set y4=(a/|.y.|)*y;
   reconsider w4=y4 as Point of TOP-REAL n by EUCLID:25;
   A5: |.y.|>=0 by EUCLID:12;
A6: now assume |.y.|=0;
    then A7:y=0*n by EUCLID:11;
      0 *y1=0 *w1 by EUCLID:def 11 .=0.REAL n by EUCLID:33
    .=0*n by EUCLID:def 9;
   hence contradiction by A4,A7;
   end;
   then A8: a/|.y.|>0 by A3,A5,REAL_2:127;
   A9: |.w4.|=|.y4.| by JGRAPH_1:def 5 .=abs(a/|.y.|)*|.y.| by EUCLID:14
   .= a/|.y.|*|.y.| by A8,ABSVALUE:def 1 .=a by A6,XCMPLX_1:88;
   A10:now assume w4 in {q : (|.q.|) < a };
    then ex q st q=w4 & (|.q.|) < a;
    hence contradiction by A9;
   end;
   then A11:w4 in Q by A1,XBOOLE_0:def 4;
   A12:now given r being Real such that
      A13: w1=r*w4 or w4=r*w1;
      reconsider y'=y,y1'=y1 as Element of n-tuples_on REAL by EUCLID:def 1;
        y1=r*((a/|.y.|)*y) or (a/|.y.|)*y=r*y1 by A13,EUCLID:def 11;
      then y1=(r*(a/|.y.|))*y or (a/|.y.|)"*(a/|.y.|)*y'=(a/|.y.|)"*(r*y1)
                                       by RVSUM_1:71;
      then y1=(r*(a/|.y.|))*y or (a/|.y.|)"*(a/|.y.|)*y=(a/|.y.|)"*r*y1'
                                       by RVSUM_1:71;
      then y1=(r*(a/|.y.|))*y' or 1 *y=(a/|.y.|)"*r*y1
                                       by A8,XCMPLX_0:def 7;
then A14:   y1=(r*(a/|.y.|))*y or y'=(a/|.y.|)"*r*y1 by RVSUM_1:74;
      per cases by A14;
      suppose y1=(r*(a/|.y.|))*y;
       hence contradiction by A4;
      suppose y=(a/|.y.|)"*r*y1;
       hence contradiction by A4;
   end;
   then consider w2,w3 being Point of TOP-REAL n such that
   A15: w2 in Q & w3 in Q &
     LSeg(w1,w2) c=Q & LSeg(w2,w3) c= Q & LSeg(w3,w4) c= Q
                             by A1,A11,Th50;
     now given r1 being Real such that
     A16: w4=r1*w7 or w7=r1*w4;
     A17:now assume r1=0;
       then w4=0.REAL n or w7=0.REAL n by A16,EUCLID:33;
       then |.w4.|=0 or |.w7.|=0 by TOPRNS_1:24;
     then w4 in {q : (|.q.|) < a } or
     w7 in {q2 where q2 is Point of TOP-REAL n: (|.q2.|) < a } by A3;
      hence contradiction by A1,A10,XBOOLE_0:def 4;
     end;
        now per cases by A2;
      case A18: w1=r8*w7;
          now per cases by A16;
        case w4=r1*w7;
          then r1"*w4=r1"*r1*w7 by EUCLID:34;
          then r1"*w4=1 *w7 by A17,XCMPLX_0:def 7;
          then r1"*w4=w7 by EUCLID:33;
          then w1=r8*r1"*w4 by A18,EUCLID:34;
         hence contradiction by A12;
        case w7=r1*w4;
          then r1"*w7=r1"*r1*w4 by EUCLID:34;
          then r1"*w7=1 *w4 by A17,XCMPLX_0:def 7;
          then r1"*w7=w4 by EUCLID:33;
          then r1""*w4=r1""*r1"*w7 by EUCLID:34;
          then r1""*w4=1 *w7 by A17,XCMPLX_0:def 7;
          then r1""*w4=w7 by EUCLID:33;
          then w1=r8*r1""*w4 by A18,EUCLID:34;
         hence contradiction by A12;
        end;
       hence contradiction;
      case A19: w7=r8*w1;
        then A20:r8"*w7=r8"*r8*w1 by EUCLID:34;
          now assume r8=0;
         then w7=0.REAL n by A19,EUCLID:33;
        then |.w7.|=0 by TOPRNS_1:24;
        then w7 in {q : (|.q.|) < a } by A3;
        hence contradiction by A1,XBOOLE_0:def 4;
        end;
        then r8"*w7=1 *w1 by A20,XCMPLX_0:def 7;
        then A21:r8"*w7=w1 by EUCLID:33;
          now per cases by A16;
        case w4=r1*w7;
          then r1"*w4=r1"*r1*w7 by EUCLID:34;
          then r1"*w4=1 *w7 by A17,XCMPLX_0:def 7;
          then r1"*w4=w7 by EUCLID:33;
          then w1=r8"*r1"*w4 by A21,EUCLID:34;
         hence contradiction by A12;
        case w7=r1*w4;
          then r1"*w7=r1"*r1*w4 by EUCLID:34;
          then r1"*w7=1 *w4 by A17,XCMPLX_0:def 7;
          then r1"*w7=w4 by EUCLID:33;
          then r1""*w4=r1""*r1"*w7 by EUCLID:34;
          then r1""*w4=1 *w7 by A17,XCMPLX_0:def 7;
          then r1""*w4=w7 by EUCLID:33;
          then w1=r8"*r1""*w4 by A21,EUCLID:34;
         hence contradiction by A12;
        end;
       hence contradiction;
      end;
    hence contradiction;
   end;
   then consider w2',w3' being Point of TOP-REAL n such that
   A22: w2' in Q & w3' in Q &
     LSeg(w4,w2') c=Q & LSeg(w2',w3') c= Q & LSeg(w3',w7) c= Q
                             by A1,A11,Th50;
   thus thesis by A11,A15,A22;
   suppose A23:a<=0;
     A24:the carrier of TOP-REAL n=REAL n by EUCLID:25;
       REAL n c= Q
     proof let x be set;assume A25:x in REAL n;
         now assume x in {q : (|.q.|) < a };
         then consider q being Point of TOP-REAL n such that
         A26:q=x & (|.q.|) < a;
        thus contradiction by A23,A26,TOPRNS_1:26;
       end;
      hence x in Q by A1,A25,XBOOLE_0:def 4;
     end;
     then A27:Q=the carrier of TOP-REAL n by A24,XBOOLE_0:def 10;
     set w2=0.REAL n;
     A28:LSeg(w1,w2) c=Q by A27;
       LSeg(w2,w7) c=Q by A27;
 hence thesis by A27,A28;
end;

theorem Th58:for a being Real st n>=1 holds {q: |.q.| >a} <>{}
proof let a be Real;assume A1:n>=1;
        now assume not (a+1)*(1.REAL n) in {q : |.q.| > a };
        then A2: |.(a+1)*(1.REAL n).|<=a;
        A3: |.(a+1)*(1.REAL n).|=abs(a+1)*|.(1.REAL n).| by TOPRNS_1:8
        .=abs(a+1)*(sqrt n) by Th37;
        A4:abs(a+1)>=0 by ABSVALUE:5;
          sqrt 1<=sqrt n by A1,SQUARE_1:94;
then A5:     abs(a+1)*1<=abs(a+1)*sqrt n by A4,AXIOMS:25,SQUARE_1:83;
          a+1<=abs(a+1) by ABSVALUE:11;
        then A6:a+1<= |.(a+1)*(1.REAL n).| by A3,A5,AXIOMS:22;
          a<a+1 by REAL_1:69;
       hence contradiction by A2,A6,AXIOMS:22;
      end;
    hence thesis;
end;

theorem Th59: for a being Real,
P being Subset of TOP-REAL n st n>=2 & P={q : |.q.| > a }
 holds P is connected
 proof let a be Real,
   P be Subset of TOP-REAL n;
  assume A1:n>=2 & P={q : |.q.| > a };
      then n>=1 by AXIOMS:22;
      then reconsider Q=P as non empty Subset of TOP-REAL n by A1,Th58;
     for w1,w7 being Point of TOP-REAL n st w1 in Q & w7 in Q & w1<>w7
    ex f being map of I[01],((TOP-REAL n)|Q) st f is continuous &
    w1=f.0 & w7=f.1
   proof let w1,w7 be Point of TOP-REAL n;
    assume A2:w1 in Q & w7 in Q & w1<>w7;
     per cases;
     suppose not (ex r being Real st w1=r*w7 or w7=r*w1);
       then consider w2,w3 being Point of TOP-REAL n such that
       A3: w2 in Q & w3 in Q &
        LSeg(w1,w2) c=Q & LSeg(w2,w3) c= Q & LSeg(w3,w7) c= Q
                                by A1,A2,Th49;
     thus ex h being map of I[01],(TOP-REAL n)|Q st h is continuous
                        & w1=h.0 & w7=h.1 by A2,A3,Th45;
     suppose (ex r being Real st w1=r*w7 or w7=r*w1);
      then consider w2,w3,w4,w5,w6 being Point of TOP-REAL n such that
      A4:w2 in Q & w3 in Q & w4 in Q & w5 in Q & w6 in Q &
       LSeg(w1,w2) c=Q & LSeg(w2,w3) c= Q & LSeg(w3,w4) c= Q
       & LSeg(w4,w5) c= Q &
       LSeg(w5,w6) c=Q & LSeg(w6,w7) c= Q by A1,A2,Th56;
    thus thesis by A2,A4,Th46;
   end;
  hence P is connected by JORDAN1:5;
 end;

theorem Th60:for a being Real st n>=1 holds
  (REAL n) \ {q: |.q.| < a} <> {}
proof let a be Real;assume A1:n>=1;
      A2:{q:(|.q.|)>a} c= (REAL n)\{q2:(|.q2.|)<a}
      proof let x be set;assume x in {q:(|.q.|)>a};
        then consider q such that A3:q=x & (|.q.|)>a;
          not (|.q.|)<a & q in the carrier of TOP-REAL n by A3;
        then A4:not (|.q.|)<a & q in REAL n by EUCLID:25;
          now assume x in {q2:(|.q2.|)<a};
          then ex q2 st q2=x & (|.q2.|)<a;
         hence contradiction by A3;
        end;
       hence thesis by A3,A4,XBOOLE_0:def 4;
      end;
     {q:(|.q.|)>a} <>{} by A1,Th58;
  hence thesis by A2,XBOOLE_1:3;
end;

theorem Th61: for a being Real,P being Subset of TOP-REAL n
 st n>=2 & P=(REAL n)\ {q : |.q.| < a }
 holds P is connected
 proof let a be Real,
   P be Subset of TOP-REAL n;
  assume A1:n>=2 & P=(REAL n)\ {q : |.q.| < a };
      then n>=1 by AXIOMS:22;
      then reconsider Q=P as non empty Subset of TOP-REAL n by A1,Th60;
     for w1,w7 being Point of TOP-REAL n st w1 in Q & w7 in Q & w1<>w7
    ex f being map of I[01],((TOP-REAL n)|Q) st f is continuous &
    w1=f.0 & w7=f.1
   proof let w1,w7 be Point of TOP-REAL n;
    assume A2:w1 in Q & w7 in Q & w1<>w7;
     per cases;
     suppose not (ex r being Real st w1=r*w7 or w7=r*w1);
       then consider w2,w3 being Point of TOP-REAL n such that
       A3: w2 in Q & w3 in Q &
        LSeg(w1,w2) c=Q & LSeg(w2,w3) c= Q & LSeg(w3,w7) c= Q
                                by A1,A2,Th50;
     thus thesis by A2,A3,Th45;
     suppose (ex r being Real st w1=r*w7 or w7=r*w1);
      then consider w2,w3,w4,w5,w6 being Point of TOP-REAL n such that
      A4:w2 in Q & w3 in Q & w4 in Q & w5 in Q & w6 in Q &
       LSeg(w1,w2) c=Q & LSeg(w2,w3) c= Q & LSeg(w3,w4) c= Q
       & LSeg(w4,w5) c= Q &
       LSeg(w5,w6) c=Q & LSeg(w6,w7) c= Q by A1,A2,Th57;
      consider f1 being map of I[01],((TOP-REAL n)|Q)
      such that A5:f1 is continuous &
        w1=f1.0 & w7=f1.1 by A2,A4,Th46;
      thus thesis by A5;
   end;
  hence P is connected by JORDAN1:5;
 end;

theorem Th62: for a being Real,n being Nat,
P being Subset of TOP-REAL n
 st n>=1 & P=(REAL n)\ {q where q is Point of TOP-REAL n: |.q.| < a }
 holds not P is Bounded
proof let a be Real,n be Nat,P be Subset of TOP-REAL n;
 assume A1:n>=1 &
  P=(REAL n)\ {q where q is Point of TOP-REAL n : |.q.| < a };
  per cases;
  suppose A2:a>0;
    now assume P is Bounded;
   then consider r being Real such that
   A3: for q being Point of TOP-REAL n st
       q in P holds |.q.|<r by Th40;
  consider p being Element of P;
  A4:P<>{} by A1,Th60;
  then p in P;
  then reconsider p as Point of TOP-REAL n;
  A5: |.p.|<r by A3,A4;
   A6: 0 <= (|.p.|) by TOPRNS_1:26;
        now assume not (a+r+1)*(1.REAL n) in (REAL n)
         \{q where q is Point of TOP-REAL n: (|.q.|) < a };
       then A7:not ((a+r+1)*(1.REAL n) in (REAL n) & not
       (a+r+1)*(1.REAL n) in
       {q where q is Point of TOP-REAL n : (|.q.|) < a }) by XBOOLE_0:def 4;
          (a+r+1)*(1.REAL n) in the carrier of TOP-REAL n;
        then consider q being Point of TOP-REAL n such that
        A8:q= (a+r+1)*(1.REAL n) & (|.q.|) < a by A7,EUCLID:25;
        A9: |.(a+r+1)*(1.REAL n).|=abs(a+r+1)*|.(1.REAL n).| by TOPRNS_1:8
        .=abs(a+r+1)*(sqrt n) by Th37;
        A10:abs(a+r+1)>=0 by ABSVALUE:5;
          sqrt 1<=sqrt n by A1,SQUARE_1:94;
then A11:     abs(a+r+1)*1<=abs(a+r+1)*sqrt n by A10,AXIOMS:25,SQUARE_1:83;
          a+r+1<=abs(a+r+1) by ABSVALUE:11;
        then A12:a+r+1<= |.(a+r+1)*(1.REAL n).| by A9,A11,AXIOMS:22;
        A13:a+r<a+r+1 by REAL_1:69;
          a<a+r by A5,A6,REAL_1:69;
        then a<a+r+1 by A13,AXIOMS:22;
       hence contradiction by A8,A12,AXIOMS:22;
      end;
      then A14:(|.((a+r+1)*(1.REAL n)).|)<=r by A1,A3;
       A15: |.(a+r+1)*(1.REAL n).|
        =abs(a+r+1)*|.(1.REAL n).| by TOPRNS_1:8
       .=abs(a+r+1)*(sqrt n) by Th37;
        A16:abs(a+r+1)>=0 by ABSVALUE:5;
          sqrt 1<=sqrt n by A1,SQUARE_1:94;
then A17:     abs(a+r+1)*1<=abs(a+r+1)*sqrt n by A16,AXIOMS:25,SQUARE_1:83;
          a+r+1<=abs(a+r+1) by ABSVALUE:11;
        then A18: a+r+1<= |.(a+r+1)*(1.REAL n).| by A15,A17,AXIOMS:22;
          a+r<a+r+1 by REAL_1:69;
        then A19:a+r<|.((a+r+1)*(1.REAL n)).| by A18,AXIOMS:22;
          r<r+a by A2,REAL_1:69;
       hence contradiction by A14,A19,AXIOMS:22;
  end;
 hence not P is Bounded;
 suppose A20:a<=0;
    now assume
     A21:{q where q is Point of TOP-REAL n: (|.q.|) < a }<>{};
      {q where q is Point of TOP-REAL n: (|.q.|) < a }
      c= the carrier of TOP-REAL n
     proof let z;assume z in
      {q where q is Point of TOP-REAL n: (|.q.|) < a };
      then consider q being Point of TOP-REAL n such that
      A22:q=z & (|.q.|) < a;
      thus z in the carrier of TOP-REAL n by A22;
     end;
    then reconsider
    Q={q where q is Point of TOP-REAL n: (|.q.|) < a }
      as Subset of TOP-REAL n;
    consider d being Element of Q;
      d in {q where q is Point of TOP-REAL n: (|.q.|) < a } by A21;
    then consider q being Point of TOP-REAL n such that
    A23:q=d & (|.q.|) < a;
   thus contradiction by A20,A23,TOPRNS_1:26;
  end;
  then P=the carrier of TOP-REAL n by A1,EUCLID:25;
  then P=[#](TOP-REAL n) by PRE_TOPC:12;
 hence not P is Bounded by A1,Th41;
end;

theorem Th63: for a being Real,P being Subset of TOP-REAL 1
 st P={q where q is Point of TOP-REAL 1:
   ex r st q=<*r*> & r > a } holds P is convex
proof let a be Real,P be Subset of TOP-REAL 1;
 assume A1:P={q where q is Point of TOP-REAL 1: ex r st q=<*r*> & r > a };
    for w1,w2 being Point of TOP-REAL 1
    st w1 in P & w2 in P holds LSeg(w1,w2) c= P
   proof let w1,w2 be Point of TOP-REAL 1;
    assume A2:w1 in P & w2 in P;
     then consider q1 being Point of TOP-REAL 1 such that
     A3:q1=w1 & ex r st q1=<*r*> & r > a by A1;
     consider r1 such that A4: q1=<*r1*> & r1 > a by A3;
     consider q2 being Point of TOP-REAL 1 such that
     A5:q2=w2 & ex r st q2=<*r*> & r > a by A1,A2;
     consider r2 such that A6: q2=<*r2*> & r2 > a by A5;
    thus LSeg(w1,w2) c= P
    proof let x be set;assume x in LSeg(w1,w2);
      then x in {(1-r)*w1+r*w2 : 0<=r & r<=1} by TOPREAL1:def 4;
      then consider r3 being Real such that
      A7: x=(1-r3)*w1+r3*w2 & (0<=r3 & r3<=1);
      A8:1-r3>=0 by A7,SQUARE_1:12;
     per cases;
     suppose A9:r3>0;
      A10:(1-r3)*r1>=(1-r3)*a by A4,A8,AXIOMS:25;
A11:   r3*r2>r3*a by A6,A9,REAL_1:70;
        (1-r3)*a+r3*a=((1-r3)+r3)*a by XCMPLX_1:8 .=1 *a by XCMPLX_1:27 .=a;
      then A12:(1-r3)*r1+r3*r2>a by A10,A11,REAL_1:67;
      A13:<*(1-r3)*r1+r3*r2*>=|[(1-r3)*r1+r3*r2]| by JORDAN2B:def 2
      .=|[(1-r3)*r1]|+|[r3*r2]| by JORDAN2B:27
      .=(1-r3)*|[r1]|+|[r3*r2]| by JORDAN2B:26
      .=(1-r3)*|[r1]|+r3*|[r2]| by JORDAN2B:26;
      A14: |[r1]|=w1 by A3,A4,JORDAN2B:def 2;
        |[r2]|=w2 by A5,A6,JORDAN2B:def 2;
     hence x in P by A1,A7,A12,A13,A14;
     suppose r3<=0; then r3=0 by A7;
      then x=w1+0 *w2 by A7,EUCLID:33
      .=w1+0.REAL 1 by EUCLID:33 .=w1 by EUCLID:31;
     hence x in P by A2;
    end;
   end;
 hence P is convex by JORDAN1:def 1;
end;

theorem Th64: for a being Real,P being Subset of TOP-REAL 1
 st P={q where q is Point of TOP-REAL 1 :
   ex r st q=<*r*> & r < -a } holds P is convex
proof let a be Real,P be Subset of TOP-REAL 1;
 assume A1:P={q where q is Point of TOP-REAL 1 : ex r st q=<*r*> & r < -a };
    for w1,w2 being Point of TOP-REAL 1
    st w1 in P & w2 in P holds LSeg(w1,w2) c= P
   proof let w1,w2 be Point of TOP-REAL 1;
    assume A2:w1 in P & w2 in P;
     then consider q1 being Point of TOP-REAL 1 such that
     A3:q1=w1 & ex r st q1=<*r*> & r < -a by A1;
     consider r1 such that A4: q1=<*r1*> & r1 < -a by A3;
     consider q2 being Point of TOP-REAL 1 such that
     A5:q2=w2 & ex r st q2=<*r*> & r < -a by A1,A2;
     consider r2 such that A6: q2=<*r2*> & r2 < -a by A5;
    thus LSeg(w1,w2) c= P
    proof let x be set;assume x in LSeg(w1,w2);
      then x in {(1-r)*w1+r*w2 : 0<=r & r<=1} by TOPREAL1:def 4;
      then consider r3 being Real such that
      A7: x=(1-r3)*w1+r3*w2 & (0<=r3 & r3<=1);
      A8:1-r3>=0 by A7,SQUARE_1:12;
     per cases;
     suppose A9:r3>0;
      A10:(1-r3)*r1<=(1-r3)*(-a) by A4,A8,AXIOMS:25;
A11:   r3*r2<r3*(-a) by A6,A9,REAL_1:70;
        (1-r3)*(-a)+r3*(-a)=((1-r3)+r3)*(-a) by XCMPLX_1:8
      .=1 *(-a) by XCMPLX_1:27 .=-a;
      then A12:(1-r3)*r1+r3*r2< -a by A10,A11,REAL_1:67;
      A13:<*(1-r3)*r1+r3*r2*>=|[(1-r3)*r1+r3*r2]| by JORDAN2B:def 2
      .=|[(1-r3)*r1]|+|[r3*r2]| by JORDAN2B:27
      .=(1-r3)*|[r1]|+|[r3*r2]| by JORDAN2B:26
      .=(1-r3)*|[r1]|+r3*|[r2]| by JORDAN2B:26;
      A14: |[r1]|=w1 by A3,A4,JORDAN2B:def 2;
         |[r2]|=w2 by A5,A6,JORDAN2B:def 2;
      hence x in P by A1,A7,A12,A13,A14;
     suppose r3<=0; then r3=0 by A7;
      then x=w1+0 *w2 by A7,EUCLID:33
      .=w1+0.REAL 1 by EUCLID:33 .=w1 by EUCLID:31;
     hence x in P by A2;
    end;
   end;
 hence P is convex by JORDAN1:def 1;
end;

theorem Th65: for a being Real,P being Subset of TOP-REAL 1
 st P={q where q is Point of TOP-REAL 1 : ex r st q=<*r*> & r > a }
 holds P is connected
proof let a be Real,P be Subset of TOP-REAL 1;
 assume P={q where q is Point of TOP-REAL 1 : ex r st q=<*r*> & r > a };
  then P is convex by Th63;
 hence thesis by Th14;
end;

theorem Th66: for a being Real,P being Subset of TOP-REAL 1
 st P={q where q is Point of TOP-REAL 1 : ex r st q=<*r*> & r < -a }
  holds P is connected
proof let a be Real,P be Subset of TOP-REAL 1;
 assume P={q where q is Point of TOP-REAL 1 : ex r st q=<*r*> & r < -a };
  then P is convex by Th64;
 hence P is connected by Th14;
end;

theorem Th67: for W being Subset of Euclid 1,a being Real,
P being Subset of TOP-REAL 1
 st W={q where q is Point of TOP-REAL 1 : ex r st q=<*r*> & r > a }
    & P=W holds P is connected & W is not bounded
proof let W be Subset of Euclid 1,a be Real,
  P be Subset of TOP-REAL 1;
 assume
 A1:W={q where q is Point of TOP-REAL 1 : ex r st q=<*r*> & r > a } & P=W;
 hence P is connected by Th65;
   now assume W is bounded;
   then consider r such that
   A2: 0<r & for x,y being Point of Euclid 1 st
      x in W & y in W holds dist(x,y)<=r by TBSP_1:def 9;
       A3:a<=abs(a) by ABSVALUE:11;
       A4:abs(a)>=0 by ABSVALUE:5;
       then abs(a)+abs(a)>=0+abs(a) by AXIOMS:24;
       then abs(a)+abs(a)+abs(a)>=0+abs(a) by A4,AXIOMS:24;
       then A5:3*abs(a)>=abs(a) by XCMPLX_1:12;
         3*r>0 by A2,REAL_2:122;
       then 0+abs(a)<3*r+3*abs(a) by A5,REAL_1:67;
       then abs(a)<3*(r+abs(a)) by XCMPLX_1:8;
        then A6:a<3*(r+abs(a)) by A3,AXIOMS:22;
       reconsider p3=<*1 qua Real*> as Element of 1-tuples_on REAL;
       reconsider p4=p3 as Element of REAL 1 by EUCLID:def 1;
         (3*(r+abs(a)))*(1.REAL 1)=(3*(r+abs(a)))*p4 by Th36,EUCLID:def 11
       .=<*((3*(r+abs(a)))*1)*> by RVSUM_1:69;
   then A7: (3*(r+abs(a)))*(1.REAL 1) in W by A1,A6;
   then reconsider z1=(3*(r+abs(a)))*(1.REAL 1) as Point of Euclid 1;
       A8:a<=abs(a) by ABSVALUE:11;
         0+abs(a)<r+abs(a) by A2,REAL_1:53;
       then A9:a<r+abs(a) by A8,AXIOMS:22;
         (r+abs(a))*(1.REAL 1)=(r+abs(a))*p4 by Th36,EUCLID:def 11
       .=<*((r+abs(a))*1)*> by RVSUM_1:69;
   then A10: (r+abs(a))*(1.REAL 1) in W by A1,A9;
   then reconsider z2=(r+abs(a))*(1.REAL 1) as Point of Euclid 1;
         0<=abs(a) by ABSVALUE:5;
   then A11: r+0<=r+abs(a) by AXIOMS:24;
   then A12:    0<r+abs(a) by A2;
     dist(z1,z2)=|.(3*(r+abs(a)))*(1.REAL 1)-((r+abs(a))*(1.REAL 1)).|
                                       by JGRAPH_1:45
           .=|.(3*(r+abs(a))-(r+abs(a)))*(1.REAL 1).| by EUCLID:54
           .=|.((r+abs(a))+(r+abs(a))+(r+abs(a))-(r+abs(a)))*(1.REAL 1).|
                                    by XCMPLX_1:12
           .=|.((r+abs(a))+(r+abs(a)))*(1.REAL 1).|
                                    by XCMPLX_1:26
           .=abs((r+abs(a))+(r+abs(a)))*|.(1.REAL 1).| by TOPRNS_1:8
           .=abs((r+abs(a))+(r+abs(a)))*(sqrt 1) by Th37;
    then A13: (r+abs(a))+(r+abs(a))<= dist(z1,z2) by ABSVALUE:11,SQUARE_1:83;
      (r+abs(a))+0<(r+abs(a))+(r+abs(a)) by A12,REAL_1:53;
    then (r+abs(a))<dist(z1,z2) by A13,AXIOMS:22;
    then r<dist(z1,z2) by A11,AXIOMS:22;
  hence contradiction by A2,A7,A10;
 end;
 hence W is not bounded;
end;

theorem Th68: for W being Subset of Euclid 1,a being Real,
P being Subset of TOP-REAL 1
 st W={q where q is Point of TOP-REAL 1 : ex r st q=<*r*> & r < -a }
    & P=W holds P is connected & W is not bounded
proof let W be Subset of Euclid 1,a be Real,
  P be Subset of TOP-REAL 1; assume
 A1:W={q where q is Point of TOP-REAL 1 : ex r st q=<*r*> & r < -a } & P=W;
 hence P is connected by Th66;
 assume W is bounded;
   then consider r such that
   A2: 0<r & for x,y being Point of Euclid 1 st
      x in W & y in W holds dist(x,y)<=r by TBSP_1:def 9;
       A3:a<=abs(a) by ABSVALUE:11;
       A4:abs(a)>=0 by ABSVALUE:5;
       then abs(a)+abs(a)>=0+abs(a) by AXIOMS:24;
       then abs(a)+abs(a)+abs(a)>=0+abs(a) by A4,AXIOMS:24;
       then A5:3*abs(a)>=abs(a) by XCMPLX_1:12;
         3*r>0 by A2,REAL_2:122;
       then 0+abs(a)<3*r+3*abs(a) by A5,REAL_1:67;
       then abs(a)<3*(r+abs(a)) by XCMPLX_1:8;
        then a<3*(r+abs(a)) by A3,AXIOMS:22;
then A6:    -a> -(3*(r+abs(a))) by REAL_1:50;
       reconsider p3=<*1 qua Real*> as Element of 1-tuples_on REAL;
       reconsider p4=p3 as Element of REAL 1 by EUCLID:def 1;
         (-3*(r+abs(a)))*(1.REAL 1)=(-3*(r+abs(a)))*p4 by Th36,EUCLID:def 11
       .=<*((-3*(r+abs(a)))*1)*> by RVSUM_1:69;
then A7:    (-3*(r+abs(a)))*(1.REAL 1) in {q where q is Point of TOP-REAL 1:
      ex r st q=<*r*> & r< -a } by A6;
   then reconsider z1=(-3*(r+abs(a)))*(1.REAL 1) as Point of Euclid 1 by A1;
       A8:a<=abs(a) by ABSVALUE:11;
         0+abs(a)<r+abs(a) by A2,REAL_1:53;
       then a<r+abs(a) by A8,AXIOMS:22;
       then A9: -a> -(r+abs(a)) by REAL_1:50;
         (-(r+abs(a)))*(1.REAL 1)=(-(r+abs(a)))*p4 by Th36,EUCLID:def 11
       .=<* (-(r+abs(a)))*1 *> by RVSUM_1:69;
   then A10: (-(r+abs(a)))*(1.REAL 1) in W by A1,A9;
   then reconsider z2=(-(r+abs(a)))*(1.REAL 1) as Point of Euclid 1;
         0<=abs(a) by ABSVALUE:5;
   then A11: r+0<=r+abs(a) by AXIOMS:24;
   then A12:    0<r+abs(a) by A2;
     dist(z1,z2)=|.(-3*(r+abs(a)))*(1.REAL 1)-(-(r+abs(a)))*(1.REAL 1).|
                                       by JGRAPH_1:45
           .=|.(-3*(r+abs(a))--(r+abs(a)))*(1.REAL 1).| by EUCLID:54
           .=|.-((-3*(r+abs(a))--(r+abs(a)))*(1.REAL 1)).| by TOPRNS_1:27
           .=|.(-(-3*(r+abs(a))--(r+abs(a))))*(1.REAL 1).| by EUCLID:44
           .=|.(-(-3*(r+abs(a))+--(r+abs(a))))*(1.REAL 1).| by XCMPLX_0:def 8
           .=|.(-(-3*(r+abs(a)))-(r+abs(a)))*(1.REAL 1).| by XCMPLX_1:161
           .=|.((r+abs(a))+(r+abs(a))+(r+abs(a))-(r+abs(a)))*(1.REAL 1).|
                                    by XCMPLX_1:12
           .=|.((r+abs(a))+(r+abs(a)))*(1.REAL 1).| by XCMPLX_1:26
           .=abs((r+abs(a))+(r+abs(a)))*|.(1.REAL 1).| by TOPRNS_1:8
           .=abs((r+abs(a))+(r+abs(a)))*(sqrt 1) by Th37;
    then A13: (r+abs(a))+(r+abs(a))<= dist(z1,z2) by ABSVALUE:11,SQUARE_1:83;
      (r+abs(a))+0<(r+abs(a))+(r+abs(a)) by A12,REAL_1:53;
    then (r+abs(a))<dist(z1,z2) by A13,AXIOMS:22;
    then r<dist(z1,z2) by A11,AXIOMS:22;
  hence contradiction by A1,A2,A7,A10;
end;

theorem Th69: for W being Subset of Euclid n,a being Real,
P being Subset of TOP-REAL n st n>=2 & W={q : |.q.| > a }
    & P=W holds P is connected & W is not bounded
proof let W be Subset of Euclid n,a be Real,
  P be Subset of TOP-REAL n;
 assume A1:n>=2 & W={q : |.q.| > a } & P=W;
 hence P is connected by Th59;
 assume W is bounded;
   then consider r such that
   A2: 0<r & for x,y being Point of Euclid n st
      x in W & y in W holds dist(x,y)<=r by TBSP_1:def 9;
  A3: |.(r+abs(a))*(1.REAL n).|=abs(r+abs(a))*|.(1.REAL n).| by TOPRNS_1:8
        .=abs(r+abs(a))*(sqrt n) by Th37;
        A4:abs(r+abs(a))>=0 by ABSVALUE:5;
A5:     1<=n by A1,AXIOMS:22;
then A6:     1<=sqrt n by SQUARE_1:83,94;
then A7:    abs(r+abs(a))*1<=abs(r+abs(a))*sqrt n by A4,AXIOMS:25;
         (r+abs(a))<=abs(r+abs(a)) by ABSVALUE:11;
       then A8:(r+abs(a))<= |.(r+abs(a))*(1.REAL n).| by A3,A7,AXIOMS:22;
       A9:a<=abs(a) by ABSVALUE:11;
         abs(a)<r+abs(a) by A2,REAL_1:69;
then A10:     a<r+abs(a) by A9,AXIOMS:22;
       then a<|.(r+abs(a))*(1.REAL n).| by A8,AXIOMS:22;
   then A11: (r+abs(a))*(1.REAL n) in W by A1;
   then reconsider z1=(r+abs(a))*(1.REAL n) as Point of Euclid n;
        A12: |.-((r+abs(a))*(1.REAL n)).|
         = |.((r+abs(a))*(1.REAL n)).| by TOPRNS_1:27
        .=abs(r+abs a)*|.(1.REAL n).| by TOPRNS_1:8
        .=abs(r+abs a)*(sqrt n) by Th37;
        A13:abs(r+abs a)>=0 by ABSVALUE:5;
         1<=sqrt n by A5,SQUARE_1:83,94;
then A14:    abs(r+abs a)*1<=abs(r+abs a)*sqrt n by A13,AXIOMS:25;
         (r+abs(a))<=abs(r+abs a) by ABSVALUE:11;
       then (r+abs(a))<= |.-((r+abs(a))*(1.REAL n)).| by A12,A14,AXIOMS:22;
       then a<|.-((r+abs(a))*(1.REAL n)).| by A10,AXIOMS:22;
   then A15:-((r+abs(a))*(1.REAL n)) in W by A1;
   then reconsider z2=-((r+abs(a))*(1.REAL n)) as Point of Euclid n;
         0<=abs(a) by ABSVALUE:5;
   then A16: r+0<=r+abs(a) by AXIOMS:24;
   then A17:    0<r+abs(a) by A2;
   A18:dist(z1,z2)=|.(r+abs(a))*(1.REAL n)--((r+abs(a))*(1.REAL n)).|
                                       by JGRAPH_1:45
           .=|.(r+abs(a))*(1.REAL n)+--((r+abs(a))*(1.REAL n)).| by EUCLID:45
           .=|.(r+abs(a))*(1.REAL n)+((r+abs(a))*(1.REAL n)).| by EUCLID:39
           .=|.((r+abs(a))+(r+abs(a)))*(1.REAL n).| by EUCLID:37
           .=abs(((r+abs(a))+(r+abs(a))))*|.(1.REAL n).| by TOPRNS_1:8
           .=abs((r+abs(a))+(r+abs(a)))*(sqrt n) by Th37;
      abs((r+abs(a))+(r+abs(a)))>=0 by ABSVALUE:5;
then A19: abs((r+abs(a))+(r+abs(a)))*1<=abs((r+abs(a))+(r+abs(a)))*sqrt n
                                       by A6,AXIOMS:25;
      (r+abs(a))+(r+abs(a))<=abs((r+abs(a))+(r+abs(a))) by ABSVALUE:11;
    then A20: (r+abs(a))+(r+abs(a))<= dist(z1,z2) by A18,A19,AXIOMS:22;
      (r+abs(a))+0<(r+abs(a))+(r+abs(a)) by A17,REAL_1:53;
    then (r+abs(a))<dist(z1,z2) by A20,AXIOMS:22;
    then r<dist(z1,z2) by A16,AXIOMS:22;
  hence contradiction by A2,A11,A15;
end;

theorem Th70: for W being Subset of Euclid n,a being Real,
P being Subset of TOP-REAL n st n>=2 & W=(REAL n)\ {q : (|.q.|) < a }
    & P=W holds P is connected & W is not bounded
proof let W be Subset of Euclid n,a be Real,
  P be Subset of TOP-REAL n;
 assume A1:n>=2 & W=(REAL n)\ {q : (|.q.|) < a } & P=W;
 hence P is connected by Th61;
 n>=1 by A1,AXIOMS:22;
then A2:1<=sqrt n by SQUARE_1:83,94;
 assume W is bounded;
   then consider r such that
   A3: 0<r & for x,y being Point of Euclid n st
      x in W & y in W holds dist(x,y)<=r by TBSP_1:def 9;
        (r+abs(a))*(1.REAL n) in the carrier of TOP-REAL n;
      then A4:(r+abs(a))*(1.REAL n) in REAL n by EUCLID:25;
        now assume (r+abs(a))*(1.REAL n) in {q : (|.q.|) < a };
        then consider q being Point of TOP-REAL n such that
        A5:q=(r+abs(a))*(1.REAL n) & (|.q.|) < a;
        A6: |.(r+abs(a))*(1.REAL n).|=abs(r+abs a)*|.(1.REAL n).|
                                        by TOPRNS_1:8
        .=abs(r+abs a)*(sqrt n) by Th37;
          abs(r+abs a)>=0 by ABSVALUE:5;
   then A7: abs(r+abs a)*1<=abs(r+abs a)*sqrt n by A2,AXIOMS:25;
         (r+abs(a))<=abs(r+abs a) by ABSVALUE:11;
       then A8:(r+abs(a))<= |.(r+abs(a))*(1.REAL n).| by A6,A7,AXIOMS:22;
       A9:a<=abs(a) by ABSVALUE:11;
         abs(a)<r+abs(a) by A3,REAL_1:69;
       then a<r+abs(a) by A9,AXIOMS:22;
      hence contradiction by A5,A8,AXIOMS:22;
     end;
   then A10: (r+abs(a))*(1.REAL n) in W by A1,A4,XBOOLE_0:def 4;
   then reconsider z1=(r+abs(a))*(1.REAL n) as Point of Euclid n;
        -((r+abs(a))*(1.REAL n)) in the carrier of TOP-REAL n;
      then A11:-((r+abs(a))*(1.REAL n)) in REAL n by EUCLID:25;
        now assume -((r+abs(a))*(1.REAL n)) in {q : (|.q.|) < a };
        then consider q being Point of TOP-REAL n such that
        A12:q=-((r+abs(a))*(1.REAL n)) & (|.q.|) < a;
        A13: |.-((r+abs(a))*(1.REAL n)).|
         = |.((r+abs(a))*(1.REAL n)).| by TOPRNS_1:27
        .=abs(r+abs a)*|.(1.REAL n).| by TOPRNS_1:8
        .=abs(r+abs a)*(sqrt n) by Th37;
          abs(r+abs a)>=0 by ABSVALUE:5;
then A14:    abs(r+abs a)*1<=abs(r+abs a)*sqrt n by A2,AXIOMS:25;
         (r+abs(a))<=abs(r+abs a) by ABSVALUE:11;
       then A15:(r+abs(a))<= |.-((r+abs(a))*(1.REAL n)).| by A13,A14,AXIOMS:22;
       A16:a<=abs(a) by ABSVALUE:11;
         abs(a)<r+abs(a) by A3,REAL_1:69;
       then a<r+abs(a) by A16,AXIOMS:22;
      hence contradiction by A12,A15,AXIOMS:22;
     end;
   then A17:-((r+abs(a))*(1.REAL n)) in W by A1,A11,XBOOLE_0:def 4;
   then reconsider z2=-((r+abs(a))*(1.REAL n)) as Point of Euclid n;
         0<=abs(a) by ABSVALUE:5;
   then A18: r+0<=r+abs(a) by AXIOMS:24;
   then A19:    0<r+abs(a) by A3;
   A20:dist(z1,z2)=|.(r+abs(a))*(1.REAL n)--((r+abs(a))*(1.REAL n)).|
                                       by JGRAPH_1:45
           .=|.(r+abs(a))*(1.REAL n)+--((r+abs(a))*(1.REAL n)).| by EUCLID:45
           .=|.(r+abs(a))*(1.REAL n)+((r+abs(a))*(1.REAL n)).| by EUCLID:39
           .=|.((r+abs(a))+(r+abs(a)))*(1.REAL n).| by EUCLID:37
           .=abs(((r+abs(a))+(r+abs(a))))*|.(1.REAL n).| by TOPRNS_1:8
           .=abs(((r+abs(a))+(r+abs(a))))*(sqrt n) by Th37;
      abs(((r+abs(a))+(r+abs(a))))>=0 by ABSVALUE:5;
then A21: abs((r+abs(a))+(r+abs(a)))*1<=abs((r+abs(a))+(r+abs(a)))*sqrt n
                                       by A2,AXIOMS:25;
          (r+abs(a))+(r+abs(a))<=abs((r+abs(a))+(r+abs(a))) by ABSVALUE:11;
    then A22: (r+abs(a))+(r+abs(a))<= dist(z1,z2) by A20,A21,AXIOMS:22;
      (r+abs(a))+0<(r+abs(a))+(r+abs(a)) by A19,REAL_1:53;
    then (r+abs(a))<dist(z1,z2) by A22,AXIOMS:22;
    then r<dist(z1,z2) by A18,AXIOMS:22;
  hence contradiction by A3,A10,A17;
end;

theorem Th71: for P, P1 being Subset of TOP-REAL n,
Q being Subset of TOP-REAL n,
W being Subset of Euclid n st P=W & P is connected
& W is not bounded & P1=skl (Down(P,Q`)) & W misses Q holds
P1 is_outside_component_of Q
proof let P,P1 be Subset of TOP-REAL n,
  Q be Subset of TOP-REAL n,
  W be Subset of Euclid n;
  assume A1:P=W & P is connected
   & W is not bounded & P1=skl (Down(P,Q`)) & W misses Q;
then A2:W <> {}Euclid n by TBSP_1:14;
A3: W /\ Q = {} by A1,XBOOLE_0:def 7;
     now assume Q`={};
     then Q=({}(the carrier of TOP-REAL n))`;
     then Q=[#](the carrier of TOP-REAL n) by SUBSET_1:22;
     then Q=the carrier of TOP-REAL n by SUBSET_1:def 4;
    hence contradiction by A1,A2,A3,XBOOLE_1:28;
   end;
   then reconsider Q1=Q` as non empty Subset of TOP-REAL n;
A4:(TOP-REAL n)|Q1 is non empty;
   A5:(Down(P,Q`))=P /\ Q` by CONNSP_3:def 5 .=P \ Q by SUBSET_1:32
   .=P by A1,XBOOLE_1:83;
   then reconsider P0=P as Subset of (TOP-REAL n)|Q`;
     skl P0 is Subset of Euclid n by A1,A5,TOPREAL3:13;
   then reconsider W0=skl P0 as Subset of Euclid n;
       P0 c= Q` by A1,SUBSET_1:43;
     then A6:((TOP-REAL n)|Q`)|P0=(TOP-REAL n)|P by JORDAN6:47;
       (TOP-REAL n)|P is connected by A1,CONNSP_1:def 3;
     then A7:P0 is connected by A6,CONNSP_1:def 3;
   then A8:skl P0 is_a_component_of ((TOP-REAL n)|Q`) by A1,A2,A4,CONNSP_3:9;
     now assume
       for D being Subset of Euclid n st D=P1 holds D is bounded;
     then A9:W0 is bounded by A1,A5;
       W c= W0 by A1,A7,CONNSP_3:1;
    hence contradiction by A1,A9,TBSP_1:21;
   end;
 hence thesis by A1,A5,A8,Th18;
end;

theorem Th72: for A being Subset of Euclid n,
B being non empty Subset of Euclid n,
C being Subset of (Euclid n)|B st
A c= B & A=C & C is bounded holds A is bounded
proof let A be Subset of Euclid n,
  B be non empty Subset of Euclid n,
  C be Subset of (Euclid n)|B;
  assume A1:A c= B & A=C & C is bounded;
   then consider r0 being Real such that
   A2: 0<r0 & for x,y being Point of (Euclid n)|B
    st x in C & y in C holds dist(x,y)<=r0 by TBSP_1:def 9;
      for x,y being Point of (Euclid n) st x in A & y in A holds dist(x,y)<=r0
    proof let x,y be Point of (Euclid n);
      assume A3:x in A & y in A;
      then reconsider x0=x,y0=y as Point of (Euclid n)|B by A1;
      A4:dist(x0,y0)<=r0 by A1,A2,A3;
      A5:(the distance of ((Euclid n)|B)).(x0,y0)=
      (the distance of (Euclid n)).(x,y)
              by TOPMETR:def 1;
        (the distance of ((Euclid n)|B)).(x0,y0)=dist(x0,y0) by METRIC_1:def 1;
     hence dist(x,y)<=r0 by A4,A5,METRIC_1:def 1;
    end;
 hence thesis by A2,TBSP_1:def 9;
end;

theorem Th73: for A being Subset of TOP-REAL n st
A is compact holds A is Bounded
proof let A be Subset of TOP-REAL n;
 assume A1: A is compact;
     A c= the carrier of ((TOP-REAL n)|(A)) by JORDAN1:1;
   then reconsider A2=A as Subset of ((TOP-REAL n)|(A));
   per cases;
   suppose A2:A<>{};
     [#]((TOP-REAL n)|(A))=A2 by PRE_TOPC:def 10;
   then [#]((TOP-REAL n)|(A)) is compact by A1,COMPTS_1:11;
   then A3:(TOP-REAL n)|(A) is compact by COMPTS_1:10;
     A is non empty Subset of Euclid n by A2,TOPREAL3:13;
   then reconsider A1=A as non empty Subset of Euclid n;
     TopSpaceMetr((Euclid n)|A1)=(TOP-REAL n)|(A) by TOPMETR:20;
   then (Euclid n)|A1 is totally_bounded by A3,TBSP_1:12;
   then (Euclid n)|A1 is bounded by TBSP_1:26;
   then A4:[#]((Euclid n)|A1) is bounded by TBSP_1:25;
     [#]((Euclid n)|A1) =the carrier of (Euclid n)|A1 by PRE_TOPC:12
                   .=A1 by TOPMETR:def 2;
  then A1 is bounded by A4,Th72;
 hence thesis by Def2;
 suppose A={};
then A5: A = {}(Euclid n);
     A is Subset of Euclid n by TOPREAL3:13;
   then reconsider A1=A as Subset of Euclid n;
     A1 is bounded by A5,TBSP_1:14;
 hence thesis by Def2;
end;

theorem Th74: for A being Subset of TOP-REAL n st 1<=n & A is Bounded holds
 A` <> {}
proof let A be Subset of TOP-REAL n;
 assume A1:1<=n & A is Bounded;
  then consider r being Real such that
  A2: for q being Point of TOP-REAL n st q in A holds |.q.|<r by Th40;
  A3: |.r*(1.REAL n).|=abs(r)*|.1.REAL n.| by TOPRNS_1:8;
  A4:1<=|.1.REAL n.| by A1,Th38;
    abs(r)>=0 by ABSVALUE:5;
  then A5:abs(r)*|.1.REAL n.|>=abs(r)*1 by A4,AXIOMS:25;
    r<=abs(r) by ABSVALUE:11;
  then r<=|.r*(1.REAL n).| by A3,A5,AXIOMS:22;
  then not r*(1.REAL n) in A by A2;
  then r*(1.REAL n) in (the carrier of TOP-REAL n)\A by XBOOLE_0:def 4;
 hence A` <>{} by SUBSET_1:def 5;
end;

theorem Th75: for r being Real holds
(ex B being Subset of Euclid n st B={q : (|.q.|) < r }) &
for A being Subset of Euclid n st A={q1 : (|.q1.|) < r }
  holds A is bounded
proof let r be Real;
  A1:{q : (|.q.|) < r } c= the carrier of Euclid n
  proof let x be set;assume x in {q : (|.q.|) < r };
    then consider q being Point of TOP-REAL n such that
    A2:q=x & (|.q.|) < r;
      x in the carrier of TOP-REAL n by A2;
   hence x in the carrier of Euclid n by TOPREAL3:13;
  end;
 hence ex B being Subset of Euclid n
 st B={q : (|.q.|) < r };
  let A be Subset of Euclid n;
    assume A3:A={q1 : (|.q1.|) < r };
        {q1 : (|.q1.|) < r }
           is Subset of TOP-REAL n by A1,TOPREAL3:13;
      then reconsider C={q1 : (|.q1.|) < r }
           as Subset of TOP-REAL n;
        for q being Point of TOP-REAL n st
           q in C holds |.q.|<r
       proof let q be Point of TOP-REAL n;
         assume q in C;
          then consider q1 being Point of TOP-REAL n such that
          A4:q1=q & (|.q1.|) < r;
        thus |.q.|<r by A4;
       end;
     then C is Bounded by Th40;
     then ex A1 being Subset of Euclid n st A1=C & A1 is bounded by Def2;
    hence A is bounded by A3;
end;

theorem Th76:
for A being Subset of TOP-REAL n st n>=2 & A is Bounded
 ex B being Subset of TOP-REAL n st B is_outside_component_of A & B=UBD A
proof let A be Subset of TOP-REAL n;
  assume A1: n>=2 & A is Bounded;
  then consider C being Subset of Euclid n
   such that A2:C=A & C is bounded by Def2;
   A3:n>=1 by A1,AXIOMS:22;
  per cases;
  suppose A4: C<>{};
    consider x0 being Element of C;
    A5:x0 in C by A4;
    then reconsider q1=x0 as Point of TOP-REAL n by A2;
    reconsider x0 as Point of Euclid n by A5;
    consider r being Real such that
    A6: 0<r & for x,y being Point of (Euclid n) st x in C & y in C holds
                               dist(x,y) <= r by A2,TBSP_1:def 9;
    reconsider o=0.REAL n as Point of Euclid n by TOPREAL3:13;
    set R0=r+dist(o,x0)+1;
      (REAL n)\{q where q is Point of TOP-REAL n:
    (|.q.|) < R0 } c= REAL n by XBOOLE_1:36;
    then (REAL n)\{q where q is Point of TOP-REAL n:
    (|.q.|) < R0 } c= the carrier of TOP-REAL n by EUCLID:25;
    then (REAL n)\{q where q is Point of TOP-REAL n:
    (|.q.|) < R0 } is Subset of Euclid n by TOPREAL3:13;
    then reconsider W=(REAL n)\{q where q is Point of TOP-REAL n:
    (|.q.|) < R0 } as Subset of Euclid n;
    reconsider P=W as Subset of TOP-REAL n by TOPREAL3:13;
    reconsider P as Subset of TOP-REAL n;
          W=(REAL n)\ {q : (|.q.|) < R0 } & P=W;
    then A7:P is connected & W is not bounded by A1,Th70;
      the carrier of (TOP-REAL n)|A`=A` by JORDAN1:1;
    then skl (Down(P,A`)) is Subset of TOP-REAL n
      by XBOOLE_1:1;
    then reconsider P1=skl (Down(P,A`)) as Subset of TOP-REAL n;
    A8:now assume W meets A;
      then consider z being set such that
      A9:z in W & z in A by XBOOLE_0:3;
A10:  not z in {q where q is Point of TOP-REAL n:(|.q.|) < R0 }
                                             by A9,XBOOLE_0:def 4;
      reconsider z as Point of Euclid n by A9;
      reconsider q1=z as Point of TOP-REAL n by TOPREAL3:13;
      A11:(|.q1.|) >= r+dist(o,x0)+1 by A10;
        |.q1-(0.REAL n).|=dist(o,z) by JGRAPH_1:45;
      then A12:dist(o,z)>=r+dist(o,x0)+1 by A11,Th13;
      A13:dist(x0,z)<=r by A2,A6,A9;
      A14:dist(o,z)<=dist(o,x0)+dist(x0,z) by METRIC_1:4;
        dist(o,x0)+dist(x0,z)<=dist(o,x0)+r by A13,AXIOMS:24;
      then dist(o,z)<=dist(o,x0)+r by A14,AXIOMS:22;
      then r+dist(o,x0)+1<=dist(o,x0)+r by A12,AXIOMS:22;
      then r+dist(o,x0)+1-(r+dist(o,x0))<=r+dist(o,x0)-(r+dist(o,x0))
                             by REAL_1:49;
      then 1<=r+dist(o,x0)-(r+dist(o,x0)) by XCMPLX_1:26;
      then 1<=0 by XCMPLX_1:14;
     hence contradiction;
    end;
    then A15:P1 is_outside_component_of A by A7,Th71;
    then A16:P1 c= UBD A by Th27;
      UBD A c= P1
    proof let z be set;assume z in UBD A;
      then z in union{B where B is Subset of TOP-REAL n:
          B is_outside_component_of A} by Def6;
      then consider y being set such that
      A17:z in y & y in {B where B is Subset of TOP-REAL n:
          B is_outside_component_of A} by TARSKI:def 4;
      consider B being Subset of TOP-REAL n such that
      A18: y=B & B is_outside_component_of A by A17;
     consider C2 being Subset of ((TOP-REAL n)|(A`)) such that
     A19: C2=B & C2 is_a_component_of ((TOP-REAL n)|(A`)) &
       C2 is not bounded Subset of Euclid n by A18,Th18;

         consider D2 being Subset of Euclid n such that
         A20: D2={q : |.q.| < R0 } by Th75;
         reconsider D2 as Subset of Euclid n;
         A21: W misses D2 by A20,XBOOLE_1:79;
      A22:C2 is connected by A19,CONNSP_1:def 5;
      A23:now assume A24:W /\ C2={};
        A25:C2 c= {q : (|.q.|) < R0 }
         proof let x8 be set;assume A26:x8 in C2;
          then x8 in the carrier of TOP-REAL n by A19;
          then A27:x8 in REAL n by EUCLID:25;
          assume not x8 in {q : (|.q.|) < R0 };
            then x8 in W by A27,XBOOLE_0:def 4;
           hence contradiction by A24,A26,XBOOLE_0:def 3;
         end;
           C2 is Subset of Euclid n by A19,TOPREAL3:13;
         then reconsider D=C2 as Subset of Euclid n;
         A28: not D is bounded by A19;
           D2 is bounded by A20,Th75;
       hence contradiction by A20,A25,A28,TBSP_1:21;
      end;
      A29: A c= D2
        proof let z be set;assume A30:z in A;
         then reconsider q2=z as Point of TOP-REAL n;
         reconsider x1=z as Point of Euclid n by A2,A30;
         A31: |.q2.|=|.q2-q1+q1.| by EUCLID:52;
         A32: |.q2-q1+q1.|<=|.q2-q1.|+|.q1.| by TOPRNS_1:30;
         A33: |.q2-q1.|=dist(x1,x0) by JGRAPH_1:45;
         A34:dist(x1,x0)<=r by A2,A6,A30;
           |.q1.|=|.q1-0.REAL n.| by Th13
         .=dist(x0,o) by JGRAPH_1:45;
         then |.q2-q1.|+|.q1.|<=r+dist(o,x0) by A33,A34,AXIOMS:24;
         then A35: |.q2.|<=r+dist(o,x0) by A31,A32,AXIOMS:22;
           r+dist(o,x0)<r+dist(o,x0)+1 by REAL_1:69;
         then |.q2.|<r+dist(o,x0)+1 by A35,AXIOMS:22;
         hence z in D2 by A20;
        end;
         the carrier of Euclid n=the carrier of TOP-REAL n by TOPREAL3:13;
       then (the carrier of Euclid n)\D2 c= (the carrier of TOP-REAL n)\A
                           by A29,XBOOLE_1:34;
       then D2` c= (the carrier of TOP-REAL n)\A by SUBSET_1:def 5;
       then A36: D2` c= A` by SUBSET_1:def 5;
  A37:  W c= D2` by A21,SUBSET_1:43;
       A38:Down(P,A`)=P /\ A` by CONNSP_3:def 5;
       A39:P c= A` by A8,SUBSET_1:43;
       then A40:Down(P,A`)=P by A38,XBOOLE_1:28;
         (TOP-REAL n)|P is connected by A7,CONNSP_1:def 3;
       then ((TOP-REAL n)|A`)|Down(P,A`) is connected
                     by A39,A40,JORDAN6:47;
       then A41:Down(P,A`) is connected by CONNSP_1:def 3;
A42:   W /\ D2` /\ C2 <>{} by A23,A37,XBOOLE_1:28;
         P /\ D2` c= (P /\ A`) by A36,XBOOLE_1:26;
       then P /\ D2` /\ C2 c= (P /\ A`) /\ C2 by XBOOLE_1:26;
      then (P /\ A`) /\ C2 <>{} by A42,XBOOLE_1:3;
      then Down(P,A`)/\ C2 <>{} by CONNSP_3:def 5;
      then A43:Down(P,A`) meets C2 by XBOOLE_0:def 7;
    reconsider G=A` as non empty Subset of TOP-REAL n
      by A1,A3,Th74;
     (TOP-REAL n)|G is non empty TopSpace;
      then C2 c= skl (Down(P,A`)) by A22,A41,A43,CONNSP_3:16;
     hence z in P1 by A17,A18,A19;
    end;
    then P1=UBD A by A16,XBOOLE_0:def 10;
   hence ex B being Subset of TOP-REAL n st
    B is_outside_component_of A & B=UBD A by A15;
  suppose A44:C={};
    then A45:A={}(TOP-REAL n) by A2;
      REAL n = the carrier of TOP-REAL n by EUCLID:25;
    then REAL n c= the carrier of Euclid n by TOPREAL3:13;
    then reconsider W=REAL n as Subset of Euclid n;
    reconsider P=W as Subset of TOP-REAL n by TOPREAL3:13;
    reconsider P as Subset of TOP-REAL n;
    A46:P is connected by Th33;
      1<=n by A1,AXIOMS:22;
    then A47: W is not bounded by Th39;
      W /\ A={} by A2,A44;
    then A48:W misses A by XBOOLE_0:def 7;
      the carrier of (TOP-REAL n)|A`=A` by JORDAN1:1;
    then skl Down(P,A`) is Subset of TOP-REAL n by XBOOLE_1:1;
    then reconsider P1=skl Down(P,A`) as Subset of TOP-REAL n;
    A49:P1 is_outside_component_of A by A46,A47,A48,Th71;
    A50:UBD A=REAL n by A3,A45,Th42;
    A51:[#](TOP-REAL n)=the carrier of TOP-REAL n by PRE_TOPC:12 .=REAL n
                                  by EUCLID:25;
    A52: [#]((TOP-REAL n)| [#](TOP-REAL n)) = [#](TOP-REAL n) by TSEP_1:3;
      (TOP-REAL n)| [#](TOP-REAL n)=TOP-REAL n by TSEP_1:3;
    then A53:[#]((TOP-REAL n)| [#](TOP-REAL n)) is_a_component_of
    (TOP-REAL n)| [#](TOP-REAL n) by Th23;
      A`=(the carrier of TOP-REAL n) \ {} by A2,A44,SUBSET_1:def 5
    .=[#](TOP-REAL n) by PRE_TOPC:12;
    then P1=skl [#]((TOP-REAL n)| [#](TOP-REAL n)) by A51,A52,CONNSP_3:28
    .=[#]((TOP-REAL n)| [#](TOP-REAL n)) by A53,CONNSP_3:7
    .=the carrier of TOP-REAL n by A52,PRE_TOPC:12
    .=REAL n by EUCLID:25;
   hence thesis by A49,A50;
end;

theorem Th77: for a being Real, P being Subset of TOP-REAL n st
 P={q : (|.q.|) < a } holds P is convex
proof let a be Real,
 P be Subset of TOP-REAL n;
 assume A1: P={q : (|.q.|) < a };
    for p1,p2 being Point of TOP-REAL n st
  p1 in P & p2 in P holds LSeg(p1,p2) c= P
  proof let p1,p2 be Point of TOP-REAL n;
    assume A2:p1 in P & p2 in P;
     then consider q1 such that A3:q1=p1 & (|.q1.|) < a by A1;
     consider q2 such that A4:q2=p2 & (|.q2.|) < a by A1,A2;
       LSeg(p1,p2) c= P
     proof let x;assume A5:x in LSeg(p1,p2);
       then x in {(1-r)*p1+r*p2: 0<=r & r<=1} by TOPREAL1:def 4;
       then consider r such that A6:x=(1-r)*p1+r*p2 & 0<=r & r<=1;
       reconsider q=x as Point of TOP-REAL n by A5;
       A7: |.(1-r)*p1+r*p2.|<=|.(1-r)*p1.|+|.r*p2.| by TOPRNS_1:30;
       A8: |.(1-r)*p1.|=abs(1-r)*|.p1.| by TOPRNS_1:8;
       A9:1-r>=0 by A6,SQUARE_1:12;
       then A10:abs(1-r)=1-r by ABSVALUE:def 1;
      per cases;
      suppose 1-r>0;
       then A11:abs(1-r)*|.p1.|<abs(1-r)*a by A3,A10,REAL_1:70;
       A12: |.r*p2.|=abs(r)*|.p2.| by TOPRNS_1:8;
         0<=abs(r) by ABSVALUE:5;
       then A13:abs(r)*|.p2.|<=abs(r)*a by A4,AXIOMS:25;
         r=abs(r) by A6,ABSVALUE:def 1;
       then A14: |.(1-r)*p1.|+|.r*p2.|<(1-r)*a+r*a by A8,A10,A11,A12,A13,REAL_1
:67;
         (1-r)*a+r*a=((1-r)+r)*a by XCMPLX_1:8 .=1 *a by XCMPLX_1:27 .=a;
       then (|.q.|)<a by A6,A7,A14,AXIOMS:22;
      hence x in P by A1;
      suppose 1-r<=0; then 1-r+r=0+r by A9,AXIOMS:21;
       then 1=r by XCMPLX_1:27;
       then A15:r>0;
       A16:abs(1-r)*|.p1.|<=abs(1-r)*a by A3,A9,A10,AXIOMS:25;
       A17: |.r*p2.|=abs(r)*|.p2.| by TOPRNS_1:8;
         0<abs(r) by A15,ABSVALUE:def 1;
       then A18:abs(r)*|.p2.|<abs(r)*a by A4,REAL_1:70;
         r=abs(r) by A6,ABSVALUE:def 1;
       then A19: |.(1-r)*p1.|+|.r*p2.|<(1-r)*a+r*a by A8,A10,A16,A17,A18,REAL_1
:67;
         (1-r)*a+r*a=((1-r)+r)*a by XCMPLX_1:8 .=1 *a by XCMPLX_1:27 .=a;
       then (|.q.|)<a by A6,A7,A19,AXIOMS:22;
      hence x in P by A1;
     end;
   hence LSeg(p1,p2) c= P;
  end;
 hence P is convex by JORDAN1:def 1;
end;

theorem Th78: for a being Real,P being Subset of TOP-REAL n st P=Ball(u,a)
 holds P is convex
proof let a be Real,
 P be Subset of TOP-REAL n;
 assume A1: P=Ball(u,a);
    for p1,p2 being Point of TOP-REAL n st p1 in P & p2 in P holds
    LSeg(p1,p2) c= P
  proof let p1,p2 be Point of TOP-REAL n;
    assume A2:p1 in P & p2 in P;
    A3:P={q where q is Element of Euclid n : dist(u,q) < a}
                         by A1,METRIC_1:18;
     then consider q1 being Point of Euclid n
     such that A4:q1=p1 & dist(u,q1) < a by A2;
     reconsider p=u as Point of TOP-REAL n by TOPREAL3:13;
     A5: |.p-p1.|<a by A4,JGRAPH_1:45;
     consider q2 being Point of Euclid n such that
     A6:q2=p2 & dist(u,q2) < a by A2,A3;
     A7: |.p-p2.|<a by A6,JGRAPH_1:45;
     A8:for p3 being Point of TOP-REAL n st |.p-p3.|<a holds p3 in P
      proof let p3 be Point of TOP-REAL n;
        assume A9: |.p-p3.|<a;
         reconsider u3=p3 as Point of Euclid n by TOPREAL3:13;
           dist(u,u3)<a by A9,JGRAPH_1:45;
        hence p3 in P by A3;
      end;
       LSeg(p1,p2) c= P
     proof let x;assume A10:x in LSeg(p1,p2);
       then x in {(1-r)*p1+r*p2: 0<=r & r<=1} by TOPREAL1:def 4;
       then consider r such that A11:x=(1-r)*p1+r*p2 & 0<=r & r<=1;
       reconsider q=x as Point of TOP-REAL n by A10;
         (1-r)*p+r*p=((1-r)+r)*p by EUCLID:37 .=1 *p by XCMPLX_1:27
       .=p by EUCLID:33;
       then |.p-((1-r)*p1+r*p2).|=|.(1-r)*p+r*p-(1-r)*p1-r*p2.| by EUCLID:50
       .=|.(1-r)*p+r*p+-(1-r)*p1-r*p2.| by EUCLID:45
       .=|.(1-r)*p+r*p+-(1-r)*p1+-r*p2.| by EUCLID:45
       .=|.(1-r)*p+-(1-r)*p1+r*p+-r*p2.| by EUCLID:30
       .=|.(1-r)*p+-(1-r)*p1+(r*p+-r*p2).| by EUCLID:30
       .=|.(1-r)*p+(1-r)*(-p1)+(r*p+-r*p2).| by EUCLID:44
       .=|.(1-r)*(p+-p1)+(r*p+-r*p2).| by EUCLID:36
       .=|.(1-r)*(p-p1)+(r*p+-r*p2).| by EUCLID:45
       .=|.(1-r)*(p-p1)+(r*p+r*(-p2)).| by EUCLID:44
       .=|.(1-r)*(p-p1)+(r*(p+-p2)).| by EUCLID:36
       .=|.(1-r)*(p-p1)+r*(p-p2).| by EUCLID:45;
  then A12: |.p-((1-r)*p1+r*p2).|<=|.(1-r)*(p-p1).|+|.r*(p-p2).| by TOPRNS_1:30
;
       A13: |.(1-r)*(p-p1).|=abs(1-r)*|.(p-p1).| by TOPRNS_1:8;
       A14:1-r>=0 by A11,SQUARE_1:12;
       then A15:abs(1-r)=1-r by ABSVALUE:def 1;
      per cases;
      suppose 1-r>0;
       then A16:abs(1-r)*|.p-p1.|<abs(1-r)*a by A5,A15,REAL_1:70;
       A17: |.r*(p-p2).|=abs(r)*|.p-p2.| by TOPRNS_1:8;
         0<=abs(r) by ABSVALUE:5;
       then A18:abs(r)*|.p-p2.|<=abs(r)*a by A7,AXIOMS:25;
         r=abs(r) by A11,ABSVALUE:def 1;
       then A19: |.(1-r)*(p-p1).|+|.r*(p-p2).|<(1-r)*a+r*a
                          by A13,A15,A16,A17,A18,REAL_1:67;
         (1-r)*a+r*a=((1-r)+r)*a by XCMPLX_1:8 .=1 *a by XCMPLX_1:27 .=a;
       then (|.p-q.|)<a by A11,A12,A19,AXIOMS:22;
      hence x in P by A8;
      suppose 1-r<=0; then 1-r+r=0+r by A14,AXIOMS:21;
       then 1=r by XCMPLX_1:27;
       then A20:r>0;
       A21:abs(1-r)*|.p-p1.|<=abs(1-r)*a by A5,A14,A15,AXIOMS:25;
       A22: |.r*(p-p2).|=abs(r)*|.p-p2.| by TOPRNS_1:8;
         0<abs(r) by A20,ABSVALUE:def 1;
       then A23:abs(r)*|.p-p2.|<abs(r)*a by A7,REAL_1:70;
         r=abs(r) by A11,ABSVALUE:def 1;
       then A24: |.(1-r)*(p-p1).|+|.r*(p-p2).|
       <(1-r)*a+r*a by A13,A15,A21,A22,A23,REAL_1:67;
         (1-r)*a+r*a=((1-r)+r)*a by XCMPLX_1:8 .=1 *a by XCMPLX_1:27 .=a;
       then (|.p-q.|)<a by A11,A12,A24,AXIOMS:22;
      hence x in P by A8;
     end;
   hence LSeg(p1,p2) c= P;
  end;
 hence P is convex by JORDAN1:def 1;
end;

theorem Th79: for a being Real,P being Subset of TOP-REAL n
 st P={q : |.q.| < a } holds P is connected
proof let a be Real,
  P be Subset of TOP-REAL n;
  assume P={q : |.q.| < a };
  then P is convex by Th77;
 hence P is connected by Th14;
end;

reserve R for Subset of TOP-REAL n;
reserve P for Subset of TOP-REAL n;

theorem Th80:
p <> q & p in Ball(u,r) & q in Ball(u,r) implies
ex h being map of I[01],TOP-REAL n st h is continuous & h.0=p & h.1=q &
      rng h c= Ball(u,r)
 proof assume
  A1: p<>q & p in Ball(u,r) & q in Ball(u,r);
       reconsider P=LSeg(p,q) as Subset of TOP-REAL n;
         Ball(u,r) is Subset of TOP-REAL n
        by TOPREAL3:13;
       then reconsider Q=Ball(u,r) as Subset of TOP-REAL n;
         Q is convex by Th78;
       then A2:LSeg(p,q) c= Ball(u,r) by A1,JORDAN1:def 1;
         LSeg(p,q) is_an_arc_of p,q by A1,TOPREAL1:15;
       then consider f being map of I[01], (TOP-REAL n)|P such that
       A3:f is_homeomorphism & f.0 = p & f.1 = q by TOPREAL1:def 2;
       A4:dom f = [#]I[01] & rng f = [#]((TOP-REAL n)|P) & f is one-to-one &
        f is continuous & f" is continuous by A3,TOPS_2:def 5;
       reconsider h=f as map of I[01],TOP-REAL n by JORDAN6:4;
       take h;
       thus thesis by A2,A3,A4,JORDAN6:4,PRE_TOPC:def 10;
      end;

theorem Th81: for f being map of I[01],TOP-REAL n st
      f is continuous & f.0=p1 & f.1=p2 &
 p in Ball(u,r) & p2 in Ball(u,r)
ex h being map of I[01],TOP-REAL n st h is continuous & h.0=p1 & h.1=p &
      rng h c= rng f \/ Ball(u,r)
proof let f be map of I[01],TOP-REAL n;
assume A1:f is continuous & f.0=p1 & f.1=p2 &
p in Ball(u,r) & p2 in Ball(u,r);
  then A2: LSeg(p2,p) c= Ball(u,r) by TOPREAL3:28;
   per cases;
   suppose p2<>p;
    then LSeg(p2,p) is_an_arc_of p2,p by TOPREAL1:15;
    then consider f1 being map of I[01], (TOP-REAL n)|LSeg(p2,p) such that
    A3: f1 is_homeomorphism & f1.0 = p2 & f1.1 = p by TOPREAL1:def 2;
    A4:dom f1 = [#]I[01] & rng f1 = [#]((TOP-REAL n)|LSeg(p2,p))
    & f1 is one-to-one &
     f1 is continuous & f1" is continuous by A3,TOPS_2:def 5;
     reconsider f2=f1 as map of I[01],TOP-REAL n by JORDAN6:4;
     A5:f2 is continuous by A4,JORDAN6:4;
     A6:rng f2=LSeg(p2,p) by A4,PRE_TOPC:def 10;
    consider h3 being map of I[01],(TOP-REAL n) such that
    A7: h3 is continuous & p1=h3.0 & p=h3.1 &
    rng h3 c= rng f \/ rng f2 by A1,A3,A5,Th32;
      rng f \/ rng f2 c= rng f \/ Ball(u,r) by A2,A6,XBOOLE_1:9;
    then rng h3 c= rng f \/ Ball(u,r) by A7,XBOOLE_1:1;
    hence ex h being map of I[01],TOP-REAL n st h is continuous
    & h.0=p1 & h.1=p & rng h c= rng f \/ Ball(u,r) by A7;
  suppose A8:p2=p; rng f c= rng f \/ Ball(u,r) by XBOOLE_1:7;
    hence ex h being map of I[01],TOP-REAL n st h is continuous
      & h.0=p1 & h.1=p & rng h c= rng f \/ Ball(u,r) by A1,A8;
end;

theorem Th82: for f being map of I[01],TOP-REAL n st
   f is continuous & rng f c=P & f.0=p1 & f.1=p2 &
      p in Ball(u,r) & p2 in Ball(u,r) & Ball(u,r) c= P
      ex f1 being map of I[01],TOP-REAL n st
      f1 is continuous & rng f1 c= P & f1.0=p1 & f1.1=p
 proof let f be map of I[01],TOP-REAL n;
 assume that
  A1: f is continuous and
  A2: rng f c= P and
  A3: f.0=p1 & f.1=p2 and
  A4: p in Ball(u,r) and
  A5: p2 in Ball(u,r) and
  A6: Ball(u,r) c= P;
       consider f3 being map of I[01],TOP-REAL n such that
       A7:f3 is continuous & f3.0=p1 & f3.1=p
          & rng f3 c= rng f \/ Ball(u,r) by A1,A3,A4,A5,Th81;
         rng f \/ Ball(u,r) c= P by A2,A6,XBOOLE_1:8;
       then rng f3 c= P & f3.0=p1 & f3.1=p by A7,XBOOLE_1:1;
      hence ex f5 being map of I[01],TOP-REAL n st
      f5 is continuous & rng f5 c= P & f5.0=p1 & f5.1=p by A7;
 end;

theorem Th83:
for p for P being Subset of TOP-REAL n st R is connected & R is open &
P = {q: q<>p & q in R & not
      ex f being map of I[01],TOP-REAL n st f is continuous & rng f c= R &
                f.0=p & f.1=q}
  holds P is open
 proof let p; let P be Subset of TOP-REAL n; assume that
  A1: R is connected & R is open and
  A2: P = {q: q<>p & q in R & not
      ex f being map of I[01],TOP-REAL n st f is continuous & rng f c= R &
                f.0=p & f.1=q};
A3:the TopStruct of TOP-REAL n = the TopStruct of TopSpaceMetr(Euclid n)
     by EUCLID:def 8;
  A4: P c= R
    proof let x;assume x in P; then consider q such that
     A5: q=x & ( q<>p & q in R & not
      ex f being map of I[01],TOP-REAL n st f is continuous & rng f c= R &
                f.0=p & f.1=q) by A2;
     thus x in R by A5;
    end;
    reconsider P'=P as Subset of TopSpaceMetr(Euclid n) by EUCLID:def 8;
     now let u;
    reconsider p2=u as Point of TOP-REAL n by TOPREAL3:13; assume
    A6: u in P;
    reconsider R'=R as Subset of TopSpaceMetr(Euclid n) by EUCLID:def 8;
    consider r be real number such that
    A7: r>0 & Ball(u,r) c= R' by A1,A3,A4,A6,TOPMETR:22;
    reconsider r'=r as Real by XREAL_0:def 1;
    A8: p2 in Ball(u,r') by A7,TBSP_1:16;
    take r;
    thus r>0 by A7;
       Ball(u,r) c= P'
     proof assume not thesis;
      then consider x such that
      A9: x in Ball(u,r) & not x in P by TARSKI:def 3;
        x in R by A7,A9;
      then reconsider q=x as Point of TOP-REAL n;
       per cases by A2,A7,A9;
       suppose A10: q=p;
          now assume A11: q=p2;
           ex p3 st p3=p2 & p3<>p & p3 in R & not
      ex f1 being map of I[01],TOP-REAL n st f1 is continuous & rng f1 c= R &
                f1.0=p & f1.1=p3 by A2,A6;
         hence contradiction by A10,A11;
        end;
       then q<>p2 & u in Ball(u,r') by A7,TBSP_1:16;
       then consider f2 being map of I[01],TOP-REAL n such that
       A12: f2 is continuous & f2.0=q & f2.1=p2
                  & rng f2 c= Ball(u,r') by A9,Th80;
       A13: f2 is continuous & rng f2 c= R &
                f2.0=p & f2.1=p2 by A7,A10,A12,XBOOLE_1:1;
         not p2 in P
        proof assume p2 in P;
         then ex q4 st q4=p2 & q4<>p & q4 in R & not
      ex f1 being map of I[01],TOP-REAL n st f1 is continuous & rng f1 c= R &
                f1.0=p & f1.1=q4 by A2;
         hence contradiction by A13;
        end;
       hence contradiction by A6;
      suppose A14:
      ex f1 being map of I[01],TOP-REAL n st f1 is continuous & rng f1 c= R &
                f1.0=p & f1.1=q;
         not p2 in P
        proof assume p2 in P;
         then ex q4 st q4=p2 & q4<>p & q4 in R & not
      ex f1 being map of I[01],TOP-REAL n st f1 is continuous & rng f1 c= R &
                f1.0=p & f1.1=q4 by A2;
         hence contradiction by A7,A8,A9,A14,Th82;
        end;
       hence contradiction by A6;
    end;
    hence Ball(u,r) c= P';
   end;
  hence thesis by A3,TOPMETR:22;
 end;

theorem Th84:
for P being Subset of TOP-REAL n st R is connected & R is open & p in R &
P = {q: q=p or
      ex f being map of I[01],TOP-REAL n st f is continuous & rng f c= R &
                f.0=p & f.1=q} holds
P is open
 proof let P be Subset of TOP-REAL n; assume that
  A1: R is connected & R is open and
  A2: p in R and
  A3: P = {q: q=p or
      ex f being map of I[01],TOP-REAL n st f is continuous & rng f c= R &
                f.0=p & f.1=q};
A4:the TopStruct of TOP-REAL n = the TopStruct of TopSpaceMetr(Euclid n)
     by EUCLID:def 8;
    reconsider P'=P as Subset of TopSpaceMetr(Euclid n) by EUCLID:def 8;
    now let u;
    reconsider p2=u as Point of TOP-REAL n by TOPREAL3:13;
    assume u in P';
    then consider q1 such that
    A5: q1=u and
    A6: q1=p or
      ex f being map of I[01],TOP-REAL n st f is continuous & rng f c= R &
                f.0=p & f.1=q1 by A3;
     now per cases by A6;
      suppose q1=p; hence p2 in R by A2,A5;
      suppose q1<>p &
      ex f being map of I[01],TOP-REAL n st f is continuous & rng f c= R &
                f.0=p & f.1=q1;
       then consider f2 being map of I[01],TOP-REAL n such that
         f2 is continuous and
       A7: rng f2 c= R & f2.0=p & f2.1=q1;
         dom f2=[.0 qua Real,1 qua Real.] by BORSUK_1:83,FUNCT_2:def 1;
       then 1 in dom f2 by TOPREAL5:1;
       then u in rng f2 by A5,A7,FUNCT_1:def 5;
       hence p2 in R by A7;
      end;
    then consider r be real number such that
    A8: r>0 & Ball(u,r) c= R by A1,A4,TOPMETR:22;
    reconsider r'=r as Real by XREAL_0:def 1;
    A9: p2 in Ball(u,r') by A8,TBSP_1:16;
    take r;
    thus r>0 by A8;
    thus Ball(u,r) c= P
     proof let x; assume
      A10: x in Ball(u,r);
     then reconsider q=x as Point of TOP-REAL n by A8,TARSKI:def 3;
     per cases;
       suppose q=p; hence x in P by A3;
       suppose A11:q<>p;
    A12: now assume
            q1<>p;
           then ex f being map of I[01],TOP-REAL n st f is continuous & rng f
c= R &
                f.0=p & f.1=q by A5,A6,A8,A9,A10,Th82;
          hence x in P by A3;
         end;
           now assume q1=p;
           then p in Ball(u,r') by A5,A8,TBSP_1:16;
           then consider f2 being map of I[01],TOP-REAL n such that
           A13:f2 is continuous & f2.0=p & f2.1=q &
           rng f2 c= Ball(u,r') by A10,A11,Th80;
             f2 is continuous & rng f2 c= R &
                f2.0=p & f2.1=q by A8,A13,XBOOLE_1:1;
           hence x in P by A3;
         end;
         hence x in P by A12;
      end;
   end;
  hence P is open by A4,TOPMETR:22;
 end;

theorem Th85:
for R being Subset of TOP-REAL n holds
p in R & P={q: q=p or
      ex f being map of I[01],TOP-REAL n st f is continuous & rng f c= R &
                f.0=p & f.1=q} implies P c= R
 proof
 let R be Subset of TOP-REAL n;
 assume that
  A1: p in R and
  A2: P = {q: q=p or
      ex f being map of I[01],TOP-REAL n st f is continuous & rng f c= R &
                f.0=p & f.1=q};
  let x; assume x in P; then consider q such that
  A3: q=x and
  A4: q=p or
      ex f being map of I[01],TOP-REAL n st f is continuous & rng f c= R &
                f.0=p & f.1=q by A2;
   per cases by A4;
   suppose q=p; hence x in R by A1,A3;
   suppose
        ex f being map of I[01],TOP-REAL n st f is continuous & rng f c= R &
                f.0=p & f.1=q;
    then consider f1 being map of I[01],TOP-REAL n such that
    A5: f1 is continuous & rng f1 c= R & f1.0=p & f1.1=q;
    reconsider P1=rng f1 as Subset of TOP-REAL n;
      dom f1=[.0 qua Real,1 qua Real.]
                           by BORSUK_1:83,FUNCT_2:def 1;
    then 1 in dom f1 by TOPREAL5:1;
    then q in P1 by A5,FUNCT_1:def 5;
    hence x in R by A3,A5;
 end;

theorem Th86: for R being Subset of TOP-REAL n,
p being Point of TOP-REAL n st
R is connected & R is open & p in R &
P={q: q=p or
    ex f being map of I[01],TOP-REAL n st f is continuous & rng f c= R &
                f.0=p & f.1=q}
holds R c= P
 proof let R be Subset of TOP-REAL n,
   p be Point of TOP-REAL n;
  assume that
  A1: R is connected & R is open and
  A2: p in R and
  A3: P = {q: q=p or
    ex f being map of I[01],TOP-REAL n st f is continuous & rng f c= R &
                f.0=p & f.1=q};
  reconsider R' = R as non empty Subset of TOP-REAL n by A2;
  A4: P c= R by A2,A3,Th85;
  set P2 = R \ P;
  A5: P2 c= R by XBOOLE_1:36;
     now let x;
     A6: now assume
A7:     x in P2;
      then A8: x in R & not x in P by XBOOLE_0:def 4;
      reconsider q2=x as Point of TOP-REAL n by A7;
        q2<>p & q2 in R & not
      ex f being map of I[01],TOP-REAL n st f is continuous & rng f c= R &
                f.0=p & f.1=q2 by A3,A8;
      hence
         x in {q: q<>p & q in R & not
      ex f being map of I[01],TOP-REAL n st f is continuous & rng f c= R &
                f.0=p & f.1=q};
     end;
       now assume
        x in {q: q<>p & q in R & not
      ex f being map of I[01],TOP-REAL n st f is continuous & rng f c= R &
                f.0=p & f.1=q};
then A9:      ex q3 st q3=x & q3<>p & q3 in R & not
      ex f being map of I[01],TOP-REAL n st f is continuous & rng f c= R &
                f.0=p & f.1=q3;
      then not ex q st q=x & (q=p or
      ex f being map of I[01],TOP-REAL n st f is continuous & rng f c= R &
                f.0=p & f.1=q);
      then x in R & not x in P by A3,A9;
      hence x in P2 by XBOOLE_0:def 4;
     end;
    hence x in P2 iff
    x in {q: q<>p & q in R & not
      ex f being map of I[01],TOP-REAL n st f is continuous & rng f c= R &
                f.0=p & f.1=q} by A6;
   end;
  then A10: P2={q: q<>p & q in R & not
      ex f being map of I[01],TOP-REAL n st f is continuous & rng f c= R &
                f.0=p & f.1=q} by TARSKI:2;
  reconsider P22=P2 as Subset of TOP-REAL n;
  A11: P22 is open by A1,A10,Th83;
  reconsider PPP=P as Subset of TOP-REAL n;
  A12: PPP is open by A1,A2,A3,Th84;
  A13: p in P by A3;
A14:  (TOP-REAL n)|R' is connected by A1,CONNSP_1:def 3;
  A15: [#]((TOP-REAL n)|R) = R by PRE_TOPC:def 10;
  then reconsider P11 = P, P12 = P22 as Subset of (TOP-REAL n)|R
    by A4,A5,PRE_TOPC:16;
  reconsider P11, P12 as Subset of (TOP-REAL n)|R
  ;
  A16: P11 misses P12 by XBOOLE_1:79;
    P \/ (R \ P) = P \/ R by XBOOLE_1:39;
  then A17: [#]((TOP-REAL n)|R) = P11 \/ P12 by A4,A15,XBOOLE_1:12;
  A18: P22 in the topology of TOP-REAL n & P in the topology of TOP-REAL n
          by A11,A12,PRE_TOPC:def 5;
    P11 = P /\ [#]((TOP-REAL n)|R) & P12 = P22 /\ [#]((TOP-REAL n)|R)
          by A4,A5,A15,XBOOLE_1:28;
  then P11 in the topology of (TOP-REAL n)|R & P12 in the topology of (
TOP-REAL n)|R
    by A18,PRE_TOPC:def 9;
then P11 is open & P12 is open by PRE_TOPC:def 5;
  then P11 = {}((TOP-REAL n)|R) or P12 = {}((TOP-REAL n)|R)
    by A14,A16,A17,CONNSP_1:12;
  hence thesis by A13,XBOOLE_1:37;
 end;

theorem Th87:for R being Subset of TOP-REAL n,
p,q being Point of TOP-REAL n st
R is connected & R is open & p in R & q in R & p<>q holds
ex f being map of I[01],TOP-REAL n st f is continuous & rng f c= R &
f.0=p & f.1=q
 proof let R be Subset of TOP-REAL n,
  p,q be Point of TOP-REAL n;
  assume that
  A1: R is connected & R is open and
  A2: p in R and
  A3: q in R and
  A4: p<>q;
  set RR={q2: q2=p or
     ex f being map of I[01],TOP-REAL n st f is continuous & rng f c= R &
                          f.0=p & f.1=q2};
    RR c= the carrier of TOP-REAL n
   proof let x; assume x in RR;
    then ex q2 st q2=x & (q2=p or
    ex f being map of I[01],TOP-REAL n st f is continuous & rng f c= R &
                f.0=p & f.1=q2);
    hence thesis;
   end;
  then R c= RR by A1,A2,Th86;
  then q in RR by A3; then ex q2 st q=q2 &(q2=p or
  ex f being map of I[01],TOP-REAL n st f is continuous & rng f c= R &
                f.0=p & f.1=q2);
  hence thesis by A4;
 end;

theorem Th88: for A being Subset of TOP-REAL n, a being real number
st A={q: |.q.|=a} holds A` is open & A is closed
proof let A be Subset of TOP-REAL n, a be real number;
 assume A1:A={q: |.q.|=a};
 reconsider a as Real by XREAL_0:def 1;
  A2:TOP-REAL n=TopSpaceMetr(Euclid n) by EUCLID:def 8;
  reconsider P1=A` as Subset of TopSpaceMetr(Euclid n) by EUCLID:def 8;
      for p being Point of Euclid n st p in P1
    ex r be real number st r>0 & Ball(p,r) c= P1
    proof let p be Point of Euclid n;
    assume p in P1;
    then p in (the carrier of TOP-REAL n)\A by SUBSET_1:def 5;
      then A3: not p in A by XBOOLE_0:def 4;
      reconsider q1=p as Point of TOP-REAL n by TOPREAL3:13;
      A4: |.q1.|<>a by A1,A3;
        now per cases;
      case |.q1.|<=a;
        then A5: |.q1.|<a by A4,REAL_1:def 5;
        set r1=(a- |.q1.|)/2;
          a- |.q1.|>0 by A5,SQUARE_1:11;
        then A6:r1>0 by REAL_2:127;
          Ball(p,r1) c= P1
        proof let x be set;assume A7:x in Ball(p,r1);
          then reconsider p2=x as Point of Euclid n;
          reconsider q2=p2 as Point of TOP-REAL n by TOPREAL3:13;
            dist(p,p2)<r1 by A7,METRIC_1:12;
          then A8: |.q2-q1.|<r1 by JGRAPH_1:45;
            now assume q2 in A;
            then consider q such that A9:q=q2 & |.q.|=a by A1;
              |.q2-q1.| >=|.q2.|- |.q1.| by TOPRNS_1:33;
             then r1>a- |.q1.| by A8,A9,AXIOMS:22;
             then r1>2*r1 by XCMPLX_1:88;
             then r1>r1+r1 by XCMPLX_1:11;
             then r1-r1>r1 by REAL_1:86;
           hence contradiction by A6,XCMPLX_1:14;
          end;
          then q2 in (the carrier of TOP-REAL n)\A by XBOOLE_0:def 4;
         hence x in P1 by SUBSET_1:def 5;
        end;
       hence ex r be real number st r>0 & Ball(p,r) c= P1 by A6;
      case A10: |.q1.|>a;
        set r1=(|.q1.|-a)/2;
          |.q1.|-a>0 by A10,SQUARE_1:11;
        then A11:r1>0 by REAL_2:127;
          Ball(p,r1) c= P1
        proof let x be set;assume A12:x in Ball(p,r1);
          then reconsider p2=x as Point of Euclid n;
          reconsider q2=p2 as Point of TOP-REAL n by TOPREAL3:13;
            dist(p,p2)<r1 by A12,METRIC_1:12;
          then A13: |.q1-q2.|<r1 by JGRAPH_1:45;
            now assume q2 in A;
            then consider q such that A14:q=q2 & |.q.|=a by A1;
              |.q1-q2.| >=|.q1.|- |.q2.| by TOPRNS_1:33;
             then r1>|.q1.|-a by A13,A14,AXIOMS:22;
             then r1>2*r1 by XCMPLX_1:88; then r1>r1+r1 by XCMPLX_1:11;
             then r1-r1>r1 by REAL_1:86;
           hence contradiction by A11,XCMPLX_1:14;
          end;
          then q2 in (the carrier of TOP-REAL n)\A by XBOOLE_0:def 4;
         hence x in P1 by SUBSET_1:def 5;
        end;
       hence ex r be real number st r>0 & Ball(p,r) c= P1 by A11;
      end;
     hence ex r be real number st r>0 & Ball(p,r) c= P1;
    end;
  then A` is open by A2,TOPMETR:22;
 hence A` is open & A is closed by TOPS_1:29;
end;

theorem Th89: for B being non empty Subset of TOP-REAL n
st B is open holds (TOP-REAL n)|B is locally_connected
proof let B be non empty Subset of TOP-REAL n;
 assume A1:B is open;
   for A being non empty Subset of ((TOP-REAL n)|B),
     C being Subset of ((TOP-REAL n)|B) st A is open &
 C is_a_component_of A holds C is open
 proof let A be non empty Subset of (TOP-REAL n)|B,
       C be Subset of (TOP-REAL n)|B;
  assume A2:A is open & C is_a_component_of A;
    then consider C1 being Subset of ((TOP-REAL n)|B)|A such that
A3: C1 = C & C1 is_a_component_of ((TOP-REAL n)|B)|A by CONNSP_1:def 6;
    A4:[#]((TOP-REAL n)|B)=B by PRE_TOPC:def 10;
      C1 c= the carrier of ((TOP-REAL n)|B)|A;
    then C1 c= [#](((TOP-REAL n)|B)|A) by PRE_TOPC:12;
    then A5:C1 c= A by PRE_TOPC:def 10;
    A6:TOP-REAL n=TopSpaceMetr(Euclid n) by EUCLID:def 8;
      A c= the carrier of (TOP-REAL n)|B;
    then A c= B by JORDAN1:1;
    then C c= B by A3,A5,XBOOLE_1:1;
    then C c= the carrier of TOP-REAL n by XBOOLE_1:1;
    then reconsider C0=C as Subset of TOP-REAL n;
      for p being Point of Euclid n st p in C0
    ex r be real number st r>0 & Ball(p,r) c= C0
     proof let p be Point of Euclid n;assume A7:p in C0;
       consider A0 being Subset of TOP-REAL n such that
       A8: A0 is open & A0 /\ [#]((TOP-REAL n)|B) = A by A2,TOPS_2:32;
       A9:A0 /\ B=A by A8,PRE_TOPC:def 10;
         A0 /\ B is open by A1,A8,TOPS_1:38;
       then consider r1 being real number such that
        A10:r1>0 & Ball(p,r1) c= A0 /\ B by A3,A5,A6,A7,A9,TOPMETR:22;
       reconsider r1 as Real by XREAL_0:def 1;
       A11:r1>0 & Ball(p,r1) c= A by A8,A10,PRE_TOPC:def 10;
         reconsider BL1=Ball(p,r1) as Subset of TOP-REAL n
           by TOPREAL3:13;
         reconsider BL1 as Subset of TOP-REAL n;
         reconsider BL2=Ball(p,r1)
          as Subset of (TOP-REAL n)|B by A11,XBOOLE_1:1;
         reconsider BL2
          as Subset of (TOP-REAL n)|B;
         Ball(p,r1) c= [#](((TOP-REAL n)|B)|A) by A11,PRE_TOPC:def 10;
         then reconsider BL=Ball(p,r1)
          as Subset of ((TOP-REAL n)|B)|A by PRE_TOPC:12;
         reconsider BL
          as Subset of ((TOP-REAL n)|B)|A;
         now assume not Ball(p,r1) c= C0;
         then consider x being set such that
         A12:x in Ball(p,r1) & not x in C0 by TARSKI:def 3;
            BL1 is convex by Th78;
          then BL1 is connected by Th14;
          then BL2 is connected by Th15;
         then A13:BL is connected by Th15;
           p in BL by A10,GOBOARD6:4;
         then BL /\ C <>{}(((TOP-REAL n)|B)|A) by A7,XBOOLE_0:def 3;
         then BL meets C by XBOOLE_0:def 7;
         then BL c= C by A3,A13,CONNSP_1:38;
        hence contradiction by A12;
       end;
      hence ex r be real number st r>0 & Ball(p,r) c= C0 by A10;
     end;
    then A14:C0 is open by A6,TOPMETR:22;
      C c= the carrier of ((TOP-REAL n)|B);
    then C c= B by JORDAN1:1;
    then C0 /\ B=C by XBOOLE_1:28;
  hence C is open by A4,A14,TOPS_2:32;
 end;
 hence thesis by CONNSP_2:24;
end;

theorem Th90: for B being non empty Subset of TOP-REAL n,
A being Subset of TOP-REAL n,a being real number st
A={q: |.q.|=a} & A`=B holds (TOP-REAL n)|B is locally_connected
proof let B be non empty Subset of TOP-REAL n,
A be Subset of TOP-REAL n,a be real number;
 assume A1:A={q: |.q.|=a} & A`=B; then A` is open by Th88;
 hence thesis by A1,Th89;
end;

theorem Th91: for f being map of TOP-REAL n,R^1 st
 (for q holds f.q=|.q.|) holds f is continuous
proof let f be map of TOP-REAL n,R^1;
 assume A1:(for q holds f.q=|.q.|);
  A2:TOP-REAL n=TopSpaceMetr(Euclid n) by EUCLID:def 8;
  then reconsider f1=f as map of TopSpaceMetr(Euclid n),TopSpaceMetr(RealSpace)
                                  by TOPMETR:def 7;
    now let r be real number,u be Element of
       the carrier of Euclid n,u1 be Element of RealSpace;
    assume A3:r>0 & u1=f1.u; set s1=r;
         for w being Element of Euclid n,
       w1 being Element of RealSpace st w1=f1.w &
       dist(u,w)<s1 holds dist(u1,w1)<r
       proof let w be Element of Euclid n,
       w1 be Element of RealSpace;
       assume A4:w1=f1.w & dist(u,w)<s1;
         reconsider qw=w,qu=u as Point of TOP-REAL n by TOPREAL3:13;
         A5:dist(u,w)=|.qu-qw.| by JGRAPH_1:45;
         reconsider tu=u1,tw=w1 as Real by METRIC_1:def 14;
         A6:dist(u1,w1)=(the distance of RealSpace).(u1,w1) by METRIC_1:def 1
         .=abs(tu-tw) by METRIC_1:def 13,def 14;
         A7:tu=|.qu.| by A1,A3;
           w1=|.qw.| by A1,A4;
         then dist(u1,w1)<=|.qu-qw.| by A6,A7,Th11;
        hence dist(u1,w1)<r by A4,A5,AXIOMS:22;
       end;
    hence ex s being real number st s>0 &
       for w being Element of Euclid n,
       w1 being Element of RealSpace st w1=f1.w &
       dist(u,w)<s holds dist(u1,w1)<r by A3;
   end;
 hence f is continuous by A2,TOPMETR:def 7,UNIFORM1:4;
end;

theorem Th92: ex f being map of TOP-REAL n,R^1 st
 (for q holds f.q=|.q.|) & f is continuous
proof
defpred P[set,set] means (ex q st q=$1 & $2=|.q.|);
  A1:for x,y1,y2 st x in the carrier of TOP-REAL n
  & P[x,y1] & P[x,y2] holds y1 = y2;
 A2:for x st x in (the carrier of TOP-REAL n)
  ex y st P[x,y]
  proof let x;assume x in the carrier of TOP-REAL n;
    then reconsider q3=x as Point of TOP-REAL n;
    take |.q3.|;
   thus thesis;
  end;
 consider f1 being Function such that
  A3:dom f1 = (the carrier of TOP-REAL n) &
  for x st x in (the carrier of TOP-REAL n) holds
  P[x,f1.x] from FuncEx(A1,A2);
    rng f1 c= the carrier of R^1
  proof let z be set;assume z in rng f1;
    then consider xz being set such that A4:xz in dom f1 & z=f1.xz by FUNCT_1:
def 5;
    consider q4 being Point of TOP-REAL n such that
    A5:q4=xz & f1.xz=|.q4.| by A3,A4;
   thus z in the carrier of R^1 by A4,A5,TOPMETR:24;
  end;
  then f1 is Function of the carrier of TOP-REAL n,the carrier of R^1
                           by A3,FUNCT_2:def 1,RELSET_1:11;
  then reconsider f2=f1 as map of TOP-REAL n,R^1 ;
  A6:for q holds f1.q=|.q.|
   proof let q;
    consider q2 such that A7:q2=q & f1.q=|.q2.| by A3;
    thus f1.q=|.q.| by A7;
   end;
  then f2 is continuous by Th91;
 hence thesis by A6;
end;

definition
  canceled;
  let X, Y be non empty 1-sorted, f be map of X,Y, x be set;
  assume A1:x is Point of X;
 func pi(f,x) -> Point of Y equals
:Def10: f.x;
coherence
 proof reconsider x0=x as Point of X by A1;
    f.x0 is Point of Y;
  hence thesis;
 end;
end;

theorem Th93: for g being map of I[01],TOP-REAL n st
 g is continuous holds ex f being map of I[01],R^1 st
 (for t being Point of I[01] holds f.t=|.g.t.|)
 & f is continuous
proof let g be map of I[01],TOP-REAL n;
 assume A1: g is continuous;
  consider h being map of TOP-REAL n,R^1 such that
 A2:(for q holds h.q=|.q.|) & h is continuous by Th92;
  set f1=h*g;
  A3:for t being Point of I[01] holds f1.t=|.g.t.|
   proof let t be Point of I[01];
     A4:dom g = the carrier of I[01] by FUNCT_2:def 1;
     reconsider q=g.t as Point of TOP-REAL n;
       f1.t=h.(g.t) by A4,FUNCT_1:23 .= |.q.| by A2;
    hence f1.t=|.g.t.|;
   end;
    f1 is continuous by A1,A2,TOPS_2:58;
 hence thesis by A3;
end;

theorem Th94: for g being map of I[01],TOP-REAL n,a being Real st
 g is continuous & |.pi(g,0).|<=a & a<=|.pi(g,1).| holds
 ex s being Point of I[01] st |.pi(g,s).|=a
proof let g be map of I[01],TOP-REAL n,a be Real;
 assume A1:g is continuous & |.pi(g,0).|<=a & a<=|.pi(g,1).|;
  then consider f being map of I[01],R^1 such that
  A2:(for t being Point of I[01] holds f.t=|.g.t.|)
      & f is continuous by Th93;
A3:  0 in [.0 qua Real,1 qua Real .] by TOPREAL5:1;
  reconsider o=0 as Point of I[01] by BORSUK_1:83,TOPREAL5:1;
  A4:f.0=|.g.o.| by A2 .=|.pi(g,0).| by Def10;
  A5:1 in the carrier of I[01] by BORSUK_1:83,TOPREAL5:1;
  reconsider I=1 as Point of I[01] by BORSUK_1:83,TOPREAL5:1;
  A6:f.1=|.g.I.| by A2 .=|.pi(g,1).| by Def10;
  set c = |.pi(g,0).|, b=|.pi(g,1).|;
  per cases by A1,REAL_1:def 5;
  suppose c < a & a<b; then consider rc being Real
  such that A7:f.rc =a & 0<rc & rc<1 by A2,A4,A6,TOPMETR:27,TOPREAL5:12;
  A8:rc in the carrier of I[01] by A7,BORSUK_1:83,TOPREAL5:1;
  reconsider rc1=rc as Point of I[01] by A7,BORSUK_1:83,TOPREAL5:1;
    |.pi(g,rc).|= |. g.rc1 .| by Def10 .=a by A2,A7;
 hence thesis by A8;
  suppose c =a;
 hence ex s being Point of I[01] st |.pi(g,s).|=a by A3,BORSUK_1:83;
  suppose a=b;
 hence thesis by A5;
end;

theorem Th95:q=<*r*> implies |.q.|=abs(r)
proof assume A1:q=<*r*>;
 reconsider xr=<*r*> as Element of REAL 1 by EUCLID:def 1;
   q is Element of REAL n by EUCLID:25;
 then q is Element of (n-tuples_on REAL) by EUCLID:def 1;
 then len <*r*> =n by A1,FINSEQ_2:109;
 then n=1 by FINSEQ_1:56;
 then A2: |.q.|=|.xr.| by A1,JGRAPH_1:def 5
 .=sqrt Sum sqr xr by EUCLID:def 5;
   len xr=1 by FINSEQ_1:56;
 then xr.1=xr/.1 by FINSEQ_4:24;
 then A3:Proj(q,1)=xr.1 by A1,JORDAN2B:def 1;
 A4:len (sqr xr) =1 by FINSEQ_2:109;
   (sqr xr).1=(Proj(q,1))^2 by A3,RVSUM_1:79;
then A5: sqr xr=<*(Proj(q,1))^2*> by A4,FINSEQ_1:57;
   sqrt (Proj(q,1))^2 =abs(Proj(q,1)) by SQUARE_1:91
 .=abs(r) by A3,FINSEQ_1:57;
hence thesis by A2,A5,RVSUM_1:103;
end;

theorem for A being Subset of TOP-REAL n,a being Real
st n>=1 & a>0 & A={q: |.q.|=a}
holds ex B being Subset of TOP-REAL n st
B is_inside_component_of A & B=BDD A
proof let A be Subset of TOP-REAL n,a be Real;
  assume A1:n>=1 & a>0 & A={q: |.q.|=a};
      {q where q is Point of TOP-REAL n:
    (|.q.|) <a } c= the carrier of TOP-REAL n
    proof let x;assume x in {q where q is Point of TOP-REAL n:
    (|.q.|) <a }; then consider q such that
      A2:q=x & (|.q.|) <a;
     thus x in the carrier of TOP-REAL n by A2;
    end;
    then reconsider W={q where q is Point of TOP-REAL n:
    (|.q.|) <a } as Subset of Euclid n by TOPREAL3:13;
    reconsider P=W as Subset of TOP-REAL n by TOPREAL3:13;
    reconsider P as Subset of TOP-REAL n;
      the carrier of (TOP-REAL n)|A`=A` by JORDAN1:1;
then A3:  skl (Down(P,A`)) is Subset of TOP-REAL n by XBOOLE_1:1
;
    then reconsider P1=skl (Down(P,A`)) as Subset of TOP-REAL n;
 A4:Down(P,A`)=P /\ A` by CONNSP_3:def 5;
       A5:P c= A`
       proof let x;assume A6:x in P;
         then reconsider q=x as Point of TOP-REAL n;
         consider q1 such that
         A7:q1=q & |.q1.|<a by A6;
           now assume q in A;
           then ex q2 st q2=q & |.q2.|=a by A1;
          hence contradiction by A7;
         end;
         then x in (the carrier of TOP-REAL n) \A by XBOOLE_0:def 4;
        hence x in A` by SUBSET_1:def 5;
       end;
         |.0.REAL n.|=0 by TOPRNS_1:24;
       then A8:0.REAL n in P by A1;
     then reconsider G=A` as non empty Subset of TOP-REAL n by
A5;
A9: (TOP-REAL n)|G is non empty;
    A10: Down(P,A`) <>{} by A4,A5,A8,XBOOLE_0:def 3;
       A11:Down(P,A`)=P by A4,A5,XBOOLE_1:28;
         P is connected by Th79;
       then (TOP-REAL n)|P is connected by CONNSP_1:def 3;
       then ((TOP-REAL n)|A`)|Down(P,A`) is connected
                     by A5,A11,JORDAN6:47;
       then A12:Down(P,A`) is connected by CONNSP_1:def 3;
    then A13:skl (Down(P,A`)) is_a_component_of (TOP-REAL n)|A`
                                by A9,A10,CONNSP_3:9;
    then A14:skl (Down(P,A`)) is connected by CONNSP_1:def 5;
A15: P c= skl (Down(P,A`)) by A8,A9,A11,A12,CONNSP_3:13;
        skl (Down(P,A`)) is bounded Subset of Euclid n
      proof
        reconsider D2=skl (Down(P,A`)) as Subset of TOP-REAL n
           by A3;
        reconsider D=D2 as Subset of Euclid n
                           by TOPREAL3:13;
        reconsider D as Subset of Euclid n;
          now assume not D2 is Bounded;
          then consider q such that
          A16:q in D2 & |.q.|>=a by Th40;
          set p=0.REAL n;
          A17: |.p.|<a by A1,TOPRNS_1:24;
          A18:p <> q by A1,A16,TOPRNS_1:24;
            D c= the carrier of (TOP-REAL n)|A`;
          then A19:D2 c= A` by JORDAN1:1;
          then D2=D2 /\ A` by XBOOLE_1:28;
          then A20:D2=Down(D2,A`) by CONNSP_3:def 5;
          reconsider B=A` as non empty Subset of TOP-REAL n
           by A5,A8;
          reconsider RR=(TOP-REAL n)|B as non empty TopSpace;
            RR is locally_connected by A1,Th90;
          then skl (Down(P,A`)) is open by A13,CONNSP_2:21;
          then consider G being Subset of TOP-REAL n such that
          A21: G is open &
          Down(D2,A`)=G /\ [#]((TOP-REAL n)|A`) by A20,TOPS_2:32;
          A22:G /\ A` = D2 by A20,A21,PRE_TOPC:def 10;
          A` is open by A1,Th88;
          then D2 is connected & D2 is open
           by A14,A21,A22,CONNSP_3:34,TOPS_1:38;
         then consider f1 being map of I[01],TOP-REAL n such that
         A23: f1 is continuous & rng f1 c= D2 &
            f1.0=p & f1.1=q by A8,A15,A16,A18,Th87;
            A24: |.pi(f1,0).|<a by A17,A23,Def10,BORSUK_1:def 17;
              |.pi(f1,1).|>=a by A16,A23,Def10,BORSUK_1:def 18;
          then consider t0 being Point of I[01] such that
           A25: |.pi(f1,t0).|=a by A23,A24,Th94;
           reconsider q2=f1.t0 as Point of TOP-REAL n;
             |.q2.|=a by A25,Def10;
            then A26:q2 in A by A1;
              t0 in the carrier of I[01];
            then t0 in [#](I[01]) by PRE_TOPC:12;
            then t0 in dom f1 by TOPS_2:51;
            then q2 in rng f1 by FUNCT_1:def 5;
            then q2 in D2 by A23;
            then A /\ A`<>{}(the carrier of TOP-REAL n) by A19,A26,XBOOLE_0:def
3;
            then A meets A` by XBOOLE_0:def 7;
         hence contradiction by SUBSET_1:26;
        end;
        then ex C being Subset of Euclid n st C=D2 & C is bounded by Def2;
       hence skl Down(P,A`) is bounded Subset of Euclid n;
      end;
   then A27: P1 is_inside_component_of A by A13,Th17;
  A28:P1 c= union{B where B is Subset of TOP-REAL n:
           B is_inside_component_of A}
   proof let x be set;assume A29:x in P1;
      P1 in {B where B is Subset of TOP-REAL n:
    B is_inside_component_of A} by A27;
    hence x in union{B where B is Subset of TOP-REAL n:
           B is_inside_component_of A} by A29,TARSKI:def 4;
   end;
    now per cases;
  case A30:n>=2;
    union{B where B is Subset of TOP-REAL n:
     B is_inside_component_of A} c= P1
     proof let x be set;assume x in
       union{B where B is Subset of TOP-REAL n:
           B is_inside_component_of A};
       then consider y being set such that
       A31: x in y & y in {B where B is Subset of TOP-REAL n:
           B is_inside_component_of A} by TARSKI:def 4;
       consider B being Subset of TOP-REAL n such that
       A32:B=y & B is_inside_component_of A by A31;
       A33:the carrier of (TOP-REAL n)|A`=A` by JORDAN1:1;
       consider C being Subset of ((TOP-REAL n)|(A`)) such that
       A34: C=B & C is_a_component_of ((TOP-REAL n)|(A`)) &
         C is bounded Subset of Euclid n by A32,Th17;
       reconsider p=x as Point of (TOP-REAL n)|A` by A31,A32,A34;
       reconsider E=A` as non empty Subset of TOP-REAL n
                                          by A5,A8;
A35:   p in the carrier of ((TOP-REAL n)|E);
       then p in (the carrier of TOP-REAL n)\A by A33,SUBSET_1:def 5;
       then A36:p in (the carrier of TOP-REAL n) & not p in A by XBOOLE_0:def 4
;
       reconsider q2=p as Point of TOP-REAL n by A33,A35;
         |.q2.|<>a by A1,A36;
then A37:   |.q2.|<a or |.q2.|>a by REAL_1:def 5;
         now per cases by A37;
       case A38:p in {q: |.q.|>a};
          {q: |.q.|>a} c= A`
         proof let z be set;assume z in {q: |.q.|>a};
           then consider q such that A39:q=z & |.q.|>a;
             now assume q in A;
             then ex q2 st q2=q & |.q2.|=a by A1;
            hence contradiction by A39;
           end;
           then q in (the carrier of TOP-REAL n) \ A by XBOOLE_0:def 4;
          hence z in A` by A39,SUBSET_1:def 5;
         end;
        then reconsider Q={q: |.q.|>a} as
                      Subset of (TOP-REAL n)|A` by JORDAN1:1;
        reconsider Q as Subset of (TOP-REAL n)|A`;
           {q: |.q.|>a} c= the carrier of TOP-REAL n
         proof let z be set;assume z in {q: |.q.|>a};
           then consider q such that A40:q=z & |.q.|>a;
          thus z in the carrier of TOP-REAL n by A40;
         end;
         then reconsider P2={q: |.q.|>a}
                       as Subset of TOP-REAL n;
          P2 is connected by A30,Th59;
        then A41:(TOP-REAL n)|P2 is connected by CONNSP_1:def 3;
           P2 is Subset of Euclid n by TOPREAL3:13;
         then reconsider W2={q: |.q.|>a} as Subset of Euclid n;
         A42:P2 is connected by A30,Th59;
           P2=W2;
         then A43:not W2 is bounded by A30,Th69;
         A44:Up(skl Q)=skl Q by CONNSP_3:def 6;
         A45:now assume W2 meets A;
           then consider z being set such that
           A46: z in W2 & z in A by XBOOLE_0:3;
A47:       ex q st q=z & |.q.|>a by A46;
             ex q2 st q2=z & |.q2.|=a by A1,A46;
          hence contradiction by A47;
         end;
         then W2 /\ A``={} by XBOOLE_0:def 7;
         then P2\A`={} by SUBSET_1:32;
         then A48:W2 c= A` by XBOOLE_1:37;
         then A49:P2 /\ A`=P2 by XBOOLE_1:28;
           (TOP-REAL n)|P2=((TOP-REAL n)|A`)|Q by A48,JORDAN6:47;
         then A50:Q is connected by A41,CONNSP_1:def 3;
           Q=Down(P2,A`) by A49,CONNSP_3:def 5;
          then Up(skl Q) is_outside_component_of A
                                 by A42,A43,A44,A45,Th71;
         then A51:skl Q c= UBD A by A44,Th27;
           Q c= skl Q by A50,CONNSP_3:1;
then A52:     p in skl Q by A38;
           B c= BDD A by A32,Th26;
         then p in (BDD A) /\ (UBD A) by A31,A32,A51,A52,XBOOLE_0:def 3;
         then (BDD A) meets (UBD A) by XBOOLE_0:4;
        hence x in P1 by Th28;
       case A53: p in {q1: |.q1.|<a};
           Down(P,A`) c= skl (Down(P,A`)) by A12,CONNSP_3:1;
        hence x in P1 by A11,A53;
       end;
      hence x in P1;
     end;
  then P1=union{B where B is Subset of TOP-REAL n:
           B is_inside_component_of A} by A28,XBOOLE_0:def 10;
   then P1=BDD A by Def5;
 hence ex B being Subset of TOP-REAL n st
   B is_inside_component_of A & B=BDD A by A27;
 case n<2;
  then n<1+1;
  then n<=1 by NAT_1:38;
  then A54:n=1 by A1,AXIOMS:21;
    union{B where B is Subset of TOP-REAL n:
           B is_inside_component_of A} c= P1
     proof let x be set;assume x in
       union{B where B is Subset of TOP-REAL n:
           B is_inside_component_of A};
       then consider y being set such that
       A55: x in y & y in {B where B is Subset of TOP-REAL n:
           B is_inside_component_of A} by TARSKI:def 4;
       consider B being Subset of TOP-REAL n such that
       A56:B=y & B is_inside_component_of A by A55;
       A57:the carrier of (TOP-REAL n)|A`=A` by JORDAN1:1;
       consider C being Subset of ((TOP-REAL n)|(A`)) such that
       A58: C=B & C is_a_component_of ((TOP-REAL n)|(A`)) &
        C is bounded Subset of Euclid n by A56,Th17;
       reconsider p=x as Point of (TOP-REAL n)|A` by A55,A56,A58;
       reconsider E=A` as non empty Subset of TOP-REAL n
                                          by A5,A8;
A59:   p in the carrier of ((TOP-REAL n)|E);
       then p in (the carrier of TOP-REAL n)\A by A57,SUBSET_1:def 5;
       then A60:p in (the carrier of TOP-REAL n) & not p in A by XBOOLE_0:def 4
;
       reconsider q2=p as Point of TOP-REAL n
        by A57,A59;
         |.q2.|<>a by A1,A60;
then A61:    |.q2.|<a or |.q2.|>a by REAL_1:def 5;
         now per cases by A61;
       case p in {q: |.q.|>a}; then consider q0 being Point of TOP-REAL n
          such that A62:q0=p & |.q0.|>a;
             q0 is Element of REAL n by EUCLID:25;
           then q0 is Element of (n-tuples_on REAL) by EUCLID:def 1;
           then consider r0 being Real such that
           A63: q0=<*r0*> by A54,FINSEQ_2:117;
           A64: |.q0.|=abs r0 by A63,Th95;
           A65:now per cases;
           case r0>=0; then r0=abs r0 by ABSVALUE:def 1;
            hence p in {q:ex r st q=<*r*> & r>a} or
            p in {q1:ex r1 st q1=<*r1*> & r1< -a} by A62,A63,A64;
           case r0<0;
            then -r0>a by A62,A64,ABSVALUE:def 1;
            then --r0< -a by REAL_1:50;
            hence p in {q:ex r st q=<*r*> & r>a} or
            p in {q1:ex r1 st q1=<*r1*> & r1< -a} by A62,A63;
           end;
        now per cases by A65;
      case A66:p in {q:ex r st q=<*r*> & r>a};
          {q:ex r st q=<*r*> & r>a} c= A`
         proof let z be set;assume z in {q:ex r st q=<*r*> & r>a};
           then consider q such that A67:q=z & (ex r st q=<*r*> & r>a);
           consider r such that A68: q=<*r*> & r>a by A67;
           A69:r>0 by A1,A68,AXIOMS:22;
            reconsider xr=<*r*> as Element of REAL n by A54,EUCLID:def 1;
            A70: |.q.|=|.xr.| by A68,JGRAPH_1:def 5
            .=sqrt Sum sqr xr by EUCLID:def 5;
              len xr=1 by FINSEQ_1:56;
            then xr.1=xr/.1 by FINSEQ_4:24;
            then A71:Proj(q,1)=xr.1 by A68,JORDAN2B:def 1;
            A72:len sqr xr =1 by A54,FINSEQ_2:109;
              (sqr xr).1=(Proj(q,1))^2 by A71,RVSUM_1:79;
then A73:   sqr xr=<*(Proj(q,1))^2*> by A72,FINSEQ_1:57;
              sqrt ((Proj(q,1))^2) =abs(Proj(q,1)) by SQUARE_1:91
            .=abs(r) by A71,FINSEQ_1:57;
then A74:          |.q.|=abs(r) by A70,A73,RVSUM_1:103.=r by A69,ABSVALUE:def 1
;
             now assume q in A;
             then ex q2 st q2=q & |.q2.|=a by A1;
            hence contradiction by A68,A74;
           end;
           then q in (the carrier of TOP-REAL n) \ A by XBOOLE_0:def 4;
          hence z in A` by A67,SUBSET_1:def 5;
         end;
        then reconsider Q={q:ex r st q=<*r*> & r>a} as
                      Subset of (TOP-REAL n)|A` by JORDAN1:1;
        reconsider Q as Subset of (TOP-REAL n)|A`;
           {q: |.q.|>a} c= the carrier of TOP-REAL n
         proof let z be set;assume z in {q: |.q.|>a};
           then consider q such that A75:q=z & |.q.|>a;
          thus z in the carrier of TOP-REAL n by A75;
         end;
         then reconsider P2={q: |.q.|>a}
                       as Subset of TOP-REAL n;
           {q:ex r st q=<*r*> & r>a} c= the carrier of TOP-REAL n
         proof let z be set;assume z in {q:ex r st q=<*r*> & r>a};
           then consider q such that A76:q=z & ex r st q=<*r*> & r>a;
          thus z in the carrier of TOP-REAL n by A76;
         end;
         then reconsider P3={q: ex r st q=<*r*> & r>a}
                       as Subset of TOP-REAL n;
           P2 is Subset of Euclid n by TOPREAL3:13;
         then reconsider W2={q: |.q.|>a}
                       as Subset of Euclid n;
           P3 is Subset of Euclid n by TOPREAL3:13;
         then reconsider W3=P3 as Subset of Euclid n;
          W3=P3;
        then A77:P3 is connected by A54,Th67;
        then A78:(TOP-REAL n)|P3 is connected by CONNSP_1:def 3;
         A79:not W3 is bounded by A54,Th67;
         A80:Up(skl Q)=skl Q by CONNSP_3:def 6;
         A81:now assume A82:not W2 /\ A={};
           consider z being Element of W2 /\ A;
           A83: z in W2 & z in A by A82,XBOOLE_0:def 3;
then A84:       ex q st q=z & |.q.|>a;
             ex q2 st q2=z & |.q2.|=a by A1,A83;
          hence contradiction by A84;
         end;
           W3 c= W2
          proof let z be set;assume z in W3;
            then consider q such that A85: q=z & (ex r st q=<*r*> & r>a);
            consider r such that A86:q=<*r*> & r>a by A85;
            reconsider xr=<*r*> as Element of REAL n by A54,EUCLID:def 1;
            A87: |.q.|=|.xr.| by A86,JGRAPH_1:def 5
            .=sqrt Sum sqr xr by EUCLID:def 5;
              len xr=1 by FINSEQ_1:56;
            then xr.1=xr/.1 by FINSEQ_4:24;
            then A88:Proj(q,1)=xr.1 by A86,JORDAN2B:def 1;
            A89:len sqr xr =1 by A54,FINSEQ_2:109;
              (sqr xr).1=(Proj(q,1))^2 by A88,RVSUM_1:79;
then A90:         sqr xr=<*(Proj(q,1))^2*> by A89,FINSEQ_1:57;
              sqrt (Proj(q,1))^2 =abs(Proj(q,1)) by SQUARE_1:91
            .=abs(r) by A88,FINSEQ_1:57;
            then A91: |.q.|=abs(r) by A87,A90,RVSUM_1:103;
              r>0 by A1,A86,AXIOMS:22;
            then r=abs(r) by ABSVALUE:def 1;
           hence z in W2 by A85,A86,A91;
          end;
         then A92: W3 /\ A c= W2 /\ A by XBOOLE_1:26;
then W3 /\ A={} by A81,XBOOLE_1:3;
         then A93:W3 misses A by XBOOLE_0:def 7;
           W3 /\ A``={} by A81,A92,XBOOLE_1:3;
         then W3\A`={} by SUBSET_1:32;
         then A94:W3 c= A` by XBOOLE_1:37;
         then A95:P3 /\ A`=P3 by XBOOLE_1:28;
           (TOP-REAL n)|P3=((TOP-REAL n)|A`)|Q by A94,JORDAN6:47;
         then A96:Q is connected by A78,CONNSP_1:def 3;
           Q=Down(P3,A`) by A95,CONNSP_3:def 5;
         then Up(skl Q) is_outside_component_of A
                                 by A77,A79,A80,A93,Th71;
         then A97:skl Q c= UBD A by A80,Th27;
           Q c= skl Q by A96,CONNSP_3:1;
then A98:     p in skl Q by A66;
           B c= BDD A by A56,Th26;
         then (BDD A) /\ (UBD A)<>{} by A55,A56,A97,A98,XBOOLE_0:def 3;
         then (BDD A) meets (UBD A) by XBOOLE_0:def 7;
        hence x in P1 by Th28;
        case A99: p in {q1:ex r1 st q1=<*r1*> & r1< -a};
          {q:ex r st q=<*r*> & r< -a} c= A`
         proof let z be set;assume z in {q:ex r st q=<*r*> & r< -a};
           then consider q such that A100:q=z & (ex r st q=<*r*> & r< -a);
           consider r such that A101: q=<*r*> & r< -a by A100;
           A102:r<0 by A1,A101,REAL_1:26,50;
            reconsider xr=<*r*> as Element of REAL n by A54,EUCLID:def 1;
            A103: |.q.|=|.xr.| by A101,JGRAPH_1:def 5
            .=sqrt Sum sqr xr by EUCLID:def 5;
              len xr=1 by FINSEQ_1:56;
            then xr.1=xr/.1 by FINSEQ_4:24;
            then A104:Proj(q,1)=xr.1 by A101,JORDAN2B:def 1;
            A105:len (sqr xr) =1 by A54,FINSEQ_2:109;
              (sqr xr).1=(Proj(q,1))^2 by A104,RVSUM_1:79;
then A106:         sqr xr=<*(Proj(q,1))^2*> by A105,FINSEQ_1:57;
              sqrt (Proj(q,1))^2 =abs(Proj(q,1)) by SQUARE_1:91
            .=abs(r) by A104,FINSEQ_1:57;
then A107:       |.q.|=abs(r) by A103,A106,RVSUM_1:103.=-r by A102,ABSVALUE:def
1;
             now assume q in A;
             then ex q2 st q2=q & |.q2.|=a by A1;
            hence contradiction by A101,A107;
           end;
           then q in (the carrier of TOP-REAL n) \ A by XBOOLE_0:def 4;
          hence z in A` by A100,SUBSET_1:def 5;
         end;
        then reconsider Q={q:ex r st q=<*r*> & r< -a} as
                      Subset of (TOP-REAL n)|A` by JORDAN1:1;
        reconsider Q as Subset of (TOP-REAL n)|A`;
           {q: |.q.|>a} c= the carrier of TOP-REAL n
         proof let z be set;assume z in {q: |.q.|>a};
           then consider q such that A108:q=z & |.q.|>a;
          thus z in the carrier of TOP-REAL n by A108;
         end;
         then reconsider P2={q: |.q.|>a}
                       as Subset of TOP-REAL n;
           {q:ex r st q=<*r*> & r< -a} c= the carrier of TOP-REAL n
         proof let z be set;assume z in {q:ex r st q=<*r*> & r< -a};
           then consider q such that A109:q=z & ex r st q=<*r*> & r< -a;
          thus z in the carrier of TOP-REAL n by A109;
         end;
         then reconsider P3={q: ex r st q=<*r*> & r< -a}
                       as Subset of TOP-REAL n;
           P2 is Subset of Euclid n by TOPREAL3:13;
         then reconsider W2={q: |.q.|>a}
                       as Subset of Euclid n;
           P3 is Subset of Euclid n by TOPREAL3:13;
         then reconsider W3=P3 as Subset of Euclid n;
          W3=P3;
        then A110:P3 is connected by A54,Th68;
        then A111:(TOP-REAL n)|P3 is connected by CONNSP_1:def 3;
         A112:not W3 is bounded by A54,Th68;
         A113:Up(skl Q)=skl Q by CONNSP_3:def 6;
         A114:now assume A115:not W2 /\ A={};
           consider z being Element of W2 /\ A;
           A116: z in W2 & z in A by A115,XBOOLE_0:def 3;
then A117:      ex q st q=z & |.q.|>a;
             ex q2 st q2=z & |.q2.|=a by A1,A116;
          hence contradiction by A117;
         end;
           W3 c= W2
          proof let z be set;assume z in W3;
            then consider q such that A118: q=z & (ex r st q=<*r*> & r< -a);
            consider r such that A119:q=<*r*> & r< -a by A118;
           A120:r<0 by A1,A119,REAL_1:26,50;
A121:      -r>--a by A119,REAL_1:50;
            reconsider xr=<*r*> as Element of REAL n by A54,EUCLID:def 1;
            A122: |.q.|=|.xr.| by A119,JGRAPH_1:def 5
            .=sqrt Sum sqr xr by EUCLID:def 5;
              len xr=1 by FINSEQ_1:56;
            then xr.1=xr/.1 by FINSEQ_4:24;
            then A123:Proj(q,1)=xr.1 by A119,JORDAN2B:def 1;
            A124:len sqr xr =1 by A54,FINSEQ_2:109;
              (sqr xr).1=(Proj(q,1))^2 by A123,RVSUM_1:79;
then A125:         sqr xr=<*(Proj(q,1))^2*> by A124,FINSEQ_1:57;
              sqrt (Proj(q,1))^2 =abs(Proj(q,1)) by SQUARE_1:91
            .=abs(r) by A123,FINSEQ_1:57;
            then |.q.|=abs(r) by A122,A125,RVSUM_1:103;
            then |.q.|>a by A120,A121,ABSVALUE:def 1;
           hence z in W2 by A118;
          end;
         then A126: W3 /\ A c= W2 /\ A by XBOOLE_1:26;
then W3 /\ A={} by A114,XBOOLE_1:3;
         then A127:W3 misses A by XBOOLE_0:def 7;
           W3 /\ A``={} by A114,A126,XBOOLE_1:3;
         then W3\A`={} by SUBSET_1:32;
         then A128:W3 c= A` by XBOOLE_1:37;
         then A129:P3 /\ A`=P3 by XBOOLE_1:28;
           (TOP-REAL n)|P3=((TOP-REAL n)|A`)|Q by A128,JORDAN6:47;
         then A130:Q is connected by A111,CONNSP_1:def 3;
           Q=Down(P3,A`) by A129,CONNSP_3:def 5;
         then Up(skl Q) is_outside_component_of A
                                 by A110,A112,A113,A127,Th71;
         then A131:skl Q c= UBD A by A113,Th27;
           Q c= skl Q by A130,CONNSP_3:1;
then A132:      p in skl Q by A99;
           B c= BDD A by A56,Th26;
         then p in (BDD A) /\ (UBD A) by A55,A56,A131,A132,XBOOLE_0:def 3;
         then (BDD A) meets (UBD A) by XBOOLE_0:def 7;
         hence x in P1 by Th28;
        end;
        hence x in P1;
       case A133: p in {q1: |.q1.|<a};
           Down(P,A`) c= skl (Down(P,A`)) by A12,CONNSP_3:1;
        hence x in P1 by A11,A133;
       end;
      hence x in P1;
     end;
  then P1=union{B where B is Subset of TOP-REAL n:
           B is_inside_component_of A} by A28,XBOOLE_0:def 10;
   then P1=BDD A by Def5;
 hence ex B being Subset of TOP-REAL n st
   B is_inside_component_of A & B=BDD A by A27;
 end;
 hence thesis;
end;

begin ::Bounded and Unbounded Domains of Rectangles

reserve D for
  non vertical non horizontal non empty compact Subset of TOP-REAL 2;

theorem Th97:
len (GoB SpStSeq D)=2 & width (GoB SpStSeq D)=2 &
(SpStSeq D)/.1=(GoB SpStSeq D)*(1,2) & (SpStSeq D)/.2=(GoB SpStSeq D)*(2,2) &
(SpStSeq D)/.3=(GoB SpStSeq D)*(2,1) & (SpStSeq D)/.4=(GoB SpStSeq D)*(1,1) &
(SpStSeq D)/.5=(GoB SpStSeq D)*(1,2)
proof
   set f=SpStSeq D;
   A1:len f=5 by SPRECT_1:90;
   then 1 in Seg len f by FINSEQ_1:3;
then A2:1 in dom f by FINSEQ_1:def 3;
     2 in Seg len f by A1,FINSEQ_1:3;
then A3:2 in dom f by FINSEQ_1:def 3;
     3 in Seg len f by A1,FINSEQ_1:3;
then A4:3 in dom f by FINSEQ_1:def 3;
     4 in Seg len f by A1,FINSEQ_1:3;
then A5:4 in dom f by FINSEQ_1:def 3;
     5 in Seg len f by A1,FINSEQ_1:3;
then A6:5 in dom f by FINSEQ_1:def 3;
   A7:W-bound L~f < E-bound L~f by SPRECT_1:33;
   A8:S-bound L~f < N-bound L~f by SPRECT_1:34;
   A9:f/.1 = N-min L~f & f/.1 = W-max L~f by SPRECT_1:91;
   A10:f/.(1+1) = N-max L~f & f/.(1+1) = E-max L~f by SPRECT_1:92;
   A11:f/.3 = S-max L~f & f/.3 = E-min L~f by SPRECT_1:93;
   A12:f/.4 = S-min L~f & f/.4 = W-min L~f by SPRECT_1:94;
   A13:f/.5=f/.1 by A1,FINSEQ_6:def 1;
   A14:GoB f = GoB(Incr(X_axis(f)),Incr(Y_axis(f))) by GOBOARD2:def 3;
     1 in dom (X_axis(f)) by A2,SPRECT_2:19;
   then (X_axis(f)).1=(f/.1)`1 by GOBOARD1:def 3;
   then A15:(X_axis(f)).1=W-bound L~f by A9,PSCOMP_1:84;
     2 in dom (X_axis(f)) by A3,SPRECT_2:19;
   then (X_axis(f)).2=(f/.2)`1 by GOBOARD1:def 3;
   then A16:(X_axis(f)).2=E-bound L~f by A10,PSCOMP_1:104;
     3 in dom (X_axis(f)) by A4,SPRECT_2:19;
   then (X_axis(f)).3=(f/.3)`1 by GOBOARD1:def 3;
   then A17:(X_axis(f)).3=E-bound L~f by A11,PSCOMP_1:104;
     4 in dom (X_axis(f)) by A5,SPRECT_2:19;
   then (X_axis(f)).4=(f/.4)`1 by GOBOARD1:def 3;
   then A18:(X_axis(f)).4=W-bound L~f by A12,PSCOMP_1:84;
     5 in dom (X_axis(f)) by A6,SPRECT_2:19;
   then (X_axis(f)).5=(f/.5)`1 by GOBOARD1:def 3;
   then A19:(X_axis(f)).5=W-bound L~f by A9,A13,PSCOMP_1:84;
   A20:rng (X_axis(f)) c= {W-bound L~f,E-bound L~f}
    proof let z be set;assume z in rng (X_axis(f));
      then consider u being set such that
      A21:u in dom (X_axis(f)) & z=(X_axis(f)).u by FUNCT_1:def 5;
        u in dom f by A21,SPRECT_2:19;
      then A22:u in Seg len f by FINSEQ_1:def 3;
      reconsider mu=u as Nat by A21;
        1<=mu & mu<=5 by A1,A22,FINSEQ_1:3;
then A23:   mu=1 or mu=1+1 or mu=1+2 or mu=1+3 or mu=1+4 by Th4;
      per cases by A23;
      suppose mu=1; hence thesis by A15,A21,TARSKI:def 2;
      suppose mu=2; hence thesis by A16,A21,TARSKI:def 2;
      suppose mu=3; hence thesis by A17,A21,TARSKI:def 2;
      suppose mu=4; hence thesis by A18,A21,TARSKI:def 2;
      suppose mu=5; hence thesis by A19,A21,TARSKI:def 2;
    end;
     {W-bound L~f,E-bound L~f} c= rng (X_axis(f))
    proof let z be set;assume
    A24:z in {W-bound L~f,E-bound L~f};
        now per cases by A24,TARSKI:def 2;
      case A25:z=W-bound L~f;
          1 in dom (X_axis(f)) by A2,SPRECT_2:19;
       hence z in rng (X_axis(f)) by A15,A25,FUNCT_1:def 5;
      case A26:z=E-bound L~f;
          2 in dom (X_axis(f)) by A3,SPRECT_2:19;
       hence z in rng (X_axis(f)) by A16,A26,FUNCT_1:def 5;
      end;
     hence z in rng (X_axis(f));
    end;
   then A27:rng (X_axis(f))={W-bound L~f,E-bound L~f} by A20,XBOOLE_0:def 10;
   then A28:rng (Incr (X_axis(f)))={W-bound L~f,E-bound L~f} by GOBOARD2:def 2;
     3 in dom (Y_axis(f)) by A4,SPRECT_2:20;
   then (Y_axis(f)).3=(f/.3)`2 by GOBOARD1:def 4;
   then A29:(Y_axis(f)).3=S-bound L~f by A11,PSCOMP_1:114;
     2 in dom (Y_axis(f)) by A3,SPRECT_2:20;
   then (Y_axis(f)).2=(f/.2)`2 by GOBOARD1:def 4;
   then A30:(Y_axis(f)).2=N-bound L~f by A10,PSCOMP_1:94;
     1 in dom (Y_axis(f)) by A2,SPRECT_2:20;
   then (Y_axis(f)).1=(f/.1)`2 by GOBOARD1:def 4;
   then A31:(Y_axis(f)).1=N-bound L~f by A9,PSCOMP_1:94;
     4 in dom (Y_axis(f)) by A5,SPRECT_2:20;
   then (Y_axis(f)).4=(f/.4)`2 by GOBOARD1:def 4;
   then A32:(Y_axis(f)).4=S-bound L~f by A12,PSCOMP_1:114;
     5 in dom (Y_axis(f)) by A6,SPRECT_2:20;
   then (Y_axis(f)).5=(f/.5)`2 by GOBOARD1:def 4;
   then A33:(Y_axis(f)).5=N-bound L~f by A9,A13,PSCOMP_1:94;
   A34:rng (Y_axis(f)) c= {S-bound L~f,N-bound L~f}
    proof let z be set;assume z in rng (Y_axis(f));
      then consider u being set such that
      A35:u in dom (Y_axis(f)) & z=(Y_axis(f)).u by FUNCT_1:def 5;
        u in dom f by A35,SPRECT_2:20;
      then A36:u in Seg len f by FINSEQ_1:def 3;
      reconsider mu=u as Nat by A35;
        1<=mu & mu<=5 by A1,A36,FINSEQ_1:3;
then A37:   mu=1 or mu=1+1 or mu=1+2 or mu=1+3 or mu=1+4 by Th4;
      per cases by A37;
      suppose mu=1; hence thesis by A31,A35,TARSKI:def 2;
      suppose mu=2; hence thesis by A30,A35,TARSKI:def 2;
      suppose mu=3; hence thesis by A29,A35,TARSKI:def 2;
      suppose mu=4; hence thesis by A32,A35,TARSKI:def 2;
      suppose mu=5; hence thesis by A33,A35,TARSKI:def 2;
    end;
     {S-bound L~f,N-bound L~f} c= rng (Y_axis(f))
    proof let z be set;assume A38:z in {S-bound L~f,N-bound L~f};
        now per cases by A38,TARSKI:def 2;
      case A39:z=S-bound L~f;
          4 in dom (Y_axis(f)) by A5,SPRECT_2:20;
       hence z in rng (Y_axis(f)) by A32,A39,FUNCT_1:def 5;
      case A40:z=N-bound L~f;
          2 in dom (Y_axis(f)) by A3,SPRECT_2:20;
       hence z in rng (Y_axis(f)) by A30,A40,FUNCT_1:def 5;
      end;
     hence z in rng (Y_axis(f));
    end;
   then A41:rng (Y_axis(f))={S-bound L~f,N-bound L~f} by A34,XBOOLE_0:def 10;
   then A42:rng (Incr (Y_axis(f)))={S-bound L~f,N-bound L~f} by GOBOARD2:def 2;
     card rng (Y_axis(f))=2 by A8,A41,CARD_2:76;
   then A43:len (Incr (Y_axis(f)))=2 by GOBOARD2:def 2;
A44:card rng (X_axis(f))=2 by A7,A27,CARD_2:76;
   then A45:len (Incr (X_axis(f)))=2 by GOBOARD2:def 2;
   A46: len GoB(f) = card rng X_axis(f) by GOBOARD2:24 .=1+1
    by A7,A27,CARD_2:76;
   A47: width GoB(f) = card rng Y_axis(f) by GOBOARD2:24 .=1+1
    by A8,A41,CARD_2:76;
   A48:1 in Seg len GoB f by A46,FINSEQ_1:3;
   A49:len GoB f in Seg len GoB f by A46,FINSEQ_1:3;
     for p being FinSequence of the carrier of TOP-REAL 2
   st p in rng GoB f holds len p = 2
    proof let p be FinSequence of the carrier of TOP-REAL 2;
     assume A50: p in rng GoB f;
        len GoB(Incr(X_axis(f)),Incr(Y_axis(f)) )
      =len (Incr(X_axis(f))) by GOBOARD2:def 1 .=2 by A44,GOBOARD2:def 2;
      then consider s1 being FinSequence such that
      A51: s1 in rng GoB(Incr(X_axis(f)),Incr(Y_axis(f)) )
      & len s1 = width GoB(Incr(X_axis(f)),Incr(Y_axis(f)) )
                              by MATRIX_1:def 4;
      A52:s1 in rng GoB f by A51,GOBOARD2:def 3;
     consider n being Nat such that
     A53:for x st x in rng GoB f
     ex s being FinSequence st s=x & len s = n by MATRIX_1:def 1;
     consider s being FinSequence such that
     A54:s=s1 & len s = n by A52,A53;
     consider s2 being FinSequence such that
     A55:s2=p & len s2 = n by A50,A53;
     thus len p = 2 by A47,A51,A54,A55,GOBOARD2:def 3;
    end;
   then A56:GoB f is Matrix of 2,2,the carrier of TOP-REAL 2
                            by A46,MATRIX_1:def 3;
   A57:width GoB f in Seg (width GoB f) by A47,FINSEQ_1:3;
   then A58:[1,width GoB f] in [:Seg (len GoB f),Seg (width GoB f):]
                                by A48,ZFMISC_1:106;
   A59: 1 in Seg (width GoB f) by A47,FINSEQ_1:3;
     [len GoB f,width GoB f] in [:Seg (len GoB f),Seg (width GoB f):]
                                by A49,A57,ZFMISC_1:106;
   then A60: [1,width GoB f] in Indices GoB f
      & [len GoB f,width GoB f] in Indices GoB f
       by A46,A47,A56,A58,MATRIX_1:25;
     [len GoB f,1] in [:Seg (len GoB f),Seg (width GoB f):]
                                by A49,A59,ZFMISC_1:106;
   then A61: [len GoB f,1] in Indices GoB f
                                by A46,A47,A56,MATRIX_1:25;
     [1,1] in [:Seg (len GoB f),Seg (width GoB f):]
                                by A48,A59,ZFMISC_1:106;
   then A62: [1,1] in Indices GoB f by A46,A47,A56,MATRIX_1:25;
     W-bound L~f =(Incr(X_axis(f))).1 by A7,A28,A45,Th8;
   then (W-max L~f)`1 =(Incr(X_axis(f))).1 by PSCOMP_1:84;
   then (W-max L~f)`1 =(|[(Incr(X_axis(f))).1,(Incr(Y_axis(f))).(1+1)]|)`1
                                    by EUCLID:56;
   then A63: (W-max L~f)`1 =((GoB f)*(1,width (GoB f)))`1
     by A14,A47,A60,GOBOARD2:def 1;
     (N-max L~f)`1 =(NE-corner D)`1 by SPRECT_1:85
   .=E-bound D by PSCOMP_1:76 .=E-bound L~f by SPRECT_1:69
   .=(Incr(X_axis(f))).2 by A7,A28,A45,Th8;
   then (N-max L~f)`1 =(|[(Incr(X_axis(f))).(1+1),(Incr(Y_axis(f))).(1+1)]|)`1
                                    by EUCLID:56;
   then A64: (N-max L~f)`1 =((GoB f)*(len (GoB f),width (GoB f)))`1
                               by A14,A46,A47,A60,GOBOARD2:def 1;
     (S-max L~f)`1 =(SE-corner D)`1 by SPRECT_1:89
   .=E-bound D by PSCOMP_1:78 .=E-bound L~f by SPRECT_1:69
   .=(Incr(X_axis(f))).2 by A7,A28,A45,Th8;
   then (S-max L~f)`1 =(|[(Incr(X_axis(f))).2,(Incr(Y_axis(f))).1]|)`1
                                    by EUCLID:56;
   then A65: (S-max L~f)`1 =((GoB f)*(len (GoB f),1))`1
                               by A14,A46,A61,GOBOARD2:def 1;
     (S-min L~f)`1 =(SW-corner D)`1 by SPRECT_1:88
   .=W-bound D by PSCOMP_1:72 .=W-bound L~f by SPRECT_1:66
   .=(Incr(X_axis(f))).1 by A7,A28,A45,Th8;
   then (S-min L~f)`1 =(|[(Incr(X_axis(f))).1,(Incr(Y_axis(f))).1]|)`1
                                    by EUCLID:56;
   then A66: (S-min L~f)`1 =((GoB f)*(1,1))`1
                               by A14,A62,GOBOARD2:def 1;
A67: N-bound L~f =(Incr(Y_axis(f))).2 by A8,A42,A43,Th8;
   then (N-min L~f)`2 =(Incr(Y_axis(f))).2 by PSCOMP_1:94;
   then (N-min L~f)`2 =(|[(Incr(X_axis(f))).1,(Incr(Y_axis(f))).2]|)`2
                                    by EUCLID:56;
   then A68: (N-min L~f)`2 =((GoB f)*(1,width (GoB f)))`2
     by A14,A47,A60,GOBOARD2:def 1;
     (N-min L~f)`2 = N-bound (L~f) &
   (N-max L~f)`2 = N-bound (L~f) by PSCOMP_1:94;
   then (N-max L~f)`2 =(|[(Incr(X_axis(f))).2,(Incr(Y_axis(f))).2]|)`2
                                    by A67,EUCLID:56;
   then A69: (N-max L~f)`2 =((GoB f)*(len (GoB f),width (GoB f)))`2
     by A14,A46,A47,A60,GOBOARD2:def 1;
   A70:S-bound L~f =(Incr(Y_axis(f))).1 by A8,A42,A43,Th8;
   then (S-max L~f)`2 =(Incr(Y_axis(f))).1 by PSCOMP_1:114;
   then (S-max L~f)`2 =(|[(Incr(X_axis(f))).2,(Incr(Y_axis(f))).1]|)`2
                                    by EUCLID:56;
   then A71: (S-max L~f)`2 =((GoB f)*(len (GoB f),1))`2
    by A14,A46,A61,GOBOARD2:def 1;
     (S-min L~f)`2 =(Incr(Y_axis(f))).1 by A70,PSCOMP_1:114;
   then (S-min L~f)`2 =(|[(Incr(X_axis(f))).1,(Incr(Y_axis(f))).1]|)`2
                                    by EUCLID:56;
   then A72: (S-min L~f)`2 =((GoB f)*(1,1))`2 by A14,A62,GOBOARD2:def 1;
   A73:f/.1=|[(f/.1)`1,(f/.1)`2]| by EUCLID:57;
   A74:f/.(1+1)=|[(f/.(1+1))`1,(f/.(1+1))`2]| by EUCLID:57;
   A75:f/.3=|[(f/.3)`1,(f/.3)`2]| by EUCLID:57;
     f/.4=|[(f/.4)`1,(f/.4)`2]| by EUCLID:57;
 hence len (GoB f)=2 & width (GoB f)=2 &
f/.1=(GoB f)*(1,2) & f/.2=(GoB f)*(2,2) &
f/.3=(GoB f)*(2,1) & f/.4=(GoB f)*(1,1) &
f/.5=(GoB f)*(1,2) by A9,A10,A11,A12,A13,A46,A47,A63,A64,A65,A66,A68,A69,A71,
A72,A73,A74,A75,EUCLID:57;
end;

theorem Th98:
 LeftComp SpStSeq D is non Bounded
proof
  set f=SpStSeq D;
  assume LeftComp f is Bounded;
   then consider r being Real such that
A1:for q being Point of TOP-REAL 2 st q in (LeftComp f) holds |.q.|<r by Th40;
   consider q3 being Element of LeftComp f;
   reconsider q4=q3 as Point of TOP-REAL 2;
   A2: |.q4.|<r by A1;
    A3: 0<=|.q4.| by TOPRNS_1:26;
   set r1=|. (1/2)*(f/.1+f/.2).|;
   A4:r1>=0 by TOPRNS_1:26;
    A5: r1<r1+r by A2,A3,REAL_1:69;
     r1+r<r1+r+1 by REAL_1:69;
   then A6:0<r1+r+1 by A4,A5,AXIOMS:22;
   set q1=|[0,r1+r+1]|+(1/2)*(f/.1+f/.2);
   A7: |.q1.|>=|.(|[0,r1+r+1]|).|-r1 by TOPRNS_1:32;
   set p=|[0,r1+r+1]|;
   A8:p`1=0 & p`2=r1+r+1 by EUCLID:56;
     |.(|[0,r1+r+1]|).| = sqrt ((p`1)^2+(p`2)^2) by JGRAPH_1:47
   .=r1+r+1 by A6,A8,SQUARE_1:60,89;
   then A9: |.(|[0,r1+r+1]|).|-r1=r1+(r+1)-r1 by XCMPLX_1:1 .=r+1 by XCMPLX_1:
26;
     r<r+1 by REAL_1:69;
then A10: |.q1.|>r by A7,A9,AXIOMS:22;
   reconsider f1=f as non constant standard special_circular_sequence;
   A11:len f1=5 by SPRECT_1:90;
   then A12:1 in Seg len f1 by FINSEQ_1:3;
   then A13:1 in dom f1 by FINSEQ_1:def 3;
     2 in Seg len f1 by A11,FINSEQ_1:3;
   then A14:2 in dom f1 by FINSEQ_1:def 3;
     3 in Seg len f1 by A11,FINSEQ_1:3;
   then A15:3 in dom f1 by FINSEQ_1:def 3;
     4 in Seg len f1 by A11,FINSEQ_1:3;
   then A16:4 in dom f1 by FINSEQ_1:def 3;
     5 in Seg len f1 by A11,FINSEQ_1:3;
   then A17:5 in dom f1 by FINSEQ_1:def 3;
   A18:W-bound L~f1 < E-bound L~f1 by SPRECT_1:33;
   A19:S-bound L~f1 < N-bound L~f1 by SPRECT_1:34;
   A20:f1/.1 = N-min L~f1 & f1/.1 = W-max L~f1 by SPRECT_1:91;
   A21:f1/.(1+1) = N-max L~f1 & f1/.(1+1) = E-max L~f1 by SPRECT_1:92;
   A22:f1/.3 = S-max L~f1 & f1/.3 = E-min L~f1 by SPRECT_1:93;
   A23:f1/.4 = S-min L~f1 & f1/.4 = W-min L~f1 by SPRECT_1:94;
   A24:f1/.5=f1/.1 by A11,FINSEQ_6:def 1;
   A25:(f1/.1)`1 < (f1/.2)`1 by A20,A21,SPRECT_2:55;
   A26:GoB f1 = GoB(Incr(X_axis(f1)),Incr(Y_axis(f1))) by GOBOARD2:def 3;
   A27:f1/.1=W-max L~f1 by SPRECT_1:91;
     1 in dom (X_axis(f1)) by A13,SPRECT_2:19;
   then (X_axis(f1)).1=(f1/.1)`1 by GOBOARD1:def 3;
   then A28:(X_axis(f1)).1=W-bound L~f1 by A27,PSCOMP_1:84;
   A29:f1/.2=E-max L~f1 by SPRECT_1:92;
   A30:2 in dom (X_axis(f1)) by A14,SPRECT_2:19;
   then (X_axis(f1)).2=(f1/.2)`1 by GOBOARD1:def 3;
   then A31:(X_axis(f1)).2=E-bound L~f1 by A29,PSCOMP_1:104;
   A32:f1/.3=E-min L~f1 by SPRECT_1:93;
     3 in dom (X_axis(f1)) by A15,SPRECT_2:19;
   then (X_axis(f1)).3=(f1/.3)`1 by GOBOARD1:def 3;
   then A33:(X_axis(f1)).3=E-bound L~f1 by A32,PSCOMP_1:104;
   A34:f1/.4=W-min L~f1 by SPRECT_1:94;
     4 in dom (X_axis(f1)) by A16,SPRECT_2:19;
   then (X_axis(f1)).4=(f1/.4)`1 by GOBOARD1:def 3;
   then A35:(X_axis(f1)).4=W-bound L~f1 by A34,PSCOMP_1:84;
   A36:f1/.5=W-max L~f1 by A11,A20,FINSEQ_6:def 1;
     5 in dom (X_axis(f1)) by A17,SPRECT_2:19;
   then (X_axis(f1)).5=(f1/.5)`1 by GOBOARD1:def 3;
   then A37:(X_axis(f1)).5=W-bound L~f1 by A36,PSCOMP_1:84;
   A38:rng (X_axis(f1)) c= {W-bound L~f1,E-bound L~f1}
    proof let z be set;assume z in rng (X_axis(f1));
      then consider u being set such that
      A39:u in dom (X_axis(f1)) & z=(X_axis(f1)).u by FUNCT_1:def 5;
        u in dom f1 by A39,SPRECT_2:19;
      then A40:u in Seg len f1 by FINSEQ_1:def 3;
      reconsider mu=u as Nat by A39;
        1<=mu & mu<=5 by A11,A40,FINSEQ_1:3;
then A41:   mu=1 or mu=1+1 or mu=1+2 or mu=1+3 or mu=1+4 by Th4;
      per cases by A41;
      suppose mu=1; hence thesis by A28,A39,TARSKI:def 2;
      suppose mu=2; hence thesis by A31,A39,TARSKI:def 2;
      suppose mu=3; hence thesis by A33,A39,TARSKI:def 2;
      suppose mu=4; hence thesis by A35,A39,TARSKI:def 2;
      suppose mu=5; hence thesis by A37,A39,TARSKI:def 2;
    end;
     {W-bound L~f1,E-bound L~f1} c= rng (X_axis(f1))
    proof let z be set; assume
    A42:z in {W-bound L~f1,E-bound L~f1};
      per cases by A42,TARSKI:def 2;
      suppose A43:z=W-bound L~f1;
          1 in dom (X_axis(f1)) by A13,SPRECT_2:19;
       hence z in rng (X_axis(f1)) by A28,A43,FUNCT_1:def 5;
      suppose A44:z=E-bound L~f1;
          2 in dom (X_axis(f1)) by A14,SPRECT_2:19;
       hence z in rng (X_axis(f1)) by A31,A44,FUNCT_1:def 5;
    end;
   then A45:rng (X_axis(f1))={W-bound L~f1,E-bound L~f1} by A38,XBOOLE_0:def 10
;
     3 in dom (Y_axis(f1)) by A15,SPRECT_2:20;
   then (Y_axis(f1)).3=(f1/.3)`2 by GOBOARD1:def 4;
   then A46:(Y_axis(f1)).3=S-bound L~f1 by A22,PSCOMP_1:114;
     2 in dom (Y_axis(f1)) by A14,SPRECT_2:20;
   then (Y_axis(f1)).2=(f1/.2)`2 by GOBOARD1:def 4;
   then A47:(Y_axis(f1)).2=N-bound L~f1 by A21,PSCOMP_1:94;
     1 in dom (Y_axis(f1)) by A13,SPRECT_2:20;
   then (Y_axis(f1)).1=(f1/.1)`2 by GOBOARD1:def 4;
   then A48:(Y_axis(f1)).1=N-bound L~f1 by A20,PSCOMP_1:94;
     4 in dom (Y_axis(f1)) by A16,SPRECT_2:20;
   then (Y_axis(f1)).4=(f1/.4)`2 by GOBOARD1:def 4;
   then A49:(Y_axis(f1)).4=S-bound L~f1 by A23,PSCOMP_1:114;
     5 in dom (Y_axis(f1)) by A17,SPRECT_2:20;
   then (Y_axis(f1)).5=(f1/.5)`2 by GOBOARD1:def 4;
   then A50:(Y_axis(f1)).5=N-bound L~f1 by A20,A24,PSCOMP_1:94;
   A51:rng (Y_axis(f1)) c= {S-bound L~f1,N-bound L~f1}
    proof let z be set;assume z in rng (Y_axis(f1));
      then consider u being set such that
      A52:u in dom (Y_axis(f1)) & z=(Y_axis(f1)).u by FUNCT_1:def 5;
        u in dom f1 by A52,SPRECT_2:20;
      then A53:u in Seg len f1 by FINSEQ_1:def 3;
      reconsider mu=u as Nat by A52;
        1<=mu & mu<=5 by A11,A53,FINSEQ_1:3;
then A54:   mu=1 or mu=1+1 or mu=1+2 or mu=1+3 or mu=1+4 by Th4;
      per cases by A54;
      suppose mu=1; hence thesis by A48,A52,TARSKI:def 2;
      suppose mu=2; hence thesis by A47,A52,TARSKI:def 2;
      suppose mu=3; hence thesis by A46,A52,TARSKI:def 2;
      suppose mu=4; hence thesis by A49,A52,TARSKI:def 2;
      suppose mu=5; hence thesis by A50,A52,TARSKI:def 2;
    end;
     {S-bound L~f1,N-bound L~f1} c= rng (Y_axis(f1))
    proof let z be set;assume
    A55:z in {S-bound L~f1,N-bound L~f1};
      per cases by A55,TARSKI:def 2;
      suppose A56:z=S-bound L~f1;
          4 in dom (Y_axis(f1)) by A16,SPRECT_2:20;
       hence z in rng (Y_axis(f1)) by A49,A56,FUNCT_1:def 5;
      suppose A57:z=N-bound L~f1;
          2 in dom (Y_axis(f1)) by A14,SPRECT_2:20;
       hence z in rng (Y_axis(f1)) by A47,A57,FUNCT_1:def 5;
    end;
   then A58:rng (Y_axis(f1))={S-bound L~f1,N-bound L~f1} by A51,XBOOLE_0:def 10
;
   A59:card rng (X_axis(f1))=2 by A18,A45,CARD_2:76;
   A60: len GoB(f1) = card rng X_axis(f1) by GOBOARD2:24 .=1+1
     by A18,A45,CARD_2:76;
   A61: width GoB(f1) = card rng Y_axis(f1) by GOBOARD2:24 .=1+1
     by A19,A58,CARD_2:76;
   A62:1 in Seg len GoB f1 by A60,FINSEQ_1:3;
   A63:len GoB f1 in Seg len GoB f1 by A60,FINSEQ_1:3;
     for p being FinSequence of the carrier of TOP-REAL 2
   st p in rng GoB f1 holds len p = 2
    proof let p be FinSequence of the carrier of TOP-REAL 2;
     assume A64: p in rng GoB f1;
        len GoB(Incr(X_axis(f1)),Incr(Y_axis(f1)) )
      =len (Incr(X_axis(f1))) by GOBOARD2:def 1
     .=2 by A59,GOBOARD2:def 2;
      then consider s1 being FinSequence such that
      A65: s1 in rng GoB(Incr(X_axis(f1)),Incr(Y_axis(f1)) )
      & len s1 = width GoB(Incr(X_axis(f1)),Incr(Y_axis(f1)))
        by MATRIX_1:def 4;
     consider n being Nat such that
     A66:for x st x in rng GoB f1
     ex s being FinSequence st s=x & len s = n by MATRIX_1:def 1;
     consider s being FinSequence such that
     A67:s=s1 & len s = n by A26,A65,A66;
     consider s2 being FinSequence such that
     A68:s2=p & len s2 = n by A64,A66;
     thus len p = 2 by A61,A65,A67,A68,GOBOARD2:def 3;
    end;
   then A69:GoB f1 is Matrix of 2,2,the carrier of TOP-REAL 2
                            by A60,MATRIX_1:def 3;
   A70:width GoB f1 in Seg (width GoB f1) by A61,FINSEQ_1:3;
then A71: [1,width GoB f1] in [:Seg (len GoB f1),Seg (width GoB f1):]
                                by A62,ZFMISC_1:106;
     [len GoB f1,width GoB f1] in [:Seg (len GoB f1),Seg (width GoB f1):]
                                by A63,A70,ZFMISC_1:106;
   then A72: [1,width GoB f1] in Indices GoB f1
      & [len GoB f1,width GoB f1] in Indices GoB f1
       by A60,A61,A69,A71,MATRIX_1:25;
   A73:f1/.1=(GoB f1)*(1,width (GoB f1)) by A61,Th97;
   A74:f1/.(1+1)=(GoB f1)*(len (GoB f1),width (GoB f1)) by A60,A61,Th97;
     len (X_axis(f1))=len f1 by GOBOARD1:def 3;
then A75:dom (X_axis(f1))=Seg len f1 by FINSEQ_1:def 3;
then A76:(X_axis(f1)).1=(f1/.1)`1 by A12,GOBOARD1:def 3;
     (X_axis(f1)).2=(f1/.2)`1 by A30,GOBOARD1:def 3;
   then A77:(f1/.1)`1 in rng X_axis(f1) & (f1/.2)`1 in rng X_axis(f1)
                                by A12,A30,A75,A76,FUNCT_1:def 5;
     {(f1/.1)`1,(f1/.2)`1} c= rng X_axis(f1)
    proof let z be set;assume z in {(f1/.1)`1,(f1/.2)`1};
     hence z in rng X_axis(f1) by A77,TARSKI:def 2;
    end;
   then {(f1/.1)`1} \/ {(f1/.2)`1} c= rng X_axis(f1) by ENUMSET1:41;
   then A78:card ({(f1/.1)`1} \/ {(f1/.2)`1}) <= card rng X_axis(f1) by CARD_1:
80;
     not (f1/.2)`1 in {(f1/.1)`1} by A25,TARSKI:def 1;
   then card ({(f1/.1)`1} \/ {(f1/.2)`1})=card ({(f1/.1)`1})+1 by CARD_2:54
   .=1+1 by CARD_1:79 .=2;
   then card rng X_axis(f1) >1 by A78,AXIOMS:22;
   then A79:1<len (GoB f1) by GOBOARD2:24;
   then A80:1<>len (GoB f1) +1 by REAL_1:69;
     width GoB f1 <> width (GoB f1)+1 by REAL_1:69;
   then left_cell(f1,1) = cell(GoB(f1),1,width GoB(f1))
                         by A11,A72,A73,A74,A80,GOBOARD5:def 7;
   then A81:Int (left_cell(f1,1))=
    { |[r2,s]| : (GoB f1)*(1,1)`1 < r2 & r2 < (GoB f1)*(1+1,1)`1
     & (GoB f1)*(1,width (GoB f1))`2 < s } by A79,GOBOARD6:28;
   A82:f/.1=|[(W-max L~f)`1,(N-min L~f)`2]| by A20,EUCLID:57;
     f/.2=|[(E-max L~f)`1,(N-max L~f)`2]| by A21,EUCLID:57;
   then A83:f/.1+f/.2=|[(W-max L~f)`1 +(E-max L~f)`1,
                     (N-min L~f)`2+(N-max L~f)`2]| by A82,EUCLID:60;
   A84:(N-min L~f1)`2=N-bound L~f1 by PSCOMP_1:94;
   then (N-min L~f1)`2+(N-max L~f1)`2=N-bound L~f1 +N-bound L~f1 by PSCOMP_1:94
                             .=2*(N-bound L~f) by XCMPLX_1:11;
   then (1/2)*((N-min L~f)`2+(N-max L~f)`2)=(1/2)*2*(N-bound L~f) by XCMPLX_1:4
                                      .=N-bound L~f;
   then (1/2)*(f/.1+f/.2)=
 |[1/2*((W-max L~f)`1 +(E-max L~f)`1),N-bound L~f]| by A83,EUCLID:62;
   then A85:q1=|[0+1/2*((W-max L~f)`1 +(E-max L~f)`1),r1+r+1+(N-bound L~f)]|
                                         by EUCLID:60
     .=|[1/2*((W-max L~f)`1 +(E-max L~f)`1),r1+r+1+(N-bound L~f)]|;
   A86:(GoB f1)*(1,1)`1<=(W-max L~f)`1 by A11,A27,A61,JORDAN5D:7;
   A87:(GoB f1)*(1+1,1)`1>=(E-max L~f)`1 by A11,A29,A60,A61,JORDAN5D:7;
     (W-max L~f)`1=W-bound L~f by PSCOMP_1:84;
   then A88:(W-max L~f)`1<(E-max L~f)`1 by A18,PSCOMP_1:104;
   then (GoB f1)*(1,1)`1<(E-max L~f)`1 by A86,AXIOMS:22;
   then (GoB f1)*(1,1)`1 +(GoB f1)*(1,1)`1<(W-max L~f)`1+(E-max L~f)`1
                                               by A86,REAL_1:67;
   then 2*((GoB f1)*(1,1)`1)<(W-max L~f)`1+(E-max L~f)`1 by XCMPLX_1:11;
   then 1/2*(2*((GoB f1)*(1,1)`1))<1/2*((W-max L~f)`1+(E-max L~f)`1)
                         by REAL_1:70;
then A89: 1/2*2*((GoB f1)*(1,1)`1)<1/2*((W-max L~f)`1+(E-max L~f)`1)
                         by XCMPLX_1:4;
     (W-max L~f)`1 <(GoB f1)*(1+1,1)`1 by A87,A88,AXIOMS:22;
   then (W-max L~f)`1+(E-max L~f)`1< (GoB f1)*(1+1,1)`1 +(GoB f1)*(1+1,1)`1
                                               by A87,REAL_1:67;
   then (W-max L~f)`1+(E-max L~f)`1 < 2*((GoB f1)*(1+1,1)`1)
                         by XCMPLX_1:11;
   then 1/2*((W-max L~f)`1+(E-max L~f)`1)< 1/2*(2*((GoB f1)*(1+1,1)`1))
                         by REAL_1:70;
then A90: 1/2*((W-max L~f)`1+(E-max L~f)`1)< 1/2*2*((GoB f1)*(1+1,1)`1)
                         by XCMPLX_1:4;
     (GoB f1)*(1,width (GoB f1))`2< (N-bound L~f)+(r1+r+1)
     by A6,A20,A73,A84,REAL_1:69;
   then A91:q1 in Int (left_cell(f1,1)) by A81,A85,A89,A90;
     Int (left_cell(f1,1)) c= (LeftComp f) by GOBOARD9:def 1;
   hence contradiction by A1,A10,A91;
end;

theorem Th99:
LeftComp SpStSeq D c= UBD (L~SpStSeq D)
proof
  set f=SpStSeq D;
  set A=L~SpStSeq D;
  A1:LeftComp f is_a_component_of A` by GOBOARD9:def 1;
    not LeftComp f is Bounded by Th98;
  then A2:LeftComp f is_outside_component_of A by A1,Def4;
    LeftComp f c=
  union{B where B is Subset of TOP-REAL 2: B is_outside_component_of A}
   proof let x;assume A3:x in LeftComp f;
       LeftComp f in
     {B where B is Subset of TOP-REAL 2: B is_outside_component_of A} by A2;
    hence x in union{B where B is Subset of TOP-REAL 2:
                            B is_outside_component_of A} by A3,TARSKI:def 4;
   end;
  hence LeftComp f c= UBD (L~SpStSeq D) by Def6;
end;

theorem Th100:for G being TopSpace,A,B,C being Subset of G st
A is_a_component_of G & B is_a_component_of G & C is connected
& A meets C & B meets C holds A=B
proof let G be TopSpace,A,B,C be Subset of G;
 assume A1:A is_a_component_of G & B is_a_component_of G
 & C is connected & A meets C & B meets C;
  then A2: C /\ A={}G or C c= A by CONNSP_1:38;
A3: C misses B or C c= B by A1,CONNSP_1:38;
 A4: A=B or A,B are_separated by A1,CONNSP_1:36;
 per cases by A4,CONNSP_1:2;
 suppose A=B;
 hence A=B;
 suppose A misses B;
then A5:A /\ B = {} by XBOOLE_0:def 7;
    C c= A /\ B by A1,A2,A3,XBOOLE_0:def 7,XBOOLE_1:19;
  then C ={} by A5,XBOOLE_1:3;
  then C /\ A = {};
 hence A=B by A1,XBOOLE_0:def 7;
end;

theorem Th101:for B being Subset of TOP-REAL 2
st B is_a_component_of (L~SpStSeq D)` & not B is Bounded
holds B=LeftComp SpStSeq D
proof let B be Subset of TOP-REAL 2;
set f = SpStSeq D;
assume A1: B is_a_component_of (L~f)` & not B is Bounded;
  A2:LeftComp f is_a_component_of (L~f)` by GOBOARD9:def 1;
A3:not LeftComp f is Bounded by Th98;
  consider B1 being Subset of (TOP-REAL 2)|(L~f)` such that
  A4: B1 = B & B1 is_a_component_of (TOP-REAL 2)|(L~f)`
                             by A1,CONNSP_1:def 6;
  consider L1 being Subset of (TOP-REAL 2)|(L~f)` such that
  A5: L1 = LeftComp f & L1 is_a_component_of (TOP-REAL 2)|(L~f)`
                               by A2,CONNSP_1:def 6;
  L~f is Bounded by Th73;
  then consider r1 being Real such that
  A6: for q being Point of TOP-REAL 2 st q in L~f holds |.q.|<r1 by Th40;
    (REAL 2)\ {q where q is Point of TOP-REAL 2: (|.q.|) < r1 }
  c= REAL 2 by XBOOLE_1:36;
  then (REAL 2)\ {q where q is Point of TOP-REAL 2: (|.q.|) < r1 }
    c= the carrier of TOP-REAL 2 by EUCLID:25;
  then reconsider P=(REAL 2)\ {q where q is Point of TOP-REAL 2: (|.q.|) < r1 }
  as Subset of TOP-REAL 2;
  A7:P is connected by Th61;
    P c= (the carrier of TOP-REAL 2)\(L~f)
   proof let z be set;assume A8:z in P;
       now assume A9:z in L~f;
       then reconsider q3=z as Point of TOP-REAL 2;
        A10: |.q3.|<r1 by A6,A9;
          not q3 in {q where q is Point of TOP-REAL 2: (|.q.|) < r1 }
                  by A8,XBOOLE_0:def 4;
      hence contradiction by A10;
     end;
    hence z in (the carrier of TOP-REAL 2)\(L~f) by A8,XBOOLE_0:def 4;
   end;
  then P /\ ((the carrier of TOP-REAL 2)\(L~f))=P by XBOOLE_1:28;
  then P /\ (L~f)`=P by SUBSET_1:def 5;
  then A11:P=Down(P,(L~f)`) by CONNSP_3:def 5;
  then A12: Down(P,(L~f)`) is connected by A7,Th15;
  consider q3 being Point of TOP-REAL 2 such that
  A13:  q3 in LeftComp f & |.q3.|>=r1 by A3,Th40;
    q3 in the carrier of TOP-REAL 2;
  then A14:q3 in REAL 2 by EUCLID:25;
    now assume q3 in {q where q is Point of TOP-REAL 2: (|.q.|) <r1};
    then ex q being Point of TOP-REAL 2 st q=q3 & (|.q.|) <r1;
   hence contradiction by A13;
  end;
  then q3 in P by A14,XBOOLE_0:def 4;
  then A15:L1 meets P by A5,A13,XBOOLE_0:3;
  consider q4 being Point of TOP-REAL 2 such that
  A16:  q4 in B & |.q4.|>=r1 by A1,Th40;
    q4 in the carrier of TOP-REAL 2;
  then A17:q4 in REAL 2 by EUCLID:25;
    now assume q4 in {q where q is Point of TOP-REAL 2: (|.q.|) <r1};
    then ex q being Point of TOP-REAL 2 st q=q4 & (|.q.|) <r1;
   hence contradiction by A16;
  end;
  then q4 in P by A17,XBOOLE_0:def 4;
  then B meets P by A16,XBOOLE_0:3;
 hence B=LeftComp f by A4,A5,A11,A12,A15,Th100;
end;

theorem Th102:
RightComp SpStSeq D c= BDD (L~SpStSeq D) & RightComp SpStSeq D is Bounded
proof
  set f=SpStSeq D;
  set A=L~SpStSeq D;
  A1:RightComp f is_a_component_of A`by GOBOARD9:def 2;
  A2:now assume
A3:  not (RightComp f) is Bounded;
      LeftComp f misses RightComp f by SPRECT_1:96;
   hence contradiction by A1,A3,Th101;
  end;
  then A4:RightComp f is_inside_component_of A by A1,Def3;
    RightComp f c=
  union{B where B is Subset of TOP-REAL 2: B is_inside_component_of A}
   proof let x;assume A5:x in RightComp f;
       RightComp f in
     {B where B is Subset of TOP-REAL 2: B is_inside_component_of A} by A4;
    hence x in union{B where B is Subset of TOP-REAL 2:
                            B is_inside_component_of A} by A5,TARSKI:def 4;
   end;
  hence RightComp f c= BDD (L~SpStSeq D) by Def5;
  thus thesis by A2;
end;

theorem Th103:
 LeftComp SpStSeq D = UBD (L~SpStSeq D) &
 RightComp SpStSeq D = BDD (L~SpStSeq D)
proof
 set f=SpStSeq D;
  A1:(L~f)`=LeftComp f \/ RightComp f by GOBRD12:11;
  A2:RightComp f c= BDD (L~SpStSeq D) by Th102;
  A3:LeftComp f c= UBD (L~SpStSeq D) by Th99;
  A4:now assume not RightComp f = BDD (L~SpStSeq D);
    then not BDD (L~SpStSeq D) c= RightComp f by A2,XBOOLE_0:def 10;
    then consider z being set such that
    A5: z in BDD (L~SpStSeq D) & not z in RightComp f by TARSKI:def 3;
      BDD (L~f) c= (L~f)` by Th29;
    then z in LeftComp f or z in RightComp f by A1,A5,XBOOLE_0:def 2;
    then (BDD L~f) meets (UBD L~f) by A3,A5,XBOOLE_0:3;
   hence contradiction by Th28;
  end;
    now assume not LeftComp f = UBD (L~SpStSeq D);
    then not UBD (L~SpStSeq D) c= LeftComp f by A3,XBOOLE_0:def 10;
    then consider z being set such that
    A6: z in UBD (L~SpStSeq D) & not z in LeftComp f by TARSKI:def 3;
      UBD (L~f) c= (L~f)` by Th30;
    then z in LeftComp f or z in RightComp f by A1,A6,XBOOLE_0:def 2;
    then (BDD L~f) meets (UBD L~f) by A2,A6,XBOOLE_0:3;
   hence contradiction by Th28;
  end;
 hence thesis by A4;
end;

theorem Th104:
UBD (L~SpStSeq D)<>{} &
UBD (L~SpStSeq D) is_outside_component_of (L~SpStSeq D) &
BDD (L~SpStSeq D)<>{} &
BDD (L~SpStSeq D) is_inside_component_of (L~SpStSeq D)
proof
  A1:BDD (L~SpStSeq D)=RightComp (SpStSeq D) by Th103;
  A2:UBD (L~SpStSeq D)=LeftComp (SpStSeq D) by Th103;
 hence UBD (L~SpStSeq D)<>{};
  set f=SpStSeq D;
  A3:LeftComp f is_a_component_of (L~f)` by GOBOARD9:def 1;
    not LeftComp f is Bounded by Th98;
 hence UBD (L~SpStSeq D) is_outside_component_of (L~SpStSeq D)
                                          by A2,A3,Def4;
 thus BDD (L~SpStSeq D)<>{} by A1;
  A4:RightComp (SpStSeq D) is_a_component_of (L~f)` by GOBOARD9:def 2;
    RightComp (SpStSeq D) is Bounded by Th102;
 hence BDD (L~SpStSeq D) is_inside_component_of (L~SpStSeq D) by A1,A4,Def3;
end;

begin :: Jordan property and boundary property

theorem Th105:for G being non empty TopSpace,
A being Subset of G st A`<>{}
holds A is boundary iff
for x being set,V being Subset of G st x in A & x in V &
V is open
ex B being Subset of G st B is_a_component_of A` & V meets B
proof let G be non empty TopSpace,A be Subset of G;
 assume A1:A`<>{};
hereby assume A2: A is boundary;
  let x be set,V be Subset of G;
   assume A3:x in A & x in V & V is open;
      A` is dense by A2,TOPS_1:def 4;
    then A4: Cl (A`)=[#] G by TOPS_1:def 3;
    reconsider A2=A` as Subset of G;
      A c= the carrier of G;
    then A c= [#] G by PRE_TOPC:12;
then A2 meets V by A3,A4,PRE_TOPC:def 13;
    then consider z being set such that
A5:  z in A` & z in V by XBOOLE_0:3;
      z in [#]G \ A by A5,PRE_TOPC:17;
    then z in (the carrier of G) \ A by PRE_TOPC:12;
    then z in A` by SUBSET_1:def 5;
    then reconsider p=z as Point of G|A` by JORDAN1:1;
    reconsider A1=A` as non empty Subset of G by A1;
A6: G|A1 is non empty;
    then A7:skl p is_a_component_of G|A` by CONNSP_1:43;
      skl p c= the carrier of G|A`;
    then skl p c= A` by JORDAN1:1;
    then reconsider B0=skl p as Subset of G by XBOOLE_1:1;
    A8:B0 is_a_component_of A` by A7,CONNSP_1:def 6;
     p in skl p by A6,CONNSP_1:40;
    then p in V /\ B0 by A5,XBOOLE_0:def 3;
    then V meets B0 by XBOOLE_0:4;
   hence ex B being Subset of G st
     B is_a_component_of A` & V meets B by A8;
 end;
assume A9:for x being set,V being Subset of G
 st x in A & x in V & V is open
 ex B being Subset of G st B is_a_component_of A` & V meets B;
    the carrier of G c= Cl (A`)
  proof let z be set;assume A10:z in the carrier of G;
    per cases;
    suppose A11:z in A;
      for G1 being Subset of G st G1 is open holds
                        z in G1 implies (A`) meets G1
     proof let G1 be Subset of G;
      assume A12:G1 is open;
      assume z in G1;
        then consider B being Subset of G such that
A13:    B is_a_component_of A` & G1 meets B by A9,A11,A12;
        consider B1 being Subset of G|A` such that
        A14: B1 = B & B1 is_a_component_of (G|A`) by A13,CONNSP_1:def 6;
          B1 c= the carrier of (G|A`);
        then B1 c= A` by JORDAN1:1;
        then G1 /\ B c= G1 /\ A` by A14,XBOOLE_1:26;
then A15:     G1 /\ B c= (A`) /\ G1;
          G1 /\ B <> {} by A13,XBOOLE_0:def 7;
       then (A`) /\ G1 <> {}G by A15,XBOOLE_1:3;
       hence (A`) meets G1 by XBOOLE_0:def 7;
     end;
   hence z in Cl (A`) by A10,PRE_TOPC:def 13;
    suppose not z in A;
     then z in (the carrier of G) \ A by A10,XBOOLE_0:def 4;
     then z in A` by SUBSET_1:def 5;
     then A16: z in A`;
      A` c= Cl(A`) by PRE_TOPC:48;
   hence z in Cl (A`) by A16;
  end;
  then A17: [#] G c= Cl (A`) by PRE_TOPC:12;
    Cl (A`) c= the carrier of G;
  then Cl (A`) c= [#] G by PRE_TOPC:12;
  then Cl (A`)=[#] G by A17,XBOOLE_0:def 10;
  then A` is dense by TOPS_1:def 3;
 hence A is boundary by TOPS_1:def 4;
end;

theorem Th106:for A being Subset of TOP-REAL 2 st A`<>{}
holds A is boundary & A is Jordan iff
 ex A1,A2 being Subset of TOP-REAL 2 st
    A` = A1 \/ A2 & A1 misses A2 & (Cl A1) \ A1 = (Cl A2) \ A2 &
     A=(Cl A1) \ A1 &
     for C1,C2 being Subset of (TOP-REAL 2)|A`
     st C1 = A1 & C2 = A2 holds
      C1 is_a_component_of (TOP-REAL 2)|A` &
      C2 is_a_component_of (TOP-REAL 2)|A`
proof let A be Subset of TOP-REAL 2;
assume A1: A`<>{};
  hereby assume A2:A is boundary & A is Jordan;
    then consider A1,A2 being Subset of TOP-REAL 2 such that
    A3: A` = A1 \/ A2 & A1 misses A2 & (Cl A1) \ A1 = (Cl A2) \ A2 &
     for C1,C2 being Subset of (TOP-REAL 2)|A`
     st C1 = A1 & C2 = A2 holds
      C1 is_a_component_of (TOP-REAL 2)|A` &
      C2 is_a_component_of (TOP-REAL 2)|A` by JORDAN1:def 2;
         A1 c= A` by A3,XBOOLE_1:7;
       then A1 c= the carrier of (TOP-REAL 2)|A` by JORDAN1:1;
       then reconsider D1=A1 as Subset of (TOP-REAL 2)|A`;
         A2 c= A` by A3,XBOOLE_1:7;
       then A2 c= the carrier of (TOP-REAL 2)|A` by JORDAN1:1;
       then reconsider D2=A2 as Subset of (TOP-REAL 2)|A`;
        D1=A1 & D2=A2;
      then A4:D1 is_a_component_of (TOP-REAL 2)|A` &
      D2 is_a_component_of (TOP-REAL 2)|A` by A3;
       A=(A1 \/ A2)` by A3;
     then A5:A=A1` /\ A2` by SUBSET_1:29;
       (Cl A1)\ A1 c= (the carrier of (TOP-REAL 2))\ A1 by XBOOLE_1:33;
     then A6:(Cl A1) \A1 c= A1` by SUBSET_1:def 5;
       (Cl A2)\ A2 c= (the carrier of (TOP-REAL 2))\ A2 by XBOOLE_1:33;
     then (Cl A2) \A2 c= A2` by SUBSET_1:def 5;
     then A7:(Cl A1) \A1 c= A by A3,A5,A6,XBOOLE_1:19;
       A c= (Cl A1) \ A1
     proof let z be set;assume A8:z in A;
       then not z in (the carrier of (TOP-REAL 2)) \A by XBOOLE_0:def 4;
       then not z in A` by SUBSET_1:def 5;
       then A9:not (z in A1 or z in A2) by A3,XBOOLE_0:def 2;
         for G being Subset of (TOP-REAL 2) st G is open holds
                        z in G implies (A1 \/ A2) meets G
        proof let G be Subset of (TOP-REAL 2);
         assume A10: G is open;
          hereby assume z in G;
          then consider B being Subset of TOP-REAL 2 such that
          A11: B is_a_component_of A` & G meets B by A1,A2,A8,A10,Th105;
          consider B1 being Subset of (TOP-REAL 2)|A` such that
A12: B1 = B & B1 is_a_component_of (TOP-REAL 2)|A` by A11,CONNSP_1:def 6;
            now per cases by A4,A12,CONNSP_1:36;
          case B1=D1;
           hence B1 c= A1 \/ A2 by XBOOLE_1:7;
          case B1,D1 are_separated;
            then A13:Cl B1 misses D1 or
            B1 misses Cl D1 by CONNSP_1:def 1;
            A14:B1 is closed by A12,CONNSP_1:35;
              D1 is closed by A4,CONNSP_1:35;
            then B1 misses D1 by A13,A14,PRE_TOPC:52;
            then A15:B1 /\ D1={} by XBOOLE_0:def 7;
              B1 c= the carrier of (TOP-REAL 2)|A`;
            then B1 c= A` by JORDAN1:1;
            then B1 = B1 /\ A` by XBOOLE_1:28
            .=B1 /\ A1 \/ B1 /\ A2 by A3,XBOOLE_1:23
            .= B1 /\ A2 by A15;
            then A16:B1 c= A2 by XBOOLE_1:17;
              A2 c= A1 \/ A2 by XBOOLE_1:7;
           hence B1 c= A1 \/ A2 by A16,XBOOLE_1:1;
          end;
then A17:       G /\ B c= G /\ (A1 \/ A2) by A12,XBOOLE_1:26;
            G /\ B <> {} by A11,XBOOLE_0:def 7;
          then (A1 \/ A2) /\ G <> {} by A17,XBOOLE_1:3;
          hence (A1 \/ A2) meets G by XBOOLE_0:def 7;
         end;
        end;
       then z in Cl (A1 \/ A2) by A8,PRE_TOPC:def 13;
       then z in (Cl A1) \/ Cl A2 by PRE_TOPC:50;
       then z in Cl A1 or z in Cl A2 by XBOOLE_0:def 2;
      hence z in (Cl A1) \ A1 by A3,A9,XBOOLE_0:def 4;
     end;
    then A=Cl A1 \ A1 by A7,XBOOLE_0:def 10;
   hence ex A1,A2 being Subset of TOP-REAL 2
   st A` = A1 \/ A2 & A1 misses A2
     & (Cl A1) \ A1 = (Cl A2) \ A2 & A=(Cl A1) \ A1 &
     for C1,C2 being Subset of (TOP-REAL 2)|A`
     st C1 = A1 & C2 = A2 holds
      C1 is_a_component_of (TOP-REAL 2)|A` &
      C2 is_a_component_of (TOP-REAL 2)|A` by A3;
  end;
  hereby assume
  ( ex A1,A2 being Subset of TOP-REAL 2 st
    A` = A1 \/ A2 & A1 misses A2 & (Cl A1) \ A1 = (Cl A2) \ A2 &
     A=(Cl A1) \ A1 &
     for C1,C2 being Subset of (TOP-REAL 2)|A`
     st C1 = A1 & C2 = A2 holds
      C1 is_a_component_of (TOP-REAL 2)|A` &
      C2 is_a_component_of (TOP-REAL 2)|A`);
  then consider A1,A2 being Subset of TOP-REAL 2
      such that A18: A` = A1 \/ A2 & A1 misses A2
     & (Cl A1) \ A1 = (Cl A2) \ A2 & A=(Cl A1) \ A1 &
     for C1,C2 being Subset of (TOP-REAL 2)|A`
     st C1 = A1 & C2 = A2 holds
      C1 is_a_component_of (TOP-REAL 2)|A` &
      C2 is_a_component_of (TOP-REAL 2)|A`;
       for x being set,V being Subset of TOP-REAL 2 st x in A & x in V &
     V is open ex B being Subset of TOP-REAL 2 st
     B is_a_component_of A` & V meets B
     proof let x be set,V be Subset of TOP-REAL 2;
      assume A19:x in A & x in V & V is open;
        then x in Cl A1 by A18,XBOOLE_0:def 4;
       then A20: A1 meets V by A19,PRE_TOPC:def 13;
         A1 c= A` by A18,XBOOLE_1:7;
       then A1 c= the carrier of (TOP-REAL 2)|A` by JORDAN1:1;
       then reconsider D1=A1 as Subset of (TOP-REAL 2)|A`;
         A2 c= A` by A18,XBOOLE_1:7;
       then A2 c= the carrier of (TOP-REAL 2)|A` by JORDAN1:1;
       then reconsider D2=A2 as Subset of (TOP-REAL 2)|A`;
        D1=A1 & D2=A2;
      then D1 is_a_component_of (TOP-REAL 2)|A` &
      D2 is_a_component_of (TOP-REAL 2)|A` by A18;
       then A1 is_a_component_of A` by CONNSP_1:def 6;
      hence ex B being Subset of TOP-REAL 2 st
       B is_a_component_of A` & V meets B by A20;
     end;
    hence A is boundary & A is Jordan by A1,A18,Th105,JORDAN1:def 2;
   end;
end;

theorem Th107:for p being Point of TOP-REAL n,
P being Subset of TOP-REAL n st n>=1 & P={p} holds P is boundary
proof let p be Point of TOP-REAL n,
  P be Subset of TOP-REAL n;
 assume that A1:n>=1 and
  A2: P={p};
    the carrier of (TOP-REAL n) c= Cl (P`)
  proof let z be set;assume A3:z in the carrier of TOP-REAL n;
    per cases;
    suppose A4:z=p;
     reconsider ez=z as Point of Euclid n by A3,TOPREAL3:13;
      for G1 being Subset of (TOP-REAL n) st G1 is open holds
                        z in G1 implies (P`) meets G1
     proof let G1 be Subset of TOP-REAL n;
      assume A5:G1 is open;
      thus z in G1 implies (P`) meets G1
      proof assume A6:z in G1;
          TOP-REAL n=TopSpaceMetr(Euclid n) by EUCLID:def 8;
        then consider r be real number such that
        A7:r>0 & Ball(ez,r) c= G1 by A5,A6,TOPMETR:22;
        reconsider r as Real by XREAL_0:def 1;
          0<n by A1,AXIOMS:22;
        then A8:0<sqrt n by SQUARE_1:93;
        A9: r/2>0 by A7,REAL_2:127;
        set p2=p-(r/2/sqrt n)*(1.REAL n);
        reconsider ep2=p2 as Point of Euclid n by TOPREAL3:13;
        A10: |.p-p2.|=|.p-p+(r/2/sqrt n)*(1.REAL n).| by EUCLID:51
                .=|.(r/2/sqrt n)*(1.REAL n)+p-p.| by EUCLID:49
                .=|.(r/2/sqrt n)*(1.REAL n).| by EUCLID:52
                .=abs(r/2/sqrt n)*|.(1.REAL n).| by TOPRNS_1:8
                .=abs(r/2/sqrt n)*(sqrt n) by Th37
                .=abs(r/2)/abs(sqrt n)*(sqrt n) by ABSVALUE:16
                .=abs(r/2)/(sqrt n)*(sqrt n) by A8,ABSVALUE:def 1
                .=abs(r/2) by A8,XCMPLX_1:88
                .=r/2 by A9,ABSVALUE:def 1;
             r/2<r by A7,SEQ_2:4;
           then dist(ez,ep2)<r by A4,A10,JGRAPH_1:45;
           then A11:p2 in Ball(ez,r) by METRIC_1:12;
             p<>p2 by A9,A10,TOPRNS_1:29;
           then not p2 in P by A2,TARSKI:def 1;
           then p2 in (the carrier of TOP-REAL n) \P by XBOOLE_0:def 4;
           then p2 in P` by SUBSET_1:def 5;
           then p2 in P`;
       hence (P`) meets G1 by A7,A11,XBOOLE_0:3;
      end;
     end;
     hence z in Cl (P`) by A3,PRE_TOPC:def 13;
    suppose z<>p; then not z in P by A2,TARSKI:def 1;
      then z in (the carrier of (TOP-REAL n)) \ P by A3,XBOOLE_0:def 4;
      then z in P` by SUBSET_1:def 5;
      then A12: z in P`;
        P` c= Cl(P`) by PRE_TOPC:48;
     hence z in Cl (P`) by A12;
  end;
  then A13: [#] (TOP-REAL n) c= Cl (P`) by PRE_TOPC:12;
    Cl (P`) c= the carrier of (TOP-REAL n);
  then Cl (P`) c= [#] (TOP-REAL n) by PRE_TOPC:12;
  then Cl (P`)=[#] (TOP-REAL n) by A13,XBOOLE_0:def 10;
  then P` is dense by TOPS_1:def 3;
 hence P is boundary by TOPS_1:def 4;
end;

theorem Th108:for p,q being Point of TOP-REAL 2,r st
p`1=q`2 & -p`2=q`1 & p=r*q holds p`1=0 & p`2=0 & p=0.REAL 2
proof let p,q be Point of TOP-REAL 2,r;
 assume p`1=q`2 & -p`2=q`1 & p=r*q;
  then p=|[r*(-p`2),r*(p`1)]| by EUCLID:61;
  then A1:p`1=r*(-p`2)& p`2=r*(p`1) by EUCLID:56;
  then p`1=-(r*(r*(p`1))) by XCMPLX_1:175 .=-(r*r*(p`1)) by XCMPLX_1:4;
  then 1 *p`1+r*r*(p`1)=0 by XCMPLX_0:def 6;
  then A2:(1+r*r)*p`1=0 by XCMPLX_1:8;
    r*r>=0 by REAL_1:93;
then A3:1+r*r>0+0 by REAL_1:67;
 hence A4:p`1=0 by A2,XCMPLX_1:6;
    p`2=r*(-(r*(p`2))) by A1,XCMPLX_1:175
  .=-r*(r*(p`2)) by XCMPLX_1:175
  .=-(r*r*(p`2)) by XCMPLX_1:4;
  then 1 * p`2+r*r*(p`2)=0 by XCMPLX_0:def 6;
  then (1+r*r)*p`2=0 by XCMPLX_1:8;
 hence p`2=0 by A3,XCMPLX_1:6;
 hence p=0.REAL 2 by A4,EUCLID:57,58;
end;

theorem Th109:for q1,q2 being Point of TOP-REAL 2 holds
LSeg(q1,q2) is boundary
proof let q1,q2 be Point of TOP-REAL 2;
  per cases;
  suppose q1=q2; then LSeg(q1,q2)={q1} by TOPREAL1:7;
   hence LSeg(q1,q2) is boundary by Th107;
  suppose A1:q1<>q2;
    set P=LSeg(q1,q2);
    the carrier of (TOP-REAL 2) c= Cl (P`)
  proof let z be set;assume A2:z in the carrier of TOP-REAL 2;
    per cases;
    suppose z in P;
     then z in {(1-s)*q1+s*q2:0<=s & s<=1} by TOPREAL1:def 4;
     then consider s being Real such that
     A3: z=(1-s)*q1+s*q2 & (0<=s & s<=1);
     set p=(1-s)*q1+s*q2;
     set p1=q1-q2;
     A4:now assume |.p1.|=0; then p1=0.REAL 2 by TOPRNS_1:25;
      hence contradiction by A1,EUCLID:47;
     end;
     A5: |.p1.|>=0 by TOPRNS_1:26;
     reconsider ez=z as Point of Euclid 2 by A2,TOPREAL3:13;
      for G1 being Subset of (TOP-REAL 2) st G1 is open holds
                        z in G1 implies (P`) meets G1
     proof let G1 be Subset of TOP-REAL 2;
      assume A6:G1 is open;
      thus z in G1 implies (P`) meets G1
      proof assume A7:z in G1;
          TOP-REAL 2=TopSpaceMetr(Euclid 2) by EUCLID:def 8;
        then consider r be real number such that
        A8:r>0 & Ball(ez,r) c= G1 by A6,A7,TOPMETR:22;
        reconsider r as Real by XREAL_0:def 1;
        A9: r/2>0 by A8,REAL_2:127;
        set p2=(r/2/|.p1.|)*|[-p1`2,p1`1]| +p;
        A10: |[-p1`2,p1`1]|`1=-p1`2 & |[-p1`2,p1`1]|`2=p1`1 by EUCLID:56;
        reconsider ep2=p2 as Point of Euclid 2 by TOPREAL3:13;
        A11: |.p-p2.|=|.p-(r/2/|.p1.|)*|[-p1`2,p1`1]| -p.| by EUCLID:50
                .=|.p+-(r/2/|.p1.|)*|[-p1`2,p1`1]| -p.| by EUCLID:45
                .=|.-(r/2/|.p1.|)*|[-p1`2,p1`1]|.| by EUCLID:52
                .=|.(r/2/|.p1.|)*|[-p1`2,p1`1]|.| by TOPRNS_1:27
                .=abs(r/2/|.p1.|)*|.|[-p1`2,p1`1]|.| by TOPRNS_1:8
                .=abs(r/2/|.p1.|)*(sqrt ((-p1`2)^2+(p1`1)^2))
                                         by A10,JGRAPH_1:47
                .=abs(r/2/|.p1.|)*(sqrt ((p1`1)^2+(p1`2)^2))
                by SQUARE_1:61
                .=abs(r/2/|.p1.|)*|.p1.| by JGRAPH_1:47
                .=abs(r/2)/(abs(|.p1.|))*|.p1.| by ABSVALUE:16
                .=abs(r/2)/(|.p1.|)*|.p1.| by A5,ABSVALUE:def 1
                .=abs (r/2) by A4,XCMPLX_1:88
                .=r/2 by A9,ABSVALUE:def 1;
             r/2<r by A8,SEQ_2:4;
           then dist(ez,ep2)<r by A3,A11,JGRAPH_1:45;
           then A12:p2 in Ball(ez,r) by METRIC_1:12;
            now assume p2 in P;
             then p2 in {(1-s2)*q1+s2*q2:0<=s2 & s2<=1} by TOPREAL1:def 4;
             then consider s2 being Real such that
             A13: p2=(1-s2)*q1+s2*q2 & (0<=s2 & s2<=1);
             A14:p2-p=(1-s2)*q1+s2*q2 -(1-s)*q1-s*q2 by A13,EUCLID:50
                 .=(1-s2)*q1-(1-s)*q1+s2*q2 -s*q2 by Th9
                 .=((1-s2)-(1-s))*q1+s2*q2 -s*q2 by EUCLID:54
                 .=(s-s2)*q1+s2*q2 -s*q2 by UNIFORM1:1
                 .=(s-s2)*q1+(s2*q2 -s*q2) by EUCLID:49
                 .=(s-s2)*q1+(s2-s)*q2 by EUCLID:54
                 .=(s-s2)*q1+(-(s-s2))*q2 by XCMPLX_1:143
                 .=(s-s2)*q1+-(s-s2)*q2 by EUCLID:44
                 .=(s-s2)*q1-(s-s2)*q2 by EUCLID:45
                 .=(s-s2)*p1 by EUCLID:53;
               p2-p = (r/2/|.p1.|)*|[-p1`2,p1`1]| by EUCLID:52;
             then A15:1/(s-s2)*(s-s2)*p1= 1/(s-s2)*((r/2/|.p1.|)*|[-p1`2,p1`1]|
)
                                           by A14,EUCLID:34;
               now assume s-s2=0;
               then s=0+s2 by XCMPLX_1:27;
               then (r/2/|.p1.|)*|[-p1`2,p1`1]|=p-p by A13,EUCLID:52;
               then A16:(r/2/|.p1.|)*|[-p1`2,p1`1]|=0.REAL 2 by EUCLID:46;
                 (|.p1.|)" <>0 by A4,XCMPLX_1:203;
           then A17:2"*(|.p1.|)" <>0 by XCMPLX_1:6;
                 r/2=r*2" by XCMPLX_0:def 9;
               then (r/2/|.p1.|)=r*2"*(|.p1.|)" by XCMPLX_0:def 9 .=r*(2"*(|.
p1.|)")
                                        by XCMPLX_1:4;
               then (r/2/|.p1.|)<>0 by A8,A17,XCMPLX_1:6;
               then |[-p1`2,p1`1]|=0.REAL 2 by A16,EUCLID:35;
           then A18:(0.REAL 2)`1=-p1`2 & (0.REAL 2)`2=p1`1 by EUCLID:56;
                 (0.REAL 2)`1=0 & (0.REAL 2)`2=0 by EUCLID:56,58;
               then p1`2=0 & p1`1=0 by A18,XCMPLX_1:135;
               then p1=|[0,0]| by EUCLID:57;
              hence contradiction by A1,EUCLID:47,58;
             end;
             then 1 *p1= 1/(s-s2)*((r/2/|.p1.|)*|[-p1`2,p1`1]|) by A15,XCMPLX_1
:107;
             then p1= 1/(s-s2)*((r/2/|.p1.|)*|[-p1`2,p1`1]|) by EUCLID:33;
             then A19:p1= 1/(s-s2)*(r/2/|.p1.|)*|[-p1`2,p1`1]| by EUCLID:34;
               p1`1=(|[-p1`2,p1`1]|)`2 & -p1`2=(|[-p1`2,p1`1]|)`1
                                    by EUCLID:56;
             then p1=0.REAL 2 by A19,Th108;
            hence contradiction by A1,EUCLID:47;
           end;
           then p2 in (the carrier of TOP-REAL 2) \P by XBOOLE_0:def 4;
           then p2 in P` by SUBSET_1:def 5;
           then p2 in P`;
       hence (P`) meets G1 by A8,A12,XBOOLE_0:3;
      end;
     end;
     hence z in Cl (P`) by A2,PRE_TOPC:def 13;
    suppose not z in P;
      then z in (the carrier of (TOP-REAL 2)) \ P by A2,XBOOLE_0:def 4;
      then z in P` by SUBSET_1:def 5;
      then A20: z in P`;
        P` c= Cl(P`) by PRE_TOPC:48;
     hence z in Cl (P`) by A20;
  end;
    then A21: [#] (TOP-REAL 2) c= Cl (P`) by PRE_TOPC:12;
      Cl (P`) c= the carrier of (TOP-REAL 2);
    then Cl (P`) c= [#] (TOP-REAL 2) by PRE_TOPC:12;
    then Cl (P`)=[#] (TOP-REAL 2) by A21,XBOOLE_0:def 10;
    then P` is dense by TOPS_1:def 3;
   hence LSeg(q1,q2) is boundary by TOPS_1:def 4;
end;

definition let q1,q2 be Point of TOP-REAL 2;
 cluster LSeg(q1,q2) -> boundary;
coherence by Th109;
end;

theorem Th110:for f being FinSequence of TOP-REAL 2 holds L~f is boundary
proof let f be FinSequence of TOP-REAL 2;
  A1: L~f=union { LSeg(f,i) : 1 <= i & i+1 <= len f } by TOPREAL1:def 6;
    union { LSeg(f,i) : 1 <= i & i+1 <= 0 } c= {}
  proof let z be set;assume
      z in union { LSeg(f,i) : 1 <= i & i+1 <= 0 };
    then consider x such that
    A2:z in x & x in { LSeg(f,i) : 1 <= i & i+1 <= 0 } by TARSKI:def 4;
    consider i such that
    A3:x=LSeg(f,i) &(1 <= i & i+1 <= 0) by A2;
      1<i+1 by A3,NAT_1:38;
   hence z in {} by A3,AXIOMS:22;
  end;
  then A4:union { LSeg(f,i) : 1 <= i & i+1 <= 0 }={}(TOP-REAL 2) by XBOOLE_1:3;
  defpred P[Nat] means for R1 being Subset of TOP-REAL 2 st
    R1=union { LSeg(f,i) : 1 <= i & i+1 <= $1 } holds R1 is boundary;
  A5: P[0]
  proof let R1 be Subset of TOP-REAL 2;
    assume R1=union { LSeg(f,i) : 1 <= i & i+1 <= 0 };
     then Int R1={} by A4,TOPS_1:47;
   hence thesis by TOPS_1:84;
  end;
  A6:for k being Nat st P[k] holds P[k+1]
  proof let k be Nat;
    assume A7: for R1 being Subset of TOP-REAL 2 st
    R1=union { LSeg(f,i) : 1 <= i & i+1 <= k } holds R1 is boundary;
      union { LSeg(f,i2) : 1 <= i2 & i2+1 <= k } c= the carrier of TOP-REAL 2
    proof let z be set;assume z in union { LSeg(f,i2) : 1 <= i2 & i2+1 <= k };
     then consider x such that
     A8:z in x & x in { LSeg(f,i) : 1 <= i & i+1 <= k } by TARSKI:def 4;
       ex i st x=LSeg(f,i) &(1 <= i & i+1 <= k) by A8;
     hence thesis by A8;
    end;
    then reconsider R3=union { LSeg(f,i2) : 1 <= i2 & i2+1 <= k }
                     as Subset of TOP-REAL 2;
    A9:R3 is boundary by A7;
 A10:now per cases;
     case 1<=k & k+1<=len f;
       then LSeg(f,k)=LSeg(f/.k,f/.(k+1)) by TOPREAL1:def 5;
      hence LSeg(f,k) is boundary;
     case not(1<=k & k+1<=len f);
       then LSeg(f,k)={} by TOPREAL1:def 5;
       then Int LSeg(f,k)={}(TOP-REAL 2) by TOPS_1:47;
      hence LSeg(f,k) is boundary by TOPS_1:84;
     end;
    A11: LSeg(f,k) is closed by SPPOL_1:40;
   thus (for R2 being Subset of TOP-REAL 2 st
   R2=union { LSeg(f,i2) : 1 <= i2 & i2+1 <= k+1 } holds R2 is boundary)
   proof let R2 be Subset of TOP-REAL 2;
     assume A12:R2=union { LSeg(f,i2) : 1 <= i2 & i2+1 <= k+1 };
       R2=R3 \/ LSeg(f,k)
     proof
      thus R2 c= R3 \/ LSeg(f,k)
       proof let z be set;assume z in R2;
         then consider x such that
         A13:z in x & x in { LSeg(f,i2) :
         1 <= i2 & i2+1 <= k+1 } by A12,TARSKI:def 4;
         consider i2 such that
         A14:x=LSeg(f,i2) &(1 <= i2 & i2+1 <= k+1) by A13;
           now per cases;
         case i2+1<=k;
           then z in x & x in { LSeg(f,j) : 1 <= j & j+1 <= k } by A13,A14;
          hence z in R3 or z in LSeg(f,k) by TARSKI:def 4;
         case i2+1>k; then k+1<=i2+1 by NAT_1:38;
           then i2+1=k+1 by A14,AXIOMS:21;
           then i2+1-1=k by XCMPLX_1:26;
          hence z in R3 or z in LSeg(f,k) by A13,A14,XCMPLX_1:26;
         end;
        hence z in R3 \/ LSeg(f,k) by XBOOLE_0:def 2;
       end;
        R3 \/ LSeg(f,k) c= R2
      proof
       let z be set;assume A15:z in R3 \/ LSeg(f,k);
       per cases by A15,XBOOLE_0:def 2;
       suppose z in R3;
         then consider x such that
         A16:z in x & x in { LSeg(f,i2) :
         1 <= i2 & i2+1 <= k } by TARSKI:def 4;
         consider i2 such that
         A17:x=LSeg(f,i2) &(1 <= i2 & i2+1 <= k) by A16;
           i2+1<k+1 by A17,NAT_1:38;
         then z in x & x in { LSeg(f,j) : 1 <= j & j+1 <= k+1 } by A16,A17;
        hence z in R2 by A12,TARSKI:def 4;
       suppose A18:z in LSeg(f,k);
          now per cases;
        suppose 1<=k;
         then LSeg(f,k) in { LSeg(f,i2) : 1 <= i2 & i2+1 <= k+1 };
        hence z in R2 by A12,A18,TARSKI:def 4;
        suppose k<1;
        hence z in R2 by A18,TOPREAL1:def 5;
       end;
      hence z in R2;
     end;
    hence thesis;
   end;
   hence thesis by A9,A10,A11,TOPS_1:85;
   end;
  end;
    for j holds P[j]
    from Ind(A5,A6);
 hence L~f is boundary by A1;
end;

definition let f be FinSequence of TOP-REAL 2;
 cluster L~f -> boundary;
coherence by Th110;
end;

theorem Th111:for ep being Point of Euclid n,p,q being Point of TOP-REAL n st
p=ep & q in Ball(ep,r) holds |.p-q.|<r & |.q-p.|<r
proof let ep be Point of Euclid n,p,q be Point of TOP-REAL n;
 assume A1:p=ep & q in Ball(ep,r);
  reconsider eq=q as Point of Euclid n by TOPREAL3:13;
    dist(ep,eq)<r by A1,METRIC_1:12;
 hence thesis by A1,JGRAPH_1:45;
end;

theorem for a being Real,p being Point of TOP-REAL 2
st a>0 & p in L~SpStSeq D holds ex q being Point of TOP-REAL 2 st
q in UBD (L~SpStSeq D) & |.p-q.|<a
proof let a be Real,p be Point of TOP-REAL 2;
 assume A1:a>0 & p in L~SpStSeq D;
  set A=L~SpStSeq D;
    A`<>{} by SPRECT_1:def 3;
 then consider A1,A2 being Subset of TOP-REAL 2 such that
  A2: A` = A1 \/ A2 & A1 misses A2 & (Cl A1) \ A1 = (Cl A2) \ A2 &
     A=(Cl A1) \ A1 &
     for C1,C2 being Subset of (TOP-REAL 2)|A`
     st C1 = A1 & C2 = A2 holds
      C1 is_a_component_of (TOP-REAL 2)|A` &
      C2 is_a_component_of (TOP-REAL 2)|A` by Th106;
   A3:Down(A1,A`)=A1 /\ A` by CONNSP_3:def 5
   .=A1 by A2,XBOOLE_1:21;
   A4:Down(A2,A`)=A2 /\ A` by CONNSP_3:def 5
   .=A2 by A2,XBOOLE_1:21;
   then A5:Down(A1,A`) is_a_component_of (TOP-REAL 2)|A` &
   Down(A2,A`) is_a_component_of (TOP-REAL 2)|A` by A2,A3;
     A is Bounded by Th73;
   then consider B being Subset of TOP-REAL 2 such that
   A6:B is_outside_component_of A & B=UBD A by Th76;
     UBD (L~SpStSeq D) is_a_component_of A` by A6,Def4;
   then consider B1 being Subset of (TOP-REAL 2)|A` such that
   A7: B1 = UBD (L~SpStSeq D) & B1 is_a_component_of (TOP-REAL 2)|A`
                                  by CONNSP_1:def 6;
     B1 c= the carrier of (TOP-REAL 2)|A`;
   then B1 c= [#]((TOP-REAL 2)|A`) by PRE_TOPC:12;
   then A8:UBD (L~SpStSeq D) c= A1 \/ A2 by A2,A7,PRE_TOPC:def 10;
   consider q1 being Element of UBD (L~SpStSeq D);
   A9:UBD (L~SpStSeq D) <>{} by Th104;
   then A10:q1 in UBD (L~SpStSeq D);
   per cases by A8,A10,XBOOLE_0:def 2;
   suppose q1 in A1;
     then B1 /\ Down(A1,A`)<>{}((TOP-REAL 2)|A`) by A3,A7,A9,XBOOLE_0:def 3;
     then B1 meets Down(A1,A`) by XBOOLE_0:def 7;
     then B1=Down(A1,A`) by A5,A7,CONNSP_1:37;
     then A11: p in Cl(UBD (L~SpStSeq D)) by A1,A2,A3,A7,XBOOLE_0:def 4;
     reconsider ep=p as Point of Euclid 2 by TOPREAL3:13;
       Ball(ep,a) c= the carrier of Euclid 2;
     then Ball(ep,a) c= the carrier of TOP-REAL 2 by TOPREAL3:13;
     then reconsider G2=Ball(ep,a) as Subset of TOP-REAL 2;
       the distance of Euclid 2 is Reflexive by METRIC_1:def 7;
     then (the distance of Euclid 2).(ep,ep)=0 by METRIC_1:def 3;
     then dist(ep,ep)=0 by METRIC_1:def 1;
     then A12:p in Ball(ep,a) by A1,METRIC_1:12;
       G2 is open by GOBOARD6:6;
     then (UBD (L~SpStSeq D)) meets G2 by A11,A12,PRE_TOPC:def 13;
     then consider t2 being set such that
     A13:t2 in UBD (L~SpStSeq D) & t2 in G2 by XBOOLE_0:3;
    reconsider qt2=t2 as Point of TOP-REAL 2 by A13;
        t2 in UBD (L~SpStSeq D) & |.p-qt2.|<a by A13,Th111;
    hence thesis;
   suppose q1 in A2;
     then B1 /\ Down(A2,A`)<>{}((TOP-REAL 2)|A`) by A4,A7,A9,XBOOLE_0:def 3;
     then B1 meets Down(A2,A`) by XBOOLE_0:def 7;
     then B1=Down(A2,A`) by A5,A7,CONNSP_1:37;
     then A14: p in Cl(UBD (L~SpStSeq D)) by A1,A2,A4,A7,XBOOLE_0:def 4;
     reconsider ep=p as Point of Euclid 2 by TOPREAL3:13;
       Ball(ep,a) c= the carrier of Euclid 2;
     then Ball(ep,a) c= the carrier of TOP-REAL 2 by TOPREAL3:13;
     then reconsider G2=Ball(ep,a) as Subset of TOP-REAL 2;
       (the distance of (Euclid 2)) is Reflexive by METRIC_1:def 7;
     then (the distance of Euclid 2).(ep,ep)=0 by METRIC_1:def 3;
     then dist(ep,ep)=0 by METRIC_1:def 1;
     then A15:p in Ball(ep,a) by A1,METRIC_1:12;
       G2 is open by GOBOARD6:6;
     then (UBD (L~SpStSeq D)) meets G2 by A14,A15,PRE_TOPC:def 13;
     then consider t2 being set such that
     A16:t2 in (UBD (L~SpStSeq D)) & t2 in G2 by XBOOLE_0:3;
    reconsider qt2=t2 as Point of TOP-REAL 2 by A16;
        t2 in UBD (L~SpStSeq D) & |.p-qt2.|<a by A16,Th111;
    hence thesis;
end;

theorem Th113: REAL 0 = {0.REAL 0}
proof
 thus REAL 0 = 0-tuples_on REAL by EUCLID:def 1
            .= { <*>REAL } by FINSEQ_2:112
            .= { 0 |-> (0 qua Real)} by FINSEQ_2:72
            .= { 0*0 } by EUCLID:def 4
            .= {0.REAL 0} by EUCLID:def 9;
end;

theorem Th114:for A being Subset of TOP-REAL n st
A is Bounded holds BDD A is Bounded
proof let A be Subset of TOP-REAL n;
 assume A is Bounded;
  then consider r being Real such that
  A1: for q being Point of TOP-REAL n st q in A holds |.q.|<r by Th40;
  per cases;
  suppose A2:n>=1;
  set a=r;
    (REAL n)\ {q : (|.q.|) < a } c= REAL n by XBOOLE_1:36;
  then (REAL n)\ {q : (|.q.|) < a } is Subset of TOP-REAL n
                              by EUCLID:25;
  then reconsider P=(REAL n)\ {q : (|.q.|) < a } as Subset of TOP-REAL n
                            ;
  A3:P c= A`
  proof let z;assume A4:z in P;
    then A5:z in (REAL n) & not z in {q : (|.q.|) < a } by XBOOLE_0:def 4;
    reconsider q0=z as Point of TOP-REAL n by A4;
       (|.q0.|) >= a by A5;
     then not q0 in A by A1;
     then q0 in (the carrier of TOP-REAL n)\A by XBOOLE_0:def 4;
   hence z in A` by SUBSET_1:def 5;
  end;
  then A6:Down(P,A`)=P by CONNSP_3:28;
     now per cases;
   suppose n>=2;
  then P is connected by Th61;
  then A7:Down(P,A`) is connected by A6,Th15;
    now assume not BDD A is Bounded;
    then consider q being Point of TOP-REAL n such that
    A8: q in BDD A & not |.q.|<r by Th40;
      q in union{B3 where B3 is Subset of TOP-REAL n:
          B3 is_inside_component_of A} by A8,Def5;
    then consider y such that A9:q in y &
          y in {B3 where B3 is Subset of TOP-REAL n:
          B3 is_inside_component_of A} by TARSKI:def 4;
    consider B3 being Subset of TOP-REAL n such that
A10: y=B3 & B3 is_inside_component_of A by A9;
      B3 is_a_component_of A` by A10,Def3;
    then consider B4 being Subset of (TOP-REAL n)|A` such that
      A11:B4 = B3 & B4 is_a_component_of (TOP-REAL n)|A`
                           by CONNSP_1:def 6;
    A12:q in the carrier of TOP-REAL n;
      for q2 being Point of TOP-REAL n st q2=q holds |.q2.| >= a by A8;
    then q in REAL n & not q in {q2 where q2 is Point of TOP-REAL n:
                       (|.q2.|) < a } by A12,EUCLID:25;
    then q in P by XBOOLE_0:def 4;
    then P /\ B4<>{}((TOP-REAL n)|A`) by A9,A10,A11,XBOOLE_0:def 3;
    then P meets B4 by XBOOLE_0:def 7;
    then A13:P c= B4 by A6,A7,A11,CONNSP_1:38;
      B3 is Bounded by A10,Def3;
    then P is Bounded by A11,A13,Th16;
   hence contradiction by A2,Th62;
  end;
  hence BDD A is Bounded;
  suppose n<2;
    then n<1+1;
    then n<=1 by NAT_1:38;
    then A14:n=1 by A2,AXIOMS:21;
      {q where q is Point of TOP-REAL n:
     for r2 being Real st q=|[r2]| holds r2<=-a} c= the carrier of TOP-REAL n
     proof let z;assume z in {q where q is Point of TOP-REAL n:
       for r2 being Real st q=|[r2]| holds r2<=-a};
       then consider q being Point of TOP-REAL n such that
       A15:q=z & for r2 being Real st q=|[r2]| holds r2<=-a;
      thus z in the carrier of TOP-REAL n by A15;
     end;
    then reconsider P1={q where q is Point of TOP-REAL n:
    for r2 being Real st q=|[r2]| holds r2<=-a} as
    Subset of TOP-REAL n;
      {q where q is Point of TOP-REAL n:
       for r2 being Real st q=|[r2]| holds r2>=a}
       c= the carrier of TOP-REAL n
     proof let z;assume z in {q where q is Point of TOP-REAL n:
       for r2 being Real st q=|[r2]| holds r2>=a};
       then consider q being Point of TOP-REAL n such that
       A16:q=z & for r2 being Real st q=|[r2]| holds r2>=a;
      thus z in the carrier of TOP-REAL n by A16;
     end;
    then reconsider P2={q where q is Point of TOP-REAL n:
    for r2 being Real st q=|[r2]| holds r2>=a} as
    Subset of TOP-REAL n;
    A17:P c= P1 \/ P2
    proof let z;assume A18:z in P;
    then A19: z in REAL n & not z in {q : (|.q.|) < a } by XBOOLE_0:def 4;
    reconsider q0=z as Point of TOP-REAL n by A18;
     A20:(|.q0.|) >= a by A19;
     consider r3 being Real such that
     A21:q0=<*r3*> by A14,JORDAN2B:24;
     A22:q0=|[r3]| by A21,JORDAN2B:def 2;
     then A23:abs r3>=a by A14,A20,Th12;
     per cases by A23,PSCOMP_1:5;
     suppose -a>=r3;
    then for r2 being Real st q0=|[r2]| holds r2<=-a by A22,JORDAN2B:29;
     then q0 in P1;
     hence z in P1 \/ P2 by XBOOLE_0:def 2;
     suppose r3>=a;
    then for r2 being Real st q0=|[r2]| holds r2>=a by A22,JORDAN2B:29;
     then q0 in P2;
     hence z in P1 \/ P2 by XBOOLE_0:def 2;
    end;
      P1 \/ P2 c= P
    proof let z;assume A24:z in P1 \/ P2;
      per cases by A24,XBOOLE_0:def 2;
      suppose z in P1;
        then consider q being Point of TOP-REAL n such that
        A25: q=z & for r2 being Real st q=|[r2]| holds r2<=-a;
     A26: z in the carrier of TOP-REAL n by A25;
         for q2 being Point of TOP-REAL n st q2=z holds |.q2.| >= a
       proof let q2 be Point of TOP-REAL n;
        assume A27:q2=z;
         consider r3 being Real such that
         A28:q2=<*r3*> by A14,JORDAN2B:24;
         A29:q2=|[r3]| by A28,JORDAN2B:def 2;
         then A30:r3<=-a by A25,A27;
           now per cases;
         case A31:a<0;
            abs r3 >=0 by ABSVALUE:5;
          hence abs r3 >=a by A31,AXIOMS:22;
         case a>=0; then -a<=0 by REAL_1:26,50;
          then abs r3=-r3 by A30,Th1;
          hence abs r3>=a by A30,REAL_2:109;
         end;
        hence |.q2.| >= a by A14,A29,Th12;
       end;
        then z in REAL n & not z in {q2 where q2 is Point of TOP-REAL n:
                |.q2.| < a } by A26,EUCLID:25;
       hence z in P by XBOOLE_0:def 4;
      suppose z in P2;
        then consider q being Point of TOP-REAL n such that
        A32: q=z & for r2 being Real st q=|[r2]| holds r2>=a;
       A33:z in the carrier of TOP-REAL n by A32;
         for q2 being Point of TOP-REAL n st q2=z holds |.q2.| >= a
       proof let q2 be Point of TOP-REAL n;
        assume A34:q2=z;
         consider r3 being Real such that
         A35:q2=<*r3*> by A14,JORDAN2B:24;
         A36:q2=|[r3]| by A35,JORDAN2B:def 2;
         then A37:r3>=a by A32,A34;
           now per cases;
         suppose A38:a<0;
             abs r3>=0 by ABSVALUE:5;
          hence abs r3>=a by A38,AXIOMS:22;
         suppose a>=0;
          hence abs r3 >=a by A37,ABSVALUE:def 1;
         end;
        hence |.q2.| >= a by A14,A36,Th12;
       end;
        then z in REAL n & not z in {q2 where q2 is Point of TOP-REAL n:
                |.q2.| < a } by A33,EUCLID:25;
       hence z in P by XBOOLE_0:def 4;
    end;
    then A39:P=P1 \/ P2 by A17,XBOOLE_0:def 10;
    then P1 c= P by XBOOLE_1:7;
    then P1 c= A` by A3,XBOOLE_1:1;
    then A40:Down(P1,A`)=P1 by CONNSP_3:28;
      P2 c= P by A39,XBOOLE_1:7;
    then P2 c= A` by A3,XBOOLE_1:1;
    then A41:Down(P2,A`)=P2 by CONNSP_3:28;
      for w1,w2 being Point of TOP-REAL n
      st w1 in P1 & w2 in P1 holds LSeg(w1,w2) c= P1
     proof let w1,w2 be Point of TOP-REAL n;
      assume A42:w1 in P1 & w2 in P1;
        then consider q1 being Point of TOP-REAL n such that
        A43: q1=w1 & for r2 being Real st q1=|[r2]| holds r2<=-a;
        consider q2 being Point of TOP-REAL n such that
        A44: q2=w2 & for r2 being Real st q2=|[r2]| holds r2<=-a by A42;
      consider r3 being Real such that
      A45:w1=<*r3*> by A14,JORDAN2B:24;
      A46:w1=|[r3]| by A45,JORDAN2B:def 2;
      then A47:r3<=-a by A43;
      consider r4 being Real such that
      A48:w2=<*r4*> by A14,JORDAN2B:24;
      A49:w2=|[r4]| by A48,JORDAN2B:def 2;
      then A50:r4<=-a by A44;
      thus LSeg(w1,w2) c= P1
       proof let z;assume z in LSeg(w1,w2);
         then z in { (1-r2)*w1 + r2*w2 : 0 <= r2 & r2 <= 1 } by TOPREAL1:def 4
;
         then consider r2 such that
         A51:z=(1-r2)*w1 + r2*w2 &( 0 <= r2 & r2 <= 1);
         A52:(1-r2)*w1=|[(1-r2)*r3]| by A14,A46,JORDAN2B:26;
           r2*w2=|[r2*r4]| by A14,A49,JORDAN2B:26;
         then A53:z=|[(1-r2)*r3+r2*r4]| by A14,A51,A52,JORDAN2B:27;
       reconsider q4=z as Point of TOP-REAL n by A51;
         for r5 being Real st q4=|[r5]| holds r5<=-a
       proof let r5 be Real;
        assume q4=|[r5]|;
         then A54:r5=(1-r2)*r3+r2*r4 by A53,JORDAN2B:29;
           1-r2>=0 by A51,SQUARE_1:12;
         then A55:(1-r2)*r3<=(1-r2)*(-a) by A47,AXIOMS:25;
         A56:r2*r4<=r2*(-a) by A50,A51,AXIOMS:25;
           (1-r2)*(-a)+r2*(-a)=((1-r2)+r2)*(-a) by XCMPLX_1:8
           .=1 *(-a) by XCMPLX_1:27 .=-a;
        hence r5<=-a by A54,A55,A56,REAL_1:55;
       end;
        hence z in P1;
       end;
     end;
    then P1 is convex by JORDAN1:def 1;
    then A57:P1 is connected by Th14;
      for w1,w2 being Point of TOP-REAL n
      st w1 in P2 & w2 in P2 holds LSeg(w1,w2) c= P2
     proof let w1,w2 be Point of TOP-REAL n;
      assume A58:w1 in P2 & w2 in P2;
        then consider q1 being Point of TOP-REAL n such that
        A59: q1=w1 & for r2 being Real st q1=|[r2]| holds r2>=a;
        consider q2 being Point of TOP-REAL n such that
        A60: q2=w2 &
       for r2 being Real st q2=|[r2]| holds r2>=a by A58;
      consider r3 being Real such that
      A61:w1=<*r3*> by A14,JORDAN2B:24;
      A62:w1=|[r3]| by A61,JORDAN2B:def 2;
      then A63:r3>=a by A59;
      consider r4 being Real such that
      A64:w2=<*r4*> by A14,JORDAN2B:24;
      A65:w2=|[r4]| by A64,JORDAN2B:def 2;
      then A66:r4>=a by A60;
      thus LSeg(w1,w2) c= P2
       proof let z;assume z in LSeg(w1,w2);
         then z in { (1-r2)*w1 + r2*w2 : 0 <= r2 & r2 <= 1 } by TOPREAL1:def 4
;
         then consider r2 such that
         A67:z=(1-r2)*w1 + r2*w2 &( 0 <= r2 & r2 <= 1);
         A68:(1-r2)*w1=|[(1-r2)*r3]| by A14,A62,JORDAN2B:26;
           r2*w2=|[r2*r4]| by A14,A65,JORDAN2B:26;
         then A69:z=|[(1-r2)*r3+r2*r4]| by A14,A67,A68,JORDAN2B:27;
       reconsider q4=z as Point of TOP-REAL n by A67;
         for r5 being Real st q4=|[r5]| holds r5>=a
       proof let r5 be Real;
        assume q4=|[r5]|;
         then A70:r5=(1-r2)*r3+r2*r4 by A69,JORDAN2B:29;
           1-r2>=0 by A67,SQUARE_1:12;
         then A71:(1-r2)*r3>=(1-r2)*(a) by A63,AXIOMS:25;
         A72:r2*r4>=r2*(a) by A66,A67,AXIOMS:25;
           (1-r2)*(a)+r2*(a)=((1-r2)+r2)*(a) by XCMPLX_1:8
           .=1 *(a) by XCMPLX_1:27 .=a;
        hence r5>=a by A70,A71,A72,REAL_1:55;
       end;
        hence z in P2;
       end;
     end;
    then P2 is convex by JORDAN1:def 1;
    then A73:P2 is connected by Th14;
    A74:now assume P1 is Bounded;
      then consider r being Real such that
      A75: for q being Point of TOP-REAL n st q in P1 holds |.q.|<r by Th40;
      set p3=|[-(abs(r)+abs(a))]|;
        for r5 being Real st p3=|[r5]| holds r5<=-a
      proof let r5 be Real;assume p3=|[r5]|;
        then A76:r5=-(abs(r)+abs(a)) by JORDAN2B:29;
          a<=abs(a) by ABSVALUE:11;
        then A77:-abs(a)<=-a by REAL_1:50;
          abs(r)>=0 by ABSVALUE:5;
        then abs(a)<=abs(a)+abs(r) by REAL_2:173;
        then -abs(a)>= -(abs(a)+abs(r)) by REAL_1:50;
       hence r5<=-a by A76,A77,AXIOMS:22;
      end;
      then p3 in P1 by A14;
      then A78: |.p3.|<r by A14,A75;
      A79: |.p3.|=abs(-(abs(r)+abs(a))) by Th12
            .=abs ((abs(r)+abs(a))) by ABSVALUE:17;
      A80:0<=abs(r) by ABSVALUE:5;
A81:    0<=abs(a) by ABSVALUE:5;
      then 0<=abs(r)+abs(a) by A80,Th5;
      then A82:abs((abs(r)+abs(a)))=abs(r)+abs(a) by ABSVALUE:def 1;
      A83:r<=abs(r) by ABSVALUE:11;
        abs(r)<=abs(r)+abs(a) by A81,REAL_2:173;
     hence contradiction by A78,A79,A82,A83,AXIOMS:22;
    end;
    A84:now assume P2 is Bounded;
      then consider r being Real such that
      A85: for q being Point of TOP-REAL n st q in P2 holds |.q.|<r by Th40;
      set p3=|[(abs(r)+abs(a))]|;
        for r5 being Real st p3=|[r5]| holds r5>=a
      proof let r5 be Real;assume p3=|[r5]|;
        then A86:r5=(abs(r)+abs(a)) by JORDAN2B:29;
        A87: a<=abs(a) by ABSVALUE:11;
          abs(r)>=0 by ABSVALUE:5;
        then abs(a)<=abs(a)+abs(r) by REAL_2:173;
       hence r5>=a by A86,A87,AXIOMS:22;
      end;
      then p3 in P2 by A14;
      then A88: |.p3.|<r by A14,A85;
      A89: |.p3.|=abs((abs(r)+abs(a))) by Th12;
      A90:0<=abs(r) by ABSVALUE:5;
A91:    0<=abs(a) by ABSVALUE:5;
      then 0<=abs(r)+abs(a) by A90,Th5;
      then A92:abs((abs(r)+abs(a)))=abs(r)+abs(a) by ABSVALUE:def 1;
      A93:r<=abs(r) by ABSVALUE:11;
        abs(r)<=abs(r)+abs(a) by A91,REAL_2:173;
     hence contradiction by A88,A89,A92,A93,AXIOMS:22;
    end;
  A94:Down(P1,A`) is connected by A40,A57,Th15;
  A95:Down(P2,A`) is connected by A41,A73,Th15;
    now assume not BDD A is Bounded;
    then consider q being Point of TOP-REAL n such that
    A96: q in BDD A & not |.q.|<r by Th40;
      q in union{B3 where B3 is Subset of TOP-REAL n:
          B3 is_inside_component_of A} by A96,Def5;
    then consider y such that A97:q in y &
          y in {B3 where B3 is Subset of TOP-REAL n:
          B3 is_inside_component_of A} by TARSKI:def 4;
    consider B3 being Subset of TOP-REAL n
          such that A98: y=B3 & B3 is_inside_component_of A by A97;
      B3 is_a_component_of A` by A98,Def3;
    then consider B4 being Subset of (TOP-REAL n)|A` such that
      A99:B4 = B3 & B4 is_a_component_of (TOP-REAL n)|A`
                           by CONNSP_1:def 6;
    A100:q in the carrier of TOP-REAL n;
      for q2 being Point of TOP-REAL n st q2=q holds |.q2.| >= a by A96;
    then q in REAL n & not q in {q2 where q2 is Point of TOP-REAL n:
                       |.q2.| < a } by A100,EUCLID:25;
    then A101:q in P by XBOOLE_0:def 4;
    per cases by A17,A101,XBOOLE_0:def 2;
    suppose q in P1;
    then P1 /\ B4<>{}((TOP-REAL n)|A`) by A97,A98,A99,XBOOLE_0:def 3;
    then P1 meets B4 by XBOOLE_0:def 7;
    then A102:P1 c= B4 by A40,A94,A99,CONNSP_1:38;
      B3 is Bounded by A98,Def3;
   hence contradiction by A74,A99,A102,Th16;
   suppose q in P2;
    then P2 /\ B4<>{}((TOP-REAL n)|A`) by A97,A98,A99,XBOOLE_0:def 3;
    then P2 meets B4 by XBOOLE_0:def 7;
    then A103:P2 c= B4 by A41,A95,A99,CONNSP_1:38;
      B3 is Bounded by A98,Def3;
   hence contradiction by A84,A99,A103,Th16;
  end;
   hence BDD A is Bounded;
  end;
 hence BDD A is Bounded;
 suppose n<1;
   then n<0+1;
   then n<=0 by NAT_1:38;
   then A104:n=0 by NAT_1:18;
     BDD A c= the carrier of TOP-REAL n;
   then A105:BDD A c= [#](TOP-REAL n) by PRE_TOPC:12;
     for q2 being Point of TOP-REAL n holds |.q2.|<1
    proof let q2 be Point of TOP-REAL n;
         q2 in the carrier of TOP-REAL n;
       then q2 in REAL n by EUCLID:25;
       then q2=0.REAL n by A104,Th113,TARSKI:def 1;
       then |.q2.|=0 by TOPRNS_1:24;
      hence |.q2.|<1;
    end;
   then for q2 being Point of TOP-REAL n st q2 in [#]
(TOP-REAL n) holds |.q2.|<1;
   then [#](TOP-REAL n) is Bounded by Th40;
 hence BDD A is Bounded by A105,Th16;
end;

theorem Th115:for G being non empty TopSpace,A,B,C,D being Subset of G st
A is_a_component_of G & B is_a_component_of G & C is_a_component_of G
& A \/ B=the carrier of G & C misses A holds C=B
proof let G be non empty TopSpace,A,B,C,D be Subset of G;
 assume A1:A is_a_component_of G &
  B is_a_component_of G & C is_a_component_of G &
  A \/ B=the carrier of G & C misses A;
    now assume C misses B;
  then A2: C /\ B = {} by XBOOLE_0:def 7;
      C /\ (the carrier of G)=C by XBOOLE_1:28;
    then A3: C /\ A \/ C /\ B=C by A1,XBOOLE_1:23;
      C <> {}G by A1,CONNSP_1:34;
   hence contradiction by A1,A2,A3,XBOOLE_0:def 7;
  end;
 hence C=B by A1,CONNSP_1:37;
end;

theorem Th116:for A being Subset of TOP-REAL 2 st
A is Bounded & A is Jordan holds BDD A is_inside_component_of A
proof let A be Subset of TOP-REAL 2;
 assume A1:A is Bounded & A is Jordan;
  then consider A1,A2 being Subset of TOP-REAL 2 such that
  A2:A` = A1 \/ A2 & A1 misses A2 & (Cl A1) \ A1 = (Cl A2) \ A2 &
     for C1,C2 being Subset of (TOP-REAL 2)|A`
     st C1 = A1 & C2 = A2 holds
      C1 is_a_component_of (TOP-REAL 2)|A` &
      C2 is_a_component_of (TOP-REAL 2)|A` by JORDAN1:def 2;
   A3:Down(A1,A`)=A1 /\ A` by CONNSP_3:def 5
   .=A1 by A2,XBOOLE_1:21;
   A4:Down(A2,A`)=A2 /\ A` by CONNSP_3:def 5
   .=A2 by A2,XBOOLE_1:21;
   then A5:Down(A1,A`) is_a_component_of (TOP-REAL 2)|A` &
   Down(A2,A`) is_a_component_of (TOP-REAL 2)|A` by A2,A3;
   then A6:A1 is_a_component_of A` by A3,CONNSP_1:def 6;
   A7:A2 is_a_component_of A` by A4,A5,CONNSP_1:def 6;
   reconsider D=A` as non empty Subset of TOP-REAL 2
     by A1,JORDAN1:def 2;
   A8:(TOP-REAL 2)|D is non empty;
   then A9:Down(A2,A`)<>{}((TOP-REAL 2)|A`) by A5,CONNSP_1:34;
     consider B being Subset of TOP-REAL 2 such that
     A10:B is_outside_component_of A & B=UBD A by A1,Th76;
       UBD A is_a_component_of A` by A10,Def4;
     then consider B1 being Subset of (TOP-REAL 2)|A` such that
     A11:B1 = UBD A & B1 is_a_component_of (TOP-REAL 2)|A` by CONNSP_1:def 6;
   per cases by A3,A5,A11,CONNSP_1:37;
     suppose A12:B1=A1;
         now assume not A2 is Bounded;
         then A2 is_outside_component_of A by A7,Def4;
         then A2 c= UBD A by Th27;
         then A2 /\ UBD A <> {} by A4,A9,XBOOLE_1:28;
         hence contradiction by A2,A11,A12,XBOOLE_0:def 7;
       end;
       then A13:A2 is_inside_component_of A by A7,Def3;
       then A14:A2 c= BDD A by Th26;
         now assume not BDD A c= A2;
         then consider x such that A15:x in BDD A & not x in A2 by TARSKI:def 3
;
            x in union{B3 where B3 is Subset of TOP-REAL 2:
          B3 is_inside_component_of A} by A15,Def5;
          then consider y such that A16:x in y &
          y in {B3 where B3 is Subset of TOP-REAL 2:
          B3 is_inside_component_of A} by TARSKI:def 4;
          consider B3 being Subset of TOP-REAL 2
          such that A17: y=B3 & B3 is_inside_component_of A by A16;
          A18:B3 is_a_component_of A` by A17,Def3;
          then consider B4 being Subset of (TOP-REAL 2)|A` such that
          A19:B4 = B3 & B4 is_a_component_of (TOP-REAL 2)|A` by CONNSP_1:def 6;
          A20:B3<>{}((TOP-REAL 2)|A`) by A8,A19,CONNSP_1:34;
            B4=Down(A2,A`) or B4 misses Down(A2,A`) by A5,A19,CONNSP_1:37;
          then A21:B4=Down(A2,A`) or B4 /\ Down(A2,A`)={}((TOP-REAL 2)|A`)
                                     by XBOOLE_0:def 7;
           now assume B4=Down(A1,A`);
            then UBD A is Bounded by A3,A11,A12,A17,A19,Def3;
           hence contradiction by A10,Def4;
          end;
then A22:       B3 misses A1 & B3 misses A2
            by A3,A4,A5,A15,A16,A17,A19,CONNSP_1:37;
            B3 c= A` by A18,SPRECT_1:7;
          then B3=B3 /\ (A1 \/ A2) by A2,XBOOLE_1:28 .=(B3 /\ A1) \/ (B3 /\ A2
)
                  by XBOOLE_1:23
          .={} by A4,A16,A17,A19,A21,A22,XBOOLE_0:def 7;
        hence contradiction by A20;
       end;
      hence BDD A is_inside_component_of A by A13,A14,XBOOLE_0:def 10;
     suppose A23:B1 misses Down(A1,A`);
       A24:BDD A misses UBD A by Th28;
       A25:BDD A \/ UBD A=A` by Th31;
       set E1=Down(A1,A`), E2=Down(A2,A`);
         E1 \/ E2=[#]((TOP-REAL 2)|A`) by A2,A3,A4,PRE_TOPC:def 10;
       then E1 \/ E2=the carrier of (TOP-REAL 2)|A` by PRE_TOPC:12;
       then A26:UBD A=A2 by A4,A5,A8,A11,A23,Th115;
       A27:BDD A c= A1
       proof let z;assume A28:z in BDD A;
         then A29:z in A` by A25,XBOOLE_0:def 2;
           not z in UBD A by A24,A28,XBOOLE_0:3;
        hence z in A1 by A2,A26,A29,XBOOLE_0:def 2;
       end;
         A1 c= BDD A
       proof let z;assume A30:z in A1;
       then A31:z in A` by A2,XBOOLE_0:def 2;
           not z in UBD A by A2,A26,A30,XBOOLE_0:3;
        hence z in BDD A by A25,A31,XBOOLE_0:def 2;
       end;
       then A32:BDD A = A1 by A27,XBOOLE_0:def 10;
         BDD A is Bounded by A1,Th114;
      hence thesis by A6,A32,Def3;
end;

theorem for a being Real,p being Point of TOP-REAL 2
st a>0 & p in (L~SpStSeq D) holds ex q being Point of TOP-REAL 2 st
q in BDD (L~SpStSeq D) & |.p-q.|<a
proof let a be Real,p be Point of TOP-REAL 2;
 assume A1:a>0 & p in (L~SpStSeq D);
  set A=L~SpStSeq D;
    A`<>{} by SPRECT_1:def 3;
 then consider A1,A2 being Subset of TOP-REAL 2 such that
  A2: A` = A1 \/ A2 & A1 misses A2 & (Cl A1) \ A1 = (Cl A2) \ A2 &
     A=(Cl A1) \ A1 & for C1,C2 being Subset of (TOP-REAL 2)|A`
     st C1 = A1 & C2 = A2 holds
      C1 is_a_component_of (TOP-REAL 2)|A` &
      C2 is_a_component_of (TOP-REAL 2)|A` by Th106;
   A3:Down(A1,A`)=A1 /\ A` by CONNSP_3:def 5
   .=A1 by A2,XBOOLE_1:21;
   A4:Down(A2,A`)=A2 /\ A` by CONNSP_3:def 5
   .=A2 by A2,XBOOLE_1:21;
   then A5:Down(A1,A`) is_a_component_of (TOP-REAL 2)|A` &
   Down(A2,A`) is_a_component_of (TOP-REAL 2)|A` by A2,A3;
     A is Bounded by Th73;
   then BDD A is_inside_component_of A by Th116;
   then BDD (L~SpStSeq D) is_a_component_of A` by Def3;
   then consider B1 being Subset of (TOP-REAL 2)|A` such that
   A6: B1 = BDD (L~SpStSeq D) & B1 is_a_component_of (TOP-REAL 2)|A`
                                  by CONNSP_1:def 6;
     B1 c= the carrier of (TOP-REAL 2)|A`;
   then A7:BDD (L~SpStSeq D) c= A1 \/ A2 by A2,A6,JORDAN1:1;
   consider q1 being Element of BDD (L~SpStSeq D);
   A8:BDD (L~SpStSeq D) <>{} by Th104;
   then A9:q1 in BDD (L~SpStSeq D);
   per cases by A7,A9,XBOOLE_0:def 2;
   suppose q1 in A1;
     then B1 /\ Down(A1,A`)<>{}((TOP-REAL 2)|A`) by A3,A6,A8,XBOOLE_0:def 3;
     then B1 meets Down(A1,A`) by XBOOLE_0:def 7;
     then B1=Down(A1,A`) by A5,A6,CONNSP_1:37;
     then A10: p in Cl(BDD (L~SpStSeq D)) by A1,A2,A3,A6,XBOOLE_0:def 4;
     reconsider ep=p as Point of Euclid 2 by TOPREAL3:13;
       Ball(ep,a) c= the carrier of Euclid 2;
     then Ball(ep,a) c= the carrier of TOP-REAL 2 by TOPREAL3:13;
     then reconsider G2=Ball(ep,a) as Subset of TOP-REAL 2;
       (the distance of (Euclid 2)) is Reflexive by METRIC_1:def 7;
     then (the distance of Euclid 2).(ep,ep)=0 by METRIC_1:def 3;
     then dist(ep,ep)=0 by METRIC_1:def 1;
     then A11:p in Ball(ep,a) by A1,METRIC_1:12;
       G2 is open by GOBOARD6:6;
     then (BDD (L~SpStSeq D)) meets G2 by A10,A11,PRE_TOPC:def 13;
     then consider t2 being set such that
     A12:t2 in (BDD (L~SpStSeq D)) & t2 in G2 by XBOOLE_0:3;
    reconsider qt2=t2 as Point of TOP-REAL 2 by A12;
        t2 in BDD (L~SpStSeq D) & |.p-qt2.|<a by A12,Th111;
    hence thesis;
   suppose q1 in A2;
     then B1 /\ Down(A2,A`)<>{}((TOP-REAL 2)|A`) by A4,A6,A8,XBOOLE_0:def 3;
     then B1 meets Down(A2,A`) by XBOOLE_0:def 7;
     then B1=Down(A2,A`) by A5,A6,CONNSP_1:37;
     then A13: p in Cl(BDD (L~SpStSeq D)) by A1,A2,A4,A6,XBOOLE_0:def 4;
     reconsider ep=p as Point of Euclid 2 by TOPREAL3:13;
       Ball(ep,a) c= the carrier of Euclid 2;
     then Ball(ep,a) c= the carrier of TOP-REAL 2 by TOPREAL3:13;
     then reconsider G2=Ball(ep,a) as Subset of TOP-REAL 2;
       (the distance of (Euclid 2)) is Reflexive by METRIC_1:def 7;
     then (the distance of Euclid 2).(ep,ep)=0 by METRIC_1:def 3;
     then dist(ep,ep)=0 by METRIC_1:def 1;
     then A14:p in Ball(ep,a) by A1,METRIC_1:12;
       G2 is open by GOBOARD6:6;
     then (BDD (L~SpStSeq D)) meets G2 by A13,A14,PRE_TOPC:def 13;
     then consider t2 being set such that
     A15:t2 in (BDD (L~SpStSeq D)) & t2 in G2 by XBOOLE_0:3;
    reconsider qt2=t2 as Point of TOP-REAL 2 by A15;
        t2 in BDD (L~SpStSeq D) & |.p-qt2.|<a by A15,Th111;
    hence thesis;
end;

begin :: Points in LeftComp

reserve f for
   clockwise_oriented (non constant standard special_circular_sequence);

theorem for p being Point of TOP-REAL 2 st f/.1 = N-min L~f &
 p`1 < W-bound (L~f) holds p in LeftComp f
proof let p be Point of TOP-REAL 2;
 assume A1:f/.1 = N-min L~f & p`1< W-bound(L~f);
  set g=SpStSeq L~f;
    W-bound L~ g=W-bound L~f by SPRECT_1:66;
  then A2:p in LeftComp g by A1,SPRECT_3:57;
    LeftComp g c= LeftComp f by A1,SPRECT_3:58;
 hence p in LeftComp f by A2;
end;

theorem for p being Point of TOP-REAL 2 st f/.1 = N-min L~f &
 p`1 > E-bound (L~f) holds p in LeftComp f
proof let p be Point of TOP-REAL 2;
 assume A1:f/.1 = N-min L~f & p`1> E-bound(L~f);
  set g=SpStSeq L~f;
    E-bound L~ g=E-bound L~f by SPRECT_1:69;
  then A2:p in LeftComp g by A1,SPRECT_3:57;
    LeftComp g c= LeftComp f by A1,SPRECT_3:58;
 hence p in LeftComp f by A2;
end;

theorem for p being Point of TOP-REAL 2 st f/.1 = N-min L~f &
 p`2 < S-bound (L~f) holds p in LeftComp f
proof let p be Point of TOP-REAL 2;
 assume A1:f/.1 = N-min L~f & p`2< S-bound(L~f);
  set g=SpStSeq L~f;
    S-bound L~ g=S-bound L~f by SPRECT_1:67;
  then A2:p in LeftComp g by A1,SPRECT_3:57;
    LeftComp g c= LeftComp f by A1,SPRECT_3:58;
 hence p in LeftComp f by A2;
end;

theorem for p being Point of TOP-REAL 2 st f/.1 = N-min L~f &
 p`2 > N-bound (L~f) holds p in LeftComp f
proof let p be Point of TOP-REAL 2;
 assume A1:f/.1 = N-min L~f & p`2>N-bound(L~f);
  set g=SpStSeq L~f;
    N-bound L~ g=N-bound L~f by SPRECT_1:68;
  then A2:p in LeftComp g by A1,SPRECT_3:57;
    LeftComp g c= LeftComp f by A1,SPRECT_3:58;
 hence p in LeftComp f by A2;
end;

Back to top