i... dont know what to commit
This commit is contained in:
29
TwoPhase.toolbox/.project
Normal file
29
TwoPhase.toolbox/.project
Normal file
@@ -0,0 +1,29 @@
|
||||
<?xml version="1.0" encoding="UTF-8"?>
|
||||
<projectDescription>
|
||||
<name>TwoPhase</name>
|
||||
<comment></comment>
|
||||
<projects>
|
||||
</projects>
|
||||
<buildSpec>
|
||||
<buildCommand>
|
||||
<name>toolbox.builder.TLAParserBuilder</name>
|
||||
<arguments>
|
||||
</arguments>
|
||||
</buildCommand>
|
||||
</buildSpec>
|
||||
<natures>
|
||||
<nature>toolbox.natures.TLANature</nature>
|
||||
</natures>
|
||||
<linkedResources>
|
||||
<link>
|
||||
<name>TCommit.tla</name>
|
||||
<type>1</type>
|
||||
<location>/Users/veitheller/Documents/Code/Github/simple/TCommit.tla</location>
|
||||
</link>
|
||||
<link>
|
||||
<name>TwoPhase.tla</name>
|
||||
<type>1</type>
|
||||
<locationURI>PARENT-1-PROJECT_LOC/TwoPhase.tla</locationURI>
|
||||
</link>
|
||||
</linkedResources>
|
||||
</projectDescription>
|
2
TwoPhase.toolbox/.settings/org.lamport.tla.toolbox.prefs
Normal file
2
TwoPhase.toolbox/.settings/org.lamport.tla.toolbox.prefs
Normal file
@@ -0,0 +1,2 @@
|
||||
ProjectRootFile=PARENT-1-PROJECT_LOC/TwoPhase.tla
|
||||
eclipse.preferences.version=1
|
21
TwoPhase.toolbox/Model_1/MC.cfg
Normal file
21
TwoPhase.toolbox/Model_1/MC.cfg
Normal file
@@ -0,0 +1,21 @@
|
||||
\* MV CONSTANT declarations
|
||||
CONSTANTS
|
||||
r1 = r1
|
||||
r2 = r2
|
||||
r3 = r3
|
||||
\* MV CONSTANT definitions
|
||||
CONSTANT
|
||||
RM <- const_154592300704034000
|
||||
\* SYMMETRY definition
|
||||
SYMMETRY symm_154592300704035000
|
||||
\* INIT definition
|
||||
INIT
|
||||
init_154592300704036000
|
||||
\* NEXT definition
|
||||
NEXT
|
||||
next_154592300704037000
|
||||
\* INVARIANT definition
|
||||
INVARIANT
|
||||
inv_154592300704038000
|
||||
inv_154592300704039000
|
||||
\* Generated on Thu Dec 27 16:03:27 CET 2018
|
37
TwoPhase.toolbox/Model_1/MC.tla
Normal file
37
TwoPhase.toolbox/Model_1/MC.tla
Normal file
@@ -0,0 +1,37 @@
|
||||
---- MODULE MC ----
|
||||
EXTENDS TwoPhase, TLC
|
||||
|
||||
\* MV CONSTANT declarations@modelParameterConstants
|
||||
CONSTANTS
|
||||
r1, r2, r3
|
||||
----
|
||||
|
||||
\* MV CONSTANT definitions RM
|
||||
const_154592300704034000 ==
|
||||
{r1, r2, r3}
|
||||
----
|
||||
|
||||
\* SYMMETRY definition
|
||||
symm_154592300704035000 ==
|
||||
Permutations(const_154592300704034000)
|
||||
----
|
||||
|
||||
\* INIT definition @modelBehaviorInit:0
|
||||
init_154592300704036000 ==
|
||||
TPInit
|
||||
----
|
||||
\* NEXT definition @modelBehaviorNext:0
|
||||
next_154592300704037000 ==
|
||||
TPNext
|
||||
----
|
||||
\* INVARIANT definition @modelCorrectnessInvariants:0
|
||||
inv_154592300704038000 ==
|
||||
TPTypeOK
|
||||
----
|
||||
\* INVARIANT definition @modelCorrectnessInvariants:1
|
||||
inv_154592300704039000 ==
|
||||
TCConsistent
|
||||
----
|
||||
=============================================================================
|
||||
\* Modification History
|
||||
\* Created Thu Dec 27 16:03:27 CET 2018 by veitheller
|
82
TwoPhase.toolbox/Model_1/TCommit.tla
Normal file
82
TwoPhase.toolbox/Model_1/TCommit.tla
Normal file
@@ -0,0 +1,82 @@
|
||||
------------------------------ MODULE TCommit ------------------------------
|
||||
|
||||
(***************************************************************************)
|
||||
(* This specification is explained in "Transaction Commit", Lecture 5 of *)
|
||||
(* the TLA+ Video Course. *)
|
||||
(***************************************************************************)
|
||||
CONSTANT RM \* The set of participating resource managers
|
||||
|
||||
VARIABLE rmState \* rmState[rm] is the state of resource manager r.
|
||||
-----------------------------------------------------------------------------
|
||||
TCTypeOK ==
|
||||
(*************************************************************************)
|
||||
(* The type-correctness invariant *)
|
||||
(*************************************************************************)
|
||||
rmState \in [RM -> {"working", "prepared", "committed", "aborted"}]
|
||||
|
||||
TCInit == rmState = [r \in RM |-> "working"]
|
||||
(*************************************************************************)
|
||||
(* The initial predicate. *)
|
||||
(*************************************************************************)
|
||||
|
||||
canCommit == \A r \in RM : rmState[r] \in {"prepared", "committed"}
|
||||
(*************************************************************************)
|
||||
(* True iff all RMs are in the "prepared" or "committed" state. *)
|
||||
(*************************************************************************)
|
||||
|
||||
notCommitted == \A r \in RM : rmState[r] # "committed"
|
||||
(*************************************************************************)
|
||||
(* True iff no resource manager has decided to commit. *)
|
||||
(*************************************************************************)
|
||||
-----------------------------------------------------------------------------
|
||||
(***************************************************************************)
|
||||
(* We now define the actions that may be performed by the RMs, and then *)
|
||||
(* define the complete next-state action of the specification to be the *)
|
||||
(* disjunction of the possible RM actions. *)
|
||||
(***************************************************************************)
|
||||
Prepare(r) == /\ rmState[r] = "working"
|
||||
/\ rmState' = [rmState EXCEPT ![r] = "prepared"]
|
||||
|
||||
Decide(r) == \/ /\ rmState[r] = "prepared"
|
||||
/\ canCommit
|
||||
/\ rmState' = [rmState EXCEPT ![r] = "committed"]
|
||||
\/ /\ rmState[r] \in {"working", "prepared"}
|
||||
/\ notCommitted
|
||||
/\ rmState' = [rmState EXCEPT ![r] = "aborted"]
|
||||
|
||||
TCNext == \E r \in RM : Prepare(r) \/ Decide(r)
|
||||
(*************************************************************************)
|
||||
(* The next-state action. *)
|
||||
(*************************************************************************)
|
||||
-----------------------------------------------------------------------------
|
||||
TCConsistent ==
|
||||
(*************************************************************************)
|
||||
(* A state predicate asserting that two RMs have not arrived at *)
|
||||
(* conflicting decisions. It is an invariant of the specification. *)
|
||||
(*************************************************************************)
|
||||
\A r1, r2 \in RM : ~ /\ rmState[r1] = "aborted"
|
||||
/\ rmState[r2] = "committed"
|
||||
-----------------------------------------------------------------------------
|
||||
(***************************************************************************)
|
||||
(* The following part of the spec is not discussed in Video Lecture 5. It *)
|
||||
(* will be explained in Video Lecture 8. *)
|
||||
(***************************************************************************)
|
||||
TCSpec == TCInit /\ [][TCNext]_rmState
|
||||
(*************************************************************************)
|
||||
(* The complete specification of the protocol written as a temporal *)
|
||||
(* formula. *)
|
||||
(*************************************************************************)
|
||||
|
||||
THEOREM TCSpec => [](TCTypeOK /\ TCConsistent)
|
||||
(*************************************************************************)
|
||||
(* This theorem asserts the truth of the temporal formula whose meaning *)
|
||||
(* is that the state predicate TCTypeOK /\ TCInvariant is an invariant *)
|
||||
(* of the specification TCSpec. Invariance of this conjunction is *)
|
||||
(* equivalent to invariance of both of the formulas TCTypeOK and *)
|
||||
(* TCConsistent. *)
|
||||
(*************************************************************************)
|
||||
|
||||
=============================================================================
|
||||
\* Modification History
|
||||
\* Last modified Fri Dec 21 17:16:06 CET 2018 by veitheller
|
||||
\* Created Fri Dec 21 17:15:44 CET 2018 by veitheller
|
182
TwoPhase.toolbox/Model_1/TwoPhase.tla
Normal file
182
TwoPhase.toolbox/Model_1/TwoPhase.tla
Normal file
@@ -0,0 +1,182 @@
|
||||
------------------------------ MODULE TwoPhase ------------------------------
|
||||
|
||||
(***************************************************************************)
|
||||
(* This specification is discussed in "Two-Phase Commit", Lecture 6 of the *)
|
||||
(* TLA+ Video Course. It describes the Two-Phase Commit protocol, in *)
|
||||
(* which a transaction manager (TM) coordinates the resource managers *)
|
||||
(* (RMs) to implement the Transaction Commit specification of module *)
|
||||
(* TCommit. In this specification, RMs spontaneously issue Prepared *)
|
||||
(* messages. We ignore the Prepare messages that the TM can send to the *)
|
||||
(* RMs. *)
|
||||
(* *)
|
||||
(* For simplicity, we also eliminate Abort messages sent by an RM when it *)
|
||||
(* decides to abort. Such a message would cause the TM to abort the *)
|
||||
(* transaction, an event represented here by the TM spontaneously deciding *)
|
||||
(* to abort. *)
|
||||
(***************************************************************************)
|
||||
CONSTANT RM \* The set of resource managers
|
||||
|
||||
VARIABLES
|
||||
rmState, \* rmState[r] is the state of resource manager r.
|
||||
tmState, \* The state of the transaction manager.
|
||||
tmPrepared, \* The set of RMs from which the TM has received "Prepared"
|
||||
\* messages.
|
||||
msgs
|
||||
(***********************************************************************)
|
||||
(* In the protocol, processes communicate with one another by sending *)
|
||||
(* messages. For simplicity, we represent message passing with the *)
|
||||
(* variable msgs whose value is the set of all messages that have been *)
|
||||
(* sent. A message is sent by adding it to the set msgs. An action *)
|
||||
(* that, in an implementation, would be enabled by the receipt of a *)
|
||||
(* certain message is here enabled by the presence of that message in *)
|
||||
(* msgs. For simplicity, messages are never removed from msgs. This *)
|
||||
(* allows a single message to be received by multiple receivers. *)
|
||||
(* Receipt of the same message twice is therefore allowed; but in this *)
|
||||
(* particular protocol, that's not a problem. *)
|
||||
(***********************************************************************)
|
||||
|
||||
Messages ==
|
||||
(*************************************************************************)
|
||||
(* The set of all possible messages. Messages of type "Prepared" are *)
|
||||
(* sent from the RM indicated by the message's rm field to the TM. *)
|
||||
(* Messages of type "Commit" and "Abort" are broadcast by the TM, to be *)
|
||||
(* received by all RMs. The set msgs contains just a single copy of *)
|
||||
(* such a message. *)
|
||||
(*************************************************************************)
|
||||
[type : {"Prepared"}, rm : RM] \cup [type : {"Commit", "Abort"}]
|
||||
|
||||
TPTypeOK ==
|
||||
(*************************************************************************)
|
||||
(* The type-correctness invariant *)
|
||||
(*************************************************************************)
|
||||
/\ rmState \in [RM -> {"working", "prepared", "committed", "aborted"}]
|
||||
/\ tmState \in {"init", "done"}
|
||||
/\ tmPrepared \subseteq RM
|
||||
/\ msgs \subseteq Messages
|
||||
|
||||
TPInit ==
|
||||
(*************************************************************************)
|
||||
(* The initial predicate. *)
|
||||
(*************************************************************************)
|
||||
/\ rmState = [r \in RM |-> "working"]
|
||||
/\ tmState = "init"
|
||||
/\ tmPrepared = {}
|
||||
/\ msgs = {}
|
||||
-----------------------------------------------------------------------------
|
||||
(***************************************************************************)
|
||||
(* We now define the actions that may be performed by the processes, first *)
|
||||
(* the TM's actions, then the RMs' actions. *)
|
||||
(***************************************************************************)
|
||||
TMRcvPrepared(r) ==
|
||||
(*************************************************************************)
|
||||
(* The TM receives a "Prepared" message from resource manager r. We *)
|
||||
(* could add the additional enabling condition r \notin tmPrepared, *)
|
||||
(* which disables the action if the TM has already received this *)
|
||||
(* message. But there is no need, because in that case the action has *)
|
||||
(* no effect; it leaves the state unchanged. *)
|
||||
(*************************************************************************)
|
||||
/\ tmState = "init"
|
||||
/\ [type |-> "Prepared", rm |-> r] \in msgs
|
||||
/\ tmPrepared' = tmPrepared \cup {r}
|
||||
/\ UNCHANGED <<rmState, tmState, msgs>>
|
||||
|
||||
TMCommit ==
|
||||
(*************************************************************************)
|
||||
(* The TM commits the transaction; enabled iff the TM is in its initial *)
|
||||
(* state and every RM has sent a "Prepared" message. *)
|
||||
(*************************************************************************)
|
||||
/\ tmState = "init"
|
||||
/\ tmPrepared = RM
|
||||
/\ tmState' = "done"
|
||||
/\ msgs' = msgs \cup {[type |-> "Commit"]}
|
||||
/\ UNCHANGED <<rmState, tmPrepared>>
|
||||
|
||||
TMAbort ==
|
||||
(*************************************************************************)
|
||||
(* The TM spontaneously aborts the transaction. *)
|
||||
(*************************************************************************)
|
||||
/\ tmState = "init"
|
||||
/\ tmState' = "done"
|
||||
/\ msgs' = msgs \cup {[type |-> "Abort"]}
|
||||
/\ UNCHANGED <<rmState, tmPrepared>>
|
||||
|
||||
RMPrepare(r) ==
|
||||
(*************************************************************************)
|
||||
(* Resource manager r prepares. *)
|
||||
(*************************************************************************)
|
||||
/\ rmState[r] = "working"
|
||||
/\ rmState' = [rmState EXCEPT ![r] = "prepared"]
|
||||
/\ msgs' = msgs \cup {[type |-> "Prepared", rm |-> r]}
|
||||
/\ UNCHANGED <<tmState, tmPrepared>>
|
||||
|
||||
RMChooseToAbort(r) ==
|
||||
(*************************************************************************)
|
||||
(* Resource manager r spontaneously decides to abort. As noted above, r *)
|
||||
(* does not send any message in our simplified spec. *)
|
||||
(*************************************************************************)
|
||||
/\ rmState[r] = "working"
|
||||
/\ rmState' = [rmState EXCEPT ![r] = "aborted"]
|
||||
/\ UNCHANGED <<tmState, tmPrepared, msgs>>
|
||||
|
||||
RMRcvCommitMsg(r) ==
|
||||
(*************************************************************************)
|
||||
(* Resource manager r is told by the TM to commit. *)
|
||||
(*************************************************************************)
|
||||
/\ [type |-> "Commit"] \in msgs
|
||||
/\ rmState' = [rmState EXCEPT ![r] = "committed"]
|
||||
/\ UNCHANGED <<tmState, tmPrepared, msgs>>
|
||||
|
||||
RMRcvAbortMsg(r) ==
|
||||
(*************************************************************************)
|
||||
(* Resource manager r is told by the TM to abort. *)
|
||||
(*************************************************************************)
|
||||
/\ [type |-> "Abort"] \in msgs
|
||||
/\ rmState' = [rmState EXCEPT ![r] = "aborted"]
|
||||
/\ UNCHANGED <<tmState, tmPrepared, msgs>>
|
||||
|
||||
TPNext ==
|
||||
\/ TMCommit \/ TMAbort
|
||||
\/ \E r \in RM :
|
||||
TMRcvPrepared(r) \/ RMPrepare(r) \/ RMChooseToAbort(r)
|
||||
\/ RMRcvCommitMsg(r) \/ RMRcvAbortMsg(r)
|
||||
-----------------------------------------------------------------------------
|
||||
(***************************************************************************)
|
||||
(* The material below this point is not discussed in Video Lecture 6. It *)
|
||||
(* will be explained in Video Lecture 8. *)
|
||||
(***************************************************************************)
|
||||
|
||||
TPSpec == TPInit /\ [][TPNext]_<<rmState, tmState, tmPrepared, msgs>>
|
||||
(*************************************************************************)
|
||||
(* The complete spec of the Two-Phase Commit protocol. *)
|
||||
(*************************************************************************)
|
||||
|
||||
THEOREM TPSpec => []TPTypeOK
|
||||
(*************************************************************************)
|
||||
(* This theorem asserts that the type-correctness predicate TPTypeOK is *)
|
||||
(* an invariant of the specification. *)
|
||||
(*************************************************************************)
|
||||
-----------------------------------------------------------------------------
|
||||
(***************************************************************************)
|
||||
(* We now assert that the Two-Phase Commit protocol implements the *)
|
||||
(* Transaction Commit protocol of module TCommit. The following statement *)
|
||||
(* imports all the definitions from module TCommit into the current *)
|
||||
(* module. *)
|
||||
(***************************************************************************)
|
||||
INSTANCE TCommit
|
||||
|
||||
THEOREM TPSpec => TCSpec
|
||||
(*************************************************************************)
|
||||
(* This theorem asserts that the specification TPSpec of the Two-Phase *)
|
||||
(* Commit protocol implements the specification TCSpec of the *)
|
||||
(* Transaction Commit protocol. *)
|
||||
(*************************************************************************)
|
||||
(***************************************************************************)
|
||||
(* The two theorems in this module have been checked with TLC for six *)
|
||||
(* RMs, a configuration with 50816 reachable states, in a little over a *)
|
||||
(* minute on a 1 GHz PC. *)
|
||||
(***************************************************************************)
|
||||
|
||||
=============================================================================
|
||||
\* Modification History
|
||||
\* Last modified Thu Dec 27 15:46:42 CET 2018 by veitheller
|
||||
\* Created Thu Dec 27 15:46:21 CET 2018 by veitheller
|
13
TwoPhase.toolbox/Model_1_SnapShot_1545922765832/MC.cfg
Normal file
13
TwoPhase.toolbox/Model_1_SnapShot_1545922765832/MC.cfg
Normal file
@@ -0,0 +1,13 @@
|
||||
\* CONSTANT definitions
|
||||
CONSTANT
|
||||
RM <- const_15459227497686000
|
||||
\* INIT definition
|
||||
INIT
|
||||
init_15459227497687000
|
||||
\* NEXT definition
|
||||
NEXT
|
||||
next_15459227497688000
|
||||
\* INVARIANT definition
|
||||
INVARIANT
|
||||
inv_15459227497689000
|
||||
\* Generated on Thu Dec 27 15:59:09 CET 2018
|
23
TwoPhase.toolbox/Model_1_SnapShot_1545922765832/MC.tla
Normal file
23
TwoPhase.toolbox/Model_1_SnapShot_1545922765832/MC.tla
Normal file
@@ -0,0 +1,23 @@
|
||||
---- MODULE MC ----
|
||||
EXTENDS TwoPhase, TLC
|
||||
|
||||
\* CONSTANT definitions @modelParameterConstants:0RM
|
||||
const_15459227497686000 ==
|
||||
{"r1", "r2", "r3"}
|
||||
----
|
||||
|
||||
\* INIT definition @modelBehaviorInit:0
|
||||
init_15459227497687000 ==
|
||||
TPInit
|
||||
----
|
||||
\* NEXT definition @modelBehaviorNext:0
|
||||
next_15459227497688000 ==
|
||||
TPNext
|
||||
----
|
||||
\* INVARIANT definition @modelCorrectnessInvariants:0
|
||||
inv_15459227497689000 ==
|
||||
TPTypeOK
|
||||
----
|
||||
=============================================================================
|
||||
\* Modification History
|
||||
\* Created Thu Dec 27 15:59:09 CET 2018 by veitheller
|
82
TwoPhase.toolbox/Model_1_SnapShot_1545922765832/TCommit.tla
Normal file
82
TwoPhase.toolbox/Model_1_SnapShot_1545922765832/TCommit.tla
Normal file
@@ -0,0 +1,82 @@
|
||||
------------------------------ MODULE TCommit ------------------------------
|
||||
|
||||
(***************************************************************************)
|
||||
(* This specification is explained in "Transaction Commit", Lecture 5 of *)
|
||||
(* the TLA+ Video Course. *)
|
||||
(***************************************************************************)
|
||||
CONSTANT RM \* The set of participating resource managers
|
||||
|
||||
VARIABLE rmState \* rmState[rm] is the state of resource manager r.
|
||||
-----------------------------------------------------------------------------
|
||||
TCTypeOK ==
|
||||
(*************************************************************************)
|
||||
(* The type-correctness invariant *)
|
||||
(*************************************************************************)
|
||||
rmState \in [RM -> {"working", "prepared", "committed", "aborted"}]
|
||||
|
||||
TCInit == rmState = [r \in RM |-> "working"]
|
||||
(*************************************************************************)
|
||||
(* The initial predicate. *)
|
||||
(*************************************************************************)
|
||||
|
||||
canCommit == \A r \in RM : rmState[r] \in {"prepared", "committed"}
|
||||
(*************************************************************************)
|
||||
(* True iff all RMs are in the "prepared" or "committed" state. *)
|
||||
(*************************************************************************)
|
||||
|
||||
notCommitted == \A r \in RM : rmState[r] # "committed"
|
||||
(*************************************************************************)
|
||||
(* True iff no resource manager has decided to commit. *)
|
||||
(*************************************************************************)
|
||||
-----------------------------------------------------------------------------
|
||||
(***************************************************************************)
|
||||
(* We now define the actions that may be performed by the RMs, and then *)
|
||||
(* define the complete next-state action of the specification to be the *)
|
||||
(* disjunction of the possible RM actions. *)
|
||||
(***************************************************************************)
|
||||
Prepare(r) == /\ rmState[r] = "working"
|
||||
/\ rmState' = [rmState EXCEPT ![r] = "prepared"]
|
||||
|
||||
Decide(r) == \/ /\ rmState[r] = "prepared"
|
||||
/\ canCommit
|
||||
/\ rmState' = [rmState EXCEPT ![r] = "committed"]
|
||||
\/ /\ rmState[r] \in {"working", "prepared"}
|
||||
/\ notCommitted
|
||||
/\ rmState' = [rmState EXCEPT ![r] = "aborted"]
|
||||
|
||||
TCNext == \E r \in RM : Prepare(r) \/ Decide(r)
|
||||
(*************************************************************************)
|
||||
(* The next-state action. *)
|
||||
(*************************************************************************)
|
||||
-----------------------------------------------------------------------------
|
||||
TCConsistent ==
|
||||
(*************************************************************************)
|
||||
(* A state predicate asserting that two RMs have not arrived at *)
|
||||
(* conflicting decisions. It is an invariant of the specification. *)
|
||||
(*************************************************************************)
|
||||
\A r1, r2 \in RM : ~ /\ rmState[r1] = "aborted"
|
||||
/\ rmState[r2] = "committed"
|
||||
-----------------------------------------------------------------------------
|
||||
(***************************************************************************)
|
||||
(* The following part of the spec is not discussed in Video Lecture 5. It *)
|
||||
(* will be explained in Video Lecture 8. *)
|
||||
(***************************************************************************)
|
||||
TCSpec == TCInit /\ [][TCNext]_rmState
|
||||
(*************************************************************************)
|
||||
(* The complete specification of the protocol written as a temporal *)
|
||||
(* formula. *)
|
||||
(*************************************************************************)
|
||||
|
||||
THEOREM TCSpec => [](TCTypeOK /\ TCConsistent)
|
||||
(*************************************************************************)
|
||||
(* This theorem asserts the truth of the temporal formula whose meaning *)
|
||||
(* is that the state predicate TCTypeOK /\ TCInvariant is an invariant *)
|
||||
(* of the specification TCSpec. Invariance of this conjunction is *)
|
||||
(* equivalent to invariance of both of the formulas TCTypeOK and *)
|
||||
(* TCConsistent. *)
|
||||
(*************************************************************************)
|
||||
|
||||
=============================================================================
|
||||
\* Modification History
|
||||
\* Last modified Fri Dec 21 17:16:06 CET 2018 by veitheller
|
||||
\* Created Fri Dec 21 17:15:44 CET 2018 by veitheller
|
182
TwoPhase.toolbox/Model_1_SnapShot_1545922765832/TwoPhase.tla
Normal file
182
TwoPhase.toolbox/Model_1_SnapShot_1545922765832/TwoPhase.tla
Normal file
@@ -0,0 +1,182 @@
|
||||
------------------------------ MODULE TwoPhase ------------------------------
|
||||
|
||||
(***************************************************************************)
|
||||
(* This specification is discussed in "Two-Phase Commit", Lecture 6 of the *)
|
||||
(* TLA+ Video Course. It describes the Two-Phase Commit protocol, in *)
|
||||
(* which a transaction manager (TM) coordinates the resource managers *)
|
||||
(* (RMs) to implement the Transaction Commit specification of module *)
|
||||
(* TCommit. In this specification, RMs spontaneously issue Prepared *)
|
||||
(* messages. We ignore the Prepare messages that the TM can send to the *)
|
||||
(* RMs. *)
|
||||
(* *)
|
||||
(* For simplicity, we also eliminate Abort messages sent by an RM when it *)
|
||||
(* decides to abort. Such a message would cause the TM to abort the *)
|
||||
(* transaction, an event represented here by the TM spontaneously deciding *)
|
||||
(* to abort. *)
|
||||
(***************************************************************************)
|
||||
CONSTANT RM \* The set of resource managers
|
||||
|
||||
VARIABLES
|
||||
rmState, \* rmState[r] is the state of resource manager r.
|
||||
tmState, \* The state of the transaction manager.
|
||||
tmPrepared, \* The set of RMs from which the TM has received "Prepared"
|
||||
\* messages.
|
||||
msgs
|
||||
(***********************************************************************)
|
||||
(* In the protocol, processes communicate with one another by sending *)
|
||||
(* messages. For simplicity, we represent message passing with the *)
|
||||
(* variable msgs whose value is the set of all messages that have been *)
|
||||
(* sent. A message is sent by adding it to the set msgs. An action *)
|
||||
(* that, in an implementation, would be enabled by the receipt of a *)
|
||||
(* certain message is here enabled by the presence of that message in *)
|
||||
(* msgs. For simplicity, messages are never removed from msgs. This *)
|
||||
(* allows a single message to be received by multiple receivers. *)
|
||||
(* Receipt of the same message twice is therefore allowed; but in this *)
|
||||
(* particular protocol, that's not a problem. *)
|
||||
(***********************************************************************)
|
||||
|
||||
Messages ==
|
||||
(*************************************************************************)
|
||||
(* The set of all possible messages. Messages of type "Prepared" are *)
|
||||
(* sent from the RM indicated by the message's rm field to the TM. *)
|
||||
(* Messages of type "Commit" and "Abort" are broadcast by the TM, to be *)
|
||||
(* received by all RMs. The set msgs contains just a single copy of *)
|
||||
(* such a message. *)
|
||||
(*************************************************************************)
|
||||
[type : {"Prepared"}, rm : RM] \cup [type : {"Commit", "Abort"}]
|
||||
|
||||
TPTypeOK ==
|
||||
(*************************************************************************)
|
||||
(* The type-correctness invariant *)
|
||||
(*************************************************************************)
|
||||
/\ rmState \in [RM -> {"working", "prepared", "committed", "aborted"}]
|
||||
/\ tmState \in {"init", "done"}
|
||||
/\ tmPrepared \subseteq RM
|
||||
/\ msgs \subseteq Messages
|
||||
|
||||
TPInit ==
|
||||
(*************************************************************************)
|
||||
(* The initial predicate. *)
|
||||
(*************************************************************************)
|
||||
/\ rmState = [r \in RM |-> "working"]
|
||||
/\ tmState = "init"
|
||||
/\ tmPrepared = {}
|
||||
/\ msgs = {}
|
||||
-----------------------------------------------------------------------------
|
||||
(***************************************************************************)
|
||||
(* We now define the actions that may be performed by the processes, first *)
|
||||
(* the TM's actions, then the RMs' actions. *)
|
||||
(***************************************************************************)
|
||||
TMRcvPrepared(r) ==
|
||||
(*************************************************************************)
|
||||
(* The TM receives a "Prepared" message from resource manager r. We *)
|
||||
(* could add the additional enabling condition r \notin tmPrepared, *)
|
||||
(* which disables the action if the TM has already received this *)
|
||||
(* message. But there is no need, because in that case the action has *)
|
||||
(* no effect; it leaves the state unchanged. *)
|
||||
(*************************************************************************)
|
||||
/\ tmState = "init"
|
||||
/\ [type |-> "Prepared", rm |-> r] \in msgs
|
||||
/\ tmPrepared' = tmPrepared \cup {r}
|
||||
/\ UNCHANGED <<rmState, tmState, msgs>>
|
||||
|
||||
TMCommit ==
|
||||
(*************************************************************************)
|
||||
(* The TM commits the transaction; enabled iff the TM is in its initial *)
|
||||
(* state and every RM has sent a "Prepared" message. *)
|
||||
(*************************************************************************)
|
||||
/\ tmState = "init"
|
||||
/\ tmPrepared = RM
|
||||
/\ tmState' = "done"
|
||||
/\ msgs' = msgs \cup {[type |-> "Commit"]}
|
||||
/\ UNCHANGED <<rmState, tmPrepared>>
|
||||
|
||||
TMAbort ==
|
||||
(*************************************************************************)
|
||||
(* The TM spontaneously aborts the transaction. *)
|
||||
(*************************************************************************)
|
||||
/\ tmState = "init"
|
||||
/\ tmState' = "done"
|
||||
/\ msgs' = msgs \cup {[type |-> "Abort"]}
|
||||
/\ UNCHANGED <<rmState, tmPrepared>>
|
||||
|
||||
RMPrepare(r) ==
|
||||
(*************************************************************************)
|
||||
(* Resource manager r prepares. *)
|
||||
(*************************************************************************)
|
||||
/\ rmState[r] = "working"
|
||||
/\ rmState' = [rmState EXCEPT ![r] = "prepared"]
|
||||
/\ msgs' = msgs \cup {[type |-> "Prepared", rm |-> r]}
|
||||
/\ UNCHANGED <<tmState, tmPrepared>>
|
||||
|
||||
RMChooseToAbort(r) ==
|
||||
(*************************************************************************)
|
||||
(* Resource manager r spontaneously decides to abort. As noted above, r *)
|
||||
(* does not send any message in our simplified spec. *)
|
||||
(*************************************************************************)
|
||||
/\ rmState[r] = "working"
|
||||
/\ rmState' = [rmState EXCEPT ![r] = "aborted"]
|
||||
/\ UNCHANGED <<tmState, tmPrepared, msgs>>
|
||||
|
||||
RMRcvCommitMsg(r) ==
|
||||
(*************************************************************************)
|
||||
(* Resource manager r is told by the TM to commit. *)
|
||||
(*************************************************************************)
|
||||
/\ [type |-> "Commit"] \in msgs
|
||||
/\ rmState' = [rmState EXCEPT ![r] = "committed"]
|
||||
/\ UNCHANGED <<tmState, tmPrepared, msgs>>
|
||||
|
||||
RMRcvAbortMsg(r) ==
|
||||
(*************************************************************************)
|
||||
(* Resource manager r is told by the TM to abort. *)
|
||||
(*************************************************************************)
|
||||
/\ [type |-> "Abort"] \in msgs
|
||||
/\ rmState' = [rmState EXCEPT ![r] = "aborted"]
|
||||
/\ UNCHANGED <<tmState, tmPrepared, msgs>>
|
||||
|
||||
TPNext ==
|
||||
\/ TMCommit \/ TMAbort
|
||||
\/ \E r \in RM :
|
||||
TMRcvPrepared(r) \/ RMPrepare(r) \/ RMChooseToAbort(r)
|
||||
\/ RMRcvCommitMsg(r) \/ RMRcvAbortMsg(r)
|
||||
-----------------------------------------------------------------------------
|
||||
(***************************************************************************)
|
||||
(* The material below this point is not discussed in Video Lecture 6. It *)
|
||||
(* will be explained in Video Lecture 8. *)
|
||||
(***************************************************************************)
|
||||
|
||||
TPSpec == TPInit /\ [][TPNext]_<<rmState, tmState, tmPrepared, msgs>>
|
||||
(*************************************************************************)
|
||||
(* The complete spec of the Two-Phase Commit protocol. *)
|
||||
(*************************************************************************)
|
||||
|
||||
THEOREM TPSpec => []TPTypeOK
|
||||
(*************************************************************************)
|
||||
(* This theorem asserts that the type-correctness predicate TPTypeOK is *)
|
||||
(* an invariant of the specification. *)
|
||||
(*************************************************************************)
|
||||
-----------------------------------------------------------------------------
|
||||
(***************************************************************************)
|
||||
(* We now assert that the Two-Phase Commit protocol implements the *)
|
||||
(* Transaction Commit protocol of module TCommit. The following statement *)
|
||||
(* imports all the definitions from module TCommit into the current *)
|
||||
(* module. *)
|
||||
(***************************************************************************)
|
||||
INSTANCE TCommit
|
||||
|
||||
THEOREM TPSpec => TCSpec
|
||||
(*************************************************************************)
|
||||
(* This theorem asserts that the specification TPSpec of the Two-Phase *)
|
||||
(* Commit protocol implements the specification TCSpec of the *)
|
||||
(* Transaction Commit protocol. *)
|
||||
(*************************************************************************)
|
||||
(***************************************************************************)
|
||||
(* The two theorems in this module have been checked with TLC for six *)
|
||||
(* RMs, a configuration with 50816 reachable states, in a little over a *)
|
||||
(* minute on a 1 GHz PC. *)
|
||||
(***************************************************************************)
|
||||
|
||||
=============================================================================
|
||||
\* Modification History
|
||||
\* Last modified Thu Dec 27 15:46:42 CET 2018 by veitheller
|
||||
\* Created Thu Dec 27 15:46:21 CET 2018 by veitheller
|
13
TwoPhase.toolbox/Model_1_SnapShot_1545922881813/MC.cfg
Normal file
13
TwoPhase.toolbox/Model_1_SnapShot_1545922881813/MC.cfg
Normal file
@@ -0,0 +1,13 @@
|
||||
\* CONSTANT definitions
|
||||
CONSTANT
|
||||
RM <- const_154592286770714000
|
||||
\* INIT definition
|
||||
INIT
|
||||
init_154592286770715000
|
||||
\* NEXT definition
|
||||
NEXT
|
||||
next_154592286770716000
|
||||
\* INVARIANT definition
|
||||
INVARIANT
|
||||
inv_154592286770717000
|
||||
\* Generated on Thu Dec 27 16:01:07 CET 2018
|
23
TwoPhase.toolbox/Model_1_SnapShot_1545922881813/MC.tla
Normal file
23
TwoPhase.toolbox/Model_1_SnapShot_1545922881813/MC.tla
Normal file
@@ -0,0 +1,23 @@
|
||||
---- MODULE MC ----
|
||||
EXTENDS TwoPhase, TLC
|
||||
|
||||
\* CONSTANT definitions @modelParameterConstants:0RM
|
||||
const_154592286770714000 ==
|
||||
{"r1", "r2", "r3", "r4", "r5"}
|
||||
----
|
||||
|
||||
\* INIT definition @modelBehaviorInit:0
|
||||
init_154592286770715000 ==
|
||||
TPInit
|
||||
----
|
||||
\* NEXT definition @modelBehaviorNext:0
|
||||
next_154592286770716000 ==
|
||||
TPNext
|
||||
----
|
||||
\* INVARIANT definition @modelCorrectnessInvariants:0
|
||||
inv_154592286770717000 ==
|
||||
TPTypeOK
|
||||
----
|
||||
=============================================================================
|
||||
\* Modification History
|
||||
\* Created Thu Dec 27 16:01:07 CET 2018 by veitheller
|
82
TwoPhase.toolbox/Model_1_SnapShot_1545922881813/TCommit.tla
Normal file
82
TwoPhase.toolbox/Model_1_SnapShot_1545922881813/TCommit.tla
Normal file
@@ -0,0 +1,82 @@
|
||||
------------------------------ MODULE TCommit ------------------------------
|
||||
|
||||
(***************************************************************************)
|
||||
(* This specification is explained in "Transaction Commit", Lecture 5 of *)
|
||||
(* the TLA+ Video Course. *)
|
||||
(***************************************************************************)
|
||||
CONSTANT RM \* The set of participating resource managers
|
||||
|
||||
VARIABLE rmState \* rmState[rm] is the state of resource manager r.
|
||||
-----------------------------------------------------------------------------
|
||||
TCTypeOK ==
|
||||
(*************************************************************************)
|
||||
(* The type-correctness invariant *)
|
||||
(*************************************************************************)
|
||||
rmState \in [RM -> {"working", "prepared", "committed", "aborted"}]
|
||||
|
||||
TCInit == rmState = [r \in RM |-> "working"]
|
||||
(*************************************************************************)
|
||||
(* The initial predicate. *)
|
||||
(*************************************************************************)
|
||||
|
||||
canCommit == \A r \in RM : rmState[r] \in {"prepared", "committed"}
|
||||
(*************************************************************************)
|
||||
(* True iff all RMs are in the "prepared" or "committed" state. *)
|
||||
(*************************************************************************)
|
||||
|
||||
notCommitted == \A r \in RM : rmState[r] # "committed"
|
||||
(*************************************************************************)
|
||||
(* True iff no resource manager has decided to commit. *)
|
||||
(*************************************************************************)
|
||||
-----------------------------------------------------------------------------
|
||||
(***************************************************************************)
|
||||
(* We now define the actions that may be performed by the RMs, and then *)
|
||||
(* define the complete next-state action of the specification to be the *)
|
||||
(* disjunction of the possible RM actions. *)
|
||||
(***************************************************************************)
|
||||
Prepare(r) == /\ rmState[r] = "working"
|
||||
/\ rmState' = [rmState EXCEPT ![r] = "prepared"]
|
||||
|
||||
Decide(r) == \/ /\ rmState[r] = "prepared"
|
||||
/\ canCommit
|
||||
/\ rmState' = [rmState EXCEPT ![r] = "committed"]
|
||||
\/ /\ rmState[r] \in {"working", "prepared"}
|
||||
/\ notCommitted
|
||||
/\ rmState' = [rmState EXCEPT ![r] = "aborted"]
|
||||
|
||||
TCNext == \E r \in RM : Prepare(r) \/ Decide(r)
|
||||
(*************************************************************************)
|
||||
(* The next-state action. *)
|
||||
(*************************************************************************)
|
||||
-----------------------------------------------------------------------------
|
||||
TCConsistent ==
|
||||
(*************************************************************************)
|
||||
(* A state predicate asserting that two RMs have not arrived at *)
|
||||
(* conflicting decisions. It is an invariant of the specification. *)
|
||||
(*************************************************************************)
|
||||
\A r1, r2 \in RM : ~ /\ rmState[r1] = "aborted"
|
||||
/\ rmState[r2] = "committed"
|
||||
-----------------------------------------------------------------------------
|
||||
(***************************************************************************)
|
||||
(* The following part of the spec is not discussed in Video Lecture 5. It *)
|
||||
(* will be explained in Video Lecture 8. *)
|
||||
(***************************************************************************)
|
||||
TCSpec == TCInit /\ [][TCNext]_rmState
|
||||
(*************************************************************************)
|
||||
(* The complete specification of the protocol written as a temporal *)
|
||||
(* formula. *)
|
||||
(*************************************************************************)
|
||||
|
||||
THEOREM TCSpec => [](TCTypeOK /\ TCConsistent)
|
||||
(*************************************************************************)
|
||||
(* This theorem asserts the truth of the temporal formula whose meaning *)
|
||||
(* is that the state predicate TCTypeOK /\ TCInvariant is an invariant *)
|
||||
(* of the specification TCSpec. Invariance of this conjunction is *)
|
||||
(* equivalent to invariance of both of the formulas TCTypeOK and *)
|
||||
(* TCConsistent. *)
|
||||
(*************************************************************************)
|
||||
|
||||
=============================================================================
|
||||
\* Modification History
|
||||
\* Last modified Fri Dec 21 17:16:06 CET 2018 by veitheller
|
||||
\* Created Fri Dec 21 17:15:44 CET 2018 by veitheller
|
182
TwoPhase.toolbox/Model_1_SnapShot_1545922881813/TwoPhase.tla
Normal file
182
TwoPhase.toolbox/Model_1_SnapShot_1545922881813/TwoPhase.tla
Normal file
@@ -0,0 +1,182 @@
|
||||
------------------------------ MODULE TwoPhase ------------------------------
|
||||
|
||||
(***************************************************************************)
|
||||
(* This specification is discussed in "Two-Phase Commit", Lecture 6 of the *)
|
||||
(* TLA+ Video Course. It describes the Two-Phase Commit protocol, in *)
|
||||
(* which a transaction manager (TM) coordinates the resource managers *)
|
||||
(* (RMs) to implement the Transaction Commit specification of module *)
|
||||
(* TCommit. In this specification, RMs spontaneously issue Prepared *)
|
||||
(* messages. We ignore the Prepare messages that the TM can send to the *)
|
||||
(* RMs. *)
|
||||
(* *)
|
||||
(* For simplicity, we also eliminate Abort messages sent by an RM when it *)
|
||||
(* decides to abort. Such a message would cause the TM to abort the *)
|
||||
(* transaction, an event represented here by the TM spontaneously deciding *)
|
||||
(* to abort. *)
|
||||
(***************************************************************************)
|
||||
CONSTANT RM \* The set of resource managers
|
||||
|
||||
VARIABLES
|
||||
rmState, \* rmState[r] is the state of resource manager r.
|
||||
tmState, \* The state of the transaction manager.
|
||||
tmPrepared, \* The set of RMs from which the TM has received "Prepared"
|
||||
\* messages.
|
||||
msgs
|
||||
(***********************************************************************)
|
||||
(* In the protocol, processes communicate with one another by sending *)
|
||||
(* messages. For simplicity, we represent message passing with the *)
|
||||
(* variable msgs whose value is the set of all messages that have been *)
|
||||
(* sent. A message is sent by adding it to the set msgs. An action *)
|
||||
(* that, in an implementation, would be enabled by the receipt of a *)
|
||||
(* certain message is here enabled by the presence of that message in *)
|
||||
(* msgs. For simplicity, messages are never removed from msgs. This *)
|
||||
(* allows a single message to be received by multiple receivers. *)
|
||||
(* Receipt of the same message twice is therefore allowed; but in this *)
|
||||
(* particular protocol, that's not a problem. *)
|
||||
(***********************************************************************)
|
||||
|
||||
Messages ==
|
||||
(*************************************************************************)
|
||||
(* The set of all possible messages. Messages of type "Prepared" are *)
|
||||
(* sent from the RM indicated by the message's rm field to the TM. *)
|
||||
(* Messages of type "Commit" and "Abort" are broadcast by the TM, to be *)
|
||||
(* received by all RMs. The set msgs contains just a single copy of *)
|
||||
(* such a message. *)
|
||||
(*************************************************************************)
|
||||
[type : {"Prepared"}, rm : RM] \cup [type : {"Commit", "Abort"}]
|
||||
|
||||
TPTypeOK ==
|
||||
(*************************************************************************)
|
||||
(* The type-correctness invariant *)
|
||||
(*************************************************************************)
|
||||
/\ rmState \in [RM -> {"working", "prepared", "committed", "aborted"}]
|
||||
/\ tmState \in {"init", "done"}
|
||||
/\ tmPrepared \subseteq RM
|
||||
/\ msgs \subseteq Messages
|
||||
|
||||
TPInit ==
|
||||
(*************************************************************************)
|
||||
(* The initial predicate. *)
|
||||
(*************************************************************************)
|
||||
/\ rmState = [r \in RM |-> "working"]
|
||||
/\ tmState = "init"
|
||||
/\ tmPrepared = {}
|
||||
/\ msgs = {}
|
||||
-----------------------------------------------------------------------------
|
||||
(***************************************************************************)
|
||||
(* We now define the actions that may be performed by the processes, first *)
|
||||
(* the TM's actions, then the RMs' actions. *)
|
||||
(***************************************************************************)
|
||||
TMRcvPrepared(r) ==
|
||||
(*************************************************************************)
|
||||
(* The TM receives a "Prepared" message from resource manager r. We *)
|
||||
(* could add the additional enabling condition r \notin tmPrepared, *)
|
||||
(* which disables the action if the TM has already received this *)
|
||||
(* message. But there is no need, because in that case the action has *)
|
||||
(* no effect; it leaves the state unchanged. *)
|
||||
(*************************************************************************)
|
||||
/\ tmState = "init"
|
||||
/\ [type |-> "Prepared", rm |-> r] \in msgs
|
||||
/\ tmPrepared' = tmPrepared \cup {r}
|
||||
/\ UNCHANGED <<rmState, tmState, msgs>>
|
||||
|
||||
TMCommit ==
|
||||
(*************************************************************************)
|
||||
(* The TM commits the transaction; enabled iff the TM is in its initial *)
|
||||
(* state and every RM has sent a "Prepared" message. *)
|
||||
(*************************************************************************)
|
||||
/\ tmState = "init"
|
||||
/\ tmPrepared = RM
|
||||
/\ tmState' = "done"
|
||||
/\ msgs' = msgs \cup {[type |-> "Commit"]}
|
||||
/\ UNCHANGED <<rmState, tmPrepared>>
|
||||
|
||||
TMAbort ==
|
||||
(*************************************************************************)
|
||||
(* The TM spontaneously aborts the transaction. *)
|
||||
(*************************************************************************)
|
||||
/\ tmState = "init"
|
||||
/\ tmState' = "done"
|
||||
/\ msgs' = msgs \cup {[type |-> "Abort"]}
|
||||
/\ UNCHANGED <<rmState, tmPrepared>>
|
||||
|
||||
RMPrepare(r) ==
|
||||
(*************************************************************************)
|
||||
(* Resource manager r prepares. *)
|
||||
(*************************************************************************)
|
||||
/\ rmState[r] = "working"
|
||||
/\ rmState' = [rmState EXCEPT ![r] = "prepared"]
|
||||
/\ msgs' = msgs \cup {[type |-> "Prepared", rm |-> r]}
|
||||
/\ UNCHANGED <<tmState, tmPrepared>>
|
||||
|
||||
RMChooseToAbort(r) ==
|
||||
(*************************************************************************)
|
||||
(* Resource manager r spontaneously decides to abort. As noted above, r *)
|
||||
(* does not send any message in our simplified spec. *)
|
||||
(*************************************************************************)
|
||||
/\ rmState[r] = "working"
|
||||
/\ rmState' = [rmState EXCEPT ![r] = "aborted"]
|
||||
/\ UNCHANGED <<tmState, tmPrepared, msgs>>
|
||||
|
||||
RMRcvCommitMsg(r) ==
|
||||
(*************************************************************************)
|
||||
(* Resource manager r is told by the TM to commit. *)
|
||||
(*************************************************************************)
|
||||
/\ [type |-> "Commit"] \in msgs
|
||||
/\ rmState' = [rmState EXCEPT ![r] = "committed"]
|
||||
/\ UNCHANGED <<tmState, tmPrepared, msgs>>
|
||||
|
||||
RMRcvAbortMsg(r) ==
|
||||
(*************************************************************************)
|
||||
(* Resource manager r is told by the TM to abort. *)
|
||||
(*************************************************************************)
|
||||
/\ [type |-> "Abort"] \in msgs
|
||||
/\ rmState' = [rmState EXCEPT ![r] = "aborted"]
|
||||
/\ UNCHANGED <<tmState, tmPrepared, msgs>>
|
||||
|
||||
TPNext ==
|
||||
\/ TMCommit \/ TMAbort
|
||||
\/ \E r \in RM :
|
||||
TMRcvPrepared(r) \/ RMPrepare(r) \/ RMChooseToAbort(r)
|
||||
\/ RMRcvCommitMsg(r) \/ RMRcvAbortMsg(r)
|
||||
-----------------------------------------------------------------------------
|
||||
(***************************************************************************)
|
||||
(* The material below this point is not discussed in Video Lecture 6. It *)
|
||||
(* will be explained in Video Lecture 8. *)
|
||||
(***************************************************************************)
|
||||
|
||||
TPSpec == TPInit /\ [][TPNext]_<<rmState, tmState, tmPrepared, msgs>>
|
||||
(*************************************************************************)
|
||||
(* The complete spec of the Two-Phase Commit protocol. *)
|
||||
(*************************************************************************)
|
||||
|
||||
THEOREM TPSpec => []TPTypeOK
|
||||
(*************************************************************************)
|
||||
(* This theorem asserts that the type-correctness predicate TPTypeOK is *)
|
||||
(* an invariant of the specification. *)
|
||||
(*************************************************************************)
|
||||
-----------------------------------------------------------------------------
|
||||
(***************************************************************************)
|
||||
(* We now assert that the Two-Phase Commit protocol implements the *)
|
||||
(* Transaction Commit protocol of module TCommit. The following statement *)
|
||||
(* imports all the definitions from module TCommit into the current *)
|
||||
(* module. *)
|
||||
(***************************************************************************)
|
||||
INSTANCE TCommit
|
||||
|
||||
THEOREM TPSpec => TCSpec
|
||||
(*************************************************************************)
|
||||
(* This theorem asserts that the specification TPSpec of the Two-Phase *)
|
||||
(* Commit protocol implements the specification TCSpec of the *)
|
||||
(* Transaction Commit protocol. *)
|
||||
(*************************************************************************)
|
||||
(***************************************************************************)
|
||||
(* The two theorems in this module have been checked with TLC for six *)
|
||||
(* RMs, a configuration with 50816 reachable states, in a little over a *)
|
||||
(* minute on a 1 GHz PC. *)
|
||||
(***************************************************************************)
|
||||
|
||||
=============================================================================
|
||||
\* Modification History
|
||||
\* Last modified Thu Dec 27 15:46:42 CET 2018 by veitheller
|
||||
\* Created Thu Dec 27 15:46:21 CET 2018 by veitheller
|
20
TwoPhase.toolbox/Model_1_SnapShot_1545922927774/MC.cfg
Normal file
20
TwoPhase.toolbox/Model_1_SnapShot_1545922927774/MC.cfg
Normal file
@@ -0,0 +1,20 @@
|
||||
\* MV CONSTANT declarations
|
||||
CONSTANTS
|
||||
r1 = r1
|
||||
r2 = r2
|
||||
r3 = r3
|
||||
\* MV CONSTANT definitions
|
||||
CONSTANT
|
||||
RM <- const_154592291369123000
|
||||
\* SYMMETRY definition
|
||||
SYMMETRY symm_154592291369124000
|
||||
\* INIT definition
|
||||
INIT
|
||||
init_154592291369125000
|
||||
\* NEXT definition
|
||||
NEXT
|
||||
next_154592291369126000
|
||||
\* INVARIANT definition
|
||||
INVARIANT
|
||||
inv_154592291369127000
|
||||
\* Generated on Thu Dec 27 16:01:53 CET 2018
|
33
TwoPhase.toolbox/Model_1_SnapShot_1545922927774/MC.tla
Normal file
33
TwoPhase.toolbox/Model_1_SnapShot_1545922927774/MC.tla
Normal file
@@ -0,0 +1,33 @@
|
||||
---- MODULE MC ----
|
||||
EXTENDS TwoPhase, TLC
|
||||
|
||||
\* MV CONSTANT declarations@modelParameterConstants
|
||||
CONSTANTS
|
||||
r1, r2, r3
|
||||
----
|
||||
|
||||
\* MV CONSTANT definitions RM
|
||||
const_154592291369123000 ==
|
||||
{r1, r2, r3}
|
||||
----
|
||||
|
||||
\* SYMMETRY definition
|
||||
symm_154592291369124000 ==
|
||||
Permutations(const_154592291369123000)
|
||||
----
|
||||
|
||||
\* INIT definition @modelBehaviorInit:0
|
||||
init_154592291369125000 ==
|
||||
TPInit
|
||||
----
|
||||
\* NEXT definition @modelBehaviorNext:0
|
||||
next_154592291369126000 ==
|
||||
TPNext
|
||||
----
|
||||
\* INVARIANT definition @modelCorrectnessInvariants:0
|
||||
inv_154592291369127000 ==
|
||||
TPTypeOK
|
||||
----
|
||||
=============================================================================
|
||||
\* Modification History
|
||||
\* Created Thu Dec 27 16:01:53 CET 2018 by veitheller
|
82
TwoPhase.toolbox/Model_1_SnapShot_1545922927774/TCommit.tla
Normal file
82
TwoPhase.toolbox/Model_1_SnapShot_1545922927774/TCommit.tla
Normal file
@@ -0,0 +1,82 @@
|
||||
------------------------------ MODULE TCommit ------------------------------
|
||||
|
||||
(***************************************************************************)
|
||||
(* This specification is explained in "Transaction Commit", Lecture 5 of *)
|
||||
(* the TLA+ Video Course. *)
|
||||
(***************************************************************************)
|
||||
CONSTANT RM \* The set of participating resource managers
|
||||
|
||||
VARIABLE rmState \* rmState[rm] is the state of resource manager r.
|
||||
-----------------------------------------------------------------------------
|
||||
TCTypeOK ==
|
||||
(*************************************************************************)
|
||||
(* The type-correctness invariant *)
|
||||
(*************************************************************************)
|
||||
rmState \in [RM -> {"working", "prepared", "committed", "aborted"}]
|
||||
|
||||
TCInit == rmState = [r \in RM |-> "working"]
|
||||
(*************************************************************************)
|
||||
(* The initial predicate. *)
|
||||
(*************************************************************************)
|
||||
|
||||
canCommit == \A r \in RM : rmState[r] \in {"prepared", "committed"}
|
||||
(*************************************************************************)
|
||||
(* True iff all RMs are in the "prepared" or "committed" state. *)
|
||||
(*************************************************************************)
|
||||
|
||||
notCommitted == \A r \in RM : rmState[r] # "committed"
|
||||
(*************************************************************************)
|
||||
(* True iff no resource manager has decided to commit. *)
|
||||
(*************************************************************************)
|
||||
-----------------------------------------------------------------------------
|
||||
(***************************************************************************)
|
||||
(* We now define the actions that may be performed by the RMs, and then *)
|
||||
(* define the complete next-state action of the specification to be the *)
|
||||
(* disjunction of the possible RM actions. *)
|
||||
(***************************************************************************)
|
||||
Prepare(r) == /\ rmState[r] = "working"
|
||||
/\ rmState' = [rmState EXCEPT ![r] = "prepared"]
|
||||
|
||||
Decide(r) == \/ /\ rmState[r] = "prepared"
|
||||
/\ canCommit
|
||||
/\ rmState' = [rmState EXCEPT ![r] = "committed"]
|
||||
\/ /\ rmState[r] \in {"working", "prepared"}
|
||||
/\ notCommitted
|
||||
/\ rmState' = [rmState EXCEPT ![r] = "aborted"]
|
||||
|
||||
TCNext == \E r \in RM : Prepare(r) \/ Decide(r)
|
||||
(*************************************************************************)
|
||||
(* The next-state action. *)
|
||||
(*************************************************************************)
|
||||
-----------------------------------------------------------------------------
|
||||
TCConsistent ==
|
||||
(*************************************************************************)
|
||||
(* A state predicate asserting that two RMs have not arrived at *)
|
||||
(* conflicting decisions. It is an invariant of the specification. *)
|
||||
(*************************************************************************)
|
||||
\A r1, r2 \in RM : ~ /\ rmState[r1] = "aborted"
|
||||
/\ rmState[r2] = "committed"
|
||||
-----------------------------------------------------------------------------
|
||||
(***************************************************************************)
|
||||
(* The following part of the spec is not discussed in Video Lecture 5. It *)
|
||||
(* will be explained in Video Lecture 8. *)
|
||||
(***************************************************************************)
|
||||
TCSpec == TCInit /\ [][TCNext]_rmState
|
||||
(*************************************************************************)
|
||||
(* The complete specification of the protocol written as a temporal *)
|
||||
(* formula. *)
|
||||
(*************************************************************************)
|
||||
|
||||
THEOREM TCSpec => [](TCTypeOK /\ TCConsistent)
|
||||
(*************************************************************************)
|
||||
(* This theorem asserts the truth of the temporal formula whose meaning *)
|
||||
(* is that the state predicate TCTypeOK /\ TCInvariant is an invariant *)
|
||||
(* of the specification TCSpec. Invariance of this conjunction is *)
|
||||
(* equivalent to invariance of both of the formulas TCTypeOK and *)
|
||||
(* TCConsistent. *)
|
||||
(*************************************************************************)
|
||||
|
||||
=============================================================================
|
||||
\* Modification History
|
||||
\* Last modified Fri Dec 21 17:16:06 CET 2018 by veitheller
|
||||
\* Created Fri Dec 21 17:15:44 CET 2018 by veitheller
|
182
TwoPhase.toolbox/Model_1_SnapShot_1545922927774/TwoPhase.tla
Normal file
182
TwoPhase.toolbox/Model_1_SnapShot_1545922927774/TwoPhase.tla
Normal file
@@ -0,0 +1,182 @@
|
||||
------------------------------ MODULE TwoPhase ------------------------------
|
||||
|
||||
(***************************************************************************)
|
||||
(* This specification is discussed in "Two-Phase Commit", Lecture 6 of the *)
|
||||
(* TLA+ Video Course. It describes the Two-Phase Commit protocol, in *)
|
||||
(* which a transaction manager (TM) coordinates the resource managers *)
|
||||
(* (RMs) to implement the Transaction Commit specification of module *)
|
||||
(* TCommit. In this specification, RMs spontaneously issue Prepared *)
|
||||
(* messages. We ignore the Prepare messages that the TM can send to the *)
|
||||
(* RMs. *)
|
||||
(* *)
|
||||
(* For simplicity, we also eliminate Abort messages sent by an RM when it *)
|
||||
(* decides to abort. Such a message would cause the TM to abort the *)
|
||||
(* transaction, an event represented here by the TM spontaneously deciding *)
|
||||
(* to abort. *)
|
||||
(***************************************************************************)
|
||||
CONSTANT RM \* The set of resource managers
|
||||
|
||||
VARIABLES
|
||||
rmState, \* rmState[r] is the state of resource manager r.
|
||||
tmState, \* The state of the transaction manager.
|
||||
tmPrepared, \* The set of RMs from which the TM has received "Prepared"
|
||||
\* messages.
|
||||
msgs
|
||||
(***********************************************************************)
|
||||
(* In the protocol, processes communicate with one another by sending *)
|
||||
(* messages. For simplicity, we represent message passing with the *)
|
||||
(* variable msgs whose value is the set of all messages that have been *)
|
||||
(* sent. A message is sent by adding it to the set msgs. An action *)
|
||||
(* that, in an implementation, would be enabled by the receipt of a *)
|
||||
(* certain message is here enabled by the presence of that message in *)
|
||||
(* msgs. For simplicity, messages are never removed from msgs. This *)
|
||||
(* allows a single message to be received by multiple receivers. *)
|
||||
(* Receipt of the same message twice is therefore allowed; but in this *)
|
||||
(* particular protocol, that's not a problem. *)
|
||||
(***********************************************************************)
|
||||
|
||||
Messages ==
|
||||
(*************************************************************************)
|
||||
(* The set of all possible messages. Messages of type "Prepared" are *)
|
||||
(* sent from the RM indicated by the message's rm field to the TM. *)
|
||||
(* Messages of type "Commit" and "Abort" are broadcast by the TM, to be *)
|
||||
(* received by all RMs. The set msgs contains just a single copy of *)
|
||||
(* such a message. *)
|
||||
(*************************************************************************)
|
||||
[type : {"Prepared"}, rm : RM] \cup [type : {"Commit", "Abort"}]
|
||||
|
||||
TPTypeOK ==
|
||||
(*************************************************************************)
|
||||
(* The type-correctness invariant *)
|
||||
(*************************************************************************)
|
||||
/\ rmState \in [RM -> {"working", "prepared", "committed", "aborted"}]
|
||||
/\ tmState \in {"init", "done"}
|
||||
/\ tmPrepared \subseteq RM
|
||||
/\ msgs \subseteq Messages
|
||||
|
||||
TPInit ==
|
||||
(*************************************************************************)
|
||||
(* The initial predicate. *)
|
||||
(*************************************************************************)
|
||||
/\ rmState = [r \in RM |-> "working"]
|
||||
/\ tmState = "init"
|
||||
/\ tmPrepared = {}
|
||||
/\ msgs = {}
|
||||
-----------------------------------------------------------------------------
|
||||
(***************************************************************************)
|
||||
(* We now define the actions that may be performed by the processes, first *)
|
||||
(* the TM's actions, then the RMs' actions. *)
|
||||
(***************************************************************************)
|
||||
TMRcvPrepared(r) ==
|
||||
(*************************************************************************)
|
||||
(* The TM receives a "Prepared" message from resource manager r. We *)
|
||||
(* could add the additional enabling condition r \notin tmPrepared, *)
|
||||
(* which disables the action if the TM has already received this *)
|
||||
(* message. But there is no need, because in that case the action has *)
|
||||
(* no effect; it leaves the state unchanged. *)
|
||||
(*************************************************************************)
|
||||
/\ tmState = "init"
|
||||
/\ [type |-> "Prepared", rm |-> r] \in msgs
|
||||
/\ tmPrepared' = tmPrepared \cup {r}
|
||||
/\ UNCHANGED <<rmState, tmState, msgs>>
|
||||
|
||||
TMCommit ==
|
||||
(*************************************************************************)
|
||||
(* The TM commits the transaction; enabled iff the TM is in its initial *)
|
||||
(* state and every RM has sent a "Prepared" message. *)
|
||||
(*************************************************************************)
|
||||
/\ tmState = "init"
|
||||
/\ tmPrepared = RM
|
||||
/\ tmState' = "done"
|
||||
/\ msgs' = msgs \cup {[type |-> "Commit"]}
|
||||
/\ UNCHANGED <<rmState, tmPrepared>>
|
||||
|
||||
TMAbort ==
|
||||
(*************************************************************************)
|
||||
(* The TM spontaneously aborts the transaction. *)
|
||||
(*************************************************************************)
|
||||
/\ tmState = "init"
|
||||
/\ tmState' = "done"
|
||||
/\ msgs' = msgs \cup {[type |-> "Abort"]}
|
||||
/\ UNCHANGED <<rmState, tmPrepared>>
|
||||
|
||||
RMPrepare(r) ==
|
||||
(*************************************************************************)
|
||||
(* Resource manager r prepares. *)
|
||||
(*************************************************************************)
|
||||
/\ rmState[r] = "working"
|
||||
/\ rmState' = [rmState EXCEPT ![r] = "prepared"]
|
||||
/\ msgs' = msgs \cup {[type |-> "Prepared", rm |-> r]}
|
||||
/\ UNCHANGED <<tmState, tmPrepared>>
|
||||
|
||||
RMChooseToAbort(r) ==
|
||||
(*************************************************************************)
|
||||
(* Resource manager r spontaneously decides to abort. As noted above, r *)
|
||||
(* does not send any message in our simplified spec. *)
|
||||
(*************************************************************************)
|
||||
/\ rmState[r] = "working"
|
||||
/\ rmState' = [rmState EXCEPT ![r] = "aborted"]
|
||||
/\ UNCHANGED <<tmState, tmPrepared, msgs>>
|
||||
|
||||
RMRcvCommitMsg(r) ==
|
||||
(*************************************************************************)
|
||||
(* Resource manager r is told by the TM to commit. *)
|
||||
(*************************************************************************)
|
||||
/\ [type |-> "Commit"] \in msgs
|
||||
/\ rmState' = [rmState EXCEPT ![r] = "committed"]
|
||||
/\ UNCHANGED <<tmState, tmPrepared, msgs>>
|
||||
|
||||
RMRcvAbortMsg(r) ==
|
||||
(*************************************************************************)
|
||||
(* Resource manager r is told by the TM to abort. *)
|
||||
(*************************************************************************)
|
||||
/\ [type |-> "Abort"] \in msgs
|
||||
/\ rmState' = [rmState EXCEPT ![r] = "aborted"]
|
||||
/\ UNCHANGED <<tmState, tmPrepared, msgs>>
|
||||
|
||||
TPNext ==
|
||||
\/ TMCommit \/ TMAbort
|
||||
\/ \E r \in RM :
|
||||
TMRcvPrepared(r) \/ RMPrepare(r) \/ RMChooseToAbort(r)
|
||||
\/ RMRcvCommitMsg(r) \/ RMRcvAbortMsg(r)
|
||||
-----------------------------------------------------------------------------
|
||||
(***************************************************************************)
|
||||
(* The material below this point is not discussed in Video Lecture 6. It *)
|
||||
(* will be explained in Video Lecture 8. *)
|
||||
(***************************************************************************)
|
||||
|
||||
TPSpec == TPInit /\ [][TPNext]_<<rmState, tmState, tmPrepared, msgs>>
|
||||
(*************************************************************************)
|
||||
(* The complete spec of the Two-Phase Commit protocol. *)
|
||||
(*************************************************************************)
|
||||
|
||||
THEOREM TPSpec => []TPTypeOK
|
||||
(*************************************************************************)
|
||||
(* This theorem asserts that the type-correctness predicate TPTypeOK is *)
|
||||
(* an invariant of the specification. *)
|
||||
(*************************************************************************)
|
||||
-----------------------------------------------------------------------------
|
||||
(***************************************************************************)
|
||||
(* We now assert that the Two-Phase Commit protocol implements the *)
|
||||
(* Transaction Commit protocol of module TCommit. The following statement *)
|
||||
(* imports all the definitions from module TCommit into the current *)
|
||||
(* module. *)
|
||||
(***************************************************************************)
|
||||
INSTANCE TCommit
|
||||
|
||||
THEOREM TPSpec => TCSpec
|
||||
(*************************************************************************)
|
||||
(* This theorem asserts that the specification TPSpec of the Two-Phase *)
|
||||
(* Commit protocol implements the specification TCSpec of the *)
|
||||
(* Transaction Commit protocol. *)
|
||||
(*************************************************************************)
|
||||
(***************************************************************************)
|
||||
(* The two theorems in this module have been checked with TLC for six *)
|
||||
(* RMs, a configuration with 50816 reachable states, in a little over a *)
|
||||
(* minute on a 1 GHz PC. *)
|
||||
(***************************************************************************)
|
||||
|
||||
=============================================================================
|
||||
\* Modification History
|
||||
\* Last modified Thu Dec 27 15:46:42 CET 2018 by veitheller
|
||||
\* Created Thu Dec 27 15:46:21 CET 2018 by veitheller
|
21
TwoPhase.toolbox/Model_1_SnapShot_1545923020132/MC.cfg
Normal file
21
TwoPhase.toolbox/Model_1_SnapShot_1545923020132/MC.cfg
Normal file
@@ -0,0 +1,21 @@
|
||||
\* MV CONSTANT declarations
|
||||
CONSTANTS
|
||||
r1 = r1
|
||||
r2 = r2
|
||||
r3 = r3
|
||||
\* MV CONSTANT definitions
|
||||
CONSTANT
|
||||
RM <- const_154592300704034000
|
||||
\* SYMMETRY definition
|
||||
SYMMETRY symm_154592300704035000
|
||||
\* INIT definition
|
||||
INIT
|
||||
init_154592300704036000
|
||||
\* NEXT definition
|
||||
NEXT
|
||||
next_154592300704037000
|
||||
\* INVARIANT definition
|
||||
INVARIANT
|
||||
inv_154592300704038000
|
||||
inv_154592300704039000
|
||||
\* Generated on Thu Dec 27 16:03:27 CET 2018
|
37
TwoPhase.toolbox/Model_1_SnapShot_1545923020132/MC.tla
Normal file
37
TwoPhase.toolbox/Model_1_SnapShot_1545923020132/MC.tla
Normal file
@@ -0,0 +1,37 @@
|
||||
---- MODULE MC ----
|
||||
EXTENDS TwoPhase, TLC
|
||||
|
||||
\* MV CONSTANT declarations@modelParameterConstants
|
||||
CONSTANTS
|
||||
r1, r2, r3
|
||||
----
|
||||
|
||||
\* MV CONSTANT definitions RM
|
||||
const_154592300704034000 ==
|
||||
{r1, r2, r3}
|
||||
----
|
||||
|
||||
\* SYMMETRY definition
|
||||
symm_154592300704035000 ==
|
||||
Permutations(const_154592300704034000)
|
||||
----
|
||||
|
||||
\* INIT definition @modelBehaviorInit:0
|
||||
init_154592300704036000 ==
|
||||
TPInit
|
||||
----
|
||||
\* NEXT definition @modelBehaviorNext:0
|
||||
next_154592300704037000 ==
|
||||
TPNext
|
||||
----
|
||||
\* INVARIANT definition @modelCorrectnessInvariants:0
|
||||
inv_154592300704038000 ==
|
||||
TPTypeOK
|
||||
----
|
||||
\* INVARIANT definition @modelCorrectnessInvariants:1
|
||||
inv_154592300704039000 ==
|
||||
TCConsistent
|
||||
----
|
||||
=============================================================================
|
||||
\* Modification History
|
||||
\* Created Thu Dec 27 16:03:27 CET 2018 by veitheller
|
82
TwoPhase.toolbox/Model_1_SnapShot_1545923020132/TCommit.tla
Normal file
82
TwoPhase.toolbox/Model_1_SnapShot_1545923020132/TCommit.tla
Normal file
@@ -0,0 +1,82 @@
|
||||
------------------------------ MODULE TCommit ------------------------------
|
||||
|
||||
(***************************************************************************)
|
||||
(* This specification is explained in "Transaction Commit", Lecture 5 of *)
|
||||
(* the TLA+ Video Course. *)
|
||||
(***************************************************************************)
|
||||
CONSTANT RM \* The set of participating resource managers
|
||||
|
||||
VARIABLE rmState \* rmState[rm] is the state of resource manager r.
|
||||
-----------------------------------------------------------------------------
|
||||
TCTypeOK ==
|
||||
(*************************************************************************)
|
||||
(* The type-correctness invariant *)
|
||||
(*************************************************************************)
|
||||
rmState \in [RM -> {"working", "prepared", "committed", "aborted"}]
|
||||
|
||||
TCInit == rmState = [r \in RM |-> "working"]
|
||||
(*************************************************************************)
|
||||
(* The initial predicate. *)
|
||||
(*************************************************************************)
|
||||
|
||||
canCommit == \A r \in RM : rmState[r] \in {"prepared", "committed"}
|
||||
(*************************************************************************)
|
||||
(* True iff all RMs are in the "prepared" or "committed" state. *)
|
||||
(*************************************************************************)
|
||||
|
||||
notCommitted == \A r \in RM : rmState[r] # "committed"
|
||||
(*************************************************************************)
|
||||
(* True iff no resource manager has decided to commit. *)
|
||||
(*************************************************************************)
|
||||
-----------------------------------------------------------------------------
|
||||
(***************************************************************************)
|
||||
(* We now define the actions that may be performed by the RMs, and then *)
|
||||
(* define the complete next-state action of the specification to be the *)
|
||||
(* disjunction of the possible RM actions. *)
|
||||
(***************************************************************************)
|
||||
Prepare(r) == /\ rmState[r] = "working"
|
||||
/\ rmState' = [rmState EXCEPT ![r] = "prepared"]
|
||||
|
||||
Decide(r) == \/ /\ rmState[r] = "prepared"
|
||||
/\ canCommit
|
||||
/\ rmState' = [rmState EXCEPT ![r] = "committed"]
|
||||
\/ /\ rmState[r] \in {"working", "prepared"}
|
||||
/\ notCommitted
|
||||
/\ rmState' = [rmState EXCEPT ![r] = "aborted"]
|
||||
|
||||
TCNext == \E r \in RM : Prepare(r) \/ Decide(r)
|
||||
(*************************************************************************)
|
||||
(* The next-state action. *)
|
||||
(*************************************************************************)
|
||||
-----------------------------------------------------------------------------
|
||||
TCConsistent ==
|
||||
(*************************************************************************)
|
||||
(* A state predicate asserting that two RMs have not arrived at *)
|
||||
(* conflicting decisions. It is an invariant of the specification. *)
|
||||
(*************************************************************************)
|
||||
\A r1, r2 \in RM : ~ /\ rmState[r1] = "aborted"
|
||||
/\ rmState[r2] = "committed"
|
||||
-----------------------------------------------------------------------------
|
||||
(***************************************************************************)
|
||||
(* The following part of the spec is not discussed in Video Lecture 5. It *)
|
||||
(* will be explained in Video Lecture 8. *)
|
||||
(***************************************************************************)
|
||||
TCSpec == TCInit /\ [][TCNext]_rmState
|
||||
(*************************************************************************)
|
||||
(* The complete specification of the protocol written as a temporal *)
|
||||
(* formula. *)
|
||||
(*************************************************************************)
|
||||
|
||||
THEOREM TCSpec => [](TCTypeOK /\ TCConsistent)
|
||||
(*************************************************************************)
|
||||
(* This theorem asserts the truth of the temporal formula whose meaning *)
|
||||
(* is that the state predicate TCTypeOK /\ TCInvariant is an invariant *)
|
||||
(* of the specification TCSpec. Invariance of this conjunction is *)
|
||||
(* equivalent to invariance of both of the formulas TCTypeOK and *)
|
||||
(* TCConsistent. *)
|
||||
(*************************************************************************)
|
||||
|
||||
=============================================================================
|
||||
\* Modification History
|
||||
\* Last modified Fri Dec 21 17:16:06 CET 2018 by veitheller
|
||||
\* Created Fri Dec 21 17:15:44 CET 2018 by veitheller
|
182
TwoPhase.toolbox/Model_1_SnapShot_1545923020132/TwoPhase.tla
Normal file
182
TwoPhase.toolbox/Model_1_SnapShot_1545923020132/TwoPhase.tla
Normal file
@@ -0,0 +1,182 @@
|
||||
------------------------------ MODULE TwoPhase ------------------------------
|
||||
|
||||
(***************************************************************************)
|
||||
(* This specification is discussed in "Two-Phase Commit", Lecture 6 of the *)
|
||||
(* TLA+ Video Course. It describes the Two-Phase Commit protocol, in *)
|
||||
(* which a transaction manager (TM) coordinates the resource managers *)
|
||||
(* (RMs) to implement the Transaction Commit specification of module *)
|
||||
(* TCommit. In this specification, RMs spontaneously issue Prepared *)
|
||||
(* messages. We ignore the Prepare messages that the TM can send to the *)
|
||||
(* RMs. *)
|
||||
(* *)
|
||||
(* For simplicity, we also eliminate Abort messages sent by an RM when it *)
|
||||
(* decides to abort. Such a message would cause the TM to abort the *)
|
||||
(* transaction, an event represented here by the TM spontaneously deciding *)
|
||||
(* to abort. *)
|
||||
(***************************************************************************)
|
||||
CONSTANT RM \* The set of resource managers
|
||||
|
||||
VARIABLES
|
||||
rmState, \* rmState[r] is the state of resource manager r.
|
||||
tmState, \* The state of the transaction manager.
|
||||
tmPrepared, \* The set of RMs from which the TM has received "Prepared"
|
||||
\* messages.
|
||||
msgs
|
||||
(***********************************************************************)
|
||||
(* In the protocol, processes communicate with one another by sending *)
|
||||
(* messages. For simplicity, we represent message passing with the *)
|
||||
(* variable msgs whose value is the set of all messages that have been *)
|
||||
(* sent. A message is sent by adding it to the set msgs. An action *)
|
||||
(* that, in an implementation, would be enabled by the receipt of a *)
|
||||
(* certain message is here enabled by the presence of that message in *)
|
||||
(* msgs. For simplicity, messages are never removed from msgs. This *)
|
||||
(* allows a single message to be received by multiple receivers. *)
|
||||
(* Receipt of the same message twice is therefore allowed; but in this *)
|
||||
(* particular protocol, that's not a problem. *)
|
||||
(***********************************************************************)
|
||||
|
||||
Messages ==
|
||||
(*************************************************************************)
|
||||
(* The set of all possible messages. Messages of type "Prepared" are *)
|
||||
(* sent from the RM indicated by the message's rm field to the TM. *)
|
||||
(* Messages of type "Commit" and "Abort" are broadcast by the TM, to be *)
|
||||
(* received by all RMs. The set msgs contains just a single copy of *)
|
||||
(* such a message. *)
|
||||
(*************************************************************************)
|
||||
[type : {"Prepared"}, rm : RM] \cup [type : {"Commit", "Abort"}]
|
||||
|
||||
TPTypeOK ==
|
||||
(*************************************************************************)
|
||||
(* The type-correctness invariant *)
|
||||
(*************************************************************************)
|
||||
/\ rmState \in [RM -> {"working", "prepared", "committed", "aborted"}]
|
||||
/\ tmState \in {"init", "done"}
|
||||
/\ tmPrepared \subseteq RM
|
||||
/\ msgs \subseteq Messages
|
||||
|
||||
TPInit ==
|
||||
(*************************************************************************)
|
||||
(* The initial predicate. *)
|
||||
(*************************************************************************)
|
||||
/\ rmState = [r \in RM |-> "working"]
|
||||
/\ tmState = "init"
|
||||
/\ tmPrepared = {}
|
||||
/\ msgs = {}
|
||||
-----------------------------------------------------------------------------
|
||||
(***************************************************************************)
|
||||
(* We now define the actions that may be performed by the processes, first *)
|
||||
(* the TM's actions, then the RMs' actions. *)
|
||||
(***************************************************************************)
|
||||
TMRcvPrepared(r) ==
|
||||
(*************************************************************************)
|
||||
(* The TM receives a "Prepared" message from resource manager r. We *)
|
||||
(* could add the additional enabling condition r \notin tmPrepared, *)
|
||||
(* which disables the action if the TM has already received this *)
|
||||
(* message. But there is no need, because in that case the action has *)
|
||||
(* no effect; it leaves the state unchanged. *)
|
||||
(*************************************************************************)
|
||||
/\ tmState = "init"
|
||||
/\ [type |-> "Prepared", rm |-> r] \in msgs
|
||||
/\ tmPrepared' = tmPrepared \cup {r}
|
||||
/\ UNCHANGED <<rmState, tmState, msgs>>
|
||||
|
||||
TMCommit ==
|
||||
(*************************************************************************)
|
||||
(* The TM commits the transaction; enabled iff the TM is in its initial *)
|
||||
(* state and every RM has sent a "Prepared" message. *)
|
||||
(*************************************************************************)
|
||||
/\ tmState = "init"
|
||||
/\ tmPrepared = RM
|
||||
/\ tmState' = "done"
|
||||
/\ msgs' = msgs \cup {[type |-> "Commit"]}
|
||||
/\ UNCHANGED <<rmState, tmPrepared>>
|
||||
|
||||
TMAbort ==
|
||||
(*************************************************************************)
|
||||
(* The TM spontaneously aborts the transaction. *)
|
||||
(*************************************************************************)
|
||||
/\ tmState = "init"
|
||||
/\ tmState' = "done"
|
||||
/\ msgs' = msgs \cup {[type |-> "Abort"]}
|
||||
/\ UNCHANGED <<rmState, tmPrepared>>
|
||||
|
||||
RMPrepare(r) ==
|
||||
(*************************************************************************)
|
||||
(* Resource manager r prepares. *)
|
||||
(*************************************************************************)
|
||||
/\ rmState[r] = "working"
|
||||
/\ rmState' = [rmState EXCEPT ![r] = "prepared"]
|
||||
/\ msgs' = msgs \cup {[type |-> "Prepared", rm |-> r]}
|
||||
/\ UNCHANGED <<tmState, tmPrepared>>
|
||||
|
||||
RMChooseToAbort(r) ==
|
||||
(*************************************************************************)
|
||||
(* Resource manager r spontaneously decides to abort. As noted above, r *)
|
||||
(* does not send any message in our simplified spec. *)
|
||||
(*************************************************************************)
|
||||
/\ rmState[r] = "working"
|
||||
/\ rmState' = [rmState EXCEPT ![r] = "aborted"]
|
||||
/\ UNCHANGED <<tmState, tmPrepared, msgs>>
|
||||
|
||||
RMRcvCommitMsg(r) ==
|
||||
(*************************************************************************)
|
||||
(* Resource manager r is told by the TM to commit. *)
|
||||
(*************************************************************************)
|
||||
/\ [type |-> "Commit"] \in msgs
|
||||
/\ rmState' = [rmState EXCEPT ![r] = "committed"]
|
||||
/\ UNCHANGED <<tmState, tmPrepared, msgs>>
|
||||
|
||||
RMRcvAbortMsg(r) ==
|
||||
(*************************************************************************)
|
||||
(* Resource manager r is told by the TM to abort. *)
|
||||
(*************************************************************************)
|
||||
/\ [type |-> "Abort"] \in msgs
|
||||
/\ rmState' = [rmState EXCEPT ![r] = "aborted"]
|
||||
/\ UNCHANGED <<tmState, tmPrepared, msgs>>
|
||||
|
||||
TPNext ==
|
||||
\/ TMCommit \/ TMAbort
|
||||
\/ \E r \in RM :
|
||||
TMRcvPrepared(r) \/ RMPrepare(r) \/ RMChooseToAbort(r)
|
||||
\/ RMRcvCommitMsg(r) \/ RMRcvAbortMsg(r)
|
||||
-----------------------------------------------------------------------------
|
||||
(***************************************************************************)
|
||||
(* The material below this point is not discussed in Video Lecture 6. It *)
|
||||
(* will be explained in Video Lecture 8. *)
|
||||
(***************************************************************************)
|
||||
|
||||
TPSpec == TPInit /\ [][TPNext]_<<rmState, tmState, tmPrepared, msgs>>
|
||||
(*************************************************************************)
|
||||
(* The complete spec of the Two-Phase Commit protocol. *)
|
||||
(*************************************************************************)
|
||||
|
||||
THEOREM TPSpec => []TPTypeOK
|
||||
(*************************************************************************)
|
||||
(* This theorem asserts that the type-correctness predicate TPTypeOK is *)
|
||||
(* an invariant of the specification. *)
|
||||
(*************************************************************************)
|
||||
-----------------------------------------------------------------------------
|
||||
(***************************************************************************)
|
||||
(* We now assert that the Two-Phase Commit protocol implements the *)
|
||||
(* Transaction Commit protocol of module TCommit. The following statement *)
|
||||
(* imports all the definitions from module TCommit into the current *)
|
||||
(* module. *)
|
||||
(***************************************************************************)
|
||||
INSTANCE TCommit
|
||||
|
||||
THEOREM TPSpec => TCSpec
|
||||
(*************************************************************************)
|
||||
(* This theorem asserts that the specification TPSpec of the Two-Phase *)
|
||||
(* Commit protocol implements the specification TCSpec of the *)
|
||||
(* Transaction Commit protocol. *)
|
||||
(*************************************************************************)
|
||||
(***************************************************************************)
|
||||
(* The two theorems in this module have been checked with TLC for six *)
|
||||
(* RMs, a configuration with 50816 reachable states, in a little over a *)
|
||||
(* minute on a 1 GHz PC. *)
|
||||
(***************************************************************************)
|
||||
|
||||
=============================================================================
|
||||
\* Modification History
|
||||
\* Last modified Thu Dec 27 15:46:42 CET 2018 by veitheller
|
||||
\* Created Thu Dec 27 15:46:21 CET 2018 by veitheller
|
48
TwoPhase.toolbox/TwoPhase___Model_1.launch
Normal file
48
TwoPhase.toolbox/TwoPhase___Model_1.launch
Normal file
@@ -0,0 +1,48 @@
|
||||
<?xml version="1.0" encoding="UTF-8" standalone="no"?>
|
||||
<launchConfiguration type="org.lamport.tla.toolbox.tool.tlc.modelCheck">
|
||||
<stringAttribute key="TLCCmdLineParameters" value=""/>
|
||||
<stringAttribute key="configurationName" value="Model_1"/>
|
||||
<booleanAttribute key="deferLiveness" value="false"/>
|
||||
<intAttribute key="dfidDepth" value="100"/>
|
||||
<booleanAttribute key="dfidMode" value="false"/>
|
||||
<intAttribute key="distributedFPSetCount" value="0"/>
|
||||
<stringAttribute key="distributedNetworkInterface" value="10.7.10.46"/>
|
||||
<intAttribute key="distributedNodesCount" value="1"/>
|
||||
<stringAttribute key="distributedTLC" value="off"/>
|
||||
<stringAttribute key="distributedTLCVMArgs" value=""/>
|
||||
<intAttribute key="fpBits" value="1"/>
|
||||
<intAttribute key="fpIndex" value="1"/>
|
||||
<intAttribute key="maxHeapSize" value="25"/>
|
||||
<intAttribute key="maxSetSize" value="1000000"/>
|
||||
<booleanAttribute key="mcMode" value="true"/>
|
||||
<stringAttribute key="modelBehaviorInit" value="TPInit"/>
|
||||
<stringAttribute key="modelBehaviorNext" value="TPNext"/>
|
||||
<stringAttribute key="modelBehaviorSpec" value=""/>
|
||||
<intAttribute key="modelBehaviorSpecType" value="2"/>
|
||||
<stringAttribute key="modelBehaviorVars" value="msgs, rmState, tmState, tmPrepared"/>
|
||||
<stringAttribute key="modelComments" value=""/>
|
||||
<booleanAttribute key="modelCorrectnessCheckDeadlock" value="false"/>
|
||||
<listAttribute key="modelCorrectnessInvariants">
|
||||
<listEntry value="1TPTypeOK"/>
|
||||
<listEntry value="1TCConsistent"/>
|
||||
</listAttribute>
|
||||
<listAttribute key="modelCorrectnessProperties"/>
|
||||
<stringAttribute key="modelExpressionEval" value=""/>
|
||||
<stringAttribute key="modelParameterActionConstraint" value=""/>
|
||||
<listAttribute key="modelParameterConstants">
|
||||
<listEntry value="RM;;{r1, r2, r3};1;1"/>
|
||||
</listAttribute>
|
||||
<stringAttribute key="modelParameterContraint" value=""/>
|
||||
<listAttribute key="modelParameterDefinitions"/>
|
||||
<stringAttribute key="modelParameterModelValues" value="{}"/>
|
||||
<stringAttribute key="modelParameterNewDefinitions" value=""/>
|
||||
<intAttribute key="numberOfWorkers" value="2"/>
|
||||
<booleanAttribute key="recover" value="false"/>
|
||||
<stringAttribute key="result.mail.address" value=""/>
|
||||
<intAttribute key="simuAril" value="-1"/>
|
||||
<intAttribute key="simuDepth" value="100"/>
|
||||
<intAttribute key="simuSeed" value="-1"/>
|
||||
<stringAttribute key="specName" value="TwoPhase"/>
|
||||
<stringAttribute key="view" value=""/>
|
||||
<booleanAttribute key="visualizeStateGraph" value="false"/>
|
||||
</launchConfiguration>
|
@@ -0,0 +1,47 @@
|
||||
<?xml version="1.0" encoding="UTF-8" standalone="no"?>
|
||||
<launchConfiguration type="org.lamport.tla.toolbox.tool.tlc.modelCheck">
|
||||
<stringAttribute key="TLCCmdLineParameters" value=""/>
|
||||
<stringAttribute key="configurationName" value="Model_1_SnapShot_1545922765832"/>
|
||||
<booleanAttribute key="deferLiveness" value="false"/>
|
||||
<intAttribute key="dfidDepth" value="100"/>
|
||||
<booleanAttribute key="dfidMode" value="false"/>
|
||||
<intAttribute key="distributedFPSetCount" value="0"/>
|
||||
<stringAttribute key="distributedNetworkInterface" value="10.7.10.46"/>
|
||||
<intAttribute key="distributedNodesCount" value="1"/>
|
||||
<stringAttribute key="distributedTLC" value="off"/>
|
||||
<stringAttribute key="distributedTLCVMArgs" value=""/>
|
||||
<intAttribute key="fpBits" value="1"/>
|
||||
<intAttribute key="fpIndex" value="1"/>
|
||||
<intAttribute key="maxHeapSize" value="25"/>
|
||||
<intAttribute key="maxSetSize" value="1000000"/>
|
||||
<booleanAttribute key="mcMode" value="true"/>
|
||||
<stringAttribute key="modelBehaviorInit" value="TPInit"/>
|
||||
<stringAttribute key="modelBehaviorNext" value="TPNext"/>
|
||||
<stringAttribute key="modelBehaviorSpec" value=""/>
|
||||
<intAttribute key="modelBehaviorSpecType" value="2"/>
|
||||
<stringAttribute key="modelBehaviorVars" value="msgs, rmState, tmState, tmPrepared"/>
|
||||
<stringAttribute key="modelComments" value=""/>
|
||||
<booleanAttribute key="modelCorrectnessCheckDeadlock" value="false"/>
|
||||
<listAttribute key="modelCorrectnessInvariants">
|
||||
<listEntry value="1TPTypeOK"/>
|
||||
</listAttribute>
|
||||
<listAttribute key="modelCorrectnessProperties"/>
|
||||
<stringAttribute key="modelExpressionEval" value=""/>
|
||||
<stringAttribute key="modelParameterActionConstraint" value=""/>
|
||||
<listAttribute key="modelParameterConstants">
|
||||
<listEntry value="RM;;{"r1", "r2", "r3"};0;0"/>
|
||||
</listAttribute>
|
||||
<stringAttribute key="modelParameterContraint" value=""/>
|
||||
<listAttribute key="modelParameterDefinitions"/>
|
||||
<stringAttribute key="modelParameterModelValues" value="{}"/>
|
||||
<stringAttribute key="modelParameterNewDefinitions" value=""/>
|
||||
<intAttribute key="numberOfWorkers" value="2"/>
|
||||
<booleanAttribute key="recover" value="false"/>
|
||||
<stringAttribute key="result.mail.address" value=""/>
|
||||
<intAttribute key="simuAril" value="-1"/>
|
||||
<intAttribute key="simuDepth" value="100"/>
|
||||
<intAttribute key="simuSeed" value="-1"/>
|
||||
<stringAttribute key="specName" value="TwoPhase"/>
|
||||
<stringAttribute key="view" value=""/>
|
||||
<booleanAttribute key="visualizeStateGraph" value="false"/>
|
||||
</launchConfiguration>
|
@@ -0,0 +1,47 @@
|
||||
<?xml version="1.0" encoding="UTF-8" standalone="no"?>
|
||||
<launchConfiguration type="org.lamport.tla.toolbox.tool.tlc.modelCheck">
|
||||
<stringAttribute key="TLCCmdLineParameters" value=""/>
|
||||
<stringAttribute key="configurationName" value="Model_1_SnapShot_1545922881813"/>
|
||||
<booleanAttribute key="deferLiveness" value="false"/>
|
||||
<intAttribute key="dfidDepth" value="100"/>
|
||||
<booleanAttribute key="dfidMode" value="false"/>
|
||||
<intAttribute key="distributedFPSetCount" value="0"/>
|
||||
<stringAttribute key="distributedNetworkInterface" value="10.7.10.46"/>
|
||||
<intAttribute key="distributedNodesCount" value="1"/>
|
||||
<stringAttribute key="distributedTLC" value="off"/>
|
||||
<stringAttribute key="distributedTLCVMArgs" value=""/>
|
||||
<intAttribute key="fpBits" value="1"/>
|
||||
<intAttribute key="fpIndex" value="1"/>
|
||||
<intAttribute key="maxHeapSize" value="25"/>
|
||||
<intAttribute key="maxSetSize" value="1000000"/>
|
||||
<booleanAttribute key="mcMode" value="true"/>
|
||||
<stringAttribute key="modelBehaviorInit" value="TPInit"/>
|
||||
<stringAttribute key="modelBehaviorNext" value="TPNext"/>
|
||||
<stringAttribute key="modelBehaviorSpec" value=""/>
|
||||
<intAttribute key="modelBehaviorSpecType" value="2"/>
|
||||
<stringAttribute key="modelBehaviorVars" value="msgs, rmState, tmState, tmPrepared"/>
|
||||
<stringAttribute key="modelComments" value=""/>
|
||||
<booleanAttribute key="modelCorrectnessCheckDeadlock" value="false"/>
|
||||
<listAttribute key="modelCorrectnessInvariants">
|
||||
<listEntry value="1TPTypeOK"/>
|
||||
</listAttribute>
|
||||
<listAttribute key="modelCorrectnessProperties"/>
|
||||
<stringAttribute key="modelExpressionEval" value=""/>
|
||||
<stringAttribute key="modelParameterActionConstraint" value=""/>
|
||||
<listAttribute key="modelParameterConstants">
|
||||
<listEntry value="RM;;{"r1", "r2", "r3", "r4", "r5"};0;0"/>
|
||||
</listAttribute>
|
||||
<stringAttribute key="modelParameterContraint" value=""/>
|
||||
<listAttribute key="modelParameterDefinitions"/>
|
||||
<stringAttribute key="modelParameterModelValues" value="{}"/>
|
||||
<stringAttribute key="modelParameterNewDefinitions" value=""/>
|
||||
<intAttribute key="numberOfWorkers" value="2"/>
|
||||
<booleanAttribute key="recover" value="false"/>
|
||||
<stringAttribute key="result.mail.address" value=""/>
|
||||
<intAttribute key="simuAril" value="-1"/>
|
||||
<intAttribute key="simuDepth" value="100"/>
|
||||
<intAttribute key="simuSeed" value="-1"/>
|
||||
<stringAttribute key="specName" value="TwoPhase"/>
|
||||
<stringAttribute key="view" value=""/>
|
||||
<booleanAttribute key="visualizeStateGraph" value="false"/>
|
||||
</launchConfiguration>
|
@@ -0,0 +1,47 @@
|
||||
<?xml version="1.0" encoding="UTF-8" standalone="no"?>
|
||||
<launchConfiguration type="org.lamport.tla.toolbox.tool.tlc.modelCheck">
|
||||
<stringAttribute key="TLCCmdLineParameters" value=""/>
|
||||
<stringAttribute key="configurationName" value="Model_1_SnapShot_1545922927774"/>
|
||||
<booleanAttribute key="deferLiveness" value="false"/>
|
||||
<intAttribute key="dfidDepth" value="100"/>
|
||||
<booleanAttribute key="dfidMode" value="false"/>
|
||||
<intAttribute key="distributedFPSetCount" value="0"/>
|
||||
<stringAttribute key="distributedNetworkInterface" value="10.7.10.46"/>
|
||||
<intAttribute key="distributedNodesCount" value="1"/>
|
||||
<stringAttribute key="distributedTLC" value="off"/>
|
||||
<stringAttribute key="distributedTLCVMArgs" value=""/>
|
||||
<intAttribute key="fpBits" value="1"/>
|
||||
<intAttribute key="fpIndex" value="1"/>
|
||||
<intAttribute key="maxHeapSize" value="25"/>
|
||||
<intAttribute key="maxSetSize" value="1000000"/>
|
||||
<booleanAttribute key="mcMode" value="true"/>
|
||||
<stringAttribute key="modelBehaviorInit" value="TPInit"/>
|
||||
<stringAttribute key="modelBehaviorNext" value="TPNext"/>
|
||||
<stringAttribute key="modelBehaviorSpec" value=""/>
|
||||
<intAttribute key="modelBehaviorSpecType" value="2"/>
|
||||
<stringAttribute key="modelBehaviorVars" value="msgs, rmState, tmState, tmPrepared"/>
|
||||
<stringAttribute key="modelComments" value=""/>
|
||||
<booleanAttribute key="modelCorrectnessCheckDeadlock" value="false"/>
|
||||
<listAttribute key="modelCorrectnessInvariants">
|
||||
<listEntry value="1TPTypeOK"/>
|
||||
</listAttribute>
|
||||
<listAttribute key="modelCorrectnessProperties"/>
|
||||
<stringAttribute key="modelExpressionEval" value=""/>
|
||||
<stringAttribute key="modelParameterActionConstraint" value=""/>
|
||||
<listAttribute key="modelParameterConstants">
|
||||
<listEntry value="RM;;{r1, r2, r3};1;1"/>
|
||||
</listAttribute>
|
||||
<stringAttribute key="modelParameterContraint" value=""/>
|
||||
<listAttribute key="modelParameterDefinitions"/>
|
||||
<stringAttribute key="modelParameterModelValues" value="{}"/>
|
||||
<stringAttribute key="modelParameterNewDefinitions" value=""/>
|
||||
<intAttribute key="numberOfWorkers" value="2"/>
|
||||
<booleanAttribute key="recover" value="false"/>
|
||||
<stringAttribute key="result.mail.address" value=""/>
|
||||
<intAttribute key="simuAril" value="-1"/>
|
||||
<intAttribute key="simuDepth" value="100"/>
|
||||
<intAttribute key="simuSeed" value="-1"/>
|
||||
<stringAttribute key="specName" value="TwoPhase"/>
|
||||
<stringAttribute key="view" value=""/>
|
||||
<booleanAttribute key="visualizeStateGraph" value="false"/>
|
||||
</launchConfiguration>
|
@@ -0,0 +1,48 @@
|
||||
<?xml version="1.0" encoding="UTF-8" standalone="no"?>
|
||||
<launchConfiguration type="org.lamport.tla.toolbox.tool.tlc.modelCheck">
|
||||
<stringAttribute key="TLCCmdLineParameters" value=""/>
|
||||
<stringAttribute key="configurationName" value="Model_1_SnapShot_1545923020132"/>
|
||||
<booleanAttribute key="deferLiveness" value="false"/>
|
||||
<intAttribute key="dfidDepth" value="100"/>
|
||||
<booleanAttribute key="dfidMode" value="false"/>
|
||||
<intAttribute key="distributedFPSetCount" value="0"/>
|
||||
<stringAttribute key="distributedNetworkInterface" value="10.7.10.46"/>
|
||||
<intAttribute key="distributedNodesCount" value="1"/>
|
||||
<stringAttribute key="distributedTLC" value="off"/>
|
||||
<stringAttribute key="distributedTLCVMArgs" value=""/>
|
||||
<intAttribute key="fpBits" value="1"/>
|
||||
<intAttribute key="fpIndex" value="1"/>
|
||||
<intAttribute key="maxHeapSize" value="25"/>
|
||||
<intAttribute key="maxSetSize" value="1000000"/>
|
||||
<booleanAttribute key="mcMode" value="true"/>
|
||||
<stringAttribute key="modelBehaviorInit" value="TPInit"/>
|
||||
<stringAttribute key="modelBehaviorNext" value="TPNext"/>
|
||||
<stringAttribute key="modelBehaviorSpec" value=""/>
|
||||
<intAttribute key="modelBehaviorSpecType" value="2"/>
|
||||
<stringAttribute key="modelBehaviorVars" value="msgs, rmState, tmState, tmPrepared"/>
|
||||
<stringAttribute key="modelComments" value=""/>
|
||||
<booleanAttribute key="modelCorrectnessCheckDeadlock" value="false"/>
|
||||
<listAttribute key="modelCorrectnessInvariants">
|
||||
<listEntry value="1TPTypeOK"/>
|
||||
<listEntry value="1TCConsistent"/>
|
||||
</listAttribute>
|
||||
<listAttribute key="modelCorrectnessProperties"/>
|
||||
<stringAttribute key="modelExpressionEval" value=""/>
|
||||
<stringAttribute key="modelParameterActionConstraint" value=""/>
|
||||
<listAttribute key="modelParameterConstants">
|
||||
<listEntry value="RM;;{r1, r2, r3};1;1"/>
|
||||
</listAttribute>
|
||||
<stringAttribute key="modelParameterContraint" value=""/>
|
||||
<listAttribute key="modelParameterDefinitions"/>
|
||||
<stringAttribute key="modelParameterModelValues" value="{}"/>
|
||||
<stringAttribute key="modelParameterNewDefinitions" value=""/>
|
||||
<intAttribute key="numberOfWorkers" value="2"/>
|
||||
<booleanAttribute key="recover" value="false"/>
|
||||
<stringAttribute key="result.mail.address" value=""/>
|
||||
<intAttribute key="simuAril" value="-1"/>
|
||||
<intAttribute key="simuDepth" value="100"/>
|
||||
<intAttribute key="simuSeed" value="-1"/>
|
||||
<stringAttribute key="specName" value="TwoPhase"/>
|
||||
<stringAttribute key="view" value=""/>
|
||||
<booleanAttribute key="visualizeStateGraph" value="false"/>
|
||||
</launchConfiguration>
|
Reference in New Issue
Block a user