| 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 | 
 | 
|  |     12 | consts
 | 
|  |     13 | 
 | 
|  |     14 | impl_sig   :: "action signature"
 | 
|  |     15 | impl_trans :: "(action, nat  * bool)transition set"
 | 
|  |     16 | impl_ioa   :: "(action, nat * bool)ioa"
 | 
|  |     17 | 
 | 
|  |     18 | defs
 | 
|  |     19 | 
 | 
| 17244 |     20 | sig_def: "impl_sig == (UN l.{Free l} Un {New},
 | 
|  |     21 |                      UN l.{Loc l},
 | 
| 6008 |     22 |                      {})"
 | 
|  |     23 | 
 | 
| 17244 |     24 | trans_def: "impl_trans ==
 | 
|  |     25 |  {tr. let s = fst(tr); k = fst s; b = snd s;
 | 
|  |     26 |           t = snd(snd(tr)); k' = fst t; b' = snd t
 | 
|  |     27 |       in
 | 
|  |     28 |       case fst(snd(tr))
 | 
|  |     29 |       of
 | 
|  |     30 |       New       => k' = k & b'  |
 | 
|  |     31 |       Loc l     => b & l= k & k'= (Suc k) & ~b' |
 | 
| 6008 |     32 |       Free l    => k'=k & b'=b}"
 | 
|  |     33 | 
 | 
| 17244 |     34 | ioa_def: "impl_ioa == (impl_sig, {(0,False)}, impl_trans,{},{})"
 | 
|  |     35 | 
 | 
|  |     36 | lemma in_impl_asig:
 | 
|  |     37 |   "New : actions(impl_sig) &
 | 
|  |     38 |     Loc l : actions(impl_sig) &
 | 
|  |     39 |     Free l : actions(impl_sig) "
 | 
|  |     40 |   by (simp add: Impl.sig_def actions_def asig_projections)
 | 
|  |     41 | 
 | 
|  |     42 | ML {* use_legacy_bindings (the_context ()) *}
 | 
| 6008 |     43 | 
 | 
|  |     44 | end
 |