:: FINSEQ_1 semantic presentation

definition
let c1 be natural number ;
func Seg c1 -> set equals :: FINSEQ_1:def 1
{ b1 where B is Nat : ( 1 <= b1 & b1 <= a1 ) } ;
correctness
coherence
{ b1 where B is Nat : ( 1 <= b1 & b1 <= c1 ) } is set
;
;
end;

:: deftheorem Def1 defines Seg FINSEQ_1:def 1 :
for b1 being natural number holds Seg b1 = { b2 where B is Nat : ( 1 <= b2 & b2 <= b1 ) } ;

definition
let c1 be natural number ;
redefine func Seg as Seg c1 -> Subset of NAT ;
coherence
Seg c1 is Subset of NAT
proof end;
end;

theorem Th1: :: FINSEQ_1:1
canceled;

theorem Th2: :: FINSEQ_1:2
canceled;

theorem Th3: :: FINSEQ_1:3
for b1, b2 being natural number holds
( b1 in Seg b2 iff ( 1 <= b1 & b1 <= b2 ) )
proof end;

theorem Th4: :: FINSEQ_1:4
( Seg 0 = {} & Seg 1 = {1} & Seg 2 = {1,2} )
proof end;

theorem Th5: :: FINSEQ_1:5
for b1 being natural number holds
( b1 = 0 or b1 in Seg b1 )
proof end;

theorem Th6: :: FINSEQ_1:6
for b1 being natural number holds b1 + 1 in Seg (b1 + 1) by Th5;

theorem Th7: :: FINSEQ_1:7
for b1, b2 being natural number holds
( b1 <= b2 iff Seg b1 c= Seg b2 )
proof end;

theorem Th8: :: FINSEQ_1:8
for b1, b2 being natural number holds
( Seg b1 = Seg b2 implies b1 = b2 )
proof end;

theorem Th9: :: FINSEQ_1:9
for b1, b2 being natural number holds
( b1 <= b2 implies Seg b1 = (Seg b2) /\ (Seg b1) )
proof end;

theorem Th10: :: FINSEQ_1:10
for b1, b2 being natural number holds
( Seg b1 = (Seg b1) /\ (Seg b2) implies b1 <= b2 )
proof end;

theorem Th11: :: FINSEQ_1:11
for b1 being natural number holds (Seg b1) \/ {(b1 + 1)} = Seg (b1 + 1)
proof end;

theorem Th12: :: FINSEQ_1:12
for b1 being natural number holds Seg b1 = (Seg (b1 + 1)) \ {(b1 + 1)}
proof end;

definition
let c1 be Relation;
attr a1 is FinSequence-like means :Def2: :: FINSEQ_1:def 2
ex b1 being Nat st dom a1 = Seg b1;
end;

:: deftheorem Def2 defines FinSequence-like FINSEQ_1:def 2 :
for b1 being Relation holds
( b1 is FinSequence-like iff ex b2 being Nat st dom b1 = Seg b2 );

registration
cluster FinSequence-like set ;
existence
ex b1 being Function st b1 is FinSequence-like
proof end;
end;

definition
mode FinSequence is FinSequence-like Function;
end;

defpred S1[ set , set ] means ex b1 being Nat st
( a1 = b1 & a2 = b1 + 1 );

registration
let c1 be natural number ;
cluster Seg a1 -> finite ;
coherence
Seg c1 is finite
proof end;
end;

registration
cluster FinSequence-like -> finite set ;
coherence
for b1 being Function holds
( b1 is FinSequence-like implies b1 is finite )
proof end;
end;

Lemma9: for b1 being Nat holds Seg b1,b1 are_equipotent
proof end;

Lemma10: for b1 being Nat holds Card (Seg b1) = Card b1
proof end;

notation
let c1 be FinSequence;
synonym len c1 for Card c1;
end;

definition
let c1 be FinSequence;
redefine func len as len c1 -> Nat means :Def3: :: FINSEQ_1:def 3
Seg a2 = dom a1;
coherence
len c1 is Nat
proof end;
compatibility
for b1 being Nat holds
( b1 = len c1 iff Seg b1 = dom c1 )
proof end;
end;

:: deftheorem Def3 defines len FINSEQ_1:def 3 :
for b1 being FinSequence
for b2 being Nat holds
( b2 = len b1 iff Seg b2 = dom b1 );

definition
let c1 be FinSequence;
redefine func dom as dom c1 -> Subset of NAT ;
coherence
dom c1 is Subset of NAT
proof end;
end;

theorem Th13: :: FINSEQ_1:13
canceled;

theorem Th14: :: FINSEQ_1:14
{} is FinSequence
proof end;

theorem Th15: :: FINSEQ_1:15
for b1 being Function holds
not ( ex b2 being Nat st dom b1 c= Seg b2 & ( for b2 being FinSequence holds
not b1 c= b2 ) )
proof end;

scheme :: FINSEQ_1:sch 1
s1{ F1() -> Nat, P1[ set , set ] } :
ex b1 being FinSequence st
( dom b1 = Seg F1() & ( for b2 being Nat holds
( b2 in Seg F1() implies P1[b2,b1 . b2] ) ) )
provided
E13: for b1 being Nat
for b2, b3 being set holds
( b1 in Seg F1() & P1[b1,b2] & P1[b1,b3] implies b2 = b3 ) and E14: for b1 being Nat holds
not ( b1 in Seg F1() & ( for b2 being set holds
not P1[b1,b2] ) )
proof end;

