:: GRAPHSP semantic presentation

theorem Th1: :: GRAPHSP:1
for b1 being FinSequence
for b2 being set holds
( ( not b2 in rng b1 & b1 is one-to-one ) iff b1 ^ <*b2*> is one-to-one )
proof end;

theorem Th2: :: GRAPHSP:2
for b1 being set
for b2 being FinSequence of b1
for b3 being Integer holds
( 1 <= b3 & b3 <= len b2 implies b2 . b3 in b1 )
proof end;

theorem Th3: :: GRAPHSP:3
for b1 being set
for b2 being FinSequence of b1
for b3 being Integer holds
( 1 <= b3 & b3 <= len b2 implies b2 /. b3 = b2 . b3 )
proof end;

theorem Th4: :: GRAPHSP:4
for b1 being Graph
for b2 being FinSequence of the Edges of b1
for b3 being Function holds
( b3 is_weight_of b1 & len b2 = 1 implies cost b2,b3 = b3 . (b2 . 1) )
proof end;

theorem Th5: :: GRAPHSP:5
for b1 being Graph
for b2 being set holds
( b2 in the Edges of b1 implies <*b2*> is oriented Simple Chain of b1 )
proof end;

Lemma6: for b1 being Nat
for b2, b3 being FinSequence holds
( 1 <= b1 & b1 <= len b2 implies b2 . b1 = (b2 ^ b3) . b1 )
proof end;

Lemma7: for b1 being Nat
for b2, b3 being FinSequence holds
( 1 <= b1 & b1 <= len b3 implies b3 . b1 = (b2 ^ b3) . ((len b2) + b1) )
proof end;

theorem Th6: :: GRAPHSP:6
for b1 being Graph
for b2, b3 being FinSequence of the Edges of b1
for b4 being oriented Simple Chain of b1 holds
( b4 = b2 ^ b3 & len b2 >= 1 & len b3 >= 1 implies ( the Target of b1 . (b4 . (len b4)) <> the Target of b1 . (b2 . (len b2)) & the Source of b1 . (b4 . 1) <> the Source of b1 . (b3 . 1) ) )
proof end;

theorem Th7: :: GRAPHSP:7
for b1 being Graph
for b2 being oriented Chain of b1
for b3 being set
for b4, b5 being Vertex of b1 holds
( b2 is_orientedpath_of b4,b5,b3 iff b2 is_orientedpath_of b4,b5,b3 \/ {b5} )
proof end;

theorem Th8: :: GRAPHSP:8
for b1 being Graph
for b2 being oriented Chain of b1
for b3 being Function
for b4 being set
for b5, b6 being Vertex of b1 holds
( b2 is_shortestpath_of b5,b6,b4,b3 iff b2 is_shortestpath_of b5,b6,b4 \/ {b6},b3 )
proof end;

theorem Th9: :: GRAPHSP:9
for b1 being Graph
for b2, b3 being oriented Chain of b1
for b4 being Function
for b5 being set
for b6, b7 being Vertex of b1 holds
( b2 is_shortestpath_of b6,b7,b5,b4 & b3 is_shortestpath_of b6,b7,b5,b4 implies cost b2,b4 = cost b3,b4 )
proof end;

theorem Th10: :: GRAPHSP:10
for b1 being oriented Graph
for b2, b3 being Vertex of b1
for b4, b5 being set holds
( b4 in the Edges of b1 & b5 in the Edges of b1 & b4 orientedly_joins b2,b3 & b5 orientedly_joins b2,b3 implies b4 = b5 )
proof end;

Lemma13: for b1 being Nat
for b2 being Graph
for b3 being FinSequence of the Edges of b2 holds
( 1 <= b1 & b1 <= len b3 implies ( the Source of b2 . (b3 . b1) in the Vertices of b2 & the Target of b2 . (b3 . b1) in the Vertices of b2 ) )
proof end;

theorem Th11: :: GRAPHSP:11
for b1 being Graph
for b2, b3 being set
for b4, b5 being Vertex of b1 holds
( the Vertices of b1 = b2 \/ b3 & b4 in b2 & b5 in b3 & ( for b6, b7 being Vertex of b1 holds
( b6 in b2 & b7 in b3 implies for b8 being set holds
not ( b8 in the Edges of b1 & b8 orientedly_joins b6,b7 ) ) ) implies for b6 being oriented Chain of b1 holds
not b6 is_orientedpath_of b4,b5 )
proof end;

Lemma15: for b1 being Nat
for b2 being Graph
for b3 being FinSequence of the Edges of b2
for b4 being Vertex of b2 holds
( 1 <= b1 & b1 <= len b3 & ( b4 = the Source of b2 . (b3 . b1) or b4 = the Target of b2 . (b3 . b1) ) implies b4 in vertices b3 )
proof end;

theorem Th12: :: GRAPHSP:12
for b1 being Graph
for b2 being oriented Chain of b1
for b3, b4 being set
for b5, b6 being Vertex of b1 holds
( the Vertices of b1 = b3 \/ b4 & b5 in b3 & ( for b7, b8 being Vertex of b1 holds
( b7 in b3 & b8 in b4 implies for b9 being set holds
not ( b9 in the Edges of b1 & b9 orientedly_joins b7,b8 ) ) ) & b2 is_orientedpath_of b5,b6 implies b2 is_orientedpath_of b5,b6,b3 )
proof end;

theorem Th13: :: GRAPHSP:13
for b1 being Function
for b2 being set
for b3 being finite Graph
for b4, b5 being oriented Chain of b3
for b6, b7, b8 being Vertex of b3 holds
( 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 & b9 orientedly_joins b7,b8 ) ) & b4 islongestInShortestpath b2,b6,b1 implies b5 is_shortestpath_of b6,b8,b2 \/ {b7},b1 )
proof end;

theorem Th14: :: GRAPHSP:14
for b1 being set
for b2 being oriented finite Graph
for b3 being oriented Chain of b2
for b4 being Function of the Edges of b2, Real>=0
for b5, b6 being Vertex of b2 holds
( b1 in the Edges of b2 & b5 <> b6 & b3 = <*b1*> & b1 orientedly_joins b5,b6 implies b3 is_shortestpath_of b5,b6,{b5},b4 )
proof end;

theorem Th15: :: GRAPHSP:15
for b1, b2 being set
for b3 being oriented finite Graph
for b4, b5 being oriented Chain of b3
for b6 being Function of the Edges of b3, Real>=0
for b7, b8, b9 being Vertex of b3 holds
( 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 holds
( b10 in b2 implies for b11 being set holds
not ( b11 in the Edges of b3 & b11 orientedly_joins b10,b9 ) ) ) implies b5 is_shortestpath_of b7,b9,b2 \/ {b8},b6 )
proof end;

