begin
theorem Th1:
for
i,
k,
l being
Nat st
i >= k + l holds
i >= k
theorem
theorem Th3:
theorem Th4:
theorem Th5:
theorem Th6:
theorem Th7:
theorem Th8:
theorem Th9:
theorem Th10:
theorem
canceled;
theorem Th12:
definition
let E be non
empty set ;
let F be
Subset of
(E ^omega);
let TS be non
empty transition-system of
F;
func _bool TS -> strict transition-system of
Lex E means :
Def1:
( the
carrier of
it = bool the
carrier of
TS & ( for
S being
Subset of
TS for
w being
Element of
E ^omega for
T being
Subset of
TS holds
(
[[S,w],T] in the
Tran of
it iff (
len w = 1 &
T = w -succ_of (
S,
TS) ) ) ) );
existence
ex b1 being strict transition-system of Lex E st
( the carrier of b1 = bool the carrier of TS & ( for S being Subset of TS
for w being Element of E ^omega
for T being Subset of TS holds
( [[S,w],T] in the Tran of b1 iff ( len w = 1 & T = w -succ_of (S,TS) ) ) ) )
uniqueness
for b1, b2 being strict transition-system of Lex E st the carrier of b1 = bool the carrier of TS & ( for S being Subset of TS
for w being Element of E ^omega
for T being Subset of TS holds
( [[S,w],T] in the Tran of b1 iff ( len w = 1 & T = w -succ_of (S,TS) ) ) ) & the carrier of b2 = bool the carrier of TS & ( for S being Subset of TS
for w being Element of E ^omega
for T being Subset of TS holds
( [[S,w],T] in the Tran of b2 iff ( len w = 1 & T = w -succ_of (S,TS) ) ) ) holds
b1 = b2
end;
:: deftheorem Def1 defines _bool FSM_3:def 1 :
for E being non empty set
for F being Subset of (E ^omega)
for TS being non empty transition-system of F
for b4 being strict transition-system of Lex E holds
( b4 = _bool TS iff ( the carrier of b4 = bool the carrier of TS & ( for S being Subset of TS
for w being Element of E ^omega
for T being Subset of TS holds
( [[S,w],T] in the Tran of b4 iff ( len w = 1 & T = w -succ_of (S,TS) ) ) ) ) );
theorem Th13:
theorem Th14:
:: deftheorem Def2 defines deterministic FSM_3:def 2 :
for E being non empty set
for F being Subset of (E ^omega)
for SA being semiautomaton of F holds
( SA is deterministic iff ( transition-system(# the carrier of SA, the Tran of SA #) is deterministic & card the InitS of SA = 1 ) );
definition
let E be non
empty set ;
let F be
Subset of
(E ^omega);
let SA be non
empty semiautomaton of
F;
func _bool SA -> strict semiautomaton of
Lex E means :
Def3:
(
transition-system(# the
carrier of
it, the
Tran of
it #)
= _bool transition-system(# the
carrier of
SA, the
Tran of
SA #) & the
InitS of
it = {((<%> E) -succ_of ( the InitS of SA,SA))} );
existence
ex b1 being strict semiautomaton of Lex E st
( transition-system(# the carrier of b1, the Tran of b1 #) = _bool transition-system(# the carrier of SA, the Tran of SA #) & the InitS of b1 = {((<%> E) -succ_of ( the InitS of SA,SA))} )
uniqueness
for b1, b2 being strict semiautomaton of Lex E st transition-system(# the carrier of b1, the Tran of b1 #) = _bool transition-system(# the carrier of SA, the Tran of SA #) & the InitS of b1 = {((<%> E) -succ_of ( the InitS of SA,SA))} & transition-system(# the carrier of b2, the Tran of b2 #) = _bool transition-system(# the carrier of SA, the Tran of SA #) & the InitS of b2 = {((<%> E) -succ_of ( the InitS of SA,SA))} holds
b1 = b2
;
end;
:: deftheorem Def3 defines _bool FSM_3:def 3 :
for E being non empty set
for F being Subset of (E ^omega)
for SA being non empty semiautomaton of F
for b4 being strict semiautomaton of Lex E holds
( b4 = _bool SA iff ( transition-system(# the carrier of b4, the Tran of b4 #) = _bool transition-system(# the carrier of SA, the Tran of SA #) & the InitS of b4 = {((<%> E) -succ_of ( the InitS of SA,SA))} ) );
theorem Th15:
:: deftheorem defines left-Lang FSM_3:def 4 :
for E being non empty set
for F being Subset of (E ^omega)
for SA being non empty semiautomaton of F
for Q being Subset of SA holds left-Lang Q = { w where w is Element of E ^omega : Q meets w -succ_of ( the InitS of SA,SA) } ;
theorem Th16:
:: deftheorem Def5 defines deterministic FSM_3:def 5 :
for E being non empty set
for F being Subset of (E ^omega)
for A being automaton of F holds
( A is deterministic iff semiautomaton(# the carrier of A, the Tran of A, the InitS of A #) is deterministic );
definition
let E be non
empty set ;
let F be
Subset of
(E ^omega);
let A be non
empty automaton of
F;
func _bool A -> strict automaton of
Lex E means :
Def6:
(
semiautomaton(# the
carrier of
it, the
Tran of
it, the
InitS of
it #)
= _bool semiautomaton(# the
carrier of
A, the
Tran of
A, the
InitS of
A #) & the
FinalS of
it = { Q where Q is Element of it : Q meets the FinalS of A } );
existence
ex b1 being strict automaton of Lex E st
( semiautomaton(# the carrier of b1, the Tran of b1, the InitS of b1 #) = _bool semiautomaton(# the carrier of A, the Tran of A, the InitS of A #) & the FinalS of b1 = { Q where Q is Element of b1 : Q meets the FinalS of A } )
uniqueness
for b1, b2 being strict automaton of Lex E st semiautomaton(# the carrier of b1, the Tran of b1, the InitS of b1 #) = _bool semiautomaton(# the carrier of A, the Tran of A, the InitS of A #) & the FinalS of b1 = { Q where Q is Element of b1 : Q meets the FinalS of A } & semiautomaton(# the carrier of b2, the Tran of b2, the InitS of b2 #) = _bool semiautomaton(# the carrier of A, the Tran of A, the InitS of A #) & the FinalS of b2 = { Q where Q is Element of b2 : Q meets the FinalS of A } holds
b1 = b2
;
end;
:: deftheorem Def6 defines _bool FSM_3:def 6 :
for E being non empty set
for F being Subset of (E ^omega)
for A being non empty automaton of F
for b4 being strict automaton of Lex E holds
( b4 = _bool A iff ( semiautomaton(# the carrier of b4, the Tran of b4, the InitS of b4 #) = _bool semiautomaton(# the carrier of A, the Tran of A, the InitS of A #) & the FinalS of b4 = { Q where Q is Element of b4 : Q meets the FinalS of A } ) );
theorem Th17:
:: deftheorem defines right-Lang FSM_3:def 7 :
for E being non empty set
for F being Subset of (E ^omega)
for A being non empty automaton of F
for Q being Subset of A holds right-Lang Q = { w where w is Element of E ^omega : w -succ_of (Q,A) meets the FinalS of A } ;
theorem Th18:
:: deftheorem defines Lang FSM_3:def 8 :
for E being non empty set
for F being Subset of (E ^omega)
for A being non empty automaton of F holds Lang A = { u where u is Element of E ^omega : ex p, q being Element of A st
( p in the InitS of A & q in the FinalS of A & p,u ==>* q,A ) } ;
theorem Th19:
theorem Th20:
theorem
theorem
theorem Th23:
theorem Th24:
theorem Th25:
:: deftheorem Def9 defines chop FSM_3:def 9 :
for E being non empty set
for u, v, b4 being Element of E ^omega holds
( ( ex w being Element of E ^omega st w ^ v = u implies ( b4 = chop (u,v) iff for w being Element of E ^omega st w ^ v = u holds
b4 = w ) ) & ( ( for w being Element of E ^omega holds not w ^ v = u ) implies ( b4 = chop (u,v) iff b4 = u ) ) );
theorem Th26:
theorem Th27:
theorem Th28:
theorem Th29:
theorem Th30:
theorem Th31:
theorem Th32:
theorem Th33:
theorem Th34:
theorem Th35:
theorem Th36:
theorem Th37:
theorem Th38:
theorem Th39:
theorem
theorem