| 8928 |      1 | (*  Title:      HOL/UNITY/AllocBase
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
 | 
|  |      4 |     Copyright   1998  University of Cambridge
 | 
|  |      5 | 
 | 
|  |      6 | Common declarations for Chandy and Charpentier's Allocator
 | 
| 8989 |      7 | 
 | 
| 9024 |      8 | add_path "../Induct";
 | 
|  |      9 | time_use_thy "AllocBase";
 | 
| 8928 |     10 | *)
 | 
|  |     11 | 
 | 
| 9024 |     12 | AllocBase = Rename + Follows + 
 | 
| 8928 |     13 | 
 | 
|  |     14 | consts
 | 
|  |     15 |   NbT      :: nat       (*Number of tokens in system*)
 | 
|  |     16 |   Nclients :: nat       (*Number of clients*)
 | 
|  |     17 | 
 | 
|  |     18 | rules
 | 
|  |     19 |   NbT_pos  "0 < NbT"
 | 
|  |     20 | 
 | 
|  |     21 | (*This function merely sums the elements of a list*)
 | 
|  |     22 | consts tokens     :: nat list => nat
 | 
|  |     23 | primrec 
 | 
|  |     24 |   "tokens [] = 0"
 | 
|  |     25 |   "tokens (x#xs) = x + tokens xs"
 | 
|  |     26 | 
 | 
| 9024 |     27 | consts
 | 
|  |     28 |   bag_of :: 'a list => 'a multiset
 | 
|  |     29 | 
 | 
|  |     30 | primrec
 | 
|  |     31 |   "bag_of []     = {#}"
 | 
|  |     32 |   "bag_of (x#xs) = {#x#} + bag_of xs"
 | 
|  |     33 | 
 | 
| 8928 |     34 | end
 |