:: Complex Sequences
:: by Agnieszka Banachowicz and Anna Winnicka
::
:: Received November 5, 1993
:: Copyright (c) 1993 Association of Mizar Users



definition
mode Complex_Sequence is sequence of COMPLEX ;
end;

theorem Th1: :: COMSEQ_1:1
for f being Function holds
( f is Complex_Sequence iff ( dom f = NAT & ( for x being set st x in NAT holds
f . x is Element of COMPLEX ) ) )
proof end;

theorem Th2: :: COMSEQ_1:2
for f being Function holds
( f is Complex_Sequence iff ( dom f = NAT & ( for n being Element of NAT holds f . n is Element of COMPLEX ) ) )
proof end;

scheme :: COMSEQ_1:sch 1
ExComplexSeq{ F1( set ) -> complex number } :
ex seq being Complex_Sequence st
for n being Element of NAT holds seq . n = F1(n)
proof end;

notation
let f be complex-valued Relation;
synonym non-zero f for non-empty ;
end;

registration
cluster non-zero Element of K21(K22(NAT ,COMPLEX ));
existence
ex b1 being Complex_Sequence st b1 is non-zero
proof end;
end;

theorem Th3: :: COMSEQ_1:3
for seq being Complex_Sequence holds
( seq is non-zero iff for x being set st x in NAT holds
seq . x <> 0c )
proof end;

theorem Th4: :: COMSEQ_1:4
for seq being Complex_Sequence holds
( seq is non-zero iff for n being Element of NAT holds seq . n <> 0c )
proof end;

theorem :: COMSEQ_1:5
for IT being non-zero Complex_Sequence holds rng IT c= COMPLEX \ {0c }
proof end;

theorem :: COMSEQ_1:6
canceled;

theorem :: COMSEQ_1:7
for r being Element of COMPLEX ex seq being Complex_Sequence st rng seq = {r}
proof end;

theorem :: COMSEQ_1:8
canceled;

theorem Th9: :: COMSEQ_1:9
for seq1, seq2, seq3 being Complex_Sequence holds (seq1 + seq2) + seq3 = seq1 + (seq2 + seq3)
proof end;

theorem :: COMSEQ_1:10
canceled;

