| author | wenzelm | 
| Mon, 26 Nov 2012 20:09:51 +0100 | |
| changeset 50234 | c97c5c34fb1d | 
| parent 45620 | f2a587696afb | 
| child 51717 | 9e7d1c139569 | 
| permissions | -rw-r--r-- | 
| 42151 | 1 | (* Title: HOL/HOLCF/IOA/NTP/Impl.thy | 
| 3073 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 2 | Author: Tobias Nipkow & Konrad Slind | 
| 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 3 | *) | 
| 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 4 | |
| 17244 | 5 | header {* The implementation *}
 | 
| 3073 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 6 | |
| 17244 | 7 | theory Impl | 
| 8 | imports Sender Receiver Abschannel | |
| 9 | begin | |
| 3073 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 10 | |
| 41476 | 11 | type_synonym 'm impl_state | 
| 19739 | 12 | = "'m sender_state * 'm receiver_state * 'm packet multiset * bool multiset" | 
| 13 | (* sender_state * receiver_state * srch_state * rsch_state *) | |
| 3073 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 14 | |
| 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 15 | |
| 27361 | 16 | definition | 
| 17 |   impl_ioa :: "('m action, 'm impl_state)ioa" where
 | |
| 18 | impl_def: "impl_ioa == (sender_ioa || receiver_ioa || srch_ioa || rsch_ioa)" | |
| 3073 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 19 | |
| 27361 | 20 | definition sen :: "'m impl_state => 'm sender_state" where "sen = fst" | 
| 21 | definition rec :: "'m impl_state => 'm receiver_state" where "rec = fst o snd" | |
| 22 | definition srch :: "'m impl_state => 'm packet multiset" where "srch = fst o snd o snd" | |
| 23 | definition rsch :: "'m impl_state => bool multiset" where "rsch = snd o snd o snd" | |
| 3073 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 24 | |
| 27361 | 25 | definition | 
| 26 | hdr_sum :: "'m packet multiset => bool => nat" where | |
| 27 | "hdr_sum M b == countm M (%pkt. hdr(pkt) = b)" | |
| 3073 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 28 | |
| 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 29 | (* Lemma 5.1 *) | 
| 27361 | 30 | definition | 
| 17244 | 31 | "inv1(s) == | 
| 32 | (!b. count (rsent(rec s)) b = count (srcvd(sen s)) b + count (rsch s) b) | |
| 33 | & (!b. count (ssent(sen s)) b | |
| 3073 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 34 | = hdr_sum (rrcvd(rec s)) b + hdr_sum (srch s) b)" | 
| 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 35 | |
| 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 36 | (* Lemma 5.2 *) | 
| 27361 | 37 | definition | 
| 38 | "inv2(s) == | |
| 17244 | 39 | (rbit(rec(s)) = sbit(sen(s)) & | 
| 40 | ssending(sen(s)) & | |
| 3073 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 41 | count (rsent(rec s)) (~sbit(sen s)) <= count (ssent(sen s)) (~sbit(sen s)) & | 
| 17244 | 42 | count (ssent(sen s)) (~sbit(sen s)) <= count (rsent(rec s)) (sbit(sen s))) | 
| 43 | | | |
| 44 | (rbit(rec(s)) = (~sbit(sen(s))) & | |
| 45 | rsending(rec(s)) & | |
| 3073 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 46 | count (ssent(sen s)) (~sbit(sen s)) <= count (rsent(rec s)) (sbit(sen s)) & | 
| 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 47 | count (rsent(rec s)) (sbit(sen s)) <= count (ssent(sen s)) (sbit(sen s)))" | 
| 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 48 | |
| 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 49 | (* Lemma 5.3 *) | 
| 27361 | 50 | definition | 
| 51 | "inv3(s) == | |
| 17244 | 52 | rbit(rec(s)) = sbit(sen(s)) | 
| 53 | --> (!m. sq(sen(s))=[] | m ~= hd(sq(sen(s))) | |
| 54 | --> count (rrcvd(rec s)) (sbit(sen(s)),m) | |
| 55 | + count (srch s) (sbit(sen(s)),m) | |
| 3073 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 56 | <= count (rsent(rec s)) (~sbit(sen s)))" | 
| 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 57 | |
| 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 58 | (* Lemma 5.4 *) | 
| 27361 | 59 | definition "inv4(s) == rbit(rec(s)) = (~sbit(sen(s))) --> sq(sen(s)) ~= []" | 
| 17244 | 60 | |
| 19739 | 61 | |
| 62 | subsection {* Invariants *}
 | |