scheme :: FINSEQ_1:sch 2
s2{ F1() -> Nat, F2( set ) -> set } :
ex b1 being FinSequence st
( len b1 = F1() & ( for b2 being Nat holds
( b2 in Seg F1() implies b1 . b2 = F2(b2) ) ) )
proof end;

theorem Th16: :: FINSEQ_1:16
for b1 being set
for b2 being FinSequence holds
not ( b1 in b2 & ( for b3 being Nat holds
not ( b3 in dom b2 & b1 = [b3,(b2 . b3)] ) ) )
proof end;

theorem Th17: :: FINSEQ_1:17
for b1, b2 being FinSequence holds
( dom b1 = dom b2 & ( for b3 being Nat holds
( b3 in dom b1 implies b1 . b3 = b2 . b3 ) ) implies b1 = b2 )
proof end;

theorem Th18: :: FINSEQ_1:18
for b1, b2 being FinSequence holds
( len b1 = len b2 & ( for b3 being Nat holds
( 1 <= b3 & b3 <= len b1 implies b1 . b3 = b2 . b3 ) ) implies b1 = b2 )
proof end;

theorem Th19: :: FINSEQ_1:19
for b1 being natural number
for b2 being FinSequence holds
b2 | (Seg b1) is FinSequence
proof end;

theorem Th20: :: FINSEQ_1:20
for b1 being Function
for b2 being FinSequence holds
( rng b2 c= dom b1 implies b1 * b2 is FinSequence )
proof end;

theorem Th21: :: FINSEQ_1:21
for b1 being natural number
for b2, b3 being FinSequence holds
( b1 <= len b2 & b3 = b2 | (Seg b1) implies ( len b3 = b1 & dom b3 = Seg b1 ) )
proof end;

definition
let c1 be set ;
mode FinSequence of c1 -> FinSequence means :Def4: :: FINSEQ_1:def 4
rng a2 c= a1;
existence
ex b1 being FinSequence st rng b1 c= c1
proof end;
end;

:: deftheorem Def4 defines FinSequence FINSEQ_1:def 4 :
for b1 being set
for b2 being FinSequence holds
( b2 is FinSequence of b1 iff rng b2 c= b1 );

Lemma17: for b1 being set
for b2 being FinSequence of b1 holds
b2 is PartFunc of NAT ,b1
proof end;

registration
cluster {} -> finite FinSequence-like ;
coherence
{} is FinSequence-like
by Th14;
end;

registration
let c1 be set ;
cluster finite FinSequence-like M5( NAT ,a1);
existence
ex b1 being PartFunc of NAT ,c1 st b1 is FinSequence-like
proof end;
end;

definition
let c1 be set ;
redefine mode FinSequence as FinSequence of c1 -> FinSequence-like PartFunc of NAT ,a1;
coherence
for b1 being FinSequence of c1 holds
b1 is FinSequence-like PartFunc of NAT ,c1
by Lemma17;
end;

theorem Th22: :: FINSEQ_1:22
canceled;

theorem Th23: :: FINSEQ_1:23
for b1 being natural number
for b2 being set
for b3 being FinSequence of b2 holds
b3 | (Seg b1) is FinSequence of b2
proof end;

theorem Th24: :: FINSEQ_1:24
for b1 being natural number
for b2 being non empty set holds
ex b3 being FinSequence of b2 st len b3 = b1
proof end;

Lemma19: for b1 being FinSequence holds
( b1 = {} iff len b1 = 0 )
proof end;

registration
cluster empty finite set ;
existence
ex b1 being FinSequence st b1 is empty
by Th14;
end;

theorem Th25: :: FINSEQ_1:25
for b1 being FinSequence holds
( len b1 = 0 iff b1 = {} ) by Lemma19;

theorem Th26: :: FINSEQ_1:26
for b1 being FinSequence holds
( b1 = {} iff dom b1 = {} ) by RELAT_1:60, RELAT_1:64;

theorem Th27: :: FINSEQ_1:27
for b1 being FinSequence holds
( b1 = {} iff rng b1 = {} ) by RELAT_1:60, RELAT_1:64;

theorem Th28: :: FINSEQ_1:28
canceled;

theorem Th29: :: FINSEQ_1:29
for b1 being set holds
{} is FinSequence of b1
proof end;

registration
let c1 be set ;
cluster empty finite FinSequence-like FinSequence of a1;
existence
ex b1 being FinSequence of c1 st b1 is empty
proof end;
end;

definition
let c1 be set ;
func <*c1*> -> set equals :: FINSEQ_1:def 5
{[1,a1]};
coherence
{[1,c1]} is set
;
end;

:: deftheorem Def5 defines <* FINSEQ_1:def 5 :
for b1 being set holds <*b1*> = {[1,b1]};

definition
let c1 be set ;
func <*> c1 -> empty FinSequence of a1 equals :: FINSEQ_1:def 6
{} ;
coherence
{} is empty FinSequence of c1
by Th29;
end;

