begin
:: deftheorem defines -succ_of FSM_1:def 1 :
for IAlph being non empty set
for fsm being non empty FSM of IAlph
for s being Element of IAlph
for q being State of fsm holds s -succ_of q = the Tran of fsm . [q,s];
:: deftheorem Def2 defines -admissible FSM_1:def 2 :
for IAlph being non empty set
for fsm being non empty FSM of IAlph
for q being State of fsm
for w being FinSequence of IAlph
for b5 being FinSequence of the carrier of fsm holds
( b5 = (q,w) -admissible iff ( b5 . 1 = q & len b5 = (len w) + 1 & ( for i being Nat st 1 <= i & i <= len w holds
ex wi being Element of IAlph ex qi, qi1 being State of fsm st
( wi = w . i & qi = b5 . i & qi1 = b5 . (i + 1) & wi -succ_of qi = qi1 ) ) ) );
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem Th16:
:: deftheorem Def3 defines -leads_to FSM_1:def 3 :
for IAlph being non empty set
for fsm being non empty FSM of IAlph
for w being FinSequence of IAlph
for q1, q2 being State of fsm holds
( q1,w -leads_to q2 iff ((q1,w) -admissible) . ((len w) + 1) = q2 );
theorem Th17:
:: deftheorem Def4 defines is_admissible_for FSM_1:def 4 :
for IAlph being non empty set
for fsm being non empty FSM of IAlph
for w being FinSequence of IAlph
for qseq being FinSequence of the carrier of fsm holds
( qseq is_admissible_for w iff ex q1 being State of fsm st
( q1 = qseq . 1 & (q1,w) -admissible = qseq ) );
theorem
:: deftheorem Def5 defines leads_to_under FSM_1:def 5 :
for IAlph being non empty set
for fsm being non empty FSM of IAlph
for q being State of fsm
for w being FinSequence of IAlph
for b5 being State of fsm holds
( b5 = q leads_to_under w iff q,w -leads_to b5 );
theorem Th19:
theorem Th20:
theorem Th21:
theorem Th22:
theorem Th23:
begin
definition
let IAlph be
set ;
let OAlph be non
empty set ;
attr c3 is
strict ;
struct Mealy-FSM of
IAlph,
OAlph -> FSM of
IAlph;
aggr Mealy-FSM(#
carrier,
Tran,
OFun,
InitS #)
-> Mealy-FSM of
IAlph,
OAlph;
sel OFun c3 -> Function of
[: the carrier of c3,IAlph:],
OAlph;
attr c3 is
strict ;
struct Moore-FSM of
IAlph,
OAlph -> FSM of
IAlph;
aggr Moore-FSM(#
carrier,
Tran,
OFun,
InitS #)
-> Moore-FSM of
IAlph,
OAlph;
sel OFun c3 -> Function of the
carrier of
c3,
OAlph;
end;
registration
let IAlph be
set ;
let OAlph be non
empty set ;
let X be non
empty finite set ;
let T be
Function of
[:X,IAlph:],
X;
let O be
Function of
[:X,IAlph:],
OAlph;
let I be
Element of
X;
cluster Mealy-FSM(#
X,
T,
O,
I #)
-> non
empty finite ;
coherence
( Mealy-FSM(# X,T,O,I #) is finite & not Mealy-FSM(# X,T,O,I #) is empty )
;
end;
registration
let IAlph be
set ;
let OAlph be non
empty set ;
let X be non
empty finite set ;
let T be
Function of
[:X,IAlph:],
X;
let O be
Function of
X,
OAlph;
let I be
Element of
X;
cluster Moore-FSM(#
X,
T,
O,
I #)
-> non
empty finite ;
coherence
( Moore-FSM(# X,T,O,I #) is finite & not Moore-FSM(# X,T,O,I #) is empty )
;
end;
definition
let IAlph,
OAlph be non
empty set ;
let tfsm be non
empty Mealy-FSM of
IAlph,
OAlph;
let qt be
State of
tfsm;
let w be
FinSequence of
IAlph;
func (
qt,
w)
-response -> FinSequence of
OAlph means :
Def6:
(
len it = len w & ( for
i being
Element of
NAT st
i in dom w holds
it . i = the
OFun of
tfsm . [(((qt,w) -admissible) . i),(w . i)] ) );
existence
ex b1 being FinSequence of OAlph st
( len b1 = len w & ( for i being Element of NAT st i in dom w holds
b1 . i = the OFun of tfsm . [(((qt,w) -admissible) . i),(w . i)] ) )
uniqueness
for b1, b2 being FinSequence of OAlph st len b1 = len w & ( for i being Element of NAT st i in dom w holds
b1 . i = the OFun of tfsm . [(((qt,w) -admissible) . i),(w . i)] ) & len b2 = len w & ( for i being Element of NAT st i in dom w holds
b2 . i = the OFun of tfsm . [(((qt,w) -admissible) . i),(w . i)] ) holds
b1 = b2
end;
:: deftheorem Def6 defines -response FSM_1:def 6 :
for IAlph, OAlph being non empty set
for tfsm being non empty Mealy-FSM of IAlph,OAlph
for qt being State of tfsm
for w being FinSequence of IAlph
for b6 being FinSequence of OAlph holds
( b6 = (qt,w) -response iff ( len b6 = len w & ( for i being Element of NAT st i in dom w holds
b6 . i = the OFun of tfsm . [(((qt,w) -admissible) . i),(w . i)] ) ) );
theorem Th24:
:: deftheorem Def7 defines -response FSM_1:def 7 :
for IAlph, OAlph being non empty set
for sfsm being non empty Moore-FSM of IAlph,OAlph
for qs being State of sfsm
for w being FinSequence of IAlph
for b6 being FinSequence of OAlph holds
( b6 = (qs,w) -response iff ( len b6 = (len w) + 1 & ( for i being Element of NAT st i in Seg ((len w) + 1) holds
b6 . i = the OFun of sfsm . (((qs,w) -admissible) . i) ) ) );
theorem Th25:
theorem Th26:
theorem Th27:
for
IAlph,
OAlph being non
empty set for
w1,
w2 being
FinSequence of
IAlph for
tfsm1,
tfsm2 being non
empty Mealy-FSM of
IAlph,
OAlph for
q11,
q12 being
State of
tfsm1 for
q21,
q22 being
State of
tfsm2 st
q11,
w1 -leads_to q12 &
q21,
w1 -leads_to q22 & (
q12,
w2)
-response <> (
q22,
w2)
-response holds
(
q11,
(w1 ^ w2))
-response <> (
q21,
(w1 ^ w2))
-response
:: deftheorem defines is_similar_to FSM_1:def 8 :
for IAlph, OAlph being non empty set
for tfsm being non empty Mealy-FSM of IAlph,OAlph
for sfsm being non empty Moore-FSM of IAlph,OAlph holds
( tfsm is_similar_to sfsm iff for tw being FinSequence of IAlph holds <*( the OFun of sfsm . the InitS of sfsm)*> ^ (( the InitS of tfsm,tw) -response) = ( the InitS of sfsm,tw) -response );
theorem
theorem
begin
definition
let IAlph,
OAlph be non
empty set ;
let tfsm1,
tfsm2 be non
empty Mealy-FSM of
IAlph,
OAlph;
pred tfsm1,
tfsm2 -are_equivalent means :
Def9:
for
w being
FinSequence of
IAlph holds ( the
InitS of
tfsm1,
w)
-response = ( the
InitS of
tfsm2,
w)
-response ;
reflexivity
for tfsm1 being non empty Mealy-FSM of IAlph,OAlph
for w being FinSequence of IAlph holds ( the InitS of tfsm1,w) -response = ( the InitS of tfsm1,w) -response
;
symmetry
for tfsm1, tfsm2 being non empty Mealy-FSM of IAlph,OAlph st ( for w being FinSequence of IAlph holds ( the InitS of tfsm1,w) -response = ( the InitS of tfsm2,w) -response ) holds
for w being FinSequence of IAlph holds ( the InitS of tfsm2,w) -response = ( the InitS of tfsm1,w) -response
;
end;
:: deftheorem Def9 defines -are_equivalent FSM_1:def 9 :
for IAlph, OAlph being non empty set
for tfsm1, tfsm2 being non empty Mealy-FSM of IAlph,OAlph holds
( tfsm1,tfsm2 -are_equivalent iff for w being FinSequence of IAlph holds ( the InitS of tfsm1,w) -response = ( the InitS of tfsm2,w) -response );
theorem Th30:
definition
let IAlph,
OAlph be non
empty set ;
let tfsm be non
empty Mealy-FSM of
IAlph,
OAlph;
let qa,
qb be
State of
tfsm;
pred qa,
qb -are_equivalent means :
Def10:
for
w being
FinSequence of
IAlph holds (
qa,
w)
-response = (
qb,
w)
-response ;
reflexivity
for qa being State of tfsm
for w being FinSequence of IAlph holds (qa,w) -response = (qa,w) -response
;
symmetry
for qa, qb being State of tfsm st ( for w being FinSequence of IAlph holds (qa,w) -response = (qb,w) -response ) holds
for w being FinSequence of IAlph holds (qb,w) -response = (qa,w) -response
;
end;
:: deftheorem Def10 defines -are_equivalent FSM_1:def 10 :
for IAlph, OAlph being non empty set
for tfsm being non empty Mealy-FSM of IAlph,OAlph
for qa, qb being State of tfsm holds
( qa,qb -are_equivalent iff for w being FinSequence of IAlph holds (qa,w) -response = (qb,w) -response );
theorem
canceled;
theorem
canceled;
theorem
theorem Th34:
theorem Th35:
theorem Th36:
theorem
:: deftheorem Def11 defines -equivalent FSM_1:def 11 :
for IAlph, OAlph being non empty set
for tfsm being non empty Mealy-FSM of IAlph,OAlph
for qa, qb being State of tfsm
for k being Nat holds
( k -equivalent qa,qb iff for w being FinSequence of IAlph st len w <= k holds
(qa,w) -response = (qb,w) -response );
theorem Th38:
theorem Th39:
theorem Th40:
theorem Th41:
theorem Th42:
theorem Th43:
theorem Th44:
definition
let IAlph,
OAlph be non
empty set ;
let tfsm be non
empty Mealy-FSM of
IAlph,
OAlph;
let i be
Nat;
func i -eq_states_EqR tfsm -> Equivalence_Relation of the
carrier of
tfsm means :
Def12:
for
qa,
qb being
State of
tfsm holds
(
[qa,qb] in it iff
i -equivalent qa,
qb );
existence
ex b1 being Equivalence_Relation of the carrier of tfsm st
for qa, qb being State of tfsm holds
( [qa,qb] in b1 iff i -equivalent qa,qb )
uniqueness
for b1, b2 being Equivalence_Relation of the carrier of tfsm st ( for qa, qb being State of tfsm holds
( [qa,qb] in b1 iff i -equivalent qa,qb ) ) & ( for qa, qb being State of tfsm holds
( [qa,qb] in b2 iff i -equivalent qa,qb ) ) holds
b1 = b2
end;
:: deftheorem Def12 defines -eq_states_EqR FSM_1:def 12 :
for IAlph, OAlph being non empty set
for tfsm being non empty Mealy-FSM of IAlph,OAlph
for i being Nat
for b5 being Equivalence_Relation of the carrier of tfsm holds
( b5 = i -eq_states_EqR tfsm iff for qa, qb being State of tfsm holds
( [qa,qb] in b5 iff i -equivalent qa,qb ) );
:: deftheorem defines -eq_states_partition FSM_1:def 13 :
for IAlph, OAlph being non empty set
for tfsm being non empty Mealy-FSM of IAlph,OAlph
for i being Nat holds i -eq_states_partition tfsm = Class (i -eq_states_EqR tfsm);
theorem Th45:
theorem Th46:
theorem
theorem Th48:
theorem Th49:
theorem Th50:
theorem
theorem Th52:
:: deftheorem Def14 defines final FSM_1:def 14 :
for IAlph, OAlph being non empty set
for tfsm being non empty Mealy-FSM of IAlph,OAlph
for IT being a_partition of the carrier of tfsm holds
( IT is final iff for qa, qb being State of tfsm holds
( qa,qb -are_equivalent iff ex X being Element of IT st
( qa in X & qb in X ) ) );
theorem
theorem Th54:
theorem
:: deftheorem Def15 defines final_states_partition FSM_1:def 15 :
for IAlph, OAlph being non empty set
for tfsm being non empty finite Mealy-FSM of IAlph,OAlph
for b4 being a_partition of the carrier of tfsm holds
( b4 = final_states_partition tfsm iff b4 is final );
theorem Th56:
begin
definition
let IAlph,
OAlph be non
empty set ;
let tfsm be non
empty finite Mealy-FSM of
IAlph,
OAlph;
let qf be
Element of
final_states_partition tfsm;
let s be
Element of
IAlph;
func (
s,
qf)
-succ_class -> Element of
final_states_partition tfsm means :
Def16:
ex
q being
State of
tfsm ex
n being
Element of
NAT st
(
q in qf &
n + 1
= card the
carrier of
tfsm &
it = Class (
(n -eq_states_EqR tfsm),
( the Tran of tfsm . [q,s])) );
existence
ex b1 being Element of final_states_partition tfsm ex q being State of tfsm ex n being Element of NAT st
( q in qf & n + 1 = card the carrier of tfsm & b1 = Class ((n -eq_states_EqR tfsm),( the Tran of tfsm . [q,s])) )
uniqueness
for b1, b2 being Element of final_states_partition tfsm st ex q being State of tfsm ex n being Element of NAT st
( q in qf & n + 1 = card the carrier of tfsm & b1 = Class ((n -eq_states_EqR tfsm),( the Tran of tfsm . [q,s])) ) & ex q being State of tfsm ex n being Element of NAT st
( q in qf & n + 1 = card the carrier of tfsm & b2 = Class ((n -eq_states_EqR tfsm),( the Tran of tfsm . [q,s])) ) holds
b1 = b2
end;
:: deftheorem Def16 defines -succ_class FSM_1:def 16 :
for IAlph, OAlph being non empty set
for tfsm being non empty finite Mealy-FSM of IAlph,OAlph
for qf being Element of final_states_partition tfsm
for s being Element of IAlph
for b6 being Element of final_states_partition tfsm holds
( b6 = (s,qf) -succ_class iff ex q being State of tfsm ex n being Element of NAT st
( q in qf & n + 1 = card the carrier of tfsm & b6 = Class ((n -eq_states_EqR tfsm),( the Tran of tfsm . [q,s])) ) );
:: deftheorem Def17 defines -class_response FSM_1:def 17 :
for IAlph, OAlph being non empty set
for tfsm being non empty finite Mealy-FSM of IAlph,OAlph
for qf being Element of final_states_partition tfsm
for s being Element of IAlph
for b6 being Element of OAlph holds
( b6 = (qf,s) -class_response iff ex q being State of tfsm st
( q in qf & b6 = the OFun of tfsm . [q,s] ) );
definition
let IAlph,
OAlph be non
empty set ;
let tfsm be non
empty finite Mealy-FSM of
IAlph,
OAlph;
func the_reduction_of tfsm -> strict Mealy-FSM of
IAlph,
OAlph means :
Def18:
( the
carrier of
it = final_states_partition tfsm & ( for
Q being
State of
it for
s being
Element of
IAlph for
q being
State of
tfsm st
q in Q holds
( the
Tran of
tfsm . (
q,
s)
in the
Tran of
it . (
Q,
s) & the
OFun of
tfsm . (
q,
s)
= the
OFun of
it . (
Q,
s) ) ) & the
InitS of
tfsm in the
InitS of
it );
existence
ex b1 being strict Mealy-FSM of IAlph,OAlph st
( the carrier of b1 = final_states_partition tfsm & ( for Q being State of b1
for s being Element of IAlph
for q being State of tfsm st q in Q holds
( the Tran of tfsm . (q,s) in the Tran of b1 . (Q,s) & the OFun of tfsm . (q,s) = the OFun of b1 . (Q,s) ) ) & the InitS of tfsm in the InitS of b1 )
uniqueness
for b1, b2 being strict Mealy-FSM of IAlph,OAlph st the carrier of b1 = final_states_partition tfsm & ( for Q being State of b1
for s being Element of IAlph
for q being State of tfsm st q in Q holds
( the Tran of tfsm . (q,s) in the Tran of b1 . (Q,s) & the OFun of tfsm . (q,s) = the OFun of b1 . (Q,s) ) ) & the InitS of tfsm in the InitS of b1 & the carrier of b2 = final_states_partition tfsm & ( for Q being State of b2
for s being Element of IAlph
for q being State of tfsm st q in Q holds
( the Tran of tfsm . (q,s) in the Tran of b2 . (Q,s) & the OFun of tfsm . (q,s) = the OFun of b2 . (Q,s) ) ) & the InitS of tfsm in the InitS of b2 holds
b1 = b2
end;
:: deftheorem Def18 defines the_reduction_of FSM_1:def 18 :
for IAlph, OAlph being non empty set
for tfsm being non empty finite Mealy-FSM of IAlph,OAlph
for b4 being strict Mealy-FSM of IAlph,OAlph holds
( b4 = the_reduction_of tfsm iff ( the carrier of b4 = final_states_partition tfsm & ( for Q being State of b4
for s being Element of IAlph
for q being State of tfsm st q in Q holds
( the Tran of tfsm . (q,s) in the Tran of b4 . (Q,s) & the OFun of tfsm . (q,s) = the OFun of b4 . (Q,s) ) ) & the InitS of tfsm in the InitS of b4 ) );
theorem Th57:
theorem Th58:
begin
definition
let IAlph,
OAlph be non
empty set ;
let tfsm1,
tfsm2 be non
empty Mealy-FSM of
IAlph,
OAlph;
pred tfsm1,
tfsm2 -are_isomorphic means :
Def19:
ex
Tf being
Function of the
carrier of
tfsm1, the
carrier of
tfsm2 st
(
Tf is
bijective &
Tf . the
InitS of
tfsm1 = the
InitS of
tfsm2 & ( for
q11 being
State of
tfsm1 for
s being
Element of
IAlph holds
(
Tf . ( the Tran of tfsm1 . (q11,s)) = the
Tran of
tfsm2 . (
(Tf . q11),
s) & the
OFun of
tfsm1 . (
q11,
s)
= the
OFun of
tfsm2 . (
(Tf . q11),
s) ) ) );
reflexivity
for tfsm1 being non empty Mealy-FSM of IAlph,OAlph ex Tf being Function of the carrier of tfsm1, the carrier of tfsm1 st
( Tf is bijective & Tf . the InitS of tfsm1 = the InitS of tfsm1 & ( for q11 being State of tfsm1
for s being Element of IAlph holds
( Tf . ( the Tran of tfsm1 . (q11,s)) = the Tran of tfsm1 . ((Tf . q11),s) & the OFun of tfsm1 . (q11,s) = the OFun of tfsm1 . ((Tf . q11),s) ) ) )
symmetry
for tfsm1, tfsm2 being non empty Mealy-FSM of IAlph,OAlph st ex Tf being Function of the carrier of tfsm1, the carrier of tfsm2 st
( Tf is bijective & Tf . the InitS of tfsm1 = the InitS of tfsm2 & ( for q11 being State of tfsm1
for s being Element of IAlph holds
( Tf . ( the Tran of tfsm1 . (q11,s)) = the Tran of tfsm2 . ((Tf . q11),s) & the OFun of tfsm1 . (q11,s) = the OFun of tfsm2 . ((Tf . q11),s) ) ) ) holds
ex Tf being Function of the carrier of tfsm2, the carrier of tfsm1 st
( Tf is bijective & Tf . the InitS of tfsm2 = the InitS of tfsm1 & ( for q11 being State of tfsm2
for s being Element of IAlph holds
( Tf . ( the Tran of tfsm2 . (q11,s)) = the Tran of tfsm1 . ((Tf . q11),s) & the OFun of tfsm2 . (q11,s) = the OFun of tfsm1 . ((Tf . q11),s) ) ) )
end;
:: deftheorem Def19 defines -are_isomorphic FSM_1:def 19 :
for IAlph, OAlph being non empty set
for tfsm1, tfsm2 being non empty Mealy-FSM of IAlph,OAlph holds
( tfsm1,tfsm2 -are_isomorphic iff ex Tf being Function of the carrier of tfsm1, the carrier of tfsm2 st
( Tf is bijective & Tf . the InitS of tfsm1 = the InitS of tfsm2 & ( for q11 being State of tfsm1
for s being Element of IAlph holds
( Tf . ( the Tran of tfsm1 . (q11,s)) = the Tran of tfsm2 . ((Tf . q11),s) & the OFun of tfsm1 . (q11,s) = the OFun of tfsm2 . ((Tf . q11),s) ) ) ) );
theorem Th59:
theorem Th60:
theorem Th61:
for
OAlph,
IAlph being non
empty set for
tfsm1,
tfsm2 being non
empty Mealy-FSM of
IAlph,
OAlph for
q11,
q12 being
State of
tfsm1 for
Tf being
Function of the
carrier of
tfsm1, the
carrier of
tfsm2 st ( for
q being
State of
tfsm1 for
s being
Element of
IAlph holds
(
Tf . ( the Tran of tfsm1 . (q,s)) = the
Tran of
tfsm2 . (
(Tf . q),
s) & the
OFun of
tfsm1 . (
q,
s)
= the
OFun of
tfsm2 . (
(Tf . q),
s) ) ) holds
(
q11,
q12 -are_equivalent iff
Tf . q11,
Tf . q12 -are_equivalent )
theorem Th62:
begin
:: deftheorem Def20 defines reduced FSM_1:def 20 :
for IAlph, OAlph being non empty set
for IT being non empty Mealy-FSM of IAlph,OAlph holds
( IT is reduced iff for qa, qb being State of IT st qa <> qb holds
not qa,qb -are_equivalent );
theorem
canceled;
theorem Th64:
theorem Th65:
:: deftheorem Def21 defines accessible FSM_1:def 21 :
for IAlph, OAlph being non empty set
for tfsm being non empty Mealy-FSM of IAlph,OAlph
for IT being State of tfsm holds
( IT is accessible iff ex w being FinSequence of IAlph st the InitS of tfsm,w -leads_to IT );
:: deftheorem Def22 defines connected FSM_1:def 22 :
for IAlph, OAlph being non empty set
for IT being non empty Mealy-FSM of IAlph,OAlph holds
( IT is connected iff for q being State of IT holds q is accessible );
theorem
canceled;
:: deftheorem defines accessibleStates FSM_1:def 23 :
for IAlph, OAlph being non empty set
for tfsm being non empty Mealy-FSM of IAlph,OAlph holds accessibleStates tfsm = { q where q is State of tfsm : q is accessible } ;
theorem Th67:
theorem Th68:
theorem
for
IAlph,
OAlph being non
empty set for
tfsm being non
empty finite Mealy-FSM of
IAlph,
OAlph for
cTran being
Function of
[:(accessibleStates tfsm),IAlph:],
(accessibleStates tfsm) for
cOFun being
Function of
[:(accessibleStates tfsm),IAlph:],
OAlph for
cInitS being
Element of
accessibleStates tfsm st
cTran = the
Tran of
tfsm | [:(accessibleStates tfsm),IAlph:] &
cOFun = the
OFun of
tfsm | [:(accessibleStates tfsm),IAlph:] &
cInitS = the
InitS of
tfsm holds
tfsm,
Mealy-FSM(#
(accessibleStates tfsm),
cTran,
cOFun,
cInitS #)
-are_equivalent
theorem
begin
:: deftheorem Def24 defines -Mealy_union FSM_1:def 24 :
for IAlph being set
for OAlph being non empty set
for tfsm1, tfsm2 being non empty Mealy-FSM of IAlph,OAlph
for b5 being strict Mealy-FSM of IAlph,OAlph holds
( b5 = tfsm1 -Mealy_union tfsm2 iff ( the carrier of b5 = the carrier of tfsm1 \/ the carrier of tfsm2 & the Tran of b5 = the Tran of tfsm1 +* the Tran of tfsm2 & the OFun of b5 = the OFun of tfsm1 +* the OFun of tfsm2 & the InitS of b5 = the InitS of tfsm1 ) );
theorem Th71:
theorem Th72:
theorem Th73:
theorem Th74:
theorem Th75:
theorem Th76:
theorem Th77:
theorem Th78:
theorem Th79:
theorem Th80:
begin
theorem Th81:
theorem Th82:
theorem Th83:
theorem Th84:
theorem