removed legacy ML scripts;
authorwenzelm
Sat, 27 May 2006 19:49:36 +0200
changeset 19738 1ac610922636
parent 19737 3b8920131be2
child 19739 c58ef2aa5430
removed legacy ML scripts;
src/HOLCF/IOA/ABP/Abschannel.thy
src/HOLCF/IOA/ABP/Abschannel_finite.thy
src/HOLCF/IOA/ABP/Action.thy
src/HOLCF/IOA/ABP/Correctness.ML
src/HOLCF/IOA/ABP/Correctness.thy
src/HOLCF/IOA/ABP/Env.thy
src/HOLCF/IOA/ABP/Impl.thy
src/HOLCF/IOA/ABP/Impl_finite.thy
src/HOLCF/IOA/ABP/Lemmas.ML
src/HOLCF/IOA/ABP/Lemmas.thy
src/HOLCF/IOA/ABP/Packet.thy
src/HOLCF/IOA/ABP/Receiver.thy
src/HOLCF/IOA/ABP/Sender.thy
src/HOLCF/IsaMakefile
--- a/src/HOLCF/IOA/ABP/Abschannel.thy	Sat May 27 19:49:07 2006 +0200
+++ b/src/HOLCF/IOA/ABP/Abschannel.thy	Sat May 27 19:49:36 2006 +0200
@@ -85,8 +85,6 @@
             S_ack(b) => None |
             R_ack(b) => None"
 
-ML {* use_legacy_bindings (the_context ()) *}
-
 end
 
 
--- a/src/HOLCF/IOA/ABP/Abschannel_finite.thy	Sat May 27 19:49:07 2006 +0200
+++ b/src/HOLCF/IOA/ABP/Abschannel_finite.thy	Sat May 27 19:49:36 2006 +0200
@@ -59,6 +59,4 @@
 
 ch_fin_ioa_def: "ch_fin_ioa == (ch_fin_asig, {[]}, ch_fin_trans,{},{})"
 
-ML {* use_legacy_bindings (the_context ()) *}
-
 end
--- a/src/HOLCF/IOA/ABP/Action.thy	Sat May 27 19:49:07 2006 +0200
+++ b/src/HOLCF/IOA/ABP/Action.thy	Sat May 27 19:49:36 2006 +0200
@@ -14,6 +14,4 @@
   | S_pkt "'m packet" | R_pkt "'m packet"
   | S_ack bool | R_ack bool
 
-ML {* use_legacy_bindings (the_context ()) *}
-
 end
