:: Concatenation of Finite Sequences Reducing Overlapping Part and an Argument of Separators of Sequential Files
:: by Hirofumi Fukura and Yatsuka Nakamura
::
:: Received March 18, 2004
:: Copyright (c) 2004-2021 Association of Mizar Users


theorem :: FINSEQ_8:1
for f, g being FinSequence st len f >= 1 holds
mid ((f ^ g),1,(len f)) = f
proof end;

theorem Th2: :: FINSEQ_8:2
for f being FinSequence
for i being Nat st i >= len f holds
f /^ i = {}
proof end;

theorem :: FINSEQ_8:3
for D being non empty set
for k1, k2 being Nat holds mid ((<*> D),k1,k2) = <*> D
proof end;

definition
let f be FinSequence;
let k1, k2 be Nat;
func smid (f,k1,k2) -> FinSequence equals :: FINSEQ_8:def 1
(f /^ (k1 -' 1)) | ((k2 + 1) -' k1);
correctness
coherence
(f /^ (k1 -' 1)) | ((k2 + 1) -' k1) is FinSequence
;
;
end;

:: deftheorem defines smid FINSEQ_8:def 1 :
for f being FinSequence
for k1, k2 being Nat holds smid (f,k1,k2) = (f /^ (k1 -' 1)) | ((k2 + 1) -' k1);

definition
let D be set ;
let f be FinSequence of D;
let k1, k2 be Nat;
:: original: smid
redefine func smid (f,k1,k2) -> FinSequence of D;
correctness
coherence
smid (f,k1,k2) is FinSequence of D
;
proof end;
end;

theorem Th4: :: FINSEQ_8:4
for f being FinSequence
for k1, k2 being Nat st k1 <= k2 holds
smid (f,k1,k2) = mid (f,k1,k2)
proof end;

theorem Th5: :: FINSEQ_8:5
for f being FinSequence
for k2 being Nat holds smid (f,1,k2) = f | k2
proof end;

theorem Th6: :: FINSEQ_8:6
for f being FinSequence
for k2 being Nat st len f <= k2 holds
smid (f,1,k2) = f
proof end;

theorem Th7: :: FINSEQ_8:7
for f being FinSequence
for k1, k2 being Nat st k1 > k2 holds
smid (f,k1,k2) = {}
proof end;

theorem :: FINSEQ_8:8
for f being FinSequence
for k2 being Nat holds smid (f,0,k2) = smid (f,1,(k2 + 1))
proof end;

theorem Th9: :: FINSEQ_8:9
for f, g being FinSequence holds smid ((f ^ g),((len f) + 1),((len f) + (len g))) = g
proof end;

:: Overlapping Part
definition
let D be non empty set ;
let f, g be FinSequence of D;
func ovlpart (f,g) -> FinSequence of D means :Def2: :: FINSEQ_8:def 2
( len it <= len g & it = smid (g,1,(len it)) & it = smid (f,(((len f) -' (len it)) + 1),(len f)) & ( for j being Nat st j <= len g & smid (g,1,j) = smid (f,(((len f) -' j) + 1),(len f)) holds
j <= len it ) );
existence
ex b1 being FinSequence of D st
( len b1 <= len g & b1 = smid (g,1,(len b1)) & b1 = smid (f,(((len f) -' (len b1)) + 1),(len f)) & ( for j being Nat st j <= len g & smid (g,1,j) = smid (f,(((len f) -' j) + 1),(len f)) holds
j <= len b1 ) )
proof end;
uniqueness
for b1, b2 being FinSequence of D st len b1 <= len g & b1 = smid (g,1,(len b1)) & b1 = smid (f,(((len f) -' (len b1)) + 1),(len f)) & ( for j being Nat st j <= len g & smid (g,1,j) = smid (f,(((len f) -' j) + 1),(len f)) holds
j <= len b1 ) & len b2 <= len g & b2 = smid (g,1,(len b2)) & b2 = smid (f,(((len f) -' (len b2)) + 1),(len f)) & ( for j being Nat st j <= len g & smid (g,1,j) = smid (f,(((len f) -' j) + 1),(len f)) holds
j <= len b2 ) holds
b1 = b2
proof end;
end;

