:: RFINSEQ semantic presentation

definition
let F be Relation, G be Relation;
pred c1,c2 are_fiberwise_equipotent means :Def1: :: RFINSEQ:def 1
for x being set holds Card (F " {x}) = Card (G " {x});
reflexivity
for F being Relation
for x being set holds Card (F " {x}) = Card (F " {x})
;
symmetry
for F, G being Relation st ( for x being set holds Card (F " {x}) = Card (G " {x}) ) holds
for x being set holds Card (G " {x}) = Card (F " {x})
;
end;

:: deftheorem Def1 defines are_fiberwise_equipotent RFINSEQ:def 1 :
for F, G being Relation holds
( F,G are_fiberwise_equipotent iff for x being set holds Card (F " {x}) = Card (G " {x}) );

Lemma30: for F being Function
for x being set st not x in rng F holds
F " {x} = {}
proof end;

theorem Th1: :: RFINSEQ:1
for F, G being Function st F,G are_fiberwise_equipotent holds
rng F = rng G
proof end;

theorem Th2: :: RFINSEQ:2
for F, G, H being Function st F,G are_fiberwise_equipotent & F,H are_fiberwise_equipotent holds
G,H are_fiberwise_equipotent
proof end;

theorem Th3: :: RFINSEQ:3
for F, G being Function holds
( F,G are_fiberwise_equipotent iff ex H being Function st
( dom H = dom F & rng H = dom G & H is one-to-one & F = G * H ) )
proof end;

theorem Th4: :: RFINSEQ:4
for F, G being Function holds
( F,G are_fiberwise_equipotent iff for X being set holds Card (F " X) = Card (G " X) )
proof end;

theorem Th5: :: RFINSEQ:5
for D being non empty set
for F, G being Function st rng F c= D & rng G c= D holds
( F,G are_fiberwise_equipotent iff for d being Element of D holds Card (F " {d}) = Card (G " {d}) )
proof end;

theorem Th6: :: RFINSEQ:6
for F, G being Function st dom F = dom G holds
( F,G are_fiberwise_equipotent iff ex P being Permutation of dom F st F = G * P )
proof end;

theorem Th7: :: RFINSEQ:7
for F, G being Function st F,G are_fiberwise_equipotent holds
Card (dom F) = Card (dom G)
proof end;

theorem Th8: :: RFINSEQ:8
canceled;

theorem Th9: :: RFINSEQ:9
for F, G being finite Function holds
( F,G are_fiberwise_equipotent iff for X being set holds card (F " X) = card (G " X) ) by ;

theorem Th10: :: RFINSEQ:10
for F, G being finite Function st F,G are_fiberwise_equipotent holds
card (dom F) = card (dom G)
proof end;

theorem Th11: :: RFINSEQ:11
for D being non empty set
for F, G being finite Function st rng F c= D & rng G c= D holds
( F,G are_fiberwise_equipotent iff for d being Element of D holds card (F " {d}) = card (G " {d}) ) by ;

theorem Th12: :: RFINSEQ:12
canceled;

theorem Th13: :: RFINSEQ:13
for f, g being FinSequence holds
( f,g are_fiberwise_equipotent iff for X being set holds card (f " X) = card (g " X) ) by ;

theorem Th14: :: RFINSEQ:14
for f, g, h being FinSequence holds
( f,g are_fiberwise_equipotent iff f ^ h,g ^ h are_fiberwise_equipotent )
proof end;

theorem Th15: :: RFINSEQ:15
for f, g being FinSequence holds f ^ g,g ^ f are_fiberwise_equipotent
proof end;

theorem Th16: :: RFINSEQ:16
for f, g being FinSequence st f,g are_fiberwise_equipotent holds
( len f = len g & dom f = dom g )
proof end;

theorem Th17: :: RFINSEQ:17
for f, g being FinSequence holds
( f,g are_fiberwise_equipotent iff ex P being Permutation of dom g st f = g * P )
proof end;

registration
let F be Function;
let X be finite set ;
cluster a1 | a2 -> Function-like finite ;
coherence
( F | X is finite & F | X is Function-like )
proof end;
end;

defpred S1[ Element of NAT ] means for X being finite set
for F being Function st card (dom (F | X)) = a1 holds
ex a being FinSequence st F | X,a are_fiberwise_equipotent ;

Lemma93: S1[0]
proof end;

Lemma95: for n being Element of NAT st S1[n] holds
S1[n + 1]
proof end;

