:: Non-contiguous Substrings and One-to-one Finite Sequences
:: by Wojciech A. Trybulec
::
:: Received April 8, 1990
:: Copyright (c) 1990-2021 Association of Mizar Users


theorem Th1: :: FINSEQ_3:1
Seg 3 = {1,2,3}
proof end;

theorem Th2: :: FINSEQ_3:2
Seg 4 = {1,2,3,4}
proof end;

theorem Th3: :: FINSEQ_3:3
Seg 5 = {1,2,3,4,5}
proof end;

theorem Th4: :: FINSEQ_3:4
Seg 6 = {1,2,3,4,5,6}
proof end;

theorem Th5: :: FINSEQ_3:5
Seg 7 = {1,2,3,4,5,6,7}
proof end;

theorem :: FINSEQ_3:6
Seg 8 = {1,2,3,4,5,6,7,8}
proof end;

theorem Th7: :: FINSEQ_3:7
for k being Nat holds
( Seg k = {} iff not k in Seg k )
proof end;

theorem Th8: :: FINSEQ_3:8
for k being Nat holds not k + 1 in Seg k
proof end;

theorem :: FINSEQ_3:9
for k, n being Nat st k <> 0 holds
k in Seg (k + n)
proof end;

theorem Th10: :: FINSEQ_3:10
for k, n being Nat st k + n in Seg k holds
n = 0
proof end;

theorem :: FINSEQ_3:11
for k, n being Nat st k < n holds
k + 1 in Seg n
proof end;

theorem Th12: :: FINSEQ_3:12
for k, m, n being Nat st k in Seg n & m < k holds
k - m in Seg n
proof end;

theorem :: FINSEQ_3:13
for k, n being Nat holds
( k - n in Seg k iff n < k )
proof end;

theorem Th14: :: FINSEQ_3:14
for k being Nat holds Seg k misses {(k + 1)}
proof end;

theorem :: FINSEQ_3:15
for k being Nat holds (Seg (k + 1)) \ (Seg k) = {(k + 1)}
proof end;

:: Theorem Seg(k + 1) \ {k + 1} = Seg k is
:: proved in RLVECT_1 and has a number 104.
theorem :: FINSEQ_3:16
for k being Nat holds Seg k <> Seg (k + 1) by Th8, FINSEQ_1:4;

theorem :: FINSEQ_3:17
for k, n being Nat st Seg k = Seg (k + n) holds
n = 0 by Th10, FINSEQ_1:3;

theorem Th18: :: FINSEQ_3:18
for k, n being Nat holds Seg k c= Seg (k + n) by FINSEQ_1:5, NAT_1:12;

theorem Th19: :: FINSEQ_3:19
for k, n being Nat holds Seg k, Seg n are_c=-comparable
proof end;

theorem Th20: :: FINSEQ_3:20
for k being Nat
for y being object st Seg k = {y} holds
( k = 1 & y = 1 )
proof end;

theorem :: FINSEQ_3:21
for x, y being object
for k being Nat st Seg k = {x,y} & x <> y holds
( k = 2 & {x,y} = {1,2} )
proof end;

theorem Th22: :: FINSEQ_3:22
for p, q being FinSequence
for x being object st x in dom p holds
x in dom (p ^ q)
proof end;

theorem :: FINSEQ_3:23
for p being FinSequence
for x being object st x in dom p holds
x is Element of NAT ;

theorem Th24: :: FINSEQ_3:24
for p being FinSequence
for x being object st x in dom p holds
x <> 0
proof end;

theorem Th25: :: FINSEQ_3:25
for p being FinSequence
for n being Nat holds
( n in dom p iff ( 1 <= n & n <= len p ) )
proof end;

theorem :: FINSEQ_3:26
for p being FinSequence
for n being Nat holds
( n in dom p iff ( n - 1 is Element of NAT & (len p) - n is Element of NAT ) )
proof end;

theorem :: FINSEQ_3:27
canceled;

theorem :: FINSEQ_3:28
canceled;

::$CT 2
theorem Th27: :: FINSEQ_3:29
for p, q being FinSequence holds
( len p = len q iff dom p = dom q )
proof end;

theorem Th28: :: FINSEQ_3:30
for p, q being FinSequence holds
( len p <= len q iff dom p c= dom q )
proof end;

theorem Th29: :: FINSEQ_3:31
for p being FinSequence
for x being object st x in rng p holds
1 in dom p
proof end;

theorem :: FINSEQ_3:32
for p being FinSequence st rng p <> {} holds
1 in dom p
proof end;

theorem :: FINSEQ_3:33
for x, y being object holds {} <> <*x,y*> ;

theorem :: FINSEQ_3:34
for x, y, z being object holds {} <> <*x,y,z*> ;

theorem Th33: :: FINSEQ_3:35
for x, y, z being object holds <*x*> <> <*y,z*>
proof end;

theorem :: FINSEQ_3:36
for u, x, y, z being object holds <*u*> <> <*x,y,z*>
proof end;

theorem :: FINSEQ_3:37
for u, v, x, y, z being object holds <*u,v*> <> <*x,y,z*>
proof end;

theorem Th36: :: FINSEQ_3:38
for p, q, r being FinSequence st len r = (len p) + (len q) & ( for k being Nat st k in dom p holds
r . k = p . k ) & ( for k being Nat st k in dom q holds
r . ((len p) + k) = q . k ) holds
r = p ^ q
proof end;

Lm1: for A being set
for k being Nat st A c= Seg k holds
Sgm A is one-to-one

proof end;

theorem Th37: :: FINSEQ_3:39
for k being Nat
for A being finite set st A c= Seg k holds
len (Sgm A) = card A
proof end;

theorem Th38: :: FINSEQ_3:40
for k being Nat
for A being finite set st A c= Seg k holds
dom (Sgm A) = Seg (card A)
proof end;

theorem Th39: :: FINSEQ_3:41
for X being set
for i, k, l, m, n being Nat st X c= Seg i & k < l & 1 <= n & m <= len (Sgm X) & (Sgm X) . m = k & (Sgm X) . n = l holds
m < n
proof end;