| 63 | ||
| 35215 
a03462cbf86f
get rid of warnings about duplicate simp rules in all HOLCF theories
 huffman parents: 
35174diff
changeset | 64 | declare le_SucI [simp] | 
| 19739 | 65 | |
| 66 | lemmas impl_ioas = | |
| 67 | impl_def sender_ioa_def receiver_ioa_def srch_ioa_thm [THEN eq_reflection] | |
| 68 | rsch_ioa_thm [THEN eq_reflection] | |
| 69 | ||
| 70 | lemmas "transitions" = | |
| 71 | sender_trans_def receiver_trans_def srch_trans_def rsch_trans_def | |
| 72 | ||
| 73 | ||
| 74 | lemmas [simp] = | |
| 75 | ioa_triple_proj starts_of_par trans_of_par4 in_sender_asig | |
| 76 | in_receiver_asig in_srch_asig in_rsch_asig | |
| 77 | ||
| 78 | declare let_weak_cong [cong] | |
| 79 | ||
| 80 | lemma [simp]: | |
| 81 | "fst(x) = sen(x)" | |
| 82 | "fst(snd(x)) = rec(x)" | |
| 83 | "fst(snd(snd(x))) = srch(x)" | |
| 84 | "snd(snd(snd(x))) = rsch(x)" | |
| 85 | by (simp_all add: sen_def rec_def srch_def rsch_def) | |
| 86 | ||
| 87 | lemma [simp]: | |
| 88 | "a:actions(sender_asig) | |
| 89 | | a:actions(receiver_asig) | |
| 90 | | a:actions(srch_asig) | |
| 91 | | a:actions(rsch_asig)" | |
| 92 | by (induct a) simp_all | |
| 93 | ||
| 94 | declare split_paired_All [simp del] | |
| 95 | ||
| 96 | ||
| 97 | (* Three Simp_sets in different sizes | |
| 98 | ---------------------------------------------- | |
| 99 | ||
| 100 | 1) simpset() does not unfold the transition relations | |
| 101 | 2) ss unfolds transition relations | |
| 102 | 3) renname_ss unfolds transitions and the abstract channel *) | |
| 103 | ||
| 104 | ML {*
 | |
| 27355 | 105 | val ss = @{simpset} addsimps @{thms "transitions"};
 | 
| 26305 | 106 | val rename_ss = ss addsimps @{thms unfold_renaming};
 | 
| 19739 | 107 | |
| 45620 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
42151diff
changeset | 108 | val tac = | 
| 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
42151diff
changeset | 109 |   asm_simp_tac (ss |> Simplifier.add_cong @{thm conj_cong} |> Splitter.add_split @{thm split_if})
 | 
| 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
42151diff
changeset | 110 | val tac_ren = | 
| 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
42151diff
changeset | 111 |   asm_simp_tac (rename_ss |> Simplifier.add_cong @{thm conj_cong} |> Splitter.add_split @{thm split_if})
 | 
| 19739 | 112 | *} | 
| 113 | ||
| 114 | ||
| 115 | subsubsection {* Invariant 1 *}
 | |
| 116 | ||
| 26305 | 117 | lemma raw_inv1: "invariant impl_ioa inv1" | 
| 19739 | 118 | |
| 119 | apply (unfold impl_ioas) | |
| 120 | apply (rule invariantI) | |
| 121 | apply (simp add: inv1_def hdr_sum_def srcvd_def ssent_def rsent_def rrcvd_def) | |
| 122 | ||
| 123 | apply (simp (no_asm) del: trans_of_par4 add: imp_conjR inv1_def) | |
| 124 | ||
| 125 | txt {* Split proof in two *}
 | |
| 126 | apply (rule conjI) | |
| 127 | ||
| 128 | (* First half *) | |
| 129 | apply (simp add: Impl.inv1_def split del: split_if) | |
| 130 | apply (induct_tac a) | |
| 131 | ||
| 132 | apply (tactic "EVERY1[tac, tac, tac, tac]") | |
| 133 | apply (tactic "tac 1") | |
| 134 | apply (tactic "tac_ren 1") | |
| 135 | ||
| 136 | txt {* 5 + 1 *}
 | |
| 137 | ||
| 138 | apply (tactic "tac 1") | |
| 139 | apply (tactic "tac_ren 1") | |
| 140 | ||
| 141 | txt {* 4 + 1 *}
 | |
