author | wenzelm |
Wed, 30 Dec 2015 22:09:44 +0100 | |
changeset 62004 | 8c6226d88ced |
parent 62002 | f1599e98c4d0 |
permissions | -rw-r--r-- |
42151 | 1 |
(* Title: HOL/HOLCF/IOA/meta_theory/LiveIOA.thy |
40945 | 2 |
Author: Olaf Müller |
17233 | 3 |
*) |
4559 | 4 |
|
62002 | 5 |
section \<open>Live I/O automata -- specified by temproal formulas\<close> |
4559 | 6 |
|
17233 | 7 |
theory LiveIOA |
8 |
imports TLS |
|
9 |
begin |
|
10 |
||
36452 | 11 |
default_sort type |
4559 | 12 |
|
41476 | 13 |
type_synonym |
17233 | 14 |
('a, 's) live_ioa = "('a,'s)ioa * ('a,'s)ioa_temp" |
4559 | 15 |
|
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
16 |
definition |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
17 |
validLIOA :: "('a,'s)live_ioa => ('a,'s)ioa_temp => bool" where |
62000 | 18 |
"validLIOA AL P = validIOA (fst AL) ((snd AL) \<^bold>\<longrightarrow> P)" |
4559 | 19 |
|
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
20 |
definition |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
21 |
WF :: "('a,'s)ioa => 'a set => ('a,'s)ioa_temp" where |
62004 | 22 |
"WF A acts = (\<diamond>\<box>\<langle>%(s,a,t). Enabled A acts s\<rangle> \<^bold>\<longrightarrow> \<box>\<diamond>\<langle>xt2 (plift (%a. a : acts))\<rangle>)" |
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
23 |
definition |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
24 |
SF :: "('a,'s)ioa => 'a set => ('a,'s)ioa_temp" where |
62004 | 25 |
"SF A acts = (\<box>\<diamond>\<langle>%(s,a,t). Enabled A acts s\<rangle> \<^bold>\<longrightarrow> \<box>\<diamond>\<langle>xt2 (plift (%a. a : acts))\<rangle>)" |
4559 | 26 |
|
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
27 |
definition |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
28 |
liveexecutions :: "('a,'s)live_ioa => ('a,'s)execution set" where |
62000 | 29 |
"liveexecutions AP = {exec. exec : executions (fst AP) & (exec \<TTurnstile> (snd AP))}" |
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
30 |
definition |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
31 |
livetraces :: "('a,'s)live_ioa => 'a trace set" where |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
32 |
"livetraces AP = {mk_trace (fst AP)$(snd ex) | ex. ex:liveexecutions AP}" |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
33 |
definition |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
34 |
live_implements :: "('a,'s1)live_ioa => ('a,'s2)live_ioa => bool" where |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
35 |
"live_implements CL AM = ((inp (fst CL) = inp (fst AM)) & |
4559 | 36 |
(out (fst CL) = out (fst AM)) & |
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
37 |
livetraces CL <= livetraces AM)" |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
38 |
definition |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
39 |
is_live_ref_map :: "('s1 => 's2) => ('a,'s1)live_ioa => ('a,'s2)live_ioa => bool" where |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
40 |
"is_live_ref_map f CL AM = |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
41 |
(is_ref_map f (fst CL ) (fst AM) & |
62000 | 42 |
(! exec : executions (fst CL). (exec \<TTurnstile> (snd CL)) --> |
43 |
((corresp_ex (fst AM) f exec) \<TTurnstile> (snd AM))))" |
|
4559 | 44 |
|
19741 | 45 |
|
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
46 |
lemma live_implements_trans: |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
47 |
"!!LC. [| live_implements (A,LA) (B,LB); live_implements (B,LB) (C,LC) |] |
19741 | 48 |
==> live_implements (A,LA) (C,LC)" |
49 |
apply (unfold live_implements_def) |
|
50 |
apply auto |
|
51 |
done |
|
52 |
||
53 |
||
54 |
subsection "Correctness of live refmap" |
|
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
55 |
|
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
56 |
lemma live_implements: "[| inp(C)=inp(A); out(C)=out(A); |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
57 |
is_live_ref_map f (C,M) (A,L) |] |
19741 | 58 |
==> live_implements (C,M) (A,L)" |
59 |
apply (simp add: is_live_ref_map_def live_implements_def livetraces_def liveexecutions_def) |
|
26359 | 60 |
apply auto |
19741 | 61 |
apply (rule_tac x = "corresp_ex A f ex" in exI) |
26359 | 62 |
apply auto |
19741 | 63 |
(* Traces coincide, Lemma 1 *) |
62002 | 64 |
apply (tactic \<open>pair_tac @{context} "ex" 1\<close>) |
19741 | 65 |
apply (erule lemma_1 [THEN spec, THEN mp]) |
66 |
apply (simp (no_asm) add: externals_def) |
|
67 |
apply (auto)[1] |
|
68 |
apply (simp add: executions_def reachable.reachable_0) |
|
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
69 |
|
19741 | 70 |
(* corresp_ex is execution, Lemma 2 *) |
62002 | 71 |
apply (tactic \<open>pair_tac @{context} "ex" 1\<close>) |
19741 | 72 |
apply (simp add: executions_def) |
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
73 |
(* start state *) |
19741 | 74 |
apply (rule conjI) |
75 |
apply (simp add: is_ref_map_def corresp_ex_def) |
|
76 |
(* is-execution-fragment *) |
|
77 |
apply (erule lemma_2 [THEN spec, THEN mp]) |
|
78 |
apply (simp add: reachable.reachable_0) |
|
79 |
||
80 |
done |
|
4559 | 81 |
|
17233 | 82 |
end |