:: PETRI semantic presentation

definition
let A be non empty set , B be non empty set ;
let r be non empty Relation of A,B;
redefine mode Element as Element of c3 -> Element of [:a1,a2:];
coherence
for b1 being Element of r holds b1 is Element of [:A,B:]
proof end;
end;

definition
attr a1 is strict;
struct PT_net_Str -> ;
aggr PT_net_Str(# Places, Transitions, S-T_Arcs, T-S_Arcs #) -> PT_net_Str ;
sel Places c1 -> non empty set ;
sel Transitions c1 -> non empty set ;
sel S-T_Arcs c1 -> non empty Relation of the Places of a1,the Transitions of a1;
sel T-S_Arcs c1 -> non empty Relation of the Transitions of a1,the Places of a1;
end;

definition
let PTN be PT_net_Str ;
mode place of a1 is Element of the Places of a1;
mode places of a1 is Element of the Places of a1;
mode transition of a1 is Element of the Transitions of a1;
mode transitions of a1 is Element of the Transitions of a1;
mode S-T_arc of a1 is Element of the S-T_Arcs of a1;
mode T-S_arc of a1 is Element of the T-S_Arcs of a1;
end;

definition
let PTN be PT_net_Str ;
let x be S-T_arc of PTN;
redefine func `1 as c2 `1 -> place of a1;
coherence
x `1 is place of PTN
proof end;
redefine func `2 as c2 `2 -> transition of a1;
coherence
x `2 is transition of PTN
proof end;
end;

definition
let PTN be PT_net_Str ;
let x be T-S_arc of PTN;
redefine func `1 as c2 `1 -> transition of a1;
coherence
x `1 is transition of PTN
proof end;
redefine func `2 as c2 `2 -> place of a1;
coherence
x `2 is place of PTN
proof end;
end;

definition
let PTN be PT_net_Str ;
let S0 be Subset of the Places of PTN;
func *' c2 -> Subset of the Transitions of a1 equals :: PETRI:def 1
{ t where t is transition of PTN : ex f being T-S_arc of PTN ex s being place of PTN st
( s in S0 & f = [t,s] )
}
;
coherence
{ t where t is transition of PTN : ex f being T-S_arc of PTN ex s being place of PTN st
( s in S0 & f = [t,s] )
}
is Subset of the Transitions of PTN
proof end;
correctness
;
func c2 *' -> Subset of the Transitions of a1 equals :: PETRI:def 2
{ t where t is transition of PTN : ex f being S-T_arc of PTN ex s being place of PTN st
( s in S0 & f = [s,t] )
}
;
coherence
{ t where t is transition of PTN : ex f being S-T_arc of PTN ex s being place of PTN st
( s in S0 & f = [s,t] )
}
is Subset of the Transitions of PTN
proof end;
correctness
;
end;

:: deftheorem Def1 defines *' PETRI:def 1 :
for PTN being PT_net_Str
for S0 being Subset of the Places of PTN holds *' S0 = { t where t is transition of PTN : ex f being T-S_arc of PTN ex s being place of PTN st
( s in S0 & f = [t,s] )
}
;

:: deftheorem Def2 defines *' PETRI:def 2 :
for PTN being PT_net_Str
for S0 being Subset of the Places of PTN holds S0 *' = { t where t is transition of PTN : ex f being S-T_arc of PTN ex s being place of PTN st
( s in S0 & f = [s,t] )
}
;

theorem Th1: :: PETRI:1
for PTN being PT_net_Str
for S0 being Subset of the Places of PTN holds *' S0 = { (f `1 ) where f is T-S_arc of PTN : f `2 in S0 }
proof end;

theorem Th2: :: PETRI:2
for PTN being PT_net_Str
for S0 being Subset of the Places of PTN
for x being set holds
( x in *' S0 iff ex f being T-S_arc of PTN ex s being place of PTN st
( s in S0 & f = [x,s] ) )
proof end;