theorem Th16: :: GRAPHSP:16
for b1, b2 being set
for b3 being oriented finite Graph
for b4 being oriented Chain of b3
for b5 being Function of the Edges of b3, Real>=0
for b6, b7 being Vertex of b3 holds
( the Vertices of b3 = b1 \/ b2 & b6 in b1 & ( for b8, b9 being Vertex of b3 holds
( b8 in b1 & b9 in b2 implies for b10 being set holds
not ( b10 in the Edges of b3 & b10 orientedly_joins b8,b9 ) ) ) implies ( b4 is_shortestpath_of b6,b7,b1,b5 iff b4 is_shortestpath_of b6,b7,b5 ) )
proof end;

notation
let c1 be Function;
let c2, c3 be set ;
synonym c1,c2 := c3 for c1 +* c2,c3;
end;

theorem Th17: :: GRAPHSP:17
for b1, b2 being set
for b3 being Function holds rng (b3,b1 := b2) c= (rng b3) \/ {b2}
proof end;

definition
let c1 be FinSequence of REAL ;
let c2 be set ;
let c3 be Real;
redefine func := as c1,c2 := c3 -> FinSequence of REAL ;
coherence
c1,c2 := c3 is FinSequence of REAL
proof end;
end;

definition
let c1, c2 be Nat;
let c3 be FinSequence of REAL ;
let c4 be Real;
func c3,c1 := c2,c4 -> FinSequence of REAL equals :: GRAPHSP:def 1
(a3,a1 := a2),a2 := a4;
coherence
(c3,c1 := c2),c2 := c4 is FinSequence of REAL
;
end;

:: deftheorem Def1 defines := GRAPHSP:def 1 :
for b1, b2 being Nat
for b3 being FinSequence of REAL
for b4 being Real holds b3,b1 := b2,b4 = (b3,b1 := b2),b2 := b4;

theorem Th18: :: GRAPHSP:18
for b1, b2 being Nat
for b3 being Element of REAL *
for b4 being Real holds
( b1 <> b2 & b1 in dom b3 implies (b3,b1 := b2,b4) . b1 = b2 )
proof end;

theorem Th19: :: GRAPHSP:19
for b1, b2, b3 being Nat
for b4 being Element of REAL *
for b5 being Real holds
( b1 <> b2 & b1 <> b3 & b1 in dom b4 implies (b4,b2 := b3,b5) . b1 = b4 . b1 )
proof end;

theorem Th20: :: GRAPHSP:20
for b1, b2 being Nat
for b3 being Element of REAL *
for b4 being Real holds
( b1 in dom b3 implies (b3,b2 := b1,b4) . b1 = b4 )
proof end;

theorem Th21: :: GRAPHSP:21
for b1, b2 being Nat
for b3 being Element of REAL *
for b4 being Real holds dom (b3,b1 := b2,b4) = dom b3
proof end;

definition
let c1 be set ;
redefine func id as id c1 -> Element of Funcs a1,a1;
coherence
id c1 is Element of Funcs c1,c1
proof end;
end;

definition
let c1 be set ;
let c2, c3 be Element of Funcs c1,c1;
redefine func * as c3 * c2 -> Element of Funcs a1,a1;
coherence
c2 * c3 is Element of Funcs c1,c1
proof end;
end;

definition
let c1 be set ;
let c2 be Element of Funcs c1,c1;
let c3 be Element of c1;
redefine func . as c2 . c3 -> Element of a1;
coherence
c2 . c3 is Element of c1
proof end;
end;

