| author | wenzelm | 
| Wed, 20 Feb 2002 00:53:53 +0100 | |
| changeset 12902 | a23dc0b7566f | 
| parent 11194 | ea13ff5a26d1 | 
| child 13786 | ab8f39f48a6f | 
| 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/TimerArray.thy | 
| 
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: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 4 | Copyright 1998 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 trivial example of reasoning about an array of processes | 
| 
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 | |
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 9 | TimerArray = PPROD + | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 10 | |
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 11 | types 'a state = "nat * 'a" (*second component allows new variables*) | 
| 
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 | constdefs | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 14 | count :: "'a state => nat" | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 15 | "count s == fst s" | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 16 | |
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 17 |   decr  :: "('a state * 'a state) set"
 | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 18 |     "decr == UN n uu. {((Suc n, uu), (n,uu))}"
 | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 19 | |
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 20 | Timer :: 'a state program | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 21 |     "Timer == mk_program (UNIV, {decr}, UNIV)"
 | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 22 | |
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 23 | end |