| 142 | apply (tactic {* EVERY1[tac, tac, tac, tac] *})
 | |
| 143 | ||
| 144 | ||
| 145 | txt {* Now the other half *}
 | |
| 146 | apply (simp add: Impl.inv1_def split del: split_if) | |
| 147 | apply (induct_tac a) | |
| 148 | apply (tactic "EVERY1 [tac, tac]") | |
| 149 | ||
| 150 | txt {* detour 1 *}
 | |
| 151 | apply (tactic "tac 1") | |
| 152 | apply (tactic "tac_ren 1") | |
| 153 | apply (rule impI) | |
| 154 | apply (erule conjE)+ | |
| 155 | apply (simp (no_asm_simp) add: hdr_sum_def Multiset.count_def Multiset.countm_nonempty_def | |
| 156 | split add: split_if) | |
| 157 | txt {* detour 2 *}
 | |
| 158 | apply (tactic "tac 1") | |
| 159 | apply (tactic "tac_ren 1") | |
| 160 | apply (rule impI) | |
| 161 | apply (erule conjE)+ | |
| 162 | apply (simp add: Impl.hdr_sum_def Multiset.count_def Multiset.countm_nonempty_def | |
| 25161 | 163 | Multiset.delm_nonempty_def split add: split_if) | 
| 19739 | 164 | apply (rule allI) | 
| 165 | apply (rule conjI) | |
| 166 | apply (rule impI) | |
| 167 | apply hypsubst | |
| 168 | apply (rule pred_suc [THEN iffD1]) | |
| 169 | apply (drule less_le_trans) | |
| 170 | apply (cut_tac eq_packet_imp_eq_hdr [unfolded Packet.hdr_def, THEN countm_props]) | |
| 171 | apply assumption | |
| 172 | apply assumption | |
| 173 | ||
| 174 | apply (rule countm_done_delm [THEN mp, symmetric]) | |
| 175 | apply (rule refl) | |
| 176 | apply (simp (no_asm_simp) add: Multiset.count_def) | |
| 177 | ||
| 178 | apply (rule impI) | |
| 179 | apply (simp add: neg_flip) | |
| 180 | apply hypsubst | |
| 181 | apply (rule countm_spurious_delm) | |
| 182 | apply (simp (no_asm)) | |
| 183 | ||
| 184 | apply (tactic "EVERY1 [tac, tac, tac, tac, tac, tac]") | |
| 185 | ||
| 186 | done | |
| 187 | ||
| 188 | ||
| 189 | ||
| 190 | subsubsection {* INVARIANT 2 *}
 | |
| 191 | ||
| 26305 | 192 | lemma raw_inv2: "invariant impl_ioa inv2" | 
| 19739 | 193 | |
| 194 | apply (rule invariantI1) | |
| 195 |   txt {* Base case *}
 | |
