begin
definition
let C be non
empty set ;
let V be non
empty addLoopStr ;
let f1,
f2 be
PartFunc of
C,
V;
deffunc H1(
set )
-> Element of the
carrier of
V =
(f1 /. $1) + (f2 /. $1);
deffunc H2(
set )
-> Element of the
carrier of
V =
(f1 /. $1) - (f2 /. $1);
defpred S1[
set ]
means $1
in (dom f1) /\ (dom f2);
set X =
(dom f1) /\ (dom f2);
func f1 + f2 -> PartFunc of
C,
V means :
Def1:
(
dom it = (dom f1) /\ (dom f2) & ( for
c being
Element of
C st
c in dom it holds
it /. c = (f1 /. c) + (f2 /. c) ) );
existence
ex b1 being PartFunc of C,V st
( dom b1 = (dom f1) /\ (dom f2) & ( for c being Element of C st c in dom b1 holds
b1 /. c = (f1 /. c) + (f2 /. c) ) )
uniqueness
for b1, b2 being PartFunc of C,V st dom b1 = (dom f1) /\ (dom f2) & ( for c being Element of C st c in dom b1 holds
b1 /. c = (f1 /. c) + (f2 /. c) ) & dom b2 = (dom f1) /\ (dom f2) & ( for c being Element of C st c in dom b2 holds
b2 /. c = (f1 /. c) + (f2 /. c) ) holds
b1 = b2
func f1 - f2 -> PartFunc of
C,
V means :
Def2:
(
dom it = (dom f1) /\ (dom f2) & ( for
c being
Element of
C st
c in dom it holds
it /. c = (f1 /. c) - (f2 /. c) ) );
existence
ex b1 being PartFunc of C,V st
( dom b1 = (dom f1) /\ (dom f2) & ( for c being Element of C st c in dom b1 holds
b1 /. c = (f1 /. c) - (f2 /. c) ) )
uniqueness
for b1, b2 being PartFunc of C,V st dom b1 = (dom f1) /\ (dom f2) & ( for c being Element of C st c in dom b1 holds
b1 /. c = (f1 /. c) - (f2 /. c) ) & dom b2 = (dom f1) /\ (dom f2) & ( for c being Element of C st c in dom b2 holds
b2 /. c = (f1 /. c) - (f2 /. c) ) holds
b1 = b2
end;
:: deftheorem Def1 defines + VFUNCT_1:def 1 :
for C being non empty set
for V being non empty addLoopStr
for f1, f2, b5 being PartFunc of C,V holds
( b5 = f1 + f2 iff ( dom b5 = (dom f1) /\ (dom f2) & ( for c being Element of C st c in dom b5 holds
b5 /. c = (f1 /. c) + (f2 /. c) ) ) );
:: deftheorem Def2 defines - VFUNCT_1:def 2 :
for C being non empty set
for V being non empty addLoopStr
for f1, f2, b5 being PartFunc of C,V holds
( b5 = f1 - f2 iff ( dom b5 = (dom f1) /\ (dom f2) & ( for c being Element of C st c in dom b5 holds
b5 /. c = (f1 /. c) - (f2 /. c) ) ) );
:: deftheorem Def3 defines (#) VFUNCT_1:def 3 :
for C being non empty set
for V being non empty RLSStruct
for f1 being PartFunc of C,REAL
for f2, b5 being PartFunc of C,V holds
( b5 = f1 (#) f2 iff ( dom b5 = (dom f1) /\ (dom f2) & ( for c being Element of C st c in dom b5 holds
b5 /. c = (f1 . c) * (f2 /. c) ) ) );
:: deftheorem Def4 defines (#) VFUNCT_1:def 4 :
for C being non empty set
for V being non empty RLSStruct
for f being PartFunc of C,V
for r being Real
for b5 being PartFunc of C,V holds
( b5 = r (#) f iff ( dom b5 = dom f & ( for c being Element of C st c in dom b5 holds
b5 /. c = r * (f /. c) ) ) );
:: deftheorem VFUNCT_1:def 5 :
canceled;
:: deftheorem Def6 defines - VFUNCT_1:def 6 :
for C being non empty set
for V being non empty addLoopStr
for f, b4 being PartFunc of C,V holds
( b4 = - f iff ( dom b4 = dom f & ( for c being Element of C st c in dom b4 holds
b4 /. c = - (f /. c) ) ) );
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
theorem
theorem
theorem Th10:
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem Th20:
theorem
theorem
theorem
theorem Th24:
theorem
theorem
theorem
theorem
theorem Th29:
theorem Th30:
theorem Th31:
theorem
theorem Th33:
theorem
theorem Th35:
theorem
theorem
theorem
theorem
theorem Th40:
theorem
theorem Th42:
theorem
theorem
theorem
theorem
:: deftheorem Def7 defines is_bounded_on VFUNCT_1:def 7 :
for C being non empty set
for V being non empty NORMSTR
for f being PartFunc of C,V
for Y being set holds
( f is_bounded_on Y iff ex r being Real st
for c being Element of C st c in Y /\ (dom f) holds
||.(f /. c).|| <= r );
theorem
canceled;
theorem
theorem
theorem
theorem Th51:
theorem Th52:
theorem Th53:
theorem
theorem
theorem
theorem
theorem
theorem Th59:
theorem Th60:
theorem Th61:
theorem
theorem
theorem