--- a/src/HOLCF/IOA/ABP/Correctness.ML	Sat May 27 19:49:07 2006 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,311 +0,0 @@
-(*  Title:      HOLCF/IOA/ABP/Correctness.ML
-    ID:         $Id$
-    Author:     Olaf Mueller
-*)
-
-
-Delsimps [split_paired_All,Collect_empty_eq];
-Addsimps 
- ([srch_asig_def, rsch_asig_def, rsch_ioa_def, srch_ioa_def, ch_ioa_def, 
-   ch_asig_def, srch_actions_def, rsch_actions_def, rename_def, rename_set_def, asig_of_def, 
-   actions_def, exis_elim, srch_trans_def, rsch_trans_def, ch_trans_def, 
-   trans_of_def] @ asig_projections @ set_lemmas);
-
-val abschannel_fin = [srch_fin_asig_def, rsch_fin_asig_def, 
-                      rsch_fin_ioa_def, srch_fin_ioa_def, 
-                      ch_fin_ioa_def,ch_fin_trans_def,ch_fin_asig_def];
-Addsimps abschannel_fin;
-
-val impl_ioas =  [sender_ioa_def,receiver_ioa_def];
-val impl_trans = [sender_trans_def,receiver_trans_def];
-val impl_asigs = [sender_asig_def,receiver_asig_def];
-Addcongs [let_weak_cong];
-Addsimps [Let_def, ioa_triple_proj, starts_of_par];
-
-val env_ioas = [env_ioa_def,env_asig_def,env_trans_def];
-val hom_ioas = env_ioas @ impl_ioas @ impl_trans @ impl_asigs @ 
-               asig_projections @ set_lemmas;
-Addsimps hom_ioas;
-
-
-
-(* auxiliary function *)
-fun rotate n i = EVERY(replicate n (etac revcut_rl i));
-
-(* lemmas about reduce *)
-
-Goal "(reduce(l)=[]) = (l=[])";  
-by (induct_tac "l" 1);
-by (auto_tac (claset(), simpset() addsplits [thm"list.split"]));
-val l_iff_red_nil = result();
-
-Goal "s~=[] --> hd(s)=hd(reduce(s))";
-by (induct_tac "s" 1);
-by (auto_tac (claset(), simpset() addsplits [thm"list.split"]));
-qed"hd_is_reduce_hd";
-
-(* to be used in the following Lemma *)
-Goal "l~=[] --> reverse(reduce(l))~=[]";
-by (induct_tac "l" 1);
-by (auto_tac (claset(), simpset() addsplits [thm"list.split"]));
-qed_spec_mp "rev_red_not_nil";
-
-(* shows applicability of the induction hypothesis of the following Lemma 1 *)
-Goal "l~=[] ==> hd(reverse(reduce(a#l))) = hd(reverse(reduce(l)))";
- by (Simp_tac 1);
- by (auto_tac (claset(), HOL_ss addsplits [thm"list.split"]
-                                addsimps (reverse.simps @ [hd_append,
-					  rev_red_not_nil])));
-qed"last_ind_on_first";
-
-val impl_ss = simpset() delsimps [reduce.reduce_Cons];
-
-(* Main Lemma 1 for S_pkt in showing that reduce is refinement  *) 
-Goal 
-   "if x=hd(reverse(reduce(l))) & reduce(l)~=[] then   \
-\      reduce(l@[x])=reduce(l) else                  \
-\      reduce(l@[x])=reduce(l)@[x]"; 
-by (stac split_if 1);
-by (rtac conjI 1);
-(* --> *)
-by (induct_tac "l" 1);
-by (Simp_tac 1);
-by (case_tac "list=[]" 1);
- by (asm_full_simp_tac (simpset() addsimps reverse.simps) 1);
- by (rtac impI 1);
-by (Simp_tac 1);
-by (cut_inst_tac [("l","list")] cons_not_nil 1);
- by (asm_full_simp_tac impl_ss 1);
- by (REPEAT (etac exE 1));
- by (hyp_subst_tac 1);
-by (asm_full_simp_tac (impl_ss addsimps [last_ind_on_first,l_iff_red_nil]) 1);
-(* <-- *)
-by (simp_tac (simpset() addsimps [and_de_morgan_and_absorbe,l_iff_red_nil]) 1);
-by (induct_tac "l" 1);
-by (Simp_tac 1);
-by (case_tac "list=[]" 1);
-by (cut_inst_tac [("l","list")] cons_not_nil 2);
-by (Asm_full_simp_tac 1);
-by (auto_tac (claset(), 
-	      impl_ss addsimps [last_ind_on_first,l_iff_red_nil] 
-                      setloop split_tac [split_if]));
-by (Asm_full_simp_tac 1);
-qed"reduce_hd";
-
-
-(* Main Lemma 2 for R_pkt in showing that reduce is refinement *)
-Goal "s~=[] ==>  \
-\    if hd(s)=hd(tl(s)) & tl(s)~=[] then    \
-\      reduce(tl(s))=reduce(s) else      \
-\      reduce(tl(s))=tl(reduce(s))"; 
-by (cut_inst_tac [("l","s")] cons_not_nil 1);
-by (Asm_full_simp_tac 1);
-by (REPEAT (etac exE 1));
-by (auto_tac (claset(), simpset() addsplits [thm"list.split"]));
-val reduce_tl =result();
-
-
-(*******************************************************************
-          C h a n n e l   A b s t r a c t i o n
- *******************************************************************)
-
-Delsplits [split_if]; 
-
-Goal "is_weak_ref_map reduce ch_ioa ch_fin_ioa";
-by (simp_tac (simpset() addsimps [is_weak_ref_map_def]) 1);
-(* ---------------- main-part ------------------- *)
-by (REPEAT (rtac allI 1));
-by (rtac imp_conj_lemma 1);
-by (induct_tac "a" 1);
-(* ----------------- 2 cases ---------------------*)
-by (ALLGOALS (simp_tac (simpset() delcongs [if_weak_cong]
-				  addsimps [externals_def])));
-(* fst case --------------------------------------*)
- by (rtac impI 1);
- by (rtac disjI2 1);
-by (rtac reduce_hd 1);
-(* snd case --------------------------------------*)
- by (rtac impI 1);
- by (REPEAT (etac conjE 1));
- by (etac disjE 1);
-by (asm_full_simp_tac (simpset() addsimps [l_iff_red_nil]) 1);
-by (etac (hd_is_reduce_hd RS mp) 1); 
-by (asm_full_simp_tac (simpset() addsimps [l_iff_red_nil]) 1);
-by (rtac conjI 1);
-by (etac (hd_is_reduce_hd RS mp) 1); 
-by (rtac (bool_if_impl_or RS mp) 1);
-by (etac reduce_tl 1);
-qed"channel_abstraction";
-
-Addsplits [split_if]; 
-
-Goal "is_weak_ref_map reduce srch_ioa srch_fin_ioa";
-by (simp_tac (HOL_ss addsimps [srch_fin_ioa_def,rsch_fin_ioa_def,
- srch_ioa_def,rsch_ioa_def,rename_through_pmap,channel_abstraction]) 1);
-qed"sender_abstraction";
-
-Goal "is_weak_ref_map reduce rsch_ioa rsch_fin_ioa";
-by (simp_tac (HOL_ss addsimps [srch_fin_ioa_def,rsch_fin_ioa_def,
- srch_ioa_def,rsch_ioa_def,rename_through_pmap,channel_abstraction]) 1);
-qed"receiver_abstraction";
-
-
-(* 3 thms that do not hold generally! The lucky restriction here is 
-   the absence of internal actions. *)
-Goal "is_weak_ref_map (%id. id) sender_ioa sender_ioa";
-by (simp_tac (simpset() addsimps [is_weak_ref_map_def]) 1);
-(* main-part *)
-by (REPEAT (rtac allI 1));
-by (induct_tac "a" 1);
-(* 7 cases *)
-by (ALLGOALS (simp_tac (simpset() addsimps [externals_def])));
-qed"sender_unchanged";
-
-(* 2 copies of before *)
-Goal "is_weak_ref_map (%id. id) receiver_ioa receiver_ioa";
-by (simp_tac (simpset() addsimps [is_weak_ref_map_def]) 1);
-(* main-part *)
-by (REPEAT (rtac allI 1));
-by (induct_tac "a" 1);
-(* 7 cases *)
-by (ALLGOALS(simp_tac (simpset() addsimps [externals_def])));
-qed"receiver_unchanged";
-
-Goal "is_weak_ref_map (%id. id) env_ioa env_ioa";
-by (simp_tac (simpset() addsimps [is_weak_ref_map_def]) 1);
-(* main-part *)
-by (REPEAT (rtac allI 1));
-by (induct_tac "a" 1);
-(* 7 cases *)
-by (ALLGOALS(simp_tac (simpset() addsimps [externals_def])));
-qed"env_unchanged";
-
-
-Goal "compatible srch_ioa rsch_ioa"; 
-by (simp_tac (simpset() addsimps [compatible_def,Int_def]) 1);
-by (rtac set_ext 1);
-by (induct_tac "x" 1);
-by (ALLGOALS(Simp_tac));
-val compat_single_ch = result();
-
-(* totally the same as before *)
-Goal "compatible srch_fin_ioa rsch_fin_ioa"; 
-by (simp_tac (simpset() addsimps [compatible_def,Int_def]) 1);
-by (rtac set_ext 1);
-by (induct_tac "x" 1);
-by (ALLGOALS(Simp_tac));
-val compat_single_fin_ch = result();
-
-val ss =
-  simpset() delsimps ([trans_of_def, srch_asig_def,rsch_asig_def,
-                      asig_of_def, actions_def, srch_trans_def, rsch_trans_def,
-                      srch_ioa_def, srch_fin_ioa_def, rsch_fin_ioa_def, 
-                      rsch_ioa_def, sender_trans_def,
-                      receiver_trans_def] @ set_lemmas);
-
-Goal "compatible receiver_ioa (srch_ioa || rsch_ioa)";
-by (simp_tac (ss addsimps [compatible_def,asig_of_par,asig_comp_def,
-			   actions_def,Int_def]) 1);
-by (Simp_tac 1);
-by (rtac set_ext 1);
-by (induct_tac "x" 1);
-by (ALLGOALS Simp_tac);
-val compat_rec = result();
-
-(* 5 proofs totally the same as before *)
-Goal "compatible receiver_ioa (srch_fin_ioa || rsch_fin_ioa)";
-by (simp_tac (ss addsimps [compatible_def,asig_of_par,asig_comp_def,
-			actions_def,Int_def]) 1);
-by (Simp_tac 1);
-by (rtac set_ext 1);
-by (induct_tac "x" 1);
-by (ALLGOALS Simp_tac);
-val compat_rec_fin =result();
-
-Goal "compatible sender_ioa \
-\      (receiver_ioa || srch_ioa || rsch_ioa)";
-by (simp_tac (ss addsimps [compatible_def,asig_of_par,asig_comp_def,
-			actions_def,Int_def]) 1);
-by (Simp_tac 1);
-by (rtac set_ext 1);
-by (induct_tac "x" 1);
-by (ALLGOALS(Simp_tac));
-val compat_sen=result();
-
-Goal "compatible sender_ioa\
-\      (receiver_ioa || srch_fin_ioa || rsch_fin_ioa)";
-by (simp_tac (ss addsimps [compatible_def,asig_of_par,asig_comp_def,
-				actions_def,Int_def]) 1);
-by (Simp_tac 1);
-by (rtac set_ext 1);
-by (induct_tac "x" 1);
-by (ALLGOALS(Simp_tac));
-val compat_sen_fin =result();
-
-Goal "compatible env_ioa\
-\      (sender_ioa || receiver_ioa || srch_ioa || rsch_ioa)";
-by (simp_tac (ss addsimps [compatible_def,asig_of_par,asig_comp_def,
-				actions_def,Int_def]) 1);
-by (Simp_tac 1);
-by (rtac set_ext 1);
-by (induct_tac "x" 1);
-by (ALLGOALS(Simp_tac));
-val compat_env=result();
-
-Goal "compatible env_ioa\
-\      (sender_ioa || receiver_ioa || srch_fin_ioa || rsch_fin_ioa)";
-by (simp_tac (ss addsimps [compatible_def,asig_of_par,asig_comp_def,
-			actions_def,Int_def]) 1);
-by (Simp_tac 1);
-by (rtac set_ext 1);
-by (induct_tac "x" 1);
-by (ALLGOALS Simp_tac);
-val compat_env_fin=result();
-
-
-(* lemmata about externals of channels *)
-Goal "externals(asig_of(srch_fin_ioa)) = externals(asig_of(srch_ioa)) &  \
-\ externals(asig_of(rsch_fin_ioa)) = externals(asig_of(rsch_ioa))";
-   by (simp_tac (simpset() addsimps [externals_def]) 1);
-val ext_single_ch = result();
-
-
-
-
-
-val ext_ss = [externals_of_par,ext_single_ch];
-val compat_ss = [compat_single_ch,compat_single_fin_ch,compat_rec,
-         compat_rec_fin,compat_sen,compat_sen_fin,compat_env,compat_env_fin];
-val abstractions = [env_unchanged,sender_unchanged,
-          receiver_unchanged,sender_abstraction,receiver_abstraction];
-
-
-(************************************************************************
-            S o u n d n e s s   o f   A b s t r a c t i o n        
-*************************************************************************)
-
-val ss = simpset() delsimps ([asig_of_def, srch_ioa_def, rsch_ioa_def,
-                             srch_fin_ioa_def, rsch_fin_ioa_def] @
-                            impl_ioas @ env_ioas);
-
-(* FIX: this proof should be done with compositionality on trace level, not on 
-        weak_ref_map level, as done here with fxg_is_weak_ref_map_of_product_IOA 
-
-Goal "is_weak_ref_map  abs  system_ioa  system_fin_ioa";
-
-by (simp_tac (impl_ss delsimps ([srch_ioa_def, rsch_ioa_def, srch_fin_ioa_def,
-                                 rsch_fin_ioa_def] @ env_ioas @ impl_ioas)
-                      addsimps [system_def, system_fin_def, abs_def,
-                                impl_ioa_def, impl_fin_ioa_def, sys_IOA,
-                                sys_fin_IOA]) 1);
-
-by (REPEAT (EVERY[rtac fxg_is_weak_ref_map_of_product_IOA 1, 
-                  simp_tac (ss addsimps abstractions) 1,
-                  rtac conjI 1]));
-
-by (ALLGOALS (simp_tac (ss addsimps ext_ss @ compat_ss))); 
-
-qed "system_refinement";
-
-
-*)
--- a/src/HOLCF/IOA/ABP/Correctness.thy	Sat May 27 19:49:07 2006 +0200
+++ b/src/HOLCF/IOA/ABP/Correctness.thy	Sat May 27 19:49:36 2006 +0200
@@ -38,6 +38,287 @@
   sys_IOA:     "IOA system_ioa"
   sys_fin_IOA: "IOA system_fin_ioa"
 
