| 
13020
 | 
     1  | 
  | 
| 
 | 
     2  | 
header {* \chapter{The Rely-Guarantee Method} 
 | 
| 
 | 
     3  | 
  | 
| 
 | 
     4  | 
\section {Abstract Syntax}
 | 
| 
 | 
     5  | 
*}
  | 
| 
 | 
     6  | 
  | 
| 
16417
 | 
     7  | 
theory RG_Com imports Main begin
  | 
| 
13020
 | 
     8  | 
  | 
| 
 | 
     9  | 
text {* Semantics of assertions and boolean expressions (bexp) as sets
 | 
| 
 | 
    10  | 
of states.  Syntax of commands @{text com} and parallel commands
 | 
| 
 | 
    11  | 
@{text par_com}. *}
 | 
| 
 | 
    12  | 
  | 
| 
42174
 | 
    13  | 
type_synonym 'a bexp = "'a set"
  | 
| 
13020
 | 
    14  | 
  | 
| 
58310
 | 
    15  | 
datatype 'a com = 
  | 
| 
13020
 | 
    16  | 
    Basic "'a \<Rightarrow>'a"
  | 
| 
 | 
    17  | 
  | Seq "'a com" "'a com"
  | 
| 
 | 
    18  | 
  | Cond "'a bexp" "'a com" "'a com"         
  | 
| 
 | 
    19  | 
  | While "'a bexp" "'a com"       
  | 
| 
 | 
    20  | 
  | Await "'a bexp" "'a com"                 
  | 
| 
 | 
    21  | 
  | 
| 
42174
 | 
    22  | 
type_synonym 'a par_com = "'a com option list"
  | 
| 
13020
 | 
    23  | 
  | 
| 
 | 
    24  | 
end  |