theorem Th11: :: COMSEQ_1:11
for seq1, seq2, seq3 being Complex_Sequence holds (seq1 (#) seq2) (#) seq3 = seq1 (#) (seq2 (#) seq3)
proof end;

theorem Th12: :: COMSEQ_1:12
for seq1, seq2, seq3 being Complex_Sequence holds (seq1 + seq2) (#) seq3 = (seq1 (#) seq3) + (seq2 (#) seq3)
proof end;

theorem :: COMSEQ_1:13
for seq3, seq1, seq2 being Complex_Sequence holds seq3 (#) (seq1 + seq2) = (seq3 (#) seq1) + (seq3 (#) seq2) by Th12;

theorem :: COMSEQ_1:14
for seq being Complex_Sequence holds - seq = (- 1r ) (#) seq ;

theorem Th15: :: COMSEQ_1:15
for r being Element of COMPLEX
for seq1, seq2 being Complex_Sequence holds r (#) (seq1 (#) seq2) = (r (#) seq1) (#) seq2
proof end;

theorem Th16: :: COMSEQ_1:16
for r being Element of COMPLEX
for seq1, seq2 being Complex_Sequence holds r (#) (seq1 (#) seq2) = seq1 (#) (r (#) seq2)
proof end;

theorem Th17: :: COMSEQ_1:17
for seq1, seq2, seq3 being Complex_Sequence holds (seq1 - seq2) (#) seq3 = (seq1 (#) seq3) - (seq2 (#) seq3)
proof end;

theorem :: COMSEQ_1:18
for seq3, seq1, seq2 being Complex_Sequence holds (seq3 (#) seq1) - (seq3 (#) seq2) = seq3 (#) (seq1 - seq2) by Th17;

theorem Th19: :: COMSEQ_1:19
for r being Element of COMPLEX
for seq1, seq2 being Complex_Sequence holds r (#) (seq1 + seq2) = (r (#) seq1) + (r (#) seq2)
proof end;

theorem Th20: :: COMSEQ_1:20
for r, p being Element of COMPLEX
for seq being Complex_Sequence holds (r * p) (#) seq = r (#) (p (#) seq)
proof end;

theorem Th21: :: COMSEQ_1:21
for r being Element of COMPLEX
for seq1, seq2 being Complex_Sequence holds r (#) (seq1 - seq2) = (r (#) seq1) - (r (#) seq2)
proof end;

theorem :: COMSEQ_1:22
for r being Element of COMPLEX
for seq1, seq being Complex_Sequence holds r (#) (seq1 /" seq) = (r (#) seq1) /" seq
proof end;

theorem :: COMSEQ_1:23
for seq1, seq2, seq3 being Complex_Sequence holds seq1 - (seq2 + seq3) = (seq1 - seq2) - seq3
proof end;

theorem :: COMSEQ_1:24
for seq being Complex_Sequence holds 1r (#) seq = seq
proof end;

theorem :: COMSEQ_1:25
for seq being Complex_Sequence holds - (- seq) = seq ;

theorem :: COMSEQ_1:26
for seq1, seq2 being Complex_Sequence holds seq1 - (- seq2) = seq1 + seq2 ;

theorem :: COMSEQ_1:27
for seq1, seq2, seq3 being Complex_Sequence holds seq1 - (seq2 - seq3) = (seq1 - seq2) + seq3
proof end;

theorem :: COMSEQ_1:28
for seq1, seq2, seq3 being Complex_Sequence holds seq1 + (seq2 - seq3) = (seq1 + seq2) - seq3
proof end;

theorem :: COMSEQ_1:29
for seq1, seq2 being Complex_Sequence holds
( (- seq1) (#) seq2 = - (seq1 (#) seq2) & seq1 (#) (- seq2) = - (seq1 (#) seq2) )
proof end;

theorem Th30: :: COMSEQ_1:30
for seq being Complex_Sequence st seq is non-zero holds
seq " is non-zero
proof end;

theorem :: COMSEQ_1:31
for seq being Complex_Sequence holds (seq " ) " = seq ;

theorem Th32: :: COMSEQ_1:32
for seq, seq1 being Complex_Sequence holds
( ( seq is non-zero & seq1 is non-zero ) iff seq (#) seq1 is non-zero )
proof end;

theorem Th33: :: COMSEQ_1:33
for seq, seq1 being Complex_Sequence holds (seq " ) (#) (seq1 " ) = (seq (#) seq1) "
proof end;

theorem :: COMSEQ_1:34
for seq, seq1 being Complex_Sequence st seq is non-zero holds
(seq1 /" seq) (#) seq = seq1
proof end;

theorem :: COMSEQ_1:35
for seq', seq, seq1', seq1 being Complex_Sequence holds (seq' /" seq) (#) (seq1' /" seq1) = (seq' (#) seq1') /" (seq (#) seq1)
proof end;

theorem :: COMSEQ_1:36
for seq, seq1 being Complex_Sequence st seq is non-zero & seq1 is non-zero holds
seq /" seq1 is non-zero
proof end;

theorem Th37: :: COMSEQ_1:37
for seq, seq1 being Complex_Sequence holds (seq /" seq1) " = seq1 /" seq
proof end;

theorem :: COMSEQ_1:38
for seq2, seq1, seq being Complex_Sequence holds seq2 (#) (seq1 /" seq) = (seq2 (#) seq1) /" seq
proof end;

theorem :: COMSEQ_1:39
for seq2, seq, seq1 being Complex_Sequence holds seq2 /" (seq /" seq1) = (seq2 (#) seq1) /" seq
proof end;

theorem Th40: :: COMSEQ_1:40
for seq1, seq2, seq being Complex_Sequence st seq1 is non-zero holds
seq2 /" seq = (seq2 (#) seq1) /" (seq (#) seq1)
proof end;

theorem Th41: :: COMSEQ_1:41
for r being Element of COMPLEX
for seq being Complex_Sequence st r <> 0c & seq is non-zero holds
r (#) seq is non-zero
proof end;

theorem :: COMSEQ_1:42
for seq being Complex_Sequence st seq is non-zero holds
- seq is non-zero
proof end;

theorem Th43: :: COMSEQ_1:43
for r being Element of COMPLEX
for seq being Complex_Sequence holds (r (#) seq) " = (r " ) (#) (seq " )
proof end;

theorem Th44: :: COMSEQ_1:44
for seq being Complex_Sequence st seq is non-zero holds
(- seq) " = (- 1r ) (#) (seq " )
proof end;

theorem :: COMSEQ_1:45
for seq, seq1 being Complex_Sequence st seq is non-zero holds
( - (seq1 /" seq) = (- seq1) /" seq & seq1 /" (- seq) = - (seq1 /" seq) )
proof end;

theorem :: COMSEQ_1:46
for seq1, seq, seq1' being Complex_Sequence holds
( (seq1 /" seq) + (seq1' /" seq) = (seq1 + seq1') /" seq & (seq1 /" seq) - (seq1' /" seq) = (seq1 - seq1') /" seq )
proof end;

theorem :: COMSEQ_1:47
for seq, seq', seq1, seq1' being Complex_Sequence st seq is non-zero & seq' is non-zero holds
( (seq1 /" seq) + (seq1' /" seq') = ((seq1 (#) seq') + (seq1' (#) seq)) /" (seq (#) seq') & (seq1 /" seq) - (seq1' /" seq') = ((seq1 (#) seq') - (seq1' (#) seq)) /" (seq (#) seq') )
proof end;

theorem :: COMSEQ_1:48
for seq1', seq, seq', seq1 being Complex_Sequence holds (seq1' /" seq) /" (seq' /" seq1) = (seq1' (#) seq1) /" (seq (#) seq')
proof end;

theorem Th49: :: COMSEQ_1:49
for seq, seq' being Complex_Sequence holds |.(seq (#) seq').| = |.seq.| (#) |.seq'.|
proof end;

theorem :: COMSEQ_1:50
for seq being Complex_Sequence st seq is non-zero holds
|.seq.| is non-zero
proof end;

theorem Th51: :: COMSEQ_1:51
for seq being Complex_Sequence holds |.seq.| " = |.(seq " ).|
proof end;

theorem :: COMSEQ_1:52
for seq', seq being Complex_Sequence holds |.(seq' /" seq).| = |.seq'.| /" |.seq.|
proof end;

theorem :: COMSEQ_1:53
for r being Element of COMPLEX
for seq being Complex_Sequence holds |.(r (#) seq).| = |.r.| (#) |.seq.|
proof end;