theorem Th18: :: RFINSEQ:18
for F being Function
for X being finite set ex f being FinSequence st F | X,f are_fiberwise_equipotent
proof end;

definition
let D be set ;
let f be FinSequence of D;
let n be Nat;
func c2 /^ c3 -> FinSequence of a1 means :Def2: :: RFINSEQ:def 2
( len it = (len f) - n & ( for m being Nat st m in dom it holds
it . m = f . (m + n) ) ) if n <= len f
otherwise it = <*> D;
existence
( ( n <= len f implies ex b1 being FinSequence of D st
( len b1 = (len f) - n & ( for m being Nat st m in dom b1 holds
b1 . m = f . (m + n) ) ) ) & ( not n <= len f implies ex b1 being FinSequence of D st b1 = <*> D ) )
proof end;
uniqueness
for b1, b2 being FinSequence of D holds
( ( n <= len f & len b1 = (len f) - n & ( for m being Nat st m in dom b1 holds
b1 . m = f . (m + n) ) & len b2 = (len f) - n & ( for m being Nat st m in dom b2 holds
b2 . m = f . (m + n) ) implies b1 = b2 ) & ( not n <= len f & b1 = <*> D & b2 = <*> D implies b1 = b2 ) )
proof end;
consistency
for b1 being FinSequence of D holds verum
;
end;

:: deftheorem Def2 defines /^ RFINSEQ:def 2 :
for D being set
for f being FinSequence of D
for n being Nat
for b4 being FinSequence of D holds
( ( n <= len f implies ( b4 = f /^ n iff ( len b4 = (len f) - n & ( for m being Nat st m in dom b4 holds
b4 . m = f . (m + n) ) ) ) ) & ( not n <= len f implies ( b4 = f /^ n iff b4 = <*> D ) ) );

Lemma106: for n being Element of NAT
for D being non empty set
for f being FinSequence of D st len f <= n holds
f | n = f
proof end;

theorem Th19: :: RFINSEQ:19
for D being non empty set
for f being FinSequence of D
for n, m being Element of NAT st n in dom f & m in Seg n holds
( (f | n) . m = f . m & m in dom f )
proof end;

theorem Th20: :: RFINSEQ:20
for D being non empty set
for f being FinSequence of D
for n being Element of NAT
for x being set st len f = n + 1 & x = f . (n + 1) holds
f = (f | n) ^ <*x*>
proof end;

theorem Th21: :: RFINSEQ:21
for D being non empty set
for f being FinSequence of D
for n being Element of NAT holds (f | n) ^ (f /^ n) = f
proof end;

theorem Th22: :: RFINSEQ:22
for R1, R2 being FinSequence of REAL st R1,R2 are_fiberwise_equipotent holds
Sum R1 = Sum R2
proof end;

