begin
Lm1:
for S being RealNormSpace
for x being Point of S
for N1, N2 being Neighbourhood of x holds N1 /\ N2 is Neighbourhood of x
Lm2:
for i, j being Element of NAT st i <= j holds
i + (j -' i) = (i + j) -' i
theorem Th1:
theorem Th2:
Lm3:
for i, j being Element of NAT st i <= j holds
(0* j) /^ i = 0* (j -' i)
Lm4:
for i, j being Element of NAT st i > j holds
(0* j) /^ i = 0* (j -' i)
theorem Th3:
theorem
theorem Th5:
theorem Th6:
theorem Th7:
Lm5:
for m being Element of NAT
for nx being Point of (REAL-NS m)
for Z being Subset of (REAL-NS m)
for i being Element of NAT st Z is open & nx in Z & 1 <= i & i <= m holds
ex N being Neighbourhood of (Proj (i,m)) . nx st
for z being Point of (REAL-NS 1) st z in N holds
(reproj (i,nx)) . z in Z
theorem Th8:
Lm6:
for m being non empty Element of NAT
for v being Element of REAL m
for x being Element of REAL
for i being Element of NAT holds len (Replace (v,i,x)) = m
Lm7:
for m being non empty Element of NAT
for x being Element of REAL
for i, j being Element of NAT st 1 <= j & j <= m holds
( ( i = j implies (Replace ((0* m),i,x)) . j = x ) & ( i <> j implies (Replace ((0* m),i,x)) . j = 0 ) )
theorem Th9:
theorem Th10:
theorem Th11:
theorem Th12:
for
m being non
empty Element of
NAT for
x,
y being
Element of
REAL for
z being
Element of
REAL m for
i being
Element of
NAT st 1
<= i &
i <= m &
y = (proj (i,m)) . z holds
(
(Replace (z,i,x)) - z = Replace (
(0* m),
i,
(x - y)) &
z - (Replace (z,i,x)) = Replace (
(0* m),
i,
(y - x)) )
theorem Th13:
theorem Th14:
theorem Th15:
theorem Th16:
theorem Th17:
theorem Th18:
theorem Th19:
theorem Th20:
theorem Th21:
theorem Th22:
:: deftheorem Def1 defines is_differentiable_in PDIFF_7:def 1 :
for n being non empty Element of NAT
for f being PartFunc of (REAL n),REAL
for x being Element of REAL n holds
( f is_differentiable_in x iff <>* f is_differentiable_in x );
:: deftheorem defines diff PDIFF_7:def 2 :
for n being non empty Element of NAT
for f being PartFunc of (REAL n),REAL
for x being Element of REAL n holds diff (f,x) = (proj (1,1)) * (diff ((<>* f),x));
theorem
for
m being non
empty Element of
NAT for
i being
Element of
NAT for
h being
PartFunc of
(REAL m),
REAL for
y being
Element of
REAL m st
h is_differentiable_in y & 1
<= i &
i <= m holds
(
h is_partial_differentiable_in y,
i &
partdiff (
h,
y,
i)
= diff (
(h * (reproj (i,y))),
((proj (i,m)) . y)) &
partdiff (
h,
y,
i)
= (diff (h,y)) . ((reproj (i,(0* m))) . 1) )
theorem Th24:
theorem Th25:
Lm8:
for m, n being non empty Element of NAT
for f being PartFunc of (REAL m),(REAL n)
for x being Element of REAL m ex L being bounded LinearOperator of m,n st
for h being Element of REAL m ex w being FinSequence of REAL n st
( dom w = Seg m & ( for i being Element of NAT st i in Seg m holds
w . i = ((proj (i,m)) . h) * (partdiff (f,x,i)) ) & L . h = Sum w )
theorem Th26:
theorem Th27:
theorem Th28:
theorem Th29:
theorem Th30:
:: deftheorem Def3 defines open PDIFF_7:def 3 :
for m being non empty Element of NAT
for X being Subset of (REAL m) holds
( X is open iff ex X0 being Subset of (REAL-NS m) st
( X0 = X & X0 is open ) );
theorem Th31:
:: deftheorem Def4 defines is_partial_differentiable_on PDIFF_7:def 4 :
for m, n being non empty Element of NAT
for i being Element of NAT
for f being PartFunc of (REAL m),(REAL n)
for X being set holds
( f is_partial_differentiable_on X,i iff ( X c= dom f & ( for x being Element of REAL m st x in X holds
f | X is_partial_differentiable_in x,i ) ) );
theorem Th32:
theorem Th33:
theorem Th34:
definition
let m,
n be non
empty Element of
NAT ;
let i be
Element of
NAT ;
let f be
PartFunc of
(REAL m),
(REAL n);
let X be
set ;
assume A1:
f is_partial_differentiable_on X,
i
;
func f `partial| (
X,
i)
-> PartFunc of
(REAL m),
(REAL n) means :
Def5:
(
dom it = X & ( for
x being
Element of
REAL m st
x in X holds
it /. x = partdiff (
f,
x,
i) ) );
existence
ex b1 being PartFunc of (REAL m),(REAL n) st
( dom b1 = X & ( for x being Element of REAL m st x in X holds
b1 /. x = partdiff (f,x,i) ) )
uniqueness
for b1, b2 being PartFunc of (REAL m),(REAL n) st dom b1 = X & ( for x being Element of REAL m st x in X holds
b1 /. x = partdiff (f,x,i) ) & dom b2 = X & ( for x being Element of REAL m st x in X holds
b2 /. x = partdiff (f,x,i) ) holds
b1 = b2
end;
:: deftheorem Def5 defines `partial| PDIFF_7:def 5 :
for m, n being non empty Element of NAT
for i being Element of NAT
for f being PartFunc of (REAL m),(REAL n)
for X being set st f is_partial_differentiable_on X,i holds
for b6 being PartFunc of (REAL m),(REAL n) holds
( b6 = f `partial| (X,i) iff ( dom b6 = X & ( for x being Element of REAL m st x in X holds
b6 /. x = partdiff (f,x,i) ) ) );
:: deftheorem Def6 defines is_continuous_in PDIFF_7:def 6 :
for m, n being non empty Element of NAT
for f being PartFunc of (REAL m),(REAL n)
for x0 being Element of REAL m holds
( f is_continuous_in x0 iff ex y0 being Point of (REAL-NS m) ex g being PartFunc of (REAL-NS m),(REAL-NS n) st
( x0 = y0 & f = g & g is_continuous_in y0 ) );
theorem Th35:
theorem
:: deftheorem Def7 defines is_continuous_on PDIFF_7:def 7 :
for m, n being non empty Element of NAT
for f being PartFunc of (REAL m),(REAL n)
for X being set holds
( f is_continuous_on X iff ( X c= dom f & ( for x0 being Element of REAL m st x0 in X holds
f | X is_continuous_in x0 ) ) );
theorem Th37:
theorem Th38:
theorem Th39:
theorem Th40:
theorem Th41:
theorem Th42:
for
m being non
empty Element of
NAT for
f being
PartFunc of
(REAL m),
REAL for
p,
q being
Real for
x being
Element of
REAL m for
i being
Element of
NAT st 1
<= i &
i <= m &
p < q & ( for
h being
Real st
h in [.p,q.] holds
(reproj (i,x)) . h in dom f ) & ( for
h being
Real st
h in [.p,q.] holds
f is_partial_differentiable_in (reproj (i,x)) . h,
i ) holds
ex
r being
Real ex
y being
Element of
REAL m st
(
r in ].p,q.[ &
y = (reproj (i,x)) . r &
(f /. ((reproj (i,x)) . q)) - (f /. ((reproj (i,x)) . p)) = (q - p) * (partdiff (f,y,i)) )
theorem Th43:
for
m being non
empty Element of
NAT for
f being
PartFunc of
(REAL m),
REAL for
p,
q being
Real for
x being
Element of
REAL m for
i being
Element of
NAT st 1
<= i &
i <= m &
p <= q & ( for
h being
Real st
h in [.p,q.] holds
(reproj (i,x)) . h in dom f ) & ( for
h being
Real st
h in [.p,q.] holds
f is_partial_differentiable_in (reproj (i,x)) . h,
i ) holds
ex
r being
Real ex
y being
Element of
REAL m st
(
r in [.p,q.] &
y = (reproj (i,x)) . r &
(f /. ((reproj (i,x)) . q)) - (f /. ((reproj (i,x)) . p)) = (q - p) * (partdiff (f,y,i)) )
theorem Th44:
for
m being non
empty Element of
NAT for
x,
y,
z,
w being
Element of
REAL m for
i being
Element of
NAT for
d,
p,
q,
r being
Real st 1
<= i &
i <= m &
|.(y - x).| < d &
|.(z - x).| < d &
p = (proj (i,m)) . y &
z = (reproj (i,y)) . q &
r in [.p,q.] &
w = (reproj (i,y)) . r holds
|.(w - x).| < d
theorem Th45:
for
m being non
empty Element of
NAT for
f being
PartFunc of
(REAL m),
REAL for
X being
Subset of
(REAL m) for
x,
y,
z being
Element of
REAL m for
i being
Element of
NAT for
d,
p,
q being
Real st 1
<= i &
i <= m &
X is
open &
x in X &
|.(y - x).| < d &
|.(z - x).| < d &
X c= dom f & ( for
x being
Element of
REAL m st
x in X holds
f is_partial_differentiable_in x,
i ) &
0 < d & ( for
z being
Element of
REAL m st
|.(z - x).| < d holds
z in X ) &
z = (reproj (i,y)) . p &
q = (proj (i,m)) . y holds
ex
w being
Element of
REAL m st
(
|.(w - x).| < d &
f is_partial_differentiable_in w,
i &
(f /. z) - (f /. y) = (p - q) * (partdiff (f,w,i)) )
theorem Th46:
theorem Th47:
theorem Th48:
theorem