author | wenzelm |
Sat, 16 Jan 2016 23:31:28 +0100 | |
changeset 62193 | 0826f6b6ba09 |
parent 62116 | bc178c0fe1a1 |
child 62195 | 799a5306e2ed |
permissions | -rw-r--r-- |
62008 | 1 |
(* Title: HOL/HOLCF/IOA/Traces.thy |
40945 | 2 |
Author: Olaf Müller |
17233 | 3 |
*) |
3071 | 4 |
|
62002 | 5 |
section \<open>Executions and Traces of I/O automata in HOLCF\<close> |
3071 | 6 |
|
17233 | 7 |
theory Traces |
8 |
imports Sequence Automata |
|
9 |
begin |
|
3071 | 10 |
|
36452 | 11 |
default_sort type |
17233 | 12 |
|
62116 | 13 |
type_synonym ('a, 's) pairs = "('a \<times> 's) Seq" |
14 |
type_synonym ('a, 's) execution = "'s \<times> ('a, 's) pairs" |
|
62005 | 15 |
type_synonym 'a trace = "'a Seq" |
62116 | 16 |
type_synonym ('a, 's) execution_module = "('a, 's) execution set \<times> 'a signature" |
17 |
type_synonym 'a schedule_module = "'a trace set \<times> 'a signature" |
|
18 |
type_synonym 'a trace_module = "'a trace set \<times> 'a signature" |
|
3071 | 19 |
|
20 |
||
62116 | 21 |
subsection \<open>Executions\<close> |
3071 | 22 |
|
62005 | 23 |
definition is_exec_fragC :: "('a, 's) ioa \<Rightarrow> ('a, 's) pairs \<rightarrow> 's \<Rightarrow> tr" |
62116 | 24 |
where "is_exec_fragC A = |
25 |
(fix $ |
|
26 |
(LAM h ex. |
|
27 |
(\<lambda>s. |
|
28 |
case ex of |
|
29 |
nil \<Rightarrow> TT |
|
30 |
| x ## xs \<Rightarrow> flift1 (\<lambda>p. Def ((s, p) \<in> trans_of A) andalso (h $ xs) (snd p)) $ x)))" |
|
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
31 |
|
62116 | 32 |
definition is_exec_frag :: "('a, 's) ioa \<Rightarrow> ('a, 's) execution \<Rightarrow> bool" |
33 |
where "is_exec_frag A ex \<longleftrightarrow> (is_exec_fragC A $ (snd ex)) (fst ex) \<noteq> FF" |
|
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
34 |
|
62005 | 35 |
definition executions :: "('a, 's) ioa \<Rightarrow> ('a, 's) execution set" |
62116 | 36 |
where "executions ioa = {e. fst e \<in> starts_of ioa \<and> is_exec_frag ioa e}" |
3071 | 37 |
|
38 |
||
62116 | 39 |
subsection \<open>Schedules\<close> |
3071 | 40 |
|
62005 | 41 |
definition filter_act :: "('a, 's) pairs \<rightarrow> 'a trace" |
42 |
where "filter_act = Map fst" |
|
3071 | 43 |
|
62116 | 44 |
definition has_schedule :: "('a, 's) ioa \<Rightarrow> 'a trace \<Rightarrow> bool" |
62005 | 45 |
where "has_schedule ioa sch \<longleftrightarrow> (\<exists>ex \<in> executions ioa. sch = filter_act $ (snd ex))" |
3071 | 46 |
|
62005 | 47 |
definition schedules :: "('a, 's) ioa \<Rightarrow> 'a trace set" |
48 |
where "schedules ioa = {sch. has_schedule ioa sch}" |
|
3071 | 49 |
|
50 |
||
62116 | 51 |
subsection \<open>Traces\<close> |
3071 | 52 |
|
62116 | 53 |
definition has_trace :: "('a, 's) ioa \<Rightarrow> 'a trace \<Rightarrow> bool" |
54 |
where "has_trace ioa tr \<longleftrightarrow> (\<exists>sch \<in> schedules ioa. tr = Filter (\<lambda>a. a \<in> ext ioa) $ sch)" |
|
17233 | 55 |
|
62005 | 56 |
definition traces :: "('a, 's) ioa \<Rightarrow> 'a trace set" |
57 |
where "traces ioa \<equiv> {tr. has_trace ioa tr}" |
|
3071 | 58 |
|
62005 | 59 |
definition mk_trace :: "('a, 's) ioa \<Rightarrow> ('a, 's) pairs \<rightarrow> 'a trace" |
60 |
where "mk_trace ioa = (LAM tr. Filter (\<lambda>a. a \<in> ext ioa) $ (filter_act $ tr))" |
|
3071 | 61 |
|
62 |
||
62116 | 63 |
subsection \<open>Fair Traces\<close> |
4559
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
3842
diff
changeset
|
64 |
|
62005 | 65 |
definition laststate :: "('a, 's) execution \<Rightarrow> 's" |
62116 | 66 |
where "laststate ex = |
67 |
(case Last $ (snd ex) of |
|
68 |
UU \<Rightarrow> fst ex |
|
69 |
| Def at \<Rightarrow> snd at)" |
|
4559
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
3842
diff
changeset
|
70 |
|
62116 | 71 |
text \<open>A predicate holds infinitely (finitely) often in a sequence.\<close> |
62005 | 72 |
definition inf_often :: "('a \<Rightarrow> bool) \<Rightarrow> 'a Seq \<Rightarrow> bool" |
73 |
where "inf_often P s \<longleftrightarrow> Infinite (Filter P $ s)" |
|
4559
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
3842
diff
changeset
|
74 |
|
62116 | 75 |
text \<open>Filtering \<open>P\<close> yields a finite or partial sequence.\<close> |
62005 | 76 |
definition fin_often :: "('a \<Rightarrow> bool) \<Rightarrow> 'a Seq \<Rightarrow> bool" |
77 |
where "fin_often P s \<longleftrightarrow> \<not> inf_often P s" |
|
78 |
||
79 |
||
62116 | 80 |
subsection \<open>Fairness of executions\<close> |
4559
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
3842
diff
changeset
|
81 |
|
62116 | 82 |
text \<open> |
83 |
Note that partial execs cannot be \<open>wfair\<close> as the inf_often predicate in the |
|
84 |
else branch prohibits it. However they can be \<open>sfair\<close> in the case when all |
|
85 |
\<open>W\<close> are only finitely often enabled: Is this the right model? |
|
86 |
||
87 |
See @{file "LiveIOA.thy"} for solution conforming with the literature and |
|
88 |
superseding this one. |
|
89 |
\<close> |
|
90 |
||
62005 | 91 |
definition is_wfair :: "('a, 's) ioa \<Rightarrow> 'a set \<Rightarrow> ('a, 's) execution \<Rightarrow> bool" |
62116 | 92 |
where "is_wfair A W ex \<longleftrightarrow> |
93 |
(inf_often (\<lambda>x. fst x \<in> W) (snd ex) \<or> |
|
94 |
inf_often (\<lambda>x. \<not> Enabled A W (snd x)) (snd ex))" |
|
95 |
||
62005 | 96 |
definition wfair_ex :: "('a, 's) ioa \<Rightarrow> ('a, 's) execution \<Rightarrow> bool" |
62116 | 97 |
where "wfair_ex A ex \<longleftrightarrow> |
98 |
(\<forall>W \<in> wfair_of A. |
|
99 |
if Finite (snd ex) |
|
100 |
then \<not> Enabled A W (laststate ex) |
|
101 |
else is_wfair A W ex)" |
|
4559
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
3842
diff
changeset
|
102 |
|
62005 | 103 |
definition is_sfair :: "('a, 's) ioa \<Rightarrow> 'a set \<Rightarrow> ('a, 's) execution \<Rightarrow> bool" |
62116 | 104 |
where "is_sfair A W ex \<longleftrightarrow> |
105 |
(inf_often (\<lambda>x. fst x \<in> W) (snd ex) \<or> |
|
106 |
fin_often (\<lambda>x. Enabled A W (snd x)) (snd ex))" |
|
107 |
||
62005 | 108 |
definition sfair_ex :: "('a, 's)ioa \<Rightarrow> ('a, 's) execution \<Rightarrow> bool" |
62116 | 109 |
where "sfair_ex A ex \<longleftrightarrow> |
110 |
(\<forall>W \<in> sfair_of A. |
|
111 |
if Finite (snd ex) |
|
112 |
then \<not> Enabled A W (laststate ex) |
|
113 |
else is_sfair A W ex)" |
|
4559
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
3842
diff
changeset
|
114 |
|
62005 | 115 |
definition fair_ex :: "('a, 's) ioa \<Rightarrow> ('a, 's) execution \<Rightarrow> bool" |
116 |
where "fair_ex A ex \<longleftrightarrow> wfair_ex A ex \<and> sfair_ex A ex" |
|
117 |
||
4559
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
3842
diff
changeset
|
118 |
|
62116 | 119 |
text \<open>Fair behavior sets.\<close> |
4559
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
3842
diff
changeset
|
120 |
|
62005 | 121 |
definition fairexecutions :: "('a, 's) ioa \<Rightarrow> ('a, 's) execution set" |
122 |
where "fairexecutions A = {ex. ex \<in> executions A \<and> fair_ex A ex}" |
|
4559
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
3842
diff
changeset
|
123 |
|
62005 | 124 |
definition fairtraces :: "('a, 's) ioa \<Rightarrow> 'a trace set" |
125 |
where "fairtraces A = {mk_trace A $ (snd ex) | ex. ex \<in> fairexecutions A}" |
|
4559
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
3842
diff
changeset
|
126 |
|
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
3842
diff
changeset
|
127 |
|
62116 | 128 |
subsection \<open>Implementation\<close> |
3071 | 129 |
|
62116 | 130 |
subsubsection \<open>Notions of implementation\<close> |
3071 | 131 |
|
62116 | 132 |
definition ioa_implements :: "('a, 's1) ioa \<Rightarrow> ('a, 's2) ioa \<Rightarrow> bool" (infixr "=<|" 12) |
133 |
where "(ioa1 =<| ioa2) \<longleftrightarrow> |
|
134 |
(inputs (asig_of ioa1) = inputs (asig_of ioa2) \<and> |
|
135 |
outputs (asig_of ioa1) = outputs (asig_of ioa2)) \<and> |
|
136 |
traces ioa1 \<subseteq> traces ioa2" |
|
62005 | 137 |
|
138 |
definition fair_implements :: "('a, 's1) ioa \<Rightarrow> ('a, 's2) ioa \<Rightarrow> bool" |
|
62116 | 139 |
where "fair_implements C A \<longleftrightarrow> |
140 |
inp C = inp A \<and> out C = out A \<and> fairtraces C \<subseteq> fairtraces A" |
|
62005 | 141 |
|
62193 | 142 |
lemma implements_trans: "A =<| B \<Longrightarrow> B =<| C \<Longrightarrow> A =<| C" |
143 |
by (auto simp add: ioa_implements_def) |
|
144 |
||
4559
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
3842
diff
changeset
|
145 |
|
62116 | 146 |
subsection \<open>Modules\<close> |
3521 | 147 |
|
62116 | 148 |
subsubsection \<open>Execution, schedule and trace modules\<close> |
62005 | 149 |
|
150 |
definition Execs :: "('a, 's) ioa \<Rightarrow> ('a, 's) execution_module" |
|
151 |
where "Execs A = (executions A, asig_of A)" |
|
3521 | 152 |
|
62005 | 153 |
definition Scheds :: "('a, 's) ioa \<Rightarrow> 'a schedule_module" |
154 |
where "Scheds A = (schedules A, asig_of A)" |
|
3521 | 155 |
|
62005 | 156 |
definition Traces :: "('a, 's) ioa \<Rightarrow> 'a trace_module" |
157 |
where "Traces A = (traces A, asig_of A)" |
|
3521 | 158 |
|
37598 | 159 |
lemmas [simp del] = HOL.ex_simps HOL.all_simps split_paired_Ex |
19741 | 160 |
declare Let_def [simp] |
62002 | 161 |
setup \<open>map_theory_claset (fn ctxt => ctxt delSWrapper "split_all_tac")\<close> |
19741 | 162 |
|
163 |
lemmas exec_rws = executions_def is_exec_frag_def |
|
164 |
||
165 |
||
62116 | 166 |
subsection \<open>Recursive equations of operators\<close> |
19741 | 167 |
|
62116 | 168 |
subsubsection \<open>\<open>filter_act\<close>\<close> |
19741 | 169 |
|
62116 | 170 |
lemma filter_act_UU: "filter_act $ UU = UU" |
62005 | 171 |
by (simp add: filter_act_def) |
19741 | 172 |
|
62116 | 173 |
lemma filter_act_nil: "filter_act $ nil = nil" |
62005 | 174 |
by (simp add: filter_act_def) |
19741 | 175 |
|
62116 | 176 |
lemma filter_act_cons: "filter_act $ (x \<leadsto> xs) = fst x \<leadsto> filter_act $ xs" |
62005 | 177 |
by (simp add: filter_act_def) |
19741 | 178 |
|
179 |
declare filter_act_UU [simp] filter_act_nil [simp] filter_act_cons [simp] |
|
180 |
||
181 |
||
62116 | 182 |
subsubsection \<open>\<open>mk_trace\<close>\<close> |
19741 | 183 |
|
62116 | 184 |
lemma mk_trace_UU: "mk_trace A $ UU = UU" |
62005 | 185 |
by (simp add: mk_trace_def) |
19741 | 186 |
|
62116 | 187 |
lemma mk_trace_nil: "mk_trace A $ nil = nil" |
62005 | 188 |
by (simp add: mk_trace_def) |
19741 | 189 |
|
62116 | 190 |
lemma mk_trace_cons: |
191 |
"mk_trace A $ (at \<leadsto> xs) = |
|
192 |
(if fst at \<in> ext A |
|
193 |
then fst at \<leadsto> mk_trace A $ xs |
|
194 |
else mk_trace A $ xs)" |
|
62005 | 195 |
by (simp add: mk_trace_def) |
19741 | 196 |
|
197 |
declare mk_trace_UU [simp] mk_trace_nil [simp] mk_trace_cons [simp] |
|
198 |
||
199 |
||
62116 | 200 |
subsubsection \<open>\<open>is_exec_fragC\<close>\<close> |
19741 | 201 |
|
62116 | 202 |
lemma is_exec_fragC_unfold: |
203 |
"is_exec_fragC A = |
|
204 |
(LAM ex. |
|
205 |
(\<lambda>s. |
|
206 |
case ex of |
|
207 |
nil \<Rightarrow> TT |
|
208 |
| x ## xs \<Rightarrow> |
|
209 |
(flift1 (\<lambda>p. Def ((s, p) \<in> trans_of A) andalso (is_exec_fragC A$xs) (snd p)) $ x)))" |
|
62005 | 210 |
apply (rule trans) |
211 |
apply (rule fix_eq4) |
|
212 |
apply (rule is_exec_fragC_def) |
|
213 |
apply (rule beta_cfun) |
|
214 |
apply (simp add: flift1_def) |
|
215 |
done |
|
19741 | 216 |
|
62116 | 217 |
lemma is_exec_fragC_UU: "(is_exec_fragC A $ UU) s = UU" |
62005 | 218 |
apply (subst is_exec_fragC_unfold) |
219 |
apply simp |
|
220 |
done |
|
19741 | 221 |
|
62116 | 222 |
lemma is_exec_fragC_nil: "(is_exec_fragC A $ nil) s = TT" |
62005 | 223 |
apply (subst is_exec_fragC_unfold) |
224 |
apply simp |
|
225 |
done |
|
19741 | 226 |
|
62116 | 227 |
lemma is_exec_fragC_cons: |
228 |
"(is_exec_fragC A $ (pr \<leadsto> xs)) s = |
|
229 |
(Def ((s, pr) \<in> trans_of A) andalso (is_exec_fragC A $ xs) (snd pr))" |
|
62005 | 230 |
apply (rule trans) |
231 |
apply (subst is_exec_fragC_unfold) |
|
232 |
apply (simp add: Consq_def flift1_def) |
|
233 |
apply simp |
|
234 |
done |
|
19741 | 235 |
|
236 |
declare is_exec_fragC_UU [simp] is_exec_fragC_nil [simp] is_exec_fragC_cons [simp] |
|
237 |
||
238 |
||
62116 | 239 |
subsubsection \<open>\<open>is_exec_frag\<close>\<close> |
19741 | 240 |
|
241 |
lemma is_exec_frag_UU: "is_exec_frag A (s, UU)" |
|
62005 | 242 |
by (simp add: is_exec_frag_def) |
19741 | 243 |
|
244 |
lemma is_exec_frag_nil: "is_exec_frag A (s, nil)" |
|
62005 | 245 |
by (simp add: is_exec_frag_def) |
19741 | 246 |
|
62116 | 247 |
lemma is_exec_frag_cons: |
248 |
"is_exec_frag A (s, (a, t) \<leadsto> ex) \<longleftrightarrow> (s, a, t) \<in> trans_of A \<and> is_exec_frag A (t, ex)" |
|
62005 | 249 |
by (simp add: is_exec_frag_def) |
19741 | 250 |
|
251 |
declare is_exec_frag_UU [simp] is_exec_frag_nil [simp] is_exec_frag_cons [simp] |
|
252 |
||
253 |
||
62116 | 254 |
subsubsection \<open>\<open>laststate\<close>\<close> |
255 |
||
256 |
lemma laststate_UU: "laststate (s, UU) = s" |
|
62005 | 257 |
by (simp add: laststate_def) |
19741 | 258 |
|
62116 | 259 |
lemma laststate_nil: "laststate (s, nil) = s" |
62005 | 260 |
by (simp add: laststate_def) |
19741 | 261 |
|
62116 | 262 |
lemma laststate_cons: "Finite ex \<Longrightarrow> laststate (s, at \<leadsto> ex) = laststate (snd at, ex)" |
263 |
apply (simp add: laststate_def) |
|
264 |
apply (cases "ex = nil") |
|
265 |
apply simp |
|
266 |
apply simp |
|
62005 | 267 |
apply (drule Finite_Last1 [THEN mp]) |
268 |
apply assumption |
|
269 |
apply defined |
|
270 |
done |
|
19741 | 271 |
|
272 |
declare laststate_UU [simp] laststate_nil [simp] laststate_cons [simp] |
|
273 |
||
62116 | 274 |
lemma exists_laststate: "Finite ex \<Longrightarrow> \<forall>s. \<exists>u. laststate (s, ex) = u" |
275 |
by (tactic "Seq_Finite_induct_tac @{context} 1") |
|
19741 | 276 |
|
277 |
||
62116 | 278 |
subsection \<open>\<open>has_trace\<close> \<open>mk_trace\<close>\<close> |
19741 | 279 |
|
62116 | 280 |
(*alternative definition of has_trace tailored for the refinement proof, as it does not |
281 |
take the detour of schedules*) |
|
282 |
lemma has_trace_def2: "has_trace A b \<longleftrightarrow> (\<exists>ex \<in> executions A. b = mk_trace A $ (snd ex))" |
|
62005 | 283 |
apply (unfold executions_def mk_trace_def has_trace_def schedules_def has_schedule_def [abs_def]) |
284 |
apply auto |
|
285 |
done |
|
19741 | 286 |
|
287 |
||
62116 | 288 |
subsection \<open>Signatures and executions, schedules\<close> |
19741 | 289 |
|
62116 | 290 |
text \<open> |
291 |
All executions of \<open>A\<close> have only actions of \<open>A\<close>. This is only true because of |
|
292 |
the predicate \<open>state_trans\<close> (part of the predicate \<open>IOA\<close>): We have no |
|
293 |
dependent types. For executions of parallel automata this assumption is not |
|
294 |
needed, as in \<open>par_def\<close> this condition is included once more. (See Lemmas |
|
295 |
1.1.1c in CompoExecs for example.) |
|
296 |
\<close> |
|
19741 | 297 |
|
62005 | 298 |
lemma execfrag_in_sig: |
62116 | 299 |
"is_trans_of A \<Longrightarrow> \<forall>s. is_exec_frag A (s, xs) \<longrightarrow> Forall (\<lambda>a. a \<in> act A) (filter_act $ xs)" |
62005 | 300 |
apply (tactic \<open>pair_induct_tac @{context} "xs" [@{thm is_exec_frag_def}, |
301 |
@{thm Forall_def}, @{thm sforall_def}] 1\<close>) |
|
62116 | 302 |
text \<open>main case\<close> |
62005 | 303 |
apply (auto simp add: is_trans_of_def) |
304 |
done |
|
19741 | 305 |
|
62005 | 306 |
lemma exec_in_sig: |
62116 | 307 |
"is_trans_of A \<Longrightarrow> x \<in> executions A \<Longrightarrow> Forall (\<lambda>a. a \<in> act A) (filter_act $ (snd x))" |
62005 | 308 |
apply (simp add: executions_def) |
309 |
apply (tactic \<open>pair_tac @{context} "x" 1\<close>) |
|
310 |
apply (rule execfrag_in_sig [THEN spec, THEN mp]) |
|
311 |
apply auto |
|
312 |
done |
|
19741 | 313 |
|
62116 | 314 |
lemma scheds_in_sig: "is_trans_of A \<Longrightarrow> x \<in> schedules A \<Longrightarrow> Forall (\<lambda>a. a \<in> act A) x" |
62005 | 315 |
apply (unfold schedules_def has_schedule_def [abs_def]) |
316 |
apply (fast intro!: exec_in_sig) |
|
317 |
done |
|
19741 | 318 |
|
319 |
||
62116 | 320 |
subsection \<open>Executions are prefix closed\<close> |
19741 | 321 |
|
62116 | 322 |
(*only admissible in y, not if done in x!*) |
323 |
lemma execfrag_prefixclosed: "\<forall>x s. is_exec_frag A (s, x) \<and> y \<sqsubseteq> x \<longrightarrow> is_exec_frag A (s, y)" |
|
62005 | 324 |
apply (tactic \<open>pair_induct_tac @{context} "y" [@{thm is_exec_frag_def}] 1\<close>) |
325 |
apply (intro strip) |
|
326 |
apply (tactic \<open>Seq_case_simp_tac @{context} "x" 1\<close>) |
|
327 |
apply (tactic \<open>pair_tac @{context} "a" 1\<close>) |
|
328 |
apply auto |
|
329 |
done |
|
19741 | 330 |
|
331 |
lemmas exec_prefixclosed = |
|
45606 | 332 |
conjI [THEN execfrag_prefixclosed [THEN spec, THEN spec, THEN mp]] |
19741 | 333 |
|
62116 | 334 |
(*second prefix notion for Finite x*) |
19741 | 335 |
lemma exec_prefix2closed [rule_format]: |
62116 | 336 |
"\<forall>y s. is_exec_frag A (s, x @@ y) \<longrightarrow> is_exec_frag A (s, x)" |
62005 | 337 |
apply (tactic \<open>pair_induct_tac @{context} "x" [@{thm is_exec_frag_def}] 1\<close>) |
338 |
apply (intro strip) |
|
339 |
apply (tactic \<open>Seq_case_simp_tac @{context} "s" 1\<close>) |
|
340 |
apply (tactic \<open>pair_tac @{context} "a" 1\<close>) |
|
341 |
apply auto |
|
342 |
done |
|
19741 | 343 |
|
17233 | 344 |
end |