definition
let c1 be set ;
let c2 be Element of Funcs c1,c1;
func repeat c2 -> Function of NAT , Funcs a1,a1 means :Def2: :: GRAPHSP:def 2
( a3 . 0 = id a1 & ( for b1 being Nat holds a3 . (b1 + 1) = a2 * (a3 . b1) ) );
existence
ex b1 being Function of NAT , Funcs c1,c1 st
( b1 . 0 = id c1 & ( for b2 being Nat holds b1 . (b2 + 1) = c2 * (b1 . b2) ) )
proof end;
uniqueness
for b1, b2 being Function of NAT , Funcs c1,c1 holds
( b1 . 0 = id c1 & ( for b3 being Nat holds b1 . (b3 + 1) = c2 * (b1 . b3) ) & b2 . 0 = id c1 & ( for b3 being Nat holds b2 . (b3 + 1) = c2 * (b2 . b3) ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def2 defines repeat GRAPHSP:def 2 :
for b1 being set
for b2 being Element of Funcs b1,b1
for b3 being Function of NAT , Funcs b1,b1 holds
( b3 = repeat b2 iff ( b3 . 0 = id b1 & ( for b4 being Nat holds b3 . (b4 + 1) = b2 * (b3 . b4) ) ) );

theorem Th22: :: GRAPHSP:22
for b1 being Element of Funcs (REAL * ),(REAL * )
for b2 being Element of REAL *
for b3, b4 being Nat holds ((repeat b1) . 0) . b2 = b2
proof end;

Lemma28: for b1 being set
for b2 being Element of Funcs b1,b1 holds dom b2 = b1
proof end;

theorem Th23: :: GRAPHSP:23
for b1, b2 being Element of Funcs (REAL * ),(REAL * )
for b3 being Element of REAL *
for b4 being Nat holds ((repeat (b1 * b2)) . (b4 + 1)) . b3 = b1 . (b2 . (((repeat (b1 * b2)) . b4) . b3))
proof end;

definition
let c1 be Element of Funcs (REAL * ),(REAL * );
let c2 be Element of REAL * ;
redefine func . as c1 . c2 -> Element of REAL * ;
coherence
c1 . c2 is Element of REAL *
proof end;
end;

definition
let c1 be Element of REAL * ;
let c2 be Nat;
func OuterVx c1,c2 -> Subset of NAT equals :: GRAPHSP:def 3
{ b1 where B is Nat : ( b1 in dom a1 & 1 <= b1 & b1 <= a2 & a1 . b1 <> - 1 & a1 . (a2 + b1) <> - 1 ) } ;
coherence
{ b1 where B is Nat : ( b1 in dom c1 & 1 <= b1 & b1 <= c2 & c1 . b1 <> - 1 & c1 . (c2 + b1) <> - 1 ) } is Subset of NAT
proof end;
end;

:: deftheorem Def3 defines OuterVx GRAPHSP:def 3 :
for b1 being Element of REAL *
for b2 being Nat holds OuterVx b1,b2 = { b3 where B is Nat : ( b3 in dom b1 & 1 <= b3 & b3 <= b2 & b1 . b3 <> - 1 & b1 . (b2 + b3) <> - 1 ) } ;

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 holds
( OuterVx (((repeat a1) . b1) . a2),a3 = {} implies a4 <= b1 ) ) );
existence
ex b1 being Nat st
( OuterVx (((repeat c1) . b1) . c2),c3 = {} & ( for b2 being Nat holds
( OuterVx (((repeat c1) . b2) . c2),c3 = {} implies b1 <= b2 ) ) )
proof end;
uniqueness
for b1, b2 being Nat holds
( OuterVx (((repeat c1) . b1) . c2),c3 = {} & ( for b3 being Nat holds
( OuterVx (((repeat c1) . b3) . c2),c3 = {} implies b1 <= b3 ) ) & OuterVx (((repeat c1) . b2) . c2),c3 = {} & ( for b3 being Nat holds
( OuterVx (((repeat c1) . b3) . c2),c3 = {} implies b2 <= b3 ) ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def4 defines LifeSpan GRAPHSP:def 4 :
for b1 being Element of Funcs (REAL * ),(REAL * )
for b2 being Element of REAL *
for b3 being Nat holds
( ex b4 being Nat st OuterVx (((repeat b1) . b4) . b2),b3 = {} implies for b4 being Nat holds
( b4 = LifeSpan b1,b2,b3 iff ( OuterVx (((repeat b1) . b4) . b2),b3 = {} & ( for b5 being Nat holds
( OuterVx (((repeat b1) . b5) . b2),b3 = {} implies b4 <= b5 ) ) ) ) );

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 ) )
proof end;
uniqueness
for b1, b2 being Element of Funcs (REAL * ),(REAL * ) holds
( 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 ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def5 defines while_do GRAPHSP:def 5 :
for b1 being Element of Funcs (REAL * ),(REAL * )
for b2 being Nat
for b3 being Element of Funcs (REAL * ),(REAL * ) holds
( b3 = while_do b1,b2 iff ( dom b3 = REAL * & ( for b4 being Element of REAL * holds b3 . b4 = ((repeat b1) . (LifeSpan b1,b4,b2)) . b4 ) ) );

definition
let c1 be oriented Graph;
let c2, c3 be Vertex of c1;
assume E32: ex b1 being set st
( b1 in the Edges of c1 & b1 orientedly_joins c2,c3 ) ;
func Edge c2,c3 -> set means :Def6: :: GRAPHSP:def 6
ex b1 being set st
( a4 = b1 & b1 in the Edges of a1 & b1 orientedly_joins a2,a3 );
existence
ex b1 being set ex b2 being set st
( b1 = b2 & b2 in the Edges of c1 & b2 orientedly_joins c2,c3 )
by E32;
uniqueness
for b1, b2 being set holds
( ex b3 being set st
( b1 = b3 & b3 in the Edges of c1 & b3 orientedly_joins c2,c3 ) & ex b3 being set st
( b2 = b3 & b3 in the Edges of c1 & b3 orientedly_joins c2,c3 ) implies b1 = b2 )
by Th10;
end;

:: deftheorem Def6 defines Edge GRAPHSP:def 6 :
for b1 being oriented Graph
for b2, b3 being Vertex of b1 holds
( ex b4 being set st
( b4 in the Edges of b1 & b4 orientedly_joins b2,b3 ) implies for b4 being set holds
( b4 = Edge b2,b3 iff ex b5 being set st
( b4 = b5 & b5 in the Edges of b1 & b5 orientedly_joins b2,b3 ) ) );

definition
let c1 be oriented Graph;
let c2, c3 be Vertex of c1;
let c4 be Function;
func Weight c2,c3,c4 -> set equals :Def7: :: GRAPHSP:def 7
a4 . (Edge a2,a3) if ex b1 being set st
( b1 in the Edges of a1 & b1 orientedly_joins a2,a3 )
otherwise - 1;
correctness
coherence
( ( ex b1 being set st
( b1 in the Edges of c1 & b1 orientedly_joins c2,c3 ) implies c4 . (Edge c2,c3) is set ) & ( ( for b1 being set holds
not ( b1 in the Edges of c1 & b1 orientedly_joins c2,c3 ) ) implies - 1 is set ) )
;
consistency
for b1 being set holds
verum
;
;
end;

:: deftheorem Def7 defines Weight GRAPHSP:def 7 :
for b1 being oriented Graph
for b2, b3 being Vertex of b1
for b4 being Function holds
( ( ex b5 being set st
( b5 in the Edges of b1 & b5 orientedly_joins b2,b3 ) implies Weight b2,b3,b4 = b4 . (Edge b2,b3) ) & ( ( for b5 being set holds
not ( b5 in the Edges of b1 & b5 orientedly_joins b2,b3 ) ) implies Weight b2,b3,b4 = - 1 ) );

definition
let c1 be oriented Graph;
let c2, c3 be Vertex of c1;
let c4 be Function of the Edges of c1, Real>=0 ;
redefine func Weight as Weight c2,c3,c4 -> Real;
coherence
Weight c2,c3,c4 is Real
proof end;
end;

theorem Th24: :: GRAPHSP:24
for b1 being oriented Graph
for b2, b3 being Vertex of b1
for b4 being Function of the Edges of b1, Real>=0 holds
( Weight b2,b3,b4 >= 0 iff ex b5 being set st
( b5 in the Edges of b1 & b5 orientedly_joins b2,b3 ) )
proof end;

theorem Th25: :: GRAPHSP:25
for b1 being oriented Graph
for b2, b3 being Vertex of b1
for b4 being Function of the Edges of b1, Real>=0 holds
( Weight b2,b3,b4 = - 1 iff for b5 being set holds
not ( b5 in the Edges of b1 & b5 orientedly_joins b2,b3 ) ) by Def7, Th24;

theorem Th26: :: GRAPHSP:26
for b1 being set
for b2 being oriented Graph
for b3, b4 being Vertex of b2
for b5 being Function of the Edges of b2, Real>=0 holds
( b1 in the Edges of b2 & b1 orientedly_joins b3,b4 implies Weight b3,b4,b5 = b5 . b1 )
proof end;

definition
let c1 be Element of REAL * ;
let c2 be Nat;
func UnusedVx c1,c2 -> Subset of NAT equals :: GRAPHSP:def 8
{ b1 where B is Nat : ( b1 in dom a1 & 1 <= b1 & b1 <= a2 & a1 . b1 <> - 1 ) } ;
coherence
{ b1 where B is Nat : ( b1 in dom c1 & 1 <= b1 & b1 <= c2 & c1 . b1 <> - 1 ) } is Subset of NAT
proof end;
end;

:: deftheorem Def8 defines UnusedVx GRAPHSP:def 8 :
for b1 being Element of REAL *
for b2 being Nat holds UnusedVx b1,b2 = { b3 where B is Nat : ( b3 in dom b1 & 1 <= b3 & b3 <= b2 & b1 . b3 <> - 1 ) } ;

definition
let c1 be Element of REAL * ;
let c2 be Nat;
func UsedVx c1,c2 -> Subset of NAT equals :: GRAPHSP:def 9
{ b1 where B is Nat : ( b1 in dom a1 & 1 <= b1 & b1 <= a2 & a1 . b1 = - 1 ) } ;
coherence
{ b1 where B is Nat : ( b1 in dom c1 & 1 <= b1 & b1 <= c2 & c1 . b1 = - 1 ) } is Subset of NAT
proof end;
end;

:: deftheorem Def9 defines UsedVx GRAPHSP:def 9 :
for b1 being Element of REAL *
for b2 being Nat holds UsedVx b1,b2 = { b3 where B is Nat : ( b3 in dom b1 & 1 <= b3 & b3 <= b2 & b1 . b3 = - 1 ) } ;

theorem Th27: :: GRAPHSP:27
for b1 being Nat
for b2 being Element of REAL * holds UnusedVx b2,b1 c= Seg b1
proof end;

registration
let c1 be Element of REAL * ;
let c2 be Nat;
cluster UnusedVx a1,a2 -> finite ;
coherence
UnusedVx c1,c2 is finite
proof end;
end;

theorem Th28: :: GRAPHSP:28
for b1 being Nat
for b2 being Element of REAL * holds OuterVx b2,b1 c= UnusedVx b2,b1
proof end;

theorem Th29: :: GRAPHSP:29
for b1 being Nat
for b2 being Element of REAL * holds OuterVx b2,b1 c= Seg b1
proof end;

registration
let c1 be Element of REAL * ;
let c2 be Nat;
cluster OuterVx a1,a2 -> finite ;
coherence
OuterVx c1,c2 is finite
proof end;
end;

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
( not ( a1 <> {} & ( for b1 being Nat holds
not ( b1 = a4 & b1 in a1 & ( for b2 being Nat holds
( b2 in a1 implies a2 /. ((2 * a3) + b1) <= a2 /. ((2 * a3) + b2) ) ) & ( for b2 being Nat holds
( b2 in a1 & a2 /. ((2 * a3) + b1) = a2 /. ((2 * a3) + b2) implies b1 <= b2 ) ) ) ) ) & ( a1 = {} implies a4 = 0 ) );
existence
ex b1 being Nat st
( not ( c1 <> {} & ( for b2 being Nat holds
not ( b2 = b1 & b2 in c1 & ( for b3 being Nat holds
( b3 in c1 implies c2 /. ((2 * c3) + b2) <= c2 /. ((2 * c3) + b3) ) ) & ( for b3 being Nat holds
( b3 in c1 & c2 /. ((2 * c3) + b2) = c2 /. ((2 * c3) + b3) implies b2 <= b3 ) ) ) ) ) & ( c1 = {} implies b1 = 0 ) )
proof end;
uniqueness
for b1, b2 being Nat holds
( not ( c1 <> {} & ( for b3 being Nat holds
not ( b3 = b1 & b3 in c1 & ( for b4 being Nat holds
( b4 in c1 implies c2 /. ((2 * c3) + b3) <= c2 /. ((2 * c3) + b4) ) ) & ( for b4 being Nat holds
( b4 in c1 & c2 /. ((2 * c3) + b3) = c2 /. ((2 * c3) + b4) implies b3 <= b4 ) ) ) ) ) & ( c1 = {} implies b1 = 0 ) & not ( c1 <> {} & ( for b3 being Nat holds
not ( b3 = b2 & b3 in c1 & ( for b4 being Nat holds
( b4 in c1 implies c2 /. ((2 * c3) + b3) <= c2 /. ((2 * c3) + b4) ) ) & ( for b4 being Nat holds
( b4 in c1 & c2 /. ((2 * c3) + b3) = c2 /. ((2 * c3) + b4) implies b3 <= b4 ) ) ) ) ) & ( c1 = {} implies b2 = 0 ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def10 defines Argmin GRAPHSP:def 10 :
for b1 being finite Subset of NAT
for b2 being Element of REAL *
for b3, b4 being Nat holds
( b4 = Argmin b1,b2,b3 iff ( not ( b1 <> {} & ( for b5 being Nat holds
not ( b5 = b4 & b5 in b1 & ( for b6 being Nat holds
( b6 in b1 implies b2 /. ((2 * b3) + b5) <= b2 /. ((2 * b3) + b6) ) ) & ( for b6 being Nat holds
( b6 in b1 & b2 /. ((2 * b3) + b5) = b2 /. ((2 * b3) + b6) implies b5 <= b6 ) ) ) ) ) & ( b1 = {} implies b4 = 0 ) ) );

theorem Th30: :: GRAPHSP:30
for b1, b2 being Nat
for b3 being Element of REAL * holds
( OuterVx b3,b1 <> {} & b2 = Argmin (OuterVx b3,b1),b3,b1 implies ( b2 in dom b3 & 1 <= b2 & b2 <= b1 & b3 . b2 <> - 1 & b3 . (b1 + b2) <> - 1 ) )
proof end;

theorem Th31: :: GRAPHSP:31
for b1 being Nat
for b2 being Element of REAL * holds Argmin (OuterVx b2,b1),b2,b1 <= b1
proof end;

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) ) )
proof end;
uniqueness
for b1, b2 being Element of Funcs (REAL * ),(REAL * ) holds
( 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) ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def11 defines findmin GRAPHSP:def 11 :
for b1 being Nat
for b2 being Element of Funcs (REAL * ),(REAL * ) holds
( b2 = findmin b1 iff ( dom b2 = REAL * & ( for b3 being Element of REAL * holds b2 . b3 = b3,(((b1 * b1) + (3 * b1)) + 1) := (Argmin (OuterVx b3,b1),b3,b1),(- 1) ) ) );

