|
4565
|
1 |
(* Title: HOLCF/IOA/meta_theory/Simulations.thy
|
|
|
2 |
ID: $Id$
|
|
12218
|
3 |
Author: Olaf Müller
|
|
4565
|
4 |
*)
|
|
|
5 |
|
|
17233
|
6 |
header {* Simulations in HOLCF/IOA *}
|
|
4565
|
7 |
|
|
17233
|
8 |
theory Simulations
|
|
|
9 |
imports RefCorrectness
|
|
|
10 |
begin
|
|
|
11 |
|
|
|
12 |
defaultsort type
|
|
4565
|
13 |
|
|
|
14 |
consts
|
|
|
15 |
|
|
|
16 |
is_simulation ::"[('s1 * 's2)set,('a,'s1)ioa,('a,'s2)ioa] => bool"
|
|
|
17 |
is_backward_simulation ::"[('s1 * 's2)set,('a,'s1)ioa,('a,'s2)ioa] => bool"
|
|
|
18 |
is_forw_back_simulation ::"[('s1 * 's2 set)set,('a,'s1)ioa,('a,'s2)ioa] => bool"
|
|
|
19 |
is_back_forw_simulation ::"[('s1 * 's2 set)set,('a,'s1)ioa,('a,'s2)ioa] => bool"
|
|
|
20 |
is_history_relation ::"[('s1 * 's2)set,('a,'s1)ioa,('a,'s2)ioa] => bool"
|
|
|
21 |
is_prophecy_relation ::"[('s1 * 's2)set,('a,'s1)ioa,('a,'s2)ioa] => bool"
|
|
|
22 |
|
|
|
23 |
defs
|
|
|
24 |
|
|
17233
|
25 |
is_simulation_def:
|
|
|
26 |
"is_simulation R C A ==
|
|
|
27 |
(!s:starts_of C. R``{s} Int starts_of A ~= {}) &
|
|
|
28 |
(!s s' t a. reachable C s &
|
|
4565
|
29 |
s -a--C-> t &
|
|
17233
|
30 |
(s,s') : R
|
|
4565
|
31 |
--> (? t' ex. (t,t'):R & move A ex s' a t'))"
|
|
|
32 |
|
|
17233
|
33 |
is_backward_simulation_def:
|
|
|
34 |
"is_backward_simulation R C A ==
|
|
|
35 |
(!s:starts_of C. R``{s} <= starts_of A) &
|
|
|
36 |
(!s t t' a. reachable C s &
|
|
4565
|
37 |
s -a--C-> t &
|
|
17233
|
38 |
(t,t') : R
|
|
4565
|
39 |
--> (? ex s'. (s,s'):R & move A ex s' a t'))"
|
|
|
40 |
|
|
17233
|
41 |
is_forw_back_simulation_def:
|
|
|
42 |
"is_forw_back_simulation R C A ==
|
|
|
43 |
(!s:starts_of C. ? S'. (s,S'):R & S'<= starts_of A) &
|
|
|
44 |
(!s S' t a. reachable C s &
|
|
4565
|
45 |
s -a--C-> t &
|
|
17233
|
46 |
(s,S') : R
|
|
4565
|
47 |
--> (? T'. (t,T'):R & (! t':T'. ? s':S'. ? ex. move A ex s' a t')))"
|
|
|
48 |
|
|
17233
|
49 |
is_back_forw_simulation_def:
|
|
|
50 |
"is_back_forw_simulation R C A ==
|
|
|
51 |
(!s:starts_of C. ! S'. (s,S'):R --> S' Int starts_of A ~={}) &
|
|
|
52 |
(!s t T' a. reachable C s &
|
|
4565
|
53 |
s -a--C-> t &
|
|
17233
|
54 |
(t,T') : R
|
|
4565
|
55 |
--> (? S'. (s,S'):R & (! s':S'. ? t':T'. ? ex. move A ex s' a t')))"
|
|
|
56 |
|
|
17233
|
57 |
is_history_relation_def:
|
|
|
58 |
"is_history_relation R C A == is_simulation R C A &
|
|
4565
|
59 |
is_ref_map (%x.(@y. (x,y):(R^-1))) A C"
|
|
|
60 |
|
|
17233
|
61 |
is_prophecy_relation_def:
|
|
|
62 |
"is_prophecy_relation R C A == is_backward_simulation R C A &
|
|
4565
|
63 |
is_ref_map (%x.(@y. (x,y):(R^-1))) A C"
|
|
17233
|
64 |
|
|
19741
|
65 |
|
|
|
66 |
lemma set_non_empty: "(A~={}) = (? x. x:A)"
|
|
|
67 |
apply auto
|
|
|
68 |
done
|
|
|
69 |
|
|
|
70 |
lemma Int_non_empty: "(A Int B ~= {}) = (? x. x: A & x:B)"
|
|
|
71 |
apply (simp add: set_non_empty)
|
|
|
72 |
done
|
|
|
73 |
|
|
|
74 |
|
|
|
75 |
lemma Sim_start_convert:
|
|
|
76 |
"(R``{x} Int S ~= {}) = (? y. (x,y):R & y:S)"
|
|
|
77 |
apply (unfold Image_def)
|
|
|
78 |
apply (simp add: Int_non_empty)
|
|
|
79 |
done
|
|
|
80 |
|
|
|
81 |
declare Sim_start_convert [simp]
|
|
|
82 |
|
|
|
83 |
|
|
|
84 |
lemma ref_map_is_simulation:
|
|
|
85 |
"!! f. is_ref_map f C A ==> is_simulation {p. (snd p) = f (fst p)} C A"
|
|
|
86 |
|
|
|
87 |
apply (unfold is_ref_map_def is_simulation_def)
|
|
|
88 |
apply simp
|
|
|
89 |
done
|
|
17233
|
90 |
|
|
4565
|
91 |
end
|