author | wenzelm |
Sat, 16 Jan 2016 23:31:28 +0100 | |
changeset 62193 | 0826f6b6ba09 |
parent 62156 | 7355fd313cf8 |
child 62195 | 799a5306e2ed |
permissions | -rw-r--r-- |
62008 | 1 |
(* Title: HOL/HOLCF/IOA/RefCorrectness.thy |
40945 | 2 |
Author: Olaf Müller |
3071 | 3 |
*) |
4 |
||
62002 | 5 |
section \<open>Correctness of Refinement Mappings in HOLCF/IOA\<close> |
3071 | 6 |
|
17233 | 7 |
theory RefCorrectness |
8 |
imports RefMappings |
|
9 |
begin |
|
10 |
||
62156 | 11 |
definition corresp_exC :: |
12 |
"('a, 's2) ioa \<Rightarrow> ('s1 \<Rightarrow> 's2) \<Rightarrow> ('a, 's1) pairs \<rightarrow> ('s1 \<Rightarrow> ('a, 's2) pairs)" |
|
13 |
where "corresp_exC A f = |
|
14 |
(fix $ |
|
15 |
(LAM h ex. |
|
16 |
(\<lambda>s. case ex of |
|
17 |
nil \<Rightarrow> nil |
|
18 |
| x ## xs \<Rightarrow> |
|
19 |
flift1 (\<lambda>pr. |
|
20 |
(SOME cex. move A cex (f s) (fst pr) (f (snd pr))) @@ ((h $ xs) (snd pr))) $ x)))" |
|
17233 | 21 |
|
62156 | 22 |
definition corresp_ex :: |
23 |
"('a, 's2) ioa \<Rightarrow> ('s1 \<Rightarrow> 's2) \<Rightarrow> ('a, 's1) execution \<Rightarrow> ('a, 's2) execution" |
|
24 |
where "corresp_ex A f ex = (f (fst ex), (corresp_exC A f $ (snd ex)) (fst ex))" |
|
4559
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
3433
diff
changeset
|
25 |
|
62156 | 26 |
definition is_fair_ref_map :: |
27 |
"('s1 \<Rightarrow> 's2) \<Rightarrow> ('a, 's1) ioa \<Rightarrow> ('a, 's2) ioa \<Rightarrow> bool" |
|
28 |
where "is_fair_ref_map f C A \<longleftrightarrow> |
|
29 |
is_ref_map f C A \<and> (\<forall>ex \<in> executions C. fair_ex C ex \<longrightarrow> fair_ex A (corresp_ex A f ex))" |
|
30 |
||
31 |
text \<open> |
|
32 |
Axioms for fair trace inclusion proof support, not for the correctness proof |
|
33 |
of refinement mappings! |
|
34 |
||
35 |
Note: Everything is superseded by @{file "LiveIOA.thy"}. |
|
36 |
\<close> |
|
4559
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
3433
diff
changeset
|
37 |
|
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
38 |
axiomatization where |
62156 | 39 |
corresp_laststate: |
40 |
"Finite ex \<Longrightarrow> laststate (corresp_ex A f (s, ex)) = f (laststate (s, ex))" |
|
41 |
||
42 |
axiomatization where |
|
43 |
corresp_Finite: "Finite (snd (corresp_ex A f (s, ex))) = Finite ex" |
|
4559
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
3433
diff
changeset
|
44 |
|
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
45 |
axiomatization where |
62156 | 46 |
FromAtoC: |
47 |
"fin_often (\<lambda>x. P (snd x)) (snd (corresp_ex A f (s, ex))) \<Longrightarrow> |
|
48 |
fin_often (\<lambda>y. P (f (snd y))) ex" |
|
4559
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
3433
diff
changeset
|
49 |
|
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
50 |
axiomatization where |
62156 | 51 |
FromCtoA: |
52 |
"inf_often (\<lambda>y. P (fst y)) ex \<Longrightarrow> |
|
53 |
inf_often (\<lambda>x. P (fst x)) (snd (corresp_ex A f (s,ex)))" |
|
4559
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
3433
diff
changeset
|
54 |
|
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
3433
diff
changeset
|
55 |
|
62156 | 56 |
text \<open> |
57 |
Proof by case on \<open>inf W\<close> in ex: If so, ok. If not, only \<open>fin W\<close> in ex, ie. |
|
58 |
there is an index \<open>i\<close> from which on no \<open>W\<close> in ex. But \<open>W\<close> inf enabled, ie at |
|
59 |
least once after \<open>i\<close> \<open>W\<close> is enabled. As \<open>W\<close> does not occur after \<open>i\<close> and \<open>W\<close> |
|
60 |
is \<open>enabling_persistent\<close>, \<open>W\<close> keeps enabled until infinity, ie. indefinitely |
|
61 |
\<close> |
|
4559
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
3433
diff
changeset
|
62 |
|
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
63 |
axiomatization where |
62156 | 64 |
persistent: |
65 |
"inf_often (\<lambda>x. Enabled A W (snd x)) ex \<Longrightarrow> en_persistent A W \<Longrightarrow> |
|
66 |
inf_often (\<lambda>x. fst x \<in> W) ex \<or> fin_often (\<lambda>x. \<not> Enabled A W (snd x)) ex" |
|
67 |
||
68 |
axiomatization where |
|
69 |
infpostcond: |
|
70 |
"is_exec_frag A (s,ex) \<Longrightarrow> inf_often (\<lambda>x. fst x \<in> W) ex \<Longrightarrow> |
|
71 |
inf_often (\<lambda>x. set_was_enabled A W (snd x)) ex" |
|
4559
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
3433
diff
changeset
|
72 |
|
19741 | 73 |
|
62156 | 74 |
subsection \<open>\<open>corresp_ex\<close>\<close> |
19741 | 75 |
|
62156 | 76 |
lemma corresp_exC_unfold: |
77 |
"corresp_exC A f = |
|
78 |
(LAM ex. |
|
79 |
(\<lambda>s. |
|
80 |
case ex of |
|
81 |
nil \<Rightarrow> nil |
|
82 |
| x ## xs \<Rightarrow> |
|
83 |
(flift1 (\<lambda>pr. |
|
84 |
(SOME cex. move A cex (f s) (fst pr) (f (snd pr))) @@ |
|
85 |
((corresp_exC A f $ xs) (snd pr))) $ x)))" |
|
86 |
apply (rule trans) |
|
87 |
apply (rule fix_eq2) |
|
88 |
apply (simp only: corresp_exC_def) |
|
89 |
apply (rule beta_cfun) |
|
90 |
apply (simp add: flift1_def) |
|
91 |
done |
|
19741 | 92 |
|
62156 | 93 |
lemma corresp_exC_UU: "(corresp_exC A f $ UU) s = UU" |
94 |
apply (subst corresp_exC_unfold) |
|
95 |
apply simp |
|
96 |
done |
|
19741 | 97 |
|
62156 | 98 |
lemma corresp_exC_nil: "(corresp_exC A f $ nil) s = nil" |
99 |
apply (subst corresp_exC_unfold) |
|
100 |
apply simp |
|
101 |
done |
|
19741 | 102 |
|
62156 | 103 |
lemma corresp_exC_cons: |
104 |
"(corresp_exC A f $ (at \<leadsto> xs)) s = |
|
105 |
(SOME cex. move A cex (f s) (fst at) (f (snd at))) @@ |
|
106 |
((corresp_exC A f $ xs) (snd at))" |
|
107 |
apply (rule trans) |
|
108 |
apply (subst corresp_exC_unfold) |
|
109 |
apply (simp add: Consq_def flift1_def) |
|
110 |
apply simp |
|
111 |
done |
|
19741 | 112 |
|
113 |
declare corresp_exC_UU [simp] corresp_exC_nil [simp] corresp_exC_cons [simp] |
|
114 |
||
115 |
||
62156 | 116 |
subsection \<open>Properties of move\<close> |
19741 | 117 |
|
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
118 |
lemma move_is_move: |
62156 | 119 |
"is_ref_map f C A \<Longrightarrow> reachable C s \<Longrightarrow> (s, a, t) \<in> trans_of C \<Longrightarrow> |
120 |
move A (SOME x. move A x (f s) a (f t)) (f s) a (f t)" |
|
121 |
apply (unfold is_ref_map_def) |
|
122 |
apply (subgoal_tac "\<exists>ex. move A ex (f s) a (f t) ") |
|
123 |
prefer 2 |
|
124 |
apply simp |
|
125 |
apply (erule exE) |
|
126 |
apply (rule someI) |
|
127 |
apply assumption |
|
128 |
done |
|
19741 | 129 |
|
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
130 |
lemma move_subprop1: |
62156 | 131 |
"is_ref_map f C A \<Longrightarrow> reachable C s \<Longrightarrow> (s, a, t) \<in> trans_of C \<Longrightarrow> |
132 |
is_exec_frag A (f s, SOME x. move A x (f s) a (f t))" |
|
133 |
apply (cut_tac move_is_move) |
|
134 |
defer |
|
135 |
apply assumption+ |
|
136 |
apply (simp add: move_def) |
|
137 |
done |
|
19741 | 138 |
|
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
139 |
lemma move_subprop2: |
62156 | 140 |
"is_ref_map f C A \<Longrightarrow> reachable C s \<Longrightarrow> (s, a, t) \<in> trans_of C \<Longrightarrow> |
141 |
Finite ((SOME x. move A x (f s) a (f t)))" |
|
142 |
apply (cut_tac move_is_move) |
|
143 |
defer |
|
144 |
apply assumption+ |
|
145 |
apply (simp add: move_def) |
|
146 |
done |
|
19741 | 147 |
|
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
148 |
lemma move_subprop3: |
62156 | 149 |
"is_ref_map f C A \<Longrightarrow> reachable C s \<Longrightarrow> (s, a, t) \<in> trans_of C \<Longrightarrow> |
19741 | 150 |
laststate (f s,@x. move A x (f s) a (f t)) = (f t)" |
62156 | 151 |
apply (cut_tac move_is_move) |
152 |
defer |
|
153 |
apply assumption+ |
|
154 |
apply (simp add: move_def) |
|
155 |
done |
|
19741 | 156 |
|
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
157 |
lemma move_subprop4: |
62156 | 158 |
"is_ref_map f C A \<Longrightarrow> reachable C s \<Longrightarrow> (s, a, t) \<in> trans_of C \<Longrightarrow> |
159 |
mk_trace A $ ((SOME x. move A x (f s) a (f t))) = |
|
160 |
(if a \<in> ext A then a \<leadsto> nil else nil)" |
|
161 |
apply (cut_tac move_is_move) |
|
162 |
defer |
|
163 |
apply assumption+ |
|
164 |
apply (simp add: move_def) |
|
165 |
done |
|
19741 | 166 |
|
167 |
||
62156 | 168 |
subsection \<open>TRACE INCLUSION Part 1: Traces coincide\<close> |
19741 | 169 |
|
62156 | 170 |
subsubsection \<open>Lemmata for \<open>\<Longleftarrow>\<close>\<close> |
19741 | 171 |
|
62156 | 172 |
text \<open>Lemma 1.1: Distribution of \<open>mk_trace\<close> and \<open>@@\<close>\<close> |
19741 | 173 |
|
62156 | 174 |
lemma mk_traceConc: |
175 |
"mk_trace C $ (ex1 @@ ex2) = (mk_trace C $ ex1) @@ (mk_trace C $ ex2)" |
|
176 |
by (simp add: mk_trace_def filter_act_def MapConc) |
|
19741 | 177 |
|
178 |
||
62156 | 179 |
text \<open>Lemma 1 : Traces coincide\<close> |
19741 | 180 |
|
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
181 |
lemma lemma_1: |
62156 | 182 |
"is_ref_map f C A \<Longrightarrow> ext C = ext A \<Longrightarrow> |
183 |
\<forall>s. reachable C s \<and> is_exec_frag C (s, xs) \<longrightarrow> |
|
184 |
mk_trace C $ xs = mk_trace A $ (snd (corresp_ex A f (s, xs)))" |
|
185 |
supply split_if [split del] |
|
186 |
apply (unfold corresp_ex_def) |
|
187 |
apply (tactic \<open>pair_induct_tac @{context} "xs" [@{thm is_exec_frag_def}] 1\<close>) |
|
188 |
text \<open>cons case\<close> |
|
189 |
apply (auto simp add: mk_traceConc) |
|
190 |
apply (frule reachable.reachable_n) |
|
191 |
apply assumption |
|
192 |
apply (auto simp add: move_subprop4 split add: split_if) |
|
193 |
done |
|
19741 | 194 |
|
195 |
||
62156 | 196 |
subsection \<open>TRACE INCLUSION Part 2: corresp_ex is execution\<close> |
197 |
||
198 |
subsubsection \<open>Lemmata for \<open>==>\<close>\<close> |
|
199 |
||
200 |
text \<open>Lemma 2.1\<close> |
|
201 |
||
202 |
lemma lemma_2_1 [rule_format]: |
|
203 |
"Finite xs \<longrightarrow> |
|
204 |
(\<forall>s. is_exec_frag A (s, xs) \<and> is_exec_frag A (t, ys) \<and> |
|
205 |
t = laststate (s, xs) \<longrightarrow> is_exec_frag A (s, xs @@ ys))" |
|
206 |
apply (rule impI) |
|
207 |
apply (tactic \<open>Seq_Finite_induct_tac @{context} 1\<close>) |
|
208 |
text \<open>main case\<close> |
|
209 |
apply (auto simp add: split_paired_all) |
|
210 |
done |
|
19741 | 211 |
|
212 |
||
62156 | 213 |
text \<open>Lemma 2 : \<open>corresp_ex\<close> is execution\<close> |
19741 | 214 |
|
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
215 |
lemma lemma_2: |
62156 | 216 |
"is_ref_map f C A \<Longrightarrow> |
217 |
\<forall>s. reachable C s \<and> is_exec_frag C (s, xs) \<longrightarrow> |
|
218 |
is_exec_frag A (corresp_ex A f (s, xs))" |
|
219 |
apply (unfold corresp_ex_def) |
|
19741 | 220 |
|
62156 | 221 |
apply simp |
222 |
apply (tactic \<open>pair_induct_tac @{context} "xs" [@{thm is_exec_frag_def}] 1\<close>) |
|
19741 | 223 |
|
62156 | 224 |
text \<open>main case\<close> |
225 |
apply auto |
|
226 |
apply (rule_tac t = "f x2" in lemma_2_1) |
|
19741 | 227 |
|
62156 | 228 |
text \<open>\<open>Finite\<close>\<close> |
229 |
apply (erule move_subprop2) |
|
230 |
apply assumption+ |
|
231 |
apply (rule conjI) |
|
19741 | 232 |
|
62156 | 233 |
text \<open>\<open>is_exec_frag\<close>\<close> |
234 |
apply (erule move_subprop1) |
|
235 |
apply assumption+ |
|
236 |
apply (rule conjI) |
|
19741 | 237 |
|
62156 | 238 |
text \<open>Induction hypothesis\<close> |
239 |
text \<open>\<open>reachable_n\<close> looping, therefore apply it manually\<close> |
|
240 |
apply (erule_tac x = "x2" in allE) |
|
241 |
apply simp |
|
242 |
apply (frule reachable.reachable_n) |
|
243 |
apply assumption |
|
244 |
apply simp |
|
245 |
||
246 |
text \<open>\<open>laststate\<close>\<close> |
|
247 |
apply (erule move_subprop3 [symmetric]) |
|
248 |
apply assumption+ |
|
249 |
done |
|
19741 | 250 |
|
251 |
||
62156 | 252 |
subsection \<open>Main Theorem: TRACE -- INCLUSION\<close> |
19741 | 253 |
|
62156 | 254 |
theorem trace_inclusion: |
255 |
"ext C = ext A \<Longrightarrow> is_ref_map f C A \<Longrightarrow> traces C \<subseteq> traces A" |
|
19741 | 256 |
|
257 |
apply (unfold traces_def) |
|
258 |
||
62156 | 259 |
apply (simp add: has_trace_def2) |
26359 | 260 |
apply auto |
19741 | 261 |
|
62156 | 262 |
text \<open>give execution of abstract automata\<close> |
19741 | 263 |
apply (rule_tac x = "corresp_ex A f ex" in bexI) |
264 |
||
62156 | 265 |
text \<open>Traces coincide, Lemma 1\<close> |
62002 | 266 |
apply (tactic \<open>pair_tac @{context} "ex" 1\<close>) |
19741 | 267 |
apply (erule lemma_1 [THEN spec, THEN mp]) |
268 |
apply assumption+ |
|
269 |
apply (simp add: executions_def reachable.reachable_0) |
|
270 |
||
62156 | 271 |
text \<open>\<open>corresp_ex\<close> is execution, Lemma 2\<close> |
62002 | 272 |
apply (tactic \<open>pair_tac @{context} "ex" 1\<close>) |
19741 | 273 |
apply (simp add: executions_def) |
62156 | 274 |
text \<open>start state\<close> |
19741 | 275 |
apply (rule conjI) |
276 |
apply (simp add: is_ref_map_def corresp_ex_def) |
|
62156 | 277 |
text \<open>\<open>is_execution_fragment\<close>\<close> |
19741 | 278 |
apply (erule lemma_2 [THEN spec, THEN mp]) |
279 |
apply (simp add: reachable.reachable_0) |
|
280 |
done |
|
281 |
||
282 |
||
62156 | 283 |
subsection \<open>Corollary: FAIR TRACE -- INCLUSION\<close> |
19741 | 284 |
|
285 |
lemma fininf: "(~inf_often P s) = fin_often P s" |
|
62156 | 286 |
by (auto simp: fin_often_def) |
19741 | 287 |
|
62156 | 288 |
lemma WF_alt: "is_wfair A W (s, ex) = |
289 |
(fin_often (\<lambda>x. \<not> Enabled A W (snd x)) ex \<longrightarrow> inf_often (\<lambda>x. fst x \<in> W) ex)" |
|
290 |
by (auto simp add: is_wfair_def fin_often_def) |
|
19741 | 291 |
|
62156 | 292 |
lemma WF_persistent: |
293 |
"is_wfair A W (s, ex) \<Longrightarrow> inf_often (\<lambda>x. Enabled A W (snd x)) ex \<Longrightarrow> |
|
294 |
en_persistent A W \<Longrightarrow> inf_often (\<lambda>x. fst x \<in> W) ex" |
|
295 |
apply (drule persistent) |
|
296 |
apply assumption |
|
297 |
apply (simp add: WF_alt) |
|
298 |
apply auto |
|
299 |
done |
|
19741 | 300 |
|
62156 | 301 |
lemma fair_trace_inclusion: |
302 |
assumes "is_ref_map f C A" |
|
303 |
and "ext C = ext A" |
|
304 |
and "\<And>ex. ex \<in> executions C \<Longrightarrow> fair_ex C ex \<Longrightarrow> |
|
305 |
fair_ex A (corresp_ex A f ex)" |
|
306 |
shows "fairtraces C \<subseteq> fairtraces A" |
|
307 |
apply (insert assms) |
|
308 |
apply (simp add: fairtraces_def fairexecutions_def) |
|
309 |
apply auto |
|
310 |
apply (rule_tac x = "corresp_ex A f ex" in exI) |
|
311 |
apply auto |
|
19741 | 312 |
|
62156 | 313 |
text \<open>Traces coincide, Lemma 1\<close> |
62002 | 314 |
apply (tactic \<open>pair_tac @{context} "ex" 1\<close>) |
19741 | 315 |
apply (erule lemma_1 [THEN spec, THEN mp]) |
316 |
apply assumption+ |
|
317 |
apply (simp add: executions_def reachable.reachable_0) |
|
318 |
||
62156 | 319 |
text \<open>\<open>corresp_ex\<close> is execution, Lemma 2\<close> |
62002 | 320 |
apply (tactic \<open>pair_tac @{context} "ex" 1\<close>) |
19741 | 321 |
apply (simp add: executions_def) |
62156 | 322 |
text \<open>start state\<close> |
19741 | 323 |
apply (rule conjI) |
324 |
apply (simp add: is_ref_map_def corresp_ex_def) |
|
62156 | 325 |
text \<open>\<open>is_execution_fragment\<close>\<close> |
19741 | 326 |
apply (erule lemma_2 [THEN spec, THEN mp]) |
327 |
apply (simp add: reachable.reachable_0) |
|
62156 | 328 |
done |
19741 | 329 |
|
62156 | 330 |
lemma fair_trace_inclusion2: |
331 |
"inp C = inp A \<Longrightarrow> out C = out A \<Longrightarrow> is_fair_ref_map f C A \<Longrightarrow> |
|
332 |
fair_implements C A" |
|
333 |
apply (simp add: is_fair_ref_map_def fair_implements_def fairtraces_def fairexecutions_def) |
|
334 |
apply auto |
|
335 |
apply (rule_tac x = "corresp_ex A f ex" in exI) |
|
336 |
apply auto |
|
26359 | 337 |
|
62156 | 338 |
text \<open>Traces coincide, Lemma 1\<close> |
62002 | 339 |
apply (tactic \<open>pair_tac @{context} "ex" 1\<close>) |
19741 | 340 |
apply (erule lemma_1 [THEN spec, THEN mp]) |
341 |
apply (simp (no_asm) add: externals_def) |
|
342 |
apply (auto)[1] |
|
343 |
apply (simp add: executions_def reachable.reachable_0) |
|
344 |
||
62156 | 345 |
text \<open>\<open>corresp_ex\<close> is execution, Lemma 2\<close> |
62002 | 346 |
apply (tactic \<open>pair_tac @{context} "ex" 1\<close>) |
19741 | 347 |
apply (simp add: executions_def) |
62156 | 348 |
text \<open>start state\<close> |
19741 | 349 |
apply (rule conjI) |
350 |
apply (simp add: is_ref_map_def corresp_ex_def) |
|
62156 | 351 |
text \<open>\<open>is_execution_fragment\<close>\<close> |
19741 | 352 |
apply (erule lemma_2 [THEN spec, THEN mp]) |
353 |
apply (simp add: reachable.reachable_0) |
|
62156 | 354 |
done |
19741 | 355 |
|
17233 | 356 |
end |