theorem Th40: :: FINSEQ_3:42
for X, Y being set
for i, j being Nat st X c= Seg i & Y c= Seg j holds
( ( for m, n being Nat st m in X & n in Y holds
m < n ) iff Sgm (X \/ Y) = (Sgm X) ^ (Sgm Y) )
proof end;

theorem Th41: :: FINSEQ_3:43
Sgm {} = {}
proof end;

:: The other way of the one above - FINSEQ_1:72.
theorem Th42: :: FINSEQ_3:44
for n being Nat st 0 <> n holds
Sgm {n} = <*n*>
proof end;

theorem :: FINSEQ_3:45
for m, n being Nat st 0 < n & n < m holds
Sgm {n,m} = <*n,m*>
proof end;

theorem Th44: :: FINSEQ_3:46
for k being Nat holds len (Sgm (Seg k)) = k
proof end;

Lm2: for k being Nat holds (Sgm (Seg (k + 0))) | (Seg k) = Sgm (Seg k)
proof end;

Lm3: now :: thesis: for n being Nat st ( for k being Nat holds (Sgm (Seg (k + n))) | (Seg k) = Sgm (Seg k) ) holds
for k being Nat holds (Sgm (Seg (k + (n + 1)))) | (Seg k) = Sgm (Seg k)
let n be Nat; :: thesis: ( ( for k being Nat holds (Sgm (Seg (k + n))) | (Seg k) = Sgm (Seg k) ) implies for k being Nat holds (Sgm (Seg (k + (n + 1)))) | (Seg k) = Sgm (Seg k) )
assume A1: for k being Nat holds (Sgm (Seg (k + n))) | (Seg k) = Sgm (Seg k) ; :: thesis: for k being Nat holds (Sgm (Seg (k + (n + 1)))) | (Seg k) = Sgm (Seg k)
let k be Nat; :: thesis: (Sgm (Seg (k + (n + 1)))) | (Seg k) = Sgm (Seg k)
set X = Sgm (Seg (k + (n + 1)));
set Y = Sgm (Seg (k + 1));
A2: (Sgm (Seg (k + 1))) | (Seg k) = Sgm (Seg k)
proof
reconsider p = (Sgm (Seg (k + 1))) | (Seg k) as FinSequence of NAT by FINSEQ_1:18;
A3: len (Sgm (Seg (k + 1))) = k + 1 by Th44;
then A4: dom (Sgm (Seg (k + 1))) = Seg (k + 1) by FINSEQ_1:def 3;
A5: k <= k + 1 by NAT_1:12;
then A6: dom p = Seg k by A3, FINSEQ_1:17;
A7: rng (Sgm (Seg (k + 1))) = Seg (k + 1) by FINSEQ_1:def 13;
A8: (Sgm (Seg (k + 1))) . (k + 1) = k + 1
proof
k + 1 in dom (Sgm (Seg (k + 1))) by A4, FINSEQ_1:4;
then A9: (Sgm (Seg (k + 1))) . (k + 1) in Seg (k + 1) by A7, FUNCT_1:def 3;
then reconsider n = (Sgm (Seg (k + 1))) . (k + 1) as Element of NAT ;
A10: k < k + 1 by XREAL_1:29;
k + 1 in rng (Sgm (Seg (k + 1))) by A7, FINSEQ_1:4;
then consider x being object such that
A11: x in dom (Sgm (Seg (k + 1))) and
A12: (Sgm (Seg (k + 1))) . x = k + 1 by FUNCT_1:def 3;
reconsider x = x as Element of NAT by A11;
assume not (Sgm (Seg (k + 1))) . (k + 1) = k + 1 ; :: thesis: contradiction
then not (Sgm (Seg (k + 1))) . (k + 1) in {(k + 1)} by TARSKI:def 1;
then (Sgm (Seg (k + 1))) . (k + 1) in (Seg (k + 1)) \ {(k + 1)} by A9, XBOOLE_0:def 5;
then A13: (Sgm (Seg (k + 1))) . (k + 1) in Seg k by FINSEQ_1:10;
A14: x <> k + 1 by A12, A13, FINSEQ_1:1, XREAL_1:29;
x <= k + 1 by A3, A11, Th25;
then A15: x < k + 1 by A14, XXREAL_0:1;
1 <= x by A11, Th25;
then A16: k + 1 < n by A3, A12, A15, FINSEQ_1:def 13;
n <= k by A13, FINSEQ_1:1;
hence contradiction by A16, A10, XXREAL_0:2; :: thesis: verum
end;
A17: Sgm (Seg (k + 1)) is one-to-one by Lm1;
rng p = (rng (Sgm (Seg (k + 1)))) \ {((Sgm (Seg (k + 1))) . (k + 1))}
proof
thus rng p c= (rng (Sgm (Seg (k + 1)))) \ {((Sgm (Seg (k + 1))) . (k + 1))} :: according to XBOOLE_0:def 10 :: thesis: (rng (Sgm (Seg (k + 1)))) \ {((Sgm (Seg (k + 1))) . (k + 1))} c= rng p
proof
let x be object ; :: according to TARSKI:def 3 :: thesis: ( not x in rng p or x in (rng (Sgm (Seg (k + 1)))) \ {((Sgm (Seg (k + 1))) . (k + 1))} )
assume A18: x in rng p ; :: thesis: x in (rng (Sgm (Seg (k + 1)))) \ {((Sgm (Seg (k + 1))) . (k + 1))}
A19: now :: thesis: not x in {(k + 1)}
assume x in {(k + 1)} ; :: thesis: contradiction
then A20: x = k + 1 by TARSKI:def 1;
A21: k < k + 1 by XREAL_1:29;
A22: k + 1 in dom (Sgm (Seg (k + 1))) by A4, FINSEQ_1:4;
A23: Seg k c= Seg (k + 1) by Th18;
consider y being object such that
A24: y in dom p and
A25: p . y = x by A18, FUNCT_1:def 3;
reconsider y = y as Element of NAT by A24;
A26: (Sgm (Seg (k + 1))) . y = p . y by A24, FUNCT_1:47;
y <= k by A6, A24, FINSEQ_1:1;
hence contradiction by A17, A4, A6, A8, A24, A25, A23, A26, A20, A22, A21; :: thesis: verum
end;
rng p c= rng (Sgm (Seg (k + 1))) by RELAT_1:70;
hence x in (rng (Sgm (Seg (k + 1)))) \ {((Sgm (Seg (k + 1))) . (k + 1))} by A8, A18, A19, XBOOLE_0:def 5; :: thesis: verum
end;
let x be object ; :: according to TARSKI:def 3 :: thesis: ( not x in (rng (Sgm (Seg (k + 1)))) \ {((Sgm (Seg (k + 1))) . (k + 1))} or x in rng p )
assume A27: x in (rng (Sgm (Seg (k + 1)))) \ {((Sgm (Seg (k + 1))) . (k + 1))} ; :: thesis: x in rng p
then x in rng (Sgm (Seg (k + 1))) by XBOOLE_0:def 5;
then consider y being object such that
A28: y in dom (Sgm (Seg (k + 1))) and
A29: (Sgm (Seg (k + 1))) . y = x by FUNCT_1:def 3;
now :: thesis: not y in {(k + 1)}
assume y in {(k + 1)} ; :: thesis: contradiction
then A30: x = k + 1 by A8, A29, TARSKI:def 1;
not x in {(k + 1)} by A8, A27, XBOOLE_0:def 5;
hence contradiction by A30, TARSKI:def 1; :: thesis: verum
end;
then y in (Seg (k + 1)) \ {(k + 1)} by A4, A28, XBOOLE_0:def 5;
then A31: y in dom p by A6, FINSEQ_1:10;
then p . y = (Sgm (Seg (k + 1))) . y by FUNCT_1:47;
hence x in rng p by A29, A31, FUNCT_1:def 3; :: thesis: verum
end;
then A32: rng p = Seg k by A7, A8, FINSEQ_1:10;
now :: thesis: for i, j, l, m being Nat st 1 <= i & i < j & j <= len p & l = p . i & m = p . j holds
l < m
A33: len p = k by A3, A5, FINSEQ_1:17;
let i, j, l, m be Nat; :: thesis: ( 1 <= i & i < j & j <= len p & l = p . i & m = p . j implies l < m )
assume that
A34: 1 <= i and
A35: i < j and
A36: j <= len p and
A37: l = p . i and
A38: m = p . j ; :: thesis: l < m
i <= len p by A35, A36, XXREAL_0:2;
then i in dom p by A6, A34, A33, FINSEQ_1:1;
then A39: p . i = (Sgm (Seg (k + 1))) . i by FUNCT_1:47;
1 <= j by A34, A35, XXREAL_0:2;
then j in dom p by A6, A36, A33, FINSEQ_1:1;
then A40: p . j = (Sgm (Seg (k + 1))) . j by FUNCT_1:47;
len (Sgm (Seg (k + 1))) = k + 1 by Th44;
then j <= len (Sgm (Seg (k + 1))) by A36, A33, NAT_1:12;
hence l < m by A34, A35, A37, A38, A39, A40, FINSEQ_1:def 13; :: thesis: verum
end;
hence (Sgm (Seg (k + 1))) | (Seg k) = Sgm (Seg k) by A32, FINSEQ_1:def 13; :: thesis: verum
end;
Sgm (Seg (k + (n + 1))) = Sgm (Seg ((k + 1) + n)) ;
then (Sgm (Seg (k + (n + 1)))) | (Seg (k + 1)) = Sgm (Seg (k + 1)) by A1;
then Sgm (Seg k) = (Sgm (Seg (k + (n + 1)))) | ((Seg (k + 1)) /\ (Seg k)) by A2, RELAT_1:71;
hence (Sgm (Seg (k + (n + 1)))) | (Seg k) = Sgm (Seg k) by FINSEQ_1:7, NAT_1:12; :: thesis: verum
end;

