definition
let X be non
empty TopSpace;
func C_Algebra_of_ContinuousFunctions X -> ComplexAlgebra equals
ComplexAlgebraStr(#
(CContinuousFunctions X),
(mult_ ((CContinuousFunctions X),(CAlgebra the carrier of X))),
(Add_ ((CContinuousFunctions X),(CAlgebra the carrier of X))),
(Mult_ ((CContinuousFunctions X),(CAlgebra the carrier of X))),
(One_ ((CContinuousFunctions X),(CAlgebra the carrier of X))),
(Zero_ ((CContinuousFunctions X),(CAlgebra the carrier of X))) #);
coherence
ComplexAlgebraStr(# (CContinuousFunctions X),(mult_ ((CContinuousFunctions X),(CAlgebra the carrier of X))),(Add_ ((CContinuousFunctions X),(CAlgebra the carrier of X))),(Mult_ ((CContinuousFunctions X),(CAlgebra the carrier of X))),(One_ ((CContinuousFunctions X),(CAlgebra the carrier of X))),(Zero_ ((CContinuousFunctions X),(CAlgebra the carrier of X))) #) is ComplexAlgebra
by CC0SP1:2;
end;
::
deftheorem defines
C_Algebra_of_ContinuousFunctions CC0SP2:def 4 :
for X being non empty TopSpace holds C_Algebra_of_ContinuousFunctions X = ComplexAlgebraStr(# (CContinuousFunctions X),(mult_ ((CContinuousFunctions X),(CAlgebra the carrier of X))),(Add_ ((CContinuousFunctions X),(CAlgebra the carrier of X))),(Mult_ ((CContinuousFunctions X),(CAlgebra the carrier of X))),(One_ ((CContinuousFunctions X),(CAlgebra the carrier of X))),(Zero_ ((CContinuousFunctions X),(CAlgebra the carrier of X))) #);
Lm1:
for X being non empty compact TopSpace
for x being set st x in CContinuousFunctions X holds
x in ComplexBoundedFunctions the carrier of X
definition
let X be non
empty compact TopSpace;
func C_Normed_Algebra_of_ContinuousFunctions X -> Normed_Complex_AlgebraStr equals
Normed_Complex_AlgebraStr(#
(CContinuousFunctions X),
(mult_ ((CContinuousFunctions X),(CAlgebra the carrier of X))),
(Add_ ((CContinuousFunctions X),(CAlgebra the carrier of X))),
(Mult_ ((CContinuousFunctions X),(CAlgebra the carrier of X))),
(One_ ((CContinuousFunctions X),(CAlgebra the carrier of X))),
(Zero_ ((CContinuousFunctions X),(CAlgebra the carrier of X))),
(CContinuousFunctionsNorm X) #);
correctness
coherence
Normed_Complex_AlgebraStr(# (CContinuousFunctions X),(mult_ ((CContinuousFunctions X),(CAlgebra the carrier of X))),(Add_ ((CContinuousFunctions X),(CAlgebra the carrier of X))),(Mult_ ((CContinuousFunctions X),(CAlgebra the carrier of X))),(One_ ((CContinuousFunctions X),(CAlgebra the carrier of X))),(Zero_ ((CContinuousFunctions X),(CAlgebra the carrier of X))),(CContinuousFunctionsNorm X) #) is Normed_Complex_AlgebraStr ;
;
end;
::
deftheorem defines
C_Normed_Algebra_of_ContinuousFunctions CC0SP2:def 6 :
for X being non empty compact TopSpace holds C_Normed_Algebra_of_ContinuousFunctions X = Normed_Complex_AlgebraStr(# (CContinuousFunctions X),(mult_ ((CContinuousFunctions X),(CAlgebra the carrier of X))),(Add_ ((CContinuousFunctions X),(CAlgebra the carrier of X))),(Mult_ ((CContinuousFunctions X),(CAlgebra the carrier of X))),(One_ ((CContinuousFunctions X),(CAlgebra the carrier of X))),(Zero_ ((CContinuousFunctions X),(CAlgebra the carrier of X))),(CContinuousFunctionsNorm X) #);
Lm2:
now for X being non empty compact TopSpace
for x, e being Element of (C_Normed_Algebra_of_ContinuousFunctions X) st e = One_ ((CContinuousFunctions X),(CAlgebra the carrier of X)) holds
( x * e = x & e * x = x )
let X be non
empty compact TopSpace;
for x, e being Element of (C_Normed_Algebra_of_ContinuousFunctions X) st e = One_ ((CContinuousFunctions X),(CAlgebra the carrier of X)) holds
( x * e = x & e * x = x )set F =
C_Normed_Algebra_of_ContinuousFunctions X;
let x,
e be
Element of
(C_Normed_Algebra_of_ContinuousFunctions X);
( e = One_ ((CContinuousFunctions X),(CAlgebra the carrier of X)) implies ( x * e = x & e * x = x ) )assume A1:
e = One_ (
(CContinuousFunctions X),
(CAlgebra the carrier of X))
;
( x * e = x & e * x = x )set X1 =
CContinuousFunctions X;
reconsider f =
x as
Element of
CContinuousFunctions X ;
1_ (CAlgebra the carrier of X) =
X --> 1
.=
1_ (C_Algebra_of_ContinuousFunctions X)
by Th14
;
then A2:
(
[f,(1_ (CAlgebra the carrier of X))] in [:(CContinuousFunctions X),(CContinuousFunctions X):] &
[(1_ (CAlgebra the carrier of X)),f] in [:(CContinuousFunctions X),(CContinuousFunctions X):] )
by ZFMISC_1:87;
x * e = (mult_ ((CContinuousFunctions X),(CAlgebra the carrier of X))) . (
f,
(1_ (CAlgebra the carrier of X)))
by A1, C0SP1:def 8;
then
x * e = ( the multF of (CAlgebra the carrier of X) || (CContinuousFunctions X)) . (
f,
(1_ (CAlgebra the carrier of X)))
by C0SP1:def 6;
then
x * e = f * (1_ (CAlgebra the carrier of X))
by A2, FUNCT_1:49;
hence
x * e = x
;
e * x = x
e * x = (mult_ ((CContinuousFunctions X),(CAlgebra the carrier of X))) . (
(1_ (CAlgebra the carrier of X)),
f)
by A1, C0SP1:def 8;
then
e * x = ( the multF of (CAlgebra the carrier of X) || (CContinuousFunctions X)) . (
(1_ (CAlgebra the carrier of X)),
f)
by C0SP1:def 6;
then
e * x = (1_ (CAlgebra the carrier of X)) * f
by A2, FUNCT_1:49;
hence
e * x = x
;
verum
end;
Lm3:
for X being non empty compact TopSpace
for x being Point of (C_Normed_Algebra_of_ContinuousFunctions X)
for y being Point of (C_Normed_Algebra_of_BoundedFunctions the carrier of X) st x = y holds
||.x.|| = ||.y.||
by FUNCT_1:49;
Lm4:
for X being non empty compact TopSpace
for x1, x2 being Point of (C_Normed_Algebra_of_ContinuousFunctions X)
for y1, y2 being Point of (C_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 Complex
for x being Point of (C_Normed_Algebra_of_ContinuousFunctions X)
for y being Point of (C_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 (C_Normed_Algebra_of_ContinuousFunctions X)
for y1, y2 being Point of (C_Normed_Algebra_of_BoundedFunctions the carrier of X) st x1 = y1 & x2 = y2 holds
x1 * x2 = y1 * y2
Lm7:
for X being non empty compact TopSpace holds 0. (C_Normed_Algebra_of_ContinuousFunctions X) = 0. (C_Normed_Algebra_of_BoundedFunctions the carrier of X)
Lm8:
for X being non empty compact TopSpace holds 1. (C_Normed_Algebra_of_ContinuousFunctions X) = 1. (C_Normed_Algebra_of_BoundedFunctions the carrier of X)
Lm9:
for X being non empty compact TopSpace
for x1, x2 being Point of (C_Normed_Algebra_of_ContinuousFunctions X)
for y1, y2 being Point of (C_Normed_Algebra_of_BoundedFunctions the carrier of X) st x1 = y1 & x2 = y2 holds
x1 - x2 = y1 - y2
Lm10:
for X being non empty TopSpace
for v, u being Element of (CAlgebra the carrier of X) st v in CC_0_Functions X & u in CC_0_Functions X holds
v + u in CC_0_Functions X
Lm11:
for X being non empty TopSpace
for a being Complex
for u being Element of (CAlgebra the carrier of X) st u in CC_0_Functions X holds
a * u in CC_0_Functions X
Lm12:
for X being non empty TopSpace
for u being Element of (CAlgebra the carrier of X) st u in CC_0_Functions X holds
- u in CC_0_Functions X
Lm13:
for X being non empty TopSpace
for x1, x2 being Point of (C_Normed_Space_of_C_0_Functions X)
for y1, y2 being Point of (C_Normed_Algebra_of_BoundedFunctions the carrier of X) st x1 = y1 & x2 = y2 holds
x1 + x2 = y1 + y2
Lm14:
for X being non empty TopSpace
for a being Complex
for x being Point of (C_Normed_Space_of_C_0_Functions X)
for y being Point of (C_Normed_Algebra_of_BoundedFunctions the carrier of X) st x = y holds
a * x = a * y
Lm15:
for X being non empty TopSpace holds C_Normed_Space_of_C_0_Functions X is ComplexNormSpace-like
by Th46;