src/HOL/HoareParallel/RG_Com.thy
author berghofe
Wed, 07 May 2008 10:56:36 +0200
changeset 26794 354c3844dfde
parent 16417 9bc16273c2d4
permissions -rw-r--r--
- Now imports Fun rather than Orderings - Moved "Set as lattice" section behind "Fun as lattice" section, since sets are just functions. - The instantiations instantiation set :: (type) distrib_lattice instantiation set :: (type) complete_lattice are no longer needed, and the former definitions inf_set_eq, sup_set_eq, Inf_set_def, and Sup_set_def can now be derived from abstract properties of sup, inf, etc.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13020
791e3b4c4039 HoareParallel Theories
prensani
parents:
diff changeset
     1
791e3b4c4039 HoareParallel Theories
prensani
parents:
diff changeset
     2
header {* \chapter{The Rely-Guarantee Method} 
791e3b4c4039 HoareParallel Theories
prensani
parents:
diff changeset
     3
791e3b4c4039 HoareParallel Theories
prensani
parents:
diff changeset
     4
\section {Abstract Syntax}
791e3b4c4039 HoareParallel Theories
prensani
parents:
diff changeset
     5
*}
791e3b4c4039 HoareParallel Theories
prensani
parents:
diff changeset
     6
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 13020
diff changeset
     7
theory RG_Com imports Main begin
13020
791e3b4c4039 HoareParallel Theories
prensani
parents:
diff changeset
     8
791e3b4c4039 HoareParallel Theories
prensani
parents:
diff changeset
     9
text {* Semantics of assertions and boolean expressions (bexp) as sets
791e3b4c4039 HoareParallel Theories
prensani
parents:
diff changeset
    10
of states.  Syntax of commands @{text com} and parallel commands
791e3b4c4039 HoareParallel Theories
prensani
parents:
diff changeset
    11
@{text par_com}. *}
791e3b4c4039 HoareParallel Theories
prensani
parents:
diff changeset
    12
791e3b4c4039 HoareParallel Theories
prensani
parents:
diff changeset
    13
types
791e3b4c4039 HoareParallel Theories
prensani
parents:
diff changeset
    14
  'a bexp = "'a set"
791e3b4c4039 HoareParallel Theories
prensani
parents:
diff changeset
    15
791e3b4c4039 HoareParallel Theories
prensani
parents:
diff changeset
    16
datatype 'a com = 
791e3b4c4039 HoareParallel Theories
prensani
parents:
diff changeset
    17
    Basic "'a \<Rightarrow>'a"
791e3b4c4039 HoareParallel Theories
prensani
parents:
diff changeset
    18
  | Seq "'a com" "'a com"
791e3b4c4039 HoareParallel Theories
prensani
parents:
diff changeset
    19
  | Cond "'a bexp" "'a com" "'a com"         
791e3b4c4039 HoareParallel Theories
prensani
parents:
diff changeset
    20
  | While "'a bexp" "'a com"       
791e3b4c4039 HoareParallel Theories
prensani
parents:
diff changeset
    21
  | Await "'a bexp" "'a com"                 
791e3b4c4039 HoareParallel Theories
prensani
parents:
diff changeset
    22
791e3b4c4039 HoareParallel Theories
prensani
parents:
diff changeset
    23
types 'a par_com = "(('a com) option) list"
791e3b4c4039 HoareParallel Theories
prensani
parents:
diff changeset
    24
791e3b4c4039 HoareParallel Theories
prensani
parents:
diff changeset
    25
end