begin
Lm1:
for f being Function st f is one-to-one holds
card (dom f) = card (rng f)
Lm2:
for F being Function
for x, y being set holds dom (F +* (x .--> y)) = (dom F) \/ {x}
Lm3:
for f being one-to-one Function
for X, Y being set st X c= Y holds
for x being set st x in dom ((f | X) ") holds
((f | X) ") . x = ((f | Y) ") . x
theorem
Lm4:
for a, b, c being real number st a < b holds
(c - b) + 1 < (c - a) + 1
theorem Th2:
for
n,
m,
k being
Nat st
m <= k &
n < m holds
k -' m < k -' n
:: deftheorem Def1 defines with_finite-elements LEXBFS:def 1 :
for S being set holds
( S is with_finite-elements iff for x being Element of S holds x is finite );
:: deftheorem Def2 defines .\/ LEXBFS:def 2 :
for f, g, b3 being Function holds
( b3 = f .\/ g iff ( dom b3 = (dom f) \/ (dom g) & ( for x being set st x in (dom f) \/ (dom g) holds
b3 . x = (f . x) \/ (g . x) ) ) );
theorem Th3:
theorem Th4:
theorem Th5:
:: deftheorem Def3 defines natsubset-yielding LEXBFS:def 3 :
for f being Relation holds
( f is natsubset-yielding iff rng f c= bool NAT );
Lm5:
for F being Function st ( for x being set st x in rng F holds
x is finite ) holds
F is finite-yielding
theorem Th6:
:: deftheorem Def4 defines .incSubset LEXBFS:def 4 :
for F being natural-valued Function
for S being set
for k being Nat
for b4 being natural-valued Function holds
( b4 = F .incSubset (S,k) iff ( dom b4 = dom F & ( for y being set holds
( ( y in S & y in dom F implies b4 . y = (F . y) + k ) & ( not y in S implies b4 . y = F . y ) ) ) ) );
:: deftheorem Def5 defines max LEXBFS:def 5 :
for n being Ordinal
for T being connected TermOrder of n
for B being non empty finite Subset of (Bags n)
for b4 being bag of n holds
( b4 = max (B,T) iff ( b4 in B & ( for x being bag of n st x in B holds
x <= b4,T ) ) );
Lm6:
for G being _Graph
for W being Walk of G
for e, v being set st e Joins W .last() ,v,G holds
(W .addEdge e) .length() = (W .length()) + 1
Lm7:
for G being _Graph
for W being Walk of G holds W .length() = (W .reverse()) .length()
Lm8:
for G being _Graph
for W being Walk of G
for e, x being set st e Joins W .last() ,x,G holds
for n being Nat st n in dom W holds
( (W .addEdge e) . n = W . n & n in dom (W .addEdge e) )
begin
:: deftheorem Def6 defines iterative LEXBFS:def 6 :
for s being ManySortedSet of NAT holds
( s is iterative iff for k, n being Nat st s . k = s . n holds
s . (k + 1) = s . (n + 1) );
:: deftheorem Def7 defines eventually-constant LEXBFS:def 7 :
for S being ManySortedSet of NAT holds
( S is eventually-constant iff ex n being Nat st
for m being Nat st n <= m holds
S . n = S . m );
theorem Th7:
theorem Th8:
theorem Th9:
theorem Th10:
begin
:: deftheorem Def8 defines preVNumberingSeq LEXBFS:def 8 :
for G being finite _Graph
for b2 being ManySortedSet of NAT holds
( b2 is preVNumberingSeq of G iff for i being Nat holds b2 . i is PartFunc of (the_Vertices_of G),NAT );
:: deftheorem Def9 defines vertex-numbering LEXBFS:def 9 :
for G being finite _Graph
for S being preVNumberingSeq of G holds
( S is vertex-numbering iff ( S . 0 = {} & S is iterative & S is halting & S .Lifespan() = G .order() & ( for n being Nat st n < S .Lifespan() holds
ex w being Vertex of G st
( not w in dom (S . n) & S . (n + 1) = (S . n) +* (w .--> ((S .Lifespan()) -' n)) ) ) ) );
:: deftheorem Def10 defines .PickedAt LEXBFS:def 10 :
for G being finite _Graph
for S being VNumberingSeq of G
for n being Nat
for b4 being set holds
( ( n >= S .Lifespan() implies ( b4 = S .PickedAt n iff b4 = choose (the_Vertices_of G) ) ) & ( not n >= S .Lifespan() implies ( b4 = S .PickedAt n iff ( not b4 in dom (S . n) & S . (n + 1) = (S . n) +* (b4 .--> ((S .Lifespan()) -' n)) ) ) ) );
theorem Th11:
theorem Th12:
theorem
theorem Th14:
theorem Th15:
theorem Th16:
theorem Th17:
theorem Th18:
theorem Th19:
theorem Th20:
theorem Th21:
theorem Th22:
theorem Th23:
theorem Th24:
begin
:: deftheorem defines LexBFS:Init LEXBFS:def 11 :
for G being _Graph holds LexBFS:Init G = [{},((the_Vertices_of G) --> {})];
definition
let G be
finite _Graph;
let L be
LexBFS:Labeling of
G;
func LexBFS:PickUnnumbered L -> Vertex of
G means :
Def12:
it = choose (the_Vertices_of G) if dom (L `1) = the_Vertices_of G otherwise ex
S being non
empty finite Subset of
(bool NAT) ex
B being non
empty finite Subset of
(Bags NAT) ex
F being
Function st
(
S = rng F &
F = (L `2) | ((the_Vertices_of G) \ (dom (L `1))) & ( for
x being
finite Subset of
NAT st
x in S holds
(
x,1)
-bag in B ) & ( for
x being
set st
x in B holds
ex
y being
finite Subset of
NAT st
(
y in S &
x = (
y,1)
-bag ) ) &
it = choose (F " {(support (max (B,(InvLexOrder NAT))))}) );
existence
( ( dom (L `1) = the_Vertices_of G implies ex b1 being Vertex of G st b1 = choose (the_Vertices_of G) ) & ( not dom (L `1) = the_Vertices_of G implies ex b1 being Vertex of G ex S being non empty finite Subset of (bool NAT) ex B being non empty finite Subset of (Bags NAT) ex F being Function st
( S = rng F & F = (L `2) | ((the_Vertices_of G) \ (dom (L `1))) & ( for x being finite Subset of NAT st x in S holds
(x,1) -bag in B ) & ( for x being set st x in B holds
ex y being finite Subset of NAT st
( y in S & x = (y,1) -bag ) ) & b1 = choose (F " {(support (max (B,(InvLexOrder NAT))))}) ) ) )
uniqueness
for b1, b2 being Vertex of G holds
( ( dom (L `1) = the_Vertices_of G & b1 = choose (the_Vertices_of G) & b2 = choose (the_Vertices_of G) implies b1 = b2 ) & ( not dom (L `1) = the_Vertices_of G & ex S being non empty finite Subset of (bool NAT) ex B being non empty finite Subset of (Bags NAT) ex F being Function st
( S = rng F & F = (L `2) | ((the_Vertices_of G) \ (dom (L `1))) & ( for x being finite Subset of NAT st x in S holds
(x,1) -bag in B ) & ( for x being set st x in B holds
ex y being finite Subset of NAT st
( y in S & x = (y,1) -bag ) ) & b1 = choose (F " {(support (max (B,(InvLexOrder NAT))))}) ) & ex S being non empty finite Subset of (bool NAT) ex B being non empty finite Subset of (Bags NAT) ex F being Function st
( S = rng F & F = (L `2) | ((the_Vertices_of G) \ (dom (L `1))) & ( for x being finite Subset of NAT st x in S holds
(x,1) -bag in B ) & ( for x being set st x in B holds
ex y being finite Subset of NAT st
( y in S & x = (y,1) -bag ) ) & b2 = choose (F " {(support (max (B,(InvLexOrder NAT))))}) ) implies b1 = b2 ) )
consistency
for b1 being Vertex of G holds verum
;
end;
:: deftheorem Def12 defines LexBFS:PickUnnumbered LEXBFS:def 12 :
for G being finite _Graph
for L being LexBFS:Labeling of G
for b3 being Vertex of G holds
( ( dom (L `1) = the_Vertices_of G implies ( b3 = LexBFS:PickUnnumbered L iff b3 = choose (the_Vertices_of G) ) ) & ( not dom (L `1) = the_Vertices_of G implies ( b3 = LexBFS:PickUnnumbered L iff ex S being non empty finite Subset of (bool NAT) ex B being non empty finite Subset of (Bags NAT) ex F being Function st
( S = rng F & F = (L `2) | ((the_Vertices_of G) \ (dom (L `1))) & ( for x being finite Subset of NAT st x in S holds
(x,1) -bag in B ) & ( for x being set st x in B holds
ex y being finite Subset of NAT st
( y in S & x = (y,1) -bag ) ) & b3 = choose (F " {(support (max (B,(InvLexOrder NAT))))}) ) ) ) );
:: deftheorem defines LexBFS:Update LEXBFS:def 13 :
for G being finite _Graph
for L being LexBFS:Labeling of G
for v being Vertex of G
for n being Nat holds LexBFS:Update (L,v,n) = [((L `1) +* (v .--> ((G .order()) -' n))),((L `2) .\/ (((G .AdjacentSet {v}) \ (dom (L `1))) --> {((G .order()) -' n)}))];
theorem Th25:
theorem Th26:
theorem Th27:
:: deftheorem Def14 defines LexBFS:Step LEXBFS:def 14 :
for G being finite _Graph
for L being LexBFS:Labeling of G holds
( ( G .order() <= card (dom (L `1)) implies LexBFS:Step L = L ) & ( not G .order() <= card (dom (L `1)) implies LexBFS:Step L = LexBFS:Update (L,(LexBFS:PickUnnumbered L),(card (dom (L `1)))) ) );
:: deftheorem Def15 defines LexBFS:LabelingSeq LEXBFS:def 15 :
for G being _Graph
for b2 being ManySortedSet of NAT holds
( b2 is LexBFS:LabelingSeq of G iff for n being Nat holds b2 . n is LexBFS:Labeling of G );
:: deftheorem Def16 defines ``1 LEXBFS:def 16 :
for G being finite _Graph
for S being LexBFS:LabelingSeq of G
for b3 being preVNumberingSeq of G holds
( b3 = S ``1 iff for n being Nat holds b3 . n = (S . n) `1 );
:: deftheorem Def17 defines LexBFS:CSeq LEXBFS:def 17 :
for G being finite _Graph
for b2 being LexBFS:LabelingSeq of G holds
( b2 = LexBFS:CSeq G iff ( b2 . 0 = LexBFS:Init G & ( for n being Nat holds b2 . (n + 1) = LexBFS:Step (b2 . n) ) ) );
theorem Th28:
Lm9:
for G being _Graph
for v being set holds
( dom ((LexBFS:Init G) `2) = the_Vertices_of G & ((LexBFS:Init G) `2) . v = {} )
theorem Th29:
theorem Th30:
theorem Th31:
theorem Th32:
theorem Th33:
theorem Th34:
theorem Th35:
theorem Th36:
theorem Th37:
theorem Th38:
theorem Th39:
theorem Th40:
theorem Th41:
theorem Th42:
theorem Th43:
theorem Th44:
theorem Th45:
theorem Th46:
theorem Th47:
theorem Th48:
theorem Th49:
theorem Th50:
theorem Th51:
theorem Th52:
:: deftheorem Def18 defines with_property_L3 LEXBFS:def 18 :
for G being _Graph
for F being PartFunc of (the_Vertices_of G),NAT holds
( F is with_property_L3 iff for a, b, c being Vertex of G st a in dom F & b in dom F & c in dom F & F . a < F . b & F . b < F . c & a,c are_adjacent & not b,c are_adjacent holds
ex d being Vertex of G st
( d in dom F & F . c < F . d & b,d are_adjacent & not a,d are_adjacent & ( for e being Vertex of G st e <> d & e,b are_adjacent & not e,a are_adjacent holds
F . e < F . d ) ) );
theorem Th53:
theorem Th54:
theorem
begin
:: deftheorem defines MCS:Init LEXBFS:def 19 :
for G being finite _Graph holds MCS:Init G = [{},((the_Vertices_of G) --> 0)];
:: deftheorem Def20 defines MCS:PickUnnumbered LEXBFS:def 20 :
for G being finite _Graph
for L being MCS:Labeling of G
for b3 being Vertex of G holds
( ( dom (L `1) = the_Vertices_of G implies ( b3 = MCS:PickUnnumbered L iff b3 = choose (the_Vertices_of G) ) ) & ( not dom (L `1) = the_Vertices_of G implies ( b3 = MCS:PickUnnumbered L iff ex S being non empty natural-membered finite set ex F being Function st
( S = rng F & F = (L `2) | ((the_Vertices_of G) \ (dom (L `1))) & b3 = choose (F " {(max S)}) ) ) ) );
:: deftheorem defines MCS:LabelAdjacent LEXBFS:def 21 :
for G being finite _Graph
for L being MCS:Labeling of G
for v being set holds MCS:LabelAdjacent (L,v) = [(L `1),((L `2) .incSubset (((G .AdjacentSet {v}) \ (dom (L `1))),1))];
definition
let G be
finite _Graph;
let L be
MCS:Labeling of
G;
let v be
Vertex of
G;
let n be
Nat;
func MCS:Update (
L,
v,
n)
-> MCS:Labeling of
G means :
Def22:
ex
M being
MCS:Labeling of
G st
(
M = [((L `1) +* (v .--> ((G .order()) -' n))),(L `2)] &
it = MCS:LabelAdjacent (
M,
v) );
existence
ex b1, M being MCS:Labeling of G st
( M = [((L `1) +* (v .--> ((G .order()) -' n))),(L `2)] & b1 = MCS:LabelAdjacent (M,v) )
uniqueness
for b1, b2 being MCS:Labeling of G st ex M being MCS:Labeling of G st
( M = [((L `1) +* (v .--> ((G .order()) -' n))),(L `2)] & b1 = MCS:LabelAdjacent (M,v) ) & ex M being MCS:Labeling of G st
( M = [((L `1) +* (v .--> ((G .order()) -' n))),(L `2)] & b2 = MCS:LabelAdjacent (M,v) ) holds
b1 = b2
;
end;
:: deftheorem Def22 defines MCS:Update LEXBFS:def 22 :
for G being finite _Graph
for L being MCS:Labeling of G
for v being Vertex of G
for n being Nat
for b5 being MCS:Labeling of G holds
( b5 = MCS:Update (L,v,n) iff ex M being MCS:Labeling of G st
( M = [((L `1) +* (v .--> ((G .order()) -' n))),(L `2)] & b5 = MCS:LabelAdjacent (M,v) ) );
:: deftheorem Def23 defines MCS:Step LEXBFS:def 23 :
for G being finite _Graph
for L being MCS:Labeling of G holds
( ( G .order() <= card (dom (L `1)) implies MCS:Step L = L ) & ( not G .order() <= card (dom (L `1)) implies MCS:Step L = MCS:Update (L,(MCS:PickUnnumbered L),(card (dom (L `1)))) ) );
:: deftheorem Def24 defines MCS:LabelingSeq LEXBFS:def 24 :
for G being _Graph
for b2 being ManySortedSet of NAT holds
( b2 is MCS:LabelingSeq of G iff for n being Nat holds b2 . n is MCS:Labeling of G );
:: deftheorem Def25 defines ``1 LEXBFS:def 25 :
for G being finite _Graph
for S being MCS:LabelingSeq of G
for b3 being preVNumberingSeq of G holds
( b3 = S ``1 iff for n being Nat holds b3 . n = (S . n) `1 );
:: deftheorem Def26 defines MCS:CSeq LEXBFS:def 26 :
for G being finite _Graph
for b2 being MCS:LabelingSeq of G holds
( b2 = MCS:CSeq G iff ( b2 . 0 = MCS:Init G & ( for n being Nat holds b2 . (n + 1) = MCS:Step (b2 . n) ) ) );
theorem Th56:
theorem Th57:
theorem Th58:
theorem Th59:
theorem Th60:
theorem Th61:
theorem Th62:
theorem
theorem Th64:
theorem Th65:
theorem Th66:
theorem Th67:
theorem Th68:
theorem Th69:
theorem Th70:
theorem Th71:
theorem Th72:
theorem Th73:
theorem Th74:
theorem Th75:
theorem Th76:
:: deftheorem Def27 defines with_property_T LEXBFS:def 27 :
for G being _Graph
for F being PartFunc of (the_Vertices_of G),NAT holds
( F is with_property_T iff for a, b, c being Vertex of G st a in dom F & b in dom F & c in dom F & F . a < F . b & F . b < F . c & a,c are_adjacent & not b,c are_adjacent holds
ex d being Vertex of G st
( d in dom F & F . b < F . d & b,d are_adjacent & not a,d are_adjacent ) );
theorem
theorem
theorem