-ML {* use_legacy_bindings (the_context ()) *}
+
+
+declare split_paired_All [simp del] Collect_empty_eq [simp del]
+
+lemmas [simp] =
+  srch_asig_def rsch_asig_def rsch_ioa_def srch_ioa_def ch_ioa_def
+  ch_asig_def srch_actions_def rsch_actions_def rename_def rename_set_def asig_of_def
+  actions_def exis_elim srch_trans_def rsch_trans_def ch_trans_def
+  trans_of_def asig_projections set_lemmas
+
+lemmas abschannel_fin [simp] =
+  srch_fin_asig_def rsch_fin_asig_def
+  rsch_fin_ioa_def srch_fin_ioa_def
+  ch_fin_ioa_def ch_fin_trans_def ch_fin_asig_def
+
+lemmas impl_ioas = sender_ioa_def receiver_ioa_def
+  and impl_trans = sender_trans_def receiver_trans_def
+  and impl_asigs = sender_asig_def receiver_asig_def
+
+declare let_weak_cong [cong]
+declare Let_def [simp] ioa_triple_proj [simp] starts_of_par [simp]
+
+lemmas env_ioas = env_ioa_def env_asig_def env_trans_def
+  and hom_ioas [simp] = env_ioas impl_ioas impl_trans impl_asigs asig_projections set_lemmas
+
+
+subsection {* lemmas about reduce *}
+
+lemma l_iff_red_nil: "(reduce l = []) = (l = [])"
+  by (induct l) (auto split: list.split)
+
+lemma hd_is_reduce_hd: "s ~= [] --> hd s = hd (reduce s)"
+  by (induct s) (auto split: list.split)
+
+text {* to be used in the following Lemma *}
+lemma rev_red_not_nil [rule_format]:
+    "l ~= [] --> reverse (reduce l) ~= []"
+  by (induct l) (auto split: list.split)
+
+text {* shows applicability of the induction hypothesis of the following Lemma 1 *}
+lemma last_ind_on_first:
+    "l ~= [] ==> hd (reverse (reduce (a # l))) = hd (reverse (reduce l))"
+  apply simp
+  apply (tactic {* auto_tac (claset(),
+    HOL_ss addsplits [thm"list.split"]
+    addsimps (thms "reverse.simps" @ [thm "hd_append", thm "rev_red_not_nil"])) *})
+  done
+
+text {* Main Lemma 1 for @{text "S_pkt"} in showing that reduce is refinement. *}
+lemma reduce_hd:
+   "if x=hd(reverse(reduce(l))) & reduce(l)~=[] then
+       reduce(l@[x])=reduce(l) else
+       reduce(l@[x])=reduce(l)@[x]"
+apply (simplesubst split_if)
+apply (rule conjI)
+txt {* @{text "-->"} *}
+apply (induct_tac "l")
+apply (simp (no_asm))
+apply (case_tac "list=[]")
+ apply (simp add: reverse.simps)
+ apply (rule impI)
+apply (simp (no_asm))
+apply (cut_tac l = "list" in cons_not_nil)
+ apply (simp del: reduce_Cons)
+ apply (erule exE)+
+ apply hypsubst
+apply (simp del: reduce_Cons add: last_ind_on_first l_iff_red_nil)
+txt {* @{text "<--"} *}
+apply (simp (no_asm) add: and_de_morgan_and_absorbe l_iff_red_nil)
+apply (induct_tac "l")
+apply (simp (no_asm))
+apply (case_tac "list=[]")
+apply (cut_tac [2] l = "list" in cons_not_nil)
+apply simp
+apply (auto simp del: reduce_Cons simp add: last_ind_on_first l_iff_red_nil split: split_if)
+apply simp
+done
+
+
+text {* Main Lemma 2 for R_pkt in showing that reduce is refinement. *}
+lemma reduce_tl: "s~=[] ==>
+     if hd(s)=hd(tl(s)) & tl(s)~=[] then
+       reduce(tl(s))=reduce(s) else
+       reduce(tl(s))=tl(reduce(s))"
+apply (cut_tac l = "s" in cons_not_nil)
+apply simp
+apply (erule exE)+
+apply (auto split: list.split)
+done
+
+
+subsection {* Channel Abstraction *}
+
+declare split_if [split del]
+
+lemma channel_abstraction: "is_weak_ref_map reduce ch_ioa ch_fin_ioa"
+apply (simp (no_asm) add: is_weak_ref_map_def)
+txt {* main-part *}
+apply (rule allI)+
+apply (rule imp_conj_lemma)
+apply (induct_tac "a")
+txt {* 2 cases *}
+apply (simp_all (no_asm) cong del: if_weak_cong add: externals_def)
+txt {* fst case *}
+ apply (rule impI)
+ apply (rule disjI2)
+apply (rule reduce_hd)
+txt {* snd case *}
+ apply (rule impI)
+ apply (erule conjE)+
+ apply (erule disjE)
+apply (simp add: l_iff_red_nil)
+apply (erule hd_is_reduce_hd [THEN mp])
+apply (simp add: l_iff_red_nil)
+apply (rule conjI)
+apply (erule hd_is_reduce_hd [THEN mp])
+apply (rule bool_if_impl_or [THEN mp])
+apply (erule reduce_tl)
+done
+
+declare split_if [split]
+
+lemma sender_abstraction: "is_weak_ref_map reduce srch_ioa srch_fin_ioa"
+apply (tactic {*
+  simp_tac (HOL_ss addsimps [thm "srch_fin_ioa_def", thm "rsch_fin_ioa_def",
+    thm "srch_ioa_def", thm "rsch_ioa_def", thm "rename_through_pmap",
+    thm "channel_abstraction"]) 1 *})
+done
+
+lemma receiver_abstraction: "is_weak_ref_map reduce rsch_ioa rsch_fin_ioa"
+apply (tactic {*
+  simp_tac (HOL_ss addsimps [thm "srch_fin_ioa_def", thm "rsch_fin_ioa_def",
+    thm "srch_ioa_def", thm "rsch_ioa_def", thm "rename_through_pmap",
+    thm "channel_abstraction"]) 1 *})
+done
+
+
+text {* 3 thms that do not hold generally! The lucky restriction here is
+   the absence of internal actions. *}
+lemma sender_unchanged: "is_weak_ref_map (%id. id) sender_ioa sender_ioa"
+apply (simp (no_asm) add: is_weak_ref_map_def)
+txt {* main-part *}
+apply (rule allI)+
+apply (induct_tac a)
+txt {* 7 cases *}
+apply (simp_all (no_asm) add: externals_def)
+done
+
+text {* 2 copies of before *}
+lemma receiver_unchanged: "is_weak_ref_map (%id. id) receiver_ioa receiver_ioa"
+apply (simp (no_asm) add: is_weak_ref_map_def)
+txt {* main-part *}
+apply (rule allI)+
+apply (induct_tac a)
+txt {* 7 cases *}
+apply (simp_all (no_asm) add: externals_def)
+done
+
+lemma env_unchanged: "is_weak_ref_map (%id. id) env_ioa env_ioa"
+apply (simp (no_asm) add: is_weak_ref_map_def)
+txt {* main-part *}
+apply (rule allI)+
+apply (induct_tac a)
+txt {* 7 cases *}
+apply (simp_all (no_asm) add: externals_def)
+done
+
+
+lemma compat_single_ch: "compatible srch_ioa rsch_ioa"
+apply (simp add: compatible_def Int_def)
+apply (rule set_ext)
+apply (induct_tac x)
+apply simp_all
+done
+
+text {* totally the same as before *}
+lemma compat_single_fin_ch: "compatible srch_fin_ioa rsch_fin_ioa"
+apply (simp add: compatible_def Int_def)
+apply (rule set_ext)
+apply (induct_tac x)
+apply simp_all
+done
+
+lemmas del_simps = trans_of_def srch_asig_def rsch_asig_def
+  asig_of_def actions_def srch_trans_def rsch_trans_def srch_ioa_def
+  srch_fin_ioa_def rsch_fin_ioa_def rsch_ioa_def sender_trans_def
+  receiver_trans_def set_lemmas
+
+lemma compat_rec: "compatible receiver_ioa (srch_ioa || rsch_ioa)"
+apply (simp del: del_simps
+  add: compatible_def asig_of_par asig_comp_def actions_def Int_def)
+apply simp
+apply (rule set_ext)
+apply (induct_tac x)
+apply simp_all
+done
+
+text {* 5 proofs totally the same as before *}
+lemma compat_rec_fin: "compatible receiver_ioa (srch_fin_ioa || rsch_fin_ioa)"
+apply (simp del: del_simps
+  add: compatible_def asig_of_par asig_comp_def actions_def Int_def)
+apply simp
+apply (rule set_ext)
+apply (induct_tac x)
+apply simp_all
+done
+
+lemma compat_sen: "compatible sender_ioa
+       (receiver_ioa || srch_ioa || rsch_ioa)"
+apply (simp del: del_simps
+  add: compatible_def asig_of_par asig_comp_def actions_def Int_def)
+apply simp
+apply (rule set_ext)
+apply (induct_tac x)
+apply simp_all
+done
+
+lemma compat_sen_fin: "compatible sender_ioa
+       (receiver_ioa || srch_fin_ioa || rsch_fin_ioa)"
+apply (simp del: del_simps
+  add: compatible_def asig_of_par asig_comp_def actions_def Int_def)
+apply simp
+apply (rule set_ext)
+apply (induct_tac x)
+apply simp_all
+done
+
+lemma compat_env: "compatible env_ioa
+       (sender_ioa || receiver_ioa || srch_ioa || rsch_ioa)"
+apply (simp del: del_simps
+  add: compatible_def asig_of_par asig_comp_def actions_def Int_def)
+apply simp
+apply (rule set_ext)
+apply (induct_tac x)
+apply simp_all
+done
+
+lemma compat_env_fin: "compatible env_ioa
+       (sender_ioa || receiver_ioa || srch_fin_ioa || rsch_fin_ioa)"
+apply (simp del: del_simps
+  add: compatible_def asig_of_par asig_comp_def actions_def Int_def)
+apply simp
+apply (rule set_ext)
+apply (induct_tac x)
+apply simp_all
+done
+
+
+text {* lemmata about externals of channels *}
+lemma ext_single_ch: "externals(asig_of(srch_fin_ioa)) = externals(asig_of(srch_ioa)) &
+    externals(asig_of(rsch_fin_ioa)) = externals(asig_of(rsch_ioa))"
+  by (simp add: externals_def)
+
+
+subsection {* Soundness of Abstraction *}
+
+lemmas ext_simps = externals_of_par ext_single_ch
+  and compat_simps = compat_single_ch compat_single_fin_ch compat_rec
+    compat_rec_fin compat_sen compat_sen_fin compat_env compat_env_fin
+  and abstractions = env_unchanged sender_unchanged
+    receiver_unchanged sender_abstraction receiver_abstraction
+
+
+(* FIX: this proof should be done with compositionality on trace level, not on
+        weak_ref_map level, as done here with fxg_is_weak_ref_map_of_product_IOA
+
+Goal "is_weak_ref_map  abs  system_ioa  system_fin_ioa"
+
+by (simp_tac (impl_ss delsimps ([srch_ioa_def, rsch_ioa_def, srch_fin_ioa_def,
+                                 rsch_fin_ioa_def] @ env_ioas @ impl_ioas)
+                      addsimps [system_def, system_fin_def, abs_def,
+                                impl_ioa_def, impl_fin_ioa_def, sys_IOA,
+                                sys_fin_IOA]) 1);
+
+by (REPEAT (EVERY[rtac fxg_is_weak_ref_map_of_product_IOA 1,
+                  simp_tac (ss addsimps abstractions) 1,
+                  rtac conjI 1]));
+
+by (ALLGOALS (simp_tac (ss addsimps ext_ss @ compat_ss)));
+
+qed "system_refinement";
+*)
 
 end
