author | haftmann |
Fri, 27 Jun 2025 08:09:26 +0200 | |
changeset 82775 | 61c39a9e5415 |
parent 67613 | ce654b0e6d69 |
permissions | -rw-r--r-- |
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 |
66453
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
62009
diff
changeset
|
8 |
imports IOA.IOA Action |
17244 | 9 |
begin |
6008 | 10 |
|
27361 | 11 |
definition |
12 |
impl_sig :: "action signature" where |
|
67613 | 13 |
"impl_sig = (\<Union>l.{Free l} \<union> {New}, |
14 |
\<Union>l.{Loc l}, |
|
27361 | 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 |
|
67613 | 25 |
New \<Rightarrow> k' = k \<and> b' | |
26 |
Loc l \<Rightarrow> b \<and> l= k \<and> k'= (Suc k) \<and> \<not>b' | |
|
27 |
Free l \<Rightarrow> k'=k \<and> 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: |
|
67613 | 34 |
"New \<in> actions(impl_sig) \<and> |
35 |
Loc l \<in> actions(impl_sig) \<and> |
|
36 |
Free l \<in> actions(impl_sig) " |
|
27361 | 37 |
by (simp add: impl_sig_def actions_def asig_projections) |
17244 | 38 |
|
6008 | 39 |
end |