begin
reconsider D = DYADIC as Subset of I[01] by BORSUK_1:83, URYSOHN2:32;
Lm1:
Cl D = [#] I[01]
theorem
theorem Th2:
theorem Th3:
theorem Th4:
theorem Th5:
theorem Th6:
theorem Th7:
theorem Th8:
theorem Th9:
theorem Th10:
:: deftheorem Def1 defines isometric MAZURULM:def 1 :
for E, F being non empty NORMSTR
for f being Function of E,F holds
( f is isometric iff for a, b being Point of E holds ||.((f . a) - (f . b)).|| = ||.(a - b).|| );
:: deftheorem Def2 defines Affine MAZURULM:def 2 :
for E, F being non empty RLSStruct
for f being Function of E,F holds
( f is Affine iff for a, b being Point of E
for t being real number st 0 <= t & t <= 1 holds
f . (((1 - t) * a) + (t * b)) = ((1 - t) * (f . a)) + (t * (f . b)) );
:: deftheorem Def3 defines midpoints-preserving MAZURULM:def 3 :
for E, F being non empty RLSStruct
for f being Function of E,F holds
( f is midpoints-preserving iff for a, b being Point of E holds f . ((1 / 2) * (a + b)) = (1 / 2) * ((f . a) + (f . b)) );
theorem Th11:
Lm2:
now
let E,
F be
RealNormSpace;
for f being Function of E,F st f is bijective holds
for a being Point of F holds f . ((f /") . a) = alet f be
Function of
E,
F;
( f is bijective implies for a being Point of F holds f . ((f /") . a) = a )assume A1:
f is
bijective
;
for a being Point of F holds f . ((f /") . a) = aset g =
f /" ;
let a be
Point of
F;
f . ((f /") . a) = aA2:
rng f = [#] F
by A1, FUNCT_2:def 3;
then A3:
(f /") /" = f
by A1, TOPS_2:64;
A4:
f /" = f "
by A1, A2, TOPS_2:def 4;
A5:
f /" is
bijective
by A1, PENCIL_2:12;
then
rng (f /") = [#] E
by FUNCT_2:def 3;
then
f = (f /") "
by A3, A5, TOPS_2:def 4;
hence
f . ((f /") . a) = a
by A1, A4, FUNCT_2:32;
verum
end;
theorem Th12:
theorem Th13:
theorem Th14:
theorem Th15:
theorem Th16:
:: deftheorem Def4 defines -reflection MAZURULM:def 4 :
for E being non empty RLSStruct
for a being Point of E
for b3 being UnOp of E holds
( b3 = a -reflection iff for b being Point of E holds b3 . b = (2 * a) - b );
theorem Th17:
theorem Th18:
theorem Th19:
theorem Th20:
theorem Th21:
theorem Th22:
theorem Th23:
deffunc H1( RealNormSpace, Point of $1, Point of $1) -> set = { g where g is UnOp of $1 : ( g is bijective & g is isometric & g . $2 = $2 & g . $3 = $3 ) } ;
deffunc H2( RealNormSpace, Point of $1, Point of $1) -> set = { ||.((g . ((1 / 2) * ($2 + $3))) - ((1 / 2) * ($2 + $3))).|| where g is UnOp of $1 : g in H1($1,$2,$3) } ;
Lm4:
now
let E be
RealNormSpace;
for a, b being Point of E
for l being real-membered set st l = H2(E,a,b) holds
2 * ||.(a - ((1 / 2) * (a + b))).|| is UpperBound of llet a,
b be
Point of
E;
for l being real-membered set st l = H2(E,a,b) holds
2 * ||.(a - ((1 / 2) * (a + b))).|| is UpperBound of llet l be
real-membered set ;
( l = H2(E,a,b) implies 2 * ||.(a - ((1 / 2) * (a + b))).|| is UpperBound of l )assume A1:
l = H2(
E,
a,
b)
;
2 * ||.(a - ((1 / 2) * (a + b))).|| is UpperBound of lset z =
(1 / 2) * (a + b);
thus
2
* ||.(a - ((1 / 2) * (a + b))).|| is
UpperBound of
l
verum
end;
Lm5:
now
let E be
RealNormSpace;
for a, b being Point of E
for h being UnOp of E st h in H1(E,a,b) holds
(((((1 / 2) * (a + b)) -reflection) * (h /")) * (((1 / 2) * (a + b)) -reflection)) * h in H1(E,a,b)let a,
b be
Point of
E;
for h being UnOp of E st h in H1(E,a,b) holds
(((((1 / 2) * (a + b)) -reflection) * (h /")) * (((1 / 2) * (a + b)) -reflection)) * h in H1(E,a,b)let h be
UnOp of
E;
( h in H1(E,a,b) implies (((((1 / 2) * (a + b)) -reflection) * (h /")) * (((1 / 2) * (a + b)) -reflection)) * h in H1(E,a,b) )assume A1:
h in H1(
E,
a,
b)
;
(((((1 / 2) * (a + b)) -reflection) * (h /")) * (((1 / 2) * (a + b)) -reflection)) * h in H1(E,a,b)set z =
(1 / 2) * (a + b);
set R =
((1 / 2) * (a + b)) -reflection ;
set gs =
(((((1 / 2) * (a + b)) -reflection) * (h /")) * (((1 / 2) * (a + b)) -reflection)) * h;
consider g being
UnOp of
E such that A2:
g = h
and A3:
g is
bijective
and A4:
g is
isometric
and A5:
g . a = a
and A6:
g . b = b
by A1;
rng g = [#] E
by A3, FUNCT_2:def 3;
then A7:
g /" = g "
by A3, TOPS_2:def 4;
then A8:
g /" is
bijective
by A3, GROUP_6:73;
A9: 2
* ((1 / 2) * (a + b)) =
(2 * (1 / 2)) * (a + b)
by RLVECT_1:def 10
.=
a + b
by RLVECT_1:def 11
;
then A10:
(2 * ((1 / 2) * (a + b))) - a = b
by Th4;
A11:
(2 * ((1 / 2) * (a + b))) - b = a
by A9, Th4;
A12:
dom g = [#] E
by FUNCT_2:def 1;
A13:
(g /") . b = b
by A7, A6, A3, A12, FUNCT_1:54;
A14:
(g /") . a = a
by A7, A5, A3, A12, FUNCT_1:54;
A15:
((((((1 / 2) * (a + b)) -reflection) * (h /")) * (((1 / 2) * (a + b)) -reflection)) * h) . a =
(((((1 / 2) * (a + b)) -reflection) * (g /")) * (((1 / 2) * (a + b)) -reflection)) . a
by A5, A2, FUNCT_2:21
.=
((((1 / 2) * (a + b)) -reflection) * (g /")) . ((((1 / 2) * (a + b)) -reflection) . a)
by FUNCT_2:21
.=
((((1 / 2) * (a + b)) -reflection) * (g /")) . b
by A10, Def4
.=
(((1 / 2) * (a + b)) -reflection) . ((g /") . b)
by FUNCT_2:21
.=
(2 * ((1 / 2) * (a + b))) - b
by A13, Def4
.=
a
by A9, Th4
;
((((((1 / 2) * (a + b)) -reflection) * (h /")) * (((1 / 2) * (a + b)) -reflection)) * h) . b =
(((((1 / 2) * (a + b)) -reflection) * (g /")) * (((1 / 2) * (a + b)) -reflection)) . b
by A6, A2, FUNCT_2:21
.=
((((1 / 2) * (a + b)) -reflection) * (g /")) . ((((1 / 2) * (a + b)) -reflection) . b)
by FUNCT_2:21
.=
((((1 / 2) * (a + b)) -reflection) * (g /")) . a
by A11, Def4
.=
(((1 / 2) * (a + b)) -reflection) . ((g /") . a)
by FUNCT_2:21
.=
(2 * ((1 / 2) * (a + b))) - a
by A14, Def4
.=
b
by A9, Th4
;
hence
(((((1 / 2) * (a + b)) -reflection) * (h /")) * (((1 / 2) * (a + b)) -reflection)) * h in H1(
E,
a,
b)
by A2, A3, A4, A8, A15;
verum
end;
Lm6:
now
let E be
RealNormSpace;
for a, b being Point of E
for l being non empty bounded_above Subset of REAL st l = H2(E,a,b) holds
sup l is UpperBound of 2 ** llet a,
b be
Point of
E;
for l being non empty bounded_above Subset of REAL st l = H2(E,a,b) holds
sup l is UpperBound of 2 ** llet l be non
empty bounded_above Subset of
REAL;
( l = H2(E,a,b) implies sup l is UpperBound of 2 ** l )assume A1:
l = H2(
E,
a,
b)
;
sup l is UpperBound of 2 ** lthus
sup l is
UpperBound of 2
** l
verum
proof
set z =
(1 / 2) * (a + b);
set R =
((1 / 2) * (a + b)) -reflection ;
let x be
ext-real number ;
XXREAL_2:def 1 ( not x in 2 ** l or x <= sup l )
assume
x in 2
** l
;
x <= sup l
then consider w being
Element of
ExtREAL such that A2:
x = 2
* w
and A3:
w in l
by MEMBER_1:188;
consider h being
UnOp of
E such that A4:
w = ||.((h . ((1 / 2) * (a + b))) - ((1 / 2) * (a + b))).||
and A5:
h in H1(
E,
a,
b)
by A3, A1;
consider g being
UnOp of
E such that A6:
g = h
and A7:
g is
bijective
and A8:
g is
isometric
and
(
g . a = a &
g . b = b )
by A5;
A9:
(((1 / 2) * (a + b)) -reflection) * ((g /") * ((((1 / 2) * (a + b)) -reflection) * g)) =
((((1 / 2) * (a + b)) -reflection) * (g /")) * ((((1 / 2) * (a + b)) -reflection) * g)
by RELAT_1:55
.=
(((((1 / 2) * (a + b)) -reflection) * (g /")) * (((1 / 2) * (a + b)) -reflection)) * g
by RELAT_1:55
;
A10:
(((1 / 2) * (a + b)) -reflection) . ((g /") . ((((1 / 2) * (a + b)) -reflection) . (g . ((1 / 2) * (a + b))))) =
(((1 / 2) * (a + b)) -reflection) . ((g /") . (((((1 / 2) * (a + b)) -reflection) * g) . ((1 / 2) * (a + b))))
by FUNCT_2:21
.=
(((1 / 2) * (a + b)) -reflection) . (((g /") * ((((1 / 2) * (a + b)) -reflection) * g)) . ((1 / 2) * (a + b)))
by FUNCT_2:21
.=
((((1 / 2) * (a + b)) -reflection) * ((g /") * ((((1 / 2) * (a + b)) -reflection) * g))) . ((1 / 2) * (a + b))
by FUNCT_2:21
;
rng g = [#] E
by A7, FUNCT_2:def 3;
then A11:
g /" = g "
by A7, TOPS_2:def 4;
(((((1 / 2) * (a + b)) -reflection) * (g /")) * (((1 / 2) * (a + b)) -reflection)) * g in H1(
E,
a,
b)
by A5, A6, Lm5;
then A12:
||.((((((((1 / 2) * (a + b)) -reflection) * (g /")) * (((1 / 2) * (a + b)) -reflection)) * g) . ((1 / 2) * (a + b))) - ((1 / 2) * (a + b))).|| in l
by A1;
reconsider d = 2 as
R_eal by XXREAL_0:def 1;
||.((g . ((1 / 2) * (a + b))) - ((1 / 2) * (a + b))).|| in ExtREAL
by XXREAL_0:def 1;
then A13:
d * ||.((g . ((1 / 2) * (a + b))) - ((1 / 2) * (a + b))).|| = 2
* ||.((g . ((1 / 2) * (a + b))) - ((1 / 2) * (a + b))).||
by EXTREAL1:13;
A14:
(g /") . (g . ((1 / 2) * (a + b))) = (1 / 2) * (a + b)
by A7, A11, FUNCT_2:32;
2
* ||.((g . ((1 / 2) * (a + b))) - ((1 / 2) * (a + b))).|| =
||.(((((1 / 2) * (a + b)) -reflection) . (g . ((1 / 2) * (a + b)))) - (g . ((1 / 2) * (a + b)))).||
by Th22
.=
||.(((g /") . ((((1 / 2) * (a + b)) -reflection) . (g . ((1 / 2) * (a + b))))) - ((g /") . (g . ((1 / 2) * (a + b))))).||
by A7, A8, Def1
.=
||.(((((1 / 2) * (a + b)) -reflection) . ((g /") . ((((1 / 2) * (a + b)) -reflection) . (g . ((1 / 2) * (a + b)))))) - ((1 / 2) * (a + b))).||
by A14, Th20
;
hence
x <= sup l
by A2, A4, A9, A10, A12, A6, A13, XXREAL_2:4;
verum
end;
end;
Lm7:
now
let E be
RealNormSpace;
for a, b being Point of E
for l being real-membered set st l = H2(E,a,b) holds
for g being UnOp of E st g in H1(E,a,b) holds
g . ((1 / 2) * (a + b)) = (1 / 2) * (a + b)let a,
b be
Point of
E;
for l being real-membered set st l = H2(E,a,b) holds
for g being UnOp of E st g in H1(E,a,b) holds
g . ((1 / 2) * (a + b)) = (1 / 2) * (a + b)let l be
real-membered set ;
( l = H2(E,a,b) implies for g being UnOp of E st g in H1(E,a,b) holds
g . ((1 / 2) * (a + b)) = (1 / 2) * (a + b) )assume A1:
l = H2(
E,
a,
b)
;
for g being UnOp of E st g in H1(E,a,b) holds
g . ((1 / 2) * (a + b)) = (1 / 2) * (a + b)let g be
UnOp of
E;
( g in H1(E,a,b) implies g . ((1 / 2) * (a + b)) = (1 / 2) * (a + b) )assume A2:
g in H1(
E,
a,
b)
;
g . ((1 / 2) * (a + b)) = (1 / 2) * (a + b)set z =
(1 / 2) * (a + b);
set R =
((1 / 2) * (a + b)) -reflection ;
A3:
l c= REAL
(
(id E) . a = a &
(id E) . b = b )
by FUNCT_1:35;
then
id E in H1(
E,
a,
b)
;
then A4:
||.(((id E) . ((1 / 2) * (a + b))) - ((1 / 2) * (a + b))).|| in l
by A1;
2
* ||.(a - ((1 / 2) * (a + b))).|| is
UpperBound of
l
by A1, Lm4;
then reconsider A =
l as non
empty bounded_above Subset of
REAL by A3, A4, XXREAL_2:def 10;
set lambda =
sup A;
reconsider d = 2 as
R_eal by XXREAL_0:def 1;
A5:
d * (sup A) = sup (2 ** A)
by URYSOHN2:20;
A6:
||.((g . ((1 / 2) * (a + b))) - ((1 / 2) * (a + b))).|| in A
by A1, A2;
A7:
d * (sup A) = 2
* (sup A)
by XXREAL_3:def 5;
sup A is
UpperBound of 2
** A
by A1, Lm6;
then
sup (2 ** A) <= sup A
by XXREAL_2:def 3;
then
(2 * (sup A)) - (sup A) <= (sup A) - (sup A)
by A7, A5, XREAL_1:11;
then
||.((g . ((1 / 2) * (a + b))) - ((1 / 2) * (a + b))).|| = 0
by A6, XXREAL_2:4;
hence
g . ((1 / 2) * (a + b)) = (1 / 2) * (a + b)
by NORMSP_1:10;
verum
end;
theorem Th24:
Lm8:
for E being RealNormSpace
for a, b being Point of E
for t being real number holds ((1 - t) * a) + (t * b) = a + (t * (b - a))
Lm9:
now
let E,
F be
RealNormSpace;
for f being Function of E,F
for a, b being Point of E
for n, j being natural number st f is midpoints-preserving & j <= 2 |^ n holds
ex x being natural number st
( x = (2 |^ n) - j & f . (a + ((j / (2 |^ (n + 1))) * (b - a))) = (1 / 2) * ((f . a) + (f . ((1 / (2 |^ n)) * ((x * a) + (j * b))))) )let f be
Function of
E,
F;
for a, b being Point of E
for n, j being natural number st f is midpoints-preserving & j <= 2 |^ n holds
ex x being natural number st
( x = (2 |^ n) - j & f . (a + ((j / (2 |^ (n + 1))) * (b - a))) = (1 / 2) * ((f . a) + (f . ((1 / (2 |^ n)) * ((x * a) + (j * b))))) )let a,
b be
Point of
E;
for n, j being natural number st f is midpoints-preserving & j <= 2 |^ n holds
ex x being natural number st
( x = (2 |^ n) - j & f . (a + ((j / (2 |^ (n + 1))) * (b - a))) = (1 / 2) * ((f . a) + (f . ((1 / (2 |^ n)) * ((x * a) + (j * b))))) )let n,
j be
natural number ;
( f is midpoints-preserving & j <= 2 |^ n implies ex x being natural number st
( x = (2 |^ n) - j & f . (a + ((j / (2 |^ (n + 1))) * (b - a))) = (1 / 2) * ((f . a) + (f . ((1 / (2 |^ n)) * ((x * a) + (j * b))))) ) )set m = 2
|^ (n + 1);
set k = 2
|^ n;
set x =
(2 |^ n) - j;
assume A1:
f is
midpoints-preserving
;
( j <= 2 |^ n implies ex x being natural number st
( x = (2 |^ n) - j & f . (a + ((j / (2 |^ (n + 1))) * (b - a))) = (1 / 2) * ((f . a) + (f . ((1 / (2 |^ n)) * ((x * a) + (j * b))))) ) )assume
j <= 2
|^ n
;
ex x being natural number st
( x = (2 |^ n) - j & f . (a + ((j / (2 |^ (n + 1))) * (b - a))) = (1 / 2) * ((f . a) + (f . ((1 / (2 |^ n)) * ((x * a) + (j * b))))) )then
(2 |^ n) - j in NAT
by INT_1:18;
then reconsider x =
(2 |^ n) - j as
natural number ;
take x =
x;
( x = (2 |^ n) - j & f . (a + ((j / (2 |^ (n + 1))) * (b - a))) = (1 / 2) * ((f . a) + (f . ((1 / (2 |^ n)) * ((x * a) + (j * b))))) )thus
x = (2 |^ n) - j
;
f . (a + ((j / (2 |^ (n + 1))) * (b - a))) = (1 / 2) * ((f . a) + (f . ((1 / (2 |^ n)) * ((x * a) + (j * b)))))set z =
(1 - (j / (2 |^ (n + 1)))) - (1 / 2);
A2:
2
|^ n <> 0
by NEWTON:102;
A3:
2
* (2 |^ n) = 2
|^ (n + 1)
by NEWTON:11;
A4:
(1 / 2) * (1 / (2 |^ n)) = 1
/ (2 * (2 |^ n))
by XCMPLX_1:103;
x / (2 |^ (n + 1)) =
((1 * (2 |^ n)) / (2 |^ (n + 1))) - (j / (2 |^ (n + 1)))
.=
(1 / 2) - (j / (2 |^ (n + 1)))
by A2, A3, XCMPLX_1:92
.=
(1 - (j / (2 |^ (n + 1)))) - (1 / 2)
;
then A5:
(((1 - (j / (2 |^ (n + 1)))) - (1 / 2)) * a) + ((j / (2 |^ (n + 1))) * b) =
((1 / (2 |^ (n + 1))) * (x * a)) + (((1 / (2 |^ (n + 1))) * j) * b)
by RLVECT_1:def 10
.=
((1 / (2 |^ (n + 1))) * (x * a)) + ((1 / (2 |^ (n + 1))) * (j * b))
by RLVECT_1:def 10
.=
(1 / (2 |^ (n + 1))) * ((x * a) + (j * b))
by RLVECT_1:def 8
.=
(1 / 2) * ((1 / (2 |^ n)) * ((x * a) + (j * b)))
by A4, A3, RLVECT_1:def 10
;
a + ((j / (2 |^ (n + 1))) * (b - a)) =
a + (((j / (2 |^ (n + 1))) * b) - ((j / (2 |^ (n + 1))) * a))
by RLVECT_1:48
.=
(a + ((j / (2 |^ (n + 1))) * b)) - ((j / (2 |^ (n + 1))) * a)
by RLVECT_1:def 6
.=
(a - ((j / (2 |^ (n + 1))) * a)) + ((j / (2 |^ (n + 1))) * b)
by RLVECT_1:def 6
.=
((1 * a) - ((j / (2 |^ (n + 1))) * a)) + ((j / (2 |^ (n + 1))) * b)
by RLVECT_1:def 11
.=
((1 - (j / (2 |^ (n + 1)))) * a) + ((j / (2 |^ (n + 1))) * b)
by RLVECT_1:49
.=
((((1 - (j / (2 |^ (n + 1)))) * a) + (((1 - (j / (2 |^ (n + 1)))) - (1 / 2)) * a)) - (((1 - (j / (2 |^ (n + 1)))) - (1 / 2)) * a)) + ((j / (2 |^ (n + 1))) * b)
by Th4
.=
((((1 - (j / (2 |^ (n + 1)))) * a) - (((1 - (j / (2 |^ (n + 1)))) - (1 / 2)) * a)) + (((1 - (j / (2 |^ (n + 1)))) - (1 / 2)) * a)) + ((j / (2 |^ (n + 1))) * b)
by RLVECT_1:def 6
.=
((((1 - (j / (2 |^ (n + 1)))) - ((1 - (j / (2 |^ (n + 1)))) - (1 / 2))) * a) + (((1 - (j / (2 |^ (n + 1)))) - (1 / 2)) * a)) + ((j / (2 |^ (n + 1))) * b)
by RLVECT_1:49
.=
((1 / 2) * a) + ((((1 - (j / (2 |^ (n + 1)))) - (1 / 2)) * a) + ((j / (2 |^ (n + 1))) * b))
by RLVECT_1:def 6
.=
(1 / 2) * (a + ((1 / (2 |^ n)) * ((x * a) + (j * b))))
by A5, RLVECT_1:def 8
;
hence
f . (a + ((j / (2 |^ (n + 1))) * (b - a))) = (1 / 2) * ((f . a) + (f . ((1 / (2 |^ n)) * ((x * a) + (j * b)))))
by A1, Def3;
verum
end;
Lm10:
now
let E,
F be
RealNormSpace;
for f being Function of E,F
for a, b being Point of E
for n, j being natural number st f is midpoints-preserving & j >= 2 |^ n holds
ex x being natural number st
( x = ((2 |^ n) + j) - (2 |^ (n + 1)) & f . (a + ((j / (2 |^ (n + 1))) * (b - a))) = (1 / 2) * ((f . b) + (f . ((1 / (2 |^ n)) * ((((2 |^ (n + 1)) - j) * a) + (x * b))))) )let f be
Function of
E,
F;
for a, b being Point of E
for n, j being natural number st f is midpoints-preserving & j >= 2 |^ n holds
ex x being natural number st
( x = ((2 |^ n) + j) - (2 |^ (n + 1)) & f . (a + ((j / (2 |^ (n + 1))) * (b - a))) = (1 / 2) * ((f . b) + (f . ((1 / (2 |^ n)) * ((((2 |^ (n + 1)) - j) * a) + (x * b))))) )let a,
b be
Point of
E;
for n, j being natural number st f is midpoints-preserving & j >= 2 |^ n holds
ex x being natural number st
( x = ((2 |^ n) + j) - (2 |^ (n + 1)) & f . (a + ((j / (2 |^ (n + 1))) * (b - a))) = (1 / 2) * ((f . b) + (f . ((1 / (2 |^ n)) * ((((2 |^ (n + 1)) - j) * a) + (x * b))))) )let n,
j be
natural number ;
( f is midpoints-preserving & j >= 2 |^ n implies ex x being natural number st
( x = ((2 |^ n) + j) - (2 |^ (n + 1)) & f . (a + ((j / (2 |^ (n + 1))) * (b - a))) = (1 / 2) * ((f . b) + (f . ((1 / (2 |^ n)) * ((((2 |^ (n + 1)) - j) * a) + (x * b))))) ) )set m = 2
|^ (n + 1);
set k = 2
|^ n;
set x =
((2 |^ n) + j) - (2 |^ (n + 1));
assume A1:
f is
midpoints-preserving
;
( j >= 2 |^ n implies ex x being natural number st
( x = ((2 |^ n) + j) - (2 |^ (n + 1)) & f . (a + ((j / (2 |^ (n + 1))) * (b - a))) = (1 / 2) * ((f . b) + (f . ((1 / (2 |^ n)) * ((((2 |^ (n + 1)) - j) * a) + (x * b))))) ) )A2:
2
* (2 |^ n) = 2
|^ (n + 1)
by NEWTON:11;
assume
j >= 2
|^ n
;
ex x being natural number st
( x = ((2 |^ n) + j) - (2 |^ (n + 1)) & f . (a + ((j / (2 |^ (n + 1))) * (b - a))) = (1 / 2) * ((f . b) + (f . ((1 / (2 |^ n)) * ((((2 |^ (n + 1)) - j) * a) + (x * b))))) )then
(2 |^ n) + (2 |^ n) <= (2 |^ n) + j
by XREAL_1:8;
then
((2 |^ n) + j) - (2 |^ (n + 1)) in NAT
by A2, INT_1:18;
then reconsider x =
((2 |^ n) + j) - (2 |^ (n + 1)) as
natural number ;
take x =
x;
( x = ((2 |^ n) + j) - (2 |^ (n + 1)) & f . (a + ((j / (2 |^ (n + 1))) * (b - a))) = (1 / 2) * ((f . b) + (f . ((1 / (2 |^ n)) * ((((2 |^ (n + 1)) - j) * a) + (x * b))))) )thus
x = ((2 |^ n) + j) - (2 |^ (n + 1))
;
f . (a + ((j / (2 |^ (n + 1))) * (b - a))) = (1 / 2) * ((f . b) + (f . ((1 / (2 |^ n)) * ((((2 |^ (n + 1)) - j) * a) + (x * b)))))set z =
(j / (2 |^ (n + 1))) - (1 / 2);
A3:
2
|^ n <> 0
by NEWTON:102;
A4:
2
|^ (n + 1) <> 0
by NEWTON:102;
A5:
(2 |^ (n + 1)) / (2 |^ (n + 1)) = 1
by A4, XCMPLX_1:60;
then A6:
1
- (j / (2 |^ (n + 1))) = (1 / (2 |^ (n + 1))) * ((2 |^ (n + 1)) - j)
;
A7:
(2 |^ n) / (2 |^ (n + 1)) =
(1 * (2 |^ n)) / (2 * (2 |^ n))
by NEWTON:11
.=
1
/ 2
by A3, XCMPLX_1:92
;
A8:
(1 / 2) * (1 / (2 |^ n)) = 1
/ (2 * (2 |^ n))
by XCMPLX_1:103;
x / (2 |^ (n + 1)) =
(((2 |^ n) + j) / (2 |^ (n + 1))) - ((2 |^ (n + 1)) / (2 |^ (n + 1)))
.=
(j / (2 |^ (n + 1))) - (1 / 2)
by A5, A7
;
then A9:
(((j / (2 |^ (n + 1))) - (1 / 2)) * b) + ((1 - (j / (2 |^ (n + 1)))) * a) =
((1 / (2 |^ (n + 1))) * (((2 |^ (n + 1)) - j) * a)) + (((1 / (2 |^ (n + 1))) * x) * b)
by A6, RLVECT_1:def 10
.=
((1 / (2 |^ (n + 1))) * (((2 |^ (n + 1)) - j) * a)) + ((1 / (2 |^ (n + 1))) * (x * b))
by RLVECT_1:def 10
.=
(1 / (2 |^ (n + 1))) * ((((2 |^ (n + 1)) - j) * a) + (x * b))
by RLVECT_1:def 8
.=
(1 / 2) * ((1 / (2 |^ n)) * ((((2 |^ (n + 1)) - j) * a) + (x * b)))
by A8, A2, RLVECT_1:def 10
;
a + ((j / (2 |^ (n + 1))) * (b - a)) =
a + (((j / (2 |^ (n + 1))) * b) - ((j / (2 |^ (n + 1))) * a))
by RLVECT_1:48
.=
(a + ((j / (2 |^ (n + 1))) * b)) - ((j / (2 |^ (n + 1))) * a)
by RLVECT_1:def 6
.=
(a - ((j / (2 |^ (n + 1))) * a)) + ((j / (2 |^ (n + 1))) * b)
by RLVECT_1:def 6
.=
((1 * a) - ((j / (2 |^ (n + 1))) * a)) + ((j / (2 |^ (n + 1))) * b)
by RLVECT_1:def 11
.=
((1 - (j / (2 |^ (n + 1)))) * a) + ((j / (2 |^ (n + 1))) * b)
by RLVECT_1:49
.=
((((j / (2 |^ (n + 1))) * b) + (((j / (2 |^ (n + 1))) - (1 / 2)) * b)) - (((j / (2 |^ (n + 1))) - (1 / 2)) * b)) + ((1 - (j / (2 |^ (n + 1)))) * a)
by Th4
.=
((((j / (2 |^ (n + 1))) * b) - (((j / (2 |^ (n + 1))) - (1 / 2)) * b)) + (((j / (2 |^ (n + 1))) - (1 / 2)) * b)) + ((1 - (j / (2 |^ (n + 1)))) * a)
by RLVECT_1:def 6
.=
((((j / (2 |^ (n + 1))) - ((j / (2 |^ (n + 1))) - (1 / 2))) * b) + (((j / (2 |^ (n + 1))) - (1 / 2)) * b)) + ((1 - (j / (2 |^ (n + 1)))) * a)
by RLVECT_1:49
.=
((1 / 2) * b) + ((((j / (2 |^ (n + 1))) - (1 / 2)) * b) + ((1 - (j / (2 |^ (n + 1)))) * a))
by RLVECT_1:def 6
.=
(1 / 2) * (b + ((1 / (2 |^ n)) * ((((2 |^ (n + 1)) - j) * a) + (x * b))))
by A9, RLVECT_1:def 8
;
hence
f . (a + ((j / (2 |^ (n + 1))) * (b - a))) = (1 / 2) * ((f . b) + (f . ((1 / (2 |^ n)) * ((((2 |^ (n + 1)) - j) * a) + (x * b)))))
by A1, Def3;
verum
end;
Lm11:
now
let E,
F be
RealNormSpace;
for f being Function of E,F
for a, b being Point of E
for t being natural number st f is midpoints-preserving holds
for n being natural number st t <= 2 |^ n holds
f . (((1 - (t / (2 |^ n))) * a) + ((t / (2 |^ n)) * b)) = ((1 - (t / (2 |^ n))) * (f . a)) + ((t / (2 |^ n)) * (f . b))let f be
Function of
E,
F;
for a, b being Point of E
for t being natural number st f is midpoints-preserving holds
for n being natural number st t <= 2 |^ n holds
f . (((1 - (t / (2 |^ n))) * a) + ((t / (2 |^ n)) * b)) = ((1 - (t / (2 |^ n))) * (f . a)) + ((t / (2 |^ n)) * (f . b))let a,
b be
Point of
E;
for t being natural number st f is midpoints-preserving holds
for n being natural number st t <= 2 |^ n holds
f . (((1 - (t / (2 |^ n))) * a) + ((t / (2 |^ n)) * b)) = ((1 - (t / (2 |^ n))) * (f . a)) + ((t / (2 |^ n)) * (f . b))let t be
natural number ;
( f is midpoints-preserving implies for n being natural number st t <= 2 |^ n holds
f . (((1 - (t / (2 |^ n))) * a) + ((t / (2 |^ n)) * b)) = ((1 - (t / (2 |^ n))) * (f . a)) + ((t / (2 |^ n)) * (f . b)) )assume A1:
f is
midpoints-preserving
;
for n being natural number st t <= 2 |^ n holds
f . (((1 - (t / (2 |^ n))) * a) + ((t / (2 |^ n)) * b)) = ((1 - (t / (2 |^ n))) * (f . a)) + ((t / (2 |^ n)) * (f . b))thus
for
n being
natural number st
t <= 2
|^ n holds
f . (((1 - (t / (2 |^ n))) * a) + ((t / (2 |^ n)) * b)) = ((1 - (t / (2 |^ n))) * (f . a)) + ((t / (2 |^ n)) * (f . b))
verum
proof
defpred S1[
natural number ]
means for
w being
natural number st
w <= 2
|^ $1 holds
f . (((1 - (w / (2 |^ $1))) * a) + ((w / (2 |^ $1)) * b)) = ((1 - (w / (2 |^ $1))) * (f . a)) + ((w / (2 |^ $1)) * (f . b));
A2:
S1[
0 ]
A7:
for
n being
natural number st
S1[
n] holds
S1[
n + 1]
proof
let n be
natural number ;
( S1[n] implies S1[n + 1] )
set m = 2
|^ n;
set k = 2
|^ (n + 1);
assume A8:
S1[
n]
;
S1[n + 1]
let t be
natural number ;
( t <= 2 |^ (n + 1) implies f . (((1 - (t / (2 |^ (n + 1)))) * a) + ((t / (2 |^ (n + 1))) * b)) = ((1 - (t / (2 |^ (n + 1)))) * (f . a)) + ((t / (2 |^ (n + 1))) * (f . b)) )
assume A9:
t <= 2
|^ (n + 1)
;
f . (((1 - (t / (2 |^ (n + 1)))) * a) + ((t / (2 |^ (n + 1))) * b)) = ((1 - (t / (2 |^ (n + 1)))) * (f . a)) + ((t / (2 |^ (n + 1))) * (f . b))
A10:
2
|^ (n + 1) = (2 |^ n) * 2
by NEWTON:11;
A11:
2
|^ n <> 0
by NEWTON:102;
A12:
(1 / 2) * (t / (2 |^ n)) =
(1 * t) / (2 * (2 |^ n))
by XCMPLX_1:77
.=
t / (2 |^ (n + 1))
by NEWTON:11
;
per cases
( t <= 2 |^ n or t >= 2 |^ n )
;
suppose
t >= 2
|^ n
;
f . (((1 - (t / (2 |^ (n + 1)))) * a) + ((t / (2 |^ (n + 1))) * b)) = ((1 - (t / (2 |^ (n + 1)))) * (f . a)) + ((t / (2 |^ (n + 1))) * (f . b))
then consider x being
natural number such that A18:
x = ((2 |^ n) + t) - (2 |^ (n + 1))
and A19:
f . (a + ((t / (2 |^ (n + 1))) * (b - a))) = (1 / 2) * ((f . b) + (f . ((1 / (2 |^ n)) * ((((2 |^ (n + 1)) - t) * a) + (x * b)))))
by A1, Lm10;
set w =
t - (2 |^ n);
A20:
t - (2 |^ n) <= (2 * (2 |^ n)) - (2 |^ n)
by A9, A10, XREAL_1:11;
A21:
(2 |^ n) / (2 |^ n) = 1
by A11, XCMPLX_1:60;
A22:
(2 |^ (n + 1)) / (2 |^ n) =
(2 * (2 |^ n)) / (1 * (2 |^ n))
by NEWTON:11
.=
2
/ 1
by A11, XCMPLX_1:92
;
A23:
(1 / 2) * (1 - ((t - (2 |^ n)) / (2 |^ n))) =
1
- ((1 / 2) * (t / (2 |^ n)))
by A21
.=
1
- ((1 * t) / (2 * (2 |^ n)))
by XCMPLX_1:77
.=
1
- (t / (2 |^ (n + 1)))
by NEWTON:11
;
thus f . (((1 - (t / (2 |^ (n + 1)))) * a) + ((t / (2 |^ (n + 1))) * b)) =
f . (a + ((t / (2 |^ (n + 1))) * (b - a)))
by Lm8
.=
((1 / 2) * (f . b)) + ((1 / 2) * (f . ((1 / (2 |^ n)) * ((((2 |^ (n + 1)) - t) * a) + (x * b)))))
by A19, RLVECT_1:def 8
.=
((1 / 2) * (f . b)) + ((1 / 2) * (f . (((1 / (2 |^ n)) * (((2 |^ (n + 1)) - t) * a)) + ((1 / (2 |^ n)) * (x * b)))))
by RLVECT_1:def 8
.=
((1 / 2) * (f . b)) + ((1 / 2) * (f . ((((1 / (2 |^ n)) * ((2 |^ (n + 1)) - t)) * a) + ((1 / (2 |^ n)) * (x * b)))))
by RLVECT_1:def 10
.=
((1 / 2) * (f . b)) + ((1 / 2) * (f . ((((1 / (2 |^ n)) * ((2 |^ (n + 1)) - t)) * a) + (((1 / (2 |^ n)) * x) * b))))
by RLVECT_1:def 10
.=
((1 / 2) * (f . b)) + ((1 / 2) * (((1 - ((t - (2 |^ n)) / (2 |^ n))) * (f . a)) + (((t - (2 |^ n)) / (2 |^ n)) * (f . b))))
by A20, A21, A22, A18, A10, A8
.=
((1 / 2) * (f . b)) + (((1 / 2) * ((1 - ((t - (2 |^ n)) / (2 |^ n))) * (f . a))) + ((1 / 2) * (((t - (2 |^ n)) / (2 |^ n)) * (f . b))))
by RLVECT_1:def 8
.=
((1 / 2) * ((1 - ((t - (2 |^ n)) / (2 |^ n))) * (f . a))) + (((1 / 2) * (f . b)) + ((1 / 2) * (((t - (2 |^ n)) / (2 |^ n)) * (f . b))))
by RLVECT_1:def 6
.=
((1 / 2) * ((1 - ((t - (2 |^ n)) / (2 |^ n))) * (f . a))) + ((1 / 2) * ((f . b) + (((t - (2 |^ n)) / (2 |^ n)) * (f . b))))
by RLVECT_1:def 8
.=
((1 / 2) * ((1 - ((t - (2 |^ n)) / (2 |^ n))) * (f . a))) + ((1 / 2) * ((1 * (f . b)) + (((t - (2 |^ n)) / (2 |^ n)) * (f . b))))
by RLVECT_1:def 11
.=
((1 / 2) * ((1 - ((t - (2 |^ n)) / (2 |^ n))) * (f . a))) + ((1 / 2) * ((1 + ((t - (2 |^ n)) / (2 |^ n))) * (f . b)))
by RLVECT_1:def 9
.=
((1 / 2) * ((1 - ((t - (2 |^ n)) / (2 |^ n))) * (f . a))) + (((1 / 2) * (1 + ((t - (2 |^ n)) / (2 |^ n)))) * (f . b))
by RLVECT_1:def 10
.=
((1 - (t / (2 |^ (n + 1)))) * (f . a)) + ((t / (2 |^ (n + 1))) * (f . b))
by A23, RLVECT_1:def 10
;
verum
end;
end;
end;
for
n being
natural number holds
S1[
n]
from NAT_1:sch 2(A2, A7);
hence
for
n being
natural number st
t <= 2
|^ n holds
f . (((1 - (t / (2 |^ n))) * a) + ((t / (2 |^ n)) * b)) = ((1 - (t / (2 |^ n))) * (f . a)) + ((t / (2 |^ n)) * (f . b))
;
verum
end;
end;