begin
:: deftheorem Def1 defines orientedly_joins GRAPH_4:def 1 :
for G being Graph
for x, y being Element of G
for e being set holds
( e orientedly_joins x,y iff ( the Source of G . e = x & the Target of G . e = y ) );
theorem Th1:
:: deftheorem defines are_orientedly_incident GRAPH_4:def 2 :
for G being Graph
for x, y being Element of the carrier of G holds
( x,y are_orientedly_incident iff ex v being set st
( v in the carrier' of G & v orientedly_joins x,y ) );
theorem Th2:
:: deftheorem defines -SVSet GRAPH_4:def 3 :
for G being Graph
for X being set holds G -SVSet 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 ) } ;
:: deftheorem defines -TVSet GRAPH_4:def 4 :
for G being Graph
for X being set holds G -TVSet X = { v where v is Element of G : ex e being Element of the carrier' of G st
( e in X & v = the Target of G . e ) } ;
theorem
canceled;
theorem
:: deftheorem Def5 defines is_oriented_vertex_seq_of GRAPH_4:def 5 :
for G being Graph
for vs being FinSequence of the carrier of G
for c being FinSequence holds
( vs is_oriented_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 orientedly_joins vs /. n,vs /. (n + 1) ) ) );
theorem Th5:
theorem
theorem
theorem
begin
theorem
theorem Th10:
theorem Th11:
:: deftheorem defines oriented-vertex-seq GRAPH_4:def 6 :
for G being Graph
for c being oriented Chain of G st c <> {} holds
for b3 being FinSequence of the carrier of G holds
( b3 = oriented-vertex-seq c iff b3 is_oriented_vertex_seq_of c );
theorem
theorem Th13:
theorem Th14:
theorem Th15:
theorem Th16:
begin
Lm1:
for G being Graph
for v being Element of G holds <*v*> is_oriented_vertex_seq_of {}
:: deftheorem Def7 defines Simple GRAPH_4:def 7 :
for G being Graph
for IT being oriented Chain of G holds
( IT is Simple iff ex vs being FinSequence of the carrier of G st
( vs is_oriented_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 Th18:
theorem
theorem Th20:
theorem Th21:
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 Th22:
theorem
theorem
theorem Th25:
theorem