src/HOL/UNITY/Alloc.thy
author wenzelm
Wed, 30 Jun 1999 12:22:31 +0200
changeset 6857 6e6eb8d92377
parent 6841 5a557122bb62
child 7188 2bc63a44721b
permissions -rw-r--r--
added sync marker;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
6815
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/Alloc
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
     2
    ID:         $Id$
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
     4
    Copyright   1998  University of Cambridge
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
     5
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
     6
Specification of Chandy and Charpentier's Allocator
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
     7
*)
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
     8
6827
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
     9
Alloc = Follows + Extend + PPROD +
6815
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    10
6837
1bd850260747 another snapshot, still not working
paulson
parents: 6827
diff changeset
    11
constdefs
1bd850260747 another snapshot, still not working
paulson
parents: 6827
diff changeset
    12
  (**TOO STRONG: DELETE**)
1bd850260747 another snapshot, still not working
paulson
parents: 6827
diff changeset
    13
  modular :: "['a, ('a=>'b) program set] => bool"
1bd850260747 another snapshot, still not working
paulson
parents: 6827
diff changeset
    14
    "modular i X ==
1bd850260747 another snapshot, still not working
paulson
parents: 6827
diff changeset
    15
       ALL F G. F : X --> drop_prog i F = drop_prog i G --> G : X"
1bd850260747 another snapshot, still not working
paulson
parents: 6827
diff changeset
    16
1bd850260747 another snapshot, still not working
paulson
parents: 6827
diff changeset
    17
6841
5a557122bb62 renamed PPI to plam
paulson
parents: 6837
diff changeset
    18
 (*UNITY.thy????????????????*)
6837
1bd850260747 another snapshot, still not working
paulson
parents: 6827
diff changeset
    19
  (*An Idle program can do nothing*)
1bd850260747 another snapshot, still not working
paulson
parents: 6827
diff changeset
    20
  Idle :: 'a program set
1bd850260747 another snapshot, still not working
paulson
parents: 6827
diff changeset
    21
    "Idle == {F. Union (Acts F) <= Id}"
1bd850260747 another snapshot, still not working
paulson
parents: 6827
diff changeset
    22
