|
1 (* Title: HOL/IOA/NTP/Impl.ML |
|
2 ID: $Id$ |
|
3 Author: Tobias Nipkow & Konrad Slind |
|
4 Copyright 1994 TU Muenchen |
|
5 |
|
6 The implementation --- Invariants |
|
7 *) |
|
8 |
|
9 |
|
10 Addsimps [Let_def]; |
|
11 |
|
12 |
|
13 open Abschannel Impl; |
|
14 |
|
15 val impl_ioas = |
|
16 [Impl.impl_def, |
|
17 Sender.sender_ioa_def, |
|
18 Receiver.receiver_ioa_def, |
|
19 srch_ioa_thm RS eq_reflection, |
|
20 rsch_ioa_thm RS eq_reflection]; |
|
21 |
|
22 val transitions = [Sender.sender_trans_def, Receiver.receiver_trans_def, |
|
23 srch_trans_def, rsch_trans_def]; |
|
24 |
|
25 |
|
26 Addsimps [ioa_triple_proj, starts_of_par, trans_of_par4, |
|
27 in_sender_asig, in_receiver_asig, in_srch_asig, |
|
28 in_rsch_asig]; |
|
29 Addcongs [let_weak_cong]; |
|
30 |
|
31 goal Impl.thy |
|
32 "fst(x) = sen(x) & \ |
|
33 \ fst(snd(x)) = rec(x) & \ |
|
34 \ fst(snd(snd(x))) = srch(x) & \ |
|
35 \ snd(snd(snd(x))) = rsch(x)"; |
|
36 by(simp_tac (!simpset addsimps |
|
37 [sen_def,rec_def,srch_def,rsch_def]) 1); |
|
38 Addsimps [result()]; |
|
39 |
|
40 goal Impl.thy "a:actions(sender_asig) \ |
|
41 \ | a:actions(receiver_asig) \ |
|
42 \ | a:actions(srch_asig) \ |
|
43 \ | a:actions(rsch_asig)"; |
|
44 by(Action.action.induct_tac "a" 1); |
|
45 by(ALLGOALS (Simp_tac)); |
|
46 Addsimps [result()]; |
|
47 |
|
48 Delsimps [split_paired_All]; |
|
49 |
|
50 |
|
51 (* Three Simp_sets in different sizes |
|
52 ---------------------------------------------- |
|
53 |
|
54 1) !simpset does not unfold the transition relations |
|
55 2) ss unfolds transition relations |
|
56 3) renname_ss unfolds transitions and the abstract channel *) |
|
57 |
|
58 |
|
59 val ss = (!simpset addcongs [if_weak_cong] |
|
60 addsimps transitions); |
|
61 val rename_ss = (ss addsimps unfold_renaming); |
|
62 |
|
63 |
|
64 |
|
65 val tac = asm_simp_tac ((ss addcongs [conj_cong]) |
|
66 setloop (split_tac [expand_if])); |
|
67 val tac_ren = asm_simp_tac ((rename_ss addcongs [conj_cong]) |
|
68 setloop (split_tac [expand_if])); |
|
69 |
|
70 |
|
71 |
|
72 (* INVARIANT 1 *) |
|
73 |
|
74 goalw Impl.thy impl_ioas "invariant impl_ioa inv1"; |
|
75 by (rtac invariantI 1); |
|
76 by(asm_full_simp_tac (!simpset |
|
77 addsimps [inv1_def, hdr_sum_def, Sender.srcvd_def, |
|
78 Sender.ssent_def, Receiver.rsent_def,Receiver.rrcvd_def]) 1); |
|
79 |
|
80 by(simp_tac (!simpset delsimps [trans_of_par4] |
|
81 addsimps [fork_lemma,inv1_def]) 1); |
|
82 |
|
83 (* Split proof in two *) |
|
84 by (rtac conjI 1); |
|
85 |
|
86 (* First half *) |
|
87 by(asm_full_simp_tac (!simpset addsimps [Impl.inv1_def]) 1); |
|
88 by (rtac Action.action.induct 1); |
|
89 |
|
90 by (EVERY1[tac, tac, tac, tac]); |
|
91 by (tac 1); |
|
92 by (tac_ren 1); |
|
93 |
|
94 (* 5 + 1 *) |
|
95 |
|
96 by (tac 1); |
|
97 by (tac_ren 1); |
|
98 |
|
99 (* 4 + 1 *) |
|
100 by(EVERY1[tac, tac, tac, tac]); |
|
101 |
|
102 |
|
103 (* Now the other half *) |
|
104 by(asm_full_simp_tac (!simpset addsimps [Impl.inv1_def]) 1); |
|
105 by (rtac Action.action.induct 1); |
|
106 by(EVERY1[tac, tac]); |
|
107 |
|
108 (* detour 1 *) |
|
109 by (tac 1); |
|
110 by (tac_ren 1); |
|
111 by (rtac impI 1); |
|
112 by (REPEAT (etac conjE 1)); |
|
113 by (asm_simp_tac (!simpset addsimps [hdr_sum_def, Multiset.count_def, |
|
114 Multiset.countm_nonempty_def] |
|
115 setloop (split_tac [expand_if])) 1); |
|
116 (* detour 2 *) |
|
117 by (tac 1); |
|
118 by (tac_ren 1); |
|
119 by (rtac impI 1); |
|
120 by (REPEAT (etac conjE 1)); |
|
121 by (asm_full_simp_tac (!simpset addsimps [Impl.hdr_sum_def, |
|
122 Multiset.count_def, |
|
123 Multiset.countm_nonempty_def, |
|
124 Multiset.delm_nonempty_def, |
|
125 left_plus_cancel, |
|
126 left_plus_cancel_inside_succ, |
|
127 unzero_less] |
|
128 setloop (split_tac [expand_if])) 1); |
|
129 by (rtac allI 1); |
|
130 by (rtac conjI 1); |
|
131 by (rtac impI 1); |
|
132 by (hyp_subst_tac 1); |
|
133 by (rtac (pred_suc RS mp RS sym RS iffD2) 1); |
|
134 by (dtac less_le_trans 1); |
|
135 by (cut_facts_tac [rewrite_rule[Packet.hdr_def] |
|
136 eq_packet_imp_eq_hdr RS countm_props] 1);; |
|
137 by (assume_tac 1); |
|
138 by (assume_tac 1); |
|
139 |
|
140 by (rtac (countm_done_delm RS mp RS sym) 1); |
|
141 by (rtac refl 1); |
|
142 by (asm_simp_tac (!simpset addsimps [Multiset.count_def]) 1); |
|
143 |
|
144 by (rtac impI 1); |
|
145 by (asm_full_simp_tac (!simpset addsimps [neg_flip]) 1); |
|
146 by (hyp_subst_tac 1); |
|
147 by (rtac countm_spurious_delm 1); |
|
148 by (Simp_tac 1); |
|
149 |
|
150 by (EVERY1[tac, tac, tac, tac, tac, tac]); |
|
151 |
|
152 qed "inv1"; |
|
153 |
|
154 |
|
155 |
|
156 (* INVARIANT 2 *) |
|
157 |
|
158 goal Impl.thy "invariant impl_ioa inv2"; |
|
159 |
|
160 by (rtac invariantI1 1); |
|
161 (* Base case *) |
|
162 by (asm_full_simp_tac (!simpset addsimps (inv2_def :: |
|
163 (receiver_projections |
|
164 @ sender_projections @ impl_ioas))) |
|
165 1); |
|
166 |
|
167 by (asm_simp_tac (!simpset addsimps impl_ioas) 1); |
|
168 by (Action.action.induct_tac "a" 1); |
|
169 |
|
170 (* 10 cases. First 4 are simple, since state doesn't change *) |
|
171 |
|
172 val tac2 = asm_full_simp_tac (ss addsimps [inv2_def]); |
|
173 |
|
174 (* 10 - 7 *) |
|
175 by (EVERY1[tac2,tac2,tac2,tac2]); |
|
176 (* 6 *) |
|
177 by(forward_tac [rewrite_rule [Impl.inv1_def] |
|
178 (inv1 RS invariantE) RS conjunct1] 1); |
|
179 (* 6 - 5 *) |
|
180 by (EVERY1[tac2,tac2]); |
|
181 |
|
182 (* 4 *) |
|
183 by (forward_tac [rewrite_rule [Impl.inv1_def] |
|
184 (inv1 RS invariantE) RS conjunct1] 1); |
|
185 by (tac2 1); |
|
186 by (fast_tac (!claset addDs [add_leD1,leD]) 1); |
|
187 |
|
188 (* 3 *) |
|
189 by (forward_tac [rewrite_rule [Impl.inv1_def] (inv1 RS invariantE)] 1); |
|
190 |
|
191 by (tac2 1); |
|
192 by (fold_tac [rewrite_rule [Packet.hdr_def]Impl.hdr_sum_def]); |
|
193 by (fast_tac (!claset addDs [add_leD1,leD]) 1); |
|
194 |
|
195 (* 2 *) |
|
196 by (tac2 1); |
|
197 by(forward_tac [rewrite_rule [Impl.inv1_def] |
|
198 (inv1 RS invariantE) RS conjunct1] 1); |
|
199 by (rtac impI 1); |
|
200 by (rtac impI 1); |
|
201 by (REPEAT (etac conjE 1)); |
|
202 by (dres_inst_tac [("k","count (rsch s) (~sbit(sen s))")] |
|
203 (standard(leq_add_leq RS mp)) 1); |
|
204 by (Asm_full_simp_tac 1); |
|
205 |
|
206 (* 1 *) |
|
207 by (tac2 1); |
|
208 by(forward_tac [rewrite_rule [Impl.inv1_def] |
|
209 (inv1 RS invariantE) RS conjunct2] 1); |
|
210 by (rtac impI 1); |
|
211 by (rtac impI 1); |
|
212 by (REPEAT (etac conjE 1)); |
|
213 by (fold_tac [rewrite_rule[Packet.hdr_def]Impl.hdr_sum_def]); |
|
214 by (dres_inst_tac [("k","hdr_sum (srch s) (sbit(sen s))")] |
|
215 (standard(leq_add_leq RS mp)) 1); |
|
216 by (Asm_full_simp_tac 1); |
|
217 qed "inv2"; |
|
218 |
|
219 |
|
220 (* INVARIANT 3 *) |
|
221 |
|
222 goal Impl.thy "invariant impl_ioa inv3"; |
|
223 |
|
224 by (rtac invariantI 1); |
|
225 (* Base case *) |
|
226 by (asm_full_simp_tac (!simpset addsimps |
|
227 (Impl.inv3_def :: (receiver_projections |
|
228 @ sender_projections @ impl_ioas))) 1); |
|
229 |
|
230 by (asm_simp_tac (!simpset addsimps impl_ioas) 1); |
|
231 by (Action.action.induct_tac "a" 1); |
|
232 |
|
233 val tac3 = asm_full_simp_tac (ss addsimps [inv3_def] |
|
234 setloop (split_tac [expand_if])); |
|
235 |
|
236 (* 10 - 8 *) |
|
237 |
|
238 by (EVERY1[tac3,tac3,tac3]); |
|
239 |
|
240 by (tac_ren 1); |
|
241 by (strip_tac 1 THEN REPEAT (etac conjE 1)); |
|
242 by (hyp_subst_tac 1); |
|
243 by (etac exE 1); |
|
244 by (Asm_full_simp_tac 1); |
|
245 |
|
246 (* 7 *) |
|
247 by (tac3 1); |
|
248 by (tac_ren 1); |
|
249 by (Fast_tac 1); |
|
250 |
|
251 (* 6 - 3 *) |
|
252 |
|
253 by (EVERY1[tac3,tac3,tac3,tac3]); |
|
254 |
|
255 (* 2 *) |
|
256 by (asm_full_simp_tac ss 1); |
|
257 by (simp_tac (!simpset addsimps [inv3_def]) 1); |
|
258 by (strip_tac 1 THEN REPEAT (etac conjE 1)); |
|
259 by (rtac (imp_or_lem RS iffD2) 1); |
|
260 by (rtac impI 1); |
|
261 by (forward_tac [rewrite_rule [Impl.inv2_def] (inv2 RS invariantE)] 1); |
|
262 by (Asm_full_simp_tac 1); |
|
263 by (REPEAT (etac conjE 1)); |
|
264 by (res_inst_tac [("j","count (ssent(sen s)) (~sbit(sen s))"), |
|
265 ("k","count (rsent(rec s)) (sbit(sen s))")] le_trans 1); |
|
266 by (forward_tac [rewrite_rule [inv1_def] |
|
267 (inv1 RS invariantE) RS conjunct2] 1); |
|
268 by (asm_full_simp_tac (!simpset addsimps |
|
269 [hdr_sum_def, Multiset.count_def]) 1); |
|
270 by (rtac (less_eq_add_cong RS mp RS mp) 1); |
|
271 by (rtac countm_props 1); |
|
272 by (Simp_tac 1); |
|
273 by (rtac countm_props 1); |
|
274 by (Simp_tac 1); |
|
275 by (assume_tac 1); |
|
276 |
|
277 (* 1 *) |
|
278 by (tac3 1); |
|
279 by (strip_tac 1 THEN REPEAT (etac conjE 1)); |
|
280 by (rtac (imp_or_lem RS iffD2) 1); |
|
281 by (rtac impI 1); |
|
282 by (forward_tac [rewrite_rule [Impl.inv2_def] (inv2 RS invariantE)] 1); |
|
283 by (Asm_full_simp_tac 1); |
|
284 by (REPEAT (etac conjE 1)); |
|
285 by (dtac mp 1); |
|
286 by (assume_tac 1); |
|
287 by (etac allE 1); |
|
288 by (dtac (imp_or_lem RS iffD1) 1); |
|
289 by (dtac mp 1); |
|
290 by (assume_tac 1); |
|
291 by (assume_tac 1); |
|
292 qed "inv3"; |
|
293 |
|
294 |
|
295 (* INVARIANT 4 *) |
|
296 |
|
297 goal Impl.thy "invariant impl_ioa inv4"; |
|
298 |
|
299 by (rtac invariantI 1); |
|
300 (* Base case *) |
|
301 by (asm_full_simp_tac (!simpset addsimps |
|
302 (Impl.inv4_def :: (receiver_projections |
|
303 @ sender_projections @ impl_ioas))) 1); |
|
304 |
|
305 by (asm_simp_tac (!simpset addsimps impl_ioas) 1); |
|
306 by (Action.action.induct_tac "a" 1); |
|
307 |
|
308 val tac4 = asm_full_simp_tac (ss addsimps [inv4_def] |
|
309 setloop (split_tac [expand_if])); |
|
310 |
|
311 (* 10 - 2 *) |
|
312 |
|
313 by (EVERY1[tac4,tac4,tac4,tac4,tac4,tac4,tac4,tac4,tac4]); |
|
314 |
|
315 (* 2 b *) |
|
316 |
|
317 by (strip_tac 1 THEN REPEAT (etac conjE 1)); |
|
318 by(forward_tac [rewrite_rule [Impl.inv2_def] |
|
319 (inv2 RS invariantE)] 1); |
|
320 by (tac4 1); |
|
321 by (Asm_full_simp_tac 1); |
|
322 |
|
323 (* 1 *) |
|
324 by (tac4 1); |
|
325 by (strip_tac 1 THEN REPEAT (etac conjE 1)); |
|
326 by (rtac ccontr 1); |
|
327 by(forward_tac [rewrite_rule [Impl.inv2_def] |
|
328 (inv2 RS invariantE)] 1); |
|
329 by(forward_tac [rewrite_rule [Impl.inv3_def] |
|
330 (inv3 RS invariantE)] 1); |
|
331 by (Asm_full_simp_tac 1); |
|
332 by (eres_inst_tac [("x","m")] allE 1); |
|
333 by (dtac less_le_trans 1); |
|
334 by (dtac (left_add_leq RS mp) 1); |
|
335 by (Asm_full_simp_tac 1); |
|
336 by (Asm_full_simp_tac 1); |
|
337 qed "inv4"; |
|
338 |
|
339 |
|
340 (* rebind them *) |
|
341 |
|
342 val inv1 = rewrite_rule [Impl.inv1_def] (inv1 RS invariantE); |
|
343 val inv2 = rewrite_rule [Impl.inv2_def] (inv2 RS invariantE); |
|
344 val inv3 = rewrite_rule [Impl.inv3_def] (inv3 RS invariantE); |
|
345 val inv4 = rewrite_rule [Impl.inv4_def] (inv4 RS invariantE); |