| author | wenzelm | 
| Thu, 22 Feb 2024 13:12:10 +0100 | |
| changeset 79694 | ab7ec4a29b9c | 
| parent 69605 | a96320074298 | 
| permissions | -rw-r--r-- | 
| 42151 | 1 | (* Title: HOL/HOLCF/IOA/ABP/Correctness.thy | 
| 40945 | 2 | Author: Olaf Müller | 
| 3072 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 3 | *) | 
| 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 4 | |
| 62002 | 5 | section \<open>The main correctness proof: System_fin implements System\<close> | 
| 17244 | 6 | |
| 7 | theory Correctness | |
| 66453 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 wenzelm parents: 
62390diff
changeset | 8 | imports IOA.IOA Env Impl Impl_finite | 
| 17244 | 9 | begin | 
| 3072 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 10 | |
| 69605 | 11 | ML_file \<open>Check.ML\<close> | 
| 48891 | 12 | |
| 27361 | 13 | primrec reduce :: "'a list => 'a list" | 
| 14 | where | |
| 17244 | 15 | reduce_Nil: "reduce [] = []" | 
| 27361 | 16 | | reduce_Cons: "reduce(x#xs) = | 
| 17244 | 17 | (case xs of | 
| 18 | [] => [x] | |
| 19 | | y#ys => (if (x=y) | |
| 20 | then reduce xs | |
| 3072 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 21 | else (x#(reduce xs))))" | 
| 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 22 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
20918diff
changeset | 23 | definition | 
| 19689 | 24 | abs where | 
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
20918diff
changeset | 25 | "abs = | 
| 19689 | 26 | (%p.(fst(p),(fst(snd(p)),(fst(snd(snd(p))), | 
| 27 | (reduce(fst(snd(snd(snd(p))))),reduce(snd(snd(snd(snd(p))))))))))" | |
| 17244 | 28 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
20918diff
changeset | 29 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
20918diff
changeset | 30 |   system_ioa :: "('m action, bool * 'm impl_state)ioa" where
 | 
| 61999 | 31 | "system_ioa = (env_ioa \<parallel> impl_ioa)" | 
| 3072 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 32 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
20918diff
changeset | 33 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
20918diff
changeset | 34 |   system_fin_ioa :: "('m action, bool * 'm impl_state)ioa" where
 | 
| 61999 | 35 | "system_fin_ioa = (env_ioa \<parallel> impl_fin_ioa)" | 
| 17244 | 36 | |
| 3072 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 37 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
20918diff
changeset | 38 | axiomatization where | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
20918diff
changeset | 39 | sys_IOA: "IOA system_ioa" and | 
| 17244 | 40 | sys_fin_IOA: "IOA system_fin_ioa" | 
| 41 | ||
| 19738 | 42 | |
| 43 | ||
| 44 | declare split_paired_All [simp del] Collect_empty_eq [simp del] | |
| 45 | ||
| 46 | lemmas [simp] = | |
| 47 | srch_asig_def rsch_asig_def rsch_ioa_def srch_ioa_def ch_ioa_def | |
| 48 | ch_asig_def srch_actions_def rsch_actions_def rename_def rename_set_def asig_of_def | |
| 49 | actions_def exis_elim srch_trans_def rsch_trans_def ch_trans_def | |
| 50 | trans_of_def asig_projections set_lemmas | |
| 51 | ||
| 52 | lemmas abschannel_fin [simp] = | |
| 53 | srch_fin_asig_def rsch_fin_asig_def | |
| 54 | rsch_fin_ioa_def srch_fin_ioa_def | |
| 55 | ch_fin_ioa_def ch_fin_trans_def ch_fin_asig_def | |
| 56 | ||
| 57 | lemmas impl_ioas = sender_ioa_def receiver_ioa_def | |
| 58 | and impl_trans = sender_trans_def receiver_trans_def | |
| 59 | and impl_asigs = sender_asig_def receiver_asig_def | |
| 60 | ||
| 61 | declare let_weak_cong [cong] | |
| 35215 
a03462cbf86f
get rid of warnings about duplicate simp rules in all HOLCF theories
 huffman parents: 
35174diff
changeset | 62 | declare ioa_triple_proj [simp] starts_of_par [simp] | 
| 19738 | 63 | |
| 64 | lemmas env_ioas = env_ioa_def env_asig_def env_trans_def | |
| 35215 
a03462cbf86f
get rid of warnings about duplicate simp rules in all HOLCF theories
 huffman parents: 
35174diff
changeset | 65 | lemmas hom_ioas = | 
| 
a03462cbf86f
get rid of warnings about duplicate simp rules in all HOLCF theories
 huffman parents: 
35174diff
changeset | 66 | env_ioas [simp] impl_ioas [simp] impl_trans [simp] impl_asigs [simp] | 
| 
a03462cbf86f
get rid of warnings about duplicate simp rules in all HOLCF theories
 huffman parents: 
35174diff
changeset | 67 | asig_projections set_lemmas | 
| 19738 | 68 | |
| 69 | ||
| 62002 | 70 | subsection \<open>lemmas about reduce\<close> | 
| 19738 | 71 | |
| 72 | lemma l_iff_red_nil: "(reduce l = []) = (l = [])" | |
| 73 | by (induct l) (auto split: list.split) | |
| 74 | ||
| 75 | lemma hd_is_reduce_hd: "s ~= [] --> hd s = hd (reduce s)" | |
| 76 | by (induct s) (auto split: list.split) | |
| 77 | ||
| 62002 | 78 | text \<open>to be used in the following Lemma\<close> | 
| 19738 | 79 | lemma rev_red_not_nil [rule_format]: | 
| 80 | "l ~= [] --> reverse (reduce l) ~= []" | |
| 81 | by (induct l) (auto split: list.split) | |
| 82 | ||
| 62002 | 83 | text \<open>shows applicability of the induction hypothesis of the following Lemma 1\<close> | 
| 19738 | 84 | lemma last_ind_on_first: | 
| 85 | "l ~= [] ==> hd (reverse (reduce (a # l))) = hd (reverse (reduce l))" | |
| 86 | apply simp | |
| 69597 | 87 | apply (tactic \<open>auto_tac (put_simpset HOL_ss \<^context> | 
| 45620 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
42795diff
changeset | 88 |     addsimps (@{thms reverse.simps} @ [@{thm hd_append}, @{thm rev_red_not_nil}])
 | 
| 62002 | 89 |     |> Splitter.add_split @{thm list.split})\<close>)
 | 
| 19738 | 90 | done | 
| 91 | ||
| 62002 | 92 | text \<open>Main Lemma 1 for \<open>S_pkt\<close> in showing that reduce is refinement.\<close> | 
| 19738 | 93 | lemma reduce_hd: | 
| 94 | "if x=hd(reverse(reduce(l))) & reduce(l)~=[] then | |
| 95 | reduce(l@[x])=reduce(l) else | |
| 96 | reduce(l@[x])=reduce(l)@[x]" | |
| 62390 | 97 | apply (simplesubst if_split) | 
| 19738 | 98 | apply (rule conjI) | 
| 62002 | 99 | txt \<open>\<open>-->\<close>\<close> | 
| 19738 | 100 | apply (induct_tac "l") | 
| 101 | apply (simp (no_asm)) | |
| 102 | apply (case_tac "list=[]") | |
| 35215 
a03462cbf86f
get rid of warnings about duplicate simp rules in all HOLCF theories
 huffman parents: 
35174diff
changeset | 103 | apply simp | 
| 19738 | 104 | apply (rule impI) | 
| 105 | apply (simp (no_asm)) | |
| 106 | apply (cut_tac l = "list" in cons_not_nil) | |
| 107 | apply (simp del: reduce_Cons) | |
| 108 | apply (erule exE)+ | |
| 109 | apply hypsubst | |
| 110 | apply (simp del: reduce_Cons add: last_ind_on_first l_iff_red_nil) | |
| 62002 | 111 | txt \<open>\<open><--\<close>\<close> | 
| 19738 | 112 | apply (simp (no_asm) add: and_de_morgan_and_absorbe l_iff_red_nil) | 
| 113 | apply (induct_tac "l") | |
| 114 | apply (simp (no_asm)) | |
| 115 | apply (case_tac "list=[]") | |
| 116 | apply (cut_tac [2] l = "list" in cons_not_nil) | |
| 117 | apply simp | |
| 62390 | 118 | apply (auto simp del: reduce_Cons simp add: last_ind_on_first l_iff_red_nil split: if_split) | 
| 19738 | 119 | apply simp | 
| 120 | done | |
| 121 | ||
| 122 | ||
| 62002 | 123 | text \<open>Main Lemma 2 for R_pkt in showing that reduce is refinement.\<close> | 
| 19738 | 124 | lemma reduce_tl: "s~=[] ==> | 
| 125 | if hd(s)=hd(tl(s)) & tl(s)~=[] then | |
| 126 | reduce(tl(s))=reduce(s) else | |
| 127 | reduce(tl(s))=tl(reduce(s))" | |
| 128 | apply (cut_tac l = "s" in cons_not_nil) | |
| 129 | apply simp | |
| 130 | apply (erule exE)+ | |
| 131 | apply (auto split: list.split) | |
| 132 | done | |
| 133 | ||
| 134 | ||
| 62002 | 135 | subsection \<open>Channel Abstraction\<close> | 
| 19738 | 136 | |
| 62390 | 137 | declare if_split [split del] | 
| 19738 | 138 | |
| 139 | lemma channel_abstraction: "is_weak_ref_map reduce ch_ioa ch_fin_ioa" | |
| 140 | apply (simp (no_asm) add: is_weak_ref_map_def) | |
| 62002 | 141 | txt \<open>main-part\<close> | 
| 19738 | 142 | apply (rule allI)+ | 
| 143 | apply (rule imp_conj_lemma) | |
| 144 | apply (induct_tac "a") | |
| 62002 | 145 | txt \<open>2 cases\<close> | 
| 19738 | 146 | apply (simp_all (no_asm) cong del: if_weak_cong add: externals_def) | 
| 62002 | 147 | txt \<open>fst case\<close> | 
| 19738 | 148 | apply (rule impI) | 
| 149 | apply (rule disjI2) | |
| 150 | apply (rule reduce_hd) | |
| 62002 | 151 | txt \<open>snd case\<close> | 
| 19738 | 152 | apply (rule impI) | 
| 153 | apply (erule conjE)+ | |
| 154 | apply (erule disjE) | |
| 155 | apply (simp add: l_iff_red_nil) | |
| 156 | apply (erule hd_is_reduce_hd [THEN mp]) | |
| 157 | apply (simp add: l_iff_red_nil) | |
| 158 | apply (rule conjI) | |
| 159 | apply (erule hd_is_reduce_hd [THEN mp]) | |
| 160 | apply (rule bool_if_impl_or [THEN mp]) | |
| 161 | apply (erule reduce_tl) | |
| 162 | done | |
| 163 | ||
| 62390 | 164 | declare if_split [split] | 
| 19738 | 165 | |
| 166 | lemma sender_abstraction: "is_weak_ref_map reduce srch_ioa srch_fin_ioa" | |
| 62002 | 167 | apply (tactic \<open> | 
| 69597 | 168 | simp_tac (put_simpset HOL_ss \<^context> | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
48891diff
changeset | 169 |     addsimps [@{thm srch_fin_ioa_def}, @{thm rsch_fin_ioa_def},
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
48891diff
changeset | 170 |       @{thm srch_ioa_def}, @{thm rsch_ioa_def}, @{thm rename_through_pmap},
 | 
| 62002 | 171 |       @{thm channel_abstraction}]) 1\<close>)
 | 
| 19738 | 172 | done | 
| 173 | ||
| 174 | lemma receiver_abstraction: "is_weak_ref_map reduce rsch_ioa rsch_fin_ioa" | |
| 62002 | 175 | apply (tactic \<open> | 
| 69597 | 176 | simp_tac (put_simpset HOL_ss \<^context> | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
48891diff
changeset | 177 |     addsimps [@{thm srch_fin_ioa_def}, @{thm rsch_fin_ioa_def},
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
48891diff
changeset | 178 |       @{thm srch_ioa_def}, @{thm rsch_ioa_def}, @{thm rename_through_pmap},
 | 
| 62002 | 179 |       @{thm channel_abstraction}]) 1\<close>)
 | 
| 19738 | 180 | done | 
| 181 | ||
| 182 | ||
| 62002 | 183 | text \<open>3 thms that do not hold generally! The lucky restriction here is | 
| 184 | the absence of internal actions.\<close> | |
| 19738 | 185 | lemma sender_unchanged: "is_weak_ref_map (%id. id) sender_ioa sender_ioa" | 
| 186 | apply (simp (no_asm) add: is_weak_ref_map_def) | |
| 62002 | 187 | txt \<open>main-part\<close> | 
| 19738 | 188 | apply (rule allI)+ | 
| 189 | apply (induct_tac a) | |
| 62002 | 190 | txt \<open>7 cases\<close> | 
| 19738 | 191 | apply (simp_all (no_asm) add: externals_def) | 
| 192 | done | |
| 193 | ||
| 62002 | 194 | text \<open>2 copies of before\<close> | 
| 19738 | 195 | lemma receiver_unchanged: "is_weak_ref_map (%id. id) receiver_ioa receiver_ioa" | 
| 196 | apply (simp (no_asm) add: is_weak_ref_map_def) | |
| 62002 | 197 | txt \<open>main-part\<close> | 
| 19738 | 198 | apply (rule allI)+ | 
| 199 | apply (induct_tac a) | |
| 62002 | 200 | txt \<open>7 cases\<close> | 
| 19738 | 201 | apply (simp_all (no_asm) add: externals_def) | 
| 202 | done | |
| 203 | ||
| 204 | lemma env_unchanged: "is_weak_ref_map (%id. id) env_ioa env_ioa" | |
| 205 | apply (simp (no_asm) add: is_weak_ref_map_def) | |
| 62002 | 206 | txt \<open>main-part\<close> | 
| 19738 | 207 | apply (rule allI)+ | 
| 208 | apply (induct_tac a) | |
| 62002 | 209 | txt \<open>7 cases\<close> | 
| 19738 | 210 | apply (simp_all (no_asm) add: externals_def) | 
| 211 | done | |
| 212 | ||
| 213 | ||
| 214 | lemma compat_single_ch: "compatible srch_ioa rsch_ioa" | |
| 215 | apply (simp add: compatible_def Int_def) | |
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39159diff
changeset | 216 | apply (rule set_eqI) | 
| 19738 | 217 | apply (induct_tac x) | 
| 218 | apply simp_all | |
| 219 | done | |
| 220 | ||
| 62002 | 221 | text \<open>totally the same as before\<close> | 
| 19738 | 222 | lemma compat_single_fin_ch: "compatible srch_fin_ioa rsch_fin_ioa" | 
| 223 | apply (simp add: compatible_def Int_def) | |
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39159diff
changeset | 224 | apply (rule set_eqI) | 
| 19738 | 225 | apply (induct_tac x) | 
| 226 | apply simp_all | |
| 227 | done | |
| 228 | ||
| 229 | lemmas del_simps = trans_of_def srch_asig_def rsch_asig_def | |
| 230 | asig_of_def actions_def srch_trans_def rsch_trans_def srch_ioa_def | |
| 231 | srch_fin_ioa_def rsch_fin_ioa_def rsch_ioa_def sender_trans_def | |
| 232 | receiver_trans_def set_lemmas | |
| 233 | ||
| 61999 | 234 | lemma compat_rec: "compatible receiver_ioa (srch_ioa \<parallel> rsch_ioa)" | 
| 19738 | 235 | apply (simp del: del_simps | 
| 236 | add: compatible_def asig_of_par asig_comp_def actions_def Int_def) | |
| 237 | apply simp | |
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39159diff
changeset | 238 | apply (rule set_eqI) | 
| 19738 | 239 | apply (induct_tac x) | 
| 240 | apply simp_all | |
| 241 | done | |
| 242 | ||
| 62002 | 243 | text \<open>5 proofs totally the same as before\<close> | 
| 61999 | 244 | lemma compat_rec_fin: "compatible receiver_ioa (srch_fin_ioa \<parallel> rsch_fin_ioa)" | 
| 19738 | 245 | apply (simp del: del_simps | 
| 246 | add: compatible_def asig_of_par asig_comp_def actions_def Int_def) | |
| 247 | apply simp | |
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39159diff
changeset | 248 | apply (rule set_eqI) | 
| 19738 | 249 | apply (induct_tac x) | 
| 250 | apply simp_all | |
| 251 | done | |
| 252 | ||
| 253 | lemma compat_sen: "compatible sender_ioa | |
| 61999 | 254 | (receiver_ioa \<parallel> srch_ioa \<parallel> rsch_ioa)" | 
| 19738 | 255 | apply (simp del: del_simps | 
| 256 | add: compatible_def asig_of_par asig_comp_def actions_def Int_def) | |
| 257 | apply simp | |
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39159diff
changeset | 258 | apply (rule set_eqI) | 
| 19738 | 259 | apply (induct_tac x) | 
| 260 | apply simp_all | |
| 261 | done | |
| 262 | ||
| 263 | lemma compat_sen_fin: "compatible sender_ioa | |
| 61999 | 264 | (receiver_ioa \<parallel> srch_fin_ioa \<parallel> rsch_fin_ioa)" | 
| 19738 | 265 | apply (simp del: del_simps | 
| 266 | add: compatible_def asig_of_par asig_comp_def actions_def Int_def) | |
| 267 | apply simp | |
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39159diff
changeset | 268 | apply (rule set_eqI) | 
| 19738 | 269 | apply (induct_tac x) | 
| 270 | apply simp_all | |
| 271 | done | |
| 272 | ||
| 273 | lemma compat_env: "compatible env_ioa | |
| 61999 | 274 | (sender_ioa \<parallel> receiver_ioa \<parallel> srch_ioa \<parallel> rsch_ioa)" | 
| 19738 | 275 | apply (simp del: del_simps | 
| 276 | add: compatible_def asig_of_par asig_comp_def actions_def Int_def) | |
| 277 | apply simp | |
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39159diff
changeset | 278 | apply (rule set_eqI) | 
| 19738 | 279 | apply (induct_tac x) | 
| 280 | apply simp_all | |
| 281 | done | |
| 282 | ||
| 283 | lemma compat_env_fin: "compatible env_ioa | |
| 61999 | 284 | (sender_ioa \<parallel> receiver_ioa \<parallel> srch_fin_ioa \<parallel> rsch_fin_ioa)" | 
| 19738 | 285 | apply (simp del: del_simps | 
| 286 | add: compatible_def asig_of_par asig_comp_def actions_def Int_def) | |
| 287 | apply simp | |
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39159diff
changeset | 288 | apply (rule set_eqI) | 
| 19738 | 289 | apply (induct_tac x) | 
| 290 | apply simp_all | |
| 291 | done | |
| 292 | ||
| 293 | ||
| 62002 | 294 | text \<open>lemmata about externals of channels\<close> | 
| 19738 | 295 | lemma ext_single_ch: "externals(asig_of(srch_fin_ioa)) = externals(asig_of(srch_ioa)) & | 
| 296 | externals(asig_of(rsch_fin_ioa)) = externals(asig_of(rsch_ioa))" | |
| 297 | by (simp add: externals_def) | |
| 298 | ||
| 299 | ||
| 62002 | 300 | subsection \<open>Soundness of Abstraction\<close> | 
| 19738 | 301 | |
| 302 | lemmas ext_simps = externals_of_par ext_single_ch | |
| 303 | and compat_simps = compat_single_ch compat_single_fin_ch compat_rec | |
| 304 | compat_rec_fin compat_sen compat_sen_fin compat_env compat_env_fin | |
| 305 | and abstractions = env_unchanged sender_unchanged | |
| 306 | receiver_unchanged sender_abstraction receiver_abstraction | |
| 307 | ||
| 308 | ||
| 309 | (* FIX: this proof should be done with compositionality on trace level, not on | |
| 310 | weak_ref_map level, as done here with fxg_is_weak_ref_map_of_product_IOA | |
| 311 | ||
| 312 | Goal "is_weak_ref_map abs system_ioa system_fin_ioa" | |
| 313 | ||
| 314 | by (simp_tac (impl_ss delsimps ([srch_ioa_def, rsch_ioa_def, srch_fin_ioa_def, | |
| 315 | rsch_fin_ioa_def] @ env_ioas @ impl_ioas) | |
| 316 | addsimps [system_def, system_fin_def, abs_def, | |
| 317 | impl_ioa_def, impl_fin_ioa_def, sys_IOA, | |
| 318 | sys_fin_IOA]) 1); | |
| 319 | ||
| 320 | by (REPEAT (EVERY[rtac fxg_is_weak_ref_map_of_product_IOA 1, | |
| 321 | simp_tac (ss addsimps abstractions) 1, | |
| 322 | rtac conjI 1])); | |
| 323 | ||
| 324 | by (ALLGOALS (simp_tac (ss addsimps ext_ss @ compat_ss))); | |
| 325 | ||
| 326 | qed "system_refinement"; | |
| 327 | *) | |
| 17244 | 328 | |
| 3072 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 329 | end |