definition
let R be FinSequence of REAL ;
func MIM c1 -> FinSequence of REAL means :Def3: :: RFINSEQ:def 3
( len it = len R & it . (len it) = R . (len R) & ( for n being Element of NAT st 1 <= n & n <= (len it) - 1 holds
it . n = (R . n) - (R . (n + 1)) ) );
existence
ex b1 being FinSequence of REAL st
( len b1 = len R & b1 . (len b1) = R . (len R) & ( for n being Element of NAT st 1 <= n & n <= (len b1) - 1 holds
b1 . n = (R . n) - (R . (n + 1)) ) )
proof end;
uniqueness
for b1, b2 being FinSequence of REAL st len b1 = len R & b1 . (len b1) = R . (len R) & ( for n being Element of NAT st 1 <= n & n <= (len b1) - 1 holds
b1 . n = (R . n) - (R . (n + 1)) ) & len b2 = len R & b2 . (len b2) = R . (len R) & ( for n being Element of NAT st 1 <= n & n <= (len b2) - 1 holds
b2 . n = (R . n) - (R . (n + 1)) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def3 defines MIM RFINSEQ:def 3 :
for R, b2 being FinSequence of REAL holds
( b2 = MIM R iff ( len b2 = len R & b2 . (len b2) = R . (len R) & ( for n being Element of NAT st 1 <= n & n <= (len b2) - 1 holds
b2 . n = (R . n) - (R . (n + 1)) ) ) );

theorem Th23: :: RFINSEQ:23
for R being FinSequence of REAL
for r being Real
for n being Element of NAT st len R = n + 2 & R . (n + 1) = r holds
MIM (R | (n + 1)) = ((MIM R) | n) ^ <*r*>
proof end;

theorem Th24: :: RFINSEQ:24
for R being FinSequence of REAL
for r, s being Real
for n being Element of NAT st len R = n + 2 & R . (n + 1) = r & R . (n + 2) = s holds
MIM R = ((MIM R) | n) ^ <*(r - s),s*>
proof end;

theorem Th25: :: RFINSEQ:25
MIM (<*> REAL ) = <*> REAL
proof end;

theorem Th26: :: RFINSEQ:26
for r being Real holds MIM <*r*> = <*r*>
proof end;

theorem Th27: :: RFINSEQ:27
for r, s being Real holds MIM <*r,s*> = <*(r - s),s*>
proof end;

theorem Th28: :: RFINSEQ:28
for R being FinSequence of REAL
for n being Element of NAT holds (MIM R) /^ n = MIM (R /^ n)
proof end;

theorem Th29: :: RFINSEQ:29
for R being FinSequence of REAL st len R <> 0 holds
Sum (MIM R) = R . 1
proof end;

theorem Th30: :: RFINSEQ:30
for R being FinSequence of REAL
for n being Element of NAT st 1 <= n & n < len R holds
Sum (MIM (R /^ n)) = R . (n + 1)
proof end;

definition
let IT be FinSequence of REAL ;
attr a1 is non-increasing means :Def4: :: RFINSEQ:def 4
for n being Element of NAT st n in dom IT & n + 1 in dom IT holds
IT . n >= IT . (n + 1);
end;

:: deftheorem Def4 defines non-increasing RFINSEQ:def 4 :
for IT being FinSequence of REAL holds
( IT is non-increasing iff for n being Element of NAT st n in dom IT & n + 1 in dom IT holds
IT . n >= IT . (n + 1) );

registration
cluster non-increasing FinSequence of REAL ;
existence
ex b1 being FinSequence of REAL st b1 is non-increasing
proof end;
end;

theorem Th31: :: RFINSEQ:31
for R being FinSequence of REAL st ( len R = 0 or len R = 1 ) holds
R is non-increasing
proof end;

theorem Th32: :: RFINSEQ:32
for R being FinSequence of REAL holds
( R is non-increasing iff for n, m being Element of NAT st n in dom R & m in dom R & n < m holds
R . n >= R . m )
proof end;

theorem Th33: :: RFINSEQ:33
for R being non-increasing FinSequence of REAL
for n being Element of NAT holds R | n is non-increasing FinSequence of REAL
proof end;

theorem Th34: :: RFINSEQ:34
for R being non-increasing FinSequence of REAL
for n being Element of NAT holds R /^ n is non-increasing FinSequence of REAL
proof end;

Lemma135: for f, g being non-increasing FinSequence of REAL
for n being Element of NAT st len f = n + 1 & len f = len g & f,g are_fiberwise_equipotent holds
( f . (len f) = g . (len g) & f | n,g | n are_fiberwise_equipotent )
proof end;

defpred S2[ Element of NAT ] means for R being FinSequence of REAL st a1 = len R holds
ex b being non-increasing FinSequence of REAL st R,b are_fiberwise_equipotent ;

Lemma137: S2[0]
proof end;

Lemma138: for n being Element of NAT st S2[n] holds
S2[n + 1]
proof end;

theorem Th35: :: RFINSEQ:35
for R being FinSequence of REAL ex R1 being non-increasing FinSequence of REAL st R,R1 are_fiberwise_equipotent
proof end;

Lemma163: for n being Element of NAT
for g1, g2 being non-increasing FinSequence of REAL st n = len g1 & g1,g2 are_fiberwise_equipotent holds
g1 = g2
proof end;

theorem Th36: :: RFINSEQ:36
for R1, R2 being non-increasing FinSequence of REAL st R1,R2 are_fiberwise_equipotent holds
R1 = R2
proof end;

theorem Th37: :: RFINSEQ:37
for R being FinSequence of REAL
for r, s being Real st r <> 0 holds
R " {(s / r)} = (r * R) " {s}
proof end;

theorem Th38: :: RFINSEQ:38
for R being FinSequence of REAL holds (0 * R) " {0} = dom R
proof end;

theorem Th39: :: RFINSEQ:39
for f, g being Function st rng f = rng g & f is one-to-one & g is one-to-one holds
f,g are_fiberwise_equipotent
proof end;