begin
theorem Th1:
theorem Th2:
theorem Th3:
theorem Th4:
theorem Th5:
begin
:: deftheorem Def1 defines is_strictly_convex_on RFUNCT_4:def 1 :
for f being PartFunc of REAL,REAL
for X being set holds
( f is_strictly_convex_on X iff ( X c= dom f & ( for p being Real st 0 < p & p < 1 holds
for r, s being Real st r in X & s in X & (p * r) + ((1 - p) * s) in X & r <> s holds
f . ((p * r) + ((1 - p) * s)) < (p * (f . r)) + ((1 - p) * (f . s)) ) ) );
theorem
theorem
theorem
theorem
theorem
Lm1:
for r being Real
for f being PartFunc of REAL,REAL
for X being set st f is_strictly_convex_on X holds
f - r is_strictly_convex_on X
theorem
Lm2:
for r being Real
for f being PartFunc of REAL,REAL
for X being set st 0 < r & f is_strictly_convex_on X holds
r (#) f is_strictly_convex_on X
theorem Th12:
theorem Th13:
theorem Th14:
theorem
theorem Th16:
theorem
theorem
theorem
begin
:: deftheorem Def2 defines is_quasiconvex_on RFUNCT_4:def 2 :
for f being PartFunc of REAL,REAL
for X being set holds
( f is_quasiconvex_on X iff ( X c= dom f & ( for p being Real st 0 < p & p < 1 holds
for r, s being Real st r in X & s in X & (p * r) + ((1 - p) * s) in X holds
f . ((p * r) + ((1 - p) * s)) <= max ((f . r),(f . s)) ) ) );
:: deftheorem Def3 defines is_strictly_quasiconvex_on RFUNCT_4:def 3 :
for f being PartFunc of REAL,REAL
for X being set holds
( f is_strictly_quasiconvex_on X iff ( X c= dom f & ( for p being Real st 0 < p & p < 1 holds
for r, s being Real st r in X & s in X & (p * r) + ((1 - p) * s) in X & f . r <> f . s holds
f . ((p * r) + ((1 - p) * s)) < max ((f . r),(f . s)) ) ) );
:: deftheorem Def4 defines is_strongly_quasiconvex_on RFUNCT_4:def 4 :
for f being PartFunc of REAL,REAL
for X being set holds
( f is_strongly_quasiconvex_on X iff ( X c= dom f & ( for p being Real st 0 < p & p < 1 holds
for r, s being Real st r in X & s in X & (p * r) + ((1 - p) * s) in X & r <> s holds
f . ((p * r) + ((1 - p) * s)) < max ((f . r),(f . s)) ) ) );
:: deftheorem Def5 defines is_upper_semicontinuous_in RFUNCT_4:def 5 :
for f being PartFunc of REAL,REAL
for x0 being real number holds
( f is_upper_semicontinuous_in x0 iff ( x0 in dom f & ( for r being Real st 0 < r holds
ex s being Real st
( 0 < s & ( for x being Real st x in dom f & abs (x - x0) < s holds
(f . x0) - (f . x) < r ) ) ) ) );
:: deftheorem Def6 defines is_upper_semicontinuous_on RFUNCT_4:def 6 :
for f being PartFunc of REAL,REAL
for X being set holds
( f is_upper_semicontinuous_on X iff ( X c= dom f & ( for x0 being Real st x0 in X holds
f | X is_upper_semicontinuous_in x0 ) ) );
:: deftheorem Def7 defines is_lower_semicontinuous_in RFUNCT_4:def 7 :
for f being PartFunc of REAL,REAL
for x0 being real number holds
( f is_lower_semicontinuous_in x0 iff ( x0 in dom f & ( for r being Real st 0 < r holds
ex s being Real st
( 0 < s & ( for x being Real st x in dom f & abs (x - x0) < s holds
(f . x) - (f . x0) < r ) ) ) ) );
:: deftheorem Def8 defines is_lower_semicontinuous_on RFUNCT_4:def 8 :
for f being PartFunc of REAL,REAL
for X being set holds
( f is_lower_semicontinuous_on X iff ( X c= dom f & ( for x0 being Real st x0 in X holds
f | X is_lower_semicontinuous_in x0 ) ) );
theorem Th20:
theorem
theorem
theorem
theorem
theorem
theorem