:: deftheorem Def6 defines <*> FINSEQ_1:def 6 :
for b1 being set holds <*> b1 = {} ;

theorem Th30: :: FINSEQ_1:30
canceled;

theorem Th31: :: FINSEQ_1:31
canceled;

theorem Th32: :: FINSEQ_1:32
for b1 being FinSequence
for b2 being set holds
( b1 = <*> b2 iff len b1 = 0 ) by Lemma19;

definition
let c1, c2 be FinSequence;
func c1 ^ c2 -> FinSequence means :Def7: :: FINSEQ_1:def 7
( dom a3 = Seg ((len a1) + (len a2)) & ( for b1 being Nat holds
( b1 in dom a1 implies a3 . b1 = a1 . b1 ) ) & ( for b1 being Nat holds
( b1 in dom a2 implies a3 . ((len a1) + b1) = a2 . b1 ) ) );
existence
ex b1 being FinSequence st
( dom b1 = Seg ((len c1) + (len c2)) & ( for b2 being Nat holds
( b2 in dom c1 implies b1 . b2 = c1 . b2 ) ) & ( for b2 being Nat holds
( b2 in dom c2 implies b1 . ((len c1) + b2) = c2 . b2 ) ) )
proof end;
uniqueness
for b1, b2 being FinSequence holds
( dom b1 = Seg ((len c1) + (len c2)) & ( for b3 being Nat holds
( b3 in dom c1 implies b1 . b3 = c1 . b3 ) ) & ( for b3 being Nat holds
( b3 in dom c2 implies b1 . ((len c1) + b3) = c2 . b3 ) ) & dom b2 = Seg ((len c1) + (len c2)) & ( for b3 being Nat holds
( b3 in dom c1 implies b2 . b3 = c1 . b3 ) ) & ( for b3 being Nat holds
( b3 in dom c2 implies b2 . ((len c1) + b3) = c2 . b3 ) ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def7 defines ^ FINSEQ_1:def 7 :
for b1, b2 being FinSequence
for b3 being FinSequence holds
( b3 = b1 ^ b2 iff ( dom b3 = Seg ((len b1) + (len b2)) & ( for b4 being Nat holds
( b4 in dom b1 implies b3 . b4 = b1 . b4 ) ) & ( for b4 being Nat holds
( b4 in dom b2 implies b3 . ((len b1) + b4) = b2 . b4 ) ) ) );

theorem Th33: :: FINSEQ_1:33
for b1, b2 being FinSequence holds b1 = (b1 ^ b2) | (dom b1)
proof end;

theorem Th34: :: FINSEQ_1:34
canceled;

theorem Th35: :: FINSEQ_1:35
for b1, b2 being FinSequence holds len (b1 ^ b2) = (len b1) + (len b2)
proof end;

theorem Th36: :: FINSEQ_1:36
for b1, b2 being FinSequence
for b3 being natural number holds
( (len b1) + 1 <= b3 & b3 <= (len b1) + (len b2) implies (b1 ^ b2) . b3 = b2 . (b3 - (len b1)) )
proof end;

theorem Th37: :: FINSEQ_1:37
for b1, b2 being FinSequence
for b3 being natural number holds
( len b1 < b3 & b3 <= len (b1 ^ b2) implies (b1 ^ b2) . b3 = b2 . (b3 - (len b1)) )
proof end;

theorem Th38: :: FINSEQ_1:38
for b1, b2 being FinSequence
for b3 being Nat holds
not ( b3 in dom (b1 ^ b2) & not b3 in dom b1 & ( for b4 being Nat holds
not ( b4 in dom b2 & b3 = (len b1) + b4 ) ) )
proof end;

theorem Th39: :: FINSEQ_1:39
for b1, b2 being FinSequence holds dom b1 c= dom (b1 ^ b2)
proof end;

theorem Th40: :: FINSEQ_1:40
for b1 being set
for b2, b3 being FinSequence holds
not ( b1 in dom b2 & ( for b4 being Nat holds
not ( b4 = b1 & (len b3) + b4 in dom (b3 ^ b2) ) ) )
proof end;

theorem Th41: :: FINSEQ_1:41
for b1 being Nat
for b2, b3 being FinSequence holds
( b1 in dom b2 implies (len b3) + b1 in dom (b3 ^ b2) )
proof end;

theorem Th42: :: FINSEQ_1:42
for b1, b2 being FinSequence holds rng b1 c= rng (b1 ^ b2)
proof end;

theorem Th43: :: FINSEQ_1:43
for b1, b2 being FinSequence holds rng b1 c= rng (b2 ^ b1)
proof end;

theorem Th44: :: FINSEQ_1:44
for b1, b2 being FinSequence holds rng (b1 ^ b2) = (rng b1) \/ (rng b2)
proof end;

theorem Th45: :: FINSEQ_1:45
for b1, b2, b3 being FinSequence holds (b1 ^ b2) ^ b3 = b1 ^ (b2 ^ b3)
proof end;

theorem Th46: :: FINSEQ_1:46
for b1, b2, b3 being FinSequence holds
( ( b1 ^ b2 = b3 ^ b2 or b2 ^ b1 = b2 ^ b3 ) implies b1 = b3 )
proof end;

theorem Th47: :: FINSEQ_1:47
for b1 being FinSequence holds
( b1 ^ {} = b1 & {} ^ b1 = b1 )
proof end;

theorem Th48: :: FINSEQ_1:48
for b1, b2 being FinSequence holds
( b1 ^ b2 = {} implies ( b1 = {} & b2 = {} ) )
proof end;

definition
let c1 be set ;
let c2, c3 be FinSequence of c1;
redefine func ^ as c2 ^ c3 -> FinSequence of a1;
coherence
c2 ^ c3 is FinSequence of c1
proof end;
end;

Lemma35: for b1, b2 being set holds
{[b1,b2]} is Function
proof end;

Lemma36: for b1, b2, b3, b4 being set holds
( [b1,b2] in {[b3,b4]} implies ( b1 = b3 & b2 = b4 ) )
proof end;

definition
let c1 be set ;
redefine func <* as <*c1*> -> Function means :Def8: :: FINSEQ_1:def 8
( dom a2 = Seg 1 & a2 . 1 = a1 );
coherence
<*c1*> is Function
by Lemma35;
compatibility
for b1 being Function holds
( b1 = <*c1*> iff ( dom b1 = Seg 1 & b1 . 1 = c1 ) )
proof end;
end;

:: deftheorem Def8 defines <* FINSEQ_1:def 8 :
for b1 being set
for b2 being Function holds
( b2 = <*b1*> iff ( dom b2 = Seg 1 & b2 . 1 = b1 ) );

registration
let c1 be set ;
cluster <*a1*> -> Relation-like Function-like ;
coherence
( <*c1*> is Function-like & <*c1*> is Relation-like )
;
end;

registration
let c1 be set ;
cluster <*a1*> -> Relation-like Function-like finite FinSequence-like ;
coherence
<*c1*> is FinSequence-like
proof end;
end;

theorem Th49: :: FINSEQ_1:49
canceled;

theorem Th50: :: FINSEQ_1:50
for b1, b2 being FinSequence
for b3 being set holds
( b1 ^ b2 is FinSequence of b3 implies ( b1 is FinSequence of b3 & b2 is FinSequence of b3 ) )
proof end;

definition
let c1, c2 be set ;
func <*c1,c2*> -> set equals :: FINSEQ_1:def 9
<*a1*> ^ <*a2*>;
correctness
coherence
<*c1*> ^ <*c2*> is set
;
;
let c3 be set ;
func <*c1,c2,c3*> -> set equals :: FINSEQ_1:def 10
(<*a1*> ^ <*a2*>) ^ <*a3*>;
correctness
coherence
(<*c1*> ^ <*c2*>) ^ <*c3*> is set
;
;
end;

:: deftheorem Def9 defines <* FINSEQ_1:def 9 :
for b1, b2 being set holds <*b1,b2*> = <*b1*> ^ <*b2*>;

:: deftheorem Def10 defines <* FINSEQ_1:def 10 :
for b1, b2, b3 being set holds <*b1,b2,b3*> = (<*b1*> ^ <*b2*>) ^ <*b3*>;

registration
let c1, c2 be set ;
cluster <*a1,a2*> -> Relation-like Function-like ;
coherence
( <*c1,c2*> is Function-like & <*c1,c2*> is Relation-like )
;
let c3 be set ;
cluster <*a1,a2,a3*> -> Relation-like Function-like ;
coherence
( <*c1,c2,c3*> is Function-like & <*c1,c2,c3*> is Relation-like )
;
end;

registration
let c1, c2 be set ;
cluster <*a1,a2*> -> Relation-like Function-like finite FinSequence-like ;
coherence
<*c1,c2*> is FinSequence-like
;
let c3 be set ;
cluster <*a1,a2,a3*> -> Relation-like Function-like finite FinSequence-like ;
coherence
<*c1,c2,c3*> is FinSequence-like
;
end;

theorem Th51: :: FINSEQ_1:51
canceled;

theorem Th52: :: FINSEQ_1:52
for b1 being set holds <*b1*> = {[1,b1]} ;

theorem Th53: :: FINSEQ_1:53
canceled;

theorem Th54: :: FINSEQ_1:54
canceled;

theorem Th55: :: FINSEQ_1:55
for b1 being set
for b2 being FinSequence holds
( b2 = <*b1*> iff ( dom b2 = Seg 1 & rng b2 = {b1} ) )
proof end;

theorem Th56: :: FINSEQ_1:56
for b1 being set
for b2 being FinSequence holds
( b2 = <*b1*> iff ( len b2 = 1 & rng b2 = {b1} ) )
proof end;

theorem Th57: :: FINSEQ_1:57
for b1 being set
for b2 being FinSequence holds
( b2 = <*b1*> iff ( len b2 = 1 & b2 . 1 = b1 ) )
proof end;

theorem Th58: :: FINSEQ_1:58
for b1 being set
for b2 being FinSequence holds (<*b1*> ^ b2) . 1 = b1
proof end;

theorem Th59: :: FINSEQ_1:59
for b1 being set
for b2 being FinSequence holds (b2 ^ <*b1*>) . ((len b2) + 1) = b1
proof end;

theorem Th60: :: FINSEQ_1:60
for b1, b2, b3 being set holds
( <*b1,b2,b3*> = <*b1*> ^ <*b2,b3*> & <*b1,b2,b3*> = <*b1,b2*> ^ <*b3*> ) by Th45;

theorem Th61: :: FINSEQ_1:61
for b1, b2 being set
for b3 being FinSequence holds
( b3 = <*b1,b2*> iff ( len b3 = 2 & b3 . 1 = b1 & b3 . 2 = b2 ) )
proof end;

theorem Th62: :: FINSEQ_1:62
for b1, b2, b3 being set
for b4 being FinSequence holds
( b4 = <*b1,b2,b3*> iff ( len b4 = 3 & b4 . 1 = b1 & b4 . 2 = b2 & b4 . 3 = b3 ) )
proof end;

theorem Th63: :: FINSEQ_1:63
for b1 being FinSequence holds
not ( b1 <> {} & ( for b2 being FinSequence
for b3 being set holds
not b1 = b2 ^ <*b3*> ) )
proof end;

definition
let c1 be non empty set ;
let c2 be Element of c1;
redefine func <* as <*c2*> -> FinSequence of a1;
coherence
<*c2*> is FinSequence of c1
proof end;
end;

scheme :: FINSEQ_1:sch 3
s3{ P1[ FinSequence] } :
for b1 being FinSequence holds P1[b1]
provided
E47: P1[ {} ] and E48: for b1 being FinSequence
for b2 being set holds
( P1[b1] implies P1[b1 ^ <*b2*>] )
proof end;

theorem Th64: :: FINSEQ_1:64
for b1, b2, b3, b4 being FinSequence holds
not ( b1 ^ b2 = b3 ^ b4 & len b1 <= len b3 & ( for b5 being FinSequence holds
not b1 ^ b5 = b3 ) )
proof end;

definition
let c1 be set ;
func c1 * -> set means :Def11: :: FINSEQ_1:def 11
for b1 being set holds
( b1 in a2 iff b1 is FinSequence of a1 );
existence
ex b1 being set st
for b2 being set holds
( b2 in b1 iff b2 is FinSequence of c1 )
proof end;
uniqueness
for b1, b2 being set holds
( ( for b3 being set holds
( b3 in b1 iff b3 is FinSequence of c1 ) ) & ( for b3 being set holds
( b3 in b2 iff b3 is FinSequence of c1 ) ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def11 defines * FINSEQ_1:def 11 :
for b1, b2 being set holds
( b2 = b1 * iff for b3 being set holds
( b3 in b2 iff b3 is FinSequence of b1 ) );

registration
let c1 be set ;
cluster a1 * -> non empty ;
coherence
not c1 * is empty
proof end;
end;

theorem Th65: :: FINSEQ_1:65
for b1, b2 being FinSequence holds
( rng b1 = rng b2 & b1 is one-to-one & b2 is one-to-one implies len b1 = len b2 )
proof end;

theorem Th66: :: FINSEQ_1:66
for b1 being set holds {} in b1 *
proof end;

scheme :: FINSEQ_1:sch 4
s4{ F1() -> non empty set , P1[ FinSequence] } :
ex b1 being set st
for b2 being set holds
( b2 in b1 iff ex b3 being FinSequence st
( b3 in F1() * & P1[b3] & b2 = b3 ) )
proof end;

definition
let c1 be Function;
attr a1 is FinSubsequence-like means :Def12: :: FINSEQ_1:def 12
ex b1 being Nat st dom a1 c= Seg b1;
end;

:: deftheorem Def12 defines FinSubsequence-like FINSEQ_1:def 12 :
for b1 being Function holds
( b1 is FinSubsequence-like iff ex b2 being Nat st dom b1 c= Seg b2 );

registration
cluster FinSubsequence-like set ;
existence
ex b1 being Function st b1 is FinSubsequence-like
proof end;
end;

definition
mode FinSubsequence is FinSubsequence-like Function;
end;

theorem Th67: :: FINSEQ_1:67
canceled;

theorem Th68: :: FINSEQ_1:68
for b1 being FinSequence holds
b1 is FinSubsequence
proof end;

theorem Th69: :: FINSEQ_1:69
for b1 being set
for b2 being FinSequence holds
( b2 | b1 is FinSubsequence & b1 | b2 is FinSubsequence )
proof end;

definition
let c1 be set ;
given c2 being natural number such that E49: c1 c= Seg c2 ;
func Sgm c1 -> FinSequence of NAT means :Def13: :: FINSEQ_1:def 13
( rng a2 = a1 & ( for b1, b2, b3, b4 being natural number holds
not ( 1 <= b1 & b1 < b2 & b2 <= len a2 & b3 = a2 . b1 & b4 = a2 . b2 & not b3 < b4 ) ) );
existence
ex b1 being FinSequence of NAT st
( rng b1 = c1 & ( for b2, b3, b4, b5 being natural number holds
not ( 1 <= b2 & b2 < b3 & b3 <= len b1 & b4 = b1 . b2 & b5 = b1 . b3 & not b4 < b5 ) ) )
proof end;
uniqueness
for b1, b2 being FinSequence of NAT holds
( rng b1 = c1 & ( for b3, b4, b5, b6 being natural number holds
not ( 1 <= b3 & b3 < b4 & b4 <= len b1 & b5 = b1 . b3 & b6 = b1 . b4 & not b5 < b6 ) ) & rng b2 = c1 & ( for b3, b4, b5, b6 being natural number holds
not ( 1 <= b3 & b3 < b4 & b4 <= len b2 & b5 = b2 . b3 & b6 = b2 . b4 & not b5 < b6 ) ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def13 defines Sgm FINSEQ_1:def 13 :
for b1 being set holds
( ex b2 being natural number st b1 c= Seg b2 implies for b2 being FinSequence of NAT holds
( b2 = Sgm b1 iff ( rng b2 = b1 & ( for b3, b4, b5, b6 being natural number holds
not ( 1 <= b3 & b3 < b4 & b4 <= len b2 & b5 = b2 . b3 & b6 = b2 . b4 & not b5 < b6 ) ) ) ) );

theorem Th70: :: FINSEQ_1:70
canceled;

theorem Th71: :: FINSEQ_1:71
for b1 being FinSubsequence holds rng (Sgm (dom b1)) = dom b1
proof end;

definition
let c1 be FinSubsequence;
func Seq c1 -> Function equals :: FINSEQ_1:def 14
a1 * (Sgm (dom a1));
coherence
c1 * (Sgm (dom c1)) is Function
;
end;

:: deftheorem Def14 defines Seq FINSEQ_1:def 14 :
for b1 being FinSubsequence holds Seq b1 = b1 * (Sgm (dom b1));

registration
let c1 be FinSubsequence;
cluster Seq a1 -> finite FinSequence-like ;
coherence
Seq c1 is FinSequence-like
proof end;
end;

theorem Th72: :: FINSEQ_1:72
for b1 being set holds
( ex b2 being Nat st b1 c= Seg b2 implies ( Sgm b1 = {} iff b1 = {} ) )
proof end;

theorem Th73: :: FINSEQ_1:73
for b1 being set holds
( b1 is finite iff ex b2 being FinSequence st b1 = rng b2 )
proof end;

registration
cluster empty finite set ;
existence
ex b1 being Function st
( b1 is finite & b1 is empty )
proof end;
end;

registration
cluster non empty finite set ;
existence
ex b1 being Function st
( b1 is finite & not b1 is empty )
proof end;
end;

registration
let c1 be finite Relation;
cluster rng a1 -> finite ;
coherence
rng c1 is finite
proof end;
end;

theorem Th74: :: FINSEQ_1:74
for b1, b2 being Nat holds
( Seg b1, Seg b2 are_equipotent implies b1 = b2 )
proof end;

theorem Th75: :: FINSEQ_1:75
for b1 being Nat holds Seg b1,b1 are_equipotent by Lemma9;

theorem Th76: :: FINSEQ_1:76
for b1 being Nat holds Card (Seg b1) = Card b1 by Lemma10;

theorem Th77: :: FINSEQ_1:77
for b1 being set holds
not ( b1 is finite & ( for b2 being Nat holds
not b1, Seg b2 are_equipotent ) )
proof end;

theorem Th78: :: FINSEQ_1:78
( ( for b1 being natural number holds card (Seg b1) = b1 ) & ( for b1 being Nat holds
( card b1 = b1 & card (Card b1) = b1 ) ) )
proof end;

registration
let c1 be set ;
cluster <*a1*> -> Relation-like Function-like non empty finite FinSequence-like ;
coherence
not <*c1*> is empty
;
end;

registration
cluster non empty finite set ;
existence
not for b1 being FinSequence holds b1 is empty
proof end;
end;

registration
let c1 be FinSequence;
let c2 be non empty FinSequence;
cluster a1 ^ a2 -> non empty finite ;
coherence
not c1 ^ c2 is empty
by Th48;
cluster a2 ^ a1 -> non empty finite ;
coherence
not c2 ^ c1 is empty
by Th48;
end;

scheme :: FINSEQ_1:sch 5
s5{ F1() -> non empty set , F2() -> Nat, P1[ set , set ] } :
ex b1 being FinSequence of F1() st
( dom b1 = Seg F2() & ( for b2 being Nat holds
( b2 in Seg F2() implies P1[b2,b1 . b2] ) ) )
provided
E51: for b1 being Nat holds
not ( b1 in Seg F2() & ( for b2 being Element of F1() holds
not P1[b1,b2] ) )
proof end;

definition
let c1 be FinSequence;
let c2 be natural number ;
func c1 | c2 -> FinSequence equals :: FINSEQ_1:def 15
a1 | (Seg a2);
coherence
c1 | (Seg c2) is FinSequence
by Th19;
end;

:: deftheorem Def15 defines | FINSEQ_1:def 15 :
for b1 being FinSequence
for b2 being natural number holds b1 | b2 = b1 | (Seg b2);

definition
let c1 be set ;
let c2 be FinSequence of c1;
let c3 be natural number ;
redefine func | as c2 | c3 -> FinSequence of a1;
coherence
c2 | c3 is FinSequence of c1
by Th23;
end;

registration
let c1 be set ;
let c2 be FinSequence of c1;
cluster a2 | 0 -> empty finite ;
coherence
c2 | 0 is empty
by Th4, RELAT_1:110;
end;

theorem Th79: :: FINSEQ_1:79
for b1 being natural number
for b2 being FinSequence holds
( len b2 <= b1 implies b2 | b1 = b2 )
proof end;

theorem Th80: :: FINSEQ_1:80
for b1 being natural number
for b2 being FinSequence holds
( b1 <= len b2 implies len (b2 | b1) = b1 )
proof end;

theorem Th81: :: FINSEQ_1:81
for b1 being Nat
for b2, b3 being natural number holds
( b2 in Seg b1 implies b2 + b3 in Seg (b1 + b3) )
proof end;

theorem Th82: :: FINSEQ_1:82
for b1 being Nat
for b2, b3 being natural number holds
( b2 > 0 & b2 + b3 in Seg (b1 + b3) implies ( b2 in Seg b1 & b2 in Seg (b1 + b3) ) )
proof end;

definition
let c1 be Relation;
func c1 [*] -> Relation means :: FINSEQ_1:def 16
for b1, b2 being set holds
( [b1,b2] in a2 iff ( b1 in field a1 & b2 in field a1 & ex b3 being FinSequence st
( len b3 >= 1 & b3 . 1 = b1 & b3 . (len b3) = b2 & ( for b4 being Nat holds
( b4 >= 1 & b4 < len b3 implies [(b3 . b4),(b3 . (b4 + 1))] in a1 ) ) ) ) );
existence
ex b1 being Relation st
for b2, b3 being set holds
( [b2,b3] in b1 iff ( b2 in field c1 & b3 in field c1 & ex b4 being FinSequence st
( len b4 >= 1 & b4 . 1 = b2 & b4 . (len b4) = b3 & ( for b5 being Nat holds
( b5 >= 1 & b5 < len b4 implies [(b4 . b5),(b4 . (b5 + 1))] in c1 ) ) ) ) )
proof end;
uniqueness
for b1, b2 being Relation holds
( ( for b3, b4 being set holds
( [b3,b4] in b1 iff ( b3 in field c1 & b4 in field c1 & ex b5 being FinSequence st
( len b5 >= 1 & b5 . 1 = b3 & b5 . (len b5) = b4 & ( for b6 being Nat holds
( b6 >= 1 & b6 < len b5 implies [(b5 . b6),(b5 . (b6 + 1))] in c1 ) ) ) ) ) ) & ( for b3, b4 being set holds
( [b3,b4] in b2 iff ( b3 in field c1 & b4 in field c1 & ex b5 being FinSequence st
( len b5 >= 1 & b5 . 1 = b3 & b5 . (len b5) = b4 & ( for b6 being Nat holds
( b6 >= 1 & b6 < len b5 implies [(b5 . b6),(b5 . (b6 + 1))] in c1 ) ) ) ) ) ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def16 defines [*] FINSEQ_1:def 16 :
for b1, b2 being Relation holds
( b2 = b1 [*] iff for b3, b4 being set holds
( [b3,b4] in b2 iff ( b3 in field b1 & b4 in field b1 & ex b5 being FinSequence st
( len b5 >= 1 & b5 . 1 = b3 & b5 . (len b5) = b4 & ( for b6 being Nat holds
( b6 >= 1 & b6 < len b5 implies [(b5 . b6),(b5 . (b6 + 1))] in b1 ) ) ) ) ) );

theorem Th83: :: FINSEQ_1:83
for b1, b2 being set holds
( b1 c= b2 implies b1 * c= b2 * )
proof end;

registration
let c1 be set ;
cluster a1 * -> non empty functional ;
coherence
c1 * is functional
proof end;
end;

registration
let c1 be finite Function;
cluster dom a1 -> finite ;
coherence
dom c1 is finite
by FINSET_1:29;
end;

registration
let c1 be finite Function;
let c2 be set ;
cluster a1 " a2 -> finite ;
coherence
c1 " c2 is finite
proof end;
end;

theorem Th84: :: FINSEQ_1:84
for b1, b2 being FinSequence holds
( b1 c= b2 implies len b1 <= len b2 )
proof end;

theorem Th85: :: FINSEQ_1:85
for b1, b2 being FinSequence
for b3 being Nat holds
( 1 <= b3 & b3 <= len b1 implies (b1 ^ b2) . b3 = b1 . b3 )
proof end;

theorem Th86: :: FINSEQ_1:86
for b1, b2 being FinSequence
for b3 being Nat holds
( 1 <= b3 & b3 <= len b2 implies (b1 ^ b2) . ((len b1) + b3) = b2 . b3 )
proof end;

scheme :: FINSEQ_1:sch 6
s6{ F1() -> Nat, F2() -> Nat, F3( set ) -> set , P1[ Nat] } :
{ F3(b1) where B is Nat : ( F1() <= b1 & b1 <= F2() & P1[b1] ) } is finite
proof end;

Lemma51: ( 1 in Seg 3 & 2 in Seg 3 & 3 in Seg 3 )
;

Lemma52: ( 1 in Seg 4 & 2 in Seg 4 & 3 in Seg 4 & 4 in Seg 4 )
;

Lemma53: ( 1 in Seg 5 & 2 in Seg 5 & 3 in Seg 5 & 4 in Seg 5 & 5 in Seg 5 )
;

Lemma54: ( 1 in Seg 6 & 2 in Seg 6 & 3 in Seg 6 & 4 in Seg 6 & 5 in Seg 6 & 6 in Seg 6 )
;

Lemma55: ( 1 in Seg 7 & 2 in Seg 7 & 3 in Seg 7 & 4 in Seg 7 & 5 in Seg 7 & 6 in Seg 7 & 7 in Seg 7 )
;

Lemma56: ( 1 in Seg 8 & 2 in Seg 8 & 3 in Seg 8 & 4 in Seg 8 & 5 in Seg 8 & 6 in Seg 8 & 7 in Seg 8 & 8 in Seg 8 )
;

theorem Th87: :: FINSEQ_1:87
for b1, b2, b3, b4 being set
for b5 being FinSequence holds
( b5 = ((<*b1*> ^ <*b2*>) ^ <*b3*>) ^ <*b4*> implies ( len b5 = 4 & b5 . 1 = b1 & b5 . 2 = b2 & b5 . 3 = b3 & b5 . 4 = b4 ) )
proof end;

theorem Th88: :: FINSEQ_1:88
for b1, b2, b3, b4, b5 being set
for b6 being FinSequence holds
( b6 = (((<*b1*> ^ <*b2*>) ^ <*b3*>) ^ <*b4*>) ^ <*b5*> implies ( len b6 = 5 & b6 . 1 = b1 & b6 . 2 = b2 & b6 . 3 = b3 & b6 . 4 = b4 & b6 . 5 = b5 ) )
proof end;

theorem Th89: :: FINSEQ_1:89
for b1, b2, b3, b4, b5, b6 being set
for b7 being FinSequence holds
( b7 = ((((<*b1*> ^ <*b2*>) ^ <*b3*>) ^ <*b4*>) ^ <*b5*>) ^ <*b6*> implies ( len b7 = 6 & b7 . 1 = b1 & b7 . 2 = b2 & b7 . 3 = b3 & b7 . 4 = b4 & b7 . 5 = b5 & b7 . 6 = b6 ) )
proof end;

theorem Th90: :: FINSEQ_1:90
for b1, b2, b3, b4, b5, b6, b7 being set
for b8 being FinSequence holds
( b8 = (((((<*b1*> ^ <*b2*>) ^ <*b3*>) ^ <*b4*>) ^ <*b5*>) ^ <*b6*>) ^ <*b7*> implies ( len b8 = 7 & b8 . 1 = b1 & b8 . 2 = b2 & b8 . 3 = b3 & b8 . 4 = b4 & b8 . 5 = b5 & b8 . 6 = b6 & b8 . 7 = b7 ) )
proof end;

theorem Th91: :: FINSEQ_1:91
for b1, b2, b3, b4, b5, b6, b7, b8 being set
for b9 being FinSequence holds
( b9 = ((((((<*b1*> ^ <*b2*>) ^ <*b3*>) ^ <*b4*>) ^ <*b5*>) ^ <*b6*>) ^ <*b7*>) ^ <*b8*> implies ( len b9 = 8 & b9 . 1 = b1 & b9 . 2 = b2 & b9 . 3 = b3 & b9 . 4 = b4 & b9 . 5 = b5 & b9 . 6 = b6 & b9 . 7 = b7 & b9 . 8 = b8 ) )
proof end;

theorem Th92: :: FINSEQ_1:92
for b1, b2, b3, b4, b5, b6, b7, b8, b9 being set
for b10 being FinSequence holds
( b10 = (((((((<*b1*> ^ <*b2*>) ^ <*b3*>) ^ <*b4*>) ^ <*b5*>) ^ <*b6*>) ^ <*b7*>) ^ <*b8*>) ^ <*b9*> implies ( len b10 = 9 & b10 . 1 = b1 & b10 . 2 = b2 & b10 . 3 = b3 & b10 . 4 = b4 & b10 . 5 = b5 & b10 . 6 = b6 & b10 . 7 = b7 & b10 . 8 = b8 & b10 . 9 = b9 ) )
proof end;

theorem Th93: :: FINSEQ_1:93
for b1 being FinSequence holds b1 | (Seg 0) = {}
proof end;

theorem Th94: :: FINSEQ_1:94
for b1, b2 being FinSequence holds b1 | (Seg 0) = b2 | (Seg 0)
proof end;

theorem Th95: :: FINSEQ_1:95
for b1 being non empty set
for b2 being Element of b1 holds
<*b2*> is FinSequence of b1 ;

theorem Th96: :: FINSEQ_1:96
for b1 being set
for b2, b3 being FinSequence of b1 holds
b2 ^ b3 is FinSequence of b1 ;

Lemma63: for b1 being Relation holds
not ( dom b1 <> {} & not b1 <> {} )
by RELAT_1:60;

registration
let c1 be Function;
let c2 be non empty Function;
cluster a1 +* a2 -> non empty ;
coherence
not c1 +* c2 is empty
proof end;
cluster a2 +* a1 -> non empty ;
coherence
not c2 +* c1 is empty
proof end;
end;

registration
let c1, c2 be finite Function;
cluster a1 +* a2 -> finite ;
coherence
c1 +* c2 is finite
proof end;
end;