--- a/src/HOLCF/IOA/ABP/Env.thy	Sat May 27 19:49:07 2006 +0200
+++ b/src/HOLCF/IOA/ABP/Env.thy	Sat May 27 19:49:36 2006 +0200
@@ -38,6 +38,4 @@
 consts
   "next"     :: "'m env_state => bool"
 
-ML {* use_legacy_bindings (the_context ()) *}
-
 end
--- a/src/HOLCF/IOA/ABP/Impl.thy	Sat May 27 19:49:07 2006 +0200
+++ b/src/HOLCF/IOA/ABP/Impl.thy	Sat May 27 19:49:36 2006 +0200
@@ -30,6 +30,4 @@
  rsch        :: "'m impl_state => bool list"
  "rsch == snd o snd o snd"
 
-ML {* use_legacy_bindings (the_context ()) *}
-
 end
--- a/src/HOLCF/IOA/ABP/Impl_finite.thy	Sat May 27 19:49:07 2006 +0200
+++ b/src/HOLCF/IOA/ABP/Impl_finite.thy	Sat May 27 19:49:36 2006 +0200
@@ -32,6 +32,4 @@
  rsch_fin        :: "'m impl_fin_state => bool list"
  "rsch_fin == snd o snd o snd"
 
-ML {* use_legacy_bindings (the_context ()) *}
-
 end
