:: GRAPHSP semantic presentation
theorem Th1: :: GRAPHSP:1
theorem Th2: :: GRAPHSP:2
theorem Th3: :: GRAPHSP:3
theorem Th4: :: GRAPHSP:4
theorem Th5: :: GRAPHSP:5
Lemma6:
for b1 being Nat
for b2, b3 being FinSequence st 1 <= b1 & b1 <= len b2 holds
b2 . b1 = (b2 ^ b3) . b1
Lemma7:
for b1 being Nat
for b2, b3 being FinSequence st 1 <= b1 & b1 <= len b3 holds
b3 . b1 = (b2 ^ b3) . ((len b2) + b1)
theorem Th6: :: GRAPHSP:6
theorem Th7: :: GRAPHSP:7
theorem Th8: :: GRAPHSP:8
theorem Th9: :: GRAPHSP:9
for
b1 being
Graphfor
b2,
b3 being
oriented Chain of
b1for
b4 being
Functionfor
b5 being
set for
b6,
b7 being
Vertex of
b1 st
b2 is_shortestpath_of b6,
b7,
b5,
b4 &
b3 is_shortestpath_of b6,
b7,
b5,
b4 holds
cost b2,
b4 = cost b3,
b4
theorem Th10: :: GRAPHSP:10
Lemma13:
for b1 being Nat
for b2 being Graph
for b3 being FinSequence of the Edges of b2 st 1 <= b1 & b1 <= len b3 holds
( the Source of b2 . (b3 . b1) in the Vertices of b2 & the Target of b2 . (b3 . b1) in the Vertices of b2 )
theorem Th11: :: GRAPHSP:11
Lemma15:
for b1 being Nat
for b2 being Graph
for b3 being FinSequence of the Edges of b2
for b4 being Vertex of b2 st 1 <= b1 & b1 <= len b3 & ( b4 = the Source of b2 . (b3 . b1) or b4 = the Target of b2 . (b3 . b1) ) holds
b4 in vertices b3
theorem Th12: :: GRAPHSP:12
theorem Th13: :: GRAPHSP:13
for
b1 being
Functionfor
b2 being
set for
b3 being
finite Graphfor
b4,
b5 being
oriented Chain of
b3for
b6,
b7,
b8 being
Vertex of
b3 st
b1 is_weight>=0of b3 &
b4 is_shortestpath_of b6,
b7,
b2,
b1 &
b6 <> b7 &
b6 <> b8 &
b5 is_shortestpath_of b6,
b8,
b2,
b1 & ( for
b9 being
set holds
( not
b9 in the
Edges of
b3 or not
b9 orientedly_joins b7,
b8 ) ) &
b4 islongestInShortestpath b2,
b6,
b1 holds
b5 is_shortestpath_of b6,
b8,
b2 \/ {b7},
b1
theorem Th14: :: GRAPHSP:14
theorem Th15: :: GRAPHSP:15
for
b1,
b2 being
set for
b3 being
oriented finite Graphfor
b4,
b5 being
oriented Chain of
b3for
b6 being
Function of the
Edges of
b3,
Real>=0 for
b7,
b8,
b9 being
Vertex of
b3 st
b1 in the
Edges of
b3 &
b4 is_shortestpath_of b7,
b8,
b2,
b6 &
b7 <> b9 &
b5 = b4 ^ <*b1*> &
b1 orientedly_joins b8,
b9 &
b7 in b2 & ( for
b10 being
Vertex of
b3 st
b10 in b2 holds
for
b11 being
set holds
( not
b11 in the
Edges of
b3 or not
b11 orientedly_joins b10,
b9 ) ) holds
b5 is_shortestpath_of b7,
b9,
b2 \/ {b8},
b6
theorem Th16: :: GRAPHSP:16
for
b1,
b2 being
set for
b3 being
oriented finite Graphfor
b4 being
oriented Chain of
b3for
b5 being
Function of the
Edges of
b3,
Real>=0 for
b6,
b7 being
Vertex of
b3 st the
Vertices of
b3 = b1 \/ b2 &
b6 in b1 & ( for
b8,
b9 being
Vertex of
b3 st
b8 in b1 &
b9 in b2 holds
for
b10 being
set holds
( not
b10 in the
Edges of
b3 or not
b10 orientedly_joins b8,
b9 ) ) holds
(
b4 is_shortestpath_of b6,
b7,
b1,
b5 iff
b4 is_shortestpath_of b6,
b7,
b5 )
theorem Th17: :: GRAPHSP:17
:: deftheorem Def1 defines := GRAPHSP:def 1 :
theorem Th18: :: GRAPHSP:18
theorem Th19: :: GRAPHSP:19
theorem Th20: :: GRAPHSP:20
theorem Th21: :: GRAPHSP:21
:: deftheorem Def2 defines repeat GRAPHSP:def 2 :
theorem Th22: :: GRAPHSP:22
Lemma28:
for b1 being set
for b2 being Element of Funcs b1,b1 holds dom b2 = b1
theorem Th23: :: GRAPHSP:23
:: deftheorem Def3 defines OuterVx GRAPHSP:def 3 :
definition
let c1 be
Element of
Funcs (REAL * ),
(REAL * );
let c2 be
Element of
REAL * ;
let c3 be
Nat;
assume E30:
ex
b1 being
Nat st
OuterVx (((repeat c1) . b1) . c2),
c3 = {}
;
func LifeSpan c1,
c2,
c3 -> Nat means :
Def4:
:: GRAPHSP:def 4
(
OuterVx (((repeat a1) . a4) . a2),
a3 = {} & ( for
b1 being
Nat st
OuterVx (((repeat a1) . b1) . a2),
a3 = {} holds
a4 <= b1 ) );
existence
ex b1 being Nat st
( OuterVx (((repeat c1) . b1) . c2),c3 = {} & ( for b2 being Nat st OuterVx (((repeat c1) . b2) . c2),c3 = {} holds
b1 <= b2 ) )
uniqueness
for b1, b2 being Nat st OuterVx (((repeat c1) . b1) . c2),c3 = {} & ( for b3 being Nat st OuterVx (((repeat c1) . b3) . c2),c3 = {} holds
b1 <= b3 ) & OuterVx (((repeat c1) . b2) . c2),c3 = {} & ( for b3 being Nat st OuterVx (((repeat c1) . b3) . c2),c3 = {} holds
b2 <= b3 ) holds
b1 = b2
end;
:: deftheorem Def4 defines LifeSpan GRAPHSP:def 4 :
definition
let c1 be
Element of
Funcs (REAL * ),
(REAL * );
let c2 be
Nat;
func while_do c1,
c2 -> Element of
Funcs (REAL * ),
(REAL * ) means :
Def5:
:: GRAPHSP:def 5
(
dom a3 = REAL * & ( for
b1 being
Element of
REAL * holds
a3 . b1 = ((repeat a1) . (LifeSpan a1,b1,a2)) . b1 ) );
existence
ex b1 being Element of Funcs (REAL * ),(REAL * ) st
( dom b1 = REAL * & ( for b2 being Element of REAL * holds b1 . b2 = ((repeat c1) . (LifeSpan c1,b2,c2)) . b2 ) )
uniqueness
for b1, b2 being Element of Funcs (REAL * ),(REAL * ) st dom b1 = REAL * & ( for b3 being Element of REAL * holds b1 . b3 = ((repeat c1) . (LifeSpan c1,b3,c2)) . b3 ) & dom b2 = REAL * & ( for b3 being Element of REAL * holds b2 . b3 = ((repeat c1) . (LifeSpan c1,b3,c2)) . b3 ) holds
b1 = b2
end;
:: deftheorem Def5 defines while_do GRAPHSP:def 5 :
:: deftheorem Def6 defines Edge GRAPHSP:def 6 :
:: deftheorem Def7 defines Weight GRAPHSP:def 7 :
theorem Th24: :: GRAPHSP:24
theorem Th25: :: GRAPHSP:25
theorem Th26: :: GRAPHSP:26
:: deftheorem Def8 defines UnusedVx GRAPHSP:def 8 :
:: deftheorem Def9 defines UsedVx GRAPHSP:def 9 :
theorem Th27: :: GRAPHSP:27
theorem Th28: :: GRAPHSP:28
theorem Th29: :: GRAPHSP:29
definition
let c1 be
finite Subset of
NAT ;
let c2 be
Element of
REAL * ;
let c3 be
Nat;
func Argmin c1,
c2,
c3 -> Nat means :
Def10:
:: GRAPHSP:def 10
( (
a1 <> {} implies ex
b1 being
Nat st
(
b1 = a4 &
b1 in a1 & ( for
b2 being
Nat st
b2 in a1 holds
a2 /. ((2 * a3) + b1) <= a2 /. ((2 * a3) + b2) ) & ( for
b2 being
Nat st
b2 in a1 &
a2 /. ((2 * a3) + b1) = a2 /. ((2 * a3) + b2) holds
b1 <= b2 ) ) ) & (
a1 = {} implies
a4 = 0 ) );
existence
ex b1 being Nat st
( ( c1 <> {} implies ex b2 being Nat st
( b2 = b1 & b2 in c1 & ( for b3 being Nat st b3 in c1 holds
c2 /. ((2 * c3) + b2) <= c2 /. ((2 * c3) + b3) ) & ( for b3 being Nat st b3 in c1 & c2 /. ((2 * c3) + b2) = c2 /. ((2 * c3) + b3) holds
b2 <= b3 ) ) ) & ( c1 = {} implies b1 = 0 ) )
uniqueness
for b1, b2 being Nat st ( c1 <> {} implies ex b3 being Nat st
( b3 = b1 & b3 in c1 & ( for b4 being Nat st b4 in c1 holds
c2 /. ((2 * c3) + b3) <= c2 /. ((2 * c3) + b4) ) & ( for b4 being Nat st b4 in c1 & c2 /. ((2 * c3) + b3) = c2 /. ((2 * c3) + b4) holds
b3 <= b4 ) ) ) & ( c1 = {} implies b1 = 0 ) & ( c1 <> {} implies ex b3 being Nat st
( b3 = b2 & b3 in c1 & ( for b4 being Nat st b4 in c1 holds
c2 /. ((2 * c3) + b3) <= c2 /. ((2 * c3) + b4) ) & ( for b4 being Nat st b4 in c1 & c2 /. ((2 * c3) + b3) = c2 /. ((2 * c3) + b4) holds
b3 <= b4 ) ) ) & ( c1 = {} implies b2 = 0 ) holds
b1 = b2
end;
:: deftheorem Def10 defines Argmin GRAPHSP:def 10 :
theorem Th30: :: GRAPHSP:30
theorem Th31: :: GRAPHSP:31
definition
let c1 be
Nat;
func findmin c1 -> Element of
Funcs (REAL * ),
(REAL * ) means :
Def11:
:: GRAPHSP:def 11
(
dom a2 = REAL * & ( for
b1 being
Element of
REAL * holds
a2 . b1 = b1,
(((a1 * a1) + (3 * a1)) + 1) := (Argmin (OuterVx b1,a1),b1,a1),
(- 1) ) );
existence
ex b1 being Element of Funcs (REAL * ),(REAL * ) st
( dom b1 = REAL * & ( for b2 being Element of REAL * holds b1 . b2 = b2,(((c1 * c1) + (3 * c1)) + 1) := (Argmin (OuterVx b2,c1),b2,c1),(- 1) ) )
uniqueness
for b1, b2 being Element of Funcs (REAL * ),(REAL * ) st dom b1 = REAL * & ( for b3 being Element of REAL * holds b1 . b3 = b3,(((c1 * c1) + (3 * c1)) + 1) := (Argmin (OuterVx b3,c1),b3,c1),(- 1) ) & dom b2 = REAL * & ( for b3 being Element of REAL * holds b2 . b3 = b3,(((c1 * c1) + (3 * c1)) + 1) := (Argmin (OuterVx b3,c1),b3,c1),(- 1) ) holds
b1 = b2
end;
:: deftheorem Def11 defines findmin GRAPHSP:def 11 :
theorem Th32: :: GRAPHSP:32
theorem Th33: :: GRAPHSP:33
theorem Th34: :: GRAPHSP:34
Lemma46:
for b1, b2 being Nat st b1 >= 1 holds
b2 <= b1 * b2
Lemma47:
for b1 being Nat holds
( 3 * b1 < ((b1 * b1) + (3 * b1)) + 1 & b1 < ((b1 * b1) + (3 * b1)) + 1 & 2 * b1 < ((b1 * b1) + (3 * b1)) + 1 )
Lemma48:
for b1, b2 being Nat holds
( ( b1 < b2 & b2 <= 2 * b1 implies ( ( not 2 * b1 < b2 or not b2 <= 3 * b1 ) & not b2 <= b1 & not b2 > 3 * b1 ) ) & ( ( b2 <= b1 or b2 > 3 * b1 ) implies ( ( not 2 * b1 < b2 or not b2 <= 3 * b1 ) & ( not b1 < b2 or not b2 <= 2 * b1 ) ) ) & ( 2 * b1 < b2 & b2 <= 3 * b1 implies ( ( not b1 < b2 or not b2 <= 2 * b1 ) & not b2 <= b1 & not b2 > 3 * b1 ) ) )
theorem Th35: :: GRAPHSP:35
:: deftheorem Def12 defines newpathcost GRAPHSP:def 12 :
:: deftheorem Def13 defines hasBetterPathAt GRAPHSP:def 13 :
definition
let c1 be
Element of
REAL * ;
let c2 be
Nat;
func Relax c1,
c2 -> Element of
REAL * means :
Def14:
:: GRAPHSP:def 14
(
dom a3 = dom a1 & ( for
b1 being
Nat st
b1 in dom a1 holds
( (
a2 < b1 &
b1 <= 2
* a2 implies ( (
a1 hasBetterPathAt a2,
b1 -' a2 implies
a3 . b1 = a1 /. (((a2 * a2) + (3 * a2)) + 1) ) & ( not
a1 hasBetterPathAt a2,
b1 -' a2 implies
a3 . b1 = a1 . b1 ) ) ) & ( 2
* a2 < b1 &
b1 <= 3
* a2 implies ( (
a1 hasBetterPathAt a2,
b1 -' (2 * a2) implies
a3 . b1 = newpathcost a1,
a2,
(b1 -' (2 * a2)) ) & ( not
a1 hasBetterPathAt a2,
b1 -' (2 * a2) implies
a3 . b1 = a1 . b1 ) ) ) & ( (
b1 <= a2 or
b1 > 3
* a2 ) implies
a3 . b1 = a1 . b1 ) ) ) );
existence
ex b1 being Element of REAL * st
( dom b1 = dom c1 & ( for b2 being Nat st b2 in dom c1 holds
( ( c2 < b2 & b2 <= 2 * c2 implies ( ( c1 hasBetterPathAt c2,b2 -' c2 implies b1 . b2 = c1 /. (((c2 * c2) + (3 * c2)) + 1) ) & ( not c1 hasBetterPathAt c2,b2 -' c2 implies b1 . b2 = c1 . b2 ) ) ) & ( 2 * c2 < b2 & b2 <= 3 * c2 implies ( ( c1 hasBetterPathAt c2,b2 -' (2 * c2) implies b1 . b2 = newpathcost c1,c2,(b2 -' (2 * c2)) ) & ( not c1 hasBetterPathAt c2,b2 -' (2 * c2) implies b1 . b2 = c1 . b2 ) ) ) & ( ( b2 <= c2 or b2 > 3 * c2 ) implies b1 . b2 = c1 . b2 ) ) ) )
uniqueness
for b1, b2 being Element of REAL * st dom b1 = dom c1 & ( for b3 being Nat st b3 in dom c1 holds
( ( c2 < b3 & b3 <= 2 * c2 implies ( ( c1 hasBetterPathAt c2,b3 -' c2 implies b1 . b3 = c1 /. (((c2 * c2) + (3 * c2)) + 1) ) & ( not c1 hasBetterPathAt c2,b3 -' c2 implies b1 . b3 = c1 . b3 ) ) ) & ( 2 * c2 < b3 & b3 <= 3 * c2 implies ( ( c1 hasBetterPathAt c2,b3 -' (2 * c2) implies b1 . b3 = newpathcost c1,c2,(b3 -' (2 * c2)) ) & ( not c1 hasBetterPathAt c2,b3 -' (2 * c2) implies b1 . b3 = c1 . b3 ) ) ) & ( ( b3 <= c2 or b3 > 3 * c2 ) implies b1 . b3 = c1 . b3 ) ) ) & dom b2 = dom c1 & ( for b3 being Nat st b3 in dom c1 holds
( ( c2 < b3 & b3 <= 2 * c2 implies ( ( c1 hasBetterPathAt c2,b3 -' c2 implies b2 . b3 = c1 /. (((c2 * c2) + (3 * c2)) + 1) ) & ( not c1 hasBetterPathAt c2,b3 -' c2 implies b2 . b3 = c1 . b3 ) ) ) & ( 2 * c2 < b3 & b3 <= 3 * c2 implies ( ( c1 hasBetterPathAt c2,b3 -' (2 * c2) implies b2 . b3 = newpathcost c1,c2,(b3 -' (2 * c2)) ) & ( not c1 hasBetterPathAt c2,b3 -' (2 * c2) implies b2 . b3 = c1 . b3 ) ) ) & ( ( b3 <= c2 or b3 > 3 * c2 ) implies b2 . b3 = c1 . b3 ) ) ) holds
b1 = b2
end;
:: deftheorem Def14 defines Relax GRAPHSP:def 14 :
:: deftheorem Def15 defines Relax GRAPHSP:def 15 :
theorem Th36: :: GRAPHSP:36
theorem Th37: :: GRAPHSP:37
theorem Th38: :: GRAPHSP:38
theorem Th39: :: GRAPHSP:39
theorem Th40: :: GRAPHSP:40
for
b1,
b2,
b3 being
Natfor
b4,
b5,
b6 being
Element of
REAL * st
b4 = ((repeat ((Relax b1) * (findmin b1))) . b2) . b5 &
b6 = ((repeat ((Relax b1) * (findmin b1))) . (b2 + 1)) . b5 &
b3 = Argmin (OuterVx b4,b1),
b4,
b1 &
OuterVx b4,
b1 <> {} holds
(
UsedVx b6,
b1 = (UsedVx b4,b1) \/ {b3} & not
b3 in UsedVx b4,
b1 )
theorem Th41: :: GRAPHSP:41
Lemma59:
for b1, b2 being Nat holds b1 - b2 <= b1
Lemma60:
for b1, b2 being FinSequence of NAT
for b3 being Element of REAL *
for b4, b5 being Nat st ( for b6 being Nat st 1 <= b6 & b6 < len b1 holds
b1 . ((len b1) - b6) = b3 . ((b1 /. (((len b1) - b6) + 1)) + b5) ) & ( for b6 being Nat st 1 <= b6 & b6 < len b2 holds
b2 . ((len b2) - b6) = b3 . ((b2 /. (((len b2) - b6) + 1)) + b5) ) & len b1 <= len b2 & b1 . (len b1) = b2 . (len b2) holds
for b6 being Nat st 1 <= b6 & b6 < len b1 holds
b1 . ((len b1) - b6) = b2 . ((len b2) - b6)
theorem Th42: :: GRAPHSP:42
:: deftheorem Def16 defines equal_at GRAPHSP:def 16 :
theorem Th43: :: GRAPHSP:43
theorem Th44: :: GRAPHSP:44
theorem Th45: :: GRAPHSP:45
theorem Th46: :: GRAPHSP:46
theorem Th47: :: GRAPHSP:47
theorem Th48: :: GRAPHSP:48
theorem Th49: :: GRAPHSP:49
:: deftheorem Def17 defines is_vertex_seq_at GRAPHSP:def 17 :
:: deftheorem Def18 defines is_simple_vertex_seq_at GRAPHSP:def 18 :
theorem Th50: :: GRAPHSP:50
:: deftheorem Def19 defines is_oriented_edge_seq_of GRAPHSP:def 19 :
theorem Th51: :: GRAPHSP:51
theorem Th52: :: GRAPHSP:52
Lemma72:
for b1, b2 being Nat st 1 <= b1 & b1 <= b2 holds
( 1 < (2 * b2) + b1 & (2 * b2) + b1 < ((b2 * b2) + (3 * b2)) + 1 & b1 < (2 * b2) + b1 )
Lemma73:
for b1, b2 being Nat st 1 <= b1 & b1 <= b2 holds
( 1 < b2 + b1 & b2 + b1 <= 2 * b2 & b2 + b1 < ((b2 * b2) + (3 * b2)) + 1 )
Lemma74:
for b1, b2, b3 being Nat st 1 <= b1 & b1 <= b2 & b3 <= b2 holds
( 1 < ((2 * b2) + (b2 * b1)) + b3 & b1 < ((2 * b2) + (b2 * b1)) + b3 & ((2 * b2) + (b2 * b1)) + b3 < ((b2 * b2) + (3 * b2)) + 1 )
Lemma75:
for b1, b2, b3 being Nat st 1 <= b1 & b1 <= b2 & 1 <= b3 & b3 <= b2 holds
( (3 * b2) + 1 <= ((2 * b2) + (b2 * b1)) + b3 & ((2 * b2) + (b2 * b1)) + b3 <= (b2 * b2) + (3 * b2) )
:: deftheorem Def20 defines is_Input_of_Dijkstra_Alg GRAPHSP:def 20 :
:: deftheorem Def21 defines DijkstraAlgorithm GRAPHSP:def 21 :
Lemma77:
for b1 being Nat
for b2, b3 being Element of REAL *
for b4 being oriented finite Graph
for b5 being Function of the Edges of b4, Real>=0
for b6 being Vertex of b4 st b2 is_Input_of_Dijkstra_Alg b4,b1,b5 & b6 = 1 & b1 >= 1 & b3 = ((repeat ((Relax b1) * (findmin b1))) . 1) . b2 holds
( ( for b7 being Vertex of b4
for b8 being Nat st b7 <> b6 & b7 = b8 & b3 . (b1 + b8) <> - 1 holds
ex b9 being FinSequence of NAT ex b10 being oriented Chain of b4 st
( b9 is_simple_vertex_seq_at b3,b8,b1 & ( for b11 being Nat st 1 <= b11 & b11 < len b9 holds
b9 . b11 in UsedVx b3,b1 ) & b10 is_oriented_edge_seq_of b9 & b10 is_shortestpath_of b6,b7, UsedVx b3,b1,b5 & cost b10,b5 = b3 . ((2 * b1) + b8) & ( not b7 in UsedVx b3,b1 implies b10 islongestInShortestpath UsedVx b3,b1,b6,b5 ) ) ) & ( for b7, b8 being Nat st b3 . (b1 + b8) = - 1 & 1 <= b8 & b8 <= b1 & b7 in UsedVx b3,b1 holds
b2 . (((2 * b1) + (b1 * b7)) + b8) = - 1 ) & ( for b7 being Nat st b7 in UsedVx b3,b1 holds
b3 . (b1 + b7) <> - 1 ) )
Lemma78:
for b1, b2, b3 being Nat
for b4, b5, b6 being Element of REAL * st b4 = ((repeat ((Relax b1) * (findmin b1))) . b2) . b5 & b6 = ((repeat ((Relax b1) * (findmin b1))) . (b2 + 1)) . b5 & OuterVx b4,b1 <> {} & b3 in UsedVx b4,b1 & len b5 = ((b1 * b1) + (3 * b1)) + 1 holds
b6 . (b1 + b3) = b4 . (b1 + b3)
Lemma79:
for b1, b2, b3 being Nat
for b4, b5, b6 being Element of REAL *
for b7 being FinSequence of NAT st b4 = ((repeat ((Relax b1) * (findmin b1))) . b2) . b5 & b6 = ((repeat ((Relax b1) * (findmin b1))) . (b2 + 1)) . b5 & OuterVx b4,b1 <> {} & len b5 = ((b1 * b1) + (3 * b1)) + 1 & b7 is_simple_vertex_seq_at b4,b3,b1 & b4 . (b1 + b3) = b6 . (b1 + b3) & ( for b8 being Nat st 1 <= b8 & b8 < len b7 holds
b7 . b8 in UsedVx b4,b1 ) holds
b7 is_simple_vertex_seq_at b6,b3,b1
Lemma80:
for b1, b2, b3, b4 being Nat
for b5, b6, b7 being Element of REAL *
for b8 being FinSequence of NAT st b5 = ((repeat ((Relax b1) * (findmin b1))) . b2) . b6 & b7 = ((repeat ((Relax b1) * (findmin b1))) . (b2 + 1)) . b6 & OuterVx b5,b1 <> {} & len b6 = ((b1 * b1) + (3 * b1)) + 1 & b8 is_simple_vertex_seq_at b5,b3,b1 & b3 = b7 . (b1 + b4) & b5 . (b1 + b3) = b7 . (b1 + b3) & b3 <> b4 & not b4 in UsedVx b5,b1 & ( for b9 being Nat st 1 <= b9 & b9 < len b8 holds
b8 . b9 in UsedVx b5,b1 ) holds
b8 ^ <*b4*> is_simple_vertex_seq_at b7,b4,b1
Lemma81:
for b1, b2 being Nat
for b3 being set
for b4, b5 being Element of REAL *
for b6 being oriented finite Graph
for b7 being oriented Chain of b6
for b8 being Function of the Edges of b6, Real>=0
for b9, b10 being Vertex of b6 st b4 is_Input_of_Dijkstra_Alg b6,b1,b8 & b8 is_weight>=0of b6 & b9 = b2 & b10 <> b9 & 1 <= b2 & b2 <= b1 & b7 is_shortestpath_of b10,b9,b3,b8 & ( for b11, b12 being Nat st b5 . (b1 + b12) = - 1 & 1 <= b12 & b12 <= b1 & b11 in b3 holds
b4 . (((2 * b1) + (b1 * b11)) + b12) = - 1 ) holds
b5 . (b1 + b2) <> - 1
Lemma82:
for b1, b2 being Nat
for b3, b4, b5 being Element of REAL *
for b6 being oriented finite Graph
for b7 being Function of the Edges of b6, Real>=0
for b8 being Vertex of b6 st b3 is_Input_of_Dijkstra_Alg b6,b1,b7 & b8 = 1 & b1 >= 1 & b4 = ((repeat ((Relax b1) * (findmin b1))) . b2) . b3 & b5 = ((repeat ((Relax b1) * (findmin b1))) . (b2 + 1)) . b3 & OuterVx b4,b1 <> {} & b2 >= 1 & 1 in UsedVx b4,b1 & ( for b9 being Vertex of b6
for b10 being Nat st b9 <> b8 & b9 = b10 & b4 . (b1 + b10) <> - 1 holds
ex b11 being FinSequence of NAT ex b12 being oriented Chain of b6 st
( b11 is_simple_vertex_seq_at b4,b10,b1 & ( for b13 being Nat st 1 <= b13 & b13 < len b11 holds
b11 . b13 in UsedVx b4,b1 ) & b12 is_oriented_edge_seq_of b11 & b12 is_shortestpath_of b8,b9, UsedVx b4,b1,b7 & cost b12,b7 = b4 . ((2 * b1) + b10) & ( not b9 in UsedVx b4,b1 implies b12 islongestInShortestpath UsedVx b4,b1,b8,b7 ) ) ) & ( for b9, b10 being Nat st b4 . (b1 + b10) = - 1 & 1 <= b10 & b10 <= b1 & b9 in UsedVx b4,b1 holds
b3 . (((2 * b1) + (b1 * b9)) + b10) = - 1 ) & ( for b9 being Nat st b9 in UsedVx b4,b1 holds
b4 . (b1 + b9) <> - 1 ) holds
( ( for b9 being Vertex of b6
for b10 being Nat st b9 <> b8 & b9 = b10 & b5 . (b1 + b10) <> - 1 holds
ex b11 being FinSequence of NAT ex b12 being oriented Chain of b6 st
( b11 is_simple_vertex_seq_at b5,b10,b1 & ( for b13 being Nat st 1 <= b13 & b13 < len b11 holds
b11 . b13 in UsedVx b5,b1 ) & b12 is_oriented_edge_seq_of b11 & b12 is_shortestpath_of b8,b9, UsedVx b5,b1,b7 & cost b12,b7 = b5 . ((2 * b1) + b10) & ( not b9 in UsedVx b5,b1 implies b12 islongestInShortestpath UsedVx b5,b1,b8,b7 ) ) ) & ( for b9, b10 being Nat st b5 . (b1 + b10) = - 1 & 1 <= b10 & b10 <= b1 & b9 in UsedVx b5,b1 holds
b3 . (((2 * b1) + (b1 * b9)) + b10) = - 1 ) & ( for b9 being Nat st b9 in UsedVx b5,b1 holds
b5 . (b1 + b9) <> - 1 ) )
theorem Th53: :: GRAPHSP:53
for
b1,
b2 being
Natfor
b3,
b4 being
Element of
REAL * for
b5 being
oriented finite Graphfor
b6 being
Function of the
Edges of
b5,
Real>=0 for
b7,
b8 being
Vertex of
b5 st
b3 is_Input_of_Dijkstra_Alg b5,
b1,
b6 &
b7 = 1 & 1
<> b8 &
b8 = b2 &
b1 >= 1 &
b4 = (DijkstraAlgorithm b1) . b3 holds
( the
Vertices of
b5 = (UsedVx b4,b1) \/ (UnusedVx b4,b1) & (
b8 in UsedVx b4,
b1 implies ex
b9 being
FinSequence of
NAT ex
b10 being
oriented Chain of
b5 st
(
b9 is_simple_vertex_seq_at b4,
b2,
b1 &
b10 is_oriented_edge_seq_of b9 &
b10 is_shortestpath_of b7,
b8,
b6 &
cost b10,
b6 = b4 . ((2 * b1) + b2) ) ) & (
b8 in UnusedVx b4,
b1 implies for
b9 being
oriented Chain of
b5 holds not
b9 is_orientedpath_of b7,
b8 ) )