theorem Th3: :: PETRI:3
for PTN being PT_net_Str
for S0 being Subset of the Places of PTN holds S0 *' = { (f `2 ) where f is S-T_arc of PTN : f `1 in S0 }
proof end;

theorem Th4: :: PETRI:4
for PTN being PT_net_Str
for S0 being Subset of the Places of PTN
for x being set holds
( x in S0 *' iff ex f being S-T_arc of PTN ex s being place of PTN st
( s in S0 & f = [s,x] ) )
proof end;

definition
let PTN be PT_net_Str ;
let T0 be Subset of the Transitions of PTN;
func *' c2 -> Subset of the Places of a1 equals :: PETRI:def 3
{ s where s is place of PTN : ex f being S-T_arc of PTN ex t being transition of PTN st
( t in T0 & f = [s,t] )
}
;
coherence
{ s where s is place of PTN : ex f being S-T_arc of PTN ex t being transition of PTN st
( t in T0 & f = [s,t] )
}
is Subset of the Places of PTN
proof end;
correctness
;
func c2 *' -> Subset of the Places of a1 equals :: PETRI:def 4
{ s where s is place of PTN : ex f being T-S_arc of PTN ex t being transition of PTN st
( t in T0 & f = [t,s] )
}
;
coherence
{ s where s is place of PTN : ex f being T-S_arc of PTN ex t being transition of PTN st
( t in T0 & f = [t,s] )
}
is Subset of the Places of PTN
proof end;
correctness
;
end;

:: deftheorem Def3 defines *' PETRI:def 3 :
for PTN being PT_net_Str
for T0 being Subset of the Transitions of PTN holds *' T0 = { s where s is place of PTN : ex f being S-T_arc of PTN ex t being transition of PTN st
( t in T0 & f = [s,t] )
}
;

:: deftheorem Def4 defines *' PETRI:def 4 :
for PTN being PT_net_Str
for T0 being Subset of the Transitions of PTN holds T0 *' = { s where s is place of PTN : ex f being T-S_arc of PTN ex t being transition of PTN st
( t in T0 & f = [t,s] )
}
;

theorem Th5: :: PETRI:5
for PTN being PT_net_Str
for T0 being Subset of the Transitions of PTN holds *' T0 = { (f `1 ) where f is S-T_arc of PTN : f `2 in T0 }
proof end;

theorem Th6: :: PETRI:6
for PTN being PT_net_Str
for T0 being Subset of the Transitions of PTN
for x being set holds
( x in *' T0 iff ex f being S-T_arc of PTN ex t being transition of PTN st
( t in T0 & f = [x,t] ) )
proof end;

theorem Th7: :: PETRI:7
for PTN being PT_net_Str
for T0 being Subset of the Transitions of PTN holds T0 *' = { (f `2 ) where f is T-S_arc of PTN : f `1 in T0 }
proof end;

theorem Th8: :: PETRI:8
for PTN being PT_net_Str
for T0 being Subset of the Transitions of PTN
for x being set holds
( x in T0 *' iff ex f being T-S_arc of PTN ex t being transition of PTN st
( t in T0 & f = [t,x] ) )
proof end;

theorem Th9: :: PETRI:9
for PTN being PT_net_Str holds *' ({} the Places of PTN) = {}
proof end;

theorem Th10: :: PETRI:10
for PTN being PT_net_Str holds ({} the Places of PTN) *' = {}
proof end;

theorem Th11: :: PETRI:11
for PTN being PT_net_Str holds *' ({} the Transitions of PTN) = {}
proof end;

theorem Th12: :: PETRI:12
for PTN being PT_net_Str holds ({} the Transitions of PTN) *' = {}
proof end;

definition
let PTN be PT_net_Str ;
let IT be Subset of the Places of PTN;
attr a2 is Deadlock-like means :: PETRI:def 5
*' IT is Subset of (IT *' );
end;

