src/HOL/UNITY/AllocBase.thy
author paulson
Wed, 14 Feb 2001 12:16:32 +0100
changeset 11115 285b31e9e026
parent 10265 4e004b548049
permissions -rw-r--r--
a new theorem from Bryan Ford
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
10265
4e004b548049 use Multiset from HOL/Library;
wenzelm
parents: 9336
diff changeset
     1
(*  Title:      HOL/UNITY/AllocBase.thy
8928
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
     2
    ID:         $Id$
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
     4
    Copyright   1998  University of Cambridge
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
     5
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
     6
Common declarations for Chandy and Charpentier's Allocator
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
     7
*)
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
     8
9024
b1f37f6819c4 new constant bag_of
paulson
parents: 8989
diff changeset
     9
AllocBase = Rename + Follows + 
8928
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    10
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    11
consts
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    12
  NbT      :: nat       (*Number of tokens in system*)
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    13
  Nclients :: nat       (*Number of clients*)
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    14
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    15
rules
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    16
  NbT_pos  "0 < NbT"
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    17
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    18
(*This function merely sums the elements of a list*)
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    19
consts tokens     :: nat list => nat
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    20
primrec 
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    21
  "tokens [] = 0"
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    22
  "tokens (x#xs) = x + tokens xs"
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    23
9024
b1f37f6819c4 new constant bag_of
paulson
parents: 8989
diff changeset
    24
consts
b1f37f6819c4 new constant bag_of
paulson
parents: 8989
diff changeset
    25
  bag_of :: 'a list => 'a multiset
b1f37f6819c4 new constant bag_of
paulson
parents: 8989
diff changeset
    26
b1f37f6819c4 new constant bag_of
paulson
parents: 8989
diff changeset
    27
primrec
b1f37f6819c4 new constant bag_of
paulson
parents: 8989
diff changeset
    28
  "bag_of []     = {#}"
b1f37f6819c4 new constant bag_of
paulson
parents: 8989
diff changeset
    29
  "bag_of (x#xs) = {#x#} + bag_of xs"
b1f37f6819c4 new constant bag_of
paulson
parents: 8989
diff changeset
    30
8928
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    31
end