theorem Th32: :: GRAPHSP:32
for b1, b2 being Nat
for b3 being Element of REAL * holds
( b1 in dom b3 & b1 > b2 & b1 <> ((b2 * b2) + (3 * b2)) + 1 implies ((findmin b2) . b3) . b1 = b3 . b1 )
proof end;

theorem Th33: :: GRAPHSP:33
for b1, b2 being Nat
for b3 being Element of REAL * holds
( b1 in dom b3 & b3 . b1 = - 1 & b1 <> ((b2 * b2) + (3 * b2)) + 1 implies ((findmin b2) . b3) . b1 = - 1 )
proof end;

theorem Th34: :: GRAPHSP:34
for b1 being Nat
for b2 being Element of REAL * holds dom ((findmin b1) . b2) = dom b2
proof end;

Lemma46: for b1, b2 being Nat holds
( b1 >= 1 implies b2 <= b1 * b2 )
proof end;

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 )
proof end;

Lemma48: for b1, b2 being Nat holds
( ( b1 < b2 & b2 <= 2 * b1 implies ( not ( 2 * b1 < b2 & b2 <= 3 * b1 ) & not b2 <= b1 & not b2 > 3 * b1 ) ) & ( not ( not b2 <= b1 & not b2 > 3 * b1 ) implies ( not ( 2 * b1 < b2 & b2 <= 3 * b1 ) & not ( b1 < b2 & b2 <= 2 * b1 ) ) ) & ( 2 * b1 < b2 & b2 <= 3 * b1 implies ( not ( b1 < b2 & b2 <= 2 * b1 ) & not b2 <= b1 & not b2 > 3 * b1 ) ) )
proof end;

