begin
theorem Th1:
theorem Th2:
theorem
theorem Th4:
theorem
:: deftheorem Def1 defines non-zero NDIFF_1:def 1 :
for X being set
for S being ZeroStr
for f being Function of X,S holds
( f is non-zero iff rng f c= NonZero S );
theorem Th6:
theorem Th7:
:: deftheorem Def2 defines (#) NDIFF_1:def 2 :
for RNS being RealLinearSpace
for S being sequence of RNS
for a being Real_Sequence
for b4 being sequence of RNS holds
( b4 = a (#) S iff for n being Element of NAT holds b4 . n = (a . n) * (S . n) );
:: deftheorem Def3 defines * NDIFF_1:def 3 :
for RNS being RealLinearSpace
for z being Point of RNS
for a being Real_Sequence
for b4 being sequence of RNS holds
( b4 = a * z iff for n being Element of NAT holds b4 . n = (a . n) * z );
theorem
theorem Th9:
theorem Th10:
theorem
theorem Th12:
theorem Th13:
theorem Th14:
theorem Th15:
theorem Th16:
theorem Th17:
theorem
canceled;
theorem
canceled;
theorem
canceled;
:: deftheorem Def4 defines convergent_to_0 NDIFF_1:def 4 :
for S being RealNormSpace
for IT being sequence of S holds
( IT is convergent_to_0 iff ( IT is non-zero & IT is convergent & lim IT = 0. S ) );
theorem Th21:
theorem Th22:
theorem Th23:
theorem
theorem
:: deftheorem Def5 defines REST-like NDIFF_1:def 5 :
for S, T being non trivial RealNormSpace
for IT being PartFunc of S,T holds
( IT is REST-like iff ( IT is total & ( for h being convergent_to_0 sequence of S holds
( (||.h.|| ") (#) (IT /* h) is convergent & lim ((||.h.|| ") (#) (IT /* h)) = 0. T ) ) ) );
theorem
theorem Th27:
theorem
canceled;
theorem
canceled;
theorem Th30:
theorem Th31:
theorem Th32:
theorem Th33:
theorem Th34:
:: deftheorem Def6 defines is_differentiable_in NDIFF_1:def 6 :
for S, T being non trivial RealNormSpace
for f being PartFunc of S,T
for x0 being Point of S holds
( f is_differentiable_in x0 iff ex N being Neighbourhood of x0 st
( N c= dom f & ex L being Point of (R_NormSpace_of_BoundedLinearOperators (S,T)) ex R being REST of S,T st
for x being Point of S st x in N holds
(f /. x) - (f /. x0) = (L . (x - x0)) + (R /. (x - x0)) ) );
definition
let S,
T be non
trivial RealNormSpace;
let f be
PartFunc of
S,
T;
let x0 be
Point of
S;
assume A1:
f is_differentiable_in x0
;
func diff (
f,
x0)
-> Point of
(R_NormSpace_of_BoundedLinearOperators (S,T)) means :
Def7:
ex
N being
Neighbourhood of
x0 st
(
N c= dom f & ex
R being
REST of
S,
T st
for
x being
Point of
S st
x in N holds
(f /. x) - (f /. x0) = (it . (x - x0)) + (R /. (x - x0)) );
existence
ex b1 being Point of (R_NormSpace_of_BoundedLinearOperators (S,T)) ex N being Neighbourhood of x0 st
( N c= dom f & ex R being REST of S,T st
for x being Point of S st x in N holds
(f /. x) - (f /. x0) = (b1 . (x - x0)) + (R /. (x - x0)) )
uniqueness
for b1, b2 being Point of (R_NormSpace_of_BoundedLinearOperators (S,T)) st ex N being Neighbourhood of x0 st
( N c= dom f & ex R being REST of S,T st
for x being Point of S st x in N holds
(f /. x) - (f /. x0) = (b1 . (x - x0)) + (R /. (x - x0)) ) & ex N being Neighbourhood of x0 st
( N c= dom f & ex R being REST of S,T st
for x being Point of S st x in N holds
(f /. x) - (f /. x0) = (b2 . (x - x0)) + (R /. (x - x0)) ) holds
b1 = b2
end;
:: deftheorem Def7 defines diff NDIFF_1:def 7 :
for S, T being non trivial RealNormSpace
for f being PartFunc of S,T
for x0 being Point of S st f is_differentiable_in x0 holds
for b5 being Point of (R_NormSpace_of_BoundedLinearOperators (S,T)) holds
( b5 = diff (f,x0) iff ex N being Neighbourhood of x0 st
( N c= dom f & ex R being REST of S,T st
for x being Point of S st x in N holds
(f /. x) - (f /. x0) = (b5 . (x - x0)) + (R /. (x - x0)) ) );
:: deftheorem Def8 defines is_differentiable_on NDIFF_1:def 8 :
for X being set
for S, T being non trivial RealNormSpace
for f being PartFunc of S,T holds
( f is_differentiable_on X iff ( X c= dom f & ( for x being Point of S st x in X holds
f | X is_differentiable_in x ) ) );
theorem Th35:
theorem Th36:
theorem
definition
let S,
T be non
trivial RealNormSpace;
let f be
PartFunc of
S,
T;
let X be
set ;
assume A1:
f is_differentiable_on X
;
func f `| X -> PartFunc of
S,
(R_NormSpace_of_BoundedLinearOperators (S,T)) means :
Def9:
(
dom it = X & ( for
x being
Point of
S st
x in X holds
it /. x = diff (
f,
x) ) );
existence
ex b1 being PartFunc of S,(R_NormSpace_of_BoundedLinearOperators (S,T)) st
( dom b1 = X & ( for x being Point of S st x in X holds
b1 /. x = diff (f,x) ) )
uniqueness
for b1, b2 being PartFunc of S,(R_NormSpace_of_BoundedLinearOperators (S,T)) st dom b1 = X & ( for x being Point of S st x in X holds
b1 /. x = diff (f,x) ) & dom b2 = X & ( for x being Point of S st x in X holds
b2 /. x = diff (f,x) ) holds
b1 = b2
end;
:: deftheorem Def9 defines `| NDIFF_1:def 9 :
for S, T being non trivial RealNormSpace
for f being PartFunc of S,T
for X being set st f is_differentiable_on X holds
for b5 being PartFunc of S,(R_NormSpace_of_BoundedLinearOperators (S,T)) holds
( b5 = f `| X iff ( dom b5 = X & ( for x being Point of S st x in X holds
b5 /. x = diff (f,x) ) ) );
theorem
theorem Th39:
theorem Th40:
theorem Th41:
theorem Th42:
theorem
theorem
theorem
theorem
theorem
theorem
theorem Th49:
theorem
theorem
theorem