Lm4: for n, k being Nat holds (Sgm (Seg (k + n))) | (Seg k) = Sgm (Seg k)
proof end;

theorem :: FINSEQ_3:47
for k, n being Nat holds (Sgm (Seg (k + n))) | (Seg k) = Sgm (Seg k) by Lm4;

Lm5: now :: thesis: for k being Nat st Sgm (Seg k) = idseq k holds
Sgm (Seg (k + 1)) = idseq (k + 1)
let k be Nat; :: thesis: ( Sgm (Seg k) = idseq k implies Sgm (Seg (k + 1)) = idseq (k + 1) )
assume A1: Sgm (Seg k) = idseq k ; :: thesis: Sgm (Seg (k + 1)) = idseq (k + 1)
A2: len (idseq (k + 1)) = k + 1 by CARD_1:def 7;
then A3: len (Sgm (Seg (k + 1))) = len (idseq (k + 1)) by Th44;
then A4: dom (Sgm (Seg (k + 1))) = Seg (k + 1) by A2, FINSEQ_1:def 3;
now :: thesis: for j being Nat st j in dom (Sgm (Seg (k + 1))) holds
(Sgm (Seg (k + 1))) . j = (idseq (k + 1)) . j
let j be Nat; :: thesis: ( j in dom (Sgm (Seg (k + 1))) implies (Sgm (Seg (k + 1))) . j = (idseq (k + 1)) . j )
assume A5: j in dom (Sgm (Seg (k + 1))) ; :: thesis: (Sgm (Seg (k + 1))) . j = (idseq (k + 1)) . j
then A6: j in (Seg k) \/ {(k + 1)} by A4, FINSEQ_1:9;
now :: thesis: (Sgm (Seg (k + 1))) . j = (idseq (k + 1)) . j
per cases ( j in Seg k or j in {(k + 1)} ) by A6, XBOOLE_0:def 3;
suppose A7: j in Seg k ; :: thesis: (Sgm (Seg (k + 1))) . j = (idseq (k + 1)) . j
A8: (idseq (k + 1)) . j = j by A4, A5, FUNCT_1:18;
A9: (Sgm (Seg (k + 1))) | (Seg k) = Sgm (Seg k) by Lm4;
(Sgm (Seg k)) . j = j by A1, A7, FUNCT_1:18;
hence (Sgm (Seg (k + 1))) . j = (idseq (k + 1)) . j by A7, A8, A9, FUNCT_1:49; :: thesis: verum
end;
suppose A10: j in {(k + 1)} ; :: thesis: (Sgm (Seg (k + 1))) . j = (idseq (k + 1)) . j
set Y = Sgm (Seg k);
set X = Sgm (Seg (k + 1));
A11: j = k + 1 by A10, TARSKI:def 1;
then A12: j in Seg (k + 1) by FINSEQ_1:4;
now :: thesis: not (Sgm (Seg (k + 1))) . j <> j
rng (Sgm (Seg (k + 1))) = Seg (k + 1) by FINSEQ_1:def 13;
then A13: (Sgm (Seg (k + 1))) . j in Seg (k + 1) by A5, FUNCT_1:def 3;
then reconsider n = (Sgm (Seg (k + 1))) . j as Element of NAT ;
assume (Sgm (Seg (k + 1))) . j <> j ; :: thesis: contradiction
then not (Sgm (Seg (k + 1))) . j in {j} by TARSKI:def 1;
then (Sgm (Seg (k + 1))) . j in (Seg (k + 1)) \ {(k + 1)} by A11, A13, XBOOLE_0:def 5;
then A14: (Sgm (Seg (k + 1))) . j in Seg k by FINSEQ_1:10;
A15: Sgm (Seg (k + 1)) is one-to-one by Lm1;
A16: dom (Sgm (Seg (k + 1))) = Seg (k + 1) by A2, A3, FINSEQ_1:def 3;
A17: k < k + 1 by XREAL_1:29;
(Sgm (Seg (k + 1))) | (Seg k) = Sgm (Seg k) by Lm4;
then A18: (Sgm (Seg (k + 1))) . n = (Sgm (Seg k)) . n by A14, FUNCT_1:49
.= n by A1, A14, FUNCT_1:18 ;
n <= k by A14, FINSEQ_1:1;
hence contradiction by A11, A12, A16, A13, A15, A18, A17; :: thesis: verum
end;
hence (Sgm (Seg (k + 1))) . j = (idseq (k + 1)) . j by A11, FINSEQ_1:4, FUNCT_1:18; :: thesis: verum
end;
end;
end;
hence (Sgm (Seg (k + 1))) . j = (idseq (k + 1)) . j ; :: thesis: verum
end;
hence Sgm (Seg (k + 1)) = idseq (k + 1) by A3, FINSEQ_2:9; :: thesis: verum
end;

