author | paulson |
Wed, 13 Nov 1996 10:47:08 +0100 | |
changeset 2183 | 8d42a7bccf0b |
parent 2083 | b56425a385b9 |
child 2513 | d708d8cdc8e8 |
permissions | -rw-r--r-- |
1138 | 1 |
(* Title: HOL/IOA/example/Correctness.ML |
1050 | 2 |
ID: $Id$ |
1138 | 3 |
Author: Tobias Nipkow & Konrad Slind |
4 |
Copyright 1994 TU Muenchen |
|
1050 | 5 |
|
6 |
*) |
|
7 |
||
8 |
open Correctness; open Abschannel; open Abschannel_finite; |
|
1138 | 9 |
open Impl; open Impl_finite; |
1050 | 10 |
|
11 |
goal Abschannel.thy "(? x.x=P & Q(x)) = Q(P)"; |
|
1894 | 12 |
by (Fast_tac 1); |
1050 | 13 |
qed"exis_elim"; |
14 |
||
1328 | 15 |
Delsimps [split_paired_All]; |
1266 | 16 |
Addsimps |
1050 | 17 |
([srch_asig_def, rsch_asig_def, rsch_ioa_def, srch_ioa_def, ch_ioa_def, |
18 |
ch_asig_def, srch_actions_def, rsch_actions_def, rename_def, asig_of_def, |
|
19 |
actions_def, exis_elim, srch_trans_def, rsch_trans_def, ch_trans_def, |
|
1266 | 20 |
trans_of_def] @ asig_projections @ set_lemmas); |
1050 | 21 |
|
22 |
val abschannel_fin = [srch_fin_asig_def, rsch_fin_asig_def, |
|
23 |
rsch_fin_ioa_def, srch_fin_ioa_def, |
|
24 |
ch_fin_ioa_def,ch_fin_trans_def,ch_fin_asig_def]; |
|
1266 | 25 |
Addsimps abschannel_fin; |
1050 | 26 |
|
27 |
val impl_ioas = [Sender.sender_ioa_def,Receiver.receiver_ioa_def]; |
|
28 |
val impl_trans = [Sender.sender_trans_def,Receiver.receiver_trans_def]; |
|
29 |
val impl_asigs = [Sender.sender_asig_def,Receiver.receiver_asig_def]; |
|
1266 | 30 |
Addcongs [let_weak_cong]; |
31 |
Addsimps [Let_def, ioa_triple_proj, starts_of_par]; |
|
1050 | 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; |
|
1266 | 36 |
Addsimps hom_ioas; |
1050 | 37 |
|
1266 | 38 |
Addsimps [reduce_Nil,reduce_Cons]; |
1050 | 39 |
|
1138 | 40 |
|
41 |
||
1050 | 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); |
|
1894 | 50 |
by (Fast_tac 1); |
1050 | 51 |
by (List.list.induct_tac "l" 1); |
1266 | 52 |
by (Simp_tac 1); |
53 |
by (Simp_tac 1); |
|
1050 | 54 |
by (rtac (expand_list_case RS iffD2) 1); |
1266 | 55 |
by (Asm_full_simp_tac 1); |
1050 | 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); |
|
1266 | 60 |
by (Asm_full_simp_tac 1); |
61 |
by (Asm_full_simp_tac 1); |
|
1050 | 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); |
|
1266 | 66 |
by (Simp_tac 1); |
1050 | 67 |
by (case_tac "list =[]" 1); |
1266 | 68 |
by (Asm_full_simp_tac 1); |
1050 | 69 |
(* main case *) |
70 |
by (rotate 1 1); |
|
71 |
by (asm_full_simp_tac list_ss 1); |
|
1266 | 72 |
by (Simp_tac 1); |
1050 | 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); |
|
1266 | 80 |
by (Simp_tac 1); |
1050 | 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); |
|
1266 | 86 |
by (Simp_tac 1); |
1050 | 87 |
by (case_tac "list =[]" 1); |
1266 | 88 |
by (asm_full_simp_tac (!simpset addsimps [reverse_Cons]) 1); |
1050 | 89 |
(* main case *) |
90 |
by (rotate 1 1); |
|
1266 | 91 |
by (Asm_full_simp_tac 1); |
1050 | 92 |
by (cut_inst_tac [("l","list")] cons_not_nil 1); |
1266 | 93 |
by (Asm_full_simp_tac 1); |
1050 | 94 |
by (REPEAT (etac exE 1)); |
1266 | 95 |
by (Asm_simp_tac 1); |
1050 | 96 |
by (rtac (expand_if RS ssubst) 1); |
97 |
by (hyp_subst_tac 1); |
|
1266 | 98 |
by (asm_full_simp_tac (!simpset addsimps [reverse_Cons]) 1); |
1050 | 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)))"; |
|
1266 | 104 |
by (Simp_tac 1); |
1050 | 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 |
||
1266 | 114 |
val impl_ss = !simpset delsimps [reduce_Cons]; |
1050 | 115 |
|
116 |
(* Main Lemma 1 for S_pkt in showing that reduce is refinement *) |
|
117 |
goal Correctness.thy |
|
118 |
"if x=hd(reverse(reduce(l))) & reduce(l)~=[] then \ |
|
119 |
\ reduce(l@[x])=reduce(l) else \ |
|
120 |
\ reduce(l@[x])=reduce(l)@[x]"; |
|
121 |
by (rtac (expand_if RS ssubst) 1); |
|
122 |
by (rtac conjI 1); |
|
123 |
(* --> *) |
|
124 |
by (List.list.induct_tac "l" 1); |
|
1266 | 125 |
by (Simp_tac 1); |
1050 | 126 |
by (case_tac "list=[]" 1); |
1266 | 127 |
by (asm_full_simp_tac (!simpset addsimps [reverse_Nil,reverse_Cons]) 1); |
1050 | 128 |
by (rtac impI 1); |
1266 | 129 |
by (Simp_tac 1); |
1050 | 130 |
by (cut_inst_tac [("l","list")] cons_not_nil 1); |
131 |
by (asm_full_simp_tac impl_ss 1); |
|
132 |
by (REPEAT (etac exE 1)); |
|
133 |
by (hyp_subst_tac 1); |
|
134 |
by (asm_full_simp_tac (impl_ss addsimps [last_ind_on_first,l_iff_red_nil]) 1); |
|
135 |
(* <-- *) |
|
1266 | 136 |
by (simp_tac (!simpset addsimps [and_de_morgan_and_absorbe,l_iff_red_nil]) 1); |
1050 | 137 |
by (List.list.induct_tac "l" 1); |
1266 | 138 |
by (Simp_tac 1); |
1050 | 139 |
by (case_tac "list=[]" 1); |
2083
b56425a385b9
Tidied some proofs: changed needed for de Morgan laws
paulson
parents:
1894
diff
changeset
|
140 |
by (cut_inst_tac [("l","list")] cons_not_nil 2); |
b56425a385b9
Tidied some proofs: changed needed for de Morgan laws
paulson
parents:
1894
diff
changeset
|
141 |
by (asm_full_simp_tac (!simpset setloop split_tac [expand_if]) 1); |
b56425a385b9
Tidied some proofs: changed needed for de Morgan laws
paulson
parents:
1894
diff
changeset
|
142 |
by (auto_tac (!claset, |
b56425a385b9
Tidied some proofs: changed needed for de Morgan laws
paulson
parents:
1894
diff
changeset
|
143 |
impl_ss addsimps [last_ind_on_first,l_iff_red_nil] |
b56425a385b9
Tidied some proofs: changed needed for de Morgan laws
paulson
parents:
1894
diff
changeset
|
144 |
setloop split_tac [expand_if])); |
b56425a385b9
Tidied some proofs: changed needed for de Morgan laws
paulson
parents:
1894
diff
changeset
|
145 |
by (asm_full_simp_tac (!simpset setloop split_tac [expand_if]) 1); |
1050 | 146 |
qed"reduce_hd"; |
147 |
||
148 |
||
149 |
(* Main Lemma 2 for R_pkt in showing that reduce is refinement *) |
|
150 |
goal Correctness.thy |
|
151 |
"!! s. [| s~=[] |] ==> \ |
|
152 |
\ if hd(s)=hd(tl(s)) & tl(s)~=[] then \ |
|
153 |
\ reduce(tl(s))=reduce(s) else \ |
|
154 |
\ reduce(tl(s))=tl(reduce(s))"; |
|
155 |
by (cut_inst_tac [("l","s")] cons_not_nil 1); |
|
1266 | 156 |
by (Asm_full_simp_tac 1); |
1050 | 157 |
by (REPEAT (etac exE 1)); |
1266 | 158 |
by (Asm_full_simp_tac 1); |
1050 | 159 |
by (rtac (expand_if RS ssubst) 1); |
160 |
by (rtac conjI 1); |
|
1266 | 161 |
by (simp_tac (!simpset addsimps [and_de_morgan_and_absorbe]) 2); |
2083
b56425a385b9
Tidied some proofs: changed needed for de Morgan laws
paulson
parents:
1894
diff
changeset
|
162 |
by (Step_tac 1); |
b56425a385b9
Tidied some proofs: changed needed for de Morgan laws
paulson
parents:
1894
diff
changeset
|
163 |
by (ALLGOALS (cut_inst_tac [("l","xs")] cons_not_nil)); |
b56425a385b9
Tidied some proofs: changed needed for de Morgan laws
paulson
parents:
1894
diff
changeset
|
164 |
by (Auto_tac()); |
1050 | 165 |
val reduce_tl =result(); |
166 |
||
167 |
||
1138 | 168 |
(******************************************************************* |
169 |
C h a n n e l A b s t r a c t i o n |
|
170 |
*******************************************************************) |
|
1050 | 171 |
|
172 |
goal Correctness.thy |
|
173 |
"is_weak_pmap reduce ch_ioa ch_fin_ioa"; |
|
1266 | 174 |
by (simp_tac (!simpset addsimps [Solve.is_weak_pmap_def]) 1); |
1050 | 175 |
by (rtac conjI 1); |
1138 | 176 |
(* -------------- start states ----------------- *) |
1266 | 177 |
by (Simp_tac 1); |
1465 | 178 |
by (rtac ballI 1); |
1266 | 179 |
by (Asm_full_simp_tac 1); |
1138 | 180 |
(* ---------------- main-part ------------------- *) |
1050 | 181 |
by (REPEAT (rtac allI 1)); |
182 |
by (rtac imp_conj_lemma 1); |
|
183 |
by (act.induct_tac "a" 1); |
|
1138 | 184 |
(* ----------------- 2 cases ---------------------*) |
1266 | 185 |
by (ALLGOALS (simp_tac (!simpset addsimps [externals_def]))); |
1138 | 186 |
(* fst case --------------------------------------*) |
1050 | 187 |
by (rtac impI 1); |
188 |
by (rtac disjI2 1); |
|
189 |
by (rtac reduce_hd 1); |
|
1138 | 190 |
(* snd case --------------------------------------*) |
1050 | 191 |
by (rtac impI 1); |
192 |
by (REPEAT (etac conjE 1)); |
|
193 |
by (etac disjE 1); |
|
1266 | 194 |
by (asm_full_simp_tac (!simpset addsimps [l_iff_red_nil]) 1); |
1050 | 195 |
by (etac (hd_is_reduce_hd RS mp) 1); |
1266 | 196 |
by (asm_full_simp_tac (!simpset addsimps [l_iff_red_nil]) 1); |
1050 | 197 |
by (rtac conjI 1); |
198 |
by (etac (hd_is_reduce_hd RS mp) 1); |
|
199 |
by (rtac (bool_if_impl_or RS mp) 1); |
|
200 |
by (etac reduce_tl 1); |
|
201 |
qed"channel_abstraction"; |
|
202 |
||
203 |
goal Correctness.thy |
|
204 |
"is_weak_pmap reduce srch_ioa srch_fin_ioa"; |
|
205 |
by (simp_tac (list_ss addsimps [srch_fin_ioa_def,rsch_fin_ioa_def, |
|
206 |
srch_ioa_def,rsch_ioa_def,rename_through_pmap,channel_abstraction]) 1); |
|
207 |
qed"sender_abstraction"; |
|
208 |
||
209 |
goal Correctness.thy |
|
210 |
"is_weak_pmap reduce rsch_ioa rsch_fin_ioa"; |
|
211 |
by (simp_tac (list_ss addsimps [srch_fin_ioa_def,rsch_fin_ioa_def, |
|
212 |
srch_ioa_def,rsch_ioa_def,rename_through_pmap,channel_abstraction]) 1); |
|
213 |
qed"receiver_abstraction"; |
|
214 |
||
215 |
||
216 |
(* 3 thms that do not hold generally! The lucky restriction here is |
|
217 |
the absence of internal actions. *) |
|
218 |
goal Correctness.thy |
|
219 |
"is_weak_pmap (%id.id) sender_ioa sender_ioa"; |
|
1266 | 220 |
by (simp_tac (!simpset addsimps [Solve.is_weak_pmap_def]) 1); |
1778 | 221 |
by (TRY( |
222 |
(rtac conjI 1) THEN |
|
1673
d22110ddd0af
repaired critical proofs depending on the order inside non-confluent SimpSets
oheimb
parents:
1667
diff
changeset
|
223 |
(* start states *) |
1894 | 224 |
(Fast_tac 1))); |
1050 | 225 |
(* main-part *) |
226 |
by (REPEAT (rtac allI 1)); |
|
227 |
by (rtac imp_conj_lemma 1); |
|
228 |
by (Action.action.induct_tac "a" 1); |
|
229 |
(* 7 cases *) |
|
1266 | 230 |
by (ALLGOALS (simp_tac ((!simpset addsimps [externals_def]) setloop (split_tac [expand_if])))); |
1050 | 231 |
qed"sender_unchanged"; |
232 |
||
233 |
(* 2 copies of before *) |
|
234 |
goal Correctness.thy |
|
235 |
"is_weak_pmap (%id.id) receiver_ioa receiver_ioa"; |
|
1266 | 236 |
by (simp_tac (!simpset addsimps [Solve.is_weak_pmap_def]) 1); |
1778 | 237 |
by (TRY( |
238 |
(rtac conjI 1) THEN |
|
1673
d22110ddd0af
repaired critical proofs depending on the order inside non-confluent SimpSets
oheimb
parents:
1667
diff
changeset
|
239 |
(* start states *) |
1894 | 240 |
(Fast_tac 1))); |
1050 | 241 |
(* main-part *) |
242 |
by (REPEAT (rtac allI 1)); |
|
243 |
by (rtac imp_conj_lemma 1); |
|
244 |
by (Action.action.induct_tac "a" 1); |
|
245 |
(* 7 cases *) |
|
1266 | 246 |
by (ALLGOALS (simp_tac ((!simpset addsimps [externals_def]) setloop (split_tac [expand_if])))); |
1050 | 247 |
qed"receiver_unchanged"; |
248 |
||
249 |
goal Correctness.thy |
|
250 |
"is_weak_pmap (%id.id) env_ioa env_ioa"; |
|
1266 | 251 |
by (simp_tac (!simpset addsimps [Solve.is_weak_pmap_def]) 1); |
1778 | 252 |
by (TRY( |
253 |
(rtac conjI 1) THEN |
|
1673
d22110ddd0af
repaired critical proofs depending on the order inside non-confluent SimpSets
oheimb
parents:
1667
diff
changeset
|
254 |
(* start states *) |
1894 | 255 |
(Fast_tac 1))); |
1050 | 256 |
(* main-part *) |
257 |
by (REPEAT (rtac allI 1)); |
|
258 |
by (rtac imp_conj_lemma 1); |
|
259 |
by (Action.action.induct_tac "a" 1); |
|
260 |
(* 7 cases *) |
|
1266 | 261 |
by (ALLGOALS (simp_tac ((!simpset addsimps [externals_def]) setloop (split_tac [expand_if])))); |
1050 | 262 |
qed"env_unchanged"; |
263 |
||
1532 | 264 |
Delsimps [Collect_False_empty]; |
1050 | 265 |
|
266 |
goal Correctness.thy "compat_ioas srch_ioa rsch_ioa"; |
|
1266 | 267 |
by (simp_tac (!simpset addsimps [compat_ioas_def,compat_asigs_def,Int_def,empty_def]) 1); |
1050 | 268 |
by (rtac set_ext 1); |
269 |
by (Action.action.induct_tac "x" 1); |
|
1266 | 270 |
by (ALLGOALS(Simp_tac)); |
1050 | 271 |
val compat_single_ch = result(); |
272 |
||
273 |
(* totally the same as before *) |
|
274 |
goal Correctness.thy "compat_ioas srch_fin_ioa rsch_fin_ioa"; |
|
1266 | 275 |
by (simp_tac (!simpset addsimps [compat_ioas_def,compat_asigs_def,Int_def,empty_def]) 1); |
1050 | 276 |
by (rtac set_ext 1); |
277 |
by (Action.action.induct_tac "x" 1); |
|
1266 | 278 |
by (ALLGOALS(Simp_tac)); |
1050 | 279 |
val compat_single_fin_ch = result(); |
280 |
||
1266 | 281 |
val ss = |
282 |
!simpset delsimps ([trans_of_def, srch_asig_def,rsch_asig_def, |
|
283 |
asig_of_def, actions_def, srch_trans_def, rsch_trans_def, |
|
284 |
srch_ioa_def, srch_fin_ioa_def, rsch_fin_ioa_def, |
|
285 |
rsch_ioa_def, Sender.sender_trans_def, |
|
286 |
Receiver.receiver_trans_def] @ set_lemmas); |
|
287 |
||
1050 | 288 |
goal Correctness.thy "compat_ioas receiver_ioa (srch_ioa || rsch_ioa)"; |
1266 | 289 |
by (simp_tac (ss addsimps [empty_def,compat_ioas_def,compat_asigs_def, |
290 |
asig_of_par,asig_comp_def,actions_def, |
|
291 |
Int_def]) 1); |
|
292 |
by (Simp_tac 1); |
|
1050 | 293 |
by (rtac set_ext 1); |
294 |
by (Action.action.induct_tac "x" 1); |
|
1266 | 295 |
by (ALLGOALS Simp_tac); |
296 |
by (ALLGOALS(simp_tac (!simpset addsimps [insert_def,Un_def]))); |
|
297 |
by (ALLGOALS(simp_tac (!simpset addsimps [de_morgan,singleton_set]))); |
|
298 |
val compat_rec = result(); |
|
1050 | 299 |
|
300 |
(* 5 proofs totally the same as before *) |
|
301 |
goal Correctness.thy "compat_ioas receiver_ioa (srch_fin_ioa || rsch_fin_ioa)"; |
|
1266 | 302 |
by (simp_tac (ss addsimps [empty_def,compat_ioas_def,compat_asigs_def,asig_of_par,asig_comp_def,actions_def,Int_def]) 1); |
303 |
by (Simp_tac 1); |
|
1050 | 304 |
by (rtac set_ext 1); |
305 |
by (Action.action.induct_tac "x" 1); |
|
1266 | 306 |
by (ALLGOALS Simp_tac); |
307 |
by (ALLGOALS(simp_tac (!simpset addsimps [insert_def,Un_def]))); |
|
308 |
by (ALLGOALS(simp_tac (!simpset addsimps [de_morgan,singleton_set]))); |
|
1050 | 309 |
val compat_rec_fin =result(); |
310 |
||
311 |
goal Correctness.thy "compat_ioas sender_ioa \ |
|
312 |
\ (receiver_ioa || srch_ioa || rsch_ioa)"; |
|
1266 | 313 |
by (simp_tac (ss addsimps [empty_def,compat_ioas_def,compat_asigs_def,asig_of_par,asig_comp_def,actions_def,Int_def]) 1); |
314 |
by (Simp_tac 1); |
|
1050 | 315 |
by (rtac set_ext 1); |
316 |
by (Action.action.induct_tac "x" 1); |
|
1266 | 317 |
by (ALLGOALS(Simp_tac)); |
318 |
by (ALLGOALS(simp_tac (!simpset addsimps [insert_def,Un_def]))); |
|
319 |
by (ALLGOALS(simp_tac (!simpset addsimps [de_morgan,singleton_set]))); |
|
1050 | 320 |
val compat_sen=result(); |
321 |
||
322 |
goal Correctness.thy "compat_ioas sender_ioa\ |
|
323 |
\ (receiver_ioa || srch_fin_ioa || rsch_fin_ioa)"; |
|
1266 | 324 |
by (simp_tac (ss addsimps [empty_def, compat_ioas_def,compat_asigs_def,asig_of_par,asig_comp_def,actions_def,Int_def]) 1); |
325 |
by (Simp_tac 1); |
|
1050 | 326 |
by (rtac set_ext 1); |
327 |
by (Action.action.induct_tac "x" 1); |
|
1266 | 328 |
by (ALLGOALS(Simp_tac)); |
329 |
by (ALLGOALS(simp_tac (!simpset addsimps [insert_def,Un_def]))); |
|
330 |
by (ALLGOALS(simp_tac (!simpset addsimps [de_morgan,singleton_set]))); |
|
1050 | 331 |
val compat_sen_fin =result(); |
332 |
||
333 |
goal Correctness.thy "compat_ioas env_ioa\ |
|
334 |
\ (sender_ioa || receiver_ioa || srch_ioa || rsch_ioa)"; |
|
1266 | 335 |
by (simp_tac (ss addsimps [empty_def,compat_ioas_def,compat_asigs_def,asig_of_par,asig_comp_def,actions_def,Int_def]) 1); |
336 |
by (Simp_tac 1); |
|
1050 | 337 |
by (rtac set_ext 1); |
338 |
by (Action.action.induct_tac "x" 1); |
|
1266 | 339 |
by (ALLGOALS(Simp_tac)); |
340 |
by (ALLGOALS(simp_tac (!simpset addsimps [insert_def,Un_def]))); |
|
341 |
by (ALLGOALS(simp_tac (!simpset addsimps [de_morgan,singleton_set]))); |
|
1050 | 342 |
val compat_env=result(); |
343 |
||
344 |
goal Correctness.thy "compat_ioas env_ioa\ |
|
345 |
\ (sender_ioa || receiver_ioa || srch_fin_ioa || rsch_fin_ioa)"; |
|
1266 | 346 |
by (simp_tac (ss addsimps [empty_def,compat_ioas_def,compat_asigs_def,asig_of_par,asig_comp_def,actions_def,Int_def]) 1); |
347 |
by (Simp_tac 1); |
|
1050 | 348 |
by (rtac set_ext 1); |
349 |
by (Action.action.induct_tac "x" 1); |
|
1266 | 350 |
by (ALLGOALS Simp_tac); |
351 |
by (ALLGOALS(simp_tac (!simpset addsimps [insert_def,Un_def]))); |
|
352 |
by (ALLGOALS(simp_tac (!simpset addsimps [de_morgan,singleton_set]))); |
|
1050 | 353 |
val compat_env_fin=result(); |
354 |
||
355 |
||
356 |
(* lemmata about externals of channels *) |
|
357 |
goal Correctness.thy |
|
358 |
"externals(asig_of(srch_fin_ioa)) = externals(asig_of(srch_ioa)) & \ |
|
359 |
\ externals(asig_of(rsch_fin_ioa)) = externals(asig_of(rsch_ioa))"; |
|
1266 | 360 |
by (simp_tac (!simpset addsimps [externals_def]) 1); |
1050 | 361 |
val ext_single_ch = result(); |
362 |
||
1138 | 363 |
|
364 |
||
365 |
||
366 |
||
367 |
val ext_ss = [externals_of_par,ext_single_ch]; |
|
368 |
val compat_ss = [compat_single_ch,compat_single_fin_ch,compat_rec, |
|
369 |
compat_rec_fin,compat_sen,compat_sen_fin,compat_env,compat_env_fin]; |
|
370 |
val abstractions = [env_unchanged,sender_unchanged, |
|
371 |
receiver_unchanged,sender_abstraction,receiver_abstraction]; |
|
372 |
||
373 |
||
374 |
(************************************************************************ |
|
375 |
S o u n d n e s s o f A b s t r a c t i o n |
|
376 |
*************************************************************************) |
|
377 |
||
1266 | 378 |
val ss = !simpset delsimps ([asig_of_def, srch_ioa_def, rsch_ioa_def, |
379 |
srch_fin_ioa_def, rsch_fin_ioa_def] @ |
|
380 |
impl_ioas @ env_ioas); |
|
1138 | 381 |
|
382 |
goal Correctness.thy |
|
383 |
"is_weak_pmap abs system_ioa system_fin_ioa"; |
|
384 |
||
1266 | 385 |
by (simp_tac (impl_ss delsimps ([srch_ioa_def, rsch_ioa_def, srch_fin_ioa_def, |
386 |
rsch_fin_ioa_def] @ env_ioas @ impl_ioas) |
|
1570 | 387 |
addsimps [system_def, system_fin_def, abs_def, |
388 |
impl_ioa_def, impl_fin_ioa_def, sys_IOA, |
|
389 |
sys_fin_IOA]) 1); |
|
1138 | 390 |
|
391 |
by (REPEAT (EVERY[rtac fxg_is_weak_pmap_of_product_IOA 1, |
|
1266 | 392 |
simp_tac (ss addsimps abstractions) 1, |
1138 | 393 |
rtac conjI 1])); |
394 |
||
1266 | 395 |
by (ALLGOALS (simp_tac (ss addsimps ext_ss @ compat_ss))); |
1138 | 396 |
|
1050 | 397 |
qed "system_refinement"; |