1050
|
1 |
(* Title: HOL/IOA/ABP/Correctness.ML
|
|
2 |
ID: $Id$
|
|
3 |
Author: Tobias Nipkow & Olaf Mueller
|
|
4 |
Copyright 1995 TU Muenchen
|
|
5 |
|
|
6 |
*)
|
|
7 |
|
|
8 |
open Correctness; open Abschannel; open Abschannel_finite;
|
|
9 |
|
|
10 |
goal Abschannel.thy "(? x.x=P & Q(x)) = Q(P)";
|
|
11 |
by (fast_tac HOL_cs 1);
|
|
12 |
qed"exis_elim";
|
|
13 |
|
|
14 |
val abschannel_ss = action_ss addsimps
|
|
15 |
([srch_asig_def, rsch_asig_def, rsch_ioa_def, srch_ioa_def, ch_ioa_def,
|
|
16 |
ch_asig_def, srch_actions_def, rsch_actions_def, rename_def, asig_of_def,
|
|
17 |
actions_def, exis_elim, srch_trans_def, rsch_trans_def, ch_trans_def,
|
|
18 |
trans_of_def]
|
|
19 |
@ Option.option.simps @ act.simps @ asig_projections @ set_lemmas);
|
|
20 |
|
|
21 |
val abschannel_fin = [srch_fin_asig_def, rsch_fin_asig_def,
|
|
22 |
rsch_fin_ioa_def, srch_fin_ioa_def,
|
|
23 |
ch_fin_ioa_def,ch_fin_trans_def,ch_fin_asig_def];
|
|
24 |
val abschannel_fin_ss = abschannel_ss addsimps abschannel_fin;
|
|
25 |
|
|
26 |
val impl_ioas = [Sender.sender_ioa_def,Receiver.receiver_ioa_def];
|
|
27 |
val impl_trans = [Sender.sender_trans_def,Receiver.receiver_trans_def];
|
|
28 |
val impl_asigs = [Sender.sender_asig_def,Receiver.receiver_asig_def];
|
|
29 |
val impl_ss = merge_ss(action_ss,list_ss)
|
|
30 |
addcongs [let_weak_cong]
|
|
31 |
addsimps [Let_def, ioa_triple_proj, starts_of_par];
|
|
32 |
|
|
33 |
val env_ioas = [Env.env_ioa_def,Env.env_asig_def,Env.env_trans_def];
|
|
34 |
val hom_ioas = env_ioas @ impl_ioas @ impl_trans @ impl_asigs @
|
|
35 |
asig_projections @ set_lemmas;
|
|
36 |
val hom_ss = (impl_ss addsimps hom_ioas);
|
|
37 |
|
|
38 |
val red_ss = impl_ss addsimps [reduce_Nil,reduce_Cons];
|
|
39 |
val red_ss_ch = merge_ss(abschannel_fin_ss,red_ss);
|
|
40 |
|
|
41 |
|
|
42 |
(* auxiliary function *)
|
|
43 |
fun rotate n i = EVERY(replicate n (etac revcut_rl i));
|
|
44 |
|
|
45 |
(* lemmas about reduce *)
|
|
46 |
|
|
47 |
goal Correctness.thy "(reduce(l)=[]) = (l=[])";
|
|
48 |
by (rtac iffI 1);
|
|
49 |
by (subgoal_tac "(l~=[]) --> (reduce(l)~=[])" 1);
|
|
50 |
by (fast_tac HOL_cs 1);
|
|
51 |
by (List.list.induct_tac "l" 1);
|
|
52 |
by (simp_tac red_ss 1);
|
|
53 |
by (simp_tac red_ss 1);
|
|
54 |
by (rtac (expand_list_case RS iffD2) 1);
|
|
55 |
by (asm_full_simp_tac list_ss 1);
|
|
56 |
by (REPEAT (rtac allI 1));
|
|
57 |
by (rtac impI 1);
|
|
58 |
by (hyp_subst_tac 1);
|
|
59 |
by (rtac (expand_if RS ssubst) 1);
|
|
60 |
by (asm_full_simp_tac list_ss 1);
|
|
61 |
by (asm_full_simp_tac red_ss 1);
|
|
62 |
val l_iff_red_nil = result();
|
|
63 |
|
|
64 |
goal Correctness.thy "s~=[] --> hd(s)=hd(reduce(s))";
|
|
65 |
by (List.list.induct_tac "s" 1);
|
|
66 |
by (simp_tac red_ss 1);
|
|
67 |
by (case_tac "list =[]" 1);
|
|
68 |
by (asm_full_simp_tac red_ss 1);
|
|
69 |
(* main case *)
|
|
70 |
by (rotate 1 1);
|
|
71 |
by (asm_full_simp_tac list_ss 1);
|
|
72 |
by (simp_tac red_ss 1);
|
|
73 |
by (rtac (expand_list_case RS iffD2) 1);
|
|
74 |
by (asm_full_simp_tac list_ss 1);
|
|
75 |
by (REPEAT (rtac allI 1));
|
|
76 |
by (rtac impI 1);
|
|
77 |
by (rtac (expand_if RS ssubst) 1);
|
|
78 |
by (REPEAT(hyp_subst_tac 1));
|
|
79 |
by (etac subst 1);
|
|
80 |
by (simp_tac list_ss 1);
|
|
81 |
qed"hd_is_reduce_hd";
|
|
82 |
|
|
83 |
(* to be used in the following Lemma *)
|
|
84 |
goal Correctness.thy "l~=[] --> reverse(reduce(l))~=[]";
|
|
85 |
by (List.list.induct_tac "l" 1);
|
|
86 |
by (simp_tac red_ss 1);
|
|
87 |
by (case_tac "list =[]" 1);
|
|
88 |
by (asm_full_simp_tac (red_ss addsimps [reverse_Cons]) 1);
|
|
89 |
(* main case *)
|
|
90 |
by (rotate 1 1);
|
|
91 |
by (asm_full_simp_tac red_ss 1);
|
|
92 |
by (cut_inst_tac [("l","list")] cons_not_nil 1);
|
|
93 |
by (asm_full_simp_tac list_ss 1);
|
|
94 |
by (REPEAT (etac exE 1));
|
|
95 |
by (asm_simp_tac list_ss 1);
|
|
96 |
by (rtac (expand_if RS ssubst) 1);
|
|
97 |
by (hyp_subst_tac 1);
|
|
98 |
by (asm_full_simp_tac (list_ss addsimps [reverse_Cons]) 1);
|
|
99 |
qed"rev_red_not_nil";
|
|
100 |
|
|
101 |
(* shows applicability of the induction hypothesis of the following Lemma 1 *)
|
|
102 |
goal Correctness.thy "!!l.[| l~=[] |] ==> \
|
|
103 |
\ hd(reverse(reduce(a#l))) = hd(reverse(reduce(l)))";
|
|
104 |
by (simp_tac red_ss 1);
|
|
105 |
by (rtac (expand_list_case RS iffD2) 1);
|
|
106 |
by (asm_full_simp_tac list_ss 1);
|
|
107 |
by (REPEAT (rtac allI 1));
|
|
108 |
by (rtac impI 1);
|
|
109 |
by (rtac (expand_if RS ssubst) 1);
|
|
110 |
by (asm_full_simp_tac (list_ss addsimps [reverse_Cons,hd_append,
|
|
111 |
(rev_red_not_nil RS mp)]) 1);
|
|
112 |
qed"last_ind_on_first";
|
|
113 |
|
|
114 |
|
|
115 |
(* Main Lemma 1 for S_pkt in showing that reduce is refinement *)
|
|
116 |
goal Correctness.thy
|
|
117 |
"if x=hd(reverse(reduce(l))) & reduce(l)~=[] then \
|
|
118 |
\ reduce(l@[x])=reduce(l) else \
|
|
119 |
\ reduce(l@[x])=reduce(l)@[x]";
|
|
120 |
by (rtac (expand_if RS ssubst) 1);
|
|
121 |
by (rtac conjI 1);
|
|
122 |
(* --> *)
|
|
123 |
by (List.list.induct_tac "l" 1);
|
|
124 |
by (simp_tac red_ss 1);
|
|
125 |
by (case_tac "list=[]" 1);
|
|
126 |
by (asm_full_simp_tac (red_ss addsimps [reverse_Nil,reverse_Cons]) 1);
|
|
127 |
by (rtac impI 1);
|
|
128 |
by (simp_tac red_ss 1);
|
|
129 |
by (cut_inst_tac [("l","list")] cons_not_nil 1);
|
|
130 |
by (asm_full_simp_tac impl_ss 1);
|
|
131 |
by (REPEAT (etac exE 1));
|
|
132 |
by (hyp_subst_tac 1);
|
|
133 |
by (asm_full_simp_tac (impl_ss addsimps [last_ind_on_first,l_iff_red_nil]) 1);
|
|
134 |
(* <-- *)
|
|
135 |
by (simp_tac (red_ss addsimps [and_de_morgan_and_absorbe,l_iff_red_nil]) 1);
|
|
136 |
by (List.list.induct_tac "l" 1);
|
|
137 |
by (simp_tac red_ss 1);
|
|
138 |
by (case_tac "list=[]" 1);
|
|
139 |
by (asm_full_simp_tac (red_ss addsimps [reverse_Nil,reverse_Cons]) 1);
|
|
140 |
by (rtac (expand_if RS ssubst) 1);
|
|
141 |
by (fast_tac HOL_cs 1);
|
|
142 |
by (rtac impI 1);
|
|
143 |
by (simp_tac red_ss 1);
|
|
144 |
by (cut_inst_tac [("l","list")] cons_not_nil 1);
|
|
145 |
by (asm_full_simp_tac impl_ss 1);
|
|
146 |
by (REPEAT (etac exE 1));
|
|
147 |
by (hyp_subst_tac 1);
|
|
148 |
by (asm_full_simp_tac (impl_ss addsimps [last_ind_on_first,l_iff_red_nil]) 1);
|
|
149 |
by (rtac (expand_if RS ssubst) 1);
|
|
150 |
by (rtac (expand_if RS ssubst) 1);
|
|
151 |
by (asm_full_simp_tac impl_ss 1);
|
|
152 |
qed"reduce_hd";
|
|
153 |
|
|
154 |
|
|
155 |
(* Main Lemma 2 for R_pkt in showing that reduce is refinement *)
|
|
156 |
goal Correctness.thy
|
|
157 |
"!! s. [| s~=[] |] ==> \
|
|
158 |
\ if hd(s)=hd(tl(s)) & tl(s)~=[] then \
|
|
159 |
\ reduce(tl(s))=reduce(s) else \
|
|
160 |
\ reduce(tl(s))=tl(reduce(s))";
|
|
161 |
by (cut_inst_tac [("l","s")] cons_not_nil 1);
|
|
162 |
by (asm_full_simp_tac red_ss 1);
|
|
163 |
by (REPEAT (etac exE 1));
|
|
164 |
by (asm_full_simp_tac red_ss 1);
|
|
165 |
by (rtac (expand_if RS ssubst) 1);
|
|
166 |
by (rtac conjI 1);
|
|
167 |
by (simp_tac (red_ss addsimps [and_de_morgan_and_absorbe]) 2);
|
|
168 |
by (ALLGOALS (EVERY'[rtac impI,etac conjE,cut_inst_tac [("l","xs")] cons_not_nil,asm_full_simp_tac red_ss]));
|
|
169 |
by (REPEAT (etac exE 1));
|
|
170 |
by (REPEAT (etac exE 2));
|
|
171 |
by (hyp_subst_tac 2);
|
|
172 |
by (ALLGOALS (asm_full_simp_tac red_ss));
|
|
173 |
val reduce_tl =result();
|
|
174 |
|
|
175 |
|
|
176 |
|
|
177 |
goal Correctness.thy
|
|
178 |
"is_weak_pmap reduce ch_ioa ch_fin_ioa";
|
|
179 |
by (simp_tac (red_ss addsimps [Solve.is_weak_pmap_def]) 1);
|
|
180 |
by (rtac conjI 1);
|
|
181 |
(* start states *)
|
|
182 |
by (simp_tac red_ss_ch 1);
|
|
183 |
br ballI 1;
|
|
184 |
by (asm_full_simp_tac red_ss_ch 1);
|
|
185 |
(* main-part *)
|
|
186 |
by (REPEAT (rtac allI 1));
|
|
187 |
by (rtac imp_conj_lemma 1);
|
|
188 |
by (act.induct_tac "a" 1);
|
|
189 |
(* 2 cases *)
|
|
190 |
by (ALLGOALS (simp_tac (red_ss_ch addsimps [externals_def])));
|
|
191 |
(* fst case *)
|
|
192 |
by (rtac impI 1);
|
|
193 |
by (rtac disjI2 1);
|
|
194 |
by (rtac reduce_hd 1);
|
|
195 |
(* snd case *)
|
|
196 |
by (rtac impI 1);
|
|
197 |
by (REPEAT (etac conjE 1));
|
|
198 |
by (etac disjE 1);
|
|
199 |
by (asm_full_simp_tac (red_ss addsimps [l_iff_red_nil]) 1);
|
|
200 |
by (etac (hd_is_reduce_hd RS mp) 1);
|
|
201 |
by (asm_full_simp_tac (red_ss addsimps [l_iff_red_nil]) 1);
|
|
202 |
by (rtac conjI 1);
|
|
203 |
by (etac (hd_is_reduce_hd RS mp) 1);
|
|
204 |
by (rtac (bool_if_impl_or RS mp) 1);
|
|
205 |
by (etac reduce_tl 1);
|
|
206 |
qed"channel_abstraction";
|
|
207 |
|
|
208 |
goal Correctness.thy
|
|
209 |
"is_weak_pmap reduce srch_ioa srch_fin_ioa";
|
|
210 |
by (simp_tac (list_ss addsimps [srch_fin_ioa_def,rsch_fin_ioa_def,
|
|
211 |
srch_ioa_def,rsch_ioa_def,rename_through_pmap,channel_abstraction]) 1);
|
|
212 |
qed"sender_abstraction";
|
|
213 |
|
|
214 |
goal Correctness.thy
|
|
215 |
"is_weak_pmap reduce rsch_ioa rsch_fin_ioa";
|
|
216 |
by (simp_tac (list_ss addsimps [srch_fin_ioa_def,rsch_fin_ioa_def,
|
|
217 |
srch_ioa_def,rsch_ioa_def,rename_through_pmap,channel_abstraction]) 1);
|
|
218 |
qed"receiver_abstraction";
|
|
219 |
|
|
220 |
|
|
221 |
(* 3 thms that do not hold generally! The lucky restriction here is
|
|
222 |
the absence of internal actions. *)
|
|
223 |
goal Correctness.thy
|
|
224 |
"is_weak_pmap (%id.id) sender_ioa sender_ioa";
|
|
225 |
by (simp_tac (red_ss addsimps [Solve.is_weak_pmap_def]) 1);
|
|
226 |
by (rtac conjI 1);
|
|
227 |
(* start states *)
|
|
228 |
br ballI 1;
|
|
229 |
by (simp_tac red_ss_ch 1);
|
|
230 |
(* main-part *)
|
|
231 |
by (REPEAT (rtac allI 1));
|
|
232 |
by (rtac imp_conj_lemma 1);
|
|
233 |
by (Action.action.induct_tac "a" 1);
|
|
234 |
(* 7 cases *)
|
|
235 |
by (ALLGOALS (simp_tac ((hom_ss addsimps [externals_def]) setloop (split_tac [expand_if]))));
|
|
236 |
qed"sender_unchanged";
|
|
237 |
|
|
238 |
(* 2 copies of before *)
|
|
239 |
goal Correctness.thy
|
|
240 |
"is_weak_pmap (%id.id) receiver_ioa receiver_ioa";
|
|
241 |
by (simp_tac (red_ss addsimps [Solve.is_weak_pmap_def]) 1);
|
|
242 |
by (rtac conjI 1);
|
|
243 |
(* start states *)
|
|
244 |
br ballI 1;
|
|
245 |
by (simp_tac red_ss_ch 1);
|
|
246 |
(* main-part *)
|
|
247 |
by (REPEAT (rtac allI 1));
|
|
248 |
by (rtac imp_conj_lemma 1);
|
|
249 |
by (Action.action.induct_tac "a" 1);
|
|
250 |
(* 7 cases *)
|
|
251 |
by (ALLGOALS (simp_tac ((hom_ss addsimps [externals_def]) setloop (split_tac [expand_if]))));
|
|
252 |
qed"receiver_unchanged";
|
|
253 |
|
|
254 |
goal Correctness.thy
|
|
255 |
"is_weak_pmap (%id.id) env_ioa env_ioa";
|
|
256 |
by (simp_tac (red_ss addsimps [Solve.is_weak_pmap_def]) 1);
|
|
257 |
by (rtac conjI 1);
|
|
258 |
(* start states *)
|
|
259 |
br ballI 1;
|
|
260 |
by (simp_tac red_ss_ch 1);
|
|
261 |
(* main-part *)
|
|
262 |
by (REPEAT (rtac allI 1));
|
|
263 |
by (rtac imp_conj_lemma 1);
|
|
264 |
by (Action.action.induct_tac "a" 1);
|
|
265 |
(* 7 cases *)
|
|
266 |
by (ALLGOALS (simp_tac ((hom_ss addsimps [externals_def]) setloop (split_tac [expand_if]))));
|
|
267 |
qed"env_unchanged";
|
|
268 |
|
|
269 |
|
|
270 |
goal Correctness.thy "compat_ioas srch_ioa rsch_ioa";
|
|
271 |
by (simp_tac (red_ss_ch addsimps [compat_ioas_def,compat_asigs_def,Int_def,empty_def]) 1);
|
|
272 |
by (rtac set_ext 1);
|
|
273 |
by (Action.action.induct_tac "x" 1);
|
|
274 |
by (ALLGOALS(simp_tac red_ss_ch));
|
|
275 |
val compat_single_ch = result();
|
|
276 |
|
|
277 |
(* totally the same as before *)
|
|
278 |
goal Correctness.thy "compat_ioas srch_fin_ioa rsch_fin_ioa";
|
|
279 |
by (simp_tac (red_ss_ch addsimps [compat_ioas_def,compat_asigs_def,Int_def,empty_def]) 1);
|
|
280 |
by (rtac set_ext 1);
|
|
281 |
by (Action.action.induct_tac "x" 1);
|
|
282 |
by (ALLGOALS(simp_tac red_ss_ch));
|
|
283 |
val compat_single_fin_ch = result();
|
|
284 |
|
|
285 |
goal Correctness.thy "compat_ioas receiver_ioa (srch_ioa || rsch_ioa)";
|
|
286 |
by (simp_tac (hom_ss addsimps [empty_def,compat_ioas_def,compat_asigs_def,asig_of_par,asig_comp_def,actions_def,Int_def]) 1);
|
|
287 |
by (simp_tac red_ss_ch 1);
|
|
288 |
by (rtac set_ext 1);
|
|
289 |
by (Action.action.induct_tac "x" 1);
|
|
290 |
by (ALLGOALS(simp_tac red_ss_ch));
|
|
291 |
by (ALLGOALS(simp_tac (red_ss addsimps [insert_def,Un_def])));
|
|
292 |
by (ALLGOALS(simp_tac (red_ss_ch addsimps [de_morgan,singleton_set])));
|
|
293 |
val compat_rec =result();
|
|
294 |
|
|
295 |
(* 5 proofs totally the same as before *)
|
|
296 |
goal Correctness.thy "compat_ioas receiver_ioa (srch_fin_ioa || rsch_fin_ioa)";
|
|
297 |
by (simp_tac (hom_ss addsimps [empty_def,compat_ioas_def,compat_asigs_def,asig_of_par,asig_comp_def,actions_def,Int_def]) 1);
|
|
298 |
by (simp_tac red_ss_ch 1);
|
|
299 |
by (rtac set_ext 1);
|
|
300 |
by (Action.action.induct_tac "x" 1);
|
|
301 |
by (ALLGOALS(simp_tac red_ss_ch));
|
|
302 |
by (ALLGOALS(simp_tac (red_ss addsimps [insert_def,Un_def])));
|
|
303 |
by (ALLGOALS(simp_tac (red_ss_ch addsimps [de_morgan,singleton_set])));
|
|
304 |
val compat_rec_fin =result();
|
|
305 |
|
|
306 |
goal Correctness.thy "compat_ioas sender_ioa \
|
|
307 |
\ (receiver_ioa || srch_ioa || rsch_ioa)";
|
|
308 |
by (simp_tac (hom_ss addsimps [empty_def,compat_ioas_def,compat_asigs_def,asig_of_par,asig_comp_def,actions_def,Int_def]) 1);
|
|
309 |
by (simp_tac red_ss_ch 1);
|
|
310 |
by (rtac set_ext 1);
|
|
311 |
by (Action.action.induct_tac "x" 1);
|
|
312 |
by (ALLGOALS(simp_tac red_ss_ch));
|
|
313 |
by (ALLGOALS(simp_tac (red_ss addsimps [insert_def,Un_def])));
|
|
314 |
by (ALLGOALS(simp_tac (red_ss_ch addsimps [de_morgan,singleton_set])));
|
|
315 |
val compat_sen=result();
|
|
316 |
|
|
317 |
goal Correctness.thy "compat_ioas sender_ioa\
|
|
318 |
\ (receiver_ioa || srch_fin_ioa || rsch_fin_ioa)";
|
|
319 |
by (simp_tac (hom_ss addsimps [empty_def, compat_ioas_def,compat_asigs_def,asig_of_par,asig_comp_def,actions_def,Int_def]) 1);
|
|
320 |
by (simp_tac red_ss_ch 1);
|
|
321 |
by (rtac set_ext 1);
|
|
322 |
by (Action.action.induct_tac "x" 1);
|
|
323 |
by (ALLGOALS(simp_tac red_ss_ch));
|
|
324 |
by (ALLGOALS(simp_tac (red_ss addsimps [insert_def,Un_def])));
|
|
325 |
by (ALLGOALS(simp_tac (red_ss_ch addsimps [de_morgan,singleton_set])));
|
|
326 |
val compat_sen_fin =result();
|
|
327 |
|
|
328 |
goal Correctness.thy "compat_ioas env_ioa\
|
|
329 |
\ (sender_ioa || receiver_ioa || srch_ioa || rsch_ioa)";
|
|
330 |
by (simp_tac (hom_ss addsimps [empty_def,compat_ioas_def,compat_asigs_def,asig_of_par,asig_comp_def,actions_def,Int_def]) 1);
|
|
331 |
by (simp_tac red_ss_ch 1);
|
|
332 |
by (rtac set_ext 1);
|
|
333 |
by (Action.action.induct_tac "x" 1);
|
|
334 |
by (ALLGOALS(simp_tac red_ss_ch));
|
|
335 |
by (ALLGOALS(simp_tac (red_ss addsimps [insert_def,Un_def])));
|
|
336 |
by (ALLGOALS(simp_tac (red_ss_ch addsimps [de_morgan,singleton_set])));
|
|
337 |
val compat_env=result();
|
|
338 |
|
|
339 |
goal Correctness.thy "compat_ioas env_ioa\
|
|
340 |
\ (sender_ioa || receiver_ioa || srch_fin_ioa || rsch_fin_ioa)";
|
|
341 |
by (simp_tac (hom_ss addsimps [empty_def,compat_ioas_def,compat_asigs_def,asig_of_par,asig_comp_def,actions_def,Int_def]) 1);
|
|
342 |
by (simp_tac red_ss_ch 1);
|
|
343 |
by (rtac set_ext 1);
|
|
344 |
by (Action.action.induct_tac "x" 1);
|
|
345 |
by (ALLGOALS(simp_tac red_ss_ch));
|
|
346 |
by (ALLGOALS(simp_tac (red_ss addsimps [insert_def,Un_def])));
|
|
347 |
by (ALLGOALS(simp_tac (red_ss_ch addsimps [de_morgan,singleton_set])));
|
|
348 |
val compat_env_fin=result();
|
|
349 |
|
|
350 |
|
|
351 |
(* lemmata about externals of channels *)
|
|
352 |
goal Correctness.thy
|
|
353 |
"externals(asig_of(srch_fin_ioa)) = externals(asig_of(srch_ioa)) & \
|
|
354 |
\ externals(asig_of(rsch_fin_ioa)) = externals(asig_of(rsch_ioa))";
|
|
355 |
by (simp_tac (red_ss_ch addsimps [externals_def]) 1);
|
|
356 |
val ext_single_ch = result();
|
|
357 |
|
|
358 |
goal Correctness.thy "is_weak_pmap \
|
|
359 |
\ (%p.(fst(p),(fst(snd(p)),(fst(snd(snd(p))), \
|
|
360 |
\ (reduce(fst(snd(snd(snd(p))))),reduce(snd(snd(snd(snd(p)))))))))) \
|
|
361 |
\ (env_ioa || impl_ioa) (env_ioa || impl_fin_ioa)";
|
|
362 |
by (simp_tac (impl_ss addsimps [Impl.impl_def,Impl_finite.impl_fin_def]) 1);
|
|
363 |
by(REPEAT(EVERY[rtac fxg_is_weak_pmap_of_product_IOA 1,
|
|
364 |
simp_tac (red_ss addsimps [env_unchanged,sender_unchanged,
|
|
365 |
receiver_unchanged,sender_abstraction,receiver_abstraction]) 1,
|
|
366 |
rtac conjI 1]));
|
|
367 |
by (ALLGOALS(simp_tac (list_ss addsimps [externals_of_par,ext_single_ch,
|
|
368 |
compat_single_ch,compat_single_fin_ch,compat_rec,compat_rec_fin,
|
|
369 |
compat_sen,compat_sen_fin,compat_env,compat_env_fin])));
|
|
370 |
qed "system_refinement";
|
|
371 |
|