begin
Lm1:
for e being set
for n being Element of NAT st e in Seg n holds
ex i being Element of NAT st
( e = i & 1 <= i & i <= n )
:: deftheorem defines nat_interval SGRAPH1:def 1 :
for m, n being natural number holds nat_interval (m,n) = { i where i is Element of NAT : ( m <= i & i <= n ) } ;
theorem
canceled;
theorem
theorem
theorem
theorem Th5:
theorem Th6:
theorem
Lm2:
for A being set
for s being Subset of A
for n being set st n in A holds
s \/ {n} is Subset of A
:: deftheorem SGRAPH1:def 2 :
canceled;
:: deftheorem SGRAPH1:def 3 :
canceled;
:: deftheorem defines TWOELEMENTSETS SGRAPH1:def 4 :
for A being set holds TWOELEMENTSETS A = { z where z is finite Subset of A : card z = 2 } ;
theorem
canceled;
theorem Th9:
theorem Th10:
theorem Th11:
theorem Th12:
theorem Th13:
theorem
theorem Th15:
theorem
theorem
begin
:: deftheorem SGRAPH1:def 5 :
canceled;
:: deftheorem defines SIMPLEGRAPHS SGRAPH1:def 6 :
for X being set holds SIMPLEGRAPHS X = { SimpleGraphStruct(# v,e #) where v is finite Subset of X, e is finite Subset of (TWOELEMENTSETS v) : verum } ;
theorem
canceled;
theorem Th19:
:: deftheorem Def7 defines SimpleGraph SGRAPH1:def 7 :
for X being set
for b2 being strict SimpleGraphStruct holds
( b2 is SimpleGraph of X iff b2 is Element of SIMPLEGRAPHS X );
theorem
canceled;
theorem Th21:
begin
theorem
canceled;
theorem Th23:
theorem
canceled;
theorem
theorem
theorem Th27:
:: deftheorem defines is_isomorphic_to SGRAPH1:def 8 :
for X being set
for G, G9 being SimpleGraph of X holds
( G is_isomorphic_to G9 iff ex Fv being Function of the carrier of G, the carrier of G9 st
( Fv is bijective & ( for v1, v2 being Element of G holds
( {v1,v2} in the SEdges of G iff {(Fv . v1),(Fv . v2)} in the SEdges of G ) ) ) );
begin
scheme
IndSimpleGraphs0{
F1()
-> set ,
P1[
set ] } :
provided
theorem
theorem
canceled;
theorem Th30:
theorem Th31:
:: deftheorem Def9 defines is_SetOfSimpleGraphs_of SGRAPH1:def 9 :
for X, GG being set holds
( GG is_SetOfSimpleGraphs_of X iff ( SimpleGraphStruct(# {},({} (TWOELEMENTSETS {})) #) in GG & ( for V being Subset of X
for E being Subset of (TWOELEMENTSETS V)
for n being set
for Evn being finite Subset of (TWOELEMENTSETS (V \/ {n})) st SimpleGraphStruct(# V,E #) in GG & n in X & not n in V holds
SimpleGraphStruct(# (V \/ {n}),Evn #) in GG ) & ( for V being Subset of X
for E being Subset of (TWOELEMENTSETS V)
for v1, v2 being set st SimpleGraphStruct(# V,E #) in GG & v1 in V & v2 in V & v1 <> v2 & not {v1,v2} in E holds
ex v1v2 being finite Subset of (TWOELEMENTSETS V) st
( v1v2 = E \/ {{v1,v2}} & SimpleGraphStruct(# V,v1v2 #) in GG ) ) ) );
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem Th35:
theorem Th36:
theorem
begin
:: deftheorem defines SubGraph SGRAPH1:def 10 :
for X being set
for G, b3 being SimpleGraph of X holds
( b3 is SubGraph of G iff ( the carrier of b3 c= the carrier of G & the SEdges of b3 c= the SEdges of G ) );
begin
:: deftheorem Def11 defines degree SGRAPH1:def 11 :
for X being set
for G being SimpleGraph of X
for v being set
for b4 being Element of NAT holds
( b4 = degree (G,v) iff ex X being finite set st
( ( for z being set holds
( z in X iff ( z in the SEdges of G & v in z ) ) ) & b4 = card X ) );
theorem
canceled;
theorem Th39:
theorem
theorem
theorem
begin
:: deftheorem Def12 defines is_path_of SGRAPH1:def 12 :
for X being set
for g being SimpleGraph of X
for v1, v2 being Element of g
for p being FinSequence of the carrier of g holds
( p is_path_of v1,v2 iff ( p . 1 = v1 & p . (len p) = v2 & ( for i being Element of NAT st 1 <= i & i < len p holds
{(p . i),(p . (i + 1))} in the SEdges of g ) & ( for i, j being Element of NAT st 1 <= i & i < len p & i < j & j < len p holds
( p . i <> p . j & {(p . i),(p . (i + 1))} <> {(p . j),(p . (j + 1))} ) ) ) );
:: deftheorem defines PATHS SGRAPH1:def 13 :
for X being set
for g being SimpleGraph of X
for v1, v2 being Element of the carrier of g holds PATHS (v1,v2) = { ss where ss is Element of the carrier of g * : ss is_path_of v1,v2 } ;
theorem
canceled;
theorem
theorem
:: deftheorem Def14 defines is_cycle_of SGRAPH1:def 14 :
for X being set
for g being SimpleGraph of X
for p being set holds
( p is_cycle_of g iff ex v being Element of the carrier of g st p in PATHS (v,v) );
begin
definition
let n,
m be
Element of
NAT ;
canceled;func K_ (
m,
n)
-> SimpleGraph of
NAT means
ex
ee being
Subset of
(TWOELEMENTSETS (Seg (m + n))) st
(
ee = { {i,j} where i, j is Element of NAT : ( i in Seg m & j in nat_interval ((m + 1),(m + n)) ) } &
it = SimpleGraphStruct(#
(Seg (m + n)),
ee #) );
existence
ex b1 being SimpleGraph of NAT ex ee being Subset of (TWOELEMENTSETS (Seg (m + n))) st
( ee = { {i,j} where i, j is Element of NAT : ( i in Seg m & j in nat_interval ((m + 1),(m + n)) ) } & b1 = SimpleGraphStruct(# (Seg (m + n)),ee #) )
uniqueness
for b1, b2 being SimpleGraph of NAT st ex ee being Subset of (TWOELEMENTSETS (Seg (m + n))) st
( ee = { {i,j} where i, j is Element of NAT : ( i in Seg m & j in nat_interval ((m + 1),(m + n)) ) } & b1 = SimpleGraphStruct(# (Seg (m + n)),ee #) ) & ex ee being Subset of (TWOELEMENTSETS (Seg (m + n))) st
( ee = { {i,j} where i, j is Element of NAT : ( i in Seg m & j in nat_interval ((m + 1),(m + n)) ) } & b2 = SimpleGraphStruct(# (Seg (m + n)),ee #) ) holds
b1 = b2
;
end;
:: deftheorem SGRAPH1:def 15 :
canceled;
:: deftheorem defines K_ SGRAPH1:def 16 :
for n, m being Element of NAT
for b3 being SimpleGraph of NAT holds
( b3 = K_ (m,n) iff ex ee being Subset of (TWOELEMENTSETS (Seg (m + n))) st
( ee = { {i,j} where i, j is Element of NAT : ( i in Seg m & j in nat_interval ((m + 1),(m + n)) ) } & b3 = SimpleGraphStruct(# (Seg (m + n)),ee #) ) );
definition
let n be
Element of
NAT ;
func K_ n -> SimpleGraph of
NAT means :
Def17:
ex
ee being
finite Subset of
(TWOELEMENTSETS (Seg n)) st
(
ee = { {i,j} where i, j is Element of NAT : ( i in Seg n & j in Seg n & i <> j ) } &
it = SimpleGraphStruct(#
(Seg n),
ee #) );
existence
ex b1 being SimpleGraph of NAT ex ee being finite Subset of (TWOELEMENTSETS (Seg n)) st
( ee = { {i,j} where i, j is Element of NAT : ( i in Seg n & j in Seg n & i <> j ) } & b1 = SimpleGraphStruct(# (Seg n),ee #) )
uniqueness
for b1, b2 being SimpleGraph of NAT st ex ee being finite Subset of (TWOELEMENTSETS (Seg n)) st
( ee = { {i,j} where i, j is Element of NAT : ( i in Seg n & j in Seg n & i <> j ) } & b1 = SimpleGraphStruct(# (Seg n),ee #) ) & ex ee being finite Subset of (TWOELEMENTSETS (Seg n)) st
( ee = { {i,j} where i, j is Element of NAT : ( i in Seg n & j in Seg n & i <> j ) } & b2 = SimpleGraphStruct(# (Seg n),ee #) ) holds
b1 = b2
;
end;
:: deftheorem Def17 defines K_ SGRAPH1:def 17 :
for n being Element of NAT
for b2 being SimpleGraph of NAT holds
( b2 = K_ n iff ex ee being finite Subset of (TWOELEMENTSETS (Seg n)) st
( ee = { {i,j} where i, j is Element of NAT : ( i in Seg n & j in Seg n & i <> j ) } & b2 = SimpleGraphStruct(# (Seg n),ee #) ) );
:: deftheorem defines TriangleGraph SGRAPH1:def 18 :
TriangleGraph = K_ 3;
theorem Th46:
theorem
theorem
theorem