theorem Th35: :: GRAPHSP:35
for b1 being Nat
for b2 being Element of REAL * holds
not ( OuterVx b2,b1 <> {} & ( for b3 being Nat holds
not ( b3 in OuterVx b2,b1 & 1 <= b3 & b3 <= b1 & ((findmin b1) . b2) . b3 = - 1 ) ) )
proof end;

definition
let c1 be Element of REAL * ;
let c2, c3 be Nat;
func newpathcost c1,c2,c3 -> Real equals :: GRAPHSP:def 12
(a1 /. ((2 * a2) + (a1 /. (((a2 * a2) + (3 * a2)) + 1)))) + (a1 /. (((2 * a2) + (a2 * (a1 /. (((a2 * a2) + (3 * a2)) + 1)))) + a3));
correctness
coherence
(c1 /. ((2 * c2) + (c1 /. (((c2 * c2) + (3 * c2)) + 1)))) + (c1 /. (((2 * c2) + (c2 * (c1 /. (((c2 * c2) + (3 * c2)) + 1)))) + c3)) is Real
;
;
end;

:: deftheorem Def12 defines newpathcost GRAPHSP:def 12 :
for b1 being Element of REAL *
for b2, b3 being Nat holds newpathcost b1,b2,b3 = (b1 /. ((2 * b2) + (b1 /. (((b2 * b2) + (3 * b2)) + 1)))) + (b1 /. (((2 * b2) + (b2 * (b1 /. (((b2 * b2) + (3 * b2)) + 1)))) + b3));

definition
let c1, c2 be Nat;
let c3 be Element of REAL * ;
pred c3 hasBetterPathAt c1,c2 means :Def13: :: GRAPHSP:def 13
( not ( not a3 . (a1 + a2) = - 1 & not a3 /. ((2 * a1) + a2) > newpathcost a3,a1,a2 ) & a3 /. (((2 * a1) + (a1 * (a3 /. (((a1 * a1) + (3 * a1)) + 1)))) + a2) >= 0 & a3 . a2 <> - 1 );
end;

