begin
Lm1:
for x, y being set holds rng <*x,y*> = {x,y}
Lm2:
for x, y, z being set holds rng <*x,y,z*> = {x,y,z}
begin
theorem
canceled;
theorem
canceled;
theorem Th3:
theorem Th4:
theorem Th5:
theorem Th6:
theorem Th7:
theorem Th8:
theorem Th9:
theorem Th10:
theorem Th11:
theorem Th12:
theorem
theorem Th14:
theorem Th15:
theorem Th16:
theorem Th17:
theorem
canceled;
theorem Th19:
theorem
theorem
canceled;
theorem Th22:
theorem Th23:
theorem Th24:
theorem Th25:
for
p1,
p2,
p3 being
set holds
p1 .. <*p1,p2,p3*> = 1
theorem Th26:
for
p1,
p2,
p3 being
set st
p1 <> p2 holds
p2 .. <*p1,p2,p3*> = 2
theorem Th27:
for
p1,
p3,
p2 being
set st
p1 <> p3 &
p2 <> p3 holds
p3 .. <*p1,p2,p3*> = 3
theorem Th28:
theorem Th29:
theorem Th30:
theorem Th31:
theorem Th32:
theorem
theorem Th34:
theorem
theorem Th36:
theorem Th37:
theorem Th38:
theorem Th39:
theorem Th40:
theorem Th41:
theorem Th42:
theorem
begin
theorem Th44:
theorem Th45:
theorem Th46:
theorem Th47:
theorem Th48:
theorem Th49:
theorem Th50:
theorem Th51:
theorem Th52:
theorem Th53:
theorem Th54:
theorem Th55:
theorem
theorem Th57:
theorem Th58:
theorem Th59:
theorem
canceled;
theorem Th61:
theorem Th62:
theorem Th63:
theorem Th64:
theorem Th65:
theorem Th66:
theorem Th67:
theorem Th68:
theorem Th69:
theorem Th70:
theorem Th71:
theorem Th72:
theorem
theorem Th74:
theorem Th75:
theorem Th76:
theorem Th77:
theorem Th78:
theorem Th79:
theorem Th80:
theorem Th81:
theorem Th82:
theorem Th83:
theorem Th84:
Lm3:
for i being Nat
for D being non empty set
for p being Element of D
for f being FinSequence of D st p in rng f & p .. f > i holds
i + (p .. (f /^ i)) = p .. f
theorem
canceled;
theorem Th86:
theorem Th87:
theorem Th88:
theorem Th89:
theorem Th90:
theorem Th91:
theorem Th92:
theorem Th93:
theorem Th94:
begin
:: deftheorem Def1 defines circular FINSEQ_6:def 1 :
for D being non empty set
for IT being FinSequence of D holds
( IT is circular iff IT /. 1 = IT /. (len IT) );
:: deftheorem Def2 defines Rotate FINSEQ_6:def 2 :
for D being non empty set
for f being FinSequence of D
for p being Element of D holds
( ( p in rng f implies Rotate (f,p) = (f :- p) ^ ((f -: p) /^ 1) ) & ( not p in rng f implies Rotate (f,p) = f ) );
theorem Th95:
theorem
theorem Th97:
theorem Th98:
theorem Th99:
theorem
theorem Th101:
theorem
theorem Th103:
theorem
theorem
theorem
theorem Th107:
theorem Th108:
theorem Th109:
theorem Th110:
theorem
theorem
begin
theorem
begin
theorem
theorem
theorem Th18:
theorem
theorem
:: deftheorem Def1 defines mid FINSEQ_6:def 3 :
for f being FinSequence
for k1, k2 being Nat holds
( ( k1 <= k2 implies mid (f,k1,k2) = (f /^ (k1 -' 1)) | ((k2 -' k1) + 1) ) & ( not k1 <= k2 implies mid (f,k1,k2) = Rev ((f /^ (k2 -' 1)) | ((k1 -' k2) + 1)) ) );
theorem
theorem Th23:
theorem Th24:
theorem
theorem
theorem Th27:
for
D being non
empty set for
f being
FinSequence of
D for
k1,
k2 being
Element of
NAT st 1
<= k1 &
k1 <= len f & 1
<= k2 &
k2 <= len f holds
(
(mid (f,k1,k2)) . 1
= f . k1 & (
k1 <= k2 implies (
len (mid (f,k1,k2)) = (k2 -' k1) + 1 & ( for
i being
Element of
NAT st 1
<= i &
i <= len (mid (f,k1,k2)) holds
(mid (f,k1,k2)) . i = f . ((i + k1) -' 1) ) ) ) & (
k1 > k2 implies (
len (mid (f,k1,k2)) = (k1 -' k2) + 1 & ( for
i being
Element of
NAT st 1
<= i &
i <= len (mid (f,k1,k2)) holds
(mid (f,k1,k2)) . i = f . ((k1 -' i) + 1) ) ) ) )
theorem
theorem
theorem
theorem Th31:
for
D being non
empty set for
f being
FinSequence of
D for
k1,
k2,
i being
Element of
NAT st 1
<= k1 &
k1 <= k2 &
k2 <= len f & 1
<= i & (
i <= (k2 -' k1) + 1 or
i <= (k2 - k1) + 1 or
i <= (k2 + 1) - k1 ) holds
(
(mid (f,k1,k2)) . i = f . ((i + k1) -' 1) &
(mid (f,k1,k2)) . i = f . ((i -' 1) + k1) &
(mid (f,k1,k2)) . i = f . ((i + k1) - 1) &
(mid (f,k1,k2)) . i = f . ((i - 1) + k1) )
theorem
theorem
theorem Th55:
theorem