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