6815
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    23
(*simplifies the expression of some specifications*)
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    24
constdefs
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    25
  sub :: ['a, 'a=>'b] => 'b
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    26
    "sub i f == f i"
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    27
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    28
(*Duplicated from HOL/ex/NatSum.thy.
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    29
  Maybe type should be  [nat=>int, nat] => int**)
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    30
consts sum     :: [nat=>nat, nat] => nat
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    31
primrec 
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    32
  "sum f 0 = 0"
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    33
  "sum f (Suc n) = f(n) + sum f n"
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    34
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    35
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    36
(*This function merely sums the elements of a list*)
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    37
consts tokens     :: nat list => nat
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    38
primrec 
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    39
  "tokens [] = 0"
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    40
  "tokens (x#xs) = x + tokens xs"
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    41
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    42
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    43
consts
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    44
  NbT      :: nat       (*Number of tokens in system*)
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    45
  Nclients :: nat       (*Number of clients*)
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    46
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    47
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    48
record clientState =
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    49
  giv :: nat list   (*client's INPUT history:  tokens GRANTED*)
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    50
  ask :: nat list   (*client's OUTPUT history: tokens REQUESTED*)
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    51
  rel :: nat list   (*client's OUTPUT history: tokens RELEASED*)
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    52
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    53
record allocState =
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    54
  allocGiv :: nat => nat list   (*allocator's local copy of "giv" for i*)
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    55
  allocAsk :: nat => nat list   (*allocator's local copy of "ask" for i*)
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    56
  allocRel :: nat => nat list   (*allocator's local copy of "rel" for i*)
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    57
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    58
record systemState = allocState +
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    59
  client :: nat => clientState  (*states of all clients*)
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    60
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    61
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    62
constdefs
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    63
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    64
(** Resource allocation system specification **)
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    65
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    66
  (*spec (1)*)
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    67
  system_safety :: systemState program set
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    68
    "system_safety ==
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    69
     Always {s. sum (%i. (tokens o giv o sub i o client) s) Nclients
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    70
	        <= NbT + sum (%i. (tokens o rel o sub i o client) s) Nclients}"
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    71
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    72
  (*spec (2)*)
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    73
  system_progress :: systemState program set
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    74
    "system_progress == INT i : lessThan Nclients.
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    75
			INT h. 
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    76
			  {s. h <= (ask o sub i o client)s} LeadsTo
6827
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
    77
			  {s. h pfixLe (giv o sub i o client) s}"
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
    78
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
    79
  system_spec :: systemState program set
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
    80
    "system_spec == system_safety Int system_progress"
6815
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    81
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    82
(** Client specification (required) ***)
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    83
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    84
  (*spec (3)*)
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    85
  client_increasing :: clientState program set
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    86
    "client_increasing ==
6827
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
    87
         UNIV guar Increasing ask Int Increasing rel"
6815
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    88
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    89
  (*spec (4)*)
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    90
  client_bounded :: clientState program set
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    91
    "client_bounded ==
6827
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
    92
         UNIV guar Always {s. ALL elt : set (ask s). elt <= NbT}"
6815
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    93
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    94
  (*spec (5)*)
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    95
  client_progress :: clientState program set
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    96
    "client_progress ==
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
    97
	 Increasing giv
6827
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
    98
	 guar
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
    99
	 (INT h. {s. h <= giv s & h pfixGe ask s}
6815
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   100
		 LeadsTo
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   101
		 {s. tokens h <= (tokens o rel) s})"
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   102
6827
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   103
  client_spec :: clientState program set
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   104
    "client_spec == client_increasing Int client_bounded Int client_progress"
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   105
6815
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   106
(** Allocator specification (required) ***)
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   107
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   108
  (*spec (6)*)
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   109
  alloc_increasing :: allocState program set
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   110
    "alloc_increasing ==
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   111
	 UNIV
6827
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   112
         guar
6815
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   113
	 (INT i : lessThan Nclients. Increasing (sub i o allocAsk))"
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   114
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   115
  (*spec (7)*)
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   116
  alloc_safety :: allocState program set
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   117
    "alloc_safety ==
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   118
	 (INT i : lessThan Nclients. Increasing (sub i o allocRel))
6827
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   119
         guar
6815
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   120
	 Always {s. sum (%i. (tokens o sub i o allocGiv) s) Nclients
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   121
	         <= NbT + sum (%i. (tokens o sub i o allocRel) s) Nclients}"
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   122
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   123
  (*spec (8)*)
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   124
  alloc_progress :: allocState program set
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   125
    "alloc_progress ==
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   126
	 (INT i : lessThan Nclients. Increasing (sub i o allocAsk) Int
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   127
	                             Increasing (sub i o allocRel))
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   128
         Int
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   129
         Always {s. ALL i : lessThan Nclients.
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   130
		    ALL k : lessThan (length (allocAsk s i)).
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   131
		    allocAsk s i ! k <= NbT}
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   132
         Int
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   133
         (INT i : lessThan Nclients. 
6827
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   134
	  INT h. {s. h <= (sub i o allocGiv)s & h pfixGe (sub i o allocAsk)s}
6815
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   135
		  LeadsTo {s. tokens h <= (tokens o sub i o allocRel)s})
6827
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   136
         guar
6815
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   137
	     (INT i : lessThan Nclients.
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   138
	      INT h. {s. h <= (sub i o allocAsk) s} LeadsTo
6827
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   139
	             {s. h pfixLe (sub i o allocGiv) s})"
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   140
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   141
  alloc_spec :: allocState program set
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   142
    "alloc_spec == alloc_increasing Int alloc_safety Int alloc_progress"
6815
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   143
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   144
(** Network specification ***)
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   145
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   146
  (*spec (9.1)*)
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   147
  network_ask :: systemState program set
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   148
    "network_ask == INT i : lessThan Nclients.
6827
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   149
                    Increasing (ask o sub i o client)
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   150
                    guar
6815
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   151
                    ((sub i o allocAsk) Fols (ask o sub i o client))"
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   152
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   153
  (*spec (9.2)*)
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   154
  network_giv :: systemState program set
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   155
    "network_giv == INT i : lessThan Nclients.
6827
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   156
                    Increasing (sub i o allocGiv)
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   157
                    guar
6815
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   158
                    ((giv o sub i o client) Fols (sub i o allocGiv))"
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   159
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   160
  (*spec (9.3)*)
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   161
  network_rel :: systemState program set
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   162
    "network_rel == INT i : lessThan Nclients.
6827
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   163
                    Increasing (rel o sub i o client)
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   164
                    guar
6815
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   165
                    ((sub i o allocRel) Fols (rel o sub i o client))"
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   166
6827
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   167
  network_spec :: systemState program set
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   168
    "network_spec == network_ask Int network_giv Int network_rel"
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   169
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   170
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   171
(** State mappings **)
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   172
  sysOfAlloc :: "(allocState * (nat => clientState)) => systemState"
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   173
    "sysOfAlloc == %(s,y). (| allocGiv = allocGiv s,
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   174
			      allocAsk = allocAsk s,
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   175
			      allocRel = allocRel s,
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   176
			      client   = y |)"
6841
5a557122bb62 renamed PPI to plam
paulson
parents: 6837
diff changeset
   177
(***    "sysOfAlloc == %(s,y). s(|client := y|)"  TYPE DOESN'T CHANGE***)
5a557122bb62 renamed PPI to plam
paulson
parents: 6837
diff changeset
   178
6827
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   179
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   180
  sysOfClient :: "((nat => clientState) * allocState ) => systemState"
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   181
    "sysOfClient == %(x,y). sysOfAlloc(y,x)"
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   182
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   183
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   184
locale System =
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   185
  fixes 
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   186
    Alloc   :: allocState program
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   187
    Client  :: clientState program
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   188
    Network :: systemState program
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   189
    System  :: systemState program
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   190
  
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   191
  assumes
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   192
    Alloc   "Alloc   : alloc_spec"
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   193
    Client  "Client  : client_spec"
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   194
    Network "Network : network_spec"
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   195
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   196
  defines
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   197
    System_def
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   198
      "System == (extend sysOfAlloc Alloc)
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   199
                 Join
6841
5a557122bb62 renamed PPI to plam
paulson
parents: 6837
diff changeset
   200
                 (extend sysOfClient (plam x: lessThan Nclients. Client))
6827
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   201
                 Join
b69a2585ec0f guar; locale for the spec
paulson
parents: 6815
diff changeset
   202
                 Network"
6815
de4d358bf01e The Allocator example: specifications
paulson
parents:
diff changeset
   203
end