Theory Introduction

(*<*)
―‹ ********************************************************************
 * Project         : HOL-CSPM - Architectural operators for HOL-CSP
 *
 * Author          : Benoît Ballenghien, Safouan Taha, Burkhart Wolff.
 *
 * This file       : Introduction
 *
 * Copyright (c) 2025 Université Paris-Saclay, France
 *
 * All rights reserved.
 *
 * Redistribution and use in source and binary forms, with or without
 * modification, are permitted provided that the following conditions are
 * met:
 *
 *     * Redistributions of source code must retain the above copyright
 *       notice, this list of conditions and the following disclaimer.
 *
 *     * Redistributions in binary form must reproduce the above
 *       copyright notice, this list of conditions and the following
 *       disclaimer in the documentation and/or other materials provided
 *       with the distribution.
 *
 *     * Neither the name of the copyright holders nor the names of its
 *       contributors may be used to endorse or promote products derived
 *       from this software without specific prior written permission.
 *
 * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS
 * "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT
 * LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR
 * A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT
 * OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL,
 * SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT
 * LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE,
 * DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY
 * THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT
 * (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
 * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
 ******************************************************************************›
(*>*)


chapter‹ Introduction ›

(*<*)
theory Introduction                                               
  imports "HOL-CSP.CSP"
begin 
  (*>*)

section‹Motivations›

text sessionHOL-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 sessionHOL-CSP are introduced the type typ('a, 'r) processptick, several 
      classic CSP operators and number of laws that govern their interactions.

      Four of them are binary operators: the non-deterministic choice termP  Q, 
      the deterministic choice termP  Q, the synchronization termP S Q and the
      sequential composition termP ; 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 terma + 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 termP  Q, this is already done in sessionHOL-CSP.
      In this session we therefore introduce the multi-operators:
         the global deterministic choice, written □ a ∈ A. P a›, generalizing termP  Q
         the multi-synchronization product, written S m ∈# M. P m›, generalizing termP 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 termP ; 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 constdeadlock_free
      and constSync: the interleaving termP ||| Q is constdeadlock_free if termP and termQ 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: 
 sessionHOL-Eisbach from the Isabelle/HOL distribution,
 sessionHOLCF from the Isabelle/HOL distribution, and
 sessionHOL-CSP 2.0 from the Isabelle Archive of Formal Proofs.
›

(*<*)
end
  (*>*)