src/HOL/UNITY/AllocImpl.thy
author paulson
Thu, 22 Jun 2000 11:35:28 +0200
changeset 9105 d9257992b845
parent 9027 daeccd9f885f
child 9403 aad13b59b8d9
permissions -rw-r--r--
fixed the merge spec (NbT -> Nclients) and added the distributor spec
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
9027
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/AllocImpl
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
     2
    ID:         $Id$
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
     4
    Copyright   1998  University of Cambridge
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
     5
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
     6
Implementation of a multiple-client allocator from a single-client allocator
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
     7
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
     8
add_path "../Induct";
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
     9
with_path "../Induct" time_use_thy "AllocImpl";
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    10
*)
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    11
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    12
AllocImpl = AllocBase + Follows + PPROD + 
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    13
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    14
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    15
(** State definitions.  OUTPUT variables are locals **)
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    16
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    17
(*Type variable 'b is the type of items being merged*)
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    18
record 'b merge =
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    19
  In   :: nat => 'b list   (*merge's INPUT histories: streams to merge*)
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    20
  Out  :: 'b list          (*merge's OUTPUT history: merged items*)
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    21
  iOut :: nat list         (*merge's OUTPUT history: origins of merged items*)
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    22
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    23
record ('a,'b) merge_u =
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    24
  'b merge +
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    25
  extra :: 'a       (*dummy field for new variables*)
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    26
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    27
constdefs
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    28
  non_extra :: ('a,'b) merge_u => 'b merge
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    29
    "non_extra s == (|In = In s, Out = Out s, iOut = iOut s|)"
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    30
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    31
record 'b distr =
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    32
  In  :: 'b list          (*items to distribute*)
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    33
  iIn :: nat list         (*destinations of items to distribute*)
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    34
  Out :: nat => 'b list   (*distributed items*)
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    35
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    36
record ('a,'b) distr_u =
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    37
  'b distr +
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    38
  extra :: 'a       (*dummy field for new variables*)
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    39
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    40
record allocState =
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    41
  giv :: nat list   (*OUTPUT history: source of tokens*)
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    42
  ask :: nat list   (*INPUT: tokens requested from allocator*)
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    43
  rel :: nat list   (*INPUT: tokens released to allocator*)
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    44
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    45
record 'a allocState_u =
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    46
  allocState +
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    47
  extra    :: 'a                (*dummy field for new variables*)
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    48
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    49
record 'a systemState =
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    50
  allocState +
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    51
  mergeRel :: nat merge
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    52
  mergeAsk :: nat merge
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    53
  distr    :: nat distr
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    54
  extra    :: 'a                  (*dummy field for new variables*)
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    55
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    56
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    57
constdefs
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    58
9105
d9257992b845 fixed the merge spec (NbT -> Nclients) and added the distributor spec
paulson
parents: 9027
diff changeset
    59
(** Merge specification (the number of inputs is Nclients) ***)
9027
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    60
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    61
  (*spec (10)*)
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    62
  merge_increasing :: ('a,'b) merge_u program set
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    63
    "merge_increasing ==
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    64
         UNIV guarantees[funPair merge.Out merge.iOut]
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    65
         (Increasing merge.Out) Int (Increasing merge.iOut)"
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    66
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    67
  (*spec (11)*)
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    68
  merge_eqOut :: ('a,'b) merge_u program set
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    69
    "merge_eqOut ==
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    70
         UNIV guarantees[funPair merge.Out merge.iOut]
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    71
         Always {s. length (merge.Out s) = length (merge.iOut s)}"
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    72
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    73
  (*spec (12)*)
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    74
  merge_bounded :: ('a,'b) merge_u program set
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    75
    "merge_bounded ==
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    76
         UNIV guarantees[merge.iOut]
9105
d9257992b845 fixed the merge spec (NbT -> Nclients) and added the distributor spec
paulson
parents: 9027
diff changeset
    77
         Always {s. ALL elt : set (merge.iOut s). elt < Nclients}"
9027
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    78
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    79
  (*spec (13)*)
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    80
  merge_follows :: ('a,'b) merge_u program set
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    81
    "merge_follows ==
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    82
	 (INT i : lessThan Nclients. Increasing (sub i o merge.In))
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    83
	 guarantees[funPair merge.Out merge.iOut]
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    84
	 (INT i : lessThan Nclients. 
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    85
	  (%s. sublist (merge.Out s) 
9105
d9257992b845 fixed the merge spec (NbT -> Nclients) and added the distributor spec
paulson
parents: 9027
diff changeset
    86
                       {k. k < size(merge.iOut s) & merge.iOut s! k = i})
9027
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    87
	  Fols (sub i o merge.In))"
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    88
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    89
  (*spec: preserves part*)
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    90
    merge_preserves :: ('a,'b) merge_u program set
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    91
    "merge_preserves == preserves (funPair merge.In merge_u.extra)"
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    92
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    93
  merge_spec :: ('a,'b) merge_u program set
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    94
    "merge_spec == merge_increasing Int merge_eqOut Int merge_bounded Int
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    95
                   merge_follows Int merge_preserves"
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
    96
9105
d9257992b845 fixed the merge spec (NbT -> Nclients) and added the distributor spec
paulson
parents: 9027
diff changeset
    97
(** Distributor specification (the number of outputs is Nclients) ***)
d9257992b845 fixed the merge spec (NbT -> Nclients) and added the distributor spec
paulson
parents: 9027
diff changeset
    98
d9257992b845 fixed the merge spec (NbT -> Nclients) and added the distributor spec
paulson
parents: 9027
diff changeset
    99
  (*spec (14)*)
d9257992b845 fixed the merge spec (NbT -> Nclients) and added the distributor spec
paulson
parents: 9027
diff changeset
   100
  distr_follows :: ('a,'b) distr_u program set
d9257992b845 fixed the merge spec (NbT -> Nclients) and added the distributor spec
paulson
parents: 9027
diff changeset
   101
    "distr_follows ==
d9257992b845 fixed the merge spec (NbT -> Nclients) and added the distributor spec
paulson
parents: 9027
diff changeset
   102
	 Increasing distr.In Int Increasing distr.iIn Int
d9257992b845 fixed the merge spec (NbT -> Nclients) and added the distributor spec
paulson
parents: 9027
diff changeset
   103
	 Always {s. ALL elt : set (distr.iIn s). elt < Nclients}
d9257992b845 fixed the merge spec (NbT -> Nclients) and added the distributor spec
paulson
parents: 9027
diff changeset
   104
	 guarantees[distr.Out]
d9257992b845 fixed the merge spec (NbT -> Nclients) and added the distributor spec
paulson
parents: 9027
diff changeset
   105
	 (INT i : lessThan Nclients. 
d9257992b845 fixed the merge spec (NbT -> Nclients) and added the distributor spec
paulson
parents: 9027
diff changeset
   106
	  (sub i o distr.Out) Fols
d9257992b845 fixed the merge spec (NbT -> Nclients) and added the distributor spec
paulson
parents: 9027
diff changeset
   107
	  (%s. sublist (distr.In s) 
d9257992b845 fixed the merge spec (NbT -> Nclients) and added the distributor spec
paulson
parents: 9027
diff changeset
   108
                       {k. k < size(distr.iIn s) & distr.iIn s ! k = i}))"
d9257992b845 fixed the merge spec (NbT -> Nclients) and added the distributor spec
paulson
parents: 9027
diff changeset
   109
d9257992b845 fixed the merge spec (NbT -> Nclients) and added the distributor spec
paulson
parents: 9027
diff changeset
   110
9027
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   111
(** Single-client allocator specification (required) ***)
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   112
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   113
  (*spec (18)*)
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   114
  alloc_increasing :: 'a allocState_u program set
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   115
    "alloc_increasing == UNIV guarantees[giv] Increasing giv"
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   116
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   117
  (*spec (19)*)
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   118
  alloc_safety :: 'a allocState_u program set
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   119
    "alloc_safety ==
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   120
	 Increasing rel
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   121
         guarantees[giv] Always {s. tokens (giv s) <= NbT + tokens (rel s)}"
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   122
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   123
  (*spec (20)*)
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   124
  alloc_progress :: 'a allocState_u program set
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   125
    "alloc_progress ==
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   126
	 Increasing ask Int Increasing rel Int
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   127
         Always {s. ALL elt : set (ask s). elt <= NbT}
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   128
         Int
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   129
         (INT h. {s. h <= giv s & h pfixGe (ask s)}
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   130
		 LeadsTo
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   131
	         {s. tokens h <= tokens (rel s)})
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   132
         guarantees[giv]
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   133
	     (INT h. {s. h <= ask s} LeadsTo {s. h pfixLe giv s})"
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   134
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   135
  (*spec: preserves part*)
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   136
    alloc_preserves :: 'a allocState_u program set
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   137
    "alloc_preserves == preserves (funPair rel
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   138
				   (funPair ask allocState_u.extra))"
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   139
  
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   140
  alloc_spec :: 'a allocState_u program set
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   141
    "alloc_spec == alloc_increasing Int alloc_safety Int alloc_progress Int
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   142
                   alloc_preserves"
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   143
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   144
(****
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   145
    (** Network specification ***)
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   146
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   147
      (*spec (9.1)*)
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   148
      network_ask :: 'a systemState program set
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   149
	"network_ask == INT i : lessThan Nclients.
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   150
			    Increasing (ask o sub i o client)
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   151
			    guarantees[ask]
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   152
			    (ask  Fols (ask o sub i o client))"
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   153
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   154
      (*spec (9.2)*)
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   155
      network_giv :: 'a systemState program set
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   156
	"network_giv == INT i : lessThan Nclients.
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   157
			    Increasing giv 
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   158
			    guarantees[giv o sub i o client]
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   159
			    ((giv o sub i o client) Fols giv )"
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   160
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   161
      (*spec (9.3)*)
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   162
      network_rel :: 'a systemState program set
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   163
	"network_rel == INT i : lessThan Nclients.
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   164
			    Increasing (rel o sub i o client)
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   165
			    guarantees[rel]
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   166
			    (rel  Fols (rel o sub i o client))"
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   167
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   168
      (*spec: preserves part*)
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   169
	network_preserves :: 'a systemState program set
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   170
	"network_preserves == preserves giv  Int
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   171
			      (INT i : lessThan Nclients.
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   172
			       preserves (funPair rel ask o sub i o client))"
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   173
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   174
      network_spec :: 'a systemState program set
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   175
	"network_spec == network_ask Int network_giv Int
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   176
			 network_rel Int network_preserves"
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   177
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   178
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   179
    (** State mappings **)
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   180
      sysOfAlloc :: "((nat => merge) * 'a) allocState_u => 'a systemState"
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   181
	"sysOfAlloc == %s. let (cl,xtr) = allocState_u.extra s
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   182
			   in (| giv = giv s,
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   183
				 ask = ask s,
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   184
				 rel = rel s,
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   185
				 client   = cl,
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   186
				 extra    = xtr|)"
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   187
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   188
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   189
      sysOfClient :: "(nat => merge) * 'a allocState_u => 'a systemState"
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   190
	"sysOfClient == %(cl,al). (| giv = giv al,
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   191
				     ask = ask al,
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   192
				     rel = rel al,
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   193
				     client   = cl,
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   194
				     systemState.extra = allocState_u.extra al|)"
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   195
****)
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   196
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   197
consts 
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   198
    Alloc  :: 'a allocState_u program
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   199
    Merge  :: ('a,'b) merge_u program
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   200
(*    
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   201
    Network :: 'a systemState program
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   202
    System  :: 'a systemState program
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   203
  *)
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   204
  
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   205
rules
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   206
    Alloc   "Alloc   : alloc_spec"
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   207
    Merge  "Merge  : merge_spec"
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   208
(**    Network "Network : network_spec"**)
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   209
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   210
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   211
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   212
(**
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   213
defs
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   214
    System_def
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   215
      "System == rename sysOfAlloc Alloc Join Network Join
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   216
                 (rename sysOfMerge
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   217
		  (plam x: lessThan Nclients. rename merge_map Merge))"
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   218
**)
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   219
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   220
daeccd9f885f The Allocator Implementation (not yet working)
paulson
parents:
diff changeset
   221
end