| 59189 |      1 | chapter \<open>The Owicki-Gries Method\<close>
 | 
| 13020 |      2 | 
 | 
| 59189 |      3 | section \<open>Abstract Syntax\<close>
 | 
| 13020 |      4 | 
 | 
| 16417 |      5 | theory OG_Com imports Main begin
 | 
| 13020 |      6 | 
 | 
| 59189 |      7 | text \<open>Type abbreviations for boolean expressions and assertions:\<close>
 | 
| 13020 |      8 | 
 | 
| 42174 |      9 | type_synonym 'a bexp = "'a set"
 | 
|  |     10 | type_synonym 'a assn = "'a set"
 | 
| 13020 |     11 | 
 | 
| 59189 |     12 | text \<open>The syntax of commands is defined by two mutually recursive
 | 
| 62042 |     13 | datatypes: \<open>'a ann_com\<close> for annotated commands and \<open>'a
 | 
|  |     14 | com\<close> for non-annotated commands.\<close>
 | 
| 13020 |     15 | 
 | 
| 59189 |     16 | datatype 'a ann_com =
 | 
|  |     17 |      AnnBasic "('a assn)"  "('a \<Rightarrow> 'a)"
 | 
|  |     18 |    | AnnSeq "('a ann_com)"  "('a ann_com)"
 | 
|  |     19 |    | AnnCond1 "('a assn)"  "('a bexp)"  "('a ann_com)"  "('a ann_com)"
 | 
|  |     20 |    | AnnCond2 "('a assn)"  "('a bexp)"  "('a ann_com)"
 | 
|  |     21 |    | AnnWhile "('a assn)"  "('a bexp)"  "('a assn)"  "('a ann_com)"
 | 
|  |     22 |    | AnnAwait "('a assn)"  "('a bexp)"  "('a com)"
 | 
|  |     23 | and 'a com =
 | 
| 13020 |     24 |      Parallel "('a ann_com option \<times> 'a assn) list"
 | 
| 59189 |     25 |    | Basic "('a \<Rightarrow> 'a)"
 | 
|  |     26 |    | Seq "('a com)"  "('a com)"
 | 
|  |     27 |    | Cond "('a bexp)"  "('a com)"  "('a com)"
 | 
| 13020 |     28 |    | While "('a bexp)"  "('a assn)"  "('a com)"
 | 
|  |     29 | 
 | 
| 62042 |     30 | text \<open>The function \<open>pre\<close> extracts the precondition of an
 | 
| 59189 |     31 | annotated command:\<close>
 | 
| 13020 |     32 | 
 | 
| 39246 |     33 | primrec pre ::"'a ann_com \<Rightarrow> 'a assn"  where
 | 
| 13020 |     34 |   "pre (AnnBasic r f) = r"
 | 
| 39246 |     35 | | "pre (AnnSeq c1 c2) = pre c1"
 | 
|  |     36 | | "pre (AnnCond1 r b c1 c2) = r"
 | 
|  |     37 | | "pre (AnnCond2 r b c) = r"
 | 
|  |     38 | | "pre (AnnWhile r b i c) = r"
 | 
|  |     39 | | "pre (AnnAwait r b c) = r"
 | 
| 13020 |     40 | 
 | 
| 59189 |     41 | text \<open>Well-formedness predicate for atomic programs:\<close>
 | 
| 13020 |     42 | 
 | 
| 39246 |     43 | primrec atom_com :: "'a com \<Rightarrow> bool" where
 | 
| 13020 |     44 |   "atom_com (Parallel Ts) = False"
 | 
| 39246 |     45 | | "atom_com (Basic f) = True"
 | 
|  |     46 | | "atom_com (Seq c1 c2) = (atom_com c1 \<and> atom_com c2)"
 | 
|  |     47 | | "atom_com (Cond b c1 c2) = (atom_com c1 \<and> atom_com c2)"
 | 
|  |     48 | | "atom_com (While b i c) = atom_com c"
 | 
| 59189 |     49 | 
 | 
| 62390 |     50 | end
 |