| author | wenzelm | 
| Tue, 26 Sep 2000 17:34:33 +0200 | |
| changeset 10086 | 5245fa2ca8d3 | 
| parent 10064 | 1a77667b21ef | 
| child 10797 | 028d22926a41 | 
| permissions | -rw-r--r-- | 
| 4776 | 1 | (* Title: HOL/UNITY/UNITY | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 4 | Copyright 1998 University of Cambridge | |
| 5 | ||
| 6 | The basic UNITY theory (revised version, based upon the "co" operator) | |
| 7 | ||
| 8 | From Misra, "A Logic for Concurrent Programming", 1994 | |
| 9 | *) | |
| 10 | ||
| 8948 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
6823diff
changeset | 11 | UNITY = Main + | 
| 6535 | 12 | |
| 13 | typedef (Program) | |
| 10064 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
8948diff
changeset | 14 |   'a program = "{(init:: 'a set, acts :: ('a * 'a)set set,
 | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
8948diff
changeset | 15 | 		  allowed :: ('a * 'a)set set). Id:acts & Id: allowed}"
 | 
| 4776 | 16 | |
| 6536 | 17 | consts | 
| 6823 | 18 | constrains :: "['a set, 'a set] => 'a program set" (infixl "co" 60) | 
| 19 | op_unless :: "['a set, 'a set] => 'a program set" (infixl "unless" 60) | |
| 6536 | 20 | |
| 4776 | 21 | constdefs | 
| 10064 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
8948diff
changeset | 22 |     mk_program :: "('a set * ('a * 'a)set set * ('a * 'a)set set)
 | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
8948diff
changeset | 23 | => 'a program" | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
8948diff
changeset | 24 | "mk_program == %(init, acts, allowed). | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
8948diff
changeset | 25 | Abs_Program (init, insert Id acts, insert Id allowed)" | 
| 6535 | 26 | |
| 27 | Init :: "'a program => 'a set" | |
| 10064 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
8948diff
changeset | 28 | "Init F == (%(init, acts, allowed). init) (Rep_Program F)" | 
| 6535 | 29 | |
| 30 |   Acts :: "'a program => ('a * 'a)set set"
 | |
| 10064 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
8948diff
changeset | 31 | "Acts F == (%(init, acts, allowed). acts) (Rep_Program F)" | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
8948diff
changeset | 32 | |
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
8948diff
changeset | 33 |   AllowedActs :: "'a program => ('a * 'a)set set"
 | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
8948diff
changeset | 34 | "AllowedActs F == (%(init, acts, allowed). allowed) (Rep_Program F)" | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
8948diff
changeset | 35 | |
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
8948diff
changeset | 36 | Allowed :: "'a program => 'a program set" | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
8948diff
changeset | 37 |     "Allowed F == {G. Acts G <= AllowedActs F}"
 | 
| 4776 | 38 | |
| 5648 | 39 | stable :: "'a set => 'a program set" | 
| 6536 | 40 | "stable A == A co A" | 
| 4776 | 41 | |
| 5648 | 42 | strongest_rhs :: "['a program, 'a set] => 'a set" | 
| 6536 | 43 |     "strongest_rhs F A == Inter {B. F : A co B}"
 | 
| 4776 | 44 | |
| 5648 | 45 | invariant :: "'a set => 'a program set" | 
| 46 |     "invariant A == {F. Init F <= A} Int stable A"
 | |
| 4776 | 47 | |
| 5648 | 48 | (*Polymorphic in both states and the meaning of <= *) | 
| 6713 | 49 |   increasing :: "['a => 'b::{order}] => 'a program set"
 | 
| 5648 | 50 |     "increasing f == INT z. stable {s. z <= f s}"
 | 
| 4776 | 51 | |
| 6536 | 52 | |
| 53 | defs | |
| 54 |   constrains_def "A co B == {F. ALL act: Acts F. act^^A <= B}"
 | |
| 55 | ||
| 56 | unless_def "A unless B == (A-B) co (A Un B)" | |
| 57 | ||
| 4776 | 58 | end |