:: deftheorem Def5 defines Deadlock-like PETRI:def 5 :
for PTN being PT_net_Str
for IT being Subset of the Places of PTN holds
( IT is Deadlock-like iff *' IT is Subset of (IT *' ) );

definition
let IT be PT_net_Str ;
attr a1 is With_Deadlocks means :: PETRI:def 6
ex S being Subset of the Places of IT st S is Deadlock-like;
end;

:: deftheorem Def6 defines With_Deadlocks PETRI:def 6 :
for IT being PT_net_Str holds
( IT is With_Deadlocks iff ex S being Subset of the Places of IT st S is Deadlock-like );

registration
cluster With_Deadlocks PT_net_Str ;
existence
ex b1 being PT_net_Str st b1 is With_Deadlocks
proof end;
end;

definition
let PTN be PT_net_Str ;
let IT be Subset of the Places of PTN;
attr a2 is Trap-like means :: PETRI:def 7
IT *' is Subset of (*' IT);
end;

:: deftheorem Def7 defines Trap-like PETRI:def 7 :
for PTN being PT_net_Str
for IT being Subset of the Places of PTN holds
( IT is Trap-like iff IT *' is Subset of (*' IT) );

definition
let IT be PT_net_Str ;
attr a1 is With_Traps means :: PETRI:def 8
ex S being Subset of the Places of IT st S is Trap-like;
end;

:: deftheorem Def8 defines With_Traps PETRI:def 8 :
for IT being PT_net_Str holds
( IT is With_Traps iff ex S being Subset of the Places of IT st S is Trap-like );

registration
cluster With_Traps PT_net_Str ;
existence
ex b1 being PT_net_Str st b1 is With_Traps
proof end;
end;

definition
let A be non empty set , B be non empty set ;
let r be non empty Relation of A,B;
redefine func ~ as c3 ~ -> non empty Relation of a2,a1;
coherence
r ~ is non empty Relation of B,A
proof end;
end;

