begin
theorem Th1:
theorem Th2:
theorem Th3:
theorem Th4:
theorem Th5:
Lm1:
for n being 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 p, q being FinSequence st 1 <= n & n <= len q holds
q . n = (p ^ q) . ((len p) + n)
theorem Th6:
begin
theorem Th7:
theorem Th8:
theorem Th9:
for
G being
Graph for
p,
q being
oriented Chain of
G for
W being
Function for
V being
set for
v1,
v2 being
Vertex of
G st
p is_shortestpath_of v1,
v2,
V,
W &
q is_shortestpath_of v1,
v2,
V,
W holds
cost (
p,
W)
= cost (
q,
W)
theorem Th10:
Lm3:
for i being Element of NAT
for G being Graph
for pe being FinSequence of the carrier' of G st 1 <= i & i <= len pe holds
( the Source of G . (pe . i) in the carrier of G & the Target of G . (pe . i) in the carrier of G )
theorem Th11:
Lm4:
for i being Element of NAT
for G being Graph
for pe being FinSequence of the carrier' of G
for v1 being Vertex of G st 1 <= i & i <= len pe & ( v1 = the Source of G . (pe . i) or v1 = the Target of G . (pe . i) ) holds
v1 in vertices pe
theorem Th12:
begin
theorem Th13:
for
W being
Function for
V being
set for
G being
finite Graph for
P,
Q being
oriented Chain of
G for
v1,
v2,
v3 being
Vertex of
G st
W is_weight>=0of G &
P is_shortestpath_of v1,
v2,
V,
W &
v1 <> v2 &
v1 <> v3 &
Q is_shortestpath_of v1,
v3,
V,
W & ( for
e being
set holds
( not
e in the
carrier' of
G or not
e orientedly_joins v2,
v3 ) ) &
P islongestInShortestpath V,
v1,
W holds
Q is_shortestpath_of v1,
v3,
V \/ {v2},
W
theorem Th14:
theorem Th15:
for
e,
V being
set for
G being
oriented finite Graph for
P,
Q being
oriented Chain of
G for
W being
Function of the
carrier' of
G,
Real>=0 for
v1,
v2,
v3 being
Vertex of
G st
e in the
carrier' of
G &
P is_shortestpath_of v1,
v2,
V,
W &
v1 <> v3 &
Q = P ^ <*e*> &
e orientedly_joins v2,
v3 &
v1 in V & ( for
v4 being
Vertex of
G st
v4 in V holds
for
ee being
set holds
( not
ee in the
carrier' of
G or not
ee orientedly_joins v4,
v3 ) ) holds
Q is_shortestpath_of v1,
v3,
V \/ {v2},
W
theorem Th16:
for
U,
V being
set for
G being
oriented finite Graph for
P being
oriented Chain of
G for
W being
Function of the
carrier' of
G,
Real>=0 for
v1,
v2 being
Vertex of
G st the
carrier of
G = U \/ V &
v1 in U & ( for
v3,
v4 being
Vertex of
G st
v3 in U &
v4 in V holds
for
e being
set holds
( not
e in the
carrier' of
G or not
e orientedly_joins v3,
v4 ) ) holds
(
P is_shortestpath_of v1,
v2,
U,
W iff
P is_shortestpath_of v1,
v2,
W )
begin
:: deftheorem defines := GRAPHSP:def 1 :
for i, k being Element of NAT
for f being FinSequence of REAL
for r being Real holds (f,i) := (k,r) = (((f,i) := k),k) := r;
theorem
canceled;
theorem Th18:
theorem Th19:
theorem Th20:
theorem Th21:
begin
:: deftheorem Def2 defines repeat GRAPHSP:def 2 :
for X being set
for f being Element of Funcs (X,X)
for b3 being Function of NAT,(Funcs (X,X)) holds
( b3 = repeat f iff ( b3 . 0 = id X & ( for i being Nat holds b3 . (i + 1) = f * (b3 . i) ) ) );
theorem Th22:
Lm5:
for X being set
for f being Element of Funcs (X,X) holds dom f = X
theorem Th23:
:: deftheorem defines OuterVx GRAPHSP:def 3 :
for f being Element of REAL *
for n being Nat holds OuterVx (f,n) = { i where i is Element of NAT : ( i in dom f & 1 <= i & i <= n & f . i <> - 1 & f . (n + i) <> - 1 ) } ;
definition
let f be
Element of
Funcs (
(REAL *),
(REAL *));
let g be
Element of
REAL * ;
let n be
Nat;
assume A1:
ex
i being
Element of
NAT st
OuterVx (
(((repeat f) . i) . g),
n)
= {}
;
func LifeSpan (
f,
g,
n)
-> Element of
NAT means :
Def4:
(
OuterVx (
(((repeat f) . it) . g),
n)
= {} & ( for
k being
Nat st
OuterVx (
(((repeat f) . k) . g),
n)
= {} holds
it <= k ) );
existence
ex b1 being Element of NAT st
( OuterVx ((((repeat f) . b1) . g),n) = {} & ( for k being Nat st OuterVx ((((repeat f) . k) . g),n) = {} holds
b1 <= k ) )
uniqueness
for b1, b2 being Element of NAT st OuterVx ((((repeat f) . b1) . g),n) = {} & ( for k being Nat st OuterVx ((((repeat f) . k) . g),n) = {} holds
b1 <= k ) & OuterVx ((((repeat f) . b2) . g),n) = {} & ( for k being Nat st OuterVx ((((repeat f) . k) . g),n) = {} holds
b2 <= k ) holds
b1 = b2
end;
:: deftheorem Def4 defines LifeSpan GRAPHSP:def 4 :
for f being Element of Funcs ((REAL *),(REAL *))
for g being Element of REAL *
for n being Nat st ex i being Element of NAT st OuterVx ((((repeat f) . i) . g),n) = {} holds
for b4 being Element of NAT holds
( b4 = LifeSpan (f,g,n) iff ( OuterVx ((((repeat f) . b4) . g),n) = {} & ( for k being Nat st OuterVx ((((repeat f) . k) . g),n) = {} holds
b4 <= k ) ) );
definition
let f be
Element of
Funcs (
(REAL *),
(REAL *));
let n be
Element of
NAT ;
func while_do (
f,
n)
-> Element of
Funcs (
(REAL *),
(REAL *))
means :
Def5:
(
dom it = REAL * & ( for
h being
Element of
REAL * holds
it . h = ((repeat f) . (LifeSpan (f,h,n))) . h ) );
existence
ex b1 being Element of Funcs ((REAL *),(REAL *)) st
( dom b1 = REAL * & ( for h being Element of REAL * holds b1 . h = ((repeat f) . (LifeSpan (f,h,n))) . h ) )
uniqueness
for b1, b2 being Element of Funcs ((REAL *),(REAL *)) st dom b1 = REAL * & ( for h being Element of REAL * holds b1 . h = ((repeat f) . (LifeSpan (f,h,n))) . h ) & dom b2 = REAL * & ( for h being Element of REAL * holds b2 . h = ((repeat f) . (LifeSpan (f,h,n))) . h ) holds
b1 = b2
end;
:: deftheorem Def5 defines while_do GRAPHSP:def 5 :
for f being Element of Funcs ((REAL *),(REAL *))
for n being Element of NAT
for b3 being Element of Funcs ((REAL *),(REAL *)) holds
( b3 = while_do (f,n) iff ( dom b3 = REAL * & ( for h being Element of REAL * holds b3 . h = ((repeat f) . (LifeSpan (f,h,n))) . h ) ) );
begin
:: deftheorem Def6 defines XEdge GRAPHSP:def 6 :
for G being oriented Graph
for v1, v2 being Vertex of G st ex e being set st
( e in the carrier' of G & e orientedly_joins v1,v2 ) holds
for b4 being set holds
( b4 = XEdge (v1,v2) iff ex e being set st
( b4 = e & e in the carrier' of G & e orientedly_joins v1,v2 ) );
:: deftheorem Def7 defines Weight GRAPHSP:def 7 :
for G being oriented Graph
for v1, v2 being Vertex of G
for W being Function holds
( ( ex e being set st
( e in the carrier' of G & e orientedly_joins v1,v2 ) implies Weight (v1,v2,W) = W . (XEdge (v1,v2)) ) & ( ( for e being set holds
( not e in the carrier' of G or not e orientedly_joins v1,v2 ) ) implies Weight (v1,v2,W) = - 1 ) );
theorem Th24:
theorem
theorem Th26:
begin
:: deftheorem defines UnusedVx GRAPHSP:def 8 :
for f being Element of REAL *
for n being Element of NAT holds UnusedVx (f,n) = { i where i is Element of NAT : ( i in dom f & 1 <= i & i <= n & f . i <> - 1 ) } ;
:: deftheorem defines UsedVx GRAPHSP:def 9 :
for f being Element of REAL *
for n being Element of NAT holds UsedVx (f,n) = { i where i is Element of NAT : ( i in dom f & 1 <= i & i <= n & f . i = - 1 ) } ;
theorem Th27:
theorem Th28:
theorem Th29:
:: deftheorem Def10 defines Argmin GRAPHSP:def 10 :
for X being finite Subset of NAT
for f being Element of REAL *
for n, b4 being Element of NAT holds
( b4 = Argmin (X,f,n) iff ( ( X <> {} implies ex i being Element of NAT st
( i = b4 & i in X & ( for k being Element of NAT st k in X holds
f /. ((2 * n) + i) <= f /. ((2 * n) + k) ) & ( for k being Element of NAT st k in X & f /. ((2 * n) + i) = f /. ((2 * n) + k) holds
i <= k ) ) ) & ( X = {} implies b4 = 0 ) ) );
theorem Th30:
theorem Th31:
definition
let n be
Element of
NAT ;
func findmin n -> Element of
Funcs (
(REAL *),
(REAL *))
means :
Def11:
(
dom it = REAL * & ( for
f being
Element of
REAL * holds
it . f = (
f,
(((n * n) + (3 * n)) + 1))
:= (
(Argmin ((OuterVx (f,n)),f,n)),
(- 1)) ) );
existence
ex b1 being Element of Funcs ((REAL *),(REAL *)) st
( dom b1 = REAL * & ( for f being Element of REAL * holds b1 . f = (f,(((n * n) + (3 * n)) + 1)) := ((Argmin ((OuterVx (f,n)),f,n)),(- 1)) ) )
uniqueness
for b1, b2 being Element of Funcs ((REAL *),(REAL *)) st dom b1 = REAL * & ( for f being Element of REAL * holds b1 . f = (f,(((n * n) + (3 * n)) + 1)) := ((Argmin ((OuterVx (f,n)),f,n)),(- 1)) ) & dom b2 = REAL * & ( for f being Element of REAL * holds b2 . f = (f,(((n * n) + (3 * n)) + 1)) := ((Argmin ((OuterVx (f,n)),f,n)),(- 1)) ) holds
b1 = b2
end;
:: deftheorem Def11 defines findmin GRAPHSP:def 11 :
for n being Element of NAT
for b2 being Element of Funcs ((REAL *),(REAL *)) holds
( b2 = findmin n iff ( dom b2 = REAL * & ( for f being Element of REAL * holds b2 . f = (f,(((n * n) + (3 * n)) + 1)) := ((Argmin ((OuterVx (f,n)),f,n)),(- 1)) ) ) );
theorem Th32:
theorem Th33:
theorem Th34:
Lm6:
for k, n being Element of NAT st k >= 1 holds
n <= k * n
Lm7:
for n being Element of NAT holds
( 3 * n < ((n * n) + (3 * n)) + 1 & n < ((n * n) + (3 * n)) + 1 & 2 * n < ((n * n) + (3 * n)) + 1 )
Lm8:
for n, k being Element of NAT holds
( ( n < k & k <= 2 * n implies ( ( not 2 * n < k or not k <= 3 * n ) & not k <= n & not k > 3 * n ) ) & ( ( k <= n or k > 3 * n ) implies ( ( not 2 * n < k or not k <= 3 * n ) & ( not n < k or not k <= 2 * n ) ) ) & ( 2 * n < k & k <= 3 * n implies ( ( not n < k or not k <= 2 * n ) & not k <= n & not k > 3 * n ) ) )
theorem Th35:
:: deftheorem defines newpathcost GRAPHSP:def 12 :
for f being Element of REAL *
for n, k being Element of NAT holds newpathcost (f,n,k) = (f /. ((2 * n) + (f /. (((n * n) + (3 * n)) + 1)))) + (f /. (((2 * n) + (n * (f /. (((n * n) + (3 * n)) + 1)))) + k));
:: deftheorem Def13 defines hasBetterPathAt GRAPHSP:def 13 :
for n, k being Element of NAT
for f being Element of REAL * holds
( f hasBetterPathAt n,k iff ( ( f . (n + k) = - 1 or f /. ((2 * n) + k) > newpathcost (f,n,k) ) & f /. (((2 * n) + (n * (f /. (((n * n) + (3 * n)) + 1)))) + k) >= 0 & f . k <> - 1 ) );
definition
let f be
Element of
REAL * ;
let n be
Element of
NAT ;
func Relax (
f,
n)
-> Element of
REAL * means :
Def14:
(
dom it = dom f & ( for
k being
Element of
NAT st
k in dom f holds
( (
n < k &
k <= 2
* n implies ( (
f hasBetterPathAt n,
k -' n implies
it . k = f /. (((n * n) + (3 * n)) + 1) ) & ( not
f hasBetterPathAt n,
k -' n implies
it . k = f . k ) ) ) & ( 2
* n < k &
k <= 3
* n implies ( (
f hasBetterPathAt n,
k -' (2 * n) implies
it . k = newpathcost (
f,
n,
(k -' (2 * n))) ) & ( not
f hasBetterPathAt n,
k -' (2 * n) implies
it . k = f . k ) ) ) & ( (
k <= n or
k > 3
* n ) implies
it . k = f . k ) ) ) );
existence
ex b1 being Element of REAL * st
( dom b1 = dom f & ( for k being Element of NAT st k in dom f holds
( ( n < k & k <= 2 * n implies ( ( f hasBetterPathAt n,k -' n implies b1 . k = f /. (((n * n) + (3 * n)) + 1) ) & ( not f hasBetterPathAt n,k -' n implies b1 . k = f . k ) ) ) & ( 2 * n < k & k <= 3 * n implies ( ( f hasBetterPathAt n,k -' (2 * n) implies b1 . k = newpathcost (f,n,(k -' (2 * n))) ) & ( not f hasBetterPathAt n,k -' (2 * n) implies b1 . k = f . k ) ) ) & ( ( k <= n or k > 3 * n ) implies b1 . k = f . k ) ) ) )
uniqueness
for b1, b2 being Element of REAL * st dom b1 = dom f & ( for k being Element of NAT st k in dom f holds
( ( n < k & k <= 2 * n implies ( ( f hasBetterPathAt n,k -' n implies b1 . k = f /. (((n * n) + (3 * n)) + 1) ) & ( not f hasBetterPathAt n,k -' n implies b1 . k = f . k ) ) ) & ( 2 * n < k & k <= 3 * n implies ( ( f hasBetterPathAt n,k -' (2 * n) implies b1 . k = newpathcost (f,n,(k -' (2 * n))) ) & ( not f hasBetterPathAt n,k -' (2 * n) implies b1 . k = f . k ) ) ) & ( ( k <= n or k > 3 * n ) implies b1 . k = f . k ) ) ) & dom b2 = dom f & ( for k being Element of NAT st k in dom f holds
( ( n < k & k <= 2 * n implies ( ( f hasBetterPathAt n,k -' n implies b2 . k = f /. (((n * n) + (3 * n)) + 1) ) & ( not f hasBetterPathAt n,k -' n implies b2 . k = f . k ) ) ) & ( 2 * n < k & k <= 3 * n implies ( ( f hasBetterPathAt n,k -' (2 * n) implies b2 . k = newpathcost (f,n,(k -' (2 * n))) ) & ( not f hasBetterPathAt n,k -' (2 * n) implies b2 . k = f . k ) ) ) & ( ( k <= n or k > 3 * n ) implies b2 . k = f . k ) ) ) holds
b1 = b2
end;
:: deftheorem Def14 defines Relax GRAPHSP:def 14 :
for f being Element of REAL *
for n being Element of NAT
for b3 being Element of REAL * holds
( b3 = Relax (f,n) iff ( dom b3 = dom f & ( for k being Element of NAT st k in dom f holds
( ( n < k & k <= 2 * n implies ( ( f hasBetterPathAt n,k -' n implies b3 . k = f /. (((n * n) + (3 * n)) + 1) ) & ( not f hasBetterPathAt n,k -' n implies b3 . k = f . k ) ) ) & ( 2 * n < k & k <= 3 * n implies ( ( f hasBetterPathAt n,k -' (2 * n) implies b3 . k = newpathcost (f,n,(k -' (2 * n))) ) & ( not f hasBetterPathAt n,k -' (2 * n) implies b3 . k = f . k ) ) ) & ( ( k <= n or k > 3 * n ) implies b3 . k = f . k ) ) ) ) );
:: deftheorem Def15 defines Relax GRAPHSP:def 15 :
for n being Element of NAT
for b2 being Element of Funcs ((REAL *),(REAL *)) holds
( b2 = Relax n iff ( dom b2 = REAL * & ( for f being Element of REAL * holds b2 . f = Relax (f,n) ) ) );
theorem Th36:
theorem Th37:
theorem Th38:
theorem Th39:
theorem Th40:
for
n,
i,
k being
Element of
NAT for
g,
f,
h being
Element of
REAL * st
g = ((repeat ((Relax n) * (findmin n))) . i) . f &
h = ((repeat ((Relax n) * (findmin n))) . (i + 1)) . f &
k = Argmin (
(OuterVx (g,n)),
g,
n) &
OuterVx (
g,
n)
<> {} holds
(
UsedVx (
h,
n)
= (UsedVx (g,n)) \/ {k} & not
k in UsedVx (
g,
n) )
theorem Th41:
Lm9:
for n, k being Element of NAT holds n - k <= n
Lm10:
for p, q being FinSequence of NAT
for f being Element of REAL *
for i, n being Element of NAT st ( for k being Element of NAT st 1 <= k & k < len p holds
p . ((len p) - k) = f . ((p /. (((len p) - k) + 1)) + n) ) & ( for k being Element of NAT st 1 <= k & k < len q holds
q . ((len q) - k) = f . ((q /. (((len q) - k) + 1)) + n) ) & len p <= len q & p . (len p) = q . (len q) holds
for k being Element of NAT st 1 <= k & k < len p holds
p . ((len p) - k) = q . ((len q) - k)
theorem Th42:
:: deftheorem Def16 defines equal_at GRAPHSP:def 16 :
for f, g being Element of REAL *
for m, n being Element of NAT holds
( f,g equal_at m,n iff ( dom f = dom g & ( for k being Element of NAT st k in dom f & m <= k & k <= n holds
f . k = g . k ) ) );
theorem Th43:
theorem Th44:
theorem Th45:
theorem
theorem Th47:
theorem Th48:
theorem Th49:
:: deftheorem Def17 defines is_vertex_seq_at GRAPHSP:def 17 :
for p being FinSequence of NAT
for f being Element of REAL *
for i, n being Element of NAT holds
( p is_vertex_seq_at f,i,n iff ( p . (len p) = i & ( for k being Element of NAT st 1 <= k & k < len p holds
p . ((len p) - k) = f . (n + (p /. (((len p) - k) + 1))) ) ) );
:: deftheorem Def18 defines is_simple_vertex_seq_at GRAPHSP:def 18 :
for p being FinSequence of NAT
for f being Element of REAL *
for i, n being Element of NAT holds
( p is_simple_vertex_seq_at f,i,n iff ( p . 1 = 1 & len p > 1 & p is_vertex_seq_at f,i,n & p is one-to-one ) );
theorem
:: deftheorem Def19 defines is_oriented_edge_seq_of GRAPHSP:def 19 :
for G being Graph
for p being FinSequence of the carrier' of G
for vs being FinSequence holds
( p is_oriented_edge_seq_of vs iff ( len vs = (len p) + 1 & ( for n being Nat st 1 <= n & n <= len p holds
( the Source of G . (p . n) = vs . n & the Target of G . (p . n) = vs . (n + 1) ) ) ) );
theorem
theorem
Lm11:
for i, n being Element of NAT st 1 <= i & i <= n holds
( 1 < (2 * n) + i & (2 * n) + i < ((n * n) + (3 * n)) + 1 & i < (2 * n) + i )
Lm12:
for i, n being Element of NAT st 1 <= i & i <= n holds
( 1 < n + i & n + i <= 2 * n & n + i < ((n * n) + (3 * n)) + 1 )
Lm13:
for i, n, j being Element of NAT st 1 <= i & i <= n & j <= n holds
( 1 < ((2 * n) + (n * i)) + j & i < ((2 * n) + (n * i)) + j & ((2 * n) + (n * i)) + j < ((n * n) + (3 * n)) + 1 )
Lm14:
for i, n, j being Element of NAT st 1 <= i & i <= n & 1 <= j & j <= n holds
( (3 * n) + 1 <= ((2 * n) + (n * i)) + j & ((2 * n) + (n * i)) + j <= (n * n) + (3 * n) )
begin
:: deftheorem Def20 defines is_Input_of_Dijkstra_Alg GRAPHSP:def 20 :
for f being Element of REAL *
for G being oriented Graph
for n being Element of NAT
for W being Function of the carrier' of G,Real>=0 holds
( f is_Input_of_Dijkstra_Alg G,n,W iff ( len f = ((n * n) + (3 * n)) + 1 & Seg n = the carrier of G & ( for i being Element of NAT st 1 <= i & i <= n holds
( f . i = 1 & f . ((2 * n) + i) = 0 ) ) & f . (n + 1) = 0 & ( for i being Element of NAT st 2 <= i & i <= n holds
f . (n + i) = - 1 ) & ( for i, j being Vertex of G
for k, m being Element of NAT st k = i & m = j holds
f . (((2 * n) + (n * k)) + m) = Weight (i,j,W) ) ) );
begin
:: deftheorem defines DijkstraAlgorithm GRAPHSP:def 21 :
for n being Element of NAT holds DijkstraAlgorithm n = while_do (((Relax n) * (findmin n)),n);
begin
Lm15:
for n being Element of NAT
for f, h being Element of REAL *
for G being oriented finite Graph
for W being Function of the carrier' of G,Real>=0
for v1 being Vertex of G st f is_Input_of_Dijkstra_Alg G,n,W & v1 = 1 & n >= 1 & h = ((repeat ((Relax n) * (findmin n))) . 1) . f holds
( ( for v3 being Vertex of G
for j being Element of NAT st v3 <> v1 & v3 = j & h . (n + j) <> - 1 holds
ex p being FinSequence of NAT ex P being oriented Chain of G st
( p is_simple_vertex_seq_at h,j,n & ( for i being Element of NAT st 1 <= i & i < len p holds
p . i in UsedVx (h,n) ) & P is_oriented_edge_seq_of p & P is_shortestpath_of v1,v3, UsedVx (h,n),W & cost (P,W) = h . ((2 * n) + j) & ( not v3 in UsedVx (h,n) implies P islongestInShortestpath UsedVx (h,n),v1,W ) ) ) & ( for m, j being Element of NAT st h . (n + j) = - 1 & 1 <= j & j <= n & m in UsedVx (h,n) holds
f . (((2 * n) + (n * m)) + j) = - 1 ) & ( for m being Element of NAT st m in UsedVx (h,n) holds
h . (n + m) <> - 1 ) )
Lm16:
for n, k, i being Element of NAT
for g, f, h being Element of REAL * st g = ((repeat ((Relax n) * (findmin n))) . k) . f & h = ((repeat ((Relax n) * (findmin n))) . (k + 1)) . f & OuterVx (g,n) <> {} & i in UsedVx (g,n) & len f = ((n * n) + (3 * n)) + 1 holds
h . (n + i) = g . (n + i)
Lm17:
for n, k, j being Element of NAT
for g, f, h being Element of REAL *
for p being FinSequence of NAT st g = ((repeat ((Relax n) * (findmin n))) . k) . f & h = ((repeat ((Relax n) * (findmin n))) . (k + 1)) . f & OuterVx (g,n) <> {} & len f = ((n * n) + (3 * n)) + 1 & p is_simple_vertex_seq_at g,j,n & g . (n + j) = h . (n + j) & ( for i being Element of NAT st 1 <= i & i < len p holds
p . i in UsedVx (g,n) ) holds
p is_simple_vertex_seq_at h,j,n
Lm18:
for n, k, m, j being Element of NAT
for g, f, h being Element of REAL *
for p being FinSequence of NAT st g = ((repeat ((Relax n) * (findmin n))) . k) . f & h = ((repeat ((Relax n) * (findmin n))) . (k + 1)) . f & OuterVx (g,n) <> {} & len f = ((n * n) + (3 * n)) + 1 & p is_simple_vertex_seq_at g,m,n & m = h . (n + j) & g . (n + m) = h . (n + m) & m <> j & not j in UsedVx (g,n) & ( for i being Element of NAT st 1 <= i & i < len p holds
p . i in UsedVx (g,n) ) holds
p ^ <*j*> is_simple_vertex_seq_at h,j,n
Lm19:
for n, i being Element of NAT
for V being set
for f, g being Element of REAL *
for G being oriented finite Graph
for P being oriented Chain of G
for W being Function of the carrier' of G,Real>=0
for v2, v1 being Vertex of G st f is_Input_of_Dijkstra_Alg G,n,W & W is_weight>=0of G & v2 = i & v1 <> v2 & 1 <= i & i <= n & P is_shortestpath_of v1,v2,V,W & ( for m, j being Element of NAT st g . (n + j) = - 1 & 1 <= j & j <= n & m in V holds
f . (((2 * n) + (n * m)) + j) = - 1 ) holds
g . (n + i) <> - 1
Lm20:
for n, k being Element of NAT
for f, g, h being Element of REAL *
for G being oriented finite Graph
for W being Function of the carrier' of G,Real>=0
for v1 being Vertex of G st f is_Input_of_Dijkstra_Alg G,n,W & v1 = 1 & n >= 1 & g = ((repeat ((Relax n) * (findmin n))) . k) . f & h = ((repeat ((Relax n) * (findmin n))) . (k + 1)) . f & OuterVx (g,n) <> {} & 1 in UsedVx (g,n) & ( for v3 being Vertex of G
for j being Element of NAT st v3 <> v1 & v3 = j & g . (n + j) <> - 1 holds
ex p being FinSequence of NAT ex P being oriented Chain of G st
( p is_simple_vertex_seq_at g,j,n & ( for i being Element of NAT st 1 <= i & i < len p holds
p . i in UsedVx (g,n) ) & P is_oriented_edge_seq_of p & P is_shortestpath_of v1,v3, UsedVx (g,n),W & cost (P,W) = g . ((2 * n) + j) & ( not v3 in UsedVx (g,n) implies P islongestInShortestpath UsedVx (g,n),v1,W ) ) ) & ( for m, j being Element of NAT st g . (n + j) = - 1 & 1 <= j & j <= n & m in UsedVx (g,n) holds
f . (((2 * n) + (n * m)) + j) = - 1 ) & ( for m being Element of NAT st m in UsedVx (g,n) holds
g . (n + m) <> - 1 ) holds
( ( for v3 being Vertex of G
for j being Element of NAT st v3 <> v1 & v3 = j & h . (n + j) <> - 1 holds
ex p being FinSequence of NAT ex P being oriented Chain of G st
( p is_simple_vertex_seq_at h,j,n & ( for i being Element of NAT st 1 <= i & i < len p holds
p . i in UsedVx (h,n) ) & P is_oriented_edge_seq_of p & P is_shortestpath_of v1,v3, UsedVx (h,n),W & cost (P,W) = h . ((2 * n) + j) & ( not v3 in UsedVx (h,n) implies P islongestInShortestpath UsedVx (h,n),v1,W ) ) ) & ( for m, j being Element of NAT st h . (n + j) = - 1 & 1 <= j & j <= n & m in UsedVx (h,n) holds
f . (((2 * n) + (n * m)) + j) = - 1 ) & ( for m being Element of NAT st m in UsedVx (h,n) holds
h . (n + m) <> - 1 ) )
theorem
for
n,
i being
Element of
NAT for
f,
g being
Element of
REAL * for
G being
oriented finite Graph for
W being
Function of the
carrier' of
G,
Real>=0 for
v1,
v2 being
Vertex of
G st
f is_Input_of_Dijkstra_Alg G,
n,
W &
v1 = 1 & 1
<> v2 &
v2 = i &
n >= 1 &
g = (DijkstraAlgorithm n) . f holds
( the
carrier of
G = (UsedVx (g,n)) \/ (UnusedVx (g,n)) & (
v2 in UsedVx (
g,
n) implies ex
p being
FinSequence of
NAT ex
P being
oriented Chain of
G st
(
p is_simple_vertex_seq_at g,
i,
n &
P is_oriented_edge_seq_of p &
P is_shortestpath_of v1,
v2,
W &
cost (
P,
W)
= g . ((2 * n) + i) ) ) & (
v2 in UnusedVx (
g,
n) implies for
Q being
oriented Chain of
G holds not
Q is_orientedpath_of v1,
v2 ) )