--- a/src/HOLCF/IOA/ABP/Lemmas.ML	Sat May 27 19:49:07 2006 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,54 +0,0 @@
-(*  Title:      HOLCF/IOA/ABP/Lemmas.ML
-    ID:         $Id$
-    Author:     Olaf Mueller
-*)
-
-(* Logic *)
-
-Goal "(~(A&B)) = ((~A)&B| ~B)";
-by (Fast_tac 1);
-qed "and_de_morgan_and_absorbe";
-
-Goal "(if C then A else B) --> (A|B)";
-by (stac split_if 1);
-by (Fast_tac 1);
-qed "bool_if_impl_or";
-
-Goal "(? x. x=P & Q(x)) = Q(P)";
-by (Fast_tac 1);
-qed"exis_elim";
-
-
-(* Sets *)
-
-val set_lemmas =
-   map (fn s => prove_goal Main.thy s (fn _ => [Fast_tac 1]))
-        ["f(x) : (UN x. {f(x)})",
-         "f x y : (UN x y. {f x y})",
-         "!!a. (!x. a ~= f(x)) ==> a ~: (UN x. {f(x)})",
-         "!!a. (!x y. a ~= f x y) ==> a ~: (UN x y. {f x y})"];
-
-(* 2 Lemmas to add to set_lemmas ... , used also for action handling, 
-   namely for Intersections and the empty list (compatibility of IOA!)  *)
-Goal "(UN b.{x. x=f(b)})= (UN b.{f(b)})"; 
- by (rtac set_ext 1);
- by (Fast_tac 1);
-qed "singleton_set";
-
-Goal "((A|B)=False) = ((~A)&(~B))";
- by (Fast_tac 1);
-qed "de_morgan";
-
-(* Lists *)
-
-Goal "hd(l@m) = (if l~=[] then hd(l) else hd(m))";
-by (induct_tac "l" 1);
-by (Simp_tac 1);
-by (Simp_tac 1);
-qed "hd_append";
-
-Goal "l ~= [] --> (? x xs. l = (x#xs))";
- by (induct_tac "l" 1);
- by (Simp_tac 1);
- by (Fast_tac 1);
-qed"cons_not_nil";
--- a/src/HOLCF/IOA/ABP/Lemmas.thy	Sat May 27 19:49:07 2006 +0200
+++ b/src/HOLCF/IOA/ABP/Lemmas.thy	Sat May 27 19:49:36 2006 +0200
@@ -7,4 +7,42 @@
 imports Main
 begin
 
