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