author | wenzelm |
Sat, 16 Jan 2016 23:31:28 +0100 | |
changeset 62193 | 0826f6b6ba09 |
parent 62116 | bc178c0fe1a1 |
child 69597 | ff784d5a5bfb |
permissions | -rw-r--r-- |
62008 | 1 |
(* Title: HOL/HOLCF/IOA/Automata.thy |
40945 | 2 |
Author: Olaf Müller, Konrad Slind, Tobias Nipkow |
17233 | 3 |
*) |
3071 | 4 |
|
62002 | 5 |
section \<open>The I/O automata of Lynch and Tuttle in HOLCF\<close> |
3071 | 6 |
|
17233 | 7 |
theory Automata |
8 |
imports Asig |
|
9 |
begin |
|
10 |
||
36452 | 11 |
default_sort type |
17233 | 12 |
|
62116 | 13 |
type_synonym ('a, 's) transition = "'s \<times> 'a \<times> 's" |
62005 | 14 |
type_synonym ('a, 's) ioa = |
62116 | 15 |
"'a signature \<times> 's set \<times> ('a, 's)transition set \<times> 'a set set \<times> 'a set set" |
3071 | 16 |
|
17 |
||
62116 | 18 |
subsection \<open>IO automata\<close> |
62005 | 19 |
|
62116 | 20 |
definition asig_of :: "('a, 's) ioa \<Rightarrow> 'a signature" |
62005 | 21 |
where "asig_of = fst" |
22 |
||
23 |
definition starts_of :: "('a, 's) ioa \<Rightarrow> 's set" |
|
62116 | 24 |
where "starts_of = fst \<circ> snd" |
62005 | 25 |
|
26 |
definition trans_of :: "('a, 's) ioa \<Rightarrow> ('a, 's) transition set" |
|
62116 | 27 |
where "trans_of = fst \<circ> snd \<circ> snd" |
62005 | 28 |
|
29 |
abbreviation trans_of_syn ("_ \<midarrow>_\<midarrow>_\<rightarrow> _" [81, 81, 81, 81] 100) |
|
30 |
where "s \<midarrow>a\<midarrow>A\<rightarrow> t \<equiv> (s, a, t) \<in> trans_of A" |
|
31 |
||
32 |
definition wfair_of :: "('a, 's) ioa \<Rightarrow> 'a set set" |
|
62116 | 33 |
where "wfair_of = fst \<circ> snd \<circ> snd \<circ> snd" |
3071 | 34 |
|
62005 | 35 |
definition sfair_of :: "('a, 's) ioa \<Rightarrow> 'a set set" |
62116 | 36 |
where "sfair_of = snd \<circ> snd \<circ> snd \<circ> snd" |
62005 | 37 |
|
38 |
definition is_asig_of :: "('a, 's) ioa \<Rightarrow> bool" |
|
39 |
where "is_asig_of A = is_asig (asig_of A)" |
|
40 |
||
41 |
definition is_starts_of :: "('a, 's) ioa \<Rightarrow> bool" |
|
42 |
where "is_starts_of A \<longleftrightarrow> starts_of A \<noteq> {}" |
|
43 |
||
44 |
definition is_trans_of :: "('a, 's) ioa \<Rightarrow> bool" |
|
45 |
where "is_trans_of A \<longleftrightarrow> |
|
46 |
(\<forall>triple. triple \<in> trans_of A \<longrightarrow> fst (snd triple) \<in> actions (asig_of A))" |
|
47 |
||
48 |
definition input_enabled :: "('a, 's) ioa \<Rightarrow> bool" |
|
49 |
where "input_enabled A \<longleftrightarrow> |
|
50 |
(\<forall>a. a \<in> inputs (asig_of A) \<longrightarrow> (\<forall>s1. \<exists>s2. (s1, a, s2) \<in> trans_of A))" |
|
51 |
||
52 |
definition IOA :: "('a, 's) ioa \<Rightarrow> bool" |
|
53 |
where "IOA A \<longleftrightarrow> |
|
54 |
is_asig_of A \<and> |
|
55 |
is_starts_of A \<and> |
|
56 |
is_trans_of A \<and> |
|
57 |
input_enabled A" |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23778
diff
changeset
|
58 |
|
62116 | 59 |
abbreviation "act A \<equiv> actions (asig_of A)" |
60 |
abbreviation "ext A \<equiv> externals (asig_of A)" |
|
61 |
abbreviation int where "int A \<equiv> internals (asig_of A)" |
|
62 |
abbreviation "inp A \<equiv> inputs (asig_of A)" |
|
63 |
abbreviation "out A \<equiv> outputs (asig_of A)" |
|
64 |
abbreviation "local A \<equiv> locals (asig_of A)" |
|
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
65 |
|
62116 | 66 |
|
67 |
text \<open>invariants\<close> |
|
68 |
||
69 |
inductive reachable :: "('a, 's) ioa \<Rightarrow> 's \<Rightarrow> bool" for C :: "('a, 's) ioa" |
|
62005 | 70 |
where |
71 |
reachable_0: "s \<in> starts_of C \<Longrightarrow> reachable C s" |
|
62116 | 72 |
| reachable_n: "reachable C s \<Longrightarrow> (s, a, t) \<in> trans_of C \<Longrightarrow> reachable C t" |
17233 | 73 |
|
62005 | 74 |
definition invariant :: "[('a, 's) ioa, 's \<Rightarrow> bool] \<Rightarrow> bool" |
75 |
where "invariant A P \<longleftrightarrow> (\<forall>s. reachable A s \<longrightarrow> P s)" |
|
3071 | 76 |
|
77 |
||
62116 | 78 |
subsection \<open>Parallel composition\<close> |
3071 | 79 |
|
62116 | 80 |
subsubsection \<open>Binary composition of action signatures and automata\<close> |
3071 | 81 |
|
62116 | 82 |
definition compatible :: "('a, 's) ioa \<Rightarrow> ('a, 't) ioa \<Rightarrow> bool" |
83 |
where "compatible A B \<longleftrightarrow> |
|
84 |
out A \<inter> out B = {} \<and> |
|
85 |
int A \<inter> act B = {} \<and> |
|
86 |
int B \<inter> act A = {}" |
|
87 |
||
88 |
definition asig_comp :: "'a signature \<Rightarrow> 'a signature \<Rightarrow> 'a signature" |
|
89 |
where "asig_comp a1 a2 = |
|
90 |
(((inputs a1 \<union> inputs a2) - (outputs a1 \<union> outputs a2), |
|
91 |
(outputs a1 \<union> outputs a2), |
|
92 |
(internals a1 \<union> internals a2)))" |
|
3071 | 93 |
|
62116 | 94 |
definition par :: "('a, 's) ioa \<Rightarrow> ('a, 't) ioa \<Rightarrow> ('a, 's * 't) ioa" (infixr "\<parallel>" 10) |
95 |
where "(A \<parallel> B) = |
|
96 |
(asig_comp (asig_of A) (asig_of B), |
|
97 |
{pr. fst pr \<in> starts_of A \<and> snd pr \<in> starts_of B}, |
|
98 |
{tr. |
|
99 |
let |
|
100 |
s = fst tr; |
|
101 |
a = fst (snd tr); |
|
102 |
t = snd (snd tr) |
|
103 |
in |
|
104 |
(a \<in> act A \<or> a \<in> act B) \<and> |
|
105 |
(if a \<in> act A then (fst s, a, fst t) \<in> trans_of A |
|
106 |
else fst t = fst s) \<and> |
|
107 |
(if a \<in> act B then (snd s, a, snd t) \<in> trans_of B |
|
108 |
else snd t = snd s)}, |
|
109 |
wfair_of A \<union> wfair_of B, |
|
110 |
sfair_of A \<union> sfair_of B)" |
|
3521 | 111 |
|
3071 | 112 |
|
62116 | 113 |
subsection \<open>Hiding\<close> |
3071 | 114 |
|
62116 | 115 |
subsubsection \<open>Hiding and restricting\<close> |
62005 | 116 |
|
62116 | 117 |
definition restrict_asig :: "'a signature \<Rightarrow> 'a set \<Rightarrow> 'a signature" |
118 |
where "restrict_asig asig actns = |
|
119 |
(inputs asig \<inter> actns, |
|
120 |
outputs asig \<inter> actns, |
|
121 |
internals asig \<union> (externals asig - actns))" |
|
3071 | 122 |
|
62116 | 123 |
text \<open> |
124 |
Notice that for \<open>wfair_of\<close> and \<open>sfair_of\<close> nothing has to be changed, as |
|
125 |
changes from the outputs to the internals does not touch the locals as a |
|
126 |
whole, which is of importance for fairness only. |
|
127 |
\<close> |
|
128 |
definition restrict :: "('a, 's) ioa \<Rightarrow> 'a set \<Rightarrow> ('a, 's) ioa" |
|
129 |
where "restrict A actns = |
|
17233 | 130 |
(restrict_asig (asig_of A) actns, |
131 |
starts_of A, |
|
3521 | 132 |
trans_of A, |
133 |
wfair_of A, |
|
134 |
sfair_of A)" |
|
135 |
||
62116 | 136 |
definition hide_asig :: "'a signature \<Rightarrow> 'a set \<Rightarrow> 'a signature" |
137 |
where "hide_asig asig actns = |
|
138 |
(inputs asig - actns, |
|
139 |
outputs asig - actns, |
|
140 |
internals asig \<union> actns)" |
|
3521 | 141 |
|
62116 | 142 |
definition hide :: "('a, 's) ioa \<Rightarrow> 'a set \<Rightarrow> ('a, 's) ioa" |
143 |
where "hide A actns = |
|
17233 | 144 |
(hide_asig (asig_of A) actns, |
145 |
starts_of A, |
|
3521 | 146 |
trans_of A, |
147 |
wfair_of A, |
|
148 |
sfair_of A)" |
|
3071 | 149 |
|
62116 | 150 |
|
151 |
subsection \<open>Renaming\<close> |
|
17233 | 152 |
|
62005 | 153 |
definition rename_set :: "'a set \<Rightarrow> ('c \<Rightarrow> 'a option) \<Rightarrow> 'c set" |
154 |
where "rename_set A ren = {b. \<exists>x. Some x = ren b \<and> x \<in> A}" |
|
3521 | 155 |
|
62005 | 156 |
definition rename :: "('a, 'b) ioa \<Rightarrow> ('c \<Rightarrow> 'a option) \<Rightarrow> ('c, 'b) ioa" |
62116 | 157 |
where "rename ioa ren = |
62005 | 158 |
((rename_set (inp ioa) ren, |
159 |
rename_set (out ioa) ren, |
|
160 |
rename_set (int ioa) ren), |
|
161 |
starts_of ioa, |
|
62116 | 162 |
{tr. |
163 |
let |
|
164 |
s = fst tr; |
|
165 |
a = fst (snd tr); |
|
166 |
t = snd (snd tr) |
|
167 |
in \<exists>x. Some x = ren a \<and> s \<midarrow>x\<midarrow>ioa\<rightarrow> t}, |
|
62005 | 168 |
{rename_set s ren | s. s \<in> wfair_of ioa}, |
169 |
{rename_set s ren | s. s \<in> sfair_of ioa})" |
|
170 |
||
3071 | 171 |
|
62116 | 172 |
subsection \<open>Fairness\<close> |
4559
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
3521
diff
changeset
|
173 |
|
62116 | 174 |
subsubsection \<open>Enabledness of actions and action sets\<close> |
62005 | 175 |
|
176 |
definition enabled :: "('a, 's) ioa \<Rightarrow> 'a \<Rightarrow> 's \<Rightarrow> bool" |
|
177 |
where "enabled A a s \<longleftrightarrow> (\<exists>t. s \<midarrow>a\<midarrow>A\<rightarrow> t)" |
|
178 |
||
179 |
definition Enabled :: "('a, 's) ioa \<Rightarrow> 'a set \<Rightarrow> 's \<Rightarrow> bool" |
|
180 |
where "Enabled A W s \<longleftrightarrow> (\<exists>w \<in> W. enabled A w s)" |
|
181 |
||
4559
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
3521
diff
changeset
|
182 |
|
62116 | 183 |
text \<open>Action set keeps enabled until probably disabled by itself.\<close> |
4559
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
3521
diff
changeset
|
184 |
|
62005 | 185 |
definition en_persistent :: "('a, 's) ioa \<Rightarrow> 'a set \<Rightarrow> bool" |
62116 | 186 |
where "en_persistent A W \<longleftrightarrow> |
62005 | 187 |
(\<forall>s a t. Enabled A W s \<and> a \<notin> W \<and> s \<midarrow>a\<midarrow>A\<rightarrow> t \<longrightarrow> Enabled A W t)" |
188 |
||
189 |
||
62116 | 190 |
text \<open>Post conditions for actions and action sets.\<close> |
4559
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
3521
diff
changeset
|
191 |
|
62005 | 192 |
definition was_enabled :: "('a, 's) ioa \<Rightarrow> 'a \<Rightarrow> 's \<Rightarrow> bool" |
193 |
where "was_enabled A a t \<longleftrightarrow> (\<exists>s. s \<midarrow>a\<midarrow>A\<rightarrow> t)" |
|
194 |
||
195 |
definition set_was_enabled :: "('a, 's) ioa \<Rightarrow> 'a set \<Rightarrow> 's \<Rightarrow> bool" |
|
196 |
where "set_was_enabled A W t \<longleftrightarrow> (\<exists>w \<in> W. was_enabled A w t)" |
|
197 |
||
198 |
||
62116 | 199 |
text \<open>Constraints for fair IOA.\<close> |
4559
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
3521
diff
changeset
|
200 |
|
62005 | 201 |
definition fairIOA :: "('a, 's) ioa \<Rightarrow> bool" |
202 |
where "fairIOA A \<longleftrightarrow> (\<forall>S \<in> wfair_of A. S \<subseteq> local A) \<and> (\<forall>S \<in> sfair_of A. S \<subseteq> local A)" |
|
4559
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
3521
diff
changeset
|
203 |
|
62005 | 204 |
definition input_resistant :: "('a, 's) ioa \<Rightarrow> bool" |
62116 | 205 |
where "input_resistant A \<longleftrightarrow> |
62005 | 206 |
(\<forall>W \<in> sfair_of A. \<forall>s a t. |
207 |
reachable A s \<and> reachable A t \<and> a \<in> inp A \<and> |
|
208 |
Enabled A W s \<and> s \<midarrow>a\<midarrow>A\<rightarrow> t \<longrightarrow> Enabled A W t)" |
|
4559
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
3521
diff
changeset
|
209 |
|
19741 | 210 |
|
211 |
declare split_paired_Ex [simp del] |
|
212 |
||
213 |
lemmas ioa_projections = asig_of_def starts_of_def trans_of_def wfair_of_def sfair_of_def |
|
214 |
||
215 |
||
62116 | 216 |
subsection "\<open>asig_of\<close>, \<open>starts_of\<close>, \<open>trans_of\<close>" |
19741 | 217 |
|
62005 | 218 |
lemma ioa_triple_proj: |
62116 | 219 |
"asig_of (x, y, z, w, s) = x \<and> |
220 |
starts_of (x, y, z, w, s) = y \<and> |
|
221 |
trans_of (x, y, z, w, s) = z \<and> |
|
222 |
wfair_of (x, y, z, w, s) = w \<and> |
|
223 |
sfair_of (x, y, z, w, s) = s" |
|
224 |
by (simp add: ioa_projections) |
|
19741 | 225 |
|
62116 | 226 |
lemma trans_in_actions: "is_trans_of A \<Longrightarrow> s1 \<midarrow>a\<midarrow>A\<rightarrow> s2 \<Longrightarrow> a \<in> act A" |
227 |
by (auto simp add: is_trans_of_def actions_def is_asig_def) |
|
19741 | 228 |
|
62116 | 229 |
lemma starts_of_par: "starts_of (A \<parallel> B) = {p. fst p \<in> starts_of A \<and> snd p \<in> starts_of B}" |
62005 | 230 |
by (simp add: par_def ioa_projections) |
19741 | 231 |
|
62005 | 232 |
lemma trans_of_par: |
62116 | 233 |
"trans_of(A \<parallel> B) = |
234 |
{tr. |
|
235 |
let |
|
236 |
s = fst tr; |
|
237 |
a = fst (snd tr); |
|
238 |
t = snd (snd tr) |
|
239 |
in |
|
240 |
(a \<in> act A \<or> a \<in> act B) \<and> |
|
241 |
(if a \<in> act A then (fst s, a, fst t) \<in> trans_of A |
|
242 |
else fst t = fst s) \<and> |
|
243 |
(if a \<in> act B then (snd s, a, snd t) \<in> trans_of B |
|
244 |
else snd t = snd s)}" |
|
62005 | 245 |
by (simp add: par_def ioa_projections) |
19741 | 246 |
|
247 |
||
62116 | 248 |
subsection \<open>\<open>actions\<close> and \<open>par\<close>\<close> |
19741 | 249 |
|
62116 | 250 |
lemma actions_asig_comp: "actions (asig_comp a b) = actions a \<union> actions b" |
62005 | 251 |
by (auto simp add: actions_def asig_comp_def asig_projections) |
252 |
||
253 |
lemma asig_of_par: "asig_of(A \<parallel> B) = asig_comp (asig_of A) (asig_of B)" |
|
254 |
by (simp add: par_def ioa_projections) |
|
255 |
||
62116 | 256 |
lemma externals_of_par: "ext (A1 \<parallel> A2) = ext A1 \<union> ext A2" |
257 |
by (auto simp add: externals_def asig_of_par asig_comp_def |
|
62005 | 258 |
asig_inputs_def asig_outputs_def Un_def set_diff_eq) |
19741 | 259 |
|
62116 | 260 |
lemma actions_of_par: "act (A1 \<parallel> A2) = act A1 \<union> act A2" |
261 |
by (auto simp add: actions_def asig_of_par asig_comp_def |
|
62005 | 262 |
asig_inputs_def asig_outputs_def asig_internals_def Un_def set_diff_eq) |
19741 | 263 |
|
62116 | 264 |
lemma inputs_of_par: "inp (A1 \<parallel> A2) = (inp A1 \<union> inp A2) - (out A1 \<union> out A2)" |
62005 | 265 |
by (simp add: actions_def asig_of_par asig_comp_def |
266 |
asig_inputs_def asig_outputs_def Un_def set_diff_eq) |
|
19741 | 267 |
|
62116 | 268 |
lemma outputs_of_par: "out (A1 \<parallel> A2) = out A1 \<union> out A2" |
62005 | 269 |
by (simp add: actions_def asig_of_par asig_comp_def |
270 |
asig_outputs_def Un_def set_diff_eq) |
|
19741 | 271 |
|
62116 | 272 |
lemma internals_of_par: "int (A1 \<parallel> A2) = int A1 \<union> int A2" |
62005 | 273 |
by (simp add: actions_def asig_of_par asig_comp_def |
274 |
asig_inputs_def asig_outputs_def asig_internals_def Un_def set_diff_eq) |
|
19741 | 275 |
|
276 |
||
62116 | 277 |
subsection \<open>Actions and compatibility\<close> |
19741 | 278 |
|
279 |
lemma compat_commute: "compatible A B = compatible B A" |
|
62005 | 280 |
by (auto simp add: compatible_def Int_commute) |
19741 | 281 |
|
62116 | 282 |
lemma ext1_is_not_int2: "compatible A1 A2 \<Longrightarrow> a \<in> ext A1 \<Longrightarrow> a \<notin> int A2" |
283 |
by (auto simp add: externals_def actions_def compatible_def) |
|
19741 | 284 |
|
62116 | 285 |
(*just commuting the previous one: better commute compatible*) |
286 |
lemma ext2_is_not_int1: "compatible A2 A1 \<Longrightarrow> a \<in> ext A1 \<Longrightarrow> a \<notin> int A2" |
|
287 |
by (auto simp add: externals_def actions_def compatible_def) |
|
19741 | 288 |
|
45606 | 289 |
lemmas ext1_ext2_is_not_act2 = ext1_is_not_int2 [THEN int_and_ext_is_act] |
290 |
lemmas ext1_ext2_is_not_act1 = ext2_is_not_int1 [THEN int_and_ext_is_act] |
|
19741 | 291 |
|
62116 | 292 |
lemma intA_is_not_extB: "compatible A B \<Longrightarrow> x \<in> int A \<Longrightarrow> x \<notin> ext B" |
293 |
by (auto simp add: externals_def actions_def compatible_def) |
|
19741 | 294 |
|
62116 | 295 |
lemma intA_is_not_actB: "compatible A B \<Longrightarrow> a \<in> int A \<Longrightarrow> a \<notin> act B" |
296 |
by (auto simp add: externals_def actions_def compatible_def is_asig_def asig_of_def) |
|
19741 | 297 |
|
62116 | 298 |
(*the only one that needs disjointness of outputs and of internals and _all_ acts*) |
299 |
lemma outAactB_is_inpB: "compatible A B \<Longrightarrow> a \<in> out A \<Longrightarrow> a \<in> act B \<Longrightarrow> a \<in> inp B" |
|
300 |
by (auto simp add: asig_outputs_def asig_internals_def actions_def asig_inputs_def |
|
62005 | 301 |
compatible_def is_asig_def asig_of_def) |
19741 | 302 |
|
62116 | 303 |
(*needed for propagation of input_enabledness from A, B to A \<parallel> B*) |
62005 | 304 |
lemma inpAAactB_is_inpBoroutB: |
62116 | 305 |
"compatible A B \<Longrightarrow> a \<in> inp A \<Longrightarrow> a \<in> act B \<Longrightarrow> a \<in> inp B \<or> a \<in> out B" |
306 |
by (auto simp add: asig_outputs_def asig_internals_def actions_def asig_inputs_def |
|
62005 | 307 |
compatible_def is_asig_def asig_of_def) |
19741 | 308 |
|
309 |
||
62116 | 310 |
subsection \<open>Input enabledness and par\<close> |
19741 | 311 |
|
62116 | 312 |
(*ugly case distinctions. Heart of proof: |
313 |
1. inpAAactB_is_inpBoroutB ie. internals are really hidden. |
|
314 |
2. inputs_of_par: outputs are no longer inputs of par. This is important here.*) |
|
62005 | 315 |
lemma input_enabled_par: |
62116 | 316 |
"compatible A B \<Longrightarrow> input_enabled A \<Longrightarrow> input_enabled B \<Longrightarrow> input_enabled (A \<parallel> B)" |
62005 | 317 |
apply (unfold input_enabled_def) |
318 |
apply (simp add: Let_def inputs_of_par trans_of_par) |
|
319 |
apply (tactic "safe_tac (Context.raw_transfer @{theory} @{theory_context Fun})") |
|
320 |
apply (simp add: inp_is_act) |
|
321 |
prefer 2 |
|
322 |
apply (simp add: inp_is_act) |
|
62116 | 323 |
text \<open>\<open>a \<in> inp A\<close>\<close> |
324 |
apply (case_tac "a \<in> act B") |
|
325 |
text \<open>\<open>a \<in> inp B\<close>\<close> |
|
62005 | 326 |
apply (erule_tac x = "a" in allE) |
327 |
apply simp |
|
328 |
apply (drule inpAAactB_is_inpBoroutB) |
|
329 |
apply assumption |
|
330 |
apply assumption |
|
331 |
apply (erule_tac x = "a" in allE) |
|
332 |
apply simp |
|
333 |
apply (erule_tac x = "aa" in allE) |
|
334 |
apply (erule_tac x = "b" in allE) |
|
335 |
apply (erule exE) |
|
336 |
apply (erule exE) |
|
62116 | 337 |
apply (rule_tac x = "(s2, s2a)" in exI) |
62005 | 338 |
apply (simp add: inp_is_act) |
62116 | 339 |
text \<open>\<open>a \<notin> act B\<close>\<close> |
62005 | 340 |
apply (simp add: inp_is_act) |
341 |
apply (erule_tac x = "a" in allE) |
|
342 |
apply simp |
|
343 |
apply (erule_tac x = "aa" in allE) |
|
344 |
apply (erule exE) |
|
345 |
apply (rule_tac x = " (s2,b) " in exI) |
|
346 |
apply simp |
|
62116 | 347 |
|
348 |
text \<open>\<open>a \<in> inp B\<close>\<close> |
|
349 |
apply (case_tac "a \<in> act A") |
|
350 |
text \<open>\<open>a \<in> act A\<close>\<close> |
|
62005 | 351 |
apply (erule_tac x = "a" in allE) |
352 |
apply (erule_tac x = "a" in allE) |
|
353 |
apply (simp add: inp_is_act) |
|
354 |
apply (frule_tac A1 = "A" in compat_commute [THEN iffD1]) |
|
355 |
apply (drule inpAAactB_is_inpBoroutB) |
|
356 |
back |
|
357 |
apply assumption |
|
358 |
apply assumption |
|
359 |
apply simp |
|
360 |
apply (erule_tac x = "aa" in allE) |
|
361 |
apply (erule_tac x = "b" in allE) |
|
362 |
apply (erule exE) |
|
363 |
apply (erule exE) |
|
62116 | 364 |
apply (rule_tac x = "(s2, s2a)" in exI) |
62005 | 365 |
apply (simp add: inp_is_act) |
62116 | 366 |
text \<open>\<open>a \<notin> act B\<close>\<close> |
62005 | 367 |
apply (simp add: inp_is_act) |
368 |
apply (erule_tac x = "a" in allE) |
|
369 |
apply (erule_tac x = "a" in allE) |
|
370 |
apply simp |
|
371 |
apply (erule_tac x = "b" in allE) |
|
372 |
apply (erule exE) |
|
62116 | 373 |
apply (rule_tac x = "(aa, s2)" in exI) |
62005 | 374 |
apply simp |
375 |
done |
|
19741 | 376 |
|
377 |
||
62116 | 378 |
subsection \<open>Invariants\<close> |
19741 | 379 |
|
380 |
lemma invariantI: |
|
62116 | 381 |
assumes "\<And>s. s \<in> starts_of A \<Longrightarrow> P s" |
382 |
and "\<And>s t a. reachable A s \<Longrightarrow> P s \<Longrightarrow> (s, a, t) \<in> trans_of A \<longrightarrow> P t" |
|
383 |
shows "invariant A P" |
|
384 |
using assms |
|
62005 | 385 |
apply (unfold invariant_def) |
386 |
apply (rule allI) |
|
387 |
apply (rule impI) |
|
388 |
apply (rule_tac x = "s" in reachable.induct) |
|
389 |
apply assumption |
|
390 |
apply blast |
|
391 |
apply blast |
|
392 |
done |
|
19741 | 393 |
|
394 |
lemma invariantI1: |
|
62116 | 395 |
assumes "\<And>s. s \<in> starts_of A \<Longrightarrow> P s" |
396 |
and "\<And>s t a. reachable A s \<Longrightarrow> P s \<longrightarrow> (s, a, t) \<in> trans_of A \<longrightarrow> P t" |
|
397 |
shows "invariant A P" |
|
398 |
using assms by (blast intro: invariantI) |
|
19741 | 399 |
|
62116 | 400 |
lemma invariantE: "invariant A P \<Longrightarrow> reachable A s \<Longrightarrow> P s" |
401 |
unfolding invariant_def by blast |
|
19741 | 402 |
|
403 |
||
62116 | 404 |
subsection \<open>\<open>restrict\<close>\<close> |
19741 | 405 |
|
406 |
lemmas reachable_0 = reachable.reachable_0 |
|
407 |
and reachable_n = reachable.reachable_n |
|
408 |
||
62116 | 409 |
lemma cancel_restrict_a: |
410 |
"starts_of (restrict ioa acts) = starts_of ioa \<and> |
|
411 |
trans_of (restrict ioa acts) = trans_of ioa" |
|
62005 | 412 |
by (simp add: restrict_def ioa_projections) |
19741 | 413 |
|
414 |
lemma cancel_restrict_b: "reachable (restrict ioa acts) s = reachable ioa s" |
|
62005 | 415 |
apply (rule iffI) |
416 |
apply (erule reachable.induct) |
|
417 |
apply (simp add: cancel_restrict_a reachable_0) |
|
418 |
apply (erule reachable_n) |
|
419 |
apply (simp add: cancel_restrict_a) |
|
62116 | 420 |
text \<open>\<open>\<longleftarrow>\<close>\<close> |
62005 | 421 |
apply (erule reachable.induct) |
422 |
apply (rule reachable_0) |
|
423 |
apply (simp add: cancel_restrict_a) |
|
424 |
apply (erule reachable_n) |
|
425 |
apply (simp add: cancel_restrict_a) |
|
426 |
done |
|
19741 | 427 |
|
428 |
lemma acts_restrict: "act (restrict A acts) = act A" |
|
62116 | 429 |
by (auto simp add: actions_def asig_internals_def |
62005 | 430 |
asig_outputs_def asig_inputs_def externals_def asig_of_def restrict_def restrict_asig_def) |
19741 | 431 |
|
62116 | 432 |
lemma cancel_restrict: |
433 |
"starts_of (restrict ioa acts) = starts_of ioa \<and> |
|
434 |
trans_of (restrict ioa acts) = trans_of ioa \<and> |
|
435 |
reachable (restrict ioa acts) s = reachable ioa s \<and> |
|
436 |
act (restrict A acts) = act A" |
|
62005 | 437 |
by (simp add: cancel_restrict_a cancel_restrict_b acts_restrict) |
19741 | 438 |
|
439 |
||
62116 | 440 |
subsection \<open>\<open>rename\<close>\<close> |
19741 | 441 |
|
62116 | 442 |
lemma trans_rename: "s \<midarrow>a\<midarrow>(rename C f)\<rightarrow> t \<Longrightarrow> (\<exists>x. Some x = f a \<and> s \<midarrow>x\<midarrow>C\<rightarrow> t)" |
62005 | 443 |
by (simp add: Let_def rename_def trans_of_def) |
19741 | 444 |
|
62116 | 445 |
lemma reachable_rename: "reachable (rename C g) s \<Longrightarrow> reachable C s" |
62005 | 446 |
apply (erule reachable.induct) |
447 |
apply (rule reachable_0) |
|
448 |
apply (simp add: rename_def ioa_projections) |
|
449 |
apply (drule trans_rename) |
|
450 |
apply (erule exE) |
|
451 |
apply (erule conjE) |
|
452 |
apply (erule reachable_n) |
|
453 |
apply assumption |
|
454 |
done |
|
19741 | 455 |
|
456 |
||
62116 | 457 |
subsection \<open>\<open>trans_of (A \<parallel> B)\<close>\<close> |
19741 | 458 |
|
62116 | 459 |
lemma trans_A_proj: |
460 |
"(s, a, t) \<in> trans_of (A \<parallel> B) \<Longrightarrow> a \<in> act A \<Longrightarrow> (fst s, a, fst t) \<in> trans_of A" |
|
62005 | 461 |
by (simp add: Let_def par_def trans_of_def) |
19741 | 462 |
|
62116 | 463 |
lemma trans_B_proj: |
464 |
"(s, a, t) \<in> trans_of (A \<parallel> B) \<Longrightarrow> a \<in> act B \<Longrightarrow> (snd s, a, snd t) \<in> trans_of B" |
|
62005 | 465 |
by (simp add: Let_def par_def trans_of_def) |
19741 | 466 |
|
62116 | 467 |
lemma trans_A_proj2: "(s, a, t) \<in> trans_of (A \<parallel> B) \<Longrightarrow> a \<notin> act A \<Longrightarrow> fst s = fst t" |
62005 | 468 |
by (simp add: Let_def par_def trans_of_def) |
19741 | 469 |
|
62116 | 470 |
lemma trans_B_proj2: "(s, a, t) \<in> trans_of (A \<parallel> B) \<Longrightarrow> a \<notin> act B \<Longrightarrow> snd s = snd t" |
471 |
by (simp add: Let_def par_def trans_of_def) |
|
472 |
||
473 |
lemma trans_AB_proj: "(s, a, t) \<in> trans_of (A \<parallel> B) \<Longrightarrow> a \<in> act A \<or> a \<in> act B" |
|
62005 | 474 |
by (simp add: Let_def par_def trans_of_def) |
19741 | 475 |
|
62116 | 476 |
lemma trans_AB: |
477 |
"a \<in> act A \<Longrightarrow> a \<in> act B \<Longrightarrow> |
|
478 |
(fst s, a, fst t) \<in> trans_of A \<Longrightarrow> |
|
479 |
(snd s, a, snd t) \<in> trans_of B \<Longrightarrow> |
|
480 |
(s, a, t) \<in> trans_of (A \<parallel> B)" |
|
62005 | 481 |
by (simp add: Let_def par_def trans_of_def) |
19741 | 482 |
|
62116 | 483 |
lemma trans_A_notB: |
484 |
"a \<in> act A \<Longrightarrow> a \<notin> act B \<Longrightarrow> |
|
485 |
(fst s, a, fst t) \<in> trans_of A \<Longrightarrow> |
|
486 |
snd s = snd t \<Longrightarrow> |
|
487 |
(s, a, t) \<in> trans_of (A \<parallel> B)" |
|
62005 | 488 |
by (simp add: Let_def par_def trans_of_def) |
19741 | 489 |
|
62116 | 490 |
lemma trans_notA_B: |
491 |
"a \<notin> act A \<Longrightarrow> a \<in> act B \<Longrightarrow> |
|
492 |
(snd s, a, snd t) \<in> trans_of B \<Longrightarrow> |
|
493 |
fst s = fst t \<Longrightarrow> |
|
494 |
(s, a, t) \<in> trans_of (A \<parallel> B)" |
|
62005 | 495 |
by (simp add: Let_def par_def trans_of_def) |
19741 | 496 |
|
497 |
lemmas trans_of_defs1 = trans_AB trans_A_notB trans_notA_B |
|
498 |
and trans_of_defs2 = trans_A_proj trans_B_proj trans_A_proj2 trans_B_proj2 trans_AB_proj |
|
499 |
||
500 |
||
62005 | 501 |
lemma trans_of_par4: |
62116 | 502 |
"(s, a, t) \<in> trans_of (A \<parallel> B \<parallel> C \<parallel> D) \<longleftrightarrow> |
503 |
((a \<in> actions (asig_of A) \<or> a \<in> actions (asig_of B) \<or> a \<in> actions (asig_of C) \<or> |
|
504 |
a \<in> actions (asig_of D)) \<and> |
|
505 |
(if a \<in> actions (asig_of A) |
|
506 |
then (fst s, a, fst t) \<in> trans_of A |
|
507 |
else fst t = fst s) \<and> |
|
508 |
(if a \<in> actions (asig_of B) |
|
509 |
then (fst (snd s), a, fst (snd t)) \<in> trans_of B |
|
510 |
else fst (snd t) = fst (snd s)) \<and> |
|
511 |
(if a \<in> actions (asig_of C) |
|
512 |
then (fst (snd (snd s)), a, fst (snd (snd t))) \<in> trans_of C |
|
513 |
else fst (snd (snd t)) = fst (snd (snd s))) \<and> |
|
514 |
(if a \<in> actions (asig_of D) |
|
515 |
then (snd (snd (snd s)), a, snd (snd (snd t))) \<in> trans_of D |
|
516 |
else snd (snd (snd t)) = snd (snd (snd s))))" |
|
62005 | 517 |
by (simp add: par_def actions_asig_comp prod_eq_iff Let_def ioa_projections) |
19741 | 518 |
|
519 |
||
62116 | 520 |
subsection \<open>Proof obligation generator for IOA requirements\<close> |
19741 | 521 |
|
62116 | 522 |
(*without assumptions on A and B because is_trans_of is also incorporated in par_def*) |
523 |
lemma is_trans_of_par: "is_trans_of (A \<parallel> B)" |
|
62005 | 524 |
by (simp add: is_trans_of_def Let_def actions_of_par trans_of_par) |
525 |
||
62116 | 526 |
lemma is_trans_of_restrict: "is_trans_of A \<Longrightarrow> is_trans_of (restrict A acts)" |
62005 | 527 |
by (simp add: is_trans_of_def cancel_restrict acts_restrict) |
19741 | 528 |
|
62116 | 529 |
lemma is_trans_of_rename: "is_trans_of A \<Longrightarrow> is_trans_of (rename A f)" |
62005 | 530 |
apply (unfold is_trans_of_def restrict_def restrict_asig_def) |
531 |
apply (simp add: Let_def actions_def trans_of_def asig_internals_def |
|
532 |
asig_outputs_def asig_inputs_def externals_def asig_of_def rename_def rename_set_def) |
|
533 |
apply blast |
|
534 |
done |
|
19741 | 535 |
|
62116 | 536 |
lemma is_asig_of_par: "is_asig_of A \<Longrightarrow> is_asig_of B \<Longrightarrow> compatible A B \<Longrightarrow> is_asig_of (A \<parallel> B)" |
62005 | 537 |
apply (simp add: is_asig_of_def asig_of_par asig_comp_def compatible_def |
538 |
asig_internals_def asig_outputs_def asig_inputs_def actions_def is_asig_def) |
|
539 |
apply (simp add: asig_of_def) |
|
540 |
apply auto |
|
541 |
done |
|
19741 | 542 |
|
62116 | 543 |
lemma is_asig_of_restrict: "is_asig_of A \<Longrightarrow> is_asig_of (restrict A f)" |
62005 | 544 |
apply (unfold is_asig_of_def is_asig_def asig_of_def restrict_def restrict_asig_def |
62116 | 545 |
asig_internals_def asig_outputs_def asig_inputs_def externals_def o_def) |
62005 | 546 |
apply simp |
547 |
apply auto |
|
548 |
done |
|
19741 | 549 |
|
62116 | 550 |
lemma is_asig_of_rename: "is_asig_of A \<Longrightarrow> is_asig_of (rename A f)" |
62005 | 551 |
apply (simp add: is_asig_of_def rename_def rename_set_def asig_internals_def |
552 |
asig_outputs_def asig_inputs_def actions_def is_asig_def asig_of_def) |
|
553 |
apply auto |
|
554 |
apply (drule_tac [!] s = "Some _" in sym) |
|
555 |
apply auto |
|
556 |
done |
|
19741 | 557 |
|
558 |
lemmas [simp] = is_asig_of_par is_asig_of_restrict |
|
559 |
is_asig_of_rename is_trans_of_par is_trans_of_restrict is_trans_of_rename |
|
560 |
||
561 |
||
62116 | 562 |
lemma compatible_par: "compatible A B \<Longrightarrow> compatible A C \<Longrightarrow> compatible A (B \<parallel> C)" |
563 |
by (auto simp add: compatible_def internals_of_par outputs_of_par actions_of_par) |
|
19741 | 564 |
|
62116 | 565 |
(*better derive by previous one and compat_commute*) |
566 |
lemma compatible_par2: "compatible A C \<Longrightarrow> compatible B C \<Longrightarrow> compatible (A \<parallel> B) C" |
|
567 |
by (auto simp add: compatible_def internals_of_par outputs_of_par actions_of_par) |
|
19741 | 568 |
|
62005 | 569 |
lemma compatible_restrict: |
62116 | 570 |
"compatible A B \<Longrightarrow> (ext B - S) \<inter> ext A = {} \<Longrightarrow> compatible A (restrict B S)" |
571 |
by (auto simp add: compatible_def ioa_triple_proj asig_triple_proj externals_def |
|
62005 | 572 |
restrict_def restrict_asig_def actions_def) |
19741 | 573 |
|
574 |
declare split_paired_Ex [simp] |
|
3071 | 575 |
|
576 |
end |