+subsection {* Logic *}
+
+lemma and_de_morgan_and_absorbe: "(~(A&B)) = ((~A)&B| ~B)"
+  by blast
+
+lemma bool_if_impl_or: "(if C then A else B) --> (A|B)"
+  by auto
+
+lemma exis_elim: "(? x. x=P & Q(x)) = Q(P)"
+  by blast
+
+
+subsection {* Sets *}
+
+lemma set_lemmas:
+    "f(x) : (UN x. {f(x)})"
+    "f x y : (UN x y. {f x y})"
+    "!!a. (!x. a ~= f(x)) ==> a ~: (UN x. {f(x)})"
+    "!!a. (!x y. a ~= f x y) ==> a ~: (UN x y. {f x y})"
+  by auto
+
+text {* 2 Lemmas to add to @{text "set_lemmas"}, used also for action handling, 
+   namely for Intersections and the empty list (compatibility of IOA!). *}
+lemma singleton_set: "(UN b.{x. x=f(b)})= (UN b.{f(b)})"
+  by blast
+
+lemma de_morgan: "((A|B)=False) = ((~A)&(~B))"
+  by blast
+
+
+subsection {* Lists *}
+
+lemma hd_append: "hd(l@m) = (if l~=[] then hd(l) else hd(m))"
+  by (induct l) simp_all
+
+lemma cons_not_nil: "l ~= [] --> (? x xs. l = (x#xs))"
+  by (induct l) simp_all
+
 end