:: deftheorem Def13 defines hasBetterPathAt GRAPHSP:def 13 :
for b1, b2 being Nat
for b3 being Element of REAL * holds
( b3 hasBetterPathAt b1,b2 iff ( not ( not b3 . (b1 + b2) = - 1 & not b3 /. ((2 * b1) + b2) > newpathcost b3,b1,b2 ) & b3 /. (((2 * b1) + (b1 * (b3 /. (((b1 * b1) + (3 * b1)) + 1)))) + b2) >= 0 & b3 . b2 <> - 1 ) );

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 holds
( b1 in dom a1 implies ( ( 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 ) ) ) & ( not ( not b1 <= a2 & not b1 > 3 * a2 ) implies a3 . b1 = a1 . b1 ) ) ) ) );
existence
ex b1 being Element of REAL * st
( dom b1 = dom c1 & ( for b2 being Nat holds
( b2 in dom c1 implies ( ( 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 ) ) ) & ( not ( not b2 <= c2 & not b2 > 3 * c2 ) implies b1 . b2 = c1 . b2 ) ) ) ) )
proof end;
uniqueness
for b1, b2 being Element of REAL * holds
( dom b1 = dom c1 & ( for b3 being Nat holds
( b3 in dom c1 implies ( ( 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 ) ) ) & ( not ( not b3 <= c2 & not b3 > 3 * c2 ) implies b1 . b3 = c1 . b3 ) ) ) ) & dom b2 = dom c1 & ( for b3 being Nat holds
( b3 in dom c1 implies ( ( 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 ) ) ) & ( not ( not b3 <= c2 & not b3 > 3 * c2 ) implies b2 . b3 = c1 . b3 ) ) ) ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def14 defines Relax GRAPHSP:def 14 :
for b1 being Element of REAL *
for b2 being Nat
for b3 being Element of REAL * holds
( b3 = Relax b1,b2 iff ( dom b3 = dom b1 & ( for b4 being Nat holds
( b4 in dom b1 implies ( ( b2 < b4 & b4 <= 2 * b2 implies ( ( b1 hasBetterPathAt b2,b4 -' b2 implies b3 . b4 = b1 /. (((b2 * b2) + (3 * b2)) + 1) ) & ( not b1 hasBetterPathAt b2,b4 -' b2 implies b3 . b4 = b1 . b4 ) ) ) & ( 2 * b2 < b4 & b4 <= 3 * b2 implies ( ( b1 hasBetterPathAt b2,b4 -' (2 * b2) implies b3 . b4 = newpathcost b1,b2,(b4 -' (2 * b2)) ) & ( not b1 hasBetterPathAt b2,b4 -' (2 * b2) implies b3 . b4 = b1 . b4 ) ) ) & ( not ( not b4 <= b2 & not b4 > 3 * b2 ) implies b3 . b4 = b1 . b4 ) ) ) ) ) );

definition
let c1 be Nat;
func Relax c1 -> Element of Funcs (REAL * ),(REAL * ) means :Def15: :: GRAPHSP:def 15
( dom a2 = REAL * & ( for b1 being Element of REAL * holds a2 . b1 = Relax b1,a1 ) );
existence
ex b1 being Element of Funcs (REAL * ),(REAL * ) st
( dom b1 = REAL * & ( for b2 being Element of REAL * holds b1 . b2 = Relax b2,c1 ) )
proof end;
uniqueness
for b1, b2 being Element of Funcs (REAL * ),(REAL * ) holds
( dom b1 = REAL * & ( for b3 being Element of REAL * holds b1 . b3 = Relax b3,c1 ) & dom b2 = REAL * & ( for b3 being Element of REAL * holds b2 . b3 = Relax b3,c1 ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def15 defines Relax GRAPHSP:def 15 :
for b1 being Nat
for b2 being Element of Funcs (REAL * ),(REAL * ) holds
( b2 = Relax b1 iff ( dom b2 = REAL * & ( for b3 being Element of REAL * holds b2 . b3 = Relax b3,b1 ) ) );

theorem Th36: :: GRAPHSP:36
for b1 being Nat
for b2 being Element of REAL * holds dom ((Relax b1) . b2) = dom b2
proof end;

theorem Th37: :: GRAPHSP:37
for b1, b2 being Nat
for b3 being Element of REAL * holds
( not ( not b1 <= b2 & not b1 > 3 * b2 ) & b1 in dom b3 implies ((Relax b2) . b3) . b1 = b3 . b1 )
proof end;

theorem Th38: :: GRAPHSP:38
for b1, b2 being Nat
for b3 being Element of REAL * holds dom (((repeat ((Relax b1) * (findmin b1))) . b2) . b3) = dom (((repeat ((Relax b1) * (findmin b1))) . (b2 + 1)) . b3)
proof end;

theorem Th39: :: GRAPHSP:39
for b1, b2 being Nat
for b3 being Element of REAL * holds
( OuterVx (((repeat ((Relax b1) * (findmin b1))) . b2) . b3),b1 <> {} implies UnusedVx (((repeat ((Relax b1) * (findmin b1))) . (b2 + 1)) . b3),b1 c< UnusedVx (((repeat ((Relax b1) * (findmin b1))) . b2) . b3),b1 )
proof end;

theorem Th40: :: GRAPHSP:40
for b1, b2, b3 being Nat
for b4, b5, b6 being Element of REAL * holds
( 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 <> {} implies ( UsedVx b6,b1 = (UsedVx b4,b1) \/ {b3} & not b3 in UsedVx b4,b1 ) )
proof end;

theorem Th41: :: GRAPHSP:41
for b1 being Nat
for b2 being Element of REAL * holds
ex b3 being Nat st
( b3 <= b1 & OuterVx (((repeat ((Relax b1) * (findmin b1))) . b3) . b2),b1 = {} )
proof end;

Lemma59: for b1, b2 being Nat holds b1 - b2 <= b1
proof end;

Lemma60: for b1, b2 being FinSequence of NAT
for b3 being Element of REAL *
for b4, b5 being Nat holds
( ( for b6 being Nat holds
( 1 <= b6 & b6 < len b1 implies b1 . ((len b1) - b6) = b3 . ((b1 /. (((len b1) - b6) + 1)) + b5) ) ) & ( for b6 being Nat holds
( 1 <= b6 & b6 < len b2 implies b2 . ((len b2) - b6) = b3 . ((b2 /. (((len b2) - b6) + 1)) + b5) ) ) & len b1 <= len b2 & b1 . (len b1) = b2 . (len b2) implies for b6 being Nat holds
( 1 <= b6 & b6 < len b1 implies b1 . ((len b1) - b6) = b2 . ((len b2) - b6) ) )
proof end;

theorem Th42: :: GRAPHSP:42
for b1, b2 being Nat
for b3 being Element of REAL * holds dom b3 = dom (((repeat ((Relax b1) * (findmin b1))) . b2) . b3)
proof end;

definition
let c1, c2 be Element of REAL * ;
let c3, c4 be Nat;
pred c1,c2 equal_at c3,c4 means :Def16: :: GRAPHSP:def 16
( dom a1 = dom a2 & ( for b1 being Nat holds
( b1 in dom a1 & a3 <= b1 & b1 <= a4 implies a1 . b1 = a2 . b1 ) ) );
end;

:: deftheorem Def16 defines equal_at GRAPHSP:def 16 :
for b1, b2 being Element of REAL *
for b3, b4 being Nat holds
( b1,b2 equal_at b3,b4 iff ( dom b1 = dom b2 & ( for b5 being Nat holds
( b5 in dom b1 & b3 <= b5 & b5 <= b4 implies b1 . b5 = b2 . b5 ) ) ) );

theorem Th43: :: GRAPHSP:43
for b1, b2 being Nat
for b3 being Element of REAL * holds b3,b3 equal_at b1,b2
proof end;

theorem Th44: :: GRAPHSP:44
for b1, b2 being Nat
for b3, b4, b5 being Element of REAL * holds
( b3,b4 equal_at b1,b2 & b4,b5 equal_at b1,b2 implies b3,b5 equal_at b1,b2 )
proof end;

theorem Th45: :: GRAPHSP:45
for b1, b2 being Nat
for b3 being Element of REAL * holds ((repeat ((Relax b1) * (findmin b1))) . b2) . b3,((repeat ((Relax b1) * (findmin b1))) . (b2 + 1)) . b3 equal_at (3 * b1) + 1,(b1 * b1) + (3 * b1)
proof end;

theorem Th46: :: GRAPHSP:46
for b1 being Element of Funcs (REAL * ),(REAL * )
for b2 being Element of REAL *
for b3, b4 being Nat holds
not ( b4 < LifeSpan b1,b2,b3 & not OuterVx (((repeat b1) . b4) . b2),b3 <> {} ) by Def4;

theorem Th47: :: GRAPHSP:47
for b1, b2 being Nat
for b3 being Element of REAL * holds b3,((repeat ((Relax b1) * (findmin b1))) . b2) . b3 equal_at (3 * b1) + 1,(b1 * b1) + (3 * b1)
proof end;

theorem Th48: :: GRAPHSP:48
for b1 being Nat
for b2 being Element of REAL * holds
( 1 <= b1 & 1 in dom b2 & b2 . (b1 + 1) <> - 1 & ( for b3 being Nat holds
( 1 <= b3 & b3 <= b1 implies b2 . b3 = 1 ) ) & ( for b3 being Nat holds
( 2 <= b3 & b3 <= b1 implies b2 . (b1 + b3) = - 1 ) ) implies ( 1 = Argmin (OuterVx b2,b1),b2,b1 & UsedVx b2,b1 = {} & {1} = UsedVx (((repeat ((Relax b1) * (findmin b1))) . 1) . b2),b1 ) )
proof end;

theorem Th49: :: GRAPHSP:49
for b1, b2, b3 being Nat
for b4, b5, b6 being Element of REAL * holds
( b4 = ((repeat ((Relax b1) * (findmin b1))) . 1) . b5 & b6 = ((repeat ((Relax b1) * (findmin b1))) . b2) . b5 & 1 <= b2 & b2 <= LifeSpan ((Relax b1) * (findmin b1)),b5,b1 & b3 in UsedVx b4,b1 implies b3 in UsedVx b6,b1 )
proof end;

definition
let c1 be FinSequence of NAT ;
let c2 be Element of REAL * ;
let c3, c4 be Nat;
pred c1 is_vertex_seq_at c2,c3,c4 means :Def17: :: GRAPHSP:def 17
( a1 . (len a1) = a3 & ( for b1 being Nat holds
( 1 <= b1 & b1 < len a1 implies a1 . ((len a1) - b1) = a2 . (a4 + (a1 /. (((len a1) - b1) + 1))) ) ) );
end;

:: deftheorem Def17 defines is_vertex_seq_at GRAPHSP:def 17 :
for b1 being FinSequence of NAT
for b2 being Element of REAL *
for b3, b4 being Nat holds
( b1 is_vertex_seq_at b2,b3,b4 iff ( b1 . (len b1) = b3 & ( for b5 being Nat holds
( 1 <= b5 & b5 < len b1 implies b1 . ((len b1) - b5) = b2 . (b4 + (b1 /. (((len b1) - b5) + 1))) ) ) ) );

definition
let c1 be FinSequence of NAT ;
let c2 be Element of REAL * ;
let c3, c4 be Nat;
pred c1 is_simple_vertex_seq_at c2,c3,c4 means :Def18: :: GRAPHSP:def 18
( a1 . 1 = 1 & len a1 > 1 & a1 is_vertex_seq_at a2,a3,a4 & a1 is one-to-one );
end;

:: deftheorem Def18 defines is_simple_vertex_seq_at GRAPHSP:def 18 :
for b1 being FinSequence of NAT
for b2 being Element of REAL *
for b3, b4 being Nat holds
( b1 is_simple_vertex_seq_at b2,b3,b4 iff ( b1 . 1 = 1 & len b1 > 1 & b1 is_vertex_seq_at b2,b3,b4 & b1 is one-to-one ) );

theorem Th50: :: GRAPHSP:50
for b1, b2 being FinSequence of NAT
for b3 being Element of REAL *
for b4, b5 being Nat holds
( b1 is_simple_vertex_seq_at b3,b4,b5 & b2 is_simple_vertex_seq_at b3,b4,b5 implies b1 = b2 )
proof end;

definition
let c1 be Graph;
let c2 be FinSequence of the Edges of c1;
let c3 be FinSequence;
pred c2 is_oriented_edge_seq_of c3 means :Def19: :: GRAPHSP:def 19
( len a3 = (len a2) + 1 & ( for b1 being Nat holds
( 1 <= b1 & b1 <= len a2 implies ( the Source of a1 . (a2 . b1) = a3 . b1 & the Target of a1 . (a2 . b1) = a3 . (b1 + 1) ) ) ) );
end;

:: deftheorem Def19 defines is_oriented_edge_seq_of GRAPHSP:def 19 :
for b1 being Graph
for b2 being FinSequence of the Edges of b1
for b3 being FinSequence holds
( b2 is_oriented_edge_seq_of b3 iff ( len b3 = (len b2) + 1 & ( for b4 being Nat holds
( 1 <= b4 & b4 <= len b2 implies ( the Source of b1 . (b2 . b4) = b3 . b4 & the Target of b1 . (b2 . b4) = b3 . (b4 + 1) ) ) ) ) );

theorem Th51: :: GRAPHSP:51
for b1 being oriented Graph
for b2 being FinSequence
for b3, b4 being oriented Chain of b1 holds
( b3 is_oriented_edge_seq_of b2 & b4 is_oriented_edge_seq_of b2 implies b3 = b4 )
proof end;

theorem Th52: :: GRAPHSP:52
for b1 being Graph
for b2, b3 being FinSequence
for b4 being oriented Chain of b1 holds
( b4 is_oriented_edge_seq_of b2 & b4 is_oriented_edge_seq_of b3 & len b4 >= 1 implies b2 = b3 )
proof end;

Lemma72: for b1, b2 being Nat holds
( 1 <= b1 & b1 <= b2 implies ( 1 < (2 * b2) + b1 & (2 * b2) + b1 < ((b2 * b2) + (3 * b2)) + 1 & b1 < (2 * b2) + b1 ) )
proof end;

Lemma73: for b1, b2 being Nat holds
( 1 <= b1 & b1 <= b2 implies ( 1 < b2 + b1 & b2 + b1 <= 2 * b2 & b2 + b1 < ((b2 * b2) + (3 * b2)) + 1 ) )
proof end;

Lemma74: for b1, b2, b3 being Nat holds
( 1 <= b1 & b1 <= b2 & b3 <= b2 implies ( 1 < ((2 * b2) + (b2 * b1)) + b3 & b1 < ((2 * b2) + (b2 * b1)) + b3 & ((2 * b2) + (b2 * b1)) + b3 < ((b2 * b2) + (3 * b2)) + 1 ) )
proof end;

Lemma75: for b1, b2, b3 being Nat holds
( 1 <= b1 & b1 <= b2 & 1 <= b3 & b3 <= b2 implies ( (3 * b2) + 1 <= ((2 * b2) + (b2 * b1)) + b3 & ((2 * b2) + (b2 * b1)) + b3 <= (b2 * b2) + (3 * b2) ) )
proof end;

definition
let c1 be Element of REAL * ;
let c2 be oriented Graph;
let c3 be Nat;
let c4 be Function of the Edges of c2, Real>=0 ;
pred c1 is_Input_of_Dijkstra_Alg c2,c3,c4 means :Def20: :: GRAPHSP:def 20
( len a1 = ((a3 * a3) + (3 * a3)) + 1 & Seg a3 = the Vertices of a2 & ( for b1 being Nat holds
( 1 <= b1 & b1 <= a3 implies ( a1 . b1 = 1 & a1 . ((2 * a3) + b1) = 0 ) ) ) & a1 . (a3 + 1) = 0 & ( for b1 being Nat holds
( 2 <= b1 & b1 <= a3 implies a1 . (a3 + b1) = - 1 ) ) & ( for b1, b2 being Vertex of a2
for b3, b4 being Nat holds
( b3 = b1 & b4 = b2 implies a1 . (((2 * a3) + (a3 * b3)) + b4) = Weight b1,b2,a4 ) ) );
end;

:: deftheorem Def20 defines is_Input_of_Dijkstra_Alg GRAPHSP:def 20 :
for b1 being Element of REAL *
for b2 being oriented Graph
for b3 being Nat
for b4 being Function of the Edges of b2, Real>=0 holds
( b1 is_Input_of_Dijkstra_Alg b2,b3,b4 iff ( len b1 = ((b3 * b3) + (3 * b3)) + 1 & Seg b3 = the Vertices of b2 & ( for b5 being Nat holds
( 1 <= b5 & b5 <= b3 implies ( b1 . b5 = 1 & b1 . ((2 * b3) + b5) = 0 ) ) ) & b1 . (b3 + 1) = 0 & ( for b5 being Nat holds
( 2 <= b5 & b5 <= b3 implies b1 . (b3 + b5) = - 1 ) ) & ( for b5, b6 being Vertex of b2
for b7, b8 being Nat holds
( b7 = b5 & b8 = b6 implies b1 . (((2 * b3) + (b3 * b7)) + b8) = Weight b5,b6,b4 ) ) ) );

definition
let c1 be Nat;
func DijkstraAlgorithm c1 -> Element of Funcs (REAL * ),(REAL * ) equals :: GRAPHSP:def 21
while_do ((Relax a1) * (findmin a1)),a1;
coherence
while_do ((Relax c1) * (findmin c1)),c1 is Element of Funcs (REAL * ),(REAL * )
;
end;

:: deftheorem Def21 defines DijkstraAlgorithm GRAPHSP:def 21 :
for b1 being Nat holds DijkstraAlgorithm b1 = while_do ((Relax b1) * (findmin b1)),b1;

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 holds
( b2 is_Input_of_Dijkstra_Alg b4,b1,b5 & b6 = 1 & b1 >= 1 & b3 = ((repeat ((Relax b1) * (findmin b1))) . 1) . b2 implies ( ( for b7 being Vertex of b4
for b8 being Nat holds
not ( b7 <> b6 & b7 = b8 & b3 . (b1 + b8) <> - 1 & ( for b9 being FinSequence of NAT
for b10 being oriented Chain of b4 holds
not ( b9 is_simple_vertex_seq_at b3,b8,b1 & ( for b11 being Nat holds
( 1 <= b11 & b11 < len b9 implies 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 holds
( b3 . (b1 + b8) = - 1 & 1 <= b8 & b8 <= b1 & b7 in UsedVx b3,b1 implies b2 . (((2 * b1) + (b1 * b7)) + b8) = - 1 ) ) & ( for b7 being Nat holds
not ( b7 in UsedVx b3,b1 & not b3 . (b1 + b7) <> - 1 ) ) ) )
proof end;

Lemma78: for b1, b2, b3 being Nat
for b4, b5, b6 being Element of REAL * holds
( 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 implies b6 . (b1 + b3) = b4 . (b1 + b3) )
proof end;

Lemma79: for b1, b2, b3 being Nat
for b4, b5, b6 being Element of REAL *
for b7 being FinSequence of NAT holds
( 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 holds
( 1 <= b8 & b8 < len b7 implies b7 . b8 in UsedVx b4,b1 ) ) implies b7 is_simple_vertex_seq_at b6,b3,b1 )
proof end;

Lemma80: for b1, b2, b3, b4 being Nat
for b5, b6, b7 being Element of REAL *
for b8 being FinSequence of NAT holds
( 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 holds
( 1 <= b9 & b9 < len b8 implies b8 . b9 in UsedVx b5,b1 ) ) implies b8 ^ <*b4*> is_simple_vertex_seq_at b7,b4,b1 )
proof end;

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 holds
not ( 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 holds
( b5 . (b1 + b12) = - 1 & 1 <= b12 & b12 <= b1 & b11 in b3 implies b4 . (((2 * b1) + (b1 * b11)) + b12) = - 1 ) ) & not b5 . (b1 + b2) <> - 1 )
proof end;

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 holds
( 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 holds
not ( b9 <> b8 & b9 = b10 & b4 . (b1 + b10) <> - 1 & ( for b11 being FinSequence of NAT
for b12 being oriented Chain of b6 holds
not ( b11 is_simple_vertex_seq_at b4,b10,b1 & ( for b13 being Nat holds
( 1 <= b13 & b13 < len b11 implies 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 holds
( b4 . (b1 + b10) = - 1 & 1 <= b10 & b10 <= b1 & b9 in UsedVx b4,b1 implies b3 . (((2 * b1) + (b1 * b9)) + b10) = - 1 ) ) & ( for b9 being Nat holds
not ( b9 in UsedVx b4,b1 & not b4 . (b1 + b9) <> - 1 ) ) implies ( ( for b9 being Vertex of b6
for b10 being Nat holds
not ( b9 <> b8 & b9 = b10 & b5 . (b1 + b10) <> - 1 & ( for b11 being FinSequence of NAT
for b12 being oriented Chain of b6 holds
not ( b11 is_simple_vertex_seq_at b5,b10,b1 & ( for b13 being Nat holds
( 1 <= b13 & b13 < len b11 implies 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 holds
( b5 . (b1 + b10) = - 1 & 1 <= b10 & b10 <= b1 & b9 in UsedVx b5,b1 implies b3 . (((2 * b1) + (b1 * b9)) + b10) = - 1 ) ) & ( for b9 being Nat holds
not ( b9 in UsedVx b5,b1 & not b5 . (b1 + b9) <> - 1 ) ) ) )
proof end;

theorem Th53: :: GRAPHSP:53
for b1, b2 being Nat
for b3, b4 being Element of REAL *
for b5 being oriented finite Graph
for b6 being Function of the Edges of b5, Real>=0
for b7, b8 being Vertex of b5 holds
( b3 is_Input_of_Dijkstra_Alg b5,b1,b6 & b7 = 1 & 1 <> b8 & b8 = b2 & b1 >= 1 & b4 = (DijkstraAlgorithm b1) . b3 implies ( the Vertices of b5 = (UsedVx b4,b1) \/ (UnusedVx b4,b1) & not ( b8 in UsedVx b4,b1 & ( for b9 being FinSequence of NAT
for b10 being oriented Chain of b5 holds
not ( 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 ) ) )
proof end;