--- a/src/HOL/IOA/IOA.ML	Mon Jun 22 17:13:09 1998 +0200
+++ b/src/HOL/IOA/IOA.ML	Mon Jun 22 17:26:46 1998 +0200
@@ -14,12 +14,12 @@
 
 val exec_rws = [executions_def,is_execution_fragment_def];
 
-goal IOA.thy
+Goal
 "asig_of((x,y,z)) = x & starts_of((x,y,z)) = y & trans_of((x,y,z)) = z";
   by (simp_tac (simpset() addsimps ioa_projections) 1);
   qed "ioa_triple_proj";
 
-goalw IOA.thy [ioa_def,state_trans_def,actions_def, is_asig_def]
+Goalw [ioa_def,state_trans_def,actions_def, is_asig_def]
   "!!A. [| IOA(A); (s1,a,s2):trans_of(A) |] ==> a:actions(asig_of(A))";
   by (REPEAT(etac conjE 1));
   by (EVERY1[etac allE, etac impE, atac]);
@@ -27,7 +27,7 @@
 qed "trans_in_actions";
 
 
-goal IOA.thy "filter_oseq p (filter_oseq p s) = filter_oseq p s";
+Goal "filter_oseq p (filter_oseq p s) = filter_oseq p s";
   by (simp_tac (simpset() addsimps [filter_oseq_def]) 1);
   by (rtac ext 1);
   by (exhaust_tac "s(i)" 1);
@@ -35,7 +35,7 @@
   by (Asm_simp_tac 1);
 qed "filter_oseq_idemp";
 
-goalw IOA.thy [mk_trace_def,filter_oseq_def]
+Goalw [mk_trace_def,filter_oseq_def]
 "(mk_trace A s n = None) =                                        \
 \  (s(n)=None | (? a. s(n)=Some(a) & a ~: externals(asig_of(A)))) \
 \  &                                                              \
@@ -46,13 +46,13 @@
   by (Fast_tac 1);
 qed "mk_trace_thm";
 
-goalw IOA.thy [reachable_def] "!!A. s:starts_of(A) ==> reachable A s";
+Goalw [reachable_def] "!!A. s:starts_of(A) ==> reachable A s";
   by (res_inst_tac [("x","(%i. None,%i. s)")] bexI 1);
   by (Simp_tac 1);
   by (asm_simp_tac (simpset() addsimps exec_rws) 1);
 qed "reachable_0";
 
-goalw IOA.thy (reachable_def::exec_rws)
+Goalw (reachable_def::exec_rws)
 "!!A. [| reachable A s; (s,a,t) : trans_of(A) |] ==> reachable A t";
   by (asm_full_simp_tac (simpset() delsimps bex_simps) 1);
   by (split_all_tac 1);
@@ -104,26 +104,26 @@
    by (rtac (p2 RS (p1 RS spec RS mp)) 1);
 qed "invariantE";
 
-goal IOA.thy 
+Goal 
 "actions(asig_comp a b) = actions(a) Un actions(b)";
   by (simp_tac (simpset() addsimps
                ([actions_def,asig_comp_def]@asig_projections)) 1);
   by (Fast_tac 1);
 qed "actions_asig_comp";
 
-goal IOA.thy
+Goal
 "starts_of(A || B) = {p. fst(p):starts_of(A) & snd(p):starts_of(B)}";
   by (simp_tac (simpset() addsimps (par_def::ioa_projections)) 1);
 qed "starts_of_par";
 
 (* Every state in an execution is reachable *)
-goalw IOA.thy [reachable_def] 
+Goalw [reachable_def] 
 "!!A. ex:executions(A) ==> !n. reachable A (snd ex n)";
   by (Fast_tac 1);
 qed "states_of_exec_reachable";
 
 
-goal IOA.thy 
+Goal 
 "(s,a,t) : trans_of(A || B || C || D) =                                      \
 \ ((a:actions(asig_of(A)) | a:actions(asig_of(B)) | a:actions(asig_of(C)) |  \
 \   a:actions(asig_of(D))) &                                                 \
@@ -141,32 +141,32 @@
                             ioa_projections)) 1);
 qed "trans_of_par4";
 
-goal IOA.thy "starts_of(restrict ioa acts) = starts_of(ioa) &     \
+Goal "starts_of(restrict ioa acts) = starts_of(ioa) &     \
 \             trans_of(restrict ioa acts) = trans_of(ioa) &       \
 \             reachable (restrict ioa acts) s = reachable ioa s";
 by (simp_tac (simpset() addsimps ([is_execution_fragment_def,executions_def,
                            reachable_def,restrict_def]@ioa_projections)) 1);
 qed "cancel_restrict";
 
-goal IOA.thy "asig_of(A || B) = asig_comp (asig_of A) (asig_of B)";
+Goal "asig_of(A || B) = asig_comp (asig_of A) (asig_of B)";
   by (simp_tac (simpset() addsimps (par_def::ioa_projections)) 1);
 qed "asig_of_par";
 
 
-goal IOA.thy "externals(asig_of(A1||A2)) =    \
+Goal "externals(asig_of(A1||A2)) =    \
 \  (externals(asig_of(A1)) Un externals(asig_of(A2)))";
 by (asm_full_simp_tac (simpset() addsimps [externals_def,asig_of_par,asig_comp_def,asig_inputs_def,asig_outputs_def,Un_def,set_diff_def]) 1);
 by (rtac set_ext 1); 
 by (Fast_tac 1);
 qed"externals_of_par"; 
 
-goalw IOA.thy [externals_def,actions_def,compat_ioas_def,compat_asigs_def]
+Goalw [externals_def,actions_def,compat_ioas_def,compat_asigs_def]
  "!! a. [| compat_ioas A1 A2; a:externals(asig_of(A1))|] ==> a~:internals(asig_of(A2))";
 by (Asm_full_simp_tac 1);
 by (best_tac (claset() addEs [equalityCE]) 1);
 qed"ext1_is_not_int2";
 
-goalw IOA.thy [externals_def,actions_def,compat_ioas_def,compat_asigs_def]
+Goalw [externals_def,actions_def,compat_ioas_def,compat_asigs_def]
  "!! a. [| compat_ioas A2 A1 ; a:externals(asig_of(A1))|] ==> a~:internals(asig_of(A2))";
 by (Asm_full_simp_tac 1);
 by (best_tac (claset() addEs [equalityCE]) 1);