Theory AsynchronousSystem
section ‹AsynchronousSystem›
text ‹
\file{AsynchronousSystem} defines a message datatype and a transition system locale to
model asynchronous distributed computation. It establishes a diamond property for a
special reachability relation within such transition systems.
›
theory AsynchronousSystem
imports Multiset
begin
text‹
The formalization is type-parameterized over
\begin{description}
\item[\var{'p} process identifiers.] Corresponds to the set $P$ in Völzer.
Finiteness is not yet demanded, but will be in \file{FLPSystem}.
\item[\var{'s} process states.] Corresponds to $S$, countability is not
imposed.
\item[\var{'v} message payloads.] Corresponds to the interprocess
communication part of $M$ from Völzer. The whole of $M$ is captured by
\isb{messageValue}.
\end{description}
›
subsection‹Messages›
text ‹
A \isb{message} is either an initial input message telling a process
which value it should introduce to the consensus negotiation, a message
to the environment communicating the consensus outcome, or a message
passed from one process to some other.
›