| author | wenzelm | 
| Sat, 24 Nov 2001 17:21:47 +0100 | |
| changeset 12294 | 2ef49890aede | 
| parent 11868 | 56db9f3a6b3e | 
| child 12338 | de0f4a63baa5 | 
| permissions | -rw-r--r-- | 
| 11194 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 1 | (* Title: HOL/UNITY/Counterc | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 2 | ID: $Id$ | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 3 | Author: Sidi O Ehmety, Cambridge University Computer Laboratory | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 4 | Copyright 2001 University of Cambridge | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 5 | |
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 6 | A family of similar counters, version with a full use of "compatibility " | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 7 | |
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 8 | From Charpentier and Chandy, | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 9 | Examples of Program Composition Illustrating the Use of Universal Properties | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 10 | In J. Rolim (editor), Parallel and Distributed Processing, | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 11 | Spriner LNCS 1586 (1999), pages 1215-1227. | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 12 | *) | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 13 | |
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 14 | Counterc = Comp + | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 15 | types state | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 16 | arities state :: term | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 17 | |
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 18 | consts | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 19 | C :: "state=>int" | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 20 | c :: "state=>nat=>int" | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 21 | |
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 22 | consts | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 23 | sum :: "[nat,state]=>int" | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 24 | sumj :: "[nat, nat, state]=>int" | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 25 | |
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 26 | primrec (* sum I s = sigma_{i<I}. c s i *)
 | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 27 | "sum 0 s = 0" | 
| 11194 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 28 | "sum (Suc i) s = (c s) i + sum i s" | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 29 | |
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 30 | primrec | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 31 | "sumj 0 i s = 0" | 
| 11194 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 32 | "sumj (Suc n) i s = (if n=i then sum n s else (c s) n + sumj n i s)" | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 33 | |
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 34 | types command = "(state*state)set" | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 35 | |
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 36 | constdefs | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 37 | a :: "nat=>command" | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 38 |  "a i == {(s, s'). (c s') i = (c s) i + 1 & (C s') = (C s) + 1}"
 | 
| 11194 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 39 | |
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 40 | Component :: "nat => state program" | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 41 |   "Component i == mk_program({s. C s = 0 & (c s) i = 0}, {a i},
 | 
| 11194 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 42 | UN G: preserves (%s. (c s) i). Acts G)" | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 43 | end |