| 196 | apply (simp add: inv2_def receiver_projections sender_projections impl_ioas) | |
| 197 | ||
| 198 | apply (simp (no_asm_simp) add: impl_ioas split del: split_if) | |
| 199 | apply (induct_tac "a") | |
| 200 | ||
| 201 |   txt {* 10 cases. First 4 are simple, since state doesn't change *}
 | |
| 202 | ||
| 28265 | 203 |   ML_prf {* val tac2 = asm_full_simp_tac (ss addsimps [@{thm inv2_def}]) *}
 | 
| 19739 | 204 | |
| 205 |   txt {* 10 - 7 *}
 | |
| 206 | apply (tactic "EVERY1 [tac2,tac2,tac2,tac2]") | |
| 207 |   txt {* 6 *}
 | |
| 26305 | 208 |   apply (tactic {* forward_tac [rewrite_rule [@{thm Impl.inv1_def}]
 | 
| 209 |                                (@{thm raw_inv1} RS @{thm invariantE}) RS conjunct1] 1 *})
 | |
| 19739 | 210 | |
| 211 |   txt {* 6 - 5 *}
 | |
| 212 | apply (tactic "EVERY1 [tac2,tac2]") | |
| 213 | ||
| 214 |   txt {* 4 *}
 | |
| 26305 | 215 |   apply (tactic {* forward_tac [rewrite_rule [@{thm Impl.inv1_def}]
 | 
| 216 |                                 (@{thm raw_inv1} RS @{thm invariantE}) RS conjunct1] 1 *})
 | |
| 19739 | 217 | apply (tactic "tac2 1") | 
| 218 | ||
| 219 |   txt {* 3 *}
 | |
| 26305 | 220 |   apply (tactic {* forward_tac [rewrite_rule [@{thm Impl.inv1_def}]
 | 
| 221 |     (@{thm raw_inv1} RS @{thm invariantE})] 1 *})
 | |
| 19739 | 222 | |
| 223 | apply (tactic "tac2 1") | |
| 28839 
32d498cf7595
eliminated rewrite_tac/fold_tac, which are not well-formed tactics due to change of main conclusion;
 wenzelm parents: 
28265diff
changeset | 224 |   apply (tactic {* fold_goals_tac [rewrite_rule [@{thm Packet.hdr_def}]
 | 
| 27361 | 225 |     (@{thm Impl.hdr_sum_def})] *})
 | 
| 19739 | 226 | apply arith | 
| 227 | ||
| 228 |   txt {* 2 *}
 | |
| 229 | apply (tactic "tac2 1") | |
| 26305 | 230 |   apply (tactic {* forward_tac [rewrite_rule [@{thm Impl.inv1_def}]
 | 
| 231 |                                (@{thm raw_inv1} RS @{thm invariantE}) RS conjunct1] 1 *})
 | |
| 19739 | 232 | apply (intro strip) | 
| 233 | apply (erule conjE)+ | |
| 234 | apply simp | |
| 235 | ||
| 236 |   txt {* 1 *}
 | |
| 237 | apply (tactic "tac2 1") | |
| 26305 | 238 |   apply (tactic {* forward_tac [rewrite_rule [@{thm Impl.inv1_def}]
 | 
| 239 |                                (@{thm raw_inv1} RS @{thm invariantE}) RS conjunct2] 1 *})
 | |
| 19739 | 240 | apply (intro strip) | 
| 241 | apply (erule conjE)+ | |
| 28839 
32d498cf7595
eliminated rewrite_tac/fold_tac, which are not well-formed tactics due to change of main conclusion;
 wenzelm parents: 
28265diff
changeset | 242 |   apply (tactic {* fold_goals_tac [rewrite_rule [@{thm Packet.hdr_def}] (@{thm Impl.hdr_sum_def})] *})
 | 
| 19739 | 243 | apply simp | 
| 244 | ||
| 245 | done | |
| 246 | ||
| 247 | ||
| 248 | subsubsection {* INVARIANT 3 *}
 | |
| 249 | ||
| 26305 | 250 | lemma raw_inv3: "invariant impl_ioa inv3" | 
| 19739 | 251 | |
| 252 | apply (rule invariantI) | |
| 253 |   txt {* Base case *}
 | |
| 254 | apply (simp add: Impl.inv3_def receiver_projections sender_projections impl_ioas) | |
| 255 | ||
| 256 | apply (simp (no_asm_simp) add: impl_ioas split del: split_if) | |
| 257 | apply (induct_tac "a") | |
| 258 | ||
| 28265 | 259 |   ML_prf {* val tac3 = asm_full_simp_tac (ss addsimps [@{thm inv3_def}]) *}
 | 
| 19739 | 260 | |
| 261 |   txt {* 10 - 8 *}
 | |
| 262 | ||
| 263 | apply (tactic "EVERY1[tac3,tac3,tac3]") | |
| 264 | ||
| 265 | apply (tactic "tac_ren 1") | |
| 266 | apply (intro strip, (erule conjE)+) | |
| 267 | apply hypsubst | |
| 268 | apply (erule exE) | |
| 269 | apply simp | |
| 270 | ||
| 271 |   txt {* 7 *}
 | |
| 272 | apply (tactic "tac3 1") | |
| 273 | apply (tactic "tac_ren 1") | |
| 274 | apply force | |
| 275 | ||
| 276 |   txt {* 6 - 3 *}
 | |
| 277 | ||
| 278 | apply (tactic "EVERY1[tac3,tac3,tac3,tac3]") | |
| 279 | ||
| 280 |   txt {* 2 *}
 | |
| 281 | apply (tactic "asm_full_simp_tac ss 1") | |
| 282 | apply (simp (no_asm) add: inv3_def) | |
| 283 | apply (intro strip, (erule conjE)+) | |
| 284 | apply (rule imp_disjL [THEN iffD1]) | |
| 285 | apply (rule impI) | |
| 26305 | 286 |   apply (tactic {* forward_tac [rewrite_rule [@{thm Impl.inv2_def}]
 | 
| 287 |     (@{thm raw_inv2} RS @{thm invariantE})] 1 *})
 | |
| 19739 | 288 | apply simp | 
| 289 | apply (erule conjE)+ | |
| 290 | apply (rule_tac j = "count (ssent (sen s)) (~sbit (sen s))" and | |
| 291 | k = "count (rsent (rec s)) (sbit (sen s))" in le_trans) | |
| 26305 | 292 |   apply (tactic {* forward_tac [rewrite_rule [@{thm inv1_def}]
 | 
| 293 |                                 (@{thm raw_inv1} RS @{thm invariantE}) RS conjunct2] 1 *})
 | |
| 19739 | 294 | apply (simp add: hdr_sum_def Multiset.count_def) | 
| 295 | apply (rule add_le_mono) | |
| 296 | apply (rule countm_props) | |
| 297 | apply (simp (no_asm)) | |
| 298 | apply (rule countm_props) | |
| 299 | apply (simp (no_asm)) | |
| 300 | apply assumption | |
| 301 | ||
| 302 |   txt {* 1 *}
 | |
| 303 | apply (tactic "tac3 1") | |
| 304 | apply (intro strip, (erule conjE)+) | |
| 305 | apply (rule imp_disjL [THEN iffD1]) | |
| 306 | apply (rule impI) | |
| 26305 | 307 |   apply (tactic {* forward_tac [rewrite_rule [@{thm Impl.inv2_def}]
 | 
| 308 |     (@{thm raw_inv2} RS @{thm invariantE})] 1 *})
 | |
| 19739 | 309 | apply simp | 
| 310 | done | |
| 311 | ||
| 312 | ||
| 313 | subsubsection {* INVARIANT 4 *}
 | |
| 314 | ||
| 26305 | 315 | lemma raw_inv4: "invariant impl_ioa inv4" | 
| 19739 | 316 | |
| 317 | apply (rule invariantI) | |
| 318 |   txt {* Base case *}
 | |
| 319 | apply (simp add: Impl.inv4_def receiver_projections sender_projections impl_ioas) | |
| 320 | ||
| 321 | apply (simp (no_asm_simp) add: impl_ioas split del: split_if) | |
| 322 | apply (induct_tac "a") | |
| 323 | ||
| 28265 | 324 |   ML_prf {* val tac4 =  asm_full_simp_tac (ss addsimps [@{thm inv4_def}]) *}
 | 
| 19739 | 325 | |
| 326 |   txt {* 10 - 2 *}
 | |
| 327 | ||
| 328 | apply (tactic "EVERY1[tac4,tac4,tac4,tac4,tac4,tac4,tac4,tac4,tac4]") | |
| 329 | ||
| 330 |   txt {* 2 b *}
 | |
| 331 | ||
| 332 | apply (intro strip, (erule conjE)+) | |
| 26305 | 333 |   apply (tactic {* forward_tac [rewrite_rule [@{thm Impl.inv2_def}]
 | 
| 334 |                                (@{thm raw_inv2} RS @{thm invariantE})] 1 *})
 | |
| 19739 | 335 | apply simp | 
| 336 | ||
| 337 |   txt {* 1 *}
 | |
| 338 | apply (tactic "tac4 1") | |
| 339 | apply (intro strip, (erule conjE)+) | |
| 340 | apply (rule ccontr) | |
| 26305 | 341 |   apply (tactic {* forward_tac [rewrite_rule [@{thm Impl.inv2_def}]
 | 
| 342 |                                (@{thm raw_inv2} RS @{thm invariantE})] 1 *})
 | |
| 343 |   apply (tactic {* forward_tac [rewrite_rule [@{thm Impl.inv3_def}]
 | |
| 344 |                                (@{thm raw_inv3} RS @{thm invariantE})] 1 *})
 | |
| 19739 | 345 | apply simp | 
| 346 | apply (erule_tac x = "m" in allE) | |
| 347 | apply simp | |
| 348 | done | |
| 349 | ||
| 350 | ||
| 351 | text {* rebind them *}
 | |
| 352 | ||
| 26305 | 353 | lemmas inv1 = raw_inv1 [THEN invariantE, unfolded inv1_def] | 
| 354 | and inv2 = raw_inv2 [THEN invariantE, unfolded inv2_def] | |
| 355 | and inv3 = raw_inv3 [THEN invariantE, unfolded inv3_def] | |
| 356 | and inv4 = raw_inv4 [THEN invariantE, unfolded inv4_def] | |
| 3073 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 357 | |
| 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 358 | end |