definition
let PTN be PT_net_Str ;
func c1 .: -> strict PT_net_Str equals :: PETRI:def 9
PT_net_Str(# the Places of PTN,the Transitions of PTN,(the T-S_Arcs of PTN ~ ),(the S-T_Arcs of PTN ~ ) #);
correctness
coherence
PT_net_Str(# the Places of PTN,the Transitions of PTN,(the T-S_Arcs of PTN ~ ),(the S-T_Arcs of PTN ~ ) #) is strict PT_net_Str
;
;
end;

:: deftheorem Def9 defines .: PETRI:def 9 :
for PTN being PT_net_Str holds PTN .: = PT_net_Str(# the Places of PTN,the Transitions of PTN,(the T-S_Arcs of PTN ~ ),(the S-T_Arcs of PTN ~ ) #);

theorem Th13: :: PETRI:13
for PTN being PT_net_Str holds (PTN .: ) .: = PT_net_Str(# the Places of PTN,the Transitions of PTN,the S-T_Arcs of PTN,the T-S_Arcs of PTN #) ;

theorem Th14: :: PETRI:14
for PTN being PT_net_Str holds
( the Places of PTN = the Places of (PTN .: ) & the Transitions of PTN = the Transitions of (PTN .: ) & the S-T_Arcs of PTN ~ = the T-S_Arcs of (PTN .: ) & the T-S_Arcs of PTN ~ = the S-T_Arcs of (PTN .: ) ) ;

definition
let PTN be PT_net_Str ;
let S0 be Subset of the Places of PTN;
func c2 .: -> Subset of the Places of (a1 .: ) equals :: PETRI:def 10
S0;
coherence
S0 is Subset of the Places of (PTN .: )
;
end;

:: deftheorem Def10 defines .: PETRI:def 10 :
for PTN being PT_net_Str
for S0 being Subset of the Places of PTN holds S0 .: = S0;

definition
let PTN be PT_net_Str ;
let s be place of PTN;
func c2 .: -> place of (a1 .: ) equals :: PETRI:def 11
s;
coherence
s is place of (PTN .: )
;
end;

:: deftheorem Def11 defines .: PETRI:def 11 :
for PTN being PT_net_Str
for s being place of PTN holds s .: = s;

definition
let PTN be PT_net_Str ;
let S0 be Subset of the Places of (PTN .: );
func .: c2 -> Subset of the Places of a1 equals :: PETRI:def 12
S0;
coherence
S0 is Subset of the Places of PTN
;
end;

:: deftheorem Def12 defines .: PETRI:def 12 :
for PTN being PT_net_Str
for S0 being Subset of the Places of (PTN .: ) holds .: S0 = S0;

definition
let PTN be PT_net_Str ;
let s be place of (PTN .: );
func .: c2 -> place of a1 equals :: PETRI:def 13
s;
coherence
s is place of PTN
;
end;

:: deftheorem Def13 defines .: PETRI:def 13 :
for PTN being PT_net_Str
for s being place of (PTN .: ) holds .: s = s;

definition
let PTN be PT_net_Str ;
let T0 be Subset of the Transitions of PTN;
func c2 .: -> Subset of the Transitions of (a1 .: ) equals :: PETRI:def 14
T0;
coherence
T0 is Subset of the Transitions of (PTN .: )
;
end;

:: deftheorem Def14 defines .: PETRI:def 14 :
for PTN being PT_net_Str
for T0 being Subset of the Transitions of PTN holds T0 .: = T0;

definition
let PTN be PT_net_Str ;
let t be transition of PTN;
func c2 .: -> transition of (a1 .: ) equals :: PETRI:def 15
t;
coherence
t is transition of (PTN .: )
;
end;

:: deftheorem Def15 defines .: PETRI:def 15 :
for PTN being PT_net_Str
for t being transition of PTN holds t .: = t;

definition
let PTN be PT_net_Str ;
let T0 be Subset of the Transitions of (PTN .: );
func .: c2 -> Subset of the Transitions of a1 equals :: PETRI:def 16
T0;
coherence
T0 is Subset of the Transitions of PTN
;
end;

:: deftheorem Def16 defines .: PETRI:def 16 :
for PTN being PT_net_Str
for T0 being Subset of the Transitions of (PTN .: ) holds .: T0 = T0;

definition
let PTN be PT_net_Str ;
let t be transition of (PTN .: );
func .: c2 -> transition of a1 equals :: PETRI:def 17
t;
coherence
t is transition of PTN
;
end;

:: deftheorem Def17 defines .: PETRI:def 17 :
for PTN being PT_net_Str
for t being transition of (PTN .: ) holds .: t = t;

theorem Th15: :: PETRI:15
for PTN being PT_net_Str
for S being Subset of the Places of PTN holds (S .: ) *' = *' S
proof end;

theorem Th16: :: PETRI:16
for PTN being PT_net_Str
for S being Subset of the Places of PTN holds *' (S .: ) = S *'
proof end;

theorem Th17: :: PETRI:17
for PTN being PT_net_Str
for S being Subset of the Places of PTN holds
( S is Deadlock-like iff S .: is Trap-like )
proof end;

theorem Th18: :: PETRI:18
for PTN being PT_net_Str
for S being Subset of the Places of PTN holds
( S is Trap-like iff S .: is Deadlock-like )
proof end;

theorem Th19: :: PETRI:19
for PTN being PT_net_Str
for t being transition of PTN
for S0 being Subset of the Places of PTN holds
( t in S0 *' iff *' {t} meets S0 )
proof end;

theorem Th20: :: PETRI:20
for PTN being PT_net_Str
for t being transition of PTN
for S0 being Subset of the Places of PTN holds
( t in *' S0 iff {t} *' meets S0 )
proof end;