begin
Lm1:
for f being FinSequence
for h being Function st dom h = dom f holds
h is FinSequence
Lm2:
for f, g being FinSequence
for h being Function st dom h = (dom f) /\ (dom g) holds
h is FinSequence
:: deftheorem Def1 defines + VALUED_1:def 1 :
for f1, f2 being complex-valued Function
for b3 being Function holds
( b3 = f1 + f2 iff ( dom b3 = (dom f1) /\ (dom f2) & ( for c being set st c in dom b3 holds
b3 . c = (f1 . c) + (f2 . c) ) ) );
theorem
begin
:: deftheorem Def2 defines + VALUED_1:def 2 :
for f being complex-valued Function
for r being complex number
for b3 being Function holds
( b3 = r + f iff ( dom b3 = dom f & ( for c being set st c in dom b3 holds
b3 . c = r + (f . c) ) ) );
theorem
begin
:: deftheorem defines - VALUED_1:def 3 :
for f being complex-valued Function
for r being complex number holds f - r = (- r) + f;
theorem
theorem
begin
:: deftheorem Def4 defines (#) VALUED_1:def 4 :
for f1, f2 being complex-valued Function
for b3 being Function holds
( b3 = f1 (#) f2 iff ( dom b3 = (dom f1) /\ (dom f2) & ( for c being set st c in dom b3 holds
b3 . c = (f1 . c) * (f2 . c) ) ) );
theorem
begin
:: deftheorem Def5 defines (#) VALUED_1:def 5 :
for f being complex-valued Function
for r being complex number
for b3 being Function holds
( b3 = r (#) f iff ( dom b3 = dom f & ( for c being set st c in dom b3 holds
b3 . c = r * (f . c) ) ) );
theorem Th6:
theorem
begin
:: deftheorem defines - VALUED_1:def 6 :
for f being complex-valued Function holds - f = (- 1) (#) f;
theorem Th8:
theorem
begin
:: deftheorem Def7 defines " VALUED_1:def 7 :
for f, b2 being complex-valued Function holds
( b2 = f " iff ( dom b2 = dom f & ( for c being set st c in dom b2 holds
b2 . c = (f . c) " ) ) );
theorem Th10:
begin
:: deftheorem defines ^2 VALUED_1:def 8 :
for f being complex-valued Function holds f ^2 = f (#) f;
theorem Th11:
begin
:: deftheorem defines - VALUED_1:def 9 :
for f1, f2 being complex-valued Function holds f1 - f2 = f1 + (- f2);
theorem Th12:
theorem
theorem
Lm3:
for C being set
for D1, D2 being non empty complex-membered set
for f1 being Function of C,D1
for f2 being Function of C,D2 holds dom (f1 - f2) = C
theorem
begin
:: deftheorem defines /" VALUED_1:def 10 :
for f1, f2 being complex-valued Function holds f1 /" f2 = f1 (#) (f2 ");
theorem Th16:
theorem
Lm4:
for C being set
for D1, D2 being non empty complex-membered set
for f1 being Function of C,D1
for f2 being Function of C,D2 holds dom (f1 /" f2) = C
begin
:: deftheorem Def11 defines |. VALUED_1:def 11 :
for f being complex-valued Function
for b2 being real-valued Function holds
( b2 = |.f.| iff ( dom b2 = dom f & ( for c being set st c in dom b2 holds
b2 . c = |.(f . c).| ) ) );
theorem
theorem
begin
:: deftheorem Def12 defines Shift VALUED_1:def 12 :
for p being Function
for k being Element of NAT
for b3 being Function holds
( b3 = Shift (p,k) iff ( dom b3 = { (m + k) where m is Element of NAT : m in dom p } & ( for m being Element of NAT st m in dom p holds
b3 . (m + k) = p . m ) ) );
theorem
canceled;
theorem
theorem
theorem
theorem
theorem Th25:
theorem Th26:
theorem
:: deftheorem defines increasing VALUED_1:def 13 :
for X being non empty ext-real-membered set
for s being sequence of X holds
( s is increasing iff for n being Nat holds s . n < s . (n + 1) );
:: deftheorem defines decreasing VALUED_1:def 14 :
for X being non empty ext-real-membered set
for s being sequence of X holds
( s is decreasing iff for n being Nat holds s . n > s . (n + 1) );
:: deftheorem defines non-decreasing VALUED_1:def 15 :
for X being non empty ext-real-membered set
for s being sequence of X holds
( s is non-decreasing iff for n being Nat holds s . n <= s . (n + 1) );
:: deftheorem defines non-increasing VALUED_1:def 16 :
for X being non empty ext-real-membered set
for s being sequence of X holds
( s is non-increasing iff for n being Nat holds s . n >= s . (n + 1) );
theorem
theorem
theorem
:: deftheorem defines LastLoc VALUED_1:def 17 :
for F being NAT -defined non empty finite Function holds LastLoc F = max (dom F);
:: deftheorem defines CutLastLoc VALUED_1:def 18 :
for F being NAT -defined non empty finite Function holds CutLastLoc F = F \ ((LastLoc F) .--> (F . (LastLoc F)));
theorem Th31:
theorem
theorem
:: deftheorem defines FirstLoc VALUED_1:def 19 :
for F being NAT -defined non empty Function holds FirstLoc F = min (dom F);
theorem
theorem
theorem
theorem Th47:
theorem Th48:
theorem Th49: