src/HOL/HoareParallel/RG_Com.thy
author berghofe
Fri, 01 Jul 2005 13:54:12 +0200
changeset 16633 208ebc9311f2
parent 16417 9bc16273c2d4
permissions -rw-r--r--
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification of premises of congruence rules.


header {* \chapter{The Rely-Guarantee Method} 

\section {Abstract Syntax}
*}

theory RG_Com imports Main begin

text {* Semantics of assertions and boolean expressions (bexp) as sets
of states.  Syntax of commands @{text com} and parallel commands
@{text par_com}. *}

types
  'a bexp = "'a set"

datatype 'a com = 
    Basic "'a \<Rightarrow>'a"
  | Seq "'a com" "'a com"
  | Cond "'a bexp" "'a com" "'a com"         
  | While "'a bexp" "'a com"       
  | Await "'a bexp" "'a com"                 

types 'a par_com = "(('a com) option) list"

end