begin
:: deftheorem defines --> C0SP2:def 1 :
for X being 1-sorted
for y being real number holds X --> y = the carrier of X --> y;
theorem Th1:
:: deftheorem defines ContinuousFunctions C0SP2:def 2 :
for X being non empty TopSpace holds ContinuousFunctions X = { f where f is continuous RealMap of X : verum } ;
definition
let X be non
empty TopSpace;
func R_Algebra_of_ContinuousFunctions X -> AlgebraStr equals
AlgebraStr(#
(ContinuousFunctions X),
(mult_ ((ContinuousFunctions X),(RAlgebra the carrier of X))),
(Add_ ((ContinuousFunctions X),(RAlgebra the carrier of X))),
(Mult_ ((ContinuousFunctions X),(RAlgebra the carrier of X))),
(One_ ((ContinuousFunctions X),(RAlgebra the carrier of X))),
(Zero_ ((ContinuousFunctions X),(RAlgebra the carrier of X))) #);
coherence
AlgebraStr(# (ContinuousFunctions X),(mult_ ((ContinuousFunctions X),(RAlgebra the carrier of X))),(Add_ ((ContinuousFunctions X),(RAlgebra the carrier of X))),(Mult_ ((ContinuousFunctions X),(RAlgebra the carrier of X))),(One_ ((ContinuousFunctions X),(RAlgebra the carrier of X))),(Zero_ ((ContinuousFunctions X),(RAlgebra the carrier of X))) #) is AlgebraStr
;
end;
:: deftheorem defines R_Algebra_of_ContinuousFunctions C0SP2:def 3 :
for X being non empty TopSpace holds R_Algebra_of_ContinuousFunctions X = AlgebraStr(# (ContinuousFunctions X),(mult_ ((ContinuousFunctions X),(RAlgebra the carrier of X))),(Add_ ((ContinuousFunctions X),(RAlgebra the carrier of X))),(Mult_ ((ContinuousFunctions X),(RAlgebra the carrier of X))),(One_ ((ContinuousFunctions X),(RAlgebra the carrier of X))),(Zero_ ((ContinuousFunctions X),(RAlgebra the carrier of X))) #);
theorem
theorem Th3:
theorem Th4:
theorem Th5:
theorem Th6:
theorem Th7:
theorem Th8:
Lm1:
for X being non empty compact TopSpace
for x being set st x in ContinuousFunctions X holds
x in BoundedFunctions the carrier of X
theorem
:: deftheorem defines ContinuousFunctionsNorm C0SP2:def 4 :
for X being non empty compact TopSpace holds ContinuousFunctionsNorm X = (BoundedFunctionsNorm the carrier of X) | (ContinuousFunctions X);
definition
let X be non
empty compact TopSpace;
func R_Normed_Algebra_of_ContinuousFunctions X -> Normed_AlgebraStr equals
Normed_AlgebraStr(#
(ContinuousFunctions X),
(mult_ ((ContinuousFunctions X),(RAlgebra the carrier of X))),
(Add_ ((ContinuousFunctions X),(RAlgebra the carrier of X))),
(Mult_ ((ContinuousFunctions X),(RAlgebra the carrier of X))),
(One_ ((ContinuousFunctions X),(RAlgebra the carrier of X))),
(Zero_ ((ContinuousFunctions X),(RAlgebra the carrier of X))),
(ContinuousFunctionsNorm X) #);
correctness
coherence
Normed_AlgebraStr(# (ContinuousFunctions X),(mult_ ((ContinuousFunctions X),(RAlgebra the carrier of X))),(Add_ ((ContinuousFunctions X),(RAlgebra the carrier of X))),(Mult_ ((ContinuousFunctions X),(RAlgebra the carrier of X))),(One_ ((ContinuousFunctions X),(RAlgebra the carrier of X))),(Zero_ ((ContinuousFunctions X),(RAlgebra the carrier of X))),(ContinuousFunctionsNorm X) #) is Normed_AlgebraStr ;
;
end;
:: deftheorem defines R_Normed_Algebra_of_ContinuousFunctions C0SP2:def 5 :
for X being non empty compact TopSpace holds R_Normed_Algebra_of_ContinuousFunctions X = Normed_AlgebraStr(# (ContinuousFunctions X),(mult_ ((ContinuousFunctions X),(RAlgebra the carrier of X))),(Add_ ((ContinuousFunctions X),(RAlgebra the carrier of X))),(Mult_ ((ContinuousFunctions X),(RAlgebra the carrier of X))),(One_ ((ContinuousFunctions X),(RAlgebra the carrier of X))),(Zero_ ((ContinuousFunctions X),(RAlgebra the carrier of X))),(ContinuousFunctionsNorm X) #);
Lm2:
now
let X be non
empty compact TopSpace;
for x, e being Element of (R_Normed_Algebra_of_ContinuousFunctions X) st e = One_ ((ContinuousFunctions X),(RAlgebra the carrier of X)) holds
( x * e = x & e * x = x )set F =
R_Normed_Algebra_of_ContinuousFunctions X;
let x,
e be
Element of
(R_Normed_Algebra_of_ContinuousFunctions X);
( e = One_ ((ContinuousFunctions X),(RAlgebra the carrier of X)) implies ( x * e = x & e * x = x ) )assume A1:
e = One_ (
(ContinuousFunctions X),
(RAlgebra the carrier of X))
;
( x * e = x & e * x = x )set X1 =
ContinuousFunctions X;
reconsider f =
x as
Element of
ContinuousFunctions X ;
1_ (RAlgebra the carrier of X) =
X --> 1
.=
1_ (R_Algebra_of_ContinuousFunctions X)
by Th7
;
then A2:
(
[f,(1_ (RAlgebra the carrier of X))] in [:(ContinuousFunctions X),(ContinuousFunctions X):] &
[(1_ (RAlgebra the carrier of X)),f] in [:(ContinuousFunctions X),(ContinuousFunctions X):] )
by ZFMISC_1:106;
x * e = (mult_ ((ContinuousFunctions X),(RAlgebra the carrier of X))) . (
f,
(1_ (RAlgebra the carrier of X)))
by A1, C0SP1:def 8;
then
x * e = ( the multF of (RAlgebra the carrier of X) || (ContinuousFunctions X)) . (
f,
(1_ (RAlgebra the carrier of X)))
by C0SP1:def 6;
then
x * e = f * (1_ (RAlgebra the carrier of X))
by A2, FUNCT_1:72;
hence
x * e = x
by VECTSP_1:def 13;
e * x = x
e * x = (mult_ ((ContinuousFunctions X),(RAlgebra the carrier of X))) . (
(1_ (RAlgebra the carrier of X)),
f)
by A1, C0SP1:def 8;
then
e * x = ( the multF of (RAlgebra the carrier of X) || (ContinuousFunctions X)) . (
(1_ (RAlgebra the carrier of X)),
f)
by C0SP1:def 6;
then
e * x = (1_ (RAlgebra the carrier of X)) * f
by A2, FUNCT_1:72;
hence
e * x = x
by VECTSP_1:def 13;
verum
end;
theorem Th10:
Lm3:
for X being non empty compact TopSpace
for x being Point of (R_Normed_Algebra_of_ContinuousFunctions X)
for y being Point of (R_Normed_Algebra_of_BoundedFunctions the carrier of X) st x = y holds
||.x.|| = ||.y.||
by FUNCT_1:72;
Lm4:
for X being non empty compact TopSpace
for x1, x2 being Point of (R_Normed_Algebra_of_ContinuousFunctions X)
for y1, y2 being Point of (R_Normed_Algebra_of_BoundedFunctions the carrier of X) st x1 = y1 & x2 = y2 holds
x1 + x2 = y1 + y2
Lm5:
for X being non empty compact TopSpace
for a being Real
for x being Point of (R_Normed_Algebra_of_ContinuousFunctions X)
for y being Point of (R_Normed_Algebra_of_BoundedFunctions the carrier of X) st x = y holds
a * x = a * y
Lm6:
for X being non empty compact TopSpace
for x1, x2 being Point of (R_Normed_Algebra_of_ContinuousFunctions X)
for y1, y2 being Point of (R_Normed_Algebra_of_BoundedFunctions the carrier of X) st x1 = y1 & x2 = y2 holds
x1 * x2 = y1 * y2
theorem Th11:
theorem Th12:
Lm7:
for X being non empty compact TopSpace holds 0. (R_Normed_Algebra_of_ContinuousFunctions X) = 0. (R_Normed_Algebra_of_BoundedFunctions the carrier of X)
Lm8:
for X being non empty compact TopSpace holds 1. (R_Normed_Algebra_of_ContinuousFunctions X) = 1. (R_Normed_Algebra_of_BoundedFunctions the carrier of X)
theorem
theorem
theorem Th15:
theorem
theorem
theorem Th18:
Lm9:
for X being non empty compact TopSpace
for x1, x2 being Point of (R_Normed_Algebra_of_ContinuousFunctions X)
for y1, y2 being Point of (R_Normed_Algebra_of_BoundedFunctions the carrier of X) st x1 = y1 & x2 = y2 holds
x1 - x2 = y1 - y2
theorem
theorem Th20:
theorem Th21:
theorem Th22:
begin
theorem Th23:
theorem Th24:
theorem
begin
:: deftheorem defines C_0_Functions C0SP2:def 6 :
for X being non empty TopSpace holds C_0_Functions X = { f where f is RealMap of X : ( f is continuous & ex Y being non empty Subset of X st
( Y is compact & ( for A being Subset of X st A = support f holds
Cl A is Subset of Y ) ) ) } ;
theorem
Lm10:
for X being non empty TopSpace
for v, u being Element of (RAlgebra the carrier of X) st v in C_0_Functions X & u in C_0_Functions X holds
v + u in C_0_Functions X
Lm11:
for X being non empty TopSpace
for a being Real
for u being Element of (RAlgebra the carrier of X) st u in C_0_Functions X holds
a * u in C_0_Functions X
Lm13:
for X being non empty TopSpace
for u being Element of (RAlgebra the carrier of X) st u in C_0_Functions X holds
- u in C_0_Functions X
theorem
theorem Th28:
definition
let X be non
empty TopSpace;
func R_VectorSpace_of_C_0_Functions X -> RealLinearSpace equals
RLSStruct(#
(C_0_Functions X),
(Zero_ ((C_0_Functions X),(RealVectSpace the carrier of X))),
(Add_ ((C_0_Functions X),(RealVectSpace the carrier of X))),
(Mult_ ((C_0_Functions X),(RealVectSpace the carrier of X))) #);
correctness
coherence
RLSStruct(# (C_0_Functions X),(Zero_ ((C_0_Functions X),(RealVectSpace the carrier of X))),(Add_ ((C_0_Functions X),(RealVectSpace the carrier of X))),(Mult_ ((C_0_Functions X),(RealVectSpace the carrier of X))) #) is RealLinearSpace;
by RSSPACE:13;
end;
:: deftheorem defines R_VectorSpace_of_C_0_Functions C0SP2:def 7 :
for X being non empty TopSpace holds R_VectorSpace_of_C_0_Functions X = RLSStruct(# (C_0_Functions X),(Zero_ ((C_0_Functions X),(RealVectSpace the carrier of X))),(Add_ ((C_0_Functions X),(RealVectSpace the carrier of X))),(Mult_ ((C_0_Functions X),(RealVectSpace the carrier of X))) #);
theorem
theorem Th30:
:: deftheorem defines C_0_FunctionsNorm C0SP2:def 8 :
for X being non empty TopSpace holds C_0_FunctionsNorm X = (BoundedFunctionsNorm the carrier of X) | (C_0_Functions X);
definition
let X be non
empty TopSpace;
func R_Normed_Space_of_C_0_Functions X -> non
empty NORMSTR equals
NORMSTR(#
(C_0_Functions X),
(Zero_ ((C_0_Functions X),(RealVectSpace the carrier of X))),
(Add_ ((C_0_Functions X),(RealVectSpace the carrier of X))),
(Mult_ ((C_0_Functions X),(RealVectSpace the carrier of X))),
(C_0_FunctionsNorm X) #);
correctness
coherence
NORMSTR(# (C_0_Functions X),(Zero_ ((C_0_Functions X),(RealVectSpace the carrier of X))),(Add_ ((C_0_Functions X),(RealVectSpace the carrier of X))),(Mult_ ((C_0_Functions X),(RealVectSpace the carrier of X))),(C_0_FunctionsNorm X) #) is non empty NORMSTR ;
;
end;
:: deftheorem defines R_Normed_Space_of_C_0_Functions C0SP2:def 9 :
for X being non empty TopSpace holds R_Normed_Space_of_C_0_Functions X = NORMSTR(# (C_0_Functions X),(Zero_ ((C_0_Functions X),(RealVectSpace the carrier of X))),(Add_ ((C_0_Functions X),(RealVectSpace the carrier of X))),(Mult_ ((C_0_Functions X),(RealVectSpace the carrier of X))),(C_0_FunctionsNorm X) #);
theorem
theorem Th32:
theorem Th33:
Lm15:
for X being non empty TopSpace
for x1, x2 being Point of (R_Normed_Space_of_C_0_Functions X)
for y1, y2 being Point of (R_Normed_Algebra_of_BoundedFunctions the carrier of X) st x1 = y1 & x2 = y2 holds
x1 + x2 = y1 + y2
Lm16:
for X being non empty TopSpace
for a being Real
for x being Point of (R_Normed_Space_of_C_0_Functions X)
for y being Point of (R_Normed_Algebra_of_BoundedFunctions the carrier of X) st x = y holds
a * x = a * y
theorem Th34:
theorem Th35:
theorem