theorem Th46: :: FINSEQ_3:48
for k being Nat holds Sgm (Seg k) = idseq k
proof end;

theorem Th47: :: FINSEQ_3:49
for p being FinSequence
for n being Nat holds
( p | (Seg n) = p iff len p <= n )
proof end;

theorem Th48: :: FINSEQ_3:50
for k, n being Nat holds (idseq (n + k)) | (Seg n) = idseq n
proof end;

theorem :: FINSEQ_3:51
for m, n being Nat holds
( (idseq n) | (Seg m) = idseq m iff m <= n )
proof end;

theorem :: FINSEQ_3:52
for m, n being Nat holds
( (idseq n) | (Seg m) = idseq n iff n <= m )
proof end;

theorem Th51: :: FINSEQ_3:53
for p, q being FinSequence
for k, l being Nat st len p = k + l & q = p | (Seg k) holds
len q = k
proof end;

theorem :: FINSEQ_3:54
for p, q being FinSequence
for k, l being Nat st len p = k + l & q = p | (Seg k) holds
dom q = Seg k
proof end;

theorem Th53: :: FINSEQ_3:55
for p, q being FinSequence
for k being Nat st len p = k + 1 & q = p | (Seg k) holds
p = q ^ <*(p . (k + 1))*>
proof end;

theorem :: FINSEQ_3:56
for p being FinSequence
for X being set holds
( p | X is FinSequence iff ex k being Element of NAT st X /\ (dom p) = Seg k )
proof end;

