| 6008 |      1 | (*  Title:      HOL/IOA/example/Spec.thy
 | 
|  |      2 |     ID:         $Id$
 | 
| 12218 |      3 |     Author:     Olaf Müller
 | 
| 6008 |      4 | *)
 | 
|  |      5 | 
 | 
| 17244 |      6 | header {* The specification of a memory *}
 | 
| 6008 |      7 | 
 | 
| 17244 |      8 | theory Spec
 | 
|  |      9 | imports IOA Action
 | 
|  |     10 | begin
 | 
| 6008 |     11 | 
 | 
| 27361 |     12 | definition
 | 
|  |     13 |   spec_sig :: "action signature" where
 | 
|  |     14 |   "spec_sig = (UN l.{Free l} Un {New},
 | 
|  |     15 |                UN l.{Loc l},
 | 
|  |     16 |                {})"
 | 
| 6008 |     17 | 
 | 
| 27361 |     18 | definition
 | 
|  |     19 |   spec_trans :: "(action, nat set * bool)transition set" where
 | 
|  |     20 |   "spec_trans =
 | 
|  |     21 |    {tr. let s = fst(tr); used = fst s; c = snd s;
 | 
|  |     22 |             t = snd(snd(tr)); used' = fst t; c' = snd t
 | 
|  |     23 |         in
 | 
|  |     24 |         case fst(snd(tr))
 | 
|  |     25 |         of
 | 
|  |     26 |         New       => used' = used & c'  |
 | 
|  |     27 |         Loc l     => c & l~:used  & used'= used Un {l} & ~c'   |
 | 
|  |     28 |         Free l    => used'=used - {l} & c'=c}"
 | 
| 6008 |     29 | 
 | 
| 27361 |     30 | definition
 | 
|  |     31 |   spec_ioa :: "(action, nat set * bool)ioa" where
 | 
|  |     32 |   "spec_ioa = (spec_sig, {({},False)}, spec_trans,{},{})"
 | 
| 17244 |     33 | 
 | 
| 6008 |     34 | end
 |