begin
definition
let S be
ManySortedSign ;
defpred S1[
set ]
means ex
op,
v being
set st
( $1
= [op,v] &
op in the
carrier' of
S &
v in the
carrier of
S & ex
n being
Nat ex
args being
Element of the
carrier of
S * st
( the
Arity of
S . op = args &
n in dom args &
args . n = v ) );
func InducedEdges S -> set means :
Def1:
for
x being
set holds
(
x in it iff ex
op,
v being
set st
(
x = [op,v] &
op in the
carrier' of
S &
v in the
carrier of
S & ex
n being
Nat ex
args being
Element of the
carrier of
S * st
( the
Arity of
S . op = args &
n in dom args &
args . n = v ) ) );
existence
ex b1 being set st
for x being set holds
( x in b1 iff ex op, v being set st
( x = [op,v] & op in the carrier' of S & v in the carrier of S & ex n being Nat ex args being Element of the carrier of S * st
( the Arity of S . op = args & n in dom args & args . n = v ) ) )
uniqueness
for b1, b2 being set st ( for x being set holds
( x in b1 iff ex op, v being set st
( x = [op,v] & op in the carrier' of S & v in the carrier of S & ex n being Nat ex args being Element of the carrier of S * st
( the Arity of S . op = args & n in dom args & args . n = v ) ) ) ) & ( for x being set holds
( x in b2 iff ex op, v being set st
( x = [op,v] & op in the carrier' of S & v in the carrier of S & ex n being Nat ex args being Element of the carrier of S * st
( the Arity of S . op = args & n in dom args & args . n = v ) ) ) ) holds
b1 = b2
end;
:: deftheorem Def1 defines InducedEdges MSSCYC_2:def 1 :
for S being ManySortedSign
for b2 being set holds
( b2 = InducedEdges S iff for x being set holds
( x in b2 iff ex op, v being set st
( x = [op,v] & op in the carrier' of S & v in the carrier of S & ex n being Nat ex args being Element of the carrier of S * st
( the Arity of S . op = args & n in dom args & args . n = v ) ) ) );
theorem Th1:
:: deftheorem Def2 defines InducedSource MSSCYC_2:def 2 :
for S being ManySortedSign
for b2 being Function of (InducedEdges S), the carrier of S holds
( b2 = InducedSource S iff for e being set st e in InducedEdges S holds
b2 . e = e `2 );
:: deftheorem Def3 defines InducedTarget MSSCYC_2:def 3 :
for S being ManySortedSign
for b2 being Function of (InducedEdges S), the carrier of S holds
( b2 = InducedTarget S iff for e being set st e in InducedEdges S holds
b2 . e = the ResultSort of S . (e `1) );
:: deftheorem defines InducedGraph MSSCYC_2:def 4 :
for S being non empty ManySortedSign holds InducedGraph S = MultiGraphStruct(# the carrier of S,(InducedEdges S),(InducedSource S),(InducedTarget S) #);
theorem Th2:
theorem
theorem
theorem Th5:
theorem