theorem Th55: :: FINSEQ_3:57
for p, q being FinSequence
for A being set holds card ((p ^ q) " A) = (card (p " A)) + (card (q " A))
proof end;

theorem Th56: :: FINSEQ_3:58
for p, q being FinSequence
for A being set holds p " A c= (p ^ q) " A
proof end;

definition
let p be FinSequence;
let A be set ;
func p - A -> FinSequence equals :: FINSEQ_3:def 1
p * (Sgm ((dom p) \ (p " A)));
coherence
p * (Sgm ((dom p) \ (p " A))) is FinSequence
proof end;
end;

:: deftheorem defines - FINSEQ_3:def 1 :
for p being FinSequence
for A being set holds p - A = p * (Sgm ((dom p) \ (p " A)));

theorem Th57: :: FINSEQ_3:59
for p being FinSequence
for A being set holds len (p - A) = (len p) - (card (p " A))
proof end;

theorem Th58: :: FINSEQ_3:60
for p being FinSequence
for A being set holds len (p - A) <= len p
proof end;

theorem Th59: :: FINSEQ_3:61
for p being FinSequence
for A being set st len (p - A) = len p holds
A misses rng p
proof end;

theorem :: FINSEQ_3:62
for p being FinSequence
for A being set
for n being Nat st n = (len p) - (card (p " A)) holds
dom (p - A) = Seg n
proof end;

theorem :: FINSEQ_3:63
for p being FinSequence
for A being set holds dom (p - A) c= dom p
proof end;

theorem :: FINSEQ_3:64
for p being FinSequence
for A being set st dom (p - A) = dom p holds
A misses rng p
proof end;

theorem Th63: :: FINSEQ_3:65
for p being FinSequence
for A being set holds rng (p - A) = (rng p) \ A
proof end;

theorem :: FINSEQ_3:66
for p being FinSequence
for A being set holds rng (p - A) c= rng p
proof end;

theorem :: FINSEQ_3:67
for p being FinSequence
for A being set st rng (p - A) = rng p holds
A misses rng p
proof end;

theorem Th66: :: FINSEQ_3:68
for p being FinSequence
for A being set holds
( p - A = {} iff rng p c= A )
proof end;

theorem Th67: :: FINSEQ_3:69
for p being FinSequence
for A being set holds
( p - A = p iff A misses rng p )
proof end;

theorem :: FINSEQ_3:70
for p being FinSequence
for x being object holds
( p - {x} = p iff not x in rng p )
proof end;

theorem :: FINSEQ_3:71
for p being FinSequence holds p - {} = p
proof end;

theorem :: FINSEQ_3:72
for p being FinSequence holds p - (rng p) = {} by Th66;

Lm6: for x being object
for A being set holds
( <*x*> - A = <*x*> iff not x in A )

proof end;

Lm7: for x being object
for A being set holds
( <*x*> - A = {} iff x in A )

proof end;

Lm8: for p being FinSequence
for A being set holds (p ^ {}) - A = (p - A) ^ ({} - A)

proof end;

Lm9: for p being FinSequence
for x being object
for A being set holds (p ^ <*x*>) - A = (p - A) ^ (<*x*> - A)

proof end;

Lm10: now :: thesis: for q being FinSequence
for x being object st ( for p being FinSequence
for A being set holds (p ^ q) - A = (p - A) ^ (q - A) ) holds
for p being FinSequence
for A being set holds (p ^ (q ^ <*x*>)) - A = (p - A) ^ ((q ^ <*x*>) - A)
let q be FinSequence; :: thesis: for x being object st ( for p being FinSequence
for A being set holds (p ^ q) - A = (p - A) ^ (q - A) ) holds
for p being FinSequence
for A being set holds (p ^ (q ^ <*x*>)) - A = (p - A) ^ ((q ^ <*x*>) - A)

let x be object ; :: thesis: ( ( for p being FinSequence
for A being set holds (p ^ q) - A = (p - A) ^ (q - A) ) implies for p being FinSequence
for A being set holds (p ^ (q ^ <*x*>)) - A = (p - A) ^ ((q ^ <*x*>) - A) )

assume A1: for p being FinSequence
for A being set holds (p ^ q) - A = (p - A) ^ (q - A) ; :: thesis: for p being FinSequence
for A being set holds (p ^ (q ^ <*x*>)) - A = (p - A) ^ ((q ^ <*x*>) - A)

let p be FinSequence; :: thesis: for A being set holds (p ^ (q ^ <*x*>)) - A = (p - A) ^ ((q ^ <*x*>) - A)
let A be set ; :: thesis: (p ^ (q ^ <*x*>)) - A = (p - A) ^ ((q ^ <*x*>) - A)
thus (p ^ (q ^ <*x*>)) - A = ((p ^ q) ^ <*x*>) - A by FINSEQ_1:32
.= ((p ^ q) - A) ^ (<*x*> - A) by Lm9
.= ((p - A) ^ (q - A)) ^ (<*x*> - A) by A1
.= (p - A) ^ ((q - A) ^ (<*x*> - A)) by FINSEQ_1:32
.= (p - A) ^ ((q ^ <*x*>) - A) by Lm9 ; :: thesis: verum
end;

Lm11: for q, p being FinSequence
for A being set holds (p ^ q) - A = (p - A) ^ (q - A)

proof end;

theorem :: FINSEQ_3:73
for p, q being FinSequence
for A being set holds (p ^ q) - A = (p - A) ^ (q - A) by Lm11;

theorem :: FINSEQ_3:74
for A being set holds {} - A = {} ;

theorem :: FINSEQ_3:75
for x being object
for A being set holds
( <*x*> - A = <*x*> iff not x in A ) by Lm6;

theorem :: FINSEQ_3:76
for x being object
for A being set holds
( <*x*> - A = {} iff x in A ) by Lm7;

theorem Th75: :: FINSEQ_3:77
for x, y being object
for A being set holds
( <*x,y*> - A = {} iff ( x in A & y in A ) )
proof end;

theorem Th76: :: FINSEQ_3:78
for x, y being object
for A being set st x in A & not y in A holds
<*x,y*> - A = <*y*>
proof end;

theorem :: FINSEQ_3:79
for x, y being object
for A being set st <*x,y*> - A = <*y*> & x <> y holds
( x in A & not y in A )
proof end;

theorem Th78: :: FINSEQ_3:80
for x, y being object
for A being set st not x in A & y in A holds
<*x,y*> - A = <*x*>
proof end;

theorem :: FINSEQ_3:81
for x, y being object
for A being set st <*x,y*> - A = <*x*> & x <> y holds
( not x in A & y in A )
proof end;

theorem :: FINSEQ_3:82
for x, y being object
for A being set holds
( <*x,y*> - A = <*x,y*> iff ( not x in A & not y in A ) )
proof end;

theorem Th81: :: FINSEQ_3:83
for p, q being FinSequence
for A being set
for k being Nat st len p = k + 1 & q = p | (Seg k) holds
( p . (k + 1) in A iff p - A = q - A )
proof end;

theorem Th82: :: FINSEQ_3:84
for p, q being FinSequence
for A being set
for k being Nat st len p = k + 1 & q = p | (Seg k) holds
( not p . (k + 1) in A iff p - A = (q - A) ^ <*(p . (k + 1))*> )
proof end;

Lm12: for l being Nat st ( for p being FinSequence
for A being set st len p = l holds
for n being Nat st n in dom p holds
for B being finite set holds
( not B = { k where k is Element of NAT : ( k in dom p & k <= n & p . k in A ) } or p . n in A or (p - A) . (n - (card B)) = p . n ) ) holds
for p being FinSequence
for A being set st len p = l + 1 holds
for n being Nat st n in dom p holds
for C being finite set holds
( not C = { m where m is Element of NAT : ( m in dom p & m <= n & p . m in A ) } or p . n in A or (p - A) . (n - (card C)) = p . n )

proof end;

Lm13: for l being Nat
for p being FinSequence
for A being set st len p = l holds
for n being Nat st n in dom p holds
for B being finite set holds
( not B = { k where k is Element of NAT : ( k in dom p & k <= n & p . k in A ) } or p . n in A or (p - A) . (n - (card B)) = p . n )

proof end;

theorem :: FINSEQ_3:85
for p being FinSequence
for A being set
for n being Nat st n in dom p holds
for B being finite set holds
( not B = { k where k is Element of NAT : ( k in dom p & k <= n & p . k in A ) } or p . n in A or (p - A) . (n - (card B)) = p . n )
proof end;

theorem :: FINSEQ_3:86
for p being FinSequence
for A, D being set st p is FinSequence of D holds
p - A is FinSequence of D
proof end;

theorem :: FINSEQ_3:87
for p being FinSequence
for A being set st p is one-to-one holds
p - A is one-to-one
proof end;

theorem Th86: :: FINSEQ_3:88
for p being FinSequence
for A being set st p is one-to-one holds
len (p - A) = (len p) - (card (A /\ (rng p)))
proof end;

theorem Th87: :: FINSEQ_3:89
for p being FinSequence
for A being finite set st p is one-to-one & A c= rng p holds
len (p - A) = (len p) - (card A)
proof end;

theorem :: FINSEQ_3:90
for p being FinSequence
for x being object st p is one-to-one & x in rng p holds
len (p - {x}) = (len p) - 1
proof end;

theorem Th89: :: FINSEQ_3:91
for p, q being FinSequence holds
( ( rng p misses rng q & p is one-to-one & q is one-to-one ) iff p ^ q is one-to-one )
proof end;

theorem Th90: :: FINSEQ_3:92
for A being set
for k being Nat st A c= Seg k holds
Sgm A is one-to-one by Lm1;

theorem Th91: :: FINSEQ_3:93
for x being object holds <*x*> is one-to-one
proof end;

theorem Th92: :: FINSEQ_3:94
for x, y being object holds
( x <> y iff <*x,y*> is one-to-one )
proof end;

theorem Th93: :: FINSEQ_3:95
for x, y, z being object holds
( ( x <> y & y <> z & z <> x ) iff <*x,y,z*> is one-to-one )
proof end;

theorem Th94: :: FINSEQ_3:96
for p being FinSequence
for x being object st p is one-to-one & rng p = {x} holds
len p = 1
proof end;

theorem :: FINSEQ_3:97
for p being FinSequence
for x being object st p is one-to-one & rng p = {x} holds
p = <*x*>
proof end;

theorem Th96: :: FINSEQ_3:98
for p being FinSequence
for x, y being object st p is one-to-one & rng p = {x,y} & x <> y holds
len p = 2
proof end;

theorem :: FINSEQ_3:99
for p being FinSequence
for x, y being object st p is one-to-one & rng p = {x,y} & x <> y & not p = <*x,y*> holds
p = <*y,x*>
proof end;

theorem Th98: :: FINSEQ_3:100
for p being FinSequence
for x, y, z being object st p is one-to-one & rng p = {x,y,z} & <*x,y,z*> is one-to-one holds
len p = 3
proof end;

theorem :: FINSEQ_3:101
for p being FinSequence
for x, y, z being object st p is one-to-one & rng p = {x,y,z} & x <> y & y <> z & x <> z holds
len p = 3
proof end;

:: from FSM_1
theorem :: FINSEQ_3:102
for D being non empty set
for df being FinSequence of D st not df is empty holds
ex d being Element of D ex df1 being FinSequence of D st
( d = df . 1 & df = <*d*> ^ df1 )
proof end;

theorem :: FINSEQ_3:103
for i being Nat
for df being FinSequence
for d being object st i in dom df holds
(<*d*> ^ df) . (i + 1) = df . i
proof end;

:: from MATRIX_2, 2005.11.16, A.T.
definition
let i be natural Number ;
let p be FinSequence;
func Del (p,i) -> FinSequence equals :: FINSEQ_3:def 2
p * (Sgm ((dom p) \ {i}));
coherence
p * (Sgm ((dom p) \ {i})) is FinSequence
proof end;
end;

:: deftheorem defines Del FINSEQ_3:def 2 :
for i being natural Number
for p being FinSequence holds Del (p,i) = p * (Sgm ((dom p) \ {i}));

theorem Th102: :: FINSEQ_3:104
for i being Nat
for p being FinSequence holds
( ( i in dom p implies ex m being Nat st
( len p = m + 1 & len (Del (p,i)) = m ) ) & ( not i in dom p implies Del (p,i) = p ) )
proof end;

theorem :: FINSEQ_3:105
for i being Nat
for D being non empty set
for p being FinSequence of D holds Del (p,i) is FinSequence of D
proof end;

:: from MATRLIN, 2005.11.16, A.T.
theorem :: FINSEQ_3:106
for i being Nat
for p being FinSequence holds rng (Del (p,i)) c= rng p by FUNCT_1:14;

:: from GOBOARD1, 2005.11.16, A.T.
theorem Th105: :: FINSEQ_3:107
for i, m, n being Nat st n = m + 1 & i in Seg n holds
len (Sgm ((Seg n) \ {i})) = m
proof end;

theorem Th106: :: FINSEQ_3:108
for i, k, m, n being Nat st n = m + 1 & k in Seg n & i in Seg m holds
( ( 1 <= i & i < k implies (Sgm ((Seg n) \ {k})) . i = i ) & ( k <= i & i <= m implies (Sgm ((Seg n) \ {k})) . i = i + 1 ) )
proof end;

theorem Th107: :: FINSEQ_3:109
for m, n being Nat
for f being FinSequence st len f = m + 1 & n in dom f holds
len (Del (f,n)) = m
proof end;

theorem :: FINSEQ_3:110
for k, n being Nat
for f being FinSequence st k < n holds
(Del (f,n)) . k = f . k
proof end;

theorem :: FINSEQ_3:111
for k, m, n being Nat
for f being FinSequence st len f = m + 1 & n in dom f & n <= k & k <= m holds
(Del (f,n)) . k = f . (k + 1)
proof end;

:: from JORDAN3, 2005.11.16, A.T.
theorem Th110: :: FINSEQ_3:112
for k, n being Nat
for f being FinSequence st k <= n holds
(f | n) . k = f . k
proof end;

:: missing, 27.11.2005, A.T.
theorem :: FINSEQ_3:113
for p, q being FinSequence st p c= q holds
q | (len p) = p
proof end;

:: from MATRLIN, 2006.01.06, A.T.
theorem Th112: :: FINSEQ_3:114
for A being set
for F being FinSequence holds (Sgm (F " A)) ^ (Sgm (F " ((rng F) \ A))) is Permutation of (dom F)
proof end;

theorem :: FINSEQ_3:115
for F being FinSequence
for A being Subset of (rng F) ex p being Permutation of (dom F) st (F - (A `)) ^ (F - A) = F * p
proof end;

:: Moved from GRAPH_2 by AK on 27.12.2006
theorem :: FINSEQ_3:116
for f being FinSubsequence st f is FinSequence holds
Seq f = f
proof end;

:: Moved from SCMBSORT by AK on 10.10.2007
theorem :: FINSEQ_3:117
for t being FinSequence of INT holds t is FinSequence of REAL
proof end;

:: missing, 2008.02.16, A.T,
theorem :: FINSEQ_3:118
for p, q being FinSequence holds
( len p < len q iff dom p c< dom q )
proof end;

:: from CATALG_1, 2009.09.08, A.T.
theorem Th117: :: FINSEQ_3:119
for A being set
for i being Nat holds
( ( i <> 0 & A = {} ) iff i -tuples_on A = {} )
proof end;

:: from AMISTD_2, 2009.09.08, A.T.
registration
let i be Nat;
let D be set ;
cluster i -tuples_on D -> with_common_domain ;
coherence
i -tuples_on D is with_common_domain
proof end;
end;

registration
let i be Nat;
let D be set ;
cluster i -tuples_on D -> product-like ;
coherence
i -tuples_on D is product-like
proof end;
end;

theorem :: FINSEQ_3:120
for D1, D2 being non empty set
for p being FinSequence of D1
for f being Function of D1,D2 holds
( dom (f * p) = dom p & len (f * p) = len p & ( for n being Nat st n in dom (f * p) holds
(f * p) . n = f . (p . n) ) )
proof end;

definition
let D be non empty set ;
let R be Relation of D;
func ExtendRel R -> Relation of (D *) means :Def3: :: FINSEQ_3:def 3
for x, y being FinSequence of D holds
( [x,y] in it iff ( len x = len y & ( for n being Nat st n in dom x holds
[(x . n),(y . n)] in R ) ) );
existence
ex b1 being Relation of (D *) st
for x, y being FinSequence of D holds
( [x,y] in b1 iff ( len x = len y & ( for n being Nat st n in dom x holds
[(x . n),(y . n)] in R ) ) )
proof end;
uniqueness
for b1, b2 being Relation of (D *) st ( for x, y being FinSequence of D holds
( [x,y] in b1 iff ( len x = len y & ( for n being Nat st n in dom x holds
[(x . n),(y . n)] in R ) ) ) ) & ( for x, y being FinSequence of D holds
( [x,y] in b2 iff ( len x = len y & ( for n being Nat st n in dom x holds
[(x . n),(y . n)] in R ) ) ) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def3 defines ExtendRel FINSEQ_3:def 3 :
for D being non empty set
for R being Relation of D
for b3 being Relation of (D *) holds
( b3 = ExtendRel R iff for x, y being FinSequence of D holds
( [x,y] in b3 iff ( len x = len y & ( for n being Nat st n in dom x holds
[(x . n),(y . n)] in R ) ) ) );

theorem :: FINSEQ_3:121
for D being non empty set holds ExtendRel (id D) = id (D *)
proof end;

definition
let D be non empty set ;
let R be Equivalence_Relation of D;
let y be FinSequence of Class R;
let x be FinSequence of D;
pred x is_representatives_FS y means :: FINSEQ_3:def 4
( len x = len y & ( for n being Nat st n in dom x holds
Class (R,(x . n)) = y . n ) );
end;

:: deftheorem defines is_representatives_FS FINSEQ_3:def 4 :
for D being non empty set
for R being Equivalence_Relation of D
for y being FinSequence of Class R
for x being FinSequence of D holds
( x is_representatives_FS y iff ( len x = len y & ( for n being Nat st n in dom x holds
Class (R,(x . n)) = y . n ) ) );

theorem :: FINSEQ_3:122
for D being non empty set
for R being Equivalence_Relation of D
for y being FinSequence of Class R ex x being FinSequence of D st x is_representatives_FS y
proof end;

theorem Th121: :: FINSEQ_3:123
for x being object
for X being set holds
( x in product <*X*> iff ex y being object st
( y in X & x = <*y*> ) )
proof end;

theorem Th122: :: FINSEQ_3:124
for z being object
for X, Y being set holds
( z in product <*X,Y*> iff ex x, y being object st
( x in X & y in Y & z = <*x,y*> ) )
proof end;

theorem Th123: :: FINSEQ_3:125
for a being object
for X, Y, Z being set holds
( a in product <*X,Y,Z*> iff ex x, y, z being object st
( x in X & y in Y & z in Z & a = <*x,y,z*> ) )
proof end;

theorem :: FINSEQ_3:126
for D being non empty set holds product <*D*> = 1 -tuples_on D
proof end;

theorem Th125: :: FINSEQ_3:127
for D1, D2 being non empty set holds product <*D1,D2*> = { <*d1,d2*> where d1 is Element of D1, d2 is Element of D2 : verum }
proof end;

theorem :: FINSEQ_3:128
for D being non empty set holds product <*D,D*> = 2 -tuples_on D
proof end;

theorem Th127: :: FINSEQ_3:129
for D1, D2, D3 being non empty set holds product <*D1,D2,D3*> = { <*d1,d2,d3*> where d1 is Element of D1, d2 is Element of D2, d3 is Element of D3 : verum }
proof end;

theorem :: FINSEQ_3:130
for D being non empty set holds product <*D,D,D*> = 3 -tuples_on D
proof end;

theorem Th129: :: FINSEQ_3:131
for i being Nat
for D being set holds product (i |-> D) = i -tuples_on D
proof end;

registration
let f be Function;
cluster <*f*> -> Function-yielding ;
coherence
<*f*> is Function-yielding
proof end;
let g be Function;
cluster <*f,g*> -> Function-yielding ;
coherence
<*f,g*> is Function-yielding
proof end;
let h be Function;
cluster <*f,g,h*> -> Function-yielding ;
coherence
<*f,g,h*> is Function-yielding
proof end;
end;

theorem Th130: :: FINSEQ_3:132
for f being Function holds
( doms <*f*> = <*(dom f)*> & rngs <*f*> = <*(rng f)*> )
proof end;

theorem Th131: :: FINSEQ_3:133
for f, g being Function holds
( doms <*f,g*> = <*(dom f),(dom g)*> & rngs <*f,g*> = <*(rng f),(rng g)*> )
proof end;

theorem :: FINSEQ_3:134
for f, g, h being Function holds
( doms <*f,g,h*> = <*(dom f),(dom g),(dom h)*> & rngs <*f,g,h*> = <*(rng f),(rng g),(rng h)*> )
proof end;

theorem Th133: :: FINSEQ_3:135
for X being set holds
( Union <*X*> = X & meet <*X*> = X )
proof end;

theorem Th134: :: FINSEQ_3:136
for X, Y being set holds
( Union <*X,Y*> = X \/ Y & meet <*X,Y*> = X /\ Y )
proof end;

theorem :: FINSEQ_3:137
for X, Y, Z being set holds
( Union <*X,Y,Z*> = (X \/ Y) \/ Z & meet <*X,Y,Z*> = (X /\ Y) /\ Z )
proof end;

theorem :: FINSEQ_3:138
for x being object
for f, g, h being Function st x in dom f holds
( <*f*> .. (1,x) = f . x & <*f,g*> .. (1,x) = f . x & <*f,g,h*> .. (1,x) = f . x )
proof end;

theorem :: FINSEQ_3:139
for x being object
for f, g, h being Function st x in dom g holds
( <*f,g*> .. (2,x) = g . x & <*f,g,h*> .. (2,x) = g . x )
proof end;

theorem :: FINSEQ_3:140
for x being object
for f, g, h being Function st x in dom h holds
<*f,g,h*> .. (3,x) = h . x
proof end;

theorem :: FINSEQ_3:141
for h being Function holds
( dom <:<*h*>:> = dom h & ( for x being object st x in dom h holds
<:<*h*>:> . x = <*(h . x)*> ) )
proof end;

theorem Th140: :: FINSEQ_3:142
for f1, f2 being Function holds
( dom <:<*f1,f2*>:> = (dom f1) /\ (dom f2) & ( for x being object st x in (dom f1) /\ (dom f2) holds
<:<*f1,f2*>:> . x = <*(f1 . x),(f2 . x)*> ) )
proof end;

theorem :: FINSEQ_3:143
for h being Function holds
( dom (Frege <*h*>) = product <*(dom h)*> & rng (Frege <*h*>) = product <*(rng h)*> & ( for x being object st x in dom h holds
(Frege <*h*>) . <*x*> = <*(h . x)*> ) )
proof end;

theorem Th142: :: FINSEQ_3:144
for f1, f2 being Function holds
( dom (Frege <*f1,f2*>) = product <*(dom f1),(dom f2)*> & rng (Frege <*f1,f2*>) = product <*(rng f1),(rng f2)*> & ( for x, y being object st x in dom f1 & y in dom f2 holds
(Frege <*f1,f2*>) . <*x,y*> = <*(f1 . x),(f2 . y)*> ) )
proof end;

theorem :: FINSEQ_3:145
for x being object
for f1, f2 being Function st x in dom f1 & x in dom f2 holds
for y1, y2 being object holds
( <:f1,f2:> . x = [y1,y2] iff <:<*f1,f2*>:> . x = <*y1,y2*> )
proof end;

theorem :: FINSEQ_3:146
for x, y being object
for f1, f2 being Function st x in dom f1 & y in dom f2 holds
for y1, y2 being object holds
( [:f1,f2:] . (x,y) = [y1,y2] iff (Frege <*f1,f2*>) . <*x,y*> = <*y1,y2*> )
proof end;

theorem :: FINSEQ_3:147
for X, Y being set holds Funcs (<*X*>,Y) = <*(Funcs (X,Y))*>
proof end;

theorem :: FINSEQ_3:148
for X, Y, Z being set holds Funcs (<*X,Y*>,Z) = <*(Funcs (X,Z)),(Funcs (Y,Z))*>
proof end;

theorem :: FINSEQ_3:149
for X, Y being set holds Funcs (X,<*Y*>) = <*(Funcs (X,Y))*>
proof end;

theorem :: FINSEQ_3:150
for X, Y, Z being set holds Funcs (X,<*Y,Z*>) = <*(Funcs (X,Y)),(Funcs (X,Z))*>
proof end;

::from JORDAN2C, 2011.07.03, A.T.
theorem :: FINSEQ_3:151
for x, y being object
for f being FinSequence st rng f = {x,y} & len f = 2 & not ( f . 1 = x & f . 2 = y ) holds
( f . 1 = y & f . 2 = x )
proof end;

:: from GLIB_001, 2011.07.30, A.T.
theorem :: FINSEQ_3:152
for X being set
for k being Element of NAT st X c= Seg k holds
for m, n being Element of NAT st m in dom (Sgm X) & n = (Sgm X) . m holds
m <= n
proof end;

registration
let i be Nat;
let D be finite set ;
cluster i -tuples_on D -> finite ;
coherence
i -tuples_on D is finite
proof end;
end;

theorem Th151: :: FINSEQ_3:153
for m being Nat
for p being b1 -element FinSequence holds len p = m by CARD_1:def 7;

theorem :: FINSEQ_3:154
for n being Nat
for p being b1 -element FinSequence
for q being FinSequence holds (p ^ q) . 1 = p . 1 & ... & (p ^ q) . n = p . n
proof end;

theorem :: FINSEQ_3:155
for m, n being Nat
for p being b2 -element FinSequence
for q being b1 -element FinSequence holds (p ^ q) . (n + 1) = q . 1 & ... & (p ^ q) . (n + m) = q . m
proof end;

theorem :: FINSEQ_3:156
for p being FinSequence
for k being Nat st k in dom p holds
for i being Nat st 1 <= i & i <= k holds
i in dom p
proof end;

:: from PNPROC_1, 2012.02.20, A.T.
theorem :: FINSEQ_3:157
for x being object
for i being Nat
for q being FinSubsequence st q = {[i,x]} holds
Seq q = <*x*>
proof end;

theorem :: FINSEQ_3:158
for p being FinSubsequence holds card p = len (Seq p)
proof end;

theorem :: FINSEQ_3:159
for x being object
for q being FinSubsequence st Seq q = <*x*> holds
ex i being Element of NAT st q = {[i,x]}
proof end;