:: deftheorem Def2 defines ovlpart FINSEQ_8:def 2 :
for D being non empty set
for f, g, b4 being FinSequence of D holds
( b4 = ovlpart (f,g) iff ( len b4 <= len g & b4 = smid (g,1,(len b4)) & b4 = smid (f,(((len f) -' (len b4)) + 1),(len f)) & ( for j being Nat st j <= len g & smid (g,1,j) = smid (f,(((len f) -' j) + 1),(len f)) holds
j <= len b4 ) ) );

theorem Th10: :: FINSEQ_8:10
for D being non empty set
for f, g being FinSequence of D holds len (ovlpart (f,g)) <= len f
proof end;

:: Concatenation Reducing Overlapping Part
definition
let D be non empty set ;
let f, g be FinSequence of D;
func ovlcon (f,g) -> FinSequence of D equals :: FINSEQ_8:def 3
f ^ (g /^ (len (ovlpart (f,g))));
coherence
f ^ (g /^ (len (ovlpart (f,g)))) is FinSequence of D
;
end;

:: deftheorem defines ovlcon FINSEQ_8:def 3 :
for D being non empty set
for f, g being FinSequence of D holds ovlcon (f,g) = f ^ (g /^ (len (ovlpart (f,g))));

theorem Th11: :: FINSEQ_8:11
for D being non empty set
for f, g being FinSequence of D holds ovlcon (f,g) = (f | ((len f) -' (len (ovlpart (f,g))))) ^ g
proof end;

::Overlapping Left Difference
definition
let D be non empty set ;
let f, g be FinSequence of D;
func ovlldiff (f,g) -> FinSequence of D equals :: FINSEQ_8:def 4
f | ((len f) -' (len (ovlpart (f,g))));
coherence
f | ((len f) -' (len (ovlpart (f,g)))) is FinSequence of D
;
end;

:: deftheorem defines ovlldiff FINSEQ_8:def 4 :
for D being non empty set
for f, g being FinSequence of D holds ovlldiff (f,g) = f | ((len f) -' (len (ovlpart (f,g))));

::Overlapping Right Difference
definition
let D be non empty set ;
let f, g be FinSequence of D;
func ovlrdiff (f,g) -> FinSequence of D equals :: FINSEQ_8:def 5
g /^ (len (ovlpart (f,g)));
coherence
g /^ (len (ovlpart (f,g))) is FinSequence of D
;
end;

:: deftheorem defines ovlrdiff FINSEQ_8:def 5 :
for D being non empty set
for f, g being FinSequence of D holds ovlrdiff (f,g) = g /^ (len (ovlpart (f,g)));

theorem :: FINSEQ_8:12
for D being non empty set
for f, g being FinSequence of D holds
( ovlcon (f,g) = ((ovlldiff (f,g)) ^ (ovlpart (f,g))) ^ (ovlrdiff (f,g)) & ovlcon (f,g) = (ovlldiff (f,g)) ^ ((ovlpart (f,g)) ^ (ovlrdiff (f,g))) )
proof end;

theorem :: FINSEQ_8:13
for D being non empty set
for f being FinSequence of D holds
( ovlcon (f,f) = f & ovlpart (f,f) = f & ovlldiff (f,f) = {} & ovlrdiff (f,f) = {} )
proof end;

theorem :: FINSEQ_8:14
for D being non empty set
for f, g being FinSequence of D holds
( ovlpart ((f ^ g),g) = g & ovlpart (f,(f ^ g)) = f )
proof end;

theorem Th15: :: FINSEQ_8:15
for D being non empty set
for f, g being FinSequence of D holds
( len (ovlcon (f,g)) = ((len f) + (len g)) - (len (ovlpart (f,g))) & len (ovlcon (f,g)) = ((len f) + (len g)) -' (len (ovlpart (f,g))) & len (ovlcon (f,g)) = (len f) + ((len g) -' (len (ovlpart (f,g)))) )
proof end;

theorem Th16: :: FINSEQ_8:16
for D being non empty set
for f, g being FinSequence of D holds
( len (ovlpart (f,g)) <= len f & len (ovlpart (f,g)) <= len g )
proof end;

definition
let D be non empty set ;
let CR be FinSequence of D;
pred CR separates_uniquely means :: FINSEQ_8:def 6
for f being FinSequence of D
for i, j being Element of NAT st 1 <= i & i < j & (j + (len CR)) -' 1 <= len f & smid (f,i,((i + (len CR)) -' 1)) = smid (f,j,((j + (len CR)) -' 1)) & smid (f,i,((i + (len CR)) -' 1)) = CR holds
j -' i >= len CR;
end;

:: deftheorem defines separates_uniquely FINSEQ_8:def 6 :
for D being non empty set
for CR being FinSequence of D holds
( CR separates_uniquely iff for f being FinSequence of D
for i, j being Element of NAT st 1 <= i & i < j & (j + (len CR)) -' 1 <= len f & smid (f,i,((i + (len CR)) -' 1)) = smid (f,j,((j + (len CR)) -' 1)) & smid (f,i,((i + (len CR)) -' 1)) = CR holds
j -' i >= len CR );

theorem :: FINSEQ_8:17
for D being non empty set
for CR being FinSequence of D holds
( CR separates_uniquely iff len (ovlpart ((CR /^ 1),CR)) = 0 )
proof end;

definition
let f, g be FinSequence;
let n be Nat;
pred g is_substring_of f,n means :: FINSEQ_8:def 7
( len g > 0 implies ex i being Nat st
( n <= i & i <= len f & mid (f,i,((i -' 1) + (len g))) = g ) );
end;

:: deftheorem defines is_substring_of FINSEQ_8:def 7 :
for f, g being FinSequence
for n being Nat holds
( g is_substring_of f,n iff ( len g > 0 implies ex i being Nat st
( n <= i & i <= len f & mid (f,i,((i -' 1) + (len g))) = g ) ) );

theorem :: FINSEQ_8:18
for f, g being FinSequence
for n, m being Nat st m >= n & g is_substring_of f,m holds
g is_substring_of f,n
proof end;

theorem Th19: :: FINSEQ_8:19
for f being FinSequence st 1 <= len f holds
f is_substring_of f,1
proof end;

theorem Th20: :: FINSEQ_8:20
for f, g being FinSequence st g is_substring_of f, 0 holds
g is_substring_of f,1
proof end;

notation
let g, f be FinSequence;
synonym g is_preposition_of f for g c= f;
end;

definition
let g, f be FinSequence;
redefine pred g c= f means :: FINSEQ_8:def 8
( len g > 0 implies ( 1 <= len f & mid (f,1,(len g)) = g ) );
compatibility
( g is_preposition_of f iff ( len g > 0 implies ( 1 <= len f & mid (f,1,(len g)) = g ) ) )
proof end;
end;

:: deftheorem defines is_preposition_of FINSEQ_8:def 8 :
for g, f being FinSequence holds
( g is_preposition_of f iff ( len g > 0 implies ( 1 <= len f & mid (f,1,(len g)) = g ) ) );

theorem :: FINSEQ_8:21
for f, g being FinSequence st len g > 0 & g is_preposition_of f holds
g . 1 = f . 1
proof end;

definition
let f, g be FinSequence;
pred g is_postposition_of f means :: FINSEQ_8:def 9
Rev g is_preposition_of Rev f;
correctness
;
end;

:: deftheorem defines is_postposition_of FINSEQ_8:def 9 :
for f, g being FinSequence holds
( g is_postposition_of f iff Rev g is_preposition_of Rev f );

theorem Th22: :: FINSEQ_8:22
for f, g being FinSequence st len g = 0 holds
g is_postposition_of f
proof end;

theorem Th23: :: FINSEQ_8:23
for D being non empty set
for f, g being FinSequence of D st g is_postposition_of f holds
len g <= len f
proof end;

theorem :: FINSEQ_8:24
for D being non empty set
for f, g being FinSequence of D st g is_postposition_of f & len g > 0 holds
( len g <= len f & mid (f,(((len f) + 1) -' (len g)),(len f)) = g )
proof end;

theorem :: FINSEQ_8:25
for D being non empty set
for f, g being FinSequence of D st ( len g > 0 implies ( len g <= len f & mid (f,(((len f) + 1) -' (len g)),(len f)) = g ) ) holds
g is_postposition_of f
proof end;

theorem :: FINSEQ_8:26
for f, g being FinSequence st 1 <= len f & g is_preposition_of f holds
g is_substring_of f,1
proof end;

theorem Th27: :: FINSEQ_8:27
for D being non empty set
for f, g being FinSequence of D
for n being Nat st not g is_substring_of f,n holds
for i being Element of NAT st n <= i & 0 < i holds
mid (f,i,((i -' 1) + (len g))) <> g
proof end;

definition
let D be non empty set ;
let f, g be FinSequence of D;
let n be Nat;
func instr (n,f,g) -> Element of NAT means :Def10: :: FINSEQ_8:def 10
( ( it <> 0 implies ( n <= it & g is_preposition_of f /^ (it -' 1) & ( for j being Element of NAT st j >= n & j > 0 & g is_preposition_of f /^ (j -' 1) holds
j >= it ) ) ) & ( it = 0 implies not g is_substring_of f,n ) );
existence
ex b1 being Element of NAT st
( ( b1 <> 0 implies ( n <= b1 & g is_preposition_of f /^ (b1 -' 1) & ( for j being Element of NAT st j >= n & j > 0 & g is_preposition_of f /^ (j -' 1) holds
j >= b1 ) ) ) & ( b1 = 0 implies not g is_substring_of f,n ) )
proof end;
uniqueness
for b1, b2 being Element of NAT st ( b1 <> 0 implies ( n <= b1 & g is_preposition_of f /^ (b1 -' 1) & ( for j being Element of NAT st j >= n & j > 0 & g is_preposition_of f /^ (j -' 1) holds
j >= b1 ) ) ) & ( b1 = 0 implies not g is_substring_of f,n ) & ( b2 <> 0 implies ( n <= b2 & g is_preposition_of f /^ (b2 -' 1) & ( for j being Element of NAT st j >= n & j > 0 & g is_preposition_of f /^ (j -' 1) holds
j >= b2 ) ) ) & ( b2 = 0 implies not g is_substring_of f,n ) holds
b1 = b2
proof end;
end;

:: deftheorem Def10 defines instr FINSEQ_8:def 10 :
for D being non empty set
for f, g being FinSequence of D
for n being Nat
for b5 being Element of NAT holds
( b5 = instr (n,f,g) iff ( ( b5 <> 0 implies ( n <= b5 & g is_preposition_of f /^ (b5 -' 1) & ( for j being Element of NAT st j >= n & j > 0 & g is_preposition_of f /^ (j -' 1) holds
j >= b5 ) ) ) & ( b5 = 0 implies not g is_substring_of f,n ) ) );

definition
let D be non empty set ;
let f, CR be FinSequence of D;
func addcr (f,CR) -> FinSequence of D equals :: FINSEQ_8:def 11
ovlcon (f,CR);
correctness
coherence
ovlcon (f,CR) is FinSequence of D
;
;
end;

:: deftheorem defines addcr FINSEQ_8:def 11 :
for D being non empty set
for f, CR being FinSequence of D holds addcr (f,CR) = ovlcon (f,CR);

definition
let D be non empty set ;
let r, CR be FinSequence of D;
pred r is_terminated_by CR means :: FINSEQ_8:def 12
( len CR > 0 implies ( len r >= len CR & instr (1,r,CR) = ((len r) + 1) -' (len CR) ) );
correctness
;
end;

:: deftheorem defines is_terminated_by FINSEQ_8:def 12 :
for D being non empty set
for r, CR being FinSequence of D holds
( r is_terminated_by CR iff ( len CR > 0 implies ( len r >= len CR & instr (1,r,CR) = ((len r) + 1) -' (len CR) ) ) );

theorem :: FINSEQ_8:28
for D being non empty set
for f being FinSequence of D holds f is_terminated_by f
proof end;

theorem :: FINSEQ_8:29
for f being FinSequence
for k1, k2 being Nat st k1 in dom f & k2 in dom f holds
smid (f,k1,k2) = (k1,k2) -cut f
proof end;