src/HOL/UNITY/TimerArray.thy
author paulson
Sat, 23 Sep 2000 16:02:01 +0200
changeset 10064 1a77667b21ef
parent 8251 9be357df93d4
permissions -rw-r--r--
added compatibility relation: AllowedActs, Allowed, ok, OK and changes to "guarantees", etc.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
7513
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/TimerArray.thy
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
     2
    ID:         $Id$
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
     4
    Copyright   1998  University of Cambridge
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
     5
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
     6
A trivial example of reasoning about an array of processes
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
     7
*)
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
     8
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
     9
TimerArray = PPROD +
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
    10
8251
9be357df93d4 New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents: 7537
diff changeset
    11
types 'a state = "nat * 'a"   (*second component allows new variables*)
9be357df93d4 New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents: 7537
diff changeset
    12
7513
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
    13
constdefs
8251
9be357df93d4 New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents: 7537
diff changeset
    14
  count  :: "'a state => nat"
9be357df93d4 New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents: 7537
diff changeset
    15
    "count s == fst s"
7537
875754b599df working snapshot
paulson
parents: 7513
diff changeset
    16
  
8251
9be357df93d4 New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents: 7537
diff changeset
    17
  decr  :: "('a state * 'a state) set"
9be357df93d4 New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents: 7537
diff changeset
    18
    "decr == UN n uu. {((Suc n, uu), (n,uu))}"
9be357df93d4 New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents: 7537
diff changeset
    19
  
9be357df93d4 New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents: 7537
diff changeset
    20
  Timer :: 'a state program
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 8251
diff changeset
    21
    "Timer == mk_program (UNIV, {decr}, UNIV)"
7513
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
    22
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
    23
end