Theory CIMP_lang
theory CIMP_lang
imports
CIMP_pred
LTL
begin
section‹CIMP syntax and semantics \label{sec:cimp-syntax-semantics}›
text‹
We define a small sequential programming language with synchronous
message passing primitives for describing the individual
processes. This has the advantage over raw transition systems in that
it is programmer-readable, includes sequential composition, supports a
program logic and VCG (\S\ref{sec:cimp-vcg}), etc. These processes are
composed in parallel at the top-level.
CIMP is inspired by IMP, as presented by \<^citet>‹"Winskel:1993"› and \<^citet>‹"ConcreteSemantics:2014"›, and the classical process algebras CCS
\<^citep>‹"Milner:1980" and "Milner:1989"› and CSP
\<^citep>‹"Hoare:1985"›. Note that the algebraic
properties of this language have not been developed.
As we operate in a concurrent setting, we need to provide a small-step
semantics (\S\ref{sec:cimp-semantics}), which we give in the style of
\emph{structural operational semantics} (SOS) as popularised by \<^citet>‹"DBLP:journals/jlp/Plotkin04"›. The semantics of a
complete system (\S\ref{sec:cimp-system-steps}) is presently taken
simply to be the states reachable by interleaving the enabled steps of
the individual processes, subject to message passing rendezvous. We
leave a trace or branching semantics to future work.
This theory contains all the trusted definitions. The soundness of the other
theories supervenes upon this one.
›
subsection‹Syntax›
text‹
Programs are represented using an explicit (deep embedding) of their
syntax, as the semantics needs to track the progress of multiple
threads of control. Each (atomic) \emph{basic command}
(\S\ref{sec:cimp-decompose}) is annotated with a @{typ "'location"},
which we use in our assertions (\S\ref{sec:cimp-control-predicates}).
These locations need not be unique, though in practice they likely
will be.
Processes maintain \emph{local states} of type @{typ "'state"}. These
can be updated with arbitrary relations of @{typ "'state ⇒ 'state
set"} with ‹LocalOp›, and conditions of type @{typ "'s ⇒
bool"} are similarly shallowly embedded. This arrangement allows the
end-user to select their own level of atomicity.
The sequential composition operator and control constructs are
standard. We add the infinite looping construct ‹Loop› so we
can construct single-state reactive systems; this has implications for
fairness assertions.
›
type_synonym 's bexp = "'s ⇒ bool"