begin
Lm1:
for r, g, r1 being real number st 0 < g & r <= r1 holds
( r - g < r1 & r < r1 + g )
Lm2:
for seq being Real_Sequence
for f1, f2 being PartFunc of REAL,REAL
for X being Subset of REAL st rng seq c= (dom (f1 (#) f2)) /\ X holds
( rng seq c= dom (f1 (#) f2) & rng seq c= X & dom (f1 (#) f2) = (dom f1) /\ (dom f2) & rng seq c= dom f1 & rng seq c= dom f2 & rng seq c= (dom f1) /\ X & rng seq c= (dom f2) /\ X )
Lm3:
for r being Real
for n being Element of NAT holds
( r - (1 / (n + 1)) < r & r < r + (1 / (n + 1)) )
Lm4:
for seq being Real_Sequence
for f1, f2 being PartFunc of REAL,REAL
for X being Subset of REAL st rng seq c= (dom (f1 + f2)) /\ X holds
( rng seq c= dom (f1 + f2) & rng seq c= X & dom (f1 + f2) = (dom f1) /\ (dom f2) & rng seq c= (dom f1) /\ X & rng seq c= (dom f2) /\ X )
theorem Th1:
theorem Th2:
Lm5:
for x0 being Real
for seq being Real_Sequence
for f being PartFunc of REAL,REAL st ( for g1 being Real ex r being Real st
( x0 < r & ( for r1 being Real st r1 < r & x0 < r1 & r1 in dom f holds
f . r1 < g1 ) ) ) & seq is convergent & lim seq = x0 & rng seq c= (dom f) /\ (right_open_halfline x0) holds
f /* seq is divergent_to-infty
theorem Th3:
theorem Th4:
theorem Th5:
theorem Th6:
:: deftheorem Def1 defines is_left_convergent_in LIMFUNC2:def 1 :
for f being PartFunc of REAL,REAL
for x0 being Real holds
( f is_left_convergent_in x0 iff ( ( for r being Real st r < x0 holds
ex g being Real st
( r < g & g < x0 & g in dom f ) ) & ex g being Real st
for seq being Real_Sequence st seq is convergent & lim seq = x0 & rng seq c= (dom f) /\ (left_open_halfline x0) holds
( f /* seq is convergent & lim (f /* seq) = g ) ) );
:: deftheorem Def2 defines is_left_divergent_to+infty_in LIMFUNC2:def 2 :
for f being PartFunc of REAL,REAL
for x0 being Real holds
( f is_left_divergent_to+infty_in x0 iff ( ( for r being Real st r < x0 holds
ex g being Real st
( r < g & g < x0 & g in dom f ) ) & ( for seq being Real_Sequence st seq is convergent & lim seq = x0 & rng seq c= (dom f) /\ (left_open_halfline x0) holds
f /* seq is divergent_to+infty ) ) );
:: deftheorem Def3 defines is_left_divergent_to-infty_in LIMFUNC2:def 3 :
for f being PartFunc of REAL,REAL
for x0 being Real holds
( f is_left_divergent_to-infty_in x0 iff ( ( for r being Real st r < x0 holds
ex g being Real st
( r < g & g < x0 & g in dom f ) ) & ( for seq being Real_Sequence st seq is convergent & lim seq = x0 & rng seq c= (dom f) /\ (left_open_halfline x0) holds
f /* seq is divergent_to-infty ) ) );
:: deftheorem Def4 defines is_right_convergent_in LIMFUNC2:def 4 :
for f being PartFunc of REAL,REAL
for x0 being Real holds
( f is_right_convergent_in x0 iff ( ( for r being Real st x0 < r holds
ex g being Real st
( g < r & x0 < g & g in dom f ) ) & ex g being Real st
for seq being Real_Sequence st seq is convergent & lim seq = x0 & rng seq c= (dom f) /\ (right_open_halfline x0) holds
( f /* seq is convergent & lim (f /* seq) = g ) ) );
:: deftheorem Def5 defines is_right_divergent_to+infty_in LIMFUNC2:def 5 :
for f being PartFunc of REAL,REAL
for x0 being Real holds
( f is_right_divergent_to+infty_in x0 iff ( ( for r being Real st x0 < r holds
ex g being Real st
( g < r & x0 < g & g in dom f ) ) & ( for seq being Real_Sequence st seq is convergent & lim seq = x0 & rng seq c= (dom f) /\ (right_open_halfline x0) holds
f /* seq is divergent_to+infty ) ) );
:: deftheorem Def6 defines is_right_divergent_to-infty_in LIMFUNC2:def 6 :
for f being PartFunc of REAL,REAL
for x0 being Real holds
( f is_right_divergent_to-infty_in x0 iff ( ( for r being Real st x0 < r holds
ex g being Real st
( g < r & x0 < g & g in dom f ) ) & ( for seq being Real_Sequence st seq is convergent & lim seq = x0 & rng seq c= (dom f) /\ (right_open_halfline x0) holds
f /* seq is divergent_to-infty ) ) );
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem Th31:
theorem
theorem Th33:
theorem
theorem Th35:
theorem
theorem Th37:
theorem
theorem Th39:
theorem Th40:
theorem Th41:
theorem Th42:
theorem
theorem
theorem
theorem
:: deftheorem Def7 defines lim_left LIMFUNC2:def 7 :
for f being PartFunc of REAL,REAL
for x0 being Real st f is_left_convergent_in x0 holds
for b3 being Real holds
( b3 = lim_left (f,x0) iff for seq being Real_Sequence st seq is convergent & lim seq = x0 & rng seq c= (dom f) /\ (left_open_halfline x0) holds
( f /* seq is convergent & lim (f /* seq) = b3 ) );
:: deftheorem Def8 defines lim_right LIMFUNC2:def 8 :
for f being PartFunc of REAL,REAL
for x0 being Real st f is_right_convergent_in x0 holds
for b3 being Real holds
( b3 = lim_right (f,x0) iff for seq being Real_Sequence st seq is convergent & lim seq = x0 & rng seq c= (dom f) /\ (right_open_halfline x0) holds
( f /* seq is convergent & lim (f /* seq) = b3 ) );
theorem
canceled;
theorem
canceled;
theorem
theorem
theorem Th51:
theorem Th52:
theorem Th53:
theorem
theorem
theorem
theorem Th57:
theorem Th58:
theorem
theorem Th60:
theorem Th61:
theorem Th62:
theorem
theorem
theorem
theorem Th66:
theorem Th67:
theorem
theorem
theorem
theorem Th71:
for
x0 being
Real for
f1,
f2,
f being
PartFunc of
REAL,
REAL st
f1 is_left_convergent_in x0 &
f2 is_left_convergent_in x0 &
lim_left (
f1,
x0)
= lim_left (
f2,
x0) & ( for
r being
Real st
r < x0 holds
ex
g being
Real st
(
r < g &
g < x0 &
g in dom f ) ) & ex
r being
Real st
(
0 < r & ( for
g being
Real st
g in (dom f) /\ ].(x0 - r),x0.[ holds
(
f1 . g <= f . g &
f . g <= f2 . g ) ) & ( (
(dom f1) /\ ].(x0 - r),x0.[ c= (dom f2) /\ ].(x0 - r),x0.[ &
(dom f) /\ ].(x0 - r),x0.[ c= (dom f1) /\ ].(x0 - r),x0.[ ) or (
(dom f2) /\ ].(x0 - r),x0.[ c= (dom f1) /\ ].(x0 - r),x0.[ &
(dom f) /\ ].(x0 - r),x0.[ c= (dom f2) /\ ].(x0 - r),x0.[ ) ) ) holds
(
f is_left_convergent_in x0 &
lim_left (
f,
x0)
= lim_left (
f1,
x0) )
theorem
theorem Th73:
for
x0 being
Real for
f1,
f2,
f being
PartFunc of
REAL,
REAL st
f1 is_right_convergent_in x0 &
f2 is_right_convergent_in x0 &
lim_right (
f1,
x0)
= lim_right (
f2,
x0) & ( for
r being
Real st
x0 < r holds
ex
g being
Real st
(
g < r &
x0 < g &
g in dom f ) ) & ex
r being
Real st
(
0 < r & ( for
g being
Real st
g in (dom f) /\ ].x0,(x0 + r).[ holds
(
f1 . g <= f . g &
f . g <= f2 . g ) ) & ( (
(dom f1) /\ ].x0,(x0 + r).[ c= (dom f2) /\ ].x0,(x0 + r).[ &
(dom f) /\ ].x0,(x0 + r).[ c= (dom f1) /\ ].x0,(x0 + r).[ ) or (
(dom f2) /\ ].x0,(x0 + r).[ c= (dom f1) /\ ].x0,(x0 + r).[ &
(dom f) /\ ].x0,(x0 + r).[ c= (dom f2) /\ ].x0,(x0 + r).[ ) ) ) holds
(
f is_right_convergent_in x0 &
lim_right (
f,
x0)
= lim_right (
f1,
x0) )
theorem
theorem
for
x0 being
Real for
f1,
f2 being
PartFunc of
REAL,
REAL st
f1 is_left_convergent_in x0 &
f2 is_left_convergent_in x0 & ex
r being
Real st
(
0 < r & ( (
(dom f1) /\ ].(x0 - r),x0.[ c= (dom f2) /\ ].(x0 - r),x0.[ & ( for
g being
Real st
g in (dom f1) /\ ].(x0 - r),x0.[ holds
f1 . g <= f2 . g ) ) or (
(dom f2) /\ ].(x0 - r),x0.[ c= (dom f1) /\ ].(x0 - r),x0.[ & ( for
g being
Real st
g in (dom f2) /\ ].(x0 - r),x0.[ holds
f1 . g <= f2 . g ) ) ) ) holds
lim_left (
f1,
x0)
<= lim_left (
f2,
x0)
theorem
for
x0 being
Real for
f1,
f2 being
PartFunc of
REAL,
REAL st
f1 is_right_convergent_in x0 &
f2 is_right_convergent_in x0 & ex
r being
Real st
(
0 < r & ( (
(dom f1) /\ ].x0,(x0 + r).[ c= (dom f2) /\ ].x0,(x0 + r).[ & ( for
g being
Real st
g in (dom f1) /\ ].x0,(x0 + r).[ holds
f1 . g <= f2 . g ) ) or (
(dom f2) /\ ].x0,(x0 + r).[ c= (dom f1) /\ ].x0,(x0 + r).[ & ( for
g being
Real st
g in (dom f2) /\ ].x0,(x0 + r).[ holds
f1 . g <= f2 . g ) ) ) ) holds
lim_right (
f1,
x0)
<= lim_right (
f2,
x0)
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem