| author | wenzelm | 
| Mon, 18 Sep 2000 14:49:51 +0200 | |
| changeset 10017 | e146bbfc38c1 | 
| parent 8948 | b797cfa3548d | 
| child 10064 | 1a77667b21ef | 
| 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) | |
| 14 |   'a program = "{(init:: 'a set, acts :: ('a * 'a)set set). Id:acts}"
 | |
| 4776 | 15 | |
| 6536 | 16 | consts | 
| 6823 | 17 | constrains :: "['a set, 'a set] => 'a program set" (infixl "co" 60) | 
| 18 | op_unless :: "['a set, 'a set] => 'a program set" (infixl "unless" 60) | |
| 6536 | 19 | |
| 4776 | 20 | constdefs | 
| 6535 | 21 |     mk_program :: "('a set * ('a * 'a)set set) => 'a program"
 | 
| 22 | "mk_program == %(init, acts). Abs_Program (init, insert Id acts)" | |
| 23 | ||
| 24 | Init :: "'a program => 'a set" | |
| 25 | "Init F == (%(init, acts). init) (Rep_Program F)" | |
| 26 | ||
| 27 |   Acts :: "'a program => ('a * 'a)set set"
 | |
| 28 | "Acts F == (%(init, acts). acts) (Rep_Program F)" | |
| 4776 | 29 | |
| 5648 | 30 | stable :: "'a set => 'a program set" | 
| 6536 | 31 | "stable A == A co A" | 
| 4776 | 32 | |
| 5648 | 33 | strongest_rhs :: "['a program, 'a set] => 'a set" | 
| 6536 | 34 |     "strongest_rhs F A == Inter {B. F : A co B}"
 | 
| 4776 | 35 | |
| 5648 | 36 | invariant :: "'a set => 'a program set" | 
| 37 |     "invariant A == {F. Init F <= A} Int stable A"
 | |
| 4776 | 38 | |
| 5648 | 39 | (*Polymorphic in both states and the meaning of <= *) | 
| 6713 | 40 |   increasing :: "['a => 'b::{order}] => 'a program set"
 | 
| 5648 | 41 |     "increasing f == INT z. stable {s. z <= f s}"
 | 
| 4776 | 42 | |
| 6536 | 43 | |
| 44 | defs | |
| 45 |   constrains_def "A co B == {F. ALL act: Acts F. act^^A <= B}"
 | |
| 46 | ||
| 47 | unless_def "A unless B == (A-B) co (A Un B)" | |
| 48 | ||
| 4776 | 49 | end |