begin
theorem Th1:
for
m,
k,
n being
Nat holds
( (
m + 1
<= k &
k <= n ) iff ex
i being
Element of
NAT st
(
m <= i &
i < n &
k = i + 1 ) )
theorem Th2:
theorem Th3:
Lm1:
for m, n being Element of NAT
for F being finite set st F = { k where k is Element of NAT : ( m <= k & k <= m + n ) } holds
card F = n + 1
theorem Th4:
theorem Th5:
begin
:: deftheorem Def1 defines -cut GRAPH_2:def 1 :
for p being FinSequence
for m, n being Nat
for b4 being FinSequence holds
( ( 1 <= m & m <= n & n <= len p implies ( b4 = (m,n) -cut p iff ( (len b4) + m = n + 1 & ( for i being Nat st i < len b4 holds
b4 . (i + 1) = p . (m + i) ) ) ) ) & ( ( not 1 <= m or not m <= n or not n <= len p ) implies ( b4 = (m,n) -cut p iff b4 = {} ) ) );
Lm2:
for p being FinSequence
for m, n being Element of NAT st 1 <= m & m <= n + 1 & n <= len p holds
( (len ((m,n) -cut p)) + m = n + 1 & ( for i being Element of NAT st i < len ((m,n) -cut p) holds
((m,n) -cut p) . (i + 1) = p . (m + i) ) )
theorem Th6:
theorem Th7:
theorem Th8:
theorem
theorem
theorem Th11:
theorem Th12:
begin
:: deftheorem defines ^' GRAPH_2:def 2 :
for p, q being FinSequence holds p ^' q = p ^ ((2,(len q)) -cut q);
theorem Th13:
theorem Th14:
theorem Th15:
theorem Th16:
theorem Th17:
theorem
begin
:: deftheorem Def3 defines TwoValued GRAPH_2:def 3 :
for f being FinSequence holds
( f is TwoValued iff card (rng f) = 2 );
theorem Th19:
then Lm4:
<*1,2*> is TwoValued
by Lm3;
Lm5:
now
let i be
Nat;
( 1 <= i & i + 1 <= len <*1,2*> implies <*1,2*> . i <> <*1,2*> . (i + 1) )set p =
<*1,2*>;
assume that A1:
1
<= i
and A2:
i + 1
<= len <*1,2*>
;
<*1,2*> . i <> <*1,2*> . (i + 1)
i + 1
<= 1
+ 1
by A2, FINSEQ_1:61;
then
i <= 1
by XREAL_1:8;
then A3:
i = 1
by A1, XXREAL_0:1;
then
<*1,2*> . i = 1
by FINSEQ_1:61;
hence
<*1,2*> . i <> <*1,2*> . (i + 1)
by A3, FINSEQ_1:61;
verum
end;
:: deftheorem Def4 defines Alternating GRAPH_2:def 4 :
for f being FinSequence holds
( f is Alternating iff for i being Nat st 1 <= i & i + 1 <= len f holds
f . i <> f . (i + 1) );
Lm6:
<*1,2*> is Alternating
by Def4, Lm5;
theorem Th20:
theorem Th21:
theorem Th22:
theorem Th23:
begin
theorem
canceled;
theorem
canceled;
theorem Th26:
theorem
theorem Th28:
theorem Th29:
theorem Th30:
begin
theorem Th31:
theorem Th32:
for
G being
Graph for
v1,
v2,
v3,
v4 being
Element of
G for
e being
set st
e joins v1,
v2 &
e joins v3,
v4 & not (
v1 = v3 &
v2 = v4 ) holds
(
v1 = v4 &
v2 = v3 )
:: deftheorem GRAPH_2:def 5 :
canceled;
:: deftheorem defines -VSet GRAPH_2:def 6 :
for G being Graph
for X being set holds G -VSet X = { v where v is Element of G : ex e being Element of the carrier' of G st
( e in X & ( v = the Source of G . e or v = the Target of G . e ) ) } ;
:: deftheorem Def7 defines is_vertex_seq_of GRAPH_2:def 7 :
for G being Graph
for vs being FinSequence of the carrier of G
for c being FinSequence holds
( vs is_vertex_seq_of c iff ( len vs = (len c) + 1 & ( for n being Element of NAT st 1 <= n & n <= len c holds
c . n joins vs /. n,vs /. (n + 1) ) ) );
theorem
canceled;
theorem Th34:
theorem Th35:
theorem Th36:
theorem Th37:
:: deftheorem Def8 defines alternates_vertices_in GRAPH_2:def 8 :
for G being Graph
for c being FinSequence holds
( c alternates_vertices_in G iff ( len c >= 1 & card (G -VSet (rng c)) = 2 & ( for n being Element of NAT st n in dom c holds
the Source of G . (c . n) <> the Target of G . (c . n) ) ) );
theorem Th38:
theorem Th39:
theorem Th40:
theorem Th41:
Lm7:
for D being non empty set st ( for x, y being set st x in D & y in D holds
x = y ) holds
card D = 1
theorem Th42:
:: deftheorem defines vertex-seq GRAPH_2:def 9 :
for G being Graph
for c being Chain of G st ( card the carrier of G = 1 or ( c <> {} & not c alternates_vertices_in G ) ) holds
for b3 being FinSequence of the carrier of G holds
( b3 = vertex-seq c iff b3 is_vertex_seq_of c );
theorem Th43:
theorem Th44:
theorem Th45:
theorem Th46:
theorem Th47:
begin
Lm8:
for G being Graph
for v being Element of G holds <*v*> is_vertex_seq_of {}
:: deftheorem Def10 defines simple GRAPH_2:def 10 :
for G being Graph
for IT being Chain of G holds
( IT is simple iff ex vs being FinSequence of the carrier of G st
( vs is_vertex_seq_of IT & ( for n, m being Element of NAT st 1 <= n & n < m & m <= len vs & vs . n = vs . m holds
( n = 1 & m = len vs ) ) ) );
theorem
canceled;
theorem
theorem Th50:
theorem
theorem Th52:
theorem
theorem
theorem
theorem
:: deftheorem defines vertex-seq GRAPH_2:def 11 :
for G being Graph
for oc being oriented Chain of G st oc <> {} holds
for b3 being FinSequence of the carrier of G holds
( b3 = vertex-seq oc iff ( b3 is_vertex_seq_of oc & b3 . 1 = the Source of G . (oc . 1) ) );
begin
theorem
theorem
theorem Th59:
theorem
theorem
theorem
definition
let F be
FinSequence of
INT ;
let m,
n be
Element of
NAT ;
assume that A1:
1
<= m
and A2:
m <= n
and A3:
n <= len F
;
func min_at (
F,
m,
n)
-> Element of
NAT means :
Def12:
ex
X being non
empty finite Subset of
INT st
(
X = rng ((m,n) -cut F) &
it + 1
= ((min X) .. ((m,n) -cut F)) + m );
existence
ex b1 being Element of NAT ex X being non empty finite Subset of INT st
( X = rng ((m,n) -cut F) & b1 + 1 = ((min X) .. ((m,n) -cut F)) + m )
uniqueness
for b1, b2 being Element of NAT st ex X being non empty finite Subset of INT st
( X = rng ((m,n) -cut F) & b1 + 1 = ((min X) .. ((m,n) -cut F)) + m ) & ex X being non empty finite Subset of INT st
( X = rng ((m,n) -cut F) & b2 + 1 = ((min X) .. ((m,n) -cut F)) + m ) holds
b1 = b2
;
end;
:: deftheorem Def12 defines min_at GRAPH_2:def 12 :
for F being FinSequence of INT
for m, n being Element of NAT st 1 <= m & m <= n & n <= len F holds
for b4 being Element of NAT holds
( b4 = min_at (F,m,n) iff ex X being non empty finite Subset of INT st
( X = rng ((m,n) -cut F) & b4 + 1 = ((min X) .. ((m,n) -cut F)) + m ) );
theorem Th63:
theorem
:: deftheorem Def13 defines is_non_decreasing_on GRAPH_2:def 13 :
for F being FinSequence of INT
for m, n being Element of NAT holds
( F is_non_decreasing_on m,n iff for i, j being Element of NAT st m <= i & i <= j & j <= n holds
F . i <= F . j );
:: deftheorem Def14 defines is_split_at GRAPH_2:def 14 :
for F being FinSequence of INT
for n being Element of NAT holds
( F is_split_at n iff for i, j being Element of NAT st 1 <= i & i <= n & n < j & j <= len F holds
F . i <= F . j );
theorem
theorem