| author | blanchet | 
| Fri, 25 Mar 2022 13:52:23 +0100 | |
| changeset 75341 | 72cbbb4d98f3 | 
| parent 67613 | ce654b0e6d69 | 
| permissions | -rw-r--r-- | 
| 42151 | 1 | (* Title: HOL/HOLCF/IOA/Storage/Spec.thy | 
| 40945 | 2 | Author: Olaf Müller | 
| 6008 | 3 | *) | 
| 4 | ||
| 62002 | 5 | section \<open>The specification of a memory\<close> | 
| 6008 | 6 | |
| 17244 | 7 | theory Spec | 
| 66453 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 wenzelm parents: 
62009diff
changeset | 8 | imports IOA.IOA Action | 
| 17244 | 9 | begin | 
| 6008 | 10 | |
| 27361 | 11 | definition | 
| 12 | spec_sig :: "action signature" where | |
| 67613 | 13 |   "spec_sig = (\<Union>l.{Free l} \<union> {New},
 | 
| 14 |                \<Union>l.{Loc l},
 | |
| 27361 | 15 |                {})"
 | 
| 6008 | 16 | |
| 27361 | 17 | definition | 
| 67613 | 18 | spec_trans :: "(action, nat set \<times> bool)transition set" where | 
| 27361 | 19 | "spec_trans = | 
| 20 |    {tr. let s = fst(tr); used = fst s; c = snd s;
 | |
| 21 | t = snd(snd(tr)); used' = fst t; c' = snd t | |
| 22 | in | |
| 23 | case fst(snd(tr)) | |
| 24 | of | |
| 67613 | 25 | New \<Rightarrow> used' = used \<and> c' | | 
| 26 |         Loc l     \<Rightarrow> c \<and> l \<notin> used \<and> used'= used \<union> {l} \<and> \<not>c'   |
 | |
| 27 |         Free l    \<Rightarrow> used'=used - {l} \<and> c'=c}"
 | |
| 6008 | 28 | |
| 27361 | 29 | definition | 
| 67613 | 30 | spec_ioa :: "(action, nat set \<times> bool)ioa" where | 
| 27361 | 31 |   "spec_ioa = (spec_sig, {({},False)}, spec_trans,{},{})"
 | 
| 17244 | 32 | |
| 6008 | 33 | end |