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