begin
theorem
canceled;
theorem Th2:
theorem Th3:
theorem Th4:
theorem
canceled;
theorem Th6:
theorem Th7:
begin
Lm1:
for n being Element of NAT
for p, q being FinSequence st 1 <= n & n <= len p holds
p . n = (p ^ q) . n
Lm2:
for n being Element of NAT
for q, p being FinSequence st 1 <= n & n <= len q holds
q . n = (p ^ q) . ((len p) + n)
theorem
canceled;
theorem Th9:
theorem Th10:
theorem
theorem Th12:
begin
theorem Th13:
theorem Th14:
theorem Th15:
begin
theorem Th16:
theorem Th17:
theorem Th18:
theorem Th19:
theorem Th20:
begin
:: deftheorem defines vertices GRAPH_5:def 1 :
for G being Graph
for e being Element of the carrier' of G holds vertices e = {( the Source of G . e),( the Target of G . e)};
:: deftheorem defines vertices GRAPH_5:def 2 :
for G being Graph
for pe being FinSequence of the carrier' of G holds vertices pe = { v where v is Vertex of G : ex i being Element of NAT st
( i in dom pe & v in vertices (pe /. i) ) } ;
theorem Th21:
theorem Th22:
theorem Th23:
theorem Th24:
theorem Th25:
theorem Th26:
theorem
canceled;
theorem Th28:
theorem Th29:
theorem Th30:
theorem Th31:
theorem Th32:
begin
:: deftheorem Def3 defines is_orientedpath_of GRAPH_5:def 3 :
for G being Graph
for p being oriented Chain of G
for v1, v2 being Element of G holds
( p is_orientedpath_of v1,v2 iff ( p <> {} & the Source of G . (p . 1) = v1 & the Target of G . (p . (len p)) = v2 ) );
:: deftheorem Def4 defines is_orientedpath_of GRAPH_5:def 4 :
for G being Graph
for v1, v2 being Element of G
for p being oriented Chain of G
for V being set holds
( p is_orientedpath_of v1,v2,V iff ( p is_orientedpath_of v1,v2 & (vertices p) \ {v2} c= V ) );
:: deftheorem defines OrientedPaths GRAPH_5:def 5 :
for G being Graph
for v1, v2 being Element of G holds OrientedPaths (v1,v2) = { p where p is oriented Chain of G : p is_orientedpath_of v1,v2 } ;
theorem Th33:
theorem
theorem Th35:
theorem Th36:
theorem Th37:
theorem Th38:
begin
:: deftheorem Def6 defines is_acyclicpath_of GRAPH_5:def 6 :
for G being Graph
for p being oriented Chain of G
for v1, v2 being Element of G holds
( p is_acyclicpath_of v1,v2 iff ( p is Simple & p is_orientedpath_of v1,v2 ) );
:: deftheorem Def7 defines is_acyclicpath_of GRAPH_5:def 7 :
for G being Graph
for p being oriented Chain of G
for v1, v2 being Element of G
for V being set holds
( p is_acyclicpath_of v1,v2,V iff ( p is Simple & p is_orientedpath_of v1,v2,V ) );
:: deftheorem defines AcyclicPaths GRAPH_5:def 8 :
for G being Graph
for v1, v2 being Element of G holds AcyclicPaths (v1,v2) = { p where p is oriented Simple Chain of G : p is_acyclicpath_of v1,v2 } ;
:: deftheorem defines AcyclicPaths GRAPH_5:def 9 :
for G being Graph
for v1, v2 being Element of G
for V being set holds AcyclicPaths (v1,v2,V) = { p where p is oriented Simple Chain of G : p is_acyclicpath_of v1,v2,V } ;
:: deftheorem defines AcyclicPaths GRAPH_5:def 10 :
for G being Graph
for p being oriented Chain of G holds AcyclicPaths p = { q where q is oriented Simple Chain of G : ( q <> {} & the Source of G . (q . 1) = the Source of G . (p . 1) & the Target of G . (q . (len q)) = the Target of G . (p . (len p)) & rng q c= rng p ) } ;
:: deftheorem defines AcyclicPaths GRAPH_5:def 11 :
for G being Graph holds AcyclicPaths G = { q where q is oriented Simple Chain of G : verum } ;
theorem
theorem
theorem
theorem Th42:
theorem Th43:
theorem Th44:
theorem Th45:
theorem
theorem Th47:
theorem Th48:
theorem Th49:
begin
:: deftheorem defines Real>=0 GRAPH_5:def 12 :
Real>=0 = { r where r is Real : r >= 0 } ;
:: deftheorem Def13 defines is_weight>=0of GRAPH_5:def 13 :
for G being Graph
for W being Function holds
( W is_weight>=0of G iff W is Function of the carrier' of G,Real>=0 );
:: deftheorem Def14 defines is_weight_of GRAPH_5:def 14 :
for G being Graph
for W being Function holds
( W is_weight_of G iff W is Function of the carrier' of G,REAL );
:: deftheorem Def15 defines RealSequence GRAPH_5:def 15 :
for G being Graph
for p being FinSequence of the carrier' of G
for W being Function st W is_weight_of G holds
for b4 being FinSequence of REAL holds
( b4 = RealSequence (p,W) iff ( dom p = dom b4 & ( for i being Element of NAT st i in dom p holds
b4 . i = W . (p . i) ) ) );
:: deftheorem defines cost GRAPH_5:def 16 :
for G being Graph
for p being FinSequence of the carrier' of G
for W being Function holds cost (p,W) = Sum (RealSequence (p,W));
theorem Th50:
theorem Th51:
theorem Th52:
Lm3:
for f being FinSequence of REAL holds
( ( for y being Real st y in rng f holds
y >= 0 ) iff for i being Nat st i in dom f holds
f . i >= 0 )
Lm4:
for p, q, r being FinSequence of REAL st r = p ^ q & ( for x being Real st x in rng r holds
x >= 0 ) holds
( ( for i being Element of NAT st i in dom p holds
p . i >= 0 ) & ( for i being Element of NAT st i in dom q holds
q . i >= 0 ) )
theorem
theorem Th54:
theorem Th55:
theorem Th56:
theorem Th57:
theorem Th58:
theorem Th59:
theorem Th60:
begin
:: deftheorem Def17 defines is_shortestpath_of GRAPH_5:def 17 :
for G being Graph
for v1, v2 being Vertex of G
for p being oriented Chain of G
for W being Function holds
( p is_shortestpath_of v1,v2,W iff ( p is_orientedpath_of v1,v2 & ( for q being oriented Chain of G st q is_orientedpath_of v1,v2 holds
cost (p,W) <= cost (q,W) ) ) );
definition
let G be
Graph;
let v1,
v2 be
Vertex of
G;
let p be
oriented Chain of
G;
let V be
set ;
let W be
Function;
pred p is_shortestpath_of v1,
v2,
V,
W means :
Def18:
(
p is_orientedpath_of v1,
v2,
V & ( for
q being
oriented Chain of
G st
q is_orientedpath_of v1,
v2,
V holds
cost (
p,
W)
<= cost (
q,
W) ) );
end;
:: deftheorem Def18 defines is_shortestpath_of GRAPH_5:def 18 :
for G being Graph
for v1, v2 being Vertex of G
for p being oriented Chain of G
for V being set
for W being Function holds
( p is_shortestpath_of v1,v2,V,W iff ( p is_orientedpath_of v1,v2,V & ( for q being oriented Chain of G st q is_orientedpath_of v1,v2,V holds
cost (p,W) <= cost (q,W) ) ) );
begin
theorem
theorem Th62:
Lm5:
for G being finite Graph holds AcyclicPaths G is finite
Lm6:
for G being finite Graph
for P being oriented Chain of G holds AcyclicPaths P is finite
Lm7:
for G being finite Graph
for v1, v2 being Element of G holds AcyclicPaths (v1,v2) is finite
Lm8:
for V being set
for G being finite Graph
for v1, v2 being Element of G holds AcyclicPaths (v1,v2,V) is finite
theorem
theorem
theorem
theorem Th66:
begin
theorem Th67:
for
V being
set for
W being
Function for
G being
finite Graph for
P being
oriented Chain of
G for
v1,
v2 being
Element of
G st
W is_weight>=0of G &
P is_shortestpath_of v1,
v2,
V,
W &
v1 <> v2 & ( for
Q being
oriented Chain of
G for
v3 being
Element of
G st not
v3 in V &
Q is_shortestpath_of v1,
v3,
V,
W holds
cost (
P,
W)
<= cost (
Q,
W) ) holds
P is_shortestpath_of v1,
v2,
W
theorem
for
V,
U being
set for
W being
Function for
G being
finite Graph for
P being
oriented Chain of
G for
v1,
v2 being
Element of
G st
W is_weight>=0of G &
P is_shortestpath_of v1,
v2,
V,
W &
v1 <> v2 &
V c= U & ( for
Q being
oriented Chain of
G for
v3 being
Element of
G st not
v3 in V &
Q is_shortestpath_of v1,
v3,
V,
W holds
cost (
P,
W)
<= cost (
Q,
W) ) holds
P is_shortestpath_of v1,
v2,
U,
W
:: deftheorem Def19 defines islongestInShortestpath GRAPH_5:def 19 :
for G being Graph
for pe being FinSequence of the carrier' of G
for V being set
for v1 being Vertex of G
for W being Function holds
( pe islongestInShortestpath V,v1,W iff for v being Vertex of G st v in V & v <> v1 holds
ex q being oriented Chain of G st
( q is_shortestpath_of v1,v,V,W & cost (q,W) <= cost (pe,W) ) );
theorem
for
e,
V being
set for
W being
Function for
G being
oriented finite Graph for
P,
Q,
R being
oriented Chain of
G for
v1,
v2,
v3 being
Element of
G st
e in the
carrier' of
G &
W is_weight>=0of G &
len P >= 1 &
P is_shortestpath_of v1,
v2,
V,
W &
v1 <> v2 &
v1 <> v3 &
R = P ^ <*e*> &
Q is_shortestpath_of v1,
v3,
V,
W &
e orientedly_joins v2,
v3 &
P islongestInShortestpath V,
v1,
W holds
( (
cost (
Q,
W)
<= cost (
R,
W) implies
Q is_shortestpath_of v1,
v3,
V \/ {v2},
W ) & (
cost (
Q,
W)
>= cost (
R,
W) implies
R is_shortestpath_of v1,
v3,
V \/ {v2},
W ) )