--- a/src/HOLCF/IOA/ABP/Packet.thy	Sat May 27 19:49:07 2006 +0200
+++ b/src/HOLCF/IOA/ABP/Packet.thy	Sat May 27 19:49:36 2006 +0200
@@ -19,6 +19,4 @@
   msg :: "'msg packet => 'msg"
   "msg == snd"
 
-ML {* use_legacy_bindings (the_context ()) *}
-
 end
--- a/src/HOLCF/IOA/ABP/Receiver.thy	Sat May 27 19:49:07 2006 +0200
+++ b/src/HOLCF/IOA/ABP/Receiver.thy	Sat May 27 19:49:36 2006 +0200
@@ -51,6 +51,4 @@
 "receiver_ioa ==
  (receiver_asig, {([],False)}, receiver_trans,{},{})"
 
-ML {* use_legacy_bindings (the_context ()) *}
-
 end
--- a/src/HOLCF/IOA/ABP/Sender.thy	Sat May 27 19:49:07 2006 +0200
+++ b/src/HOLCF/IOA/ABP/Sender.thy	Sat May 27 19:49:36 2006 +0200
@@ -49,6 +49,4 @@
 "sender_ioa ==
  (sender_asig, {([],True)}, sender_trans,{},{})"
 
-ML {* use_legacy_bindings (the_context ()) *}
-
 end
--- a/src/HOLCF/IsaMakefile	Sat May 27 19:49:07 2006 +0200
+++ b/src/HOLCF/IsaMakefile	Sat May 27 19:49:36 2006 +0200
@@ -105,9 +105,9 @@
 
 $(LOG)/IOA-ABP.gz: $(OUT)/IOA IOA/ABP/Abschannel.thy \
   IOA/ABP/Abschannel_finite.thy IOA/ABP/Action.thy \
-  IOA/ABP/Check.ML IOA/ABP/Correctness.ML IOA/ABP/Correctness.thy \
+  IOA/ABP/Check.ML IOA/ABP/Correctness.thy \
   IOA/ABP/Env.thy IOA/ABP/Impl.thy IOA/ABP/Impl_finite.thy \
-  IOA/ABP/Lemmas.ML IOA/ABP/Lemmas.thy IOA/ABP/Packet.thy \
+  IOA/ABP/Lemmas.thy IOA/ABP/Packet.thy \
   IOA/ABP/ROOT.ML IOA/ABP/Receiver.thy IOA/ABP/Sender.thy \
   IOA/ABP/Spec.thy
 	@cd IOA; $(ISATOOL) usedir $(OUT)/IOA ABP