| 6008 |      1 | (*  Title:      HOL/IOA/example/Spec.thy
 | 
|  |      2 |     ID:         $Id$
 | 
| 12218 |      3 |     Author:     Olaf Müller
 | 
| 6008 |      4 | *)
 | 
|  |      5 | 
 | 
| 17244 |      6 | header {* The implementation of a memory *}
 | 
| 6008 |      7 | 
 | 
| 17244 |      8 | theory Impl
 | 
|  |      9 | imports IOA Action
 | 
|  |     10 | begin
 | 
| 6008 |     11 | 
 | 
| 27361 |     12 | definition
 | 
|  |     13 |   impl_sig :: "action signature" where
 | 
|  |     14 |   "impl_sig = (UN l.{Free l} Un {New},
 | 
|  |     15 |                UN l.{Loc l},
 | 
|  |     16 |                {})"
 | 
| 6008 |     17 | 
 | 
| 27361 |     18 | definition
 | 
|  |     19 |   impl_trans :: "(action, nat  * bool)transition set" where
 | 
|  |     20 |   "impl_trans =
 | 
|  |     21 |     {tr. let s = fst(tr); k = fst s; b = snd s;
 | 
|  |     22 |              t = snd(snd(tr)); k' = fst t; b' = snd t
 | 
|  |     23 |          in
 | 
|  |     24 |          case fst(snd(tr))
 | 
|  |     25 |          of
 | 
|  |     26 |          New       => k' = k & b'  |
 | 
|  |     27 |          Loc l     => b & l= k & k'= (Suc k) & ~b' |
 | 
|  |     28 |          Free l    => k'=k & b'=b}"
 | 
| 6008 |     29 | 
 | 
| 27361 |     30 | definition
 | 
|  |     31 |   impl_ioa :: "(action, nat * bool)ioa" where
 | 
|  |     32 |   "impl_ioa = (impl_sig, {(0,False)}, impl_trans,{},{})"
 | 
| 17244 |     33 | 
 | 
|  |     34 | lemma in_impl_asig:
 | 
|  |     35 |   "New : actions(impl_sig) &
 | 
|  |     36 |     Loc l : actions(impl_sig) &
 | 
|  |     37 |     Free l : actions(impl_sig) "
 | 
| 27361 |     38 |   by (simp add: impl_sig_def actions_def asig_projections)
 | 
| 17244 |     39 | 
 | 
| 6008 |     40 | end
 |