Theory Introduction
chapter‹ Introduction ›
theory Introduction
imports "HOL-CSP.CSP"
begin
section‹Motivations›
text ‹\<^session>‹HOL-CSP› \<^cite>‹"HOL-CSP-AFP"› is a formalization in Isabelle/HOL of
the work of Hoare and Roscoe on the denotational semantics of the
Failure/Divergence Model of CSP. It follows essentially the presentation of CSP
in Roscoe's Book "Theory and Practice of Concurrency" \<^cite>‹"roscoe:csp:1998"›
and the semantic details in a joint Paper of Roscoe and Brooks
"An improved failures model for communicating processes" \<^cite>‹"brookes-roscoe85"›.›
text ‹In the session \<^session>‹HOL-CSP› are introduced the type \<^typ>‹('a, 'r) process⇩p⇩t⇩i⇩c⇩k›, several
classic CSP operators and number of laws that govern their interactions.
Four of them are binary operators: the non-deterministic choice \<^term>‹P ⊓ Q›,
the deterministic choice \<^term>‹P □ Q›, the synchronization \<^term>‹P ⟦S⟧ Q› and the
sequential composition \<^term>‹P ❙; Q›.›
text ‹Analogously to the finite sum
@{term [mode=latex_sum, eta_contract = false] ‹∑i = 0::nat..n. a i›} which is generalization
of the addition \<^term>‹a + b›, we define generalisations of the binary operators
of CSP.
The most straight-forward way to do so would be a fold on a list of processes.
However, in many cases, we have additional properties, like commutativity, idempotency, etc.
that allow for stronger/more abstract constructions. In particular, in several cases,
generalization to unbounded and even infinite index-sets are possible.
The notations we choose are widely inspired by the CSP$_M$ syntax of FDR:
🌐‹https://cocotec.io/fdr/manual/cspm.html›.›
text ‹For the non-deterministic choice \<^term>‹P ⊓ Q›, this is already done in \<^session>‹HOL-CSP›.
In this session we therefore introduce the multi-operators:
▪ the global deterministic choice, written ‹□ a ∈ A. P a›, generalizing \<^term>‹P □ Q›
▪ the multi-synchronization product, written ‹❙⟦S❙⟧ m ∈# M. P m›, generalizing \<^term>‹P ⟦S⟧ Q›
with the two special cases ‹❙|❙|❙| m ∈# M. P m› and ‹❙|❙| m ∈# M. P m›
▪ the multi-sequential composition, written ‹SEQ l ∈@ L. P l›, generalizing \<^term>‹P ❙; Q›.
We prove their continuity and refinements rules,
as well as some laws governing their interactions.›
text ‹We also provide the definitions of the POTS and Dining Philosophers examples,
which greatly benefit from the newly introduced generalized operators.
Since they appear naturally when modeling complex architectures,
we may call them ∗‹architectural operators›:
these multi-operators represent the heart of the architectural composition principles of CSP.›
text ‹Additionally, we developed the theory of the interrupt operators ∗‹Sliding›, ∗‹Throw›
and ∗‹Interrupt› \<^cite>‹"Roscoe2010UnderstandingCS"›.
This part of the present theory reintroduces denotational semantics for these operators and
constructs on this basis the algebraic laws for them.
In several places, our formalization efforts led to slight modifications of the original
definitions in order to achieve the goal of a combined integrated theory.
In some cases -- in particular in connection with the ∗‹Interrupt› operator definition --
some corrections have been necessary since the fundamental invariants were not respected.›
text ‹Finally, his session includes a very powerful result about \<^const>‹deadlock_free›
and \<^const>‹Sync›: the interleaving \<^term>‹P ||| Q› is \<^const>‹deadlock_free› if \<^term>‹P› and \<^term>‹Q› are,
and so is the multi-interleaving of processes ‹P m› for ‹m ∈# M›.
\newpage›
section‹The Global Architecture of HOL-CSPM›
text‹
\begin{figure}[ht]
\centering
\includegraphics[width=0.85\textwidth]{session_graph.pdf}
\caption{The overall architecture}
\label{fig:fig1}
\end{figure}
›
text‹The global architecture of HOL-CSPM is shown in \autoref{fig:fig1}.
The entire package resides on:
▸ \<^session>‹HOL-Eisbach› from the Isabelle/HOL distribution,
▸ \<^session>‹HOLCF› from the Isabelle/HOL distribution, and
▸ \<^session>‹HOL-CSP› 2.0 from the Isabelle Archive of Formal Proofs.
›
end