scheme
SCH2{
P1[
object ],
P2[
object ],
F1()
-> non
empty set ,
F2()
-> non
empty set ,
F3()
-> non
empty set ,
F4(
object )
-> object ,
F5(
object )
-> object ,
F6(
object )
-> object } :
ex
f being
Function of
F3(),
F2() st
for
a being
Element of
F1() st
a in F3() holds
( (
P1[
a] implies
f . a = F4(
a) ) & (
P1[
a] &
P2[
a] implies
f . a = F5(
a) ) & (
P1[
a] &
P2[
a] implies
f . a = F6(
a) ) )
provided
A1:
F3()
c= F1()
and A2:
for
a being
Element of
F1() st
a in F3() holds
( (
P1[
a] implies
F4(
a)
in F2() ) & (
P1[
a] &
P2[
a] implies
F5(
a)
in F2() ) & (
P1[
a] &
P2[
a] implies
F6(
a)
in F2() ) )
definition
existence
ex b1 being non empty strict TopSpace st
( the carrier of b1 = y>=0-plane & ex B being Neighborhood_System of b1 st
( ( for x being Real holds B . |[x,0]| = { ((Ball (|[x,r]|,r)) \/ {|[x,0]|}) where r is Real : r > 0 } ) & ( for x, y being Real st y > 0 holds
B . |[x,y]| = { ((Ball (|[x,y]|,r)) /\ y>=0-plane) where r is Real : r > 0 } ) ) )
uniqueness
for b1, b2 being non empty strict TopSpace st the carrier of b1 = y>=0-plane & ex B being Neighborhood_System of b1 st
( ( for x being Real holds B . |[x,0]| = { ((Ball (|[x,r]|,r)) \/ {|[x,0]|}) where r is Real : r > 0 } ) & ( for x, y being Real st y > 0 holds
B . |[x,y]| = { ((Ball (|[x,y]|,r)) /\ y>=0-plane) where r is Real : r > 0 } ) ) & the carrier of b2 = y>=0-plane & ex B being Neighborhood_System of b2 st
( ( for x being Real holds B . |[x,0]| = { ((Ball (|[x,r]|,r)) \/ {|[x,0]|}) where r is Real : r > 0 } ) & ( for x, y being Real st y > 0 holds
B . |[x,y]| = { ((Ball (|[x,y]|,r)) /\ y>=0-plane) where r is Real : r > 0 } ) ) holds
b1 = b2
end;
Lm1:
the carrier of Niemytzki-plane = y>=0-plane
by Def3;
definition
let x be
Real;
let r be
positive Real;
func + (
x,
r)
-> Function of
Niemytzki-plane,
I[01] means :
Def5:
(
it . |[x,0]| = 0 & ( for
a being
Real for
b being non
negative Real holds
( ( (
a <> x or
b <> 0 ) & not
|[a,b]| in Ball (
|[x,r]|,
r) implies
it . |[a,b]| = 1 ) & (
|[a,b]| in Ball (
|[x,r]|,
r) implies
it . |[a,b]| = (|.(|[x,0]| - |[a,b]|).| ^2) / ((2 * r) * b) ) ) ) );
existence
ex b1 being Function of Niemytzki-plane,I[01] st
( b1 . |[x,0]| = 0 & ( for a being Real
for b being non negative Real holds
( ( ( a <> x or b <> 0 ) & not |[a,b]| in Ball (|[x,r]|,r) implies b1 . |[a,b]| = 1 ) & ( |[a,b]| in Ball (|[x,r]|,r) implies b1 . |[a,b]| = (|.(|[x,0]| - |[a,b]|).| ^2) / ((2 * r) * b) ) ) ) )
uniqueness
for b1, b2 being Function of Niemytzki-plane,I[01] st b1 . |[x,0]| = 0 & ( for a being Real
for b being non negative Real holds
( ( ( a <> x or b <> 0 ) & not |[a,b]| in Ball (|[x,r]|,r) implies b1 . |[a,b]| = 1 ) & ( |[a,b]| in Ball (|[x,r]|,r) implies b1 . |[a,b]| = (|.(|[x,0]| - |[a,b]|).| ^2) / ((2 * r) * b) ) ) ) & b2 . |[x,0]| = 0 & ( for a being Real
for b being non negative Real holds
( ( ( a <> x or b <> 0 ) & not |[a,b]| in Ball (|[x,r]|,r) implies b2 . |[a,b]| = 1 ) & ( |[a,b]| in Ball (|[x,r]|,r) implies b2 . |[a,b]| = (|.(|[x,0]| - |[a,b]|).| ^2) / ((2 * r) * b) ) ) ) holds
b1 = b2
end;
::
deftheorem Def5 defines
+ TOPGEN_5:def 5 :
for x being Real
for r being positive Real
for b3 being Function of Niemytzki-plane,I[01] holds
( b3 = + (x,r) iff ( b3 . |[x,0]| = 0 & ( for a being Real
for b being non negative Real holds
( ( ( a <> x or b <> 0 ) & not |[a,b]| in Ball (|[x,r]|,r) implies b3 . |[a,b]| = 1 ) & ( |[a,b]| in Ball (|[x,r]|,r) implies b3 . |[a,b]| = (|.(|[x,0]| - |[a,b]|).| ^2) / ((2 * r) * b) ) ) ) ) );
theorem Th76:
for
U being
Subset of
Niemytzki-plane for
x being
Real for
r being
positive Real st
U = (Ball (|[x,r]|,r)) \/ {|[x,0]|} holds
ex
f being
continuous Function of
Niemytzki-plane,
I[01] st
(
f . |[x,0]| = 0 & ( for
a,
b being
Real holds
( (
|[a,b]| in U ` implies
f . |[a,b]| = 1 ) & (
|[a,b]| in U \ {|[x,0]|} implies
f . |[a,b]| = (|.(|[x,0]| - |[a,b]|).| ^2) / ((2 * r) * b) ) ) ) )
definition
let x,
y be
Real;
let r be
positive Real;
func + (
x,
y,
r)
-> Function of
Niemytzki-plane,
I[01] means :
Def6:
for
a being
Real for
b being non
negative Real holds
( ( not
|[a,b]| in Ball (
|[x,y]|,
r) implies
it . |[a,b]| = 1 ) & (
|[a,b]| in Ball (
|[x,y]|,
r) implies
it . |[a,b]| = |.(|[x,y]| - |[a,b]|).| / r ) );
existence
ex b1 being Function of Niemytzki-plane,I[01] st
for a being Real
for b being non negative Real holds
( ( not |[a,b]| in Ball (|[x,y]|,r) implies b1 . |[a,b]| = 1 ) & ( |[a,b]| in Ball (|[x,y]|,r) implies b1 . |[a,b]| = |.(|[x,y]| - |[a,b]|).| / r ) )
uniqueness
for b1, b2 being Function of Niemytzki-plane,I[01] st ( for a being Real
for b being non negative Real holds
( ( not |[a,b]| in Ball (|[x,y]|,r) implies b1 . |[a,b]| = 1 ) & ( |[a,b]| in Ball (|[x,y]|,r) implies b1 . |[a,b]| = |.(|[x,y]| - |[a,b]|).| / r ) ) ) & ( for a being Real
for b being non negative Real holds
( ( not |[a,b]| in Ball (|[x,y]|,r) implies b2 . |[a,b]| = 1 ) & ( |[a,b]| in Ball (|[x,y]|,r) implies b2 . |[a,b]| = |.(|[x,y]| - |[a,b]|).| / r ) ) ) holds
b1 = b2
end;
::
deftheorem Def6 defines
+ TOPGEN_5:def 6 :
for x, y being Real
for r being positive Real
for b4 being Function of Niemytzki-plane,I[01] holds
( b4 = + (x,y,r) iff for a being Real
for b being non negative Real holds
( ( not |[a,b]| in Ball (|[x,y]|,r) implies b4 . |[a,b]| = 1 ) & ( |[a,b]| in Ball (|[x,y]|,r) implies b4 . |[a,b]| = |.(|[x,y]| - |[a,b]|).| / r ) ) );
theorem Th80:
for
p being
Point of
(TOP-REAL 2) st
p `2 = 0 holds
for
x being
Real for
a being non
negative Real for
y,
r being
positive Real st
(+ (x,y,r)) . p > a holds
(
|.(|[x,y]| - p).| > r * a & ex
r1 being
positive Real st
(
r1 = (|.(|[x,y]| - p).| - (r * a)) / 2 &
(Ball (|[(p `1),r1]|,r1)) \/ {p} c= (+ (x,y,r)) " ].a,1.] ) )
theorem Th81:
for
U being
Subset of
Niemytzki-plane for
x,
y being
Real for
r being
positive Real st
y > 0 &
U = (Ball (|[x,y]|,r)) /\ y>=0-plane holds
ex
f being
continuous Function of
Niemytzki-plane,
I[01] st
(
f . |[x,y]| = 0 & ( for
a,
b being
Real holds
( (
|[a,b]| in U ` implies
f . |[a,b]| = 1 ) & (
|[a,b]| in U implies
f . |[a,b]| = |.(|[x,y]| - |[a,b]|).| / r ) ) ) )