isatool fixgoal;
authorwenzelm
Mon, 22 Jun 1998 17:13:09 +0200
changeset 5068 fb28eaa07e01
parent 5067 62b6288e6005
child 5069 3ea049f7979d
isatool fixgoal;
src/HOLCF/Discrete.ML
src/HOLCF/Discrete0.ML
src/HOLCF/Discrete1.ML
src/HOLCF/Fix.ML
src/HOLCF/IMP/Denotational.ML
src/HOLCF/IOA/ABP/Action.ML
src/HOLCF/IOA/ABP/Correctness.ML
src/HOLCF/IOA/NTP/Abschannel.ML
src/HOLCF/IOA/NTP/Action.ML
src/HOLCF/IOA/NTP/Correctness.ML
src/HOLCF/IOA/NTP/Impl.ML
src/HOLCF/IOA/NTP/Multiset.ML
src/HOLCF/IOA/NTP/Packet.ML
src/HOLCF/IOA/NTP/Receiver.ML
src/HOLCF/IOA/NTP/Sender.ML
src/HOLCF/IOA/meta_theory/Abstraction.ML
src/HOLCF/IOA/meta_theory/Asig.ML
src/HOLCF/IOA/meta_theory/Automata.ML
src/HOLCF/IOA/meta_theory/CompoExecs.ML
src/HOLCF/IOA/meta_theory/CompoScheds.ML
src/HOLCF/IOA/meta_theory/CompoTraces.ML
src/HOLCF/IOA/meta_theory/Compositionality.ML
src/HOLCF/IOA/meta_theory/Deadlock.ML
src/HOLCF/IOA/meta_theory/LiveIOA.ML
src/HOLCF/IOA/meta_theory/RefCorrectness.ML
src/HOLCF/IOA/meta_theory/RefMappings.ML
src/HOLCF/IOA/meta_theory/Seq.ML
src/HOLCF/IOA/meta_theory/Sequence.ML
src/HOLCF/IOA/meta_theory/ShortExecutions.ML
src/HOLCF/IOA/meta_theory/SimCorrectness.ML
src/HOLCF/IOA/meta_theory/Simulations.ML
src/HOLCF/IOA/meta_theory/TL.ML
src/HOLCF/IOA/meta_theory/TLS.ML
src/HOLCF/IOA/meta_theory/Traces.ML
src/HOLCF/IOA/meta_theory/TrivEx.ML
src/HOLCF/IOA/meta_theory/TrivEx2.ML
src/HOLCF/Lift.ML
src/HOLCF/Lift1.ML
src/HOLCF/Lift2.ML
src/HOLCF/Lift3.ML
src/HOLCF/Porder.ML
src/HOLCF/Porder0.ML
src/HOLCF/Tr.ML
src/ZF/AC/AC15_WO6.ML
src/ZF/AC/AC16_WO4.ML
src/ZF/AC/AC16_lemmas.ML
src/ZF/AC/AC17_AC1.ML
src/ZF/AC/AC18_AC19.ML
src/ZF/AC/AC1_AC17.ML
src/ZF/AC/AC1_WO2.ML
src/ZF/AC/AC2_AC6.ML
src/ZF/AC/AC7_AC9.ML
src/ZF/AC/AC_Equiv.ML
src/ZF/AC/Cardinal_aux.ML
src/ZF/AC/DC.ML
src/ZF/AC/DC_lemmas.ML
src/ZF/AC/HH.ML
src/ZF/AC/Hartog.ML
src/ZF/AC/WO1_AC.ML
src/ZF/AC/WO1_WO6.ML
src/ZF/AC/WO1_WO7.ML
src/ZF/AC/WO1_WO8.ML
src/ZF/AC/WO2_AC16.ML
src/ZF/AC/WO6_WO1.ML
src/ZF/AC/WO_AC.ML
src/ZF/AC/recfunAC16.ML
src/ZF/Coind/ECR.ML
src/ZF/Coind/MT.ML
src/ZF/Coind/Map.ML
src/ZF/Coind/Types.ML
src/ZF/Coind/Values.ML
src/ZF/IMP/Equiv.ML
src/ZF/Resid/Confluence.ML
src/ZF/Resid/Conversion.ML
src/ZF/Resid/Cube.ML
src/ZF/Resid/Redex.ML
src/ZF/Resid/Reduction.ML
src/ZF/Resid/Residuals.ML
src/ZF/Resid/SubUnion.ML
src/ZF/Resid/Substitution.ML
src/ZF/Resid/Terms.ML
src/ZF/ex/Acc.ML
src/ZF/ex/BT.ML
src/ZF/ex/Bin.ML
src/ZF/ex/Brouwer.ML
src/ZF/ex/CoUnit.ML
src/ZF/ex/Comb.ML
src/ZF/ex/Data.ML
src/ZF/ex/Enum.ML
src/ZF/ex/Integ.ML
src/ZF/ex/LList.ML
src/ZF/ex/Limit.ML
src/ZF/ex/ListN.ML
src/ZF/ex/Mutil.ML
src/ZF/ex/Ntree.ML
src/ZF/ex/Primes.ML
src/ZF/ex/Primrec.ML
src/ZF/ex/PropLog.ML
src/ZF/ex/Ramsey.ML
src/ZF/ex/Rmap.ML
src/ZF/ex/TF.ML
src/ZF/ex/Term.ML
src/ZF/ex/misc.ML
--- a/src/HOLCF/Discrete.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/Discrete.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -4,17 +4,17 @@
     Copyright   1997 TUM
 *)
 
-goalw thy [undiscr_def] "undiscr(Discr x) = x";
+Goalw [undiscr_def] "undiscr(Discr x) = x";
 by (Simp_tac 1);
 qed "undiscr_Discr";
 Addsimps [undiscr_Discr];
 
-goal thy
+Goal
  "!!S::nat=>('a::term)discr. chain(S) ==> range(%i. f(S i)) = {f(S 0)}";
 by (fast_tac (claset() addDs [discr_chain0] addEs [arg_cong]) 1);
 qed "discr_chain_f_range0";
 
-goalw thy [cont,is_lub,is_ub] "cont(%x::('a::term)discr. f x)";
+Goalw [cont,is_lub,is_ub] "cont(%x::('a::term)discr. f x)";
 by (simp_tac (simpset() addsimps [discr_chain_f_range0]) 1);
 qed "cont_discr";
 AddIffs [cont_discr];
--- a/src/HOLCF/Discrete0.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/Discrete0.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -6,17 +6,17 @@
 Proves that 'a discr is a po
 *)
 
-goalw thy [less_discr_def] "(x::('a::term)discr) << x";
+Goalw [less_discr_def] "(x::('a::term)discr) << x";
 by (rtac refl 1);
 qed "less_discr_refl";
 
-goalw thy [less_discr_def]
+Goalw [less_discr_def]
   "!!x. [| (x::('a::term)discr) << y; y << z |] ==> x <<  z";
 by (etac trans 1);
 by (assume_tac 1);
 qed "less_discr_trans";
 
-goalw thy [less_discr_def]
+Goalw [less_discr_def]
   "!!x. [| (x::('a::term)discr) << y; y << x |] ==> x=y";
 by (assume_tac 1);
 qed "less_discr_antisym";
--- a/src/HOLCF/Discrete1.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/Discrete1.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -6,12 +6,12 @@
 Proves that 'a discr is a cpo
 *)
 
-goalw thy [less_discr_def] "((x::('a::term)discr) << y) = (x=y)";
+Goalw [less_discr_def] "((x::('a::term)discr) << y) = (x=y)";
 by (rtac refl 1);
 qed "discr_less_eq";
 AddIffs [discr_less_eq];
 
-goalw thy [chain]
+Goalw [chain]
  "!!S::nat=>('a::term)discr. chain S ==> S i = S 0";
 by (nat_ind_tac "i" 1);
 by (rtac refl 1);
@@ -20,13 +20,13 @@
 by (Fast_tac 1);
 qed "discr_chain0";
 
-goal thy
+Goal
  "!!S::nat=>('a::term)discr. chain(S) ==> range(S) = {S 0}";
 by (fast_tac (claset() addEs [discr_chain0]) 1);
 qed "discr_chain_range0";
 Addsimps [discr_chain_range0];
 
-goalw thy [is_lub,is_ub]
+Goalw [is_lub,is_ub]
  "!!S. chain S ==> ? x::('a::term)discr. range(S) <<| x";
 by (Asm_simp_tac 1);
 qed "discr_cpo";
--- a/src/HOLCF/Fix.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/Fix.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -884,7 +884,7 @@
         (Simp_tac 1)
         ]);
 
-goal Fix.thy "!! P. [| adm (%x. P x --> Q x); adm (%x. Q x --> P x) |] \
+Goal "!! P. [| adm (%x. P x --> Q x); adm (%x. Q x --> P x) |] \
 \           ==> adm (%x. P x = Q x)";
 by (subgoal_tac "(%x. P x = Q x) = (%x. (P x --> Q x) & (Q x --> P x))" 1);
 by (Asm_simp_tac 1);
--- a/src/HOLCF/IMP/Denotational.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/IMP/Denotational.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -6,29 +6,29 @@
 Equivalence of Denotational Semantics in HOLCF and Evaluation Semantics in HOL
 *)
 
-goalw thy [dlift_def] "dlift f`(Def x) = f`(Discr x)";
+Goalw [dlift_def] "dlift f`(Def x) = f`(Discr x)";
 by (Simp_tac 1);
 qed "dlift_Def";
 Addsimps [dlift_Def];
 
-goalw thy [dlift_def] "cont(%f. dlift f)";
+Goalw [dlift_def] "cont(%f. dlift f)";
 by (Simp_tac 1);
 qed "cont_dlift";
 AddIffs [cont_dlift];
 
-goalw thy [dlift_def]
+Goalw [dlift_def]
  "(dlift f`l = Def y) = (? x. l = Def x & f`(Discr x) = Def y)";
 by (simp_tac (simpset() addsplits [split_lift_case]) 1);
 qed "dlift_is_Def";
 Addsimps [dlift_is_Def];
 
-goal thy "!!c. <c,s> -c-> t ==> D c`(Discr s) = (Def t)";
+Goal "!!c. <c,s> -c-> t ==> D c`(Discr s) = (Def t)";
 by (etac evalc.induct 1);
       by (ALLGOALS Asm_simp_tac);
  by (ALLGOALS (stac fix_eq THEN' Asm_full_simp_tac));
 qed_spec_mp "eval_implies_D";
 
-goal thy "!s t. D c`(Discr s) = (Def t) --> <c,s> -c-> t";
+Goal "!s t. D c`(Discr s) = (Def t) --> <c,s> -c-> t";
 by (com.induct_tac "c" 1);
     by (fast_tac ((HOL_cs addSIs evalc.intrs) addss simpset()) 1);
    by (fast_tac ((HOL_cs addSIs evalc.intrs) addss simpset()) 1);
@@ -46,6 +46,6 @@
 qed_spec_mp "D_implies_eval";
 
 
-goal thy "(D c`(Discr s) = (Def t)) = (<c,s> -c-> t)";
+Goal "(D c`(Discr s) = (Def t)) = (<c,s> -c-> t)";
 by (fast_tac (HOL_cs addSEs [D_implies_eval,eval_implies_D]) 1);
 qed "D_is_eval";
--- a/src/HOLCF/IOA/ABP/Action.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/IOA/ABP/Action.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -6,7 +6,7 @@
 Derived rules for actions
 *)
 
-goal Action.thy "!!x. x = y ==> action_case a b c d e f g x =     \
+Goal "!!x. x = y ==> action_case a b c d e f g x =     \
 \                               action_case a b c d e f g y";
 by (Asm_simp_tac 1);
 
--- a/src/HOLCF/IOA/ABP/Correctness.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/IOA/ABP/Correctness.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -42,7 +42,7 @@
 
 (* lemmas about reduce *)
 
-goal Correctness.thy "(reduce(l)=[]) = (l=[])";  
+Goal "(reduce(l)=[]) = (l=[])";  
  by (rtac iffI 1);
  by (subgoal_tac "(l~=[]) --> (reduce(l)~=[])" 1);
  by (Fast_tac 1); 
@@ -59,7 +59,7 @@
  by (Asm_full_simp_tac 1);
 val l_iff_red_nil = result();
 
-goal Correctness.thy "s~=[] --> hd(s)=hd(reduce(s))";
+Goal "s~=[] --> hd(s)=hd(reduce(s))";
 by (List.list.induct_tac "s" 1);
 by (Simp_tac 1);
 by (case_tac "list =[]" 1);
@@ -79,7 +79,7 @@
 qed"hd_is_reduce_hd";
 
 (* to be used in the following Lemma *)
-goal Correctness.thy "l~=[] --> reverse(reduce(l))~=[]";
+Goal "l~=[] --> reverse(reduce(l))~=[]";
 by (List.list.induct_tac "l" 1);
 by (Simp_tac 1);
 by (case_tac "list =[]" 1);
@@ -97,7 +97,7 @@
 qed"rev_red_not_nil";
 
 (* shows applicability of the induction hypothesis of the following Lemma 1 *)
-goal Correctness.thy "!!l.[| l~=[] |] ==>   \
+Goal "!!l.[| l~=[] |] ==>   \
 \   hd(reverse(reduce(a#l))) = hd(reverse(reduce(l)))";
  by (Simp_tac 1);
  by (rtac (split_list_case RS iffD2) 1);
@@ -112,7 +112,7 @@
 val impl_ss = simpset() delsimps [reduce_Cons];
 
 (* Main Lemma 1 for S_pkt in showing that reduce is refinement  *) 
-goal Correctness.thy 
+Goal 
    "if x=hd(reverse(reduce(l))) & reduce(l)~=[] then   \
 \      reduce(l@[x])=reduce(l) else                  \
 \      reduce(l@[x])=reduce(l)@[x]"; 
@@ -145,7 +145,7 @@
 
 
 (* Main Lemma 2 for R_pkt in showing that reduce is refinement *)
-goal Correctness.thy 
+Goal 
   "!! s. [| s~=[] |] ==>  \
 \    if hd(s)=hd(tl(s)) & tl(s)~=[] then    \
 \      reduce(tl(s))=reduce(s) else      \
@@ -168,7 +168,7 @@
  *******************************************************************)
 
 Delsplits [split_if];
-goal Correctness.thy 
+Goal 
       "is_weak_ref_map reduce ch_ioa ch_fin_ioa";
 by (simp_tac (simpset() addsimps [is_weak_ref_map_def]) 1);
 (* ---------------- main-part ------------------- *)
@@ -194,13 +194,13 @@
 by (etac reduce_tl 1);
 qed"channel_abstraction";
 
-goal Correctness.thy 
+Goal 
       "is_weak_ref_map reduce srch_ioa srch_fin_ioa";
 by (simp_tac (list_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 Correctness.thy 
+Goal 
       "is_weak_ref_map reduce rsch_ioa rsch_fin_ioa";
 by (simp_tac (list_ss addsimps [srch_fin_ioa_def,rsch_fin_ioa_def,
  srch_ioa_def,rsch_ioa_def,rename_through_pmap,channel_abstraction]) 1);
@@ -209,7 +209,7 @@
 
 (* 3 thms that do not hold generally! The lucky restriction here is 
    the absence of internal actions. *)
-goal Correctness.thy 
+Goal 
       "is_weak_ref_map (%id. id) sender_ioa sender_ioa";
 by (simp_tac (simpset() addsimps [is_weak_ref_map_def]) 1);
 by (TRY(
@@ -225,7 +225,7 @@
 qed"sender_unchanged";
 
 (* 2 copies of before *)
-goal Correctness.thy 
+Goal 
       "is_weak_ref_map (%id. id) receiver_ioa receiver_ioa";
 by (simp_tac (simpset() addsimps [is_weak_ref_map_def]) 1);
 by (TRY(
@@ -240,7 +240,7 @@
 by (ALLGOALS(simp_tac (simpset() addsimps [externals_def] addsplits [split_if])));
 qed"receiver_unchanged";
 
-goal Correctness.thy 
+Goal 
       "is_weak_ref_map (%id. id) env_ioa env_ioa";
 by (simp_tac (simpset() addsimps [is_weak_ref_map_def]) 1);
 by (TRY(
@@ -256,7 +256,7 @@
 qed"env_unchanged";
 Addsplits [split_if];
 
-goal Correctness.thy "compatible srch_ioa rsch_ioa"; 
+Goal "compatible srch_ioa rsch_ioa"; 
 by (simp_tac (simpset() addsimps [compatible_def,Int_def,empty_def]) 1);
 by (rtac set_ext 1);
 by (Action.action.induct_tac "x" 1);
@@ -264,7 +264,7 @@
 val compat_single_ch = result();
 
 (* totally the same as before *)
-goal Correctness.thy "compatible srch_fin_ioa rsch_fin_ioa"; 
+Goal "compatible srch_fin_ioa rsch_fin_ioa"; 
 by (simp_tac (simpset() addsimps [compatible_def,Int_def,empty_def]) 1);
 by (rtac set_ext 1);
 by (Action.action.induct_tac "x" 1);
@@ -278,7 +278,7 @@
                       rsch_ioa_def, Sender.sender_trans_def,
                       Receiver.receiver_trans_def] @ set_lemmas);
 
-goal Correctness.thy "compatible receiver_ioa (srch_ioa || rsch_ioa)";
+Goal "compatible receiver_ioa (srch_ioa || rsch_ioa)";
 by (simp_tac (ss addsimps [empty_def,compatible_def,
                            asig_of_par,asig_comp_def,actions_def,
                            Int_def]) 1);
@@ -289,7 +289,7 @@
 val compat_rec = result();
 
 (* 5 proofs totally the same as before *)
-goal Correctness.thy "compatible receiver_ioa (srch_fin_ioa || rsch_fin_ioa)";
+Goal "compatible receiver_ioa (srch_fin_ioa || rsch_fin_ioa)";
 by (simp_tac (ss addsimps [empty_def,compatible_def,asig_of_par,asig_comp_def,actions_def,Int_def]) 1);
 by (Simp_tac 1);
 by (rtac set_ext 1);
@@ -297,7 +297,7 @@
 by (ALLGOALS Simp_tac);
 val compat_rec_fin =result();
 
-goal Correctness.thy "compatible sender_ioa \
+Goal "compatible sender_ioa \
 \      (receiver_ioa || srch_ioa || rsch_ioa)";
 by (simp_tac (ss addsimps [empty_def,compatible_def,asig_of_par,asig_comp_def,actions_def,Int_def]) 1);
 by (Simp_tac 1);
@@ -306,7 +306,7 @@
 by (ALLGOALS(Simp_tac));
 val compat_sen=result();
 
-goal Correctness.thy "compatible sender_ioa\
+Goal "compatible sender_ioa\
 \      (receiver_ioa || srch_fin_ioa || rsch_fin_ioa)";
 by (simp_tac (ss addsimps [empty_def,  compatible_def,asig_of_par,asig_comp_def,actions_def,Int_def]) 1);
 by (Simp_tac 1);
@@ -315,7 +315,7 @@
 by (ALLGOALS(Simp_tac));
 val compat_sen_fin =result();
 
-goal Correctness.thy "compatible env_ioa\
+Goal "compatible env_ioa\
 \      (sender_ioa || receiver_ioa || srch_ioa || rsch_ioa)";
 by (simp_tac (ss addsimps [empty_def,compatible_def,asig_of_par,asig_comp_def,actions_def,Int_def]) 1);
 by (Simp_tac 1);
@@ -324,7 +324,7 @@
 by (ALLGOALS(Simp_tac));
 val compat_env=result();
 
-goal Correctness.thy "compatible env_ioa\
+Goal "compatible env_ioa\
 \      (sender_ioa || receiver_ioa || srch_fin_ioa || rsch_fin_ioa)";
 by (simp_tac (ss addsimps [empty_def,compatible_def,asig_of_par,asig_comp_def,actions_def,Int_def]) 1);
 by (Simp_tac 1);
@@ -335,7 +335,7 @@
 
 
 (* lemmata about externals of channels *)
-goal Correctness.thy 
+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);
@@ -363,7 +363,7 @@
 (* 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 Correctness.thy 
+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,
--- a/src/HOLCF/IOA/NTP/Abschannel.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/IOA/NTP/Abschannel.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -14,7 +14,7 @@
    actions_def, srch_trans_def, rsch_trans_def, ch_trans_def,starts_of_def, 
    trans_of_def] @ asig_projections;
 
-goal Abschannel.thy
+Goal
      "S_msg(m) ~: actions(srch_asig)        &    \
      \ R_msg(m) ~: actions(srch_asig)        &    \
      \ S_pkt(pkt) : actions(srch_asig)    &    \
@@ -28,7 +28,7 @@
 by (simp_tac (simpset() addsimps unfold_renaming) 1);
 qed"in_srch_asig";
 
-goal Abschannel.thy
+Goal
       "S_msg(m) ~: actions(rsch_asig)         & \
      \ R_msg(m) ~: actions(rsch_asig)         & \
      \ S_pkt(pkt) ~: actions(rsch_asig)    & \
@@ -43,14 +43,14 @@
 by (simp_tac (simpset() addsimps unfold_renaming) 1);
 qed"in_rsch_asig";
 
-goal Abschannel.thy "srch_ioa = \
+Goal "srch_ioa = \
 \   (srch_asig, {{|}}, srch_trans,srch_wfair,srch_sfair)";
 by (simp_tac (simpset() addsimps [srch_asig_def,srch_trans_def, asig_of_def, trans_of_def,
               wfair_of_def, sfair_of_def,srch_wfair_def,srch_sfair_def]) 1);
 by (simp_tac (simpset() addsimps unfold_renaming) 1);
 qed"srch_ioa_thm";
 
-goal Abschannel.thy "rsch_ioa = \
+Goal "rsch_ioa = \
 \    (rsch_asig, {{|}}, rsch_trans,rsch_wfair,rsch_sfair)";
 by (simp_tac (simpset() addsimps [rsch_asig_def,rsch_trans_def, asig_of_def, trans_of_def,
               wfair_of_def, sfair_of_def,rsch_wfair_def,rsch_sfair_def]) 1);
--- a/src/HOLCF/IOA/NTP/Action.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/IOA/NTP/Action.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -6,7 +6,7 @@
 Derived rules for actions
 *)
 
-goal Action.thy "!!x. x = y ==> action_case a b c d e f g h i j x =     \
+Goal "!!x. x = y ==> action_case a b c d e f g h i j x =     \
 \                               action_case a b c d e f g h i j y";
 by (Asm_simp_tac 1);
 
--- a/src/HOLCF/IOA/NTP/Correctness.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/IOA/NTP/Correctness.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -27,7 +27,7 @@
 (* A lemma about restricting the action signature of the implementation
  * to that of the specification.
  ****************************)
-goal Correctness.thy
+Goal
  "a:externals(asig_of(restrict impl_ioa (externals spec_sig))) = \
 \ (case a of                  \
 \     S_msg(m) => True        \
@@ -66,7 +66,7 @@
 
 
 (* Proof of correctness *)
-goalw Correctness.thy [Spec.ioa_def, is_weak_ref_map_def]
+Goalw [Spec.ioa_def, is_weak_ref_map_def]
   "is_weak_ref_map hom (restrict impl_ioa (externals spec_sig)) spec_ioa";
 by (simp_tac (simpset() delsplits [split_if] addsimps 
     [Correctness.hom_def, cancel_restrict, externals_lemma]) 1);
--- a/src/HOLCF/IOA/NTP/Impl.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/IOA/NTP/Impl.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -28,7 +28,7 @@
           in_rsch_asig];
 Addcongs [let_weak_cong];
 
-goal Impl.thy
+Goal
   "fst(x) = sen(x)            & \
 \  fst(snd(x)) = rec(x)       & \
 \  fst(snd(snd(x))) = srch(x) & \
@@ -37,7 +37,7 @@
              [sen_def,rec_def,srch_def,rsch_def]) 1);
 Addsimps [result()];
 
-goal Impl.thy "a:actions(sender_asig)   \
+Goal "a:actions(sender_asig)   \
 \            | a:actions(receiver_asig) \
 \            | a:actions(srch_asig)     \
 \            | a:actions(rsch_asig)";
@@ -67,7 +67,7 @@
 
 (* INVARIANT 1 *)
 
-goalw Impl.thy impl_ioas "invariant impl_ioa inv1";
+Goalw impl_ioas "invariant impl_ioa inv1";
 by (rtac invariantI 1);
 by (asm_full_simp_tac (simpset()
    addsimps [inv1_def, hdr_sum_def, Sender.srcvd_def,
@@ -150,7 +150,7 @@
 
 (* INVARIANT 2 *)
 
-  goal Impl.thy "invariant impl_ioa inv2";
+Goal "invariant impl_ioa inv2";
 
   by (rtac invariantI1 1); 
   (* Base case *)
@@ -212,7 +212,7 @@
 
 (* INVARIANT 3 *)
 
-goal Impl.thy "invariant impl_ioa inv3";
+Goal "invariant impl_ioa inv3";
 
   by (rtac invariantI 1); 
   (* Base case *)
@@ -278,7 +278,7 @@
 
 (* INVARIANT 4 *)
 
-goal Impl.thy "invariant impl_ioa inv4";
+Goal "invariant impl_ioa inv4";
 
   by (rtac invariantI 1); 
   (* Base case *)
--- a/src/HOLCF/IOA/NTP/Multiset.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/IOA/NTP/Multiset.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -7,22 +7,22 @@
 Should be done as a subtype and moved to a global place.
 *)
 
-goalw Multiset.thy [Multiset.count_def, Multiset.countm_empty_def]
+Goalw [Multiset.count_def, Multiset.countm_empty_def]
    "count {|} x = 0";
  by (rtac refl 1);
 qed "count_empty";
 
-goal Multiset.thy 
+Goal 
      "count (addm M x) y = (if y=x then Suc(count M y) else count M y)";
   by (asm_simp_tac (simpset() addsimps 
                     [Multiset.count_def,Multiset.countm_nonempty_def]) 1);
 qed "count_addm_simp";
 
-goal Multiset.thy "count M y <= count (addm M x) y";
+Goal "count M y <= count (addm M x) y";
   by (simp_tac (simpset() addsimps [count_addm_simp]) 1);
 qed "count_leq_addm";
 
-goalw Multiset.thy [Multiset.count_def] 
+Goalw [Multiset.count_def] 
      "count (delm M x) y = (if y=x then count M y - 1 else count M y)";
 by (res_inst_tac [("M","M")] Multiset.induction 1);
 by (asm_simp_tac (simpset() 
@@ -34,14 +34,14 @@
 by (Asm_full_simp_tac 1);
 qed "count_delm_simp";
 
-goal Multiset.thy "!!M. (!x. P(x) --> Q(x)) ==> (countm M P <= countm M Q)";
+Goal "!!M. (!x. P(x) --> Q(x)) ==> (countm M P <= countm M Q)";
 by (res_inst_tac [("M","M")] Multiset.induction 1);
  by (simp_tac (simpset() addsimps [Multiset.countm_empty_def]) 1);
 by (simp_tac (simpset() addsimps[Multiset.countm_nonempty_def]) 1);
 auto();
 qed "countm_props";
 
-goal Multiset.thy "!!P. ~P(obj) ==> countm M P = countm (delm M obj) P";
+Goal "!!P. ~P(obj) ==> countm M P = countm (delm M obj) P";
   by (res_inst_tac [("M","M")] Multiset.induction 1);
   by (simp_tac (simpset() addsimps [Multiset.delm_empty_def,
                                    Multiset.countm_empty_def]) 1);
@@ -50,7 +50,7 @@
 qed "countm_spurious_delm";
 
 
-goal Multiset.thy "!!P. P(x) ==> 0<count M x --> 0<countm M P";
+Goal "!!P. P(x) ==> 0<count M x --> 0<countm M P";
   by (res_inst_tac [("M","M")] Multiset.induction 1);
   by (simp_tac (simpset() addsimps 
                           [Multiset.delm_empty_def,Multiset.count_def,
@@ -60,7 +60,7 @@
                         Multiset.countm_nonempty_def]) 1);
 qed_spec_mp "pos_count_imp_pos_countm";
 
-goal Multiset.thy
+Goal
    "!!P. P(x) ==> 0<count M x --> countm (delm M x) P = countm M P - 1";
   by (res_inst_tac [("M","M")] Multiset.induction 1);
   by (simp_tac (simpset() addsimps 
--- a/src/HOLCF/IOA/NTP/Packet.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/IOA/NTP/Packet.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -8,7 +8,7 @@
 
 
 (* Instantiation of a tautology? *)
-goal Packet.thy "!x. x = packet --> hdr(x) = hdr(packet)";
+Goal "!x. x = packet --> hdr(x) = hdr(packet)";
  by (simp_tac (simpset() addsimps [Packet.hdr_def]) 1);
 qed "eq_packet_imp_eq_hdr"; 
 
--- a/src/HOLCF/IOA/NTP/Receiver.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/IOA/NTP/Receiver.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -4,7 +4,7 @@
     Copyright   1994  TU Muenchen
 *)
 
-goal Receiver.thy
+Goal
  "S_msg(m) ~: actions(receiver_asig)      &   \
 \ R_msg(m) : actions(receiver_asig)       &   \
 \ S_pkt(pkt) ~: actions(receiver_asig) &   \
--- a/src/HOLCF/IOA/NTP/Sender.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/IOA/NTP/Sender.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -4,7 +4,7 @@
     Copyright   1994  TU Muenchen
 *)
 
-goal Sender.thy
+Goal
  "S_msg(m) : actions(sender_asig)       &   \
 \ R_msg(m) ~: actions(sender_asig)      &   \
 \ S_pkt(pkt) : actions(sender_asig)  &   \
--- a/src/HOLCF/IOA/meta_theory/Abstraction.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/IOA/meta_theory/Abstraction.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -14,15 +14,15 @@
 (*                             cex_abs                              *)
 (* ---------------------------------------------------------------- *)
 
-goal thy "cex_abs f (s,UU) = (f s, UU)";
+Goal "cex_abs f (s,UU) = (f s, UU)";
 by (simp_tac (simpset() addsimps [cex_abs_def]) 1);
 qed"cex_abs_UU";
 
-goal thy "cex_abs f (s,nil) = (f s, nil)";
+Goal "cex_abs f (s,nil) = (f s, nil)";
 by (simp_tac (simpset() addsimps [cex_abs_def]) 1);
 qed"cex_abs_nil";
 
-goal thy "cex_abs f (s,(a,t)>>ex) = (f s, (a,f t) >> (snd (cex_abs f (t,ex))))";
+Goal "cex_abs f (s,(a,t)>>ex) = (f s, (a,f t) >> (snd (cex_abs f (t,ex))))";
 by (simp_tac (simpset() addsimps [cex_abs_def]) 1);
 qed"cex_abs_cons";
 
@@ -36,13 +36,13 @@
 (*                           Lemmas                                 *)
 (* ---------------------------------------------------------------- *)
 
-goal thy "temp_weakening Q P h = (! ex. (ex |== P) --> (cex_abs h ex |== Q))";
+Goal "temp_weakening Q P h = (! ex. (ex |== P) --> (cex_abs h ex |== Q))";
 by (simp_tac (simpset() addsimps [temp_weakening_def,temp_strengthening_def,
      NOT_def,temp_sat_def,satisfies_def]) 1);
 auto();
 qed"temp_weakening_def2";
 
-goal thy "state_weakening Q P h = (! s t a. P (s,a,t) --> Q (h(s),a,h(t)))";
+Goal "state_weakening Q P h = (! s t a. P (s,a,t) --> Q (h(s),a,h(t)))";
 by (simp_tac (simpset() addsimps [state_weakening_def,state_strengthening_def,
      NOT_def]) 1);
 auto();
@@ -56,7 +56,7 @@
 (* ---------------------------------------------------------------- *)
 
 
-goalw thy [cex_abs_def]
+Goalw [cex_abs_def]
  "!!h.[| is_abstraction h C A |] ==>\
 \ !s. reachable C s & is_exec_frag C (s,xs) \
 \ --> is_exec_frag A (cex_abs h (s,xs))"; 
@@ -72,7 +72,7 @@
 qed"exec_frag_abstraction";
 
 
-goal thy "!!A. is_abstraction h C A ==> weakeningIOA A C h";
+Goal "!!A. is_abstraction h C A ==> weakeningIOA A C h";
 by (asm_full_simp_tac (simpset() addsimps [weakeningIOA_def])1);
 auto();
 by (asm_full_simp_tac (simpset() addsimps [executions_def]) 1);
@@ -85,7 +85,7 @@
 qed"abs_is_weakening";
 
 
-goal thy "!!A. [|is_abstraction h C A; validIOA A Q; temp_strengthening Q P h |] \
+Goal "!!A. [|is_abstraction h C A; validIOA A Q; temp_strengthening Q P h |] \
 \         ==> validIOA C P";
 bd abs_is_weakening 1;
 by (asm_full_simp_tac (simpset() addsimps [weakeningIOA_def, 
@@ -97,26 +97,26 @@
 
 (* FIX: Nach TLS.ML *)
 
-goal thy "(ex |== P .--> Q) = ((ex |== P) --> (ex |== Q))";
+Goal "(ex |== P .--> Q) = ((ex |== P) --> (ex |== Q))";
 by (simp_tac (simpset() addsimps [IMPLIES_def,temp_sat_def, satisfies_def])1);
 qed"IMPLIES_temp_sat";
 
-goal thy "(ex |== P .& Q) = ((ex |== P) & (ex |== Q))";
+Goal "(ex |== P .& Q) = ((ex |== P) & (ex |== Q))";
 by (simp_tac (simpset() addsimps [AND_def,temp_sat_def, satisfies_def])1);
 qed"AND_temp_sat";
 
-goal thy "(ex |== P .| Q) = ((ex |== P) | (ex |== Q))";
+Goal "(ex |== P .| Q) = ((ex |== P) | (ex |== Q))";
 by (simp_tac (simpset() addsimps [OR_def,temp_sat_def, satisfies_def])1);
 qed"OR_temp_sat";
 
-goal thy "(ex |== .~ P) = (~ (ex |== P))";
+Goal "(ex |== .~ P) = (~ (ex |== P))";
 by (simp_tac (simpset() addsimps [NOT_def,temp_sat_def, satisfies_def])1);
 qed"NOT_temp_sat";
 
 Addsimps [IMPLIES_temp_sat,AND_temp_sat,OR_temp_sat,NOT_temp_sat];
 
 
-goalw thy [is_live_abstraction_def]
+Goalw [is_live_abstraction_def]
    "!!A. [|is_live_abstraction h (C,L) (A,M); \
 \         validLIOA (A,M) Q;  temp_strengthening Q P h |] \
 \         ==> validLIOA (C,L) P";
@@ -129,7 +129,7 @@
 qed"AbsRuleT2";
 
 
-goalw thy [is_live_abstraction_def]
+Goalw [is_live_abstraction_def]
    "!!A. [|is_live_abstraction h (C,L) (A,M); \
 \         validLIOA (A,M) (H1 .--> Q);  temp_strengthening Q P h; \
 \         temp_weakening H1 H2 h; validLIOA (C,L) H2 |] \
@@ -150,7 +150,7 @@
 (* ---------------------------------------------------------------- *)
 
 
-goalw thy [is_abstraction_def,is_ref_map_def] 
+Goalw [is_abstraction_def,is_ref_map_def] 
 "!! h. is_abstraction h C A ==> is_ref_map h C A";
 by (safe_tac set_cs);
 by (res_inst_tac[("x","(a,h t)>>nil")] exI 1);
@@ -158,7 +158,7 @@
 qed"abstraction_is_ref_map";
 
 
-goal thy "!! h. [| inp(C)=inp(A); out(C)=out(A); \
+Goal "!! h. [| inp(C)=inp(A); out(C)=out(A); \
 \                  is_abstraction h C A |] \
 \               ==> C =<| A";
 by (asm_full_simp_tac (simpset() addsimps [ioa_implements_def]) 1);
@@ -178,7 +178,7 @@
 
 (* Reduces to Filter (Map fst x) = Filter (Map fst (Map (%(a,t). (a,x)) x),
    that is to special Map Lemma *)
-goalw thy [cex_abs_def,mk_trace_def,filter_act_def]
+Goalw [cex_abs_def,mk_trace_def,filter_act_def]
   "!! f. ext C = ext A \
 \        ==> mk_trace C`xs = mk_trace A`(snd (cex_abs f (s,xs)))";
 by (Asm_full_simp_tac 1);
@@ -186,7 +186,7 @@
 qed"traces_coincide_abs";
 
 
-goalw thy [cex_abs_def]
+Goalw [cex_abs_def]
  "!!f.[| is_abstraction h C A |] ==>\
 \ !s. reachable C s & is_exec_frag C (s,xs) \
 \ --> is_exec_frag A (cex_abs h (s,xs))"; 
@@ -210,7 +210,7 @@
    is_live_abstraction includes temp_strengthening which is necessarily based
    on cex_abs and not on corresp_ex. Thus, the proof is redoone in a more specific
    way for cex_abs *)
-goal thy "!! h. [| inp(C)=inp(A); out(C)=out(A); \
+Goal "!! h. [| inp(C)=inp(A); out(C)=out(A); \
 \                  is_live_abstraction h (C,M) (A,L) |] \
 \               ==> live_implements (C,M) (A,L)";
 
@@ -242,7 +242,7 @@
 
 (* FIX: NAch Traces.ML bringen *)
 
-goalw thy [ioa_implements_def] 
+Goalw [ioa_implements_def] 
 "!! A. [| A =<| B; B =<| C|] ==> A =<| C"; 
 auto();
 qed"implements_trans";
@@ -254,7 +254,7 @@
 (*                Abstraction Rules for Automata                    *)
 (* ---------------------------------------------------------------- *)
 
-goal thy "!! C. [| inp(C)=inp(A); out(C)=out(A); \
+Goal "!! C. [| inp(C)=inp(A); out(C)=out(A); \
 \                  inp(Q)=inp(P); out(Q)=out(P); \
 \                  is_abstraction h1 C A; \
 \                  A =<| Q ; \
@@ -270,7 +270,7 @@
 qed"AbsRuleA1";
 
 
-goal thy "!! C. [| inp(C)=inp(A); out(C)=out(A); \
+Goal "!! C. [| inp(C)=inp(A); out(C)=out(A); \
 \                  inp(Q)=inp(P); out(Q)=out(P); \
 \                  is_live_abstraction h1 (C,LC) (A,LA); \
 \                  live_implements (A,LA) (Q,LQ) ; \
@@ -295,28 +295,28 @@
 (*                Localizing Temproal Strengthenings - 1               *)
 (* ---------------------------------------------------------------- *)
 
-goalw thy [temp_strengthening_def]
+Goalw [temp_strengthening_def]
 "!! h. [| temp_strengthening P1 Q1 h; \
 \         temp_strengthening P2 Q2 h |] \
 \      ==> temp_strengthening (P1 .& P2) (Q1 .& Q2) h";
 auto();
 qed"strength_AND";
 
-goalw thy [temp_strengthening_def]
+Goalw [temp_strengthening_def]
 "!! h. [| temp_strengthening P1 Q1 h; \
 \         temp_strengthening P2 Q2 h |] \
 \      ==> temp_strengthening (P1 .| P2) (Q1 .| Q2) h";
 auto();
 qed"strength_OR";
 
-goalw thy [temp_strengthening_def]
+Goalw [temp_strengthening_def]
 "!! h. [| temp_weakening P Q h |] \
 \      ==> temp_strengthening (.~ P) (.~ Q) h";
 by (asm_full_simp_tac (simpset() addsimps [temp_weakening_def2])1);
 auto();
 qed"strength_NOT";
 
-goalw thy [temp_strengthening_def]
+Goalw [temp_strengthening_def]
 "!! h. [| temp_weakening P1 Q1 h; \
 \         temp_strengthening P2 Q2 h |] \
 \      ==> temp_strengthening (P1 .--> P2) (Q1 .--> Q2) h";
@@ -329,28 +329,28 @@
 (*                Localizing Temproal Weakenings - Part 1           *)
 (* ---------------------------------------------------------------- *)
 
-goal thy
+Goal
 "!! h. [| temp_weakening P1 Q1 h; \
 \         temp_weakening P2 Q2 h |] \
 \      ==> temp_weakening (P1 .& P2) (Q1 .& Q2) h";
 by (asm_full_simp_tac (simpset() addsimps [temp_weakening_def2])1);
 qed"weak_AND";
 
-goal thy 
+Goal 
 "!! h. [| temp_weakening P1 Q1 h; \
 \         temp_weakening P2 Q2 h |] \
 \      ==> temp_weakening (P1 .| P2) (Q1 .| Q2) h";
 by (asm_full_simp_tac (simpset() addsimps [temp_weakening_def2])1);
 qed"weak_OR";
 
-goalw thy [temp_strengthening_def]
+Goalw [temp_strengthening_def]
 "!! h. [| temp_strengthening P Q h |] \
 \      ==> temp_weakening (.~ P) (.~ Q) h";
 by (asm_full_simp_tac (simpset() addsimps [temp_weakening_def2])1);
 auto();
 qed"weak_NOT";
 
-goalw thy [temp_strengthening_def]
+Goalw [temp_strengthening_def]
 "!! h. [| temp_strengthening P1 Q1 h; \
 \         temp_weakening P2 Q2 h |] \
 \      ==> temp_weakening (P1 .--> P2) (Q1 .--> Q2) h";
@@ -366,12 +366,12 @@
 (* ------------------ Box ----------------------------*)
 
 (* FIX: should be same as nil_is_Conc2 when all nils are turned to right side !! *)
-goal thy "(UU = x @@ y) = (((x::'a Seq)= UU) | (x=nil & y=UU))";
+Goal "(UU = x @@ y) = (((x::'a Seq)= UU) | (x=nil & y=UU))";
 by (Seq_case_simp_tac "x" 1);
 by Auto_tac;
 qed"UU_is_Conc";
 
-goal thy 
+Goal 
 "Finite s1 --> \
 \ (! ex. (s~=nil & s~=UU & ex2seq ex = s1 @@ s) --> (? ex'. s = ex2seq ex'))";
 by (rtac impI 1);
@@ -392,7 +392,7 @@
 
 (* important property of ex2seq: can be shiftet, as defined "pointwise" *)
 
-goalw thy [tsuffix_def,suffix_def]
+Goalw [tsuffix_def,suffix_def]
 "!!s. tsuffix s (ex2seq ex) ==> ? ex'. s = (ex2seq ex')";
 auto();
 bd ex2seqConc 1;
@@ -400,11 +400,11 @@
 qed"ex2seq_tsuffix";
 
 
-goal thy "(Map f`s = nil) = (s=nil)";
+Goal "(Map f`s = nil) = (s=nil)";
 by (Seq_case_simp_tac "s" 1);
 qed"Mapnil";
 
-goal thy "(Map f`s = UU) = (s=UU)";
+Goal "(Map f`s = UU) = (s=UU)";
 by (Seq_case_simp_tac "s" 1);
 qed"MapUU";
 
@@ -412,7 +412,7 @@
 (* important property of cex_absSeq: As it is a 1to1 correspondence, 
   properties carry over *)
 
-goalw thy [tsuffix_def,suffix_def,cex_absSeq_def]
+Goalw [tsuffix_def,suffix_def,cex_absSeq_def]
 "!! s. tsuffix s t ==> tsuffix (cex_absSeq h s) (cex_absSeq h t)";
 auto();
 by (asm_full_simp_tac (simpset() addsimps [Mapnil])1);
@@ -422,7 +422,7 @@
 qed"cex_absSeq_tsuffix";
 
 
-goalw thy [temp_strengthening_def,state_strengthening_def, temp_sat_def,
+Goalw [temp_strengthening_def,state_strengthening_def, temp_sat_def,
 satisfies_def,Box_def]
 "!! h. [| temp_strengthening P Q h |]\
 \      ==> temp_strengthening ([] P) ([] Q) h";
@@ -436,7 +436,7 @@
 
 (* ------------------ Init ----------------------------*)
 
-goalw thy [temp_strengthening_def,state_strengthening_def,
+Goalw [temp_strengthening_def,state_strengthening_def,
 temp_sat_def,satisfies_def,Init_def,unlift_def]
 "!! h. [| state_strengthening P Q h |]\
 \      ==> temp_strengthening (Init P) (Init Q) h";
@@ -449,7 +449,7 @@
 
 (* ------------------ Next ----------------------------*)
 
-goal thy 
+Goal 
 "(TL`(ex2seq (cex_abs h ex))=UU) = (TL`(ex2seq ex)=UU)";
 by (pair_tac "ex" 1);
 by (Seq_case_simp_tac "y" 1);
@@ -458,7 +458,7 @@
 by (pair_tac "a" 1);
 qed"TL_ex2seq_UU";
 
-goal thy 
+Goal 
 "(TL`(ex2seq (cex_abs h ex))=nil) = (TL`(ex2seq ex)=nil)";
 by (pair_tac "ex" 1);
 by (Seq_case_simp_tac "y" 1);
@@ -468,21 +468,21 @@
 qed"TL_ex2seq_nil";
 
 (* FIX: put to Sequence Lemmas *)
-goal thy "Map f`(TL`s) = TL`(Map f`s)";
+Goal "Map f`(TL`s) = TL`(Map f`s)";
 by (Seq_induct_tac "s" [] 1);
 qed"MapTL";
 
 (* important property of cex_absSeq: As it is a 1to1 correspondence, 
   properties carry over *)
 
-goalw thy [cex_absSeq_def]
+Goalw [cex_absSeq_def]
 "cex_absSeq h (TL`s) = (TL`(cex_absSeq h s))";
 by (simp_tac (simpset() addsimps [MapTL]) 1);
 qed"cex_absSeq_TL";
 
 (* important property of ex2seq: can be shiftet, as defined "pointwise" *)
 
-goal thy "!!ex. [| (snd ex)~=UU ; (snd ex)~=nil |] ==> (? ex'. TL`(ex2seq ex) = ex2seq ex')";
+Goal "!!ex. [| (snd ex)~=UU ; (snd ex)~=nil |] ==> (? ex'. TL`(ex2seq ex) = ex2seq ex')";
 by (pair_tac "ex" 1);
 by (Seq_case_simp_tac "y" 1);
 by (pair_tac "a" 1);
@@ -490,7 +490,7 @@
 qed"TLex2seq";
 
 
-goal thy "(TL`(ex2seq ex)~=UU) = ((snd ex)~=UU)";
+Goal "(TL`(ex2seq ex)~=UU) = ((snd ex)~=UU)";
 by (pair_tac "ex" 1);
 by (Seq_case_simp_tac "y" 1);
 by (pair_tac "a" 1);
@@ -498,7 +498,7 @@
 by (pair_tac "a" 1);
 qed"ex2seqUUTL";
 
-goal thy "(TL`(ex2seq ex)~=nil) = ((snd ex)~=nil)";
+Goal "(TL`(ex2seq ex)~=nil) = ((snd ex)~=nil)";
 by (pair_tac "ex" 1);
 by (Seq_case_simp_tac "y" 1);
 by (pair_tac "a" 1);
@@ -507,7 +507,7 @@
 qed"ex2seqnilTL";
 
 
-goalw thy [temp_strengthening_def,state_strengthening_def,
+Goalw [temp_strengthening_def,state_strengthening_def,
 temp_sat_def, satisfies_def,Next_def]
 "!! h. [| temp_strengthening P Q h |]\
 \      ==> temp_strengthening (Next P) (Next Q) h";
@@ -532,7 +532,7 @@
 (* ---------------------------------------------------------------- *)
 
 
-goal thy 
+Goal 
 "!! h. [| state_weakening P Q h |]\
 \      ==> temp_weakening (Init P) (Init Q) h";
 by (asm_full_simp_tac (simpset() addsimps [temp_weakening_def2,
@@ -548,7 +548,7 @@
 
 (* analog to strengthening thm above, with analog lemmas used  *)
 
-goalw thy [state_weakening_def]
+Goalw [state_weakening_def]
 "!! h. [| temp_weakening P Q h |]\
 \      ==> temp_weakening ([] P) ([] Q) h";
 by (asm_full_simp_tac (simpset() addsimps [temp_weakening_def2,
@@ -556,7 +556,7 @@
 
 (* analog to strengthening thm above, with analog lemmas used  *)
 
-goalw thy [state_weakening_def]
+Goalw [state_weakening_def]
 "!! h. [| temp_weakening P Q h |]\
 \      ==> temp_weakening (Next P) (Next Q) h";
 by (asm_full_simp_tac (simpset() addsimps [temp_weakening_def2,
@@ -569,7 +569,7 @@
 (* ---------------------------------------------------------------- *)
 
 
-goalw thy [Diamond_def]
+Goalw [Diamond_def]
 "!! h. [| temp_strengthening P Q h |]\
 \      ==> temp_strengthening (<> P) (<> Q) h";
 br strength_NOT 1;
@@ -577,7 +577,7 @@
 be weak_NOT 1;
 qed"strength_Diamond";
 
-goalw thy [Leadsto_def]
+Goalw [Leadsto_def]
 "!! h. [| temp_weakening P1 P2 h;\
 \         temp_strengthening Q1 Q2 h |]\
 \      ==> temp_strengthening (P1 ~> Q1) (P2 ~> Q2) h";
@@ -592,7 +592,7 @@
 (* ---------------------------------------------------------------- *)
 
 
-goalw thy [Diamond_def]
+Goalw [Diamond_def]
 "!! h. [| temp_weakening P Q h |]\
 \      ==> temp_weakening (<> P) (<> Q) h";
 br weak_NOT 1;
@@ -600,7 +600,7 @@
 be strength_NOT 1;
 qed"weak_Diamond";
 
-goalw thy [Leadsto_def]
+Goalw [Leadsto_def]
 "!! h. [| temp_strengthening P1 P2 h;\
 \         temp_weakening Q1 Q2 h |]\
 \      ==> temp_weakening (P1 ~> Q1) (P2 ~> Q2) h";
@@ -609,7 +609,7 @@
 be weak_Diamond 1;
 qed"weak_Leadsto";
 
-goalw thy [WF_def]
+Goalw [WF_def]
   " !!A. [| !! s. Enabled A acts (h s) ==> Enabled C acts s|] \ 
 \   ==> temp_weakening (WF A acts) (WF C acts) h";
 br weak_IMPLIES 1;
@@ -624,7 +624,7 @@
                              xt2_def,plift_def,option_lift_def,NOT_def]));
 qed"weak_WF";
 
-goalw thy [SF_def]
+Goalw [SF_def]
   " !!A. [| !! s. Enabled A acts (h s) ==> Enabled C acts s|] \ 
 \   ==> temp_weakening (SF A acts) (SF C acts) h";
 br weak_IMPLIES 1;
--- a/src/HOLCF/IOA/meta_theory/Asig.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/IOA/meta_theory/Asig.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -10,48 +10,48 @@
 
 val asig_projections = [asig_inputs_def, asig_outputs_def, asig_internals_def];
 
-goal thy
+Goal
  "(outputs    (a,b,c) = b)   & \
 \ (inputs     (a,b,c) = a) & \
 \ (internals  (a,b,c) = c)";
   by (simp_tac (simpset() addsimps asig_projections) 1);
 qed "asig_triple_proj";
 
-goal thy "!!a.[| a~:internals(S) ;a~:externals(S)|] ==> a~:actions(S)";
+Goal "!!a.[| a~:internals(S) ;a~:externals(S)|] ==> a~:actions(S)";
 by (asm_full_simp_tac (simpset() addsimps [externals_def,actions_def]) 1);
 qed"int_and_ext_is_act";
 
-goal thy "!!a.[|a:externals(S)|] ==> a:actions(S)";
+Goal "!!a.[|a:externals(S)|] ==> a:actions(S)";
 by (asm_full_simp_tac (simpset() addsimps [externals_def,actions_def]) 1);
 qed"ext_is_act";
 
-goal thy "!!a.[|a:internals S|] ==> a:actions S";
+Goal "!!a.[|a:internals S|] ==> a:actions S";
 by (asm_full_simp_tac (simpset() addsimps [asig_internals_def,actions_def]) 1);
 qed"int_is_act";
 
-goal thy "!!a.[|a:inputs S|] ==> a:actions S";
+Goal "!!a.[|a:inputs S|] ==> a:actions S";
 by (asm_full_simp_tac (simpset() addsimps [asig_inputs_def,actions_def]) 1);
 qed"inp_is_act";
 
-goal thy "!!a.[|a:outputs S|] ==> a:actions S";
+Goal "!!a.[|a:outputs S|] ==> a:actions S";
 by (asm_full_simp_tac (simpset() addsimps [asig_outputs_def,actions_def]) 1);
 qed"out_is_act";
 
-goal thy "(x: actions S & x : externals S) = (x: externals S)";
+Goal "(x: actions S & x : externals S) = (x: externals S)";
 by (fast_tac (claset() addSIs [ext_is_act]) 1 );
 qed"ext_and_act";
  
-goal thy "!!S. [|is_asig S;x: actions S|] ==> (x~:externals S) = (x: internals S)";
+Goal "!!S. [|is_asig S;x: actions S|] ==> (x~:externals S) = (x: internals S)";
 by (asm_full_simp_tac (simpset() addsimps [actions_def,is_asig_def,externals_def]) 1);
 by (best_tac (set_cs addEs [equalityCE]) 1);
 qed"not_ext_is_int";
 
-goal thy "!!S. is_asig S ==> (x~:externals S) = (x: internals S | x~:actions S)";
+Goal "!!S. is_asig S ==> (x~:externals S) = (x: internals S | x~:actions S)";
 by (asm_full_simp_tac (simpset() addsimps [actions_def,is_asig_def,externals_def]) 1);
 by (best_tac (set_cs addEs [equalityCE]) 1);
 qed"not_ext_is_int_or_not_act";
 
-goalw thy  [externals_def,actions_def,is_asig_def]
+Goalw  [externals_def,actions_def,is_asig_def]
  "!! x. [| is_asig (S); x:internals S |] ==> x~:externals S";
 by (Asm_full_simp_tac 1);
 by (best_tac (set_cs addEs [equalityCE]) 1);
--- a/src/HOLCF/IOA/meta_theory/Automata.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/IOA/meta_theory/Automata.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -19,7 +19,7 @@
 section "asig_of, starts_of, trans_of";
 
 
-goal thy
+Goal
  "((asig_of (x,y,z,w,s)) = x)   & \
 \ ((starts_of (x,y,z,w,s)) = y) & \
 \ ((trans_of (x,y,z,w,s)) = z)  & \
@@ -28,19 +28,19 @@
   by (simp_tac (simpset() addsimps ioa_projections) 1);
 qed "ioa_triple_proj";
 
-goalw thy [is_trans_of_def,actions_def, is_asig_def]
+Goalw [is_trans_of_def,actions_def, is_asig_def]
   "!!A. [| is_trans_of A; (s1,a,s2):trans_of(A) |] ==> a:act A";
   by (REPEAT(etac conjE 1));
   by (EVERY1[etac allE, etac impE, atac]);
   by (Asm_full_simp_tac 1);
 qed "trans_in_actions";
 
-goal 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";
 
-goal thy
+Goal
 "trans_of(A || B) = {tr. let s = fst(tr); a = fst(snd(tr)); t = snd(snd(tr)) \       
 \            in (a:act A | a:act B) & \
 \               (if a:act A then       \                
@@ -60,7 +60,7 @@
 section "actions and par";
 
 
-goal 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);
@@ -68,12 +68,12 @@
 qed "actions_asig_comp";
 
 
-goal 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 thy "ext (A1||A2) =    \
+Goal "ext (A1||A2) =    \
 \  (ext A1) Un (ext 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);
@@ -81,7 +81,7 @@
 by (fast_tac set_cs 1);
 qed"externals_of_par"; 
 
-goal thy "act (A1||A2) =    \
+Goal "act (A1||A2) =    \
 \  (act A1) Un (act A2)";
 by (asm_full_simp_tac (simpset() addsimps [actions_def,asig_of_par,asig_comp_def,
       asig_inputs_def,asig_outputs_def,asig_internals_def,Un_def,set_diff_def]) 1);
@@ -89,19 +89,19 @@
 by (fast_tac set_cs 1);
 qed"actions_of_par"; 
 
-goal thy "inp (A1||A2) =\
+Goal "inp (A1||A2) =\
 \         ((inp A1) Un (inp A2)) - ((out A1) Un (out A2))";
 by (asm_full_simp_tac (simpset() addsimps [actions_def,asig_of_par,asig_comp_def,
       asig_inputs_def,asig_outputs_def,Un_def,set_diff_def]) 1);
 qed"inputs_of_par";
   
-goal thy "out (A1||A2) =\
+Goal "out (A1||A2) =\
 \         (out A1) Un (out A2)";
 by (asm_full_simp_tac (simpset() addsimps [actions_def,asig_of_par,asig_comp_def,
       asig_outputs_def,Un_def,set_diff_def]) 1);
 qed"outputs_of_par";
 
-goal thy "int (A1||A2) =\
+Goal "int (A1||A2) =\
 \         (int A1) Un (int A2)";
 by (asm_full_simp_tac (simpset() addsimps [actions_def,asig_of_par,asig_comp_def,
       asig_inputs_def,asig_outputs_def,asig_internals_def,Un_def,set_diff_def]) 1);
@@ -111,19 +111,19 @@
 
 section "actions and compatibility"; 
 
-goal thy"compatible A B = compatible B A";
+Goal"compatible A B = compatible B A";
 by (asm_full_simp_tac (simpset() addsimps [compatible_def,Int_commute]) 1);
 by Auto_tac;
 qed"compat_commute";
 
-goalw thy [externals_def,actions_def,compatible_def]
+Goalw [externals_def,actions_def,compatible_def]
  "!! a. [| compatible A1 A2; a:ext A1|] ==> a~:int A2";
 by (Asm_full_simp_tac 1);
 by (best_tac (set_cs addEs [equalityCE]) 1);
 qed"ext1_is_not_int2";
 
 (* just commuting the previous one: better commute compatible *)
-goalw thy [externals_def,actions_def,compatible_def]
+Goalw [externals_def,actions_def,compatible_def]
  "!! a. [| compatible A2 A1 ; a:ext A1|] ==> a~:int A2";
 by (Asm_full_simp_tac 1);
 by (best_tac (set_cs addEs [equalityCE]) 1);
@@ -132,20 +132,20 @@
 bind_thm("ext1_ext2_is_not_act2",ext1_is_not_int2 RS int_and_ext_is_act);
 bind_thm("ext1_ext2_is_not_act1",ext2_is_not_int1 RS int_and_ext_is_act);
 
-goalw thy  [externals_def,actions_def,compatible_def]
+Goalw  [externals_def,actions_def,compatible_def]
  "!! x. [| compatible A B; x:int A |] ==> x~:ext B";
 by (Asm_full_simp_tac 1);
 by (best_tac (set_cs addEs [equalityCE]) 1);
 qed"intA_is_not_extB";
 
-goalw thy [externals_def,actions_def,compatible_def,is_asig_def,asig_of_def]
+Goalw [externals_def,actions_def,compatible_def,is_asig_def,asig_of_def]
 "!! a. [| compatible A B; a:int A |] ==> a ~: act B";
 by (Asm_full_simp_tac 1);
 by (best_tac (set_cs addEs [equalityCE]) 1);
 qed"intA_is_not_actB";
 
 (* the only one that needs disjointness of outputs and of internals and _all_ acts *)
-goalw thy [asig_outputs_def,asig_internals_def,actions_def,asig_inputs_def,
+Goalw [asig_outputs_def,asig_internals_def,actions_def,asig_inputs_def,
     compatible_def,is_asig_def,asig_of_def]
 "!! a. [| compatible A B; a:out A ;a:act B|] ==> a : inp B";
 by (Asm_full_simp_tac 1);
@@ -153,7 +153,7 @@
 qed"outAactB_is_inpB";
 
 (* needed for propagation of input_enabledness from A,B to A||B *)
-goalw thy [asig_outputs_def,asig_internals_def,actions_def,asig_inputs_def,
+Goalw [asig_outputs_def,asig_internals_def,actions_def,asig_inputs_def,
     compatible_def,is_asig_def,asig_of_def]
 "!! a. [| compatible A B; a:inp A ;a:act B|] ==> a : inp B | a: out B";
 by (Asm_full_simp_tac 1);
@@ -168,7 +168,7 @@
 (* ugly case distinctions. Heart of proof: 
      1. inpAAactB_is_inpBoroutB ie. internals are really hidden.
      2. inputs_of_par: outputs are no longer inputs of par. This is important here *)
-goalw thy [input_enabled_def] 
+Goalw [input_enabled_def] 
 "!!A. [| compatible A B; input_enabled A; input_enabled B|] \
 \     ==> input_enabled (A||B)";
 by (asm_full_simp_tac (simpset()addsimps[Let_def,inputs_of_par,trans_of_par])1);
@@ -267,12 +267,12 @@
 section "restrict";
 
 
-goal 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)";
 by (simp_tac (simpset() addsimps ([restrict_def]@ioa_projections)) 1);
 qed "cancel_restrict_a";
 
-goal thy "reachable (restrict ioa acts) s = reachable ioa s";
+Goal "reachable (restrict ioa acts) s = reachable ioa s";
 by (rtac iffI 1);
 by (etac reachable.induct 1);
 by (asm_full_simp_tac (simpset() addsimps [cancel_restrict_a,reachable_0]) 1);
@@ -286,14 +286,14 @@
 by (asm_full_simp_tac (simpset() addsimps [cancel_restrict_a]) 1);
 qed "cancel_restrict_b";
 
-goal thy "act (restrict A acts) = act A";
+Goal "act (restrict A acts) = act A";
 by (simp_tac (simpset() addsimps [actions_def,asig_internals_def,
         asig_outputs_def,asig_inputs_def,externals_def,asig_of_def,restrict_def,
         restrict_asig_def]) 1);
 by Auto_tac;
 qed"acts_restrict";
 
-goal 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 & \
 \         act (restrict A acts) = act A";
@@ -306,12 +306,12 @@
 
 
 
-goal thy "!!f. s -a--(rename C f)-> t ==> (? x. Some(x) = f(a) & s -x--C-> t)";
+Goal "!!f. s -a--(rename C f)-> t ==> (? x. Some(x) = f(a) & s -x--C-> t)";
 by (asm_full_simp_tac (simpset() addsimps [Let_def,rename_def,trans_of_def]) 1);
 qed"trans_rename";
 
 
-goal thy "!!s.[| reachable (rename C g) s |] ==> reachable C s";
+Goal "!!s.[| reachable (rename C g) s |] ==> reachable C s";
 by (etac reachable.induct 1);
 by (rtac reachable_0 1);
 by (asm_full_simp_tac (simpset() addsimps [rename_def]@ioa_projections) 1);
@@ -329,44 +329,44 @@
 section "trans_of(A||B)";
 
 
-goal thy "!!A.[|(s,a,t):trans_of (A||B); a:act A|] \
+Goal "!!A.[|(s,a,t):trans_of (A||B); a:act A|] \
 \             ==> (fst s,a,fst t):trans_of A";
 by (asm_full_simp_tac (simpset() addsimps [Let_def,par_def,trans_of_def]) 1);
 qed"trans_A_proj";
 
-goal thy "!!A.[|(s,a,t):trans_of (A||B); a:act B|] \
+Goal "!!A.[|(s,a,t):trans_of (A||B); a:act B|] \
 \             ==> (snd s,a,snd t):trans_of B";
 by (asm_full_simp_tac (simpset() addsimps [Let_def,par_def,trans_of_def]) 1);
 qed"trans_B_proj";
 
-goal thy "!!A.[|(s,a,t):trans_of (A||B); a~:act A|]\
+Goal "!!A.[|(s,a,t):trans_of (A||B); a~:act A|]\
 \             ==> fst s = fst t";
 by (asm_full_simp_tac (simpset() addsimps [Let_def,par_def,trans_of_def]) 1);
 qed"trans_A_proj2";
 
-goal thy "!!A.[|(s,a,t):trans_of (A||B); a~:act B|]\
+Goal "!!A.[|(s,a,t):trans_of (A||B); a~:act B|]\
 \             ==> snd s = snd t";
 by (asm_full_simp_tac (simpset() addsimps [Let_def,par_def,trans_of_def]) 1);
 qed"trans_B_proj2";
 
-goal thy "!!A.(s,a,t):trans_of (A||B) \
+Goal "!!A.(s,a,t):trans_of (A||B) \
 \              ==> a :act A | a :act B";
 by (asm_full_simp_tac (simpset() addsimps [Let_def,par_def,trans_of_def]) 1);
 qed"trans_AB_proj";
 
-goal thy "!!A. [|a:act A;a:act B;\
+Goal "!!A. [|a:act A;a:act B;\
 \      (fst s,a,fst t):trans_of A;(snd s,a,snd t):trans_of B|]\
 \  ==> (s,a,t):trans_of (A||B)";
 by (asm_full_simp_tac (simpset() addsimps [Let_def,par_def,trans_of_def]) 1);
 qed"trans_AB";
 
-goal thy "!!A. [|a:act A;a~:act B;\
+Goal "!!A. [|a:act A;a~:act B;\
 \      (fst s,a,fst t):trans_of A;snd s=snd t|]\
 \  ==> (s,a,t):trans_of (A||B)";
 by (asm_full_simp_tac (simpset() addsimps [Let_def,par_def,trans_of_def]) 1);
 qed"trans_A_notB";
 
-goal thy "!!A. [|a~:act A;a:act B;\
+Goal "!!A. [|a~:act A;a:act B;\
 \      (snd s,a,snd t):trans_of B;fst s=fst t|]\
 \  ==> (s,a,t):trans_of (A||B)";
 by (asm_full_simp_tac (simpset() addsimps [Let_def,par_def,trans_of_def]) 1);
@@ -377,7 +377,7 @@
                       trans_B_proj2,trans_AB_proj];
 
 
-goal 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))) &                                                 \
@@ -402,16 +402,16 @@
 section "proof obligation generator for IOA requirements";  
 
 (* without assumptions on A and B because is_trans_of is also incorporated in ||def *)
-goalw thy [is_trans_of_def] "is_trans_of (A||B)";
+Goalw [is_trans_of_def] "is_trans_of (A||B)";
 by (simp_tac (simpset() addsimps [Let_def,actions_of_par,trans_of_par]) 1);
 qed"is_trans_of_par";
 
-goalw thy [is_trans_of_def] 
+Goalw [is_trans_of_def] 
 "!!A. is_trans_of A ==> is_trans_of (restrict A acts)";
 by (asm_simp_tac (simpset() addsimps [cancel_restrict,acts_restrict])1);
 qed"is_trans_of_restrict";
 
-goalw thy [is_trans_of_def,restrict_def,restrict_asig_def] 
+Goalw [is_trans_of_def,restrict_def,restrict_asig_def] 
 "!!A. is_trans_of A ==> is_trans_of (rename A f)";
 by (asm_full_simp_tac
     (simpset() addsimps [Let_def,actions_def,trans_of_def, asig_internals_def,
@@ -420,7 +420,7 @@
 by (Blast_tac 1);
 qed"is_trans_of_rename";
 
-goal thy "!! A. [| is_asig_of A; is_asig_of B; compatible A B|]  \
+Goal "!! A. [| is_asig_of A; is_asig_of B; compatible A B|]  \
 \         ==> is_asig_of (A||B)";
 by (asm_full_simp_tac (simpset() addsimps [is_asig_of_def,asig_of_par,
        asig_comp_def,compatible_def,asig_internals_def,asig_outputs_def,
@@ -430,7 +430,7 @@
 by (REPEAT (best_tac (set_cs addEs [equalityCE]) 1));
 qed"is_asig_of_par";
 
-goalw thy [is_asig_of_def,is_asig_def,asig_of_def,restrict_def,restrict_asig_def,
+Goalw [is_asig_of_def,is_asig_def,asig_of_def,restrict_def,restrict_asig_def,
            asig_internals_def,asig_outputs_def,asig_inputs_def,externals_def,o_def] 
 "!! A. is_asig_of A ==> is_asig_of (restrict A f)";
 by (Asm_full_simp_tac 1);
@@ -438,7 +438,7 @@
 by (REPEAT (best_tac (set_cs addEs [equalityCE]) 1));
 qed"is_asig_of_restrict";
 
-goal thy "!! A. is_asig_of A ==> is_asig_of (rename A f)";
+Goal "!! A. is_asig_of A ==> is_asig_of (rename A f)";
 by (asm_full_simp_tac (simpset() addsimps [is_asig_of_def,
      rename_def,rename_set_def,asig_internals_def,asig_outputs_def,
      asig_inputs_def,actions_def,is_asig_def,asig_of_def]) 1);
@@ -452,7 +452,7 @@
           is_trans_of_par,is_trans_of_restrict,is_trans_of_rename];
 
 
-goalw thy [compatible_def]
+Goalw [compatible_def]
 "!! A. [|compatible A B; compatible A C |]==> compatible A (B||C)";
 by (asm_full_simp_tac (simpset() addsimps [internals_of_par, 
    outputs_of_par,actions_of_par]) 1);
@@ -461,7 +461,7 @@
 qed"compatible_par";
 
 (* FIX: better derive by previous one and compat_commute *)
-goalw thy [compatible_def]
+Goalw [compatible_def]
 "!! A. [|compatible A C; compatible B C |]==> compatible (A||B) C";
 by (asm_full_simp_tac (simpset() addsimps [internals_of_par, 
    outputs_of_par,actions_of_par]) 1);
@@ -469,7 +469,7 @@
 by (REPEAT (best_tac (set_cs addEs [equalityCE]) 1));
 qed"compatible_par2";
 
-goalw thy [compatible_def]
+Goalw [compatible_def]
 "!! A. [| compatible A B; (ext B - S) Int ext A = {}|] \
 \     ==> compatible A (restrict B S)";
 by (asm_full_simp_tac (simpset() addsimps [ioa_triple_proj,asig_triple_proj,
--- a/src/HOLCF/IOA/meta_theory/CompoExecs.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/IOA/meta_theory/CompoExecs.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -20,15 +20,15 @@
 (* ---------------------------------------------------------------- *)
 
 
-goal thy  "ProjA2`UU = UU";
+Goal  "ProjA2`UU = UU";
 by (simp_tac (simpset() addsimps [ProjA2_def]) 1);
 qed"ProjA2_UU";
 
-goal thy  "ProjA2`nil = nil";
+Goal  "ProjA2`nil = nil";
 by (simp_tac (simpset() addsimps [ProjA2_def]) 1);
 qed"ProjA2_nil";
 
-goal thy "ProjA2`((a,t)>>xs) = (a,fst t) >> ProjA2`xs";
+Goal "ProjA2`((a,t)>>xs) = (a,fst t) >> ProjA2`xs";
 by (simp_tac (simpset() addsimps [ProjA2_def]) 1);
 qed"ProjA2_cons";
 
@@ -38,15 +38,15 @@
 (* ---------------------------------------------------------------- *)
 
 
-goal thy  "ProjB2`UU = UU";
+Goal  "ProjB2`UU = UU";
 by (simp_tac (simpset() addsimps [ProjB2_def]) 1);
 qed"ProjB2_UU";
 
-goal thy  "ProjB2`nil = nil";
+Goal  "ProjB2`nil = nil";
 by (simp_tac (simpset() addsimps [ProjB2_def]) 1);
 qed"ProjB2_nil";
 
-goal thy "ProjB2`((a,t)>>xs) = (a,snd t) >> ProjB2`xs";
+Goal "ProjB2`((a,t)>>xs) = (a,snd t) >> ProjB2`xs";
 by (simp_tac (simpset() addsimps [ProjB2_def]) 1);
 qed"ProjB2_cons";
 
@@ -57,15 +57,15 @@
 (* ---------------------------------------------------------------- *)
 
 
-goal thy "Filter_ex2 sig`UU=UU";
+Goal "Filter_ex2 sig`UU=UU";
 by (simp_tac (simpset() addsimps [Filter_ex2_def]) 1);
 qed"Filter_ex2_UU";
 
-goal thy "Filter_ex2 sig`nil=nil";
+Goal "Filter_ex2 sig`nil=nil";
 by (simp_tac (simpset() addsimps [Filter_ex2_def]) 1);
 qed"Filter_ex2_nil";
 
-goal thy "Filter_ex2 sig`(at >> xs) =    \
+Goal "Filter_ex2 sig`(at >> xs) =    \
 \            (if (fst at:actions sig)    \       
 \                 then at >> (Filter_ex2 sig`xs) \   
 \                 else        Filter_ex2 sig`xs)";
@@ -79,7 +79,7 @@
 (* ---------------------------------------------------------------- *)
 
 
-goal thy "stutter2 sig = (LAM ex. (%s. case ex of \
+Goal "stutter2 sig = (LAM ex. (%s. case ex of \
 \      nil => TT \
 \    | x##xs => (flift1 \ 
 \            (%p.(If Def ((fst p)~:actions sig) \
@@ -95,17 +95,17 @@
 by (simp_tac (simpset() addsimps [flift1_def]) 1);
 qed"stutter2_unfold";
 
-goal thy "(stutter2 sig`UU) s=UU";
+Goal "(stutter2 sig`UU) s=UU";
 by (stac stutter2_unfold 1);
 by (Simp_tac 1);
 qed"stutter2_UU";
 
-goal thy "(stutter2 sig`nil) s = TT";
+Goal "(stutter2 sig`nil) s = TT";
 by (stac stutter2_unfold 1);
 by (Simp_tac 1);
 qed"stutter2_nil";
 
-goal thy "(stutter2 sig`(at>>xs)) s = \
+Goal "(stutter2 sig`(at>>xs)) s = \
 \              ((if (fst at)~:actions sig then Def (s=snd at) else TT) \
 \                andalso (stutter2 sig`xs) (snd at))"; 
 by (rtac trans 1);
@@ -122,15 +122,15 @@
 (*                             stutter                              *)
 (* ---------------------------------------------------------------- *)
 
-goal thy "stutter sig (s, UU)";
+Goal "stutter sig (s, UU)";
 by (simp_tac (simpset() addsimps [stutter_def]) 1);
 qed"stutter_UU";
 
-goal thy "stutter sig (s, nil)";
+Goal "stutter sig (s, nil)";
 by (simp_tac (simpset() addsimps [stutter_def]) 1);
 qed"stutter_nil";
 
-goal thy "stutter sig (s, (a,t)>>ex) = \
+Goal "stutter sig (s, (a,t)>>ex) = \
 \     ((a~:actions sig --> (s=t)) & stutter sig (t,ex))";
 by (simp_tac (simpset() addsimps [stutter_def]) 1);
 qed"stutter_cons";
@@ -158,7 +158,7 @@
 (*  Lemma_1_1a : is_ex_fr propagates from A||B to Projections A and B    *)
 (* --------------------------------------------------------------------- *)
 
-goal thy "!s. is_exec_frag (A||B) (s,xs) \
+Goal "!s. is_exec_frag (A||B) (s,xs) \
 \      -->  is_exec_frag A (fst s, Filter_ex2 (asig_of A)`(ProjA2`xs)) &\
 \           is_exec_frag B (snd s, Filter_ex2 (asig_of B)`(ProjB2`xs))";
 
@@ -174,7 +174,7 @@
 (*  Lemma_1_1b : is_ex_fr (A||B) implies stuttering on Projections       *)
 (* --------------------------------------------------------------------- *)
 
-goal thy "!s. is_exec_frag (A||B) (s,xs) \
+Goal "!s. is_exec_frag (A||B) (s,xs) \
 \      --> stutter (asig_of A) (fst s,ProjA2`xs)  &\
 \          stutter (asig_of B) (snd s,ProjB2`xs)"; 
 
@@ -189,7 +189,7 @@
 (*  Lemma_1_1c : Executions of A||B have only  A- or B-actions           *)
 (* --------------------------------------------------------------------- *)
 
-goal thy "!s. (is_exec_frag (A||B) (s,xs) \
+Goal "!s. (is_exec_frag (A||B) (s,xs) \
 \  --> Forall (%x. fst x:act (A||B)) xs)";
 
 by (pair_induct_tac "xs" [Forall_def,sforall_def,is_exec_frag_def] 1);
@@ -204,7 +204,7 @@
 (*  Lemma_1_2 : ex A, exB, stuttering and forall a:A|B implies ex (A||B)   *)
 (* ----------------------------------------------------------------------- *)
 
-goal thy 
+Goal 
 "!s. is_exec_frag A (fst s,Filter_ex2 (asig_of A)`(ProjA2`xs)) &\
 \    is_exec_frag B (snd s,Filter_ex2 (asig_of B)`(ProjB2`xs)) &\
 \    stutter (asig_of A) (fst s,(ProjA2`xs)) & \
@@ -230,7 +230,7 @@
 (* ------------------------------------------------------------------ *)
 
 
-goal thy 
+Goal 
 "ex:executions(A||B) =\
 \(Filter_ex (asig_of A) (ProjA ex) : executions A &\
 \ Filter_ex (asig_of B) (ProjB ex) : executions B &\
@@ -256,7 +256,7 @@
 (*                            For Modules                             *)
 (* ------------------------------------------------------------------ *)
 
-goalw thy [Execs_def,par_execs_def]
+Goalw [Execs_def,par_execs_def]
 
 "Execs (A||B) = par_execs (Execs A) (Execs B)";
 
--- a/src/HOLCF/IOA/meta_theory/CompoScheds.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/IOA/meta_theory/CompoScheds.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -56,17 +56,17 @@
 \      )))");
 
 
-goal thy "(mkex2 A B`UU`exA`exB) s t = UU";
+Goal "(mkex2 A B`UU`exA`exB) s t = UU";
 by (stac mkex2_unfold 1);
 by (Simp_tac 1);
 qed"mkex2_UU";
 
-goal thy "(mkex2 A B`nil`exA`exB) s t= nil";
+Goal "(mkex2 A B`nil`exA`exB) s t= nil";
 by (stac mkex2_unfold 1);
 by (Simp_tac 1);
 qed"mkex2_nil";
 
-goal thy "!!x.[| x:act A; x~:act B; HD`exA=Def a|] \
+Goal "!!x.[| x:act A; x~:act B; HD`exA=Def a|] \
 \   ==> (mkex2 A B`(x>>sch)`exA`exB) s t =   \
 \       (x,snd a,t) >> (mkex2 A B`sch`(TL`exA)`exB) (snd a) t";
 by (rtac trans 1);
@@ -75,7 +75,7 @@
 by (asm_full_simp_tac (simpset() addsimps [Cons_def]) 1);
 qed"mkex2_cons_1";
 
-goal thy "!!x.[| x~:act A; x:act B; HD`exB=Def b|] \
+Goal "!!x.[| x~:act A; x:act B; HD`exB=Def b|] \
 \   ==> (mkex2 A B`(x>>sch)`exA`exB) s t = \
 \       (x,s,snd b) >> (mkex2 A B`sch`exA`(TL`exB)) s (snd b)";
 by (rtac trans 1);
@@ -84,7 +84,7 @@
 by (asm_full_simp_tac (simpset() addsimps [Cons_def]) 1);
 qed"mkex2_cons_2";
 
-goal thy "!!x.[| x:act A; x:act B; HD`exA=Def a;HD`exB=Def b|] \
+Goal "!!x.[| x:act A; x:act B; HD`exA=Def a;HD`exB=Def b|] \
 \   ==> (mkex2 A B`(x>>sch)`exA`exB) s t =  \
 \        (x,snd a,snd b) >> \
 \           (mkex2 A B`sch`(TL`exA)`(TL`exB)) (snd a) (snd b)";
@@ -101,15 +101,15 @@
 (*                             mkex                                 *)
 (* ---------------------------------------------------------------- *)
 
-goal thy "mkex A B UU  (s,exA) (t,exB) = ((s,t),UU)";
+Goal "mkex A B UU  (s,exA) (t,exB) = ((s,t),UU)";
 by (simp_tac (simpset() addsimps [mkex_def]) 1);
 qed"mkex_UU";
 
-goal thy "mkex A B nil (s,exA) (t,exB) = ((s,t),nil)";
+Goal "mkex A B nil (s,exA) (t,exB) = ((s,t),nil)";
 by (simp_tac (simpset() addsimps [mkex_def]) 1);
 qed"mkex_nil";
 
-goal thy "!!x.[| x:act A; x~:act B |] \
+Goal "!!x.[| x:act A; x~:act B |] \
 \   ==> mkex A B (x>>sch) (s,a>>exA) (t,exB)  =  \
 \       ((s,t), (x,snd a,t) >> snd (mkex A B sch (snd a,exA) (t,exB)))";
 by (simp_tac (simpset() addsimps [mkex_def]) 1);
@@ -117,7 +117,7 @@
 by Auto_tac;
 qed"mkex_cons_1";
 
-goal thy "!!x.[| x~:act A; x:act B |] \
+Goal "!!x.[| x~:act A; x:act B |] \
 \   ==> mkex A B (x>>sch) (s,exA) (t,b>>exB) =  \ 
 \       ((s,t), (x,s,snd b) >> snd (mkex A B sch (s,exA) (snd b,exB)))";
 by (simp_tac (simpset() addsimps [mkex_def]) 1);
@@ -125,7 +125,7 @@
 by Auto_tac;
 qed"mkex_cons_2";
 
-goal thy "!!x.[| x:act A; x:act B |]  \
+Goal "!!x.[| x:act A; x:act B |]  \
 \   ==>  mkex A B (x>>sch) (s,a>>exA) (t,b>>exB) = \
 \        ((s,t), (x,snd a,snd b) >> snd (mkex A B sch (snd a,exA) (snd b,exB)))";
 by (simp_tac (simpset() addsimps [mkex_def]) 1);
@@ -155,7 +155,7 @@
 (*    Lemma_2_1 :  tfilter(ex) and filter_act are commutative            *)
 (* --------------------------------------------------------------------- *)
 
-goalw thy [filter_act_def,Filter_ex2_def]
+Goalw [filter_act_def,Filter_ex2_def]
    "filter_act`(Filter_ex2 (asig_of A)`xs)=\
 \   Filter (%a. a:act A)`(filter_act`xs)";
 
@@ -167,7 +167,7 @@
 (*    Lemma_2_2 : State-projections do not affect filter_act             *)
 (* --------------------------------------------------------------------- *)
 
-goal thy 
+Goal 
    "filter_act`(ProjA2`xs) =filter_act`xs &\
 \   filter_act`(ProjB2`xs) =filter_act`xs";
 
@@ -183,7 +183,7 @@
    an ex is in A or B, but after projecting it onto the action schedule. Of course, this
    is the same proposition, but we cannot change this one, when then rather lemma_1_1c  *)
 
-goal thy "!s. is_exec_frag (A||B) (s,xs) \
+Goal "!s. is_exec_frag (A||B) (s,xs) \
 \  --> Forall (%x. x:act (A||B)) (filter_act`xs)";
 
 by (pair_induct_tac "xs" [is_exec_frag_def,Forall_def,sforall_def] 1);
@@ -203,7 +203,7 @@
                              structural induction
   --------------------------------------------------------------------------- *)
 
-goal thy "! exA exB s t. \
+Goal "! exA exB s t. \
 \ Forall (%x. x:act (A||B)) sch  & \
 \ Filter (%a. a:act A)`sch << filter_act`exA &\
 \ Filter (%a. a:act B)`sch << filter_act`exB \
@@ -266,7 +266,7 @@
   --------------------------------------------------------------------------- *)
 
 
-goal thy "! exA exB s t. \
+Goal "! exA exB s t. \
 \ Forall (%x. x:act (A||B)) sch & \
 \ Filter (%a. a:act A)`sch << filter_act`exA &\
 \ Filter (%a. a:act B)`sch << filter_act`exB \
@@ -277,7 +277,7 @@
 qed"stutterA_mkex";
 
 
-goal thy "!! sch.[|  \
+Goal "!! sch.[|  \
 \ Forall (%x. x:act (A||B)) sch ; \
 \ Filter (%a. a:act A)`sch << filter_act`(snd exA) ;\
 \ Filter (%a. a:act B)`sch << filter_act`(snd exB) |] \
@@ -297,7 +297,7 @@
                              structural induction
   --------------------------------------------------------------------------- *)
 
-goal thy "! exA exB s t. \
+Goal "! exA exB s t. \
 \ Forall (%x. x:act (A||B)) sch & \
 \ Filter (%a. a:act A)`sch << filter_act`exA &\
 \ Filter (%a. a:act B)`sch << filter_act`exB \
@@ -308,7 +308,7 @@
 qed"stutterB_mkex";
 
 
-goal thy "!! sch.[|  \
+Goal "!! sch.[|  \
 \ Forall (%x. x:act (A||B)) sch ; \
 \ Filter (%a. a:act A)`sch << filter_act`(snd exA) ;\
 \ Filter (%a. a:act B)`sch << filter_act`(snd exB) |] \
@@ -330,7 +330,7 @@
                              structural induction
   --------------------------------------------------------------------------- *)
 
-goal thy "! exA exB s t. \
+Goal "! exA exB s t. \
 \ Forall (%x. x:act (A||B)) sch & \
 \ Filter (%a. a:act A)`sch << filter_act`exA  &\
 \ Filter (%a. a:act B)`sch << filter_act`exB \
@@ -346,7 +346,7 @@
                     lemma for admissibility problems         
   --------------------------------------------------------------------------- *)
 
-goal thy  "Zip`(Map fst`y)`(Map snd`y) = y";
+Goal  "Zip`(Map fst`y)`(Map snd`y) = y";
 by (Seq_induct_tac "y" [] 1);
 qed"Zip_Map_fst_snd";
 
@@ -356,7 +356,7 @@
          lemma for eliminating non admissible equations in assumptions      
   --------------------------------------------------------------------------- *)
 
-goal thy "!! sch ex. \
+Goal "!! sch ex. \
 \ Filter (%a. a:act AB)`sch = filter_act`ex  \
 \ ==> ex = Zip`(Filter (%a. a:act AB)`sch)`(Map snd`ex)";
 by (asm_full_simp_tac (simpset() addsimps [filter_act_def]) 1);
@@ -369,7 +369,7 @@
   --------------------------------------------------------------------------- *)
 
 
-goal thy "!!sch exA exB.\
+Goal "!!sch exA exB.\
 \ [| Forall (%a. a:act (A||B)) sch ; \
 \ Filter (%a. a:act A)`sch = filter_act`(snd exA)  ;\
 \ Filter (%a. a:act B)`sch = filter_act`(snd exB) |]\
@@ -394,7 +394,7 @@
   --------------------------------------------------------------------------- *)
 
 
-goal thy "! exA exB s t. \
+Goal "! exA exB s t. \
 \ Forall (%x. x:act (A||B)) sch & \
 \ Filter (%a. a:act A)`sch << filter_act`exA  &\
 \ Filter (%a. a:act B)`sch << filter_act`exB \
@@ -413,7 +413,7 @@
   --------------------------------------------------------------------------- *)
 
 
-goal thy "!!sch exA exB.\
+Goal "!!sch exA exB.\
 \ [| Forall (%a. a:act (A||B)) sch ; \
 \ Filter (%a. a:act A)`sch = filter_act`(snd exA)  ;\
 \ Filter (%a. a:act B)`sch = filter_act`(snd exB) |]\
@@ -434,7 +434,7 @@
 (* --------------------------------------------------------------------- *)
 
 
-goal thy "!s t exA exB. \
+Goal "!s t exA exB. \
 \ Forall (%x. x : act (A || B)) sch &\
 \ Filter (%a. a:act A)`sch << filter_act`exA  &\
 \ Filter (%a. a:act B)`sch << filter_act`exB \
@@ -451,7 +451,7 @@
 (*                          Main Theorem                              *)
 (* ------------------------------------------------------------------ *)
 
-goal thy  
+Goal  
 "sch : schedules (A||B) = \
 \ (Filter (%a. a:act A)`sch : schedules A &\
 \  Filter (%a. a:act B)`sch : schedules B &\
@@ -500,7 +500,7 @@
 (*                            For Modules                             *)
 (* ------------------------------------------------------------------ *)
 
-goalw thy [Scheds_def,par_scheds_def]
+Goalw [Scheds_def,par_scheds_def]
 
 "Scheds (A||B) = par_scheds (Scheds A) (Scheds B)";
 
--- a/src/HOLCF/IOA/meta_theory/CompoTraces.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/IOA/meta_theory/CompoTraces.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -50,17 +50,17 @@
 \       ))");
 
 
-goal thy "mksch A B`UU`schA`schB = UU";
+Goal "mksch A B`UU`schA`schB = UU";
 by (stac mksch_unfold 1);
 by (Simp_tac 1);
 qed"mksch_UU";
 
-goal thy "mksch A B`nil`schA`schB = nil";
+Goal "mksch A B`nil`schA`schB = nil";
 by (stac mksch_unfold 1);
 by (Simp_tac 1);
 qed"mksch_nil";
 
-goal thy "!!x.[|x:act A;x~:act B|]  \
+Goal "!!x.[|x:act A;x~:act B|]  \
 \   ==> mksch A B`(x>>tr)`schA`schB = \
 \         (Takewhile (%a. a:int A)`schA) \
 \         @@ (x>>(mksch A B`tr`(TL`(Dropwhile (%a. a:int A)`schA)) \
@@ -71,7 +71,7 @@
 by (simp_tac (simpset() addsimps [Cons_def]) 1);
 qed"mksch_cons1";
 
-goal thy "!!x.[|x~:act A;x:act B|] \
+Goal "!!x.[|x~:act A;x:act B|] \
 \   ==> mksch A B`(x>>tr)`schA`schB = \
 \        (Takewhile (%a. a:int B)`schB)  \
 \         @@ (x>>(mksch A B`tr`schA`(TL`(Dropwhile (%a. a:int B)`schB))  \
@@ -82,7 +82,7 @@
 by (simp_tac (simpset() addsimps [Cons_def]) 1);
 qed"mksch_cons2";
 
-goal thy "!!x.[|x:act A;x:act B|] \
+Goal "!!x.[|x:act A;x:act B|] \
 \   ==> mksch A B`(x>>tr)`schA`schB = \
 \            (Takewhile (%a. a:int A)`schA) \
 \         @@ ((Takewhile (%a. a:int B)`schB)  \
@@ -114,7 +114,7 @@
    the compatibility of IOA, in particular out of the condition that internals are 
    really hidden. *)
 
-goal thy "(eB & ~eA --> ~A) -->       \
+Goal "(eB & ~eA --> ~A) -->       \
 \         (A & (eA | eB)) = (eA & A)";
 by (Fast_tac 1);
 qed"compatibility_consequence1";
@@ -122,7 +122,7 @@
 
 (* very similar to above, only the commutativity of | is used to make a slight change *)
 
-goal thy "(eB & ~eA --> ~A) -->       \
+Goal "(eB & ~eA --> ~A) -->       \
 \         (A & (eB | eA)) = (eA & A)";
 by (Fast_tac 1);
 qed"compatibility_consequence2";
@@ -146,7 +146,7 @@
 qed"subst_lemma2";
 
 
-goal thy "!!A B. compatible A B ==> \
+Goal "!!A B. compatible A B ==> \
 \   ! schA schB. Forall (%x. x:act (A||B)) tr \
 \   --> Forall (%x. x:act (A||B)) (mksch A B`tr`schA`schB)";
 by (Seq_induct_tac "tr" [Forall_def,sforall_def,mksch_def] 1); 
@@ -173,7 +173,7 @@
 by Auto_tac;
 qed_spec_mp"ForallAorB_mksch";
 
-goal thy "!!A B. compatible B A  ==> \
+Goal "!!A B. compatible B A  ==> \
 \   ! schA schB.  (Forall (%x. x:act B & x~:act A) tr \
 \   --> Forall (%x. x:act B & x~:act A) (mksch A B`tr`schA`schB))";
 
@@ -184,7 +184,7 @@
                             intA_is_not_actB,int_is_act]) 1);
 qed_spec_mp "ForallBnAmksch";
 
-goal thy "!!A B. compatible A B ==> \
+Goal "!!A B. compatible A B ==> \
 \   ! schA schB.  (Forall (%x. x:act A & x~:act B) tr \
 \   --> Forall (%x. x:act A & x~:act B) (mksch A B`tr`schA`schB))";
 
@@ -199,7 +199,7 @@
 (* safe-tac makes too many case distinctions with this lemma in the next proof *)
 Delsimps [FiniteConc];
 
-goal thy "!! tr. [| Finite tr; is_asig(asig_of A); is_asig(asig_of B) |] ==> \
+Goal "!! tr. [| Finite tr; is_asig(asig_of A); is_asig(asig_of B) |] ==> \
 \   ! x y. Forall (%x. x:act A) x & Forall (%x. x:act B) y & \
 \          Filter (%a. a:ext A)`x = Filter (%a. a:act A)`tr & \
 \          Filter (%a. a:ext B)`y = Filter (%a. a:act B)`tr &\
@@ -272,7 +272,7 @@
 (* otherwise subst_lemma2 does not fit exactly, just to avoid a subst_lemma3 *)
 Delsimps [FilterConc]; 
 
-goal thy " !!bs. [| Finite bs; is_asig(asig_of A); is_asig(asig_of B);compatible A B|] ==>  \
+Goal " !!bs. [| Finite bs; is_asig(asig_of A); is_asig(asig_of B);compatible A B|] ==>  \
 \! y. Forall (%x. x:act B) y & Forall (%x. x:act B & x~:act A) bs &\
 \    Filter (%a. a:ext B)`y = Filter (%a. a:act B)`(bs @@ z) \
 \    --> (? y1 y2.  (mksch A B`(bs @@ z)`x`y) = (y1 @@ (mksch A B`z`x`y2)) & \
@@ -326,7 +326,7 @@
 Delsimps [FilterConc]; 
 
 
-goal thy " !!as. [| Finite as; is_asig(asig_of A); is_asig(asig_of B);compatible A B|] ==>  \
+Goal " !!as. [| Finite as; is_asig(asig_of A); is_asig(asig_of B);compatible A B|] ==>  \
 \! x. Forall (%x. x:act A) x & Forall (%x. x:act A & x~:act B) as &\
 \    Filter (%a. a:ext A)`x = Filter (%a. a:act A)`(as @@ z) \
 \    --> (? x1 x2.  (mksch A B`(as @@ z)`x`y) = (x1 @@ (mksch A B`z`x2`y)) & \
@@ -379,7 +379,7 @@
 (*
 
 
-goal thy "!! y.Finite y ==> ! z tr. Forall (%a.a:ext (A||B)) tr & \
+Goal "!! y.Finite y ==> ! z tr. Forall (%a.a:ext (A||B)) tr & \
 \                             y = z @@ mksch A B`tr`a`b\
 \                             --> Finite tr";
 by (etac Seq_Finite_ind  1);
@@ -438,19 +438,19 @@
 
 
 
-goal thy "(x>>xs = y @@ z) = ((y=nil & x>>xs=z) | (? ys. y=x>>ys & xs=ys@@z))";
+Goal "(x>>xs = y @@ z) = ((y=nil & x>>xs=z) | (? ys. y=x>>ys & xs=ys@@z))";
 by (Seq_case_simp_tac "y" 1);
 by Auto_tac;
 qed"Conc_Conc_eq";
 
-goal thy "!! (y::'a Seq).Finite y ==> ~ y= x@@UU";
+Goal "!! (y::'a Seq).Finite y ==> ~ y= x@@UU";
 by (etac Seq_Finite_ind 1);
 by (Seq_case_simp_tac "x" 1);
 by (Seq_case_simp_tac "x" 1);
 by Auto_tac;
 qed"FiniteConcUU1";
 
-goal thy "~ Finite ((x::'a Seq)@@UU)";
+Goal "~ Finite ((x::'a Seq)@@UU)";
 by (rtac (FiniteConcUU1 COMP rev_contrapos) 1);
 by Auto_tac;
 qed"FiniteConcUU";
@@ -466,7 +466,7 @@
 (*                             structural induction
   ------------------------------------------------------------------------------------- *)
 
-goal thy 
+Goal 
 "!! A B. [| compatible A B; compatible B A;\
 \           is_asig(asig_of A); is_asig(asig_of B) |] ==> \
 \ ! schA schB. Forall (%x. x:act A) schA & Forall (%x. x:act B) schB & \
@@ -543,7 +543,7 @@
                              take lemma
   --------------------------------------------------------------------------- *)
 
-goal thy "!! A B. [| compatible A B; compatible B A; \
+Goal "!! A B. [| compatible A B; compatible B A; \
 \ is_asig(asig_of A); is_asig(asig_of B) |] ==> \
 \ Forall (%x. x:ext (A||B)) tr & \
 \ Forall (%x. x:act A) schA & Forall (%x. x:act B) schB & \
@@ -677,7 +677,7 @@
 
 
 
-goal thy "!! A B. [| compatible A B; compatible B A; \
+Goal "!! A B. [| compatible A B; compatible B A; \
 \ is_asig(asig_of A); is_asig(asig_of B) |] ==> \
 \ Forall (%x. x:ext (A||B)) tr & \
 \ Forall (%x. x:act A) schA & Forall (%x. x:act B) schB & \
@@ -917,7 +917,7 @@
 
 
 
-goal thy "!! A B. [| compatible A B; compatible B A; \
+Goal "!! A B. [| compatible A B; compatible B A; \
 \ is_asig(asig_of A); is_asig(asig_of B) |] ==> \
 \ Forall (%x. x:ext (A||B)) tr & \
 \ Forall (%x. x:act A) schA & Forall (%x. x:act B) schB & \
@@ -1155,7 +1155,7 @@
          section"COMPOSITIONALITY on TRACE Level -- Main Theorem";
 (* ------------------------------------------------------------------ *)
  
-goal thy 
+Goal 
 "!! A B. [| is_trans_of A; is_trans_of B; compatible A B; compatible B A; \
 \           is_asig(asig_of A); is_asig(asig_of B)|] \
 \       ==>  tr: traces(A||B) = \
@@ -1218,7 +1218,7 @@
 (*                            For Modules                             *)
 (* ------------------------------------------------------------------ *)
 
-goalw thy [Traces_def,par_traces_def]
+Goalw [Traces_def,par_traces_def]
 
 "!! A B. [| is_trans_of A; is_trans_of B; compatible A B; compatible B A; \
 \           is_asig(asig_of A); is_asig(asig_of B)|] \
--- a/src/HOLCF/IOA/meta_theory/Compositionality.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/IOA/meta_theory/Compositionality.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -8,12 +8,12 @@
 
 
 
-goal thy "!! A. [|eA --> A ; eB & ~eA --> ~A|] ==> (eA | eB) --> A=eA";
+Goal "!! A. [|eA --> A ; eB & ~eA --> ~A|] ==> (eA | eB) --> A=eA";
 by Auto_tac;
 qed"compatibility_consequence3";
 
 
-goal thy 
+Goal 
 "!! A B. [| compatible A B; Forall (%a. a: ext A | a: ext B) tr |] ==> \
 \           Filter (%a. a: act A)`tr= Filter (%a. a: ext A)`tr";
 by (rtac ForallPFilterQR 1);
@@ -28,11 +28,11 @@
 (* the next two theorems are only necessary, as there is no theorem ext (A||B) = ext (B||A)
     or even better A||B= B||A, FIX *)
 
-goal thy "!! A. [|eA --> A ; eB & ~eA --> ~A|] ==> (eB | eA) --> A=eA";
+Goal "!! A. [|eA --> A ; eB & ~eA --> ~A|] ==> (eB | eA) --> A=eA";
 by Auto_tac;
 qed"compatibility_consequence4";
 
-goal thy 
+Goal 
 "!! A B. [| compatible A B; Forall (%a. a: ext B | a: ext A) tr |] ==> \
 \           Filter (%a. a: act A)`tr= Filter (%a. a: ext A)`tr";
 by (rtac ForallPFilterQR 1);
@@ -49,7 +49,7 @@
 
 
 
-goal thy "!! A1 A2 B1 B2. \
+Goal "!! A1 A2 B1 B2. \
 \         [| is_trans_of A1; is_trans_of A2; is_trans_of B1; is_trans_of B2;\
 \            is_asig_of A1; is_asig_of A2; \
 \            is_asig_of B1; is_asig_of B2; \
--- a/src/HOLCF/IOA/meta_theory/Deadlock.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/IOA/meta_theory/Deadlock.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -10,7 +10,7 @@
                input actions may always be added to a schedule
 **********************************************************************************)
 
-goal thy "!! sch. [| Filter (%x. x:act A)`sch : schedules A; a:inp A; input_enabled A; Finite sch|] \
+Goal "!! sch. [| Filter (%x. x:act A)`sch : schedules A; a:inp A; input_enabled A; Finite sch|] \
 \         ==> Filter (%x. x:act A)`sch @@ a>>nil : schedules A";
 by (asm_full_simp_tac (simpset() addsimps [schedules_def,has_schedule_def]) 1);
 by (safe_tac set_cs);
@@ -51,7 +51,7 @@
                     and distributivity of is_exec_frag over @@
 **********************************************************************************)
 Delsplits [split_if];
-goal thy "!! sch. [| a : local A; Finite sch; sch : schedules (A||B); \
+Goal "!! sch. [| a : local A; Finite sch; sch : schedules (A||B); \
 \            Filter (%x. x:act A)`(sch @@ a>>nil) : schedules A; compatible A B; input_enabled B |] \
 \          ==> (sch @@ a>>nil) : schedules (A||B)";
 
--- a/src/HOLCF/IOA/meta_theory/LiveIOA.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/IOA/meta_theory/LiveIOA.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -9,7 +9,7 @@
 
 Delsimps [split_paired_Ex];
 
-goalw thy [live_implements_def] 
+Goalw [live_implements_def] 
 "!! A. [| live_implements (A,LA) (B,LB); live_implements (B,LB) (C,LC) |] \
 \     ==> live_implements (A,LA) (C,LC)"; 
 auto();
@@ -24,7 +24,7 @@
 (* ---------------------------------------------------------------- *)
 
 
-goal thy "!! f. [| inp(C)=inp(A); out(C)=out(A); \
+Goal "!! f. [| inp(C)=inp(A); out(C)=out(A); \
 \                  is_live_ref_map f (C,M) (A,L) |] \
 \               ==> live_implements (C,M) (A,L)";
 
@@ -59,7 +59,7 @@
 
 (* Classical Fairness and Fairness by Temporal Formula coincide *)
  
-goal thy "!! ex. Finite (snd ex) ==> \
+Goal "!! ex. Finite (snd ex) ==> \
 \          (ex |== WF A acts) = (~ Enabled A acts (laststate ex))";
 
 In 3 steps:
--- a/src/HOLCF/IOA/meta_theory/RefCorrectness.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/IOA/meta_theory/RefCorrectness.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -18,7 +18,7 @@
 (* ---------------------------------------------------------------- *)
 
 
-goal thy "corresp_exC A f  = (LAM ex. (%s. case ex of \
+Goal "corresp_exC A f  = (LAM ex. (%s. case ex of \
 \      nil =>  nil   \
 \    | x##xs => (flift1 (%pr. (@cex. move A cex (f s) (fst pr) (f (snd pr)))   \
 \                              @@ ((corresp_exC A f `xs) (snd pr)))   \
@@ -30,17 +30,17 @@
 by (simp_tac (simpset() addsimps [flift1_def]) 1);
 qed"corresp_exC_unfold";
 
-goal thy "(corresp_exC A f`UU) s=UU";
+Goal "(corresp_exC A f`UU) s=UU";
 by (stac corresp_exC_unfold 1);
 by (Simp_tac 1);
 qed"corresp_exC_UU";
 
-goal thy "(corresp_exC A f`nil) s = nil";
+Goal "(corresp_exC A f`nil) s = nil";
 by (stac corresp_exC_unfold 1);
 by (Simp_tac 1);
 qed"corresp_exC_nil";
 
-goal thy "(corresp_exC A f`(at>>xs)) s = \
+Goal "(corresp_exC A f`(at>>xs)) s = \
 \          (@cex. move A cex (f s) (fst at) (f (snd at)))  \
 \          @@ ((corresp_exC A f`xs) (snd at))";
 by (rtac trans 1);
@@ -61,7 +61,7 @@
 
 section"properties of move";
 
-goalw thy [is_ref_map_def]
+Goalw [is_ref_map_def]
    "!!f. [|is_ref_map f C A; reachable C s; (s,a,t):trans_of C|] ==>\
 \     move A (@x. move A x (f s) a (f t)) (f s) a (f t)";
 
@@ -72,7 +72,7 @@
 by (assume_tac 1);
 qed"move_is_move";
 
-goal thy
+Goal
    "!!f. [|is_ref_map f C A; reachable C s; (s,a,t):trans_of C|] ==>\
 \    is_exec_frag A (f s,@x. move A x (f s) a (f t))";
 by (cut_inst_tac [] move_is_move 1);
@@ -80,7 +80,7 @@
 by (asm_full_simp_tac (simpset() addsimps [move_def]) 1);
 qed"move_subprop1";
 
-goal thy
+Goal
    "!!f. [|is_ref_map f C A; reachable C s; (s,a,t):trans_of C|] ==>\
 \    Finite ((@x. move A x (f s) a (f t)))";
 by (cut_inst_tac [] move_is_move 1);
@@ -88,7 +88,7 @@
 by (asm_full_simp_tac (simpset() addsimps [move_def]) 1);
 qed"move_subprop2";
 
-goal thy
+Goal
    "!!f. [|is_ref_map f C A; reachable C s; (s,a,t):trans_of C|] ==>\
 \    laststate (f s,@x. move A x (f s) a (f t)) = (f t)";
 by (cut_inst_tac [] move_is_move 1);
@@ -96,7 +96,7 @@
 by (asm_full_simp_tac (simpset() addsimps [move_def]) 1);
 qed"move_subprop3";
 
-goal thy
+Goal
    "!!f. [|is_ref_map f C A; reachable C s; (s,a,t):trans_of C|] ==>\
 \     mk_trace A`((@x. move A x (f s) a (f t))) = \
 \       (if a:ext A then a>>nil else nil)";
@@ -119,7 +119,7 @@
 (* --------------------------------------------------- *)
 
 
-goal thy "mk_trace C`(ex1 @@ ex2)= (mk_trace C`ex1) @@ (mk_trace C`ex2)";
+Goal "mk_trace C`(ex1 @@ ex2)= (mk_trace C`ex1) @@ (mk_trace C`ex2)";
 by (simp_tac (simpset() addsimps [mk_trace_def,filter_act_def,
                                  FilterConc,MapConc]) 1);
 qed"mk_traceConc";
@@ -130,7 +130,7 @@
                  Lemma 1 :Traces coincide  
    ------------------------------------------------------- *)
 Delsplits[split_if];
-goalw thy [corresp_ex_def]
+Goalw [corresp_ex_def]
   "!!f.[|is_ref_map f C A; ext C = ext A|] ==>  \     
 \        !s. reachable C s & is_exec_frag C (s,xs) --> \
 \            mk_trace C`xs = mk_trace A`(snd (corresp_ex A f (s,xs)))";
@@ -159,7 +159,7 @@
 (*                   Lemma 2.1                        *)
 (* -------------------------------------------------- *)
 
-goal thy 
+Goal 
 "Finite xs --> \
 \(!s .is_exec_frag A (s,xs) & is_exec_frag A (t,ys) & \ 
 \     t = laststate (s,xs) \
@@ -179,7 +179,7 @@
 
 
 
-goalw thy [corresp_ex_def]
+Goalw [corresp_ex_def]
  "!!f.[| is_ref_map f C A |] ==>\
 \ !s. reachable C s & is_exec_frag C (s,xs) \
 \ --> is_exec_frag A (corresp_ex A f (s,xs))"; 
@@ -220,7 +220,7 @@
 (* -------------------------------------------------------------------------------- *)
 
 
-goalw thy [traces_def]
+Goalw [traces_def]
   "!!f. [| ext C = ext A; is_ref_map f C A |] \
 \          ==> traces C <= traces A"; 
 
@@ -255,19 +255,19 @@
 (* -------------------------------------------------------------------------------- *)
 
 
-goalw thy [fin_often_def] "(~inf_often P s) = fin_often P s";
+Goalw [fin_often_def] "(~inf_often P s) = fin_often P s";
 auto();
 qed"fininf";
 
 
-goal thy 
+Goal 
 "is_wfair A W (s,ex) = \
 \ (fin_often (%x. ~Enabled A W (snd x)) ex --> inf_often (%x. fst x :W) ex)";
 by (asm_full_simp_tac (simpset() addsimps [is_wfair_def,fin_often_def])1);
 auto();
 qed"WF_alt";
 
-goal thy  
+Goal  
 "!! ex. [|is_wfair A W (s,ex); inf_often (%x. Enabled A W (snd x)) ex; \
 \         en_persistent A W|] \
 \   ==> inf_often (%x. fst x :W) ex";
@@ -278,7 +278,7 @@
 qed"WF_persistent";
 
 
-goal thy "!! C A. \
+Goal "!! C A. \
 \         [| is_ref_map f C A; ext C = ext A; \
 \         !! ex. [| ex:executions C; fair_ex C ex|] ==> fair_ex A (corresp_ex A f ex) |] \
 \         ==> fairtraces C <= fairtraces A";
@@ -308,7 +308,7 @@
 auto();
 qed"fair_trace_inclusion";
 
-goal thy "!! C A. \
+Goal "!! C A. \
 \         [| inp(C) = inp(A); out(C)=out(A); \
 \            is_fair_ref_map f C A |] \
 \         ==> fair_implements C A";
--- a/src/HOLCF/IOA/meta_theory/RefMappings.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/IOA/meta_theory/RefMappings.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -15,21 +15,21 @@
 section "transitions and moves";
 
 
-goal thy"!!f. s -a--A-> t ==> ? ex. move A ex s a t";
+Goal"!!f. s -a--A-> t ==> ? ex. move A ex s a t";
 
 by (res_inst_tac [("x","(a,t)>>nil")] exI 1);
 by (asm_full_simp_tac (simpset() addsimps [move_def]) 1);
 qed"transition_is_ex";
  
 
-goal thy"!!f. (~a:ext A) & s=t ==> ? ex. move A ex s a t";
+Goal"!!f. (~a:ext A) & s=t ==> ? ex. move A ex s a t";
 
 by (res_inst_tac [("x","nil")] exI 1);
 by (asm_full_simp_tac (simpset() addsimps [move_def]) 1);
 qed"nothing_is_ex";
 
 
-goal thy"!!f. (s -a--A-> s') & (s' -a'--A-> s'') & (~a':ext A) \
+Goal"!!f. (s -a--A-> s') & (s' -a'--A-> s'') & (~a':ext A) \
 \        ==> ? ex. move A ex s a s''";
 
 by (res_inst_tac [("x","(a,s')>>(a',s'')>>nil")] exI 1);
@@ -37,7 +37,7 @@
 qed"ei_transitions_are_ex";
 
 
-goal thy
+Goal
 "!!f. (s1 -a1--A-> s2) & (s2 -a2--A-> s3) & (s3 -a3--A-> s4) &\
 \     (~a2:ext A) & (~a3:ext A) ==> \
 \     ? ex. move A ex s1 a1 s4";
@@ -52,7 +52,7 @@
 section "weak_ref_map and ref_map";
 
 
-goalw thy [is_weak_ref_map_def,is_ref_map_def]
+Goalw [is_weak_ref_map_def,is_ref_map_def]
   "!!f. [| ext C = ext A; \
 \    is_weak_ref_map f C A |] ==> is_ref_map f C A";
 by (safe_tac set_cs);
@@ -69,7 +69,7 @@
 val lemma = result();
 
 Delsplits [split_if];
-goal thy "!!f.[| is_weak_ref_map f C A |]\
+Goal "!!f.[| is_weak_ref_map f C A |]\
 \                      ==> (is_weak_ref_map f (rename C g) (rename A g))";
 by (asm_full_simp_tac (simpset() addsimps [is_weak_ref_map_def]) 1);
 by (rtac conjI 1);
--- a/src/HOLCF/IOA/meta_theory/Seq.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/IOA/meta_theory/Seq.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -9,7 +9,7 @@
 
 
 (* Instead of adding UU_neq_nil every equation UU~=x could be changed to x~=UU *)
-goal thy "UU ~=nil";
+Goal "UU ~=nil";
 by (res_inst_tac [("s1","UU"),("t1","nil")]  (sym RS rev_contrapos) 1);
 by (REPEAT (Simp_tac 1));
 qed"UU_neq_nil";
@@ -36,17 +36,17 @@
  \                         nil  => nil \
  \                       | x##xs => f`x ## smap`f`xs)");
 
-goal thy "smap`f`nil=nil"; 
+Goal "smap`f`nil=nil"; 
 by (stac smap_unfold 1);
 by (Simp_tac 1);
 qed"smap_nil";
 
-goal thy "smap`f`UU=UU"; 
+Goal "smap`f`UU=UU"; 
 by (stac smap_unfold 1);
 by (Simp_tac 1);
 qed"smap_UU";
 
-goal thy 
+Goal 
 "!!x.[|x~=UU|] ==> smap`f`(x##xs)= (f`x)##smap`f`xs"; 
 by (rtac trans 1);
 by (stac smap_unfold 1);
@@ -63,17 +63,17 @@
  \         nil   => nil \
  \       | x##xs => If P`x then x##(sfilter`P`xs) else sfilter`P`xs fi)");
 
-goal thy "sfilter`P`nil=nil";
+Goal "sfilter`P`nil=nil";
 by (stac sfilter_unfold 1);
 by (Simp_tac 1);
 qed"sfilter_nil";
 
-goal thy "sfilter`P`UU=UU";
+Goal "sfilter`P`UU=UU";
 by (stac sfilter_unfold 1);
 by (Simp_tac 1);
 qed"sfilter_UU";
 
-goal thy 
+Goal 
 "!!x. x~=UU ==> sfilter`P`(x##xs)=   \
 \             (If P`x then x##(sfilter`P`xs) else sfilter`P`xs fi)";
 by (rtac trans 1);
@@ -91,17 +91,17 @@
  \                         nil   => TT \
  \                       | x##xs => (P`x andalso sforall2`P`xs))");
 
-goal thy "sforall2`P`nil=TT"; 
+Goal "sforall2`P`nil=TT"; 
 by (stac sforall2_unfold 1);
 by (Simp_tac 1);
 qed"sforall2_nil";
 
-goal thy "sforall2`P`UU=UU"; 
+Goal "sforall2`P`UU=UU"; 
 by (stac sforall2_unfold 1);
 by (Simp_tac 1);
 qed"sforall2_UU";
 
-goal thy 
+Goal 
 "!!x. x~=UU ==> sforall2`P`(x##xs)= ((P`x) andalso sforall2`P`xs)";
 by (rtac trans 1);
 by (stac sforall2_unfold 1);
@@ -120,17 +120,17 @@
  \                         nil   => nil \
  \                       | x##xs => (If P`x then x##(stakewhile`P`xs) else nil fi))");
 
-goal thy "stakewhile`P`nil=nil"; 
+Goal "stakewhile`P`nil=nil"; 
 by (stac stakewhile_unfold 1);
 by (Simp_tac 1);
 qed"stakewhile_nil";
 
-goal thy "stakewhile`P`UU=UU"; 
+Goal "stakewhile`P`UU=UU"; 
 by (stac stakewhile_unfold 1);
 by (Simp_tac 1);
 qed"stakewhile_UU";
 
-goal thy 
+Goal 
 "!!x. x~=UU ==> stakewhile`P`(x##xs) =   \
 \             (If P`x then x##(stakewhile`P`xs) else nil fi)";
 by (rtac trans 1);
@@ -149,17 +149,17 @@
  \                         nil   => nil \
  \                       | x##xs => (If P`x then sdropwhile`P`xs else tr fi))");
 
-goal thy "sdropwhile`P`nil=nil"; 
+Goal "sdropwhile`P`nil=nil"; 
 by (stac sdropwhile_unfold 1);
 by (Simp_tac 1);
 qed"sdropwhile_nil";
 
-goal thy "sdropwhile`P`UU=UU"; 
+Goal "sdropwhile`P`UU=UU"; 
 by (stac sdropwhile_unfold 1);
 by (Simp_tac 1);
 qed"sdropwhile_UU";
 
-goal thy 
+Goal 
 "!!x. x~=UU ==> sdropwhile`P`(x##xs) =   \
 \             (If P`x then sdropwhile`P`xs else x##xs fi)";
 by (rtac trans 1);
@@ -180,17 +180,17 @@
  \                       | x##xs => (If is_nil`xs then x else slast`xs fi))");
 
 
-goal thy "slast`nil=UU"; 
+Goal "slast`nil=UU"; 
 by (stac slast_unfold 1);
 by (Simp_tac 1);
 qed"slast_nil";
 
-goal thy "slast`UU=UU"; 
+Goal "slast`UU=UU"; 
 by (stac slast_unfold 1);
 by (Simp_tac 1);
 qed"slast_UU";
 
-goal thy 
+Goal 
 "!!x. x~=UU ==> slast`(x##xs)= (If is_nil`xs then x else slast`xs fi)";
 by (rtac trans 1);
 by (stac slast_unfold 1);
@@ -209,17 +209,17 @@
  \                       | x##xs => x ## (xs @@ t2))");
 
 
-goal thy "nil @@ y = y"; 
+Goal "nil @@ y = y"; 
 by (stac sconc_unfold 1);
 by (Simp_tac 1);
 qed"sconc_nil";
 
-goal thy "UU @@ y=UU"; 
+Goal "UU @@ y=UU"; 
 by (stac sconc_unfold 1);
 by (Simp_tac 1);
 qed"sconc_UU";
 
-goal thy "(x##xs) @@ y=x##(xs @@ y)";
+Goal "(x##xs) @@ y=x##(xs @@ y)";
 by (rtac trans 1);
 by (stac sconc_unfold 1);
 by (Asm_full_simp_tac 1);
@@ -239,17 +239,17 @@
  \                         nil   => nil \
  \                       | x##xs => x @@ sflat`xs)");
 
-goal thy "sflat`nil=nil"; 
+Goal "sflat`nil=nil"; 
 by (stac sflat_unfold 1);
 by (Simp_tac 1);
 qed"sflat_nil";
 
-goal thy "sflat`UU=UU"; 
+Goal "sflat`UU=UU"; 
 by (stac sflat_unfold 1);
 by (Simp_tac 1);
 qed"sflat_UU";
 
-goal thy "sflat`(x##xs)= x@@(sflat`xs)"; 
+Goal "sflat`(x##xs)= x@@(sflat`xs)"; 
 by (rtac trans 1);
 by (stac sflat_unfold 1);
 by (Asm_full_simp_tac 1);
@@ -271,30 +271,30 @@
 \                          nil => UU    \
 \                        | y##ys => <x,y>##(szip`xs`ys)))");
 
-goal thy "szip`nil`y=nil"; 
+Goal "szip`nil`y=nil"; 
 by (stac szip_unfold 1);
 by (Simp_tac 1);
 qed"szip_nil";
 
-goal thy "szip`UU`y=UU"; 
+Goal "szip`UU`y=UU"; 
 by (stac szip_unfold 1);
 by (Simp_tac 1);
 qed"szip_UU1";
 
-goal thy "!!x. x~=nil ==> szip`x`UU=UU"; 
+Goal "!!x. x~=nil ==> szip`x`UU=UU"; 
 by (stac szip_unfold 1);
 by (Asm_full_simp_tac 1);
 by (res_inst_tac [("x","x")] seq.casedist 1);
 by (REPEAT (Asm_full_simp_tac 1));
 qed"szip_UU2";
 
-goal thy "!!x. x~=UU ==> szip`(x##xs)`nil=UU";
+Goal "!!x. x~=UU ==> szip`(x##xs)`nil=UU";
 by (rtac trans 1);
 by (stac szip_unfold 1);
 by (REPEAT (Asm_full_simp_tac 1));
 qed"szip_cons_nil";
 
-goal thy "!!x.[| x~=UU; y~=UU|] ==> szip`(x##xs)`(y##ys) = <x,y>##szip`xs`ys";
+Goal "!!x.[| x~=UU; y~=UU|] ==> szip`(x##xs)`(y##ys) = <x,y>##szip`xs`ys";
 by (rtac trans 1);
 by (stac szip_unfold 1);
 by (REPEAT (Asm_full_simp_tac 1));
@@ -316,14 +316,14 @@
 
 section "scons, nil";
 
-goal thy 
+Goal 
  "!!x. [|x~=UU;y~=UU|]==> (x##xs=y##ys) = (x=y & xs=ys)";
 by (rtac iffI 1);
 by (etac (hd seq.injects) 1);
 by Auto_tac;
 qed"scons_inject_eq";
 
-goal thy "!!x. nil<<x ==> nil=x";
+Goal "!!x. nil<<x ==> nil=x";
 by (res_inst_tac [("x","x")] seq.casedist 1);
 by (hyp_subst_tac 1);
 by (etac antisym_less 1);
@@ -338,7 +338,7 @@
 section "sfilter, sforall, sconc";
 
 
-goal thy  "(if b then tr1 else tr2) @@ tr \
+Goal  "(if b then tr1 else tr2) @@ tr \
         \= (if b then tr1 @@ tr else tr2 @@ tr)";
 by (res_inst_tac [("P","b"),("Q","b")] impCE 1);
 by (fast_tac HOL_cs 1);
@@ -348,7 +348,7 @@
 Addsimps [if_and_sconc];
 
 
-goal thy "sfilter`P`(x @@ y) = (sfilter`P`x @@ sfilter`P`y)";
+Goal "sfilter`P`(x @@ y) = (sfilter`P`x @@ sfilter`P`y)";
 
 by (res_inst_tac[("x","x")] seq.ind 1);
 (* adm *)
@@ -360,7 +360,7 @@
 by (asm_full_simp_tac (simpset() setloop split_If_tac) 1);
 qed"sfiltersconc";
 
-goal thy "sforall P (stakewhile`P`x)";
+Goal "sforall P (stakewhile`P`x)";
 by (simp_tac (simpset() addsimps [sforall_def]) 1);
 by (res_inst_tac[("x","x")] seq.ind 1);
 (* adm *)
@@ -372,7 +372,7 @@
 by (asm_full_simp_tac (simpset() setloop split_If_tac) 1);
 qed"sforallPstakewhileP";
 
-goal thy "sforall P (sfilter`P`x)";
+Goal "sforall P (sfilter`P`x)";
 by (simp_tac (simpset() addsimps [sforall_def]) 1);
 by (res_inst_tac[("x","x")] seq.ind 1);
 (* adm *)
@@ -398,26 +398,26 @@
 (* 3. a~=UU==> Finite(a##x)=Finite(x)                  *)
 (* ----------------------------------------------------  *)
 
-goal thy "Finite x --> x~=UU";
+Goal "Finite x --> x~=UU";
 by (rtac impI 1);
 by (etac sfinite.induct 1);
  by (Asm_full_simp_tac 1);
 by (Asm_full_simp_tac 1);
 qed"Finite_UU_a";
 
-goal thy "~(Finite UU)";
+Goal "~(Finite UU)";
 by (cut_inst_tac [("x","UU")]Finite_UU_a  1);
 by (fast_tac HOL_cs 1);
 qed"Finite_UU";
 
-goal thy "Finite x --> a~=UU --> x=a##xs --> Finite xs";
+Goal "Finite x --> a~=UU --> x=a##xs --> Finite xs";
 by (strip_tac 1);
 by (etac sfinite.elim 1);
 by (fast_tac (HOL_cs addss simpset()) 1);
 by (fast_tac (HOL_cs addSDs seq.injects) 1);
 qed"Finite_cons_a";
 
-goal thy "!!x. a~=UU ==>(Finite (a##x)) = (Finite x)";
+Goal "!!x. a~=UU ==>(Finite (a##x)) = (Finite x)";
 by (rtac iffI 1);
 by (rtac (Finite_cons_a RS mp RS mp RS mp) 1);
 by (REPEAT (assume_tac 1));
@@ -449,7 +449,7 @@
         ]);
 
 
-goal  thy 
+Goal 
 "!!P.[|P(UU);P(nil);\
 \  !! x s1.[|x~=UU;P(s1)|] ==> P(x##s1)\
 \  |] ==> seq_finite(s) --> P(s)";
@@ -470,13 +470,13 @@
    Deshalb Finite jetzt induktiv und nicht mehr rekursiv definiert! 
    ------------------------------------------------------------------ *)
 
-goal thy "seq_finite nil";
+Goal "seq_finite nil";
 by (simp_tac (simpset() addsimps [seq.sfinite_def]) 1);
 by (res_inst_tac [("x","Suc 0")] exI 1);
 by (simp_tac (simpset() addsimps seq.rews) 1);
 qed"seq_finite_nil";
 
-goal thy "seq_finite UU";
+Goal "seq_finite UU";
 by (simp_tac (simpset() addsimps [seq.sfinite_def]) 1);
 qed"seq_finite_UU";
 
@@ -487,7 +487,7 @@
 qed"logic_lemma";
 
 
-goal thy "!!P.[| P nil; \
+Goal "!!P.[| P nil; \
 \                !!a t. [|a~=UU;Finite t; P t|] ==> P (a##t)|]\
 \            ==> Finite x --> P x";
 
@@ -500,7 +500,7 @@
 qed"Finite_ind";
 
 
-goal thy "Finite tr --> seq_finite tr";
+Goal "Finite tr --> seq_finite tr";
 by (rtac seq.ind 1);
 (* adm *)
 (* hier grosses Problem !!!!!!!!!!!!!!! *)
--- a/src/HOLCF/IOA/meta_theory/Sequence.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/IOA/meta_theory/Sequence.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -18,15 +18,15 @@
 (*                               Map                                *)
 (* ---------------------------------------------------------------- *)
 
-goal thy "Map f`UU =UU";
+Goal "Map f`UU =UU";
 by (simp_tac (simpset() addsimps [Map_def]) 1);
 qed"Map_UU";
 
-goal thy "Map f`nil =nil";
+Goal "Map f`nil =nil";
 by (simp_tac (simpset() addsimps [Map_def]) 1);
 qed"Map_nil";
 
-goal thy "Map f`(x>>xs)=(f x) >> Map f`xs";
+Goal "Map f`(x>>xs)=(f x) >> Map f`xs";
 by (simp_tac (simpset() addsimps [Map_def, Cons_def,flift2_def]) 1);
 qed"Map_cons";
 
@@ -34,15 +34,15 @@
 (*                               Filter                             *)
 (* ---------------------------------------------------------------- *)
 
-goal thy "Filter P`UU =UU";
+Goal "Filter P`UU =UU";
 by (simp_tac (simpset() addsimps [Filter_def]) 1);
 qed"Filter_UU";
 
-goal thy "Filter P`nil =nil";
+Goal "Filter P`nil =nil";
 by (simp_tac (simpset() addsimps [Filter_def]) 1);
 qed"Filter_nil";
 
-goal thy "Filter P`(x>>xs)= (if P x then x>>(Filter P`xs) else Filter P`xs)"; 
+Goal "Filter P`(x>>xs)= (if P x then x>>(Filter P`xs) else Filter P`xs)"; 
 by (simp_tac (simpset() addsimps [Filter_def, Cons_def,flift2_def,If_and_if]) 1);
 qed"Filter_cons";
 
@@ -50,15 +50,15 @@
 (*                               Forall                             *)
 (* ---------------------------------------------------------------- *)
 
-goal thy "Forall P UU";
+Goal "Forall P UU";
 by (simp_tac (simpset() addsimps [Forall_def,sforall_def]) 1);
 qed"Forall_UU";
 
-goal thy "Forall P nil";
+Goal "Forall P nil";
 by (simp_tac (simpset() addsimps [Forall_def,sforall_def]) 1);
 qed"Forall_nil";
 
-goal thy "Forall P (x>>xs)= (P x & Forall P xs)";
+Goal "Forall P (x>>xs)= (P x & Forall P xs)";
 by (simp_tac (simpset() addsimps [Forall_def, sforall_def,
                                  Cons_def,flift2_def]) 1);
 qed"Forall_cons";
@@ -68,7 +68,7 @@
 (* ---------------------------------------------------------------- *)
 
 
-goal thy "(x>>xs) @@ y = x>>(xs @@y)"; 
+Goal "(x>>xs) @@ y = x>>(xs @@y)"; 
 by (simp_tac (simpset() addsimps [Cons_def]) 1);
 qed"Conc_cons";
 
@@ -76,15 +76,15 @@
 (*                               Takewhile                          *)
 (* ---------------------------------------------------------------- *)
 
-goal thy "Takewhile P`UU =UU";
+Goal "Takewhile P`UU =UU";
 by (simp_tac (simpset() addsimps [Takewhile_def]) 1);
 qed"Takewhile_UU";
 
-goal thy "Takewhile P`nil =nil";
+Goal "Takewhile P`nil =nil";
 by (simp_tac (simpset() addsimps [Takewhile_def]) 1);
 qed"Takewhile_nil";
 
-goal thy "Takewhile P`(x>>xs)= (if P x then x>>(Takewhile P`xs) else nil)"; 
+Goal "Takewhile P`(x>>xs)= (if P x then x>>(Takewhile P`xs) else nil)"; 
 by (simp_tac (simpset() addsimps [Takewhile_def, Cons_def,flift2_def,If_and_if]) 1);
 qed"Takewhile_cons";
 
@@ -92,15 +92,15 @@
 (*                               Dropwhile                          *)
 (* ---------------------------------------------------------------- *)
 
-goal thy "Dropwhile P`UU =UU";
+Goal "Dropwhile P`UU =UU";
 by (simp_tac (simpset() addsimps [Dropwhile_def]) 1);
 qed"Dropwhile_UU";
 
-goal thy "Dropwhile P`nil =nil";
+Goal "Dropwhile P`nil =nil";
 by (simp_tac (simpset() addsimps [Dropwhile_def]) 1);
 qed"Dropwhile_nil";
 
-goal thy "Dropwhile P`(x>>xs)= (if P x then Dropwhile P`xs else x>>xs)"; 
+Goal "Dropwhile P`(x>>xs)= (if P x then Dropwhile P`xs else x>>xs)"; 
 by (simp_tac (simpset() addsimps [Dropwhile_def, Cons_def,flift2_def,If_and_if]) 1);
 qed"Dropwhile_cons";
 
@@ -109,15 +109,15 @@
 (* ---------------------------------------------------------------- *)
 
 
-goal thy "Last`UU =UU";
+Goal "Last`UU =UU";
 by (simp_tac (simpset() addsimps [Last_def]) 1);
 qed"Last_UU";
 
-goal thy "Last`nil =UU";
+Goal "Last`nil =UU";
 by (simp_tac (simpset() addsimps [Last_def]) 1);
 qed"Last_nil";
 
-goal thy "Last`(x>>xs)= (if xs=nil then Def x else Last`xs)"; 
+Goal "Last`(x>>xs)= (if xs=nil then Def x else Last`xs)"; 
 by (simp_tac (simpset() addsimps [Last_def, Cons_def]) 1);
 by (res_inst_tac [("x","xs")] seq.casedist 1);
 by (Asm_simp_tac 1);
@@ -129,15 +129,15 @@
 (*                               Flat                               *)
 (* ---------------------------------------------------------------- *)
 
-goal thy "Flat`UU =UU";
+Goal "Flat`UU =UU";
 by (simp_tac (simpset() addsimps [Flat_def]) 1);
 qed"Flat_UU";
 
-goal thy "Flat`nil =nil";
+Goal "Flat`nil =nil";
 by (simp_tac (simpset() addsimps [Flat_def]) 1);
 qed"Flat_nil";
 
-goal thy "Flat`(x##xs)= x @@ (Flat`xs)"; 
+Goal "Flat`(x##xs)= x @@ (Flat`xs)"; 
 by (simp_tac (simpset() addsimps [Flat_def, Cons_def]) 1);
 qed"Flat_cons";
 
@@ -146,7 +146,7 @@
 (*                               Zip                                *)
 (* ---------------------------------------------------------------- *)
 
-goal thy "Zip = (LAM t1 t2. case t1 of \
+Goal "Zip = (LAM t1 t2. case t1 of \
 \               nil   => nil \
 \             | x##xs => (case t2 of \ 
 \                          nil => UU  \
@@ -162,29 +162,29 @@
 by (Simp_tac 1);
 qed"Zip_unfold";
 
-goal thy "Zip`UU`y =UU";
+Goal "Zip`UU`y =UU";
 by (stac Zip_unfold 1);
 by (Simp_tac 1);
 qed"Zip_UU1";
 
-goal thy "!! x. x~=nil ==> Zip`x`UU =UU";
+Goal "!! x. x~=nil ==> Zip`x`UU =UU";
 by (stac Zip_unfold 1);
 by (Simp_tac 1);
 by (res_inst_tac [("x","x")] seq.casedist 1);
 by (REPEAT (Asm_full_simp_tac 1));
 qed"Zip_UU2";
 
-goal thy "Zip`nil`y =nil";
+Goal "Zip`nil`y =nil";
 by (stac Zip_unfold 1);
 by (Simp_tac 1);
 qed"Zip_nil";
 
-goal thy "Zip`(x>>xs)`nil= UU"; 
+Goal "Zip`(x>>xs)`nil= UU"; 
 by (stac Zip_unfold 1);
 by (simp_tac (simpset() addsimps [Cons_def]) 1);
 qed"Zip_cons_nil";
 
-goal thy "Zip`(x>>xs)`(y>>ys)= (x,y) >> Zip`xs`ys"; 
+Goal "Zip`(x>>xs)`(y>>ys)= (x,y) >> Zip`xs`ys"; 
 by (rtac trans 1);
 by (stac Zip_unfold 1);
 by (Simp_tac 1);
@@ -218,7 +218,7 @@
 
 Can Filter with HOL predicate directly be defined as fixpoint ?
 
-goal thy "Filter2 P = (LAM tr. case tr of  \
+Goal "Filter2 P = (LAM tr. case tr of  \
  \         nil   => nil \
  \       | x##xs => (case x of Undef => UU | Def y => \
 \                   (if P y then y>>(Filter2 P`xs) else Filter2 P`xs)))";
@@ -239,18 +239,18 @@
 
 section "Cons";
 
-goal thy "a>>s = (Def a)##s";
+Goal "a>>s = (Def a)##s";
 by (simp_tac (simpset() addsimps [Cons_def]) 1);
 qed"Cons_def2";
 
-goal thy "x = UU | x = nil | (? a s. x = a >> s)";
+Goal "x = UU | x = nil | (? a s. x = a >> s)";
 by (simp_tac (simpset() addsimps [Cons_def2]) 1);
 by (cut_facts_tac [seq.exhaust] 1);
 by (fast_tac (HOL_cs addDs [not_Undef_is_Def RS iffD1]) 1);
 qed"Seq_exhaust";
 
 
-goal thy "!!P. [| x = UU ==> P; x = nil ==> P; !!a s. x = a >> s  ==> P |] ==> P";
+Goal "!!P. [| x = UU ==> P; x = nil ==> P; !!a s. x = a >> s  ==> P |] ==> P";
 by (cut_inst_tac [("x","x")] Seq_exhaust 1);
 by (etac disjE 1);
 by (Asm_full_simp_tac 1);
@@ -268,36 +268,36 @@
                                              THEN Asm_full_simp_tac (i+1)
                                              THEN Asm_full_simp_tac i;
 
-goal thy "a>>s ~= UU";
+Goal "a>>s ~= UU";
 by (stac Cons_def2 1);
 by (resolve_tac seq.con_rews 1);
 by (rtac Def_not_UU 1);
 qed"Cons_not_UU";
 
 
-goal thy "~(a>>x) << UU";
+Goal "~(a>>x) << UU";
 by (rtac notI 1);
 by (dtac antisym_less 1);
 by (Simp_tac 1);
 by (asm_full_simp_tac (simpset() addsimps [Cons_not_UU]) 1);
 qed"Cons_not_less_UU";
 
-goal thy "~a>>s << nil";
+Goal "~a>>s << nil";
 by (stac Cons_def2 1);
 by (resolve_tac seq.rews 1);
 by (rtac Def_not_UU 1);
 qed"Cons_not_less_nil";
 
-goal thy "a>>s ~= nil";
+Goal "a>>s ~= nil";
 by (stac Cons_def2 1);
 by (resolve_tac seq.rews 1);
 qed"Cons_not_nil";
 
-goal thy "nil ~= a>>s";
+Goal "nil ~= a>>s";
 by (simp_tac (simpset() addsimps [Cons_def2]) 1);
 qed"Cons_not_nil2";
 
-goal thy "(a>>s = b>>t) = (a = b & s = t)";
+Goal "(a>>s = b>>t) = (a = b & s = t)";
 by (simp_tac (HOL_ss addsimps [Cons_def2]) 1);
 by (stac (hd lift.inject RS sym) 1);
 back(); back();
@@ -305,7 +305,7 @@
 by (REPEAT(rtac Def_not_UU 1));
 qed"Cons_inject_eq";
 
-goal thy "(a>>s<<b>>t) = (a = b & s<<t)";
+Goal "(a>>s<<b>>t) = (a = b & s<<t)";
 by (simp_tac (simpset() addsimps [Cons_def2]) 1);
 by (stac (Def_inject_less_eq RS sym) 1);
 back();
@@ -319,7 +319,7 @@
 by (etac monofun_cfun_arg 1);
 qed"Cons_inject_less_eq";
 
-goal thy "seq_take (Suc n)`(a>>x) = a>> (seq_take n`x)";
+Goal "seq_take (Suc n)`(a>>x) = a>> (seq_take n`x)";
 by (simp_tac (simpset() addsimps [Cons_def]) 1);
 qed"seq_take_Cons";
 
@@ -327,7 +327,7 @@
           Cons_not_UU,Cons_not_less_UU,Cons_not_less_nil,Cons_not_nil];
 
 (* Instead of adding UU_neq_Cons every equation UU~=x could be changed to x~=UU *)
-goal thy "UU ~= x>>xs";
+Goal "UU ~= x>>xs";
 by (res_inst_tac [("s1","UU"),("t1","x>>xs")]  (sym RS rev_contrapos) 1);
 by (REPEAT (Simp_tac 1));
 qed"UU_neq_Cons";
@@ -339,14 +339,14 @@
 
 section "induction";
 
-goal thy "!! P. [| adm P; P UU; P nil; !! a s. P s ==> P (a>>s)|] ==> P x";
+Goal "!! P. [| adm P; P UU; P nil; !! a s. P s ==> P (a>>s)|] ==> P x";
 by (etac seq.ind 1);
 by (REPEAT (atac 1));
 by (def_tac 1);
 by (asm_full_simp_tac (simpset() addsimps [Cons_def]) 1);
 qed"Seq_induct";
 
-goal thy "!! P.[|P UU;P nil; !! a s. P s ==> P(a>>s) |]  \
+Goal "!! P.[|P UU;P nil; !! a s. P s ==> P(a>>s) |]  \
 \               ==> seq_finite x --> P x";
 by (etac seq_finite_ind 1);
 by (REPEAT (atac 1));
@@ -354,7 +354,7 @@
 by (asm_full_simp_tac (simpset() addsimps [Cons_def]) 1);
 qed"Seq_FinitePartial_ind";
 
-goal thy "!! P.[| Finite x; P nil; !! a s. [| Finite s; P s|] ==> P (a>>s) |] ==> P x";
+Goal "!! P.[| Finite x; P nil; !! a s. [| Finite s; P s|] ==> P (a>>s) |] ==> P x";
 by (etac sfinite.induct 1);
 by (assume_tac 1);
 by (def_tac 1);
@@ -386,11 +386,11 @@
 
 section "HD,TL";
 
-goal thy "HD`(x>>y) = Def x";
+Goal "HD`(x>>y) = Def x";
 by (simp_tac (simpset() addsimps [Cons_def]) 1);
 qed"HD_Cons";
 
-goal thy "TL`(x>>y) = y";
+Goal "TL`(x>>y) = y";
 by (simp_tac (simpset() addsimps [Cons_def]) 1);
 qed"TL_Cons";
 
@@ -400,16 +400,16 @@
 
 section "Finite, Partial, Infinite";
 
-goal thy "Finite (a>>xs) = Finite xs";
+Goal "Finite (a>>xs) = Finite xs";
 by (simp_tac (simpset() addsimps [Cons_def2,Finite_cons]) 1);
 qed"Finite_Cons";
 
 Addsimps [Finite_Cons];
-goal thy "!! x. Finite (x::'a Seq) ==> Finite y --> Finite (x@@y)";
+Goal "!! x. Finite (x::'a Seq) ==> Finite y --> Finite (x@@y)";
 by (Seq_Finite_induct_tac 1);
 qed"FiniteConc_1";
 
-goal thy "!! z. Finite (z::'a Seq) ==> !x y. z= x@@y --> (Finite x & Finite y)";
+Goal "!! z. Finite (z::'a Seq) ==> !x y. z= x@@y --> (Finite x & Finite y)";
 by (Seq_Finite_induct_tac 1);
 (* nil*)
 by (strip_tac 1);
@@ -425,7 +425,7 @@
 by (Asm_full_simp_tac 1);
 qed"FiniteConc_2";
 
-goal thy "Finite(x@@y) = (Finite (x::'a Seq) & Finite y)";
+Goal "Finite(x@@y) = (Finite (x::'a Seq) & Finite y)";
 by (rtac iffI 1);
 by (etac (FiniteConc_2 RS spec RS spec RS mp) 1);
 by (rtac refl 1);
@@ -436,11 +436,11 @@
 Addsimps [FiniteConc];
 
 
-goal thy "!! s. Finite s ==> Finite (Map f`s)";
+Goal "!! s. Finite s ==> Finite (Map f`s)";
 by (Seq_Finite_induct_tac 1);
 qed"FiniteMap1";
 
-goal thy "!! s. Finite s ==> ! t. (s = Map f`t) --> Finite t";
+Goal "!! s. Finite s ==> ! t. (s = Map f`t) --> Finite t";
 by (Seq_Finite_induct_tac 1);
 by (strip_tac 1);
 by (Seq_case_simp_tac "t" 1);
@@ -451,7 +451,7 @@
 by (Asm_full_simp_tac 1);
 qed"FiniteMap2";
 
-goal thy "Finite (Map f`s) = Finite s";
+Goal "Finite (Map f`s) = Finite s";
 by Auto_tac;
 by (etac (FiniteMap2 RS spec RS mp) 1);
 by (rtac refl 1);
@@ -459,7 +459,7 @@
 qed"Map2Finite";
 
 
-goal thy "!! s. Finite s ==> Finite (Filter P`s)";
+Goal "!! s. Finite s ==> Finite (Filter P`s)";
 by (Seq_Finite_induct_tac 1);
 qed"FiniteFilter";
 
@@ -473,7 +473,7 @@
    Then the assumption that an _infinite_ chain exists (from admI2) is set to a contradiction 
    to Finite_flat *)
 
-goal thy "!! (x:: 'a Seq). Finite x ==> !y. Finite (y:: 'a Seq) & x<<y --> x=y";
+Goal "!! (x:: 'a Seq). Finite x ==> !y. Finite (y:: 'a Seq) & x<<y --> x=y";
 by (Seq_Finite_induct_tac 1);
 by (strip_tac 1);
 by (etac conjE 1);
@@ -485,7 +485,7 @@
 qed_spec_mp"Finite_flat";
 
 
-goal thy "adm(%(x:: 'a Seq).Finite x)";
+Goal "adm(%(x:: 'a Seq).Finite x)";
 by (rtac admI2 1);
 by (eres_inst_tac [("x","0")] allE 1);
 back();
@@ -506,15 +506,15 @@
 
 section "Conc";
 
-goal thy "!! x::'a Seq. Finite x ==> ((x @@ y) = (x @@ z)) = (y = z)";
+Goal "!! x::'a Seq. Finite x ==> ((x @@ y) = (x @@ z)) = (y = z)";
 by (Seq_Finite_induct_tac 1);
 qed"Conc_cong";
 
-goal thy "(x @@ y) @@ z = (x::'a Seq) @@ y @@ z";
+Goal "(x @@ y) @@ z = (x::'a Seq) @@ y @@ z";
 by (Seq_induct_tac "x" [] 1);
 qed"Conc_assoc";
 
-goal thy "s@@ nil = s";
+Goal "s@@ nil = s";
 by (res_inst_tac[("x","s")] seq.ind 1);
 by (Simp_tac 1);
 by (Simp_tac 1);
@@ -525,12 +525,12 @@
 Addsimps [nilConc];
 
 (* FIX: should be same as nil_is_Conc2 when all nils are turned to right side !! *)
-goal thy "(nil = x @@ y) = ((x::'a Seq)= nil & y = nil)";
+Goal "(nil = x @@ y) = ((x::'a Seq)= nil & y = nil)";
 by (Seq_case_simp_tac "x" 1);
 by Auto_tac;
 qed"nil_is_Conc";
 
-goal thy "(x @@ y = nil) = ((x::'a Seq)= nil & y = nil)";
+Goal "(x @@ y = nil) = ((x::'a Seq)= nil & y = nil)";
 by (Seq_case_simp_tac "x" 1);
 by Auto_tac;
 qed"nil_is_Conc2";
@@ -540,11 +540,11 @@
 
 section "Last";
 
-goal thy "!! s. Finite s ==> s~=nil --> Last`s~=UU";
+Goal "!! s. Finite s ==> s~=nil --> Last`s~=UU";
 by (Seq_Finite_induct_tac  1);
 qed"Finite_Last1";
 
-goal thy "!! s. Finite s ==> Last`s=UU --> s=nil";
+Goal "!! s. Finite s ==> Last`s=UU --> s=nil";
 by (Seq_Finite_induct_tac  1);
 by (fast_tac HOL_cs 1);
 qed"Finite_Last2";
@@ -556,11 +556,11 @@
 section "Filter, Conc";
 
 
-goal thy "Filter P`(Filter Q`s) = Filter (%x. P x & Q x)`s";
+Goal "Filter P`(Filter Q`s) = Filter (%x. P x & Q x)`s";
 by (Seq_induct_tac "s" [Filter_def] 1);
 qed"FilterPQ";
 
-goal thy "Filter P`(x @@ y) = (Filter P`x @@ Filter P`y)";
+Goal "Filter P`(x @@ y) = (Filter P`x @@ Filter P`y)";
 by (simp_tac (simpset() addsimps [Filter_def,sfiltersconc]) 1);
 qed"FilterConc";
 
@@ -568,24 +568,24 @@
 
 section "Map";
 
-goal thy "Map f`(Map g`s) = Map (f o g)`s";
+Goal "Map f`(Map g`s) = Map (f o g)`s";
 by (Seq_induct_tac "s" [] 1);
 qed"MapMap";
 
-goal thy "Map f`(x@@y) = (Map f`x) @@ (Map f`y)";
+Goal "Map f`(x@@y) = (Map f`x) @@ (Map f`y)";
 by (Seq_induct_tac "x" [] 1);
 qed"MapConc";
 
-goal thy "Filter P`(Map f`x) = Map f`(Filter (P o f)`x)";
+Goal "Filter P`(Map f`x) = Map f`(Filter (P o f)`x)";
 by (Seq_induct_tac "x" [] 1);
 qed"MapFilter";
 
-goal thy "nil = (Map f`s) --> s= nil";
+Goal "nil = (Map f`s) --> s= nil";
 by (Seq_case_simp_tac "s" 1);
 qed"nilMap";
 
 
-goal thy "Forall P (Map f`s) = Forall (P o f) s";
+Goal "Forall P (Map f`s) = Forall (P o f) s";
 by (Seq_induct_tac "s" [Forall_def,sforall_def] 1);
 qed"ForallMap";
 
@@ -597,30 +597,30 @@
 section "Forall";
 
 
-goal thy "Forall P ys & (! x. P x --> Q x) \
+Goal "Forall P ys & (! x. P x --> Q x) \
 \         --> Forall Q ys";
 by (Seq_induct_tac "ys" [Forall_def,sforall_def] 1);
 qed"ForallPForallQ1";
 
 bind_thm ("ForallPForallQ",impI RSN (2,allI RSN (2,conjI RS (ForallPForallQ1 RS mp))));
 
-goal thy "(Forall P x & Forall P y) --> Forall P (x @@ y)";
+Goal "(Forall P x & Forall P y) --> Forall P (x @@ y)";
 by (Seq_induct_tac "x" [Forall_def,sforall_def] 1);
 qed"Forall_Conc_impl";
 
-goal thy "!! x. Finite x ==> Forall P (x @@ y) = (Forall P x & Forall P y)";
+Goal "!! x. Finite x ==> Forall P (x @@ y) = (Forall P x & Forall P y)";
 by (Seq_Finite_induct_tac  1);
 qed"Forall_Conc";
 
 Addsimps [Forall_Conc];
 
-goal thy "Forall P s  --> Forall P (TL`s)";
+Goal "Forall P s  --> Forall P (TL`s)";
 by (Seq_induct_tac "s" [Forall_def,sforall_def] 1);
 qed"ForallTL1";
 
 bind_thm ("ForallTL",ForallTL1 RS mp);
 
-goal thy "Forall P s  --> Forall P (Dropwhile Q`s)";
+Goal "Forall P s  --> Forall P (Dropwhile Q`s)";
 by (Seq_induct_tac "s" [Forall_def,sforall_def] 1);
 qed"ForallDropwhile1";
 
@@ -629,7 +629,7 @@
 
 (* only admissible in t, not if done in s *)
 
-goal thy "! s. Forall P s --> t<<s --> Forall P t";
+Goal "! s. Forall P s --> t<<s --> Forall P t";
 by (Seq_induct_tac "t" [Forall_def,sforall_def] 1);
 by (strip_tac 1); 
 by (Seq_case_simp_tac "sa" 1);
@@ -640,12 +640,12 @@
 bind_thm ("Forall_prefixclosed",Forall_prefix RS spec RS mp RS mp);
 
 
-goal thy "!! h. [| Finite h; Forall P s; s= h @@ t |] ==> Forall P t";
+Goal "!! h. [| Finite h; Forall P s; s= h @@ t |] ==> Forall P t";
 by Auto_tac;
 qed"Forall_postfixclosed";
 
 
-goal thy "((! x. P x --> (Q x = R x)) & Forall P tr) --> Filter Q`tr = Filter R`tr";
+Goal "((! x. P x --> (Q x = R x)) & Forall P tr) --> Filter Q`tr = Filter R`tr";
 by (Seq_induct_tac "tr" [Forall_def,sforall_def] 1);
 qed"ForallPFilterQR1";
 
@@ -657,12 +657,12 @@
 section "Forall, Filter";
 
 
-goal thy "Forall P (Filter P`x)";
+Goal "Forall P (Filter P`x)";
 by (simp_tac (simpset() addsimps [Filter_def,Forall_def,forallPsfilterP]) 1);
 qed"ForallPFilterP";
 
 (* holds also in other direction, then equal to forallPfilterP *)
-goal thy "Forall P x --> Filter P`x = x";
+Goal "Forall P x --> Filter P`x = x";
 by (Seq_induct_tac "x" [Forall_def,sforall_def,Filter_def] 1);
 qed"ForallPFilterPid1";
 
@@ -670,7 +670,7 @@
 
 
 (* holds also in other direction *)
-goal thy "!! ys . Finite ys ==> \
+Goal "!! ys . Finite ys ==> \
 \   Forall (%x. ~P x) ys --> Filter P`ys = nil ";
 by (Seq_Finite_induct_tac 1);
 qed"ForallnPFilterPnil1";
@@ -679,7 +679,7 @@
 
 
 (* holds also in other direction *)
-goal thy   "!! P. ~Finite ys & Forall (%x. ~P x) ys \
+Goal   "!! P. ~Finite ys & Forall (%x. ~P x) ys \
 \                  --> Filter P`ys = UU ";
 by (Seq_induct_tac "ys" [Forall_def,sforall_def] 1);
 qed"ForallnPFilterPUU1";
@@ -689,7 +689,7 @@
 
 (* inverse of ForallnPFilterPnil *)
 
-goal thy "!! ys . Filter P`ys = nil --> \
+Goal "!! ys . Filter P`ys = nil --> \
 \   (Forall (%x. ~P x) ys & Finite ys)";
 by (res_inst_tac[("x","ys")] Seq_induct 1);
 (* adm *)
@@ -706,15 +706,15 @@
 
 (* inverse of ForallnPFilterPUU. proved by 2 lemmas because of adm problems *)
 
-goal thy "!! ys. Finite ys ==> Filter P`ys ~= UU";
+Goal "!! ys. Finite ys ==> Filter P`ys ~= UU";
 by (Seq_Finite_induct_tac 1);
 qed"FilterUU_nFinite_lemma1";
 
-goal thy "~ Forall (%x. ~P x) ys --> Filter P`ys ~= UU";
+Goal "~ Forall (%x. ~P x) ys --> Filter P`ys ~= UU";
 by (Seq_induct_tac "ys" [Forall_def,sforall_def] 1);
 qed"FilterUU_nFinite_lemma2";
 
-goal thy   "!! P. Filter P`ys = UU ==> \
+Goal   "!! P. Filter P`ys = UU ==> \
 \                (Forall (%x. ~P x) ys  & ~Finite ys)";
 by (rtac conjI 1);
 by (cut_inst_tac [] (FilterUU_nFinite_lemma2 RS mp COMP rev_contrapos) 1);
@@ -723,14 +723,14 @@
 qed"FilternPUUForallP";
 
 
-goal thy  "!! Q P.[| Forall Q ys; Finite ys; !!x. Q x ==> ~P x|] \
+Goal  "!! Q P.[| Forall Q ys; Finite ys; !!x. Q x ==> ~P x|] \
 \   ==> Filter P`ys = nil";
 by (etac ForallnPFilterPnil 1);
 by (etac ForallPForallQ 1);
 by Auto_tac;
 qed"ForallQFilterPnil";
 
-goal thy "!! Q P. [| ~Finite ys; Forall Q ys;  !!x. Q x ==> ~P x|] \
+Goal "!! Q P. [| ~Finite ys; Forall Q ys;  !!x. Q x ==> ~P x|] \
 \   ==> Filter P`ys = UU ";
 by (etac ForallnPFilterPUU 1);
 by (etac ForallPForallQ 1);
@@ -744,19 +744,19 @@
 section "Takewhile, Forall, Filter";
 
 
-goal thy "Forall P (Takewhile P`x)";
+Goal "Forall P (Takewhile P`x)";
 by (simp_tac (simpset() addsimps [Forall_def,Takewhile_def,sforallPstakewhileP]) 1);
 qed"ForallPTakewhileP";
 
 
-goal thy"!! P. [| !!x. Q x==> P x |] ==> Forall P (Takewhile Q`x)";
+Goal"!! P. [| !!x. Q x==> P x |] ==> Forall P (Takewhile Q`x)";
 by (rtac ForallPForallQ 1);
 by (rtac ForallPTakewhileP 1);
 by Auto_tac;
 qed"ForallPTakewhileQ";
 
 
-goal thy  "!! Q P.[| Finite (Takewhile Q`ys); !!x. Q x ==> ~P x |] \
+Goal  "!! Q P.[| Finite (Takewhile Q`ys); !!x. Q x ==> ~P x |] \
 \   ==> Filter P`(Takewhile Q`ys) = nil";
 by (etac ForallnPFilterPnil 1);
 by (rtac ForallPForallQ 1);
@@ -764,7 +764,7 @@
 by Auto_tac;
 qed"FilterPTakewhileQnil";
 
-goal thy "!! Q P. [| !!x. Q x ==> P x |] ==> \
+Goal "!! Q P. [| !!x. Q x ==> P x |] ==> \
 \            Filter P`(Takewhile Q`ys) = (Takewhile Q`ys)";
 by (rtac ForallPFilterPid 1);
 by (rtac ForallPForallQ 1);
@@ -775,28 +775,28 @@
 Addsimps [ForallPTakewhileP,ForallPTakewhileQ,
           FilterPTakewhileQnil,FilterPTakewhileQid];
 
-goal thy "Takewhile P`(Takewhile P`s) = Takewhile P`s";
+Goal "Takewhile P`(Takewhile P`s) = Takewhile P`s";
 by (Seq_induct_tac "s" [Forall_def,sforall_def] 1);
 qed"Takewhile_idempotent";
 
-goal thy "Forall P s --> Takewhile (%x. Q x | (~P x))`s = Takewhile Q`s";
+Goal "Forall P s --> Takewhile (%x. Q x | (~P x))`s = Takewhile Q`s";
 by (Seq_induct_tac "s" [Forall_def,sforall_def] 1);
 qed"ForallPTakewhileQnP";
 
-goal thy "Forall P s --> Dropwhile (%x. Q x | (~P x))`s = Dropwhile Q`s";
+Goal "Forall P s --> Dropwhile (%x. Q x | (~P x))`s = Dropwhile Q`s";
 by (Seq_induct_tac "s" [Forall_def,sforall_def] 1);
 qed"ForallPDropwhileQnP";
 
 Addsimps [ForallPTakewhileQnP RS mp, ForallPDropwhileQnP RS mp];
 
 
-goal thy "Forall P s --> Takewhile P`(s @@ t) = s @@ (Takewhile P`t)";
+Goal "Forall P s --> Takewhile P`(s @@ t) = s @@ (Takewhile P`t)";
 by (Seq_induct_tac "s" [Forall_def,sforall_def] 1);
 qed"TakewhileConc1";
 
 bind_thm("TakewhileConc",TakewhileConc1 RS mp);
 
-goal thy "!! s. Finite s ==> Forall P s --> Dropwhile P`(s @@ t) = Dropwhile P`t";
+Goal "!! s. Finite s ==> Forall P s --> Dropwhile P`(s @@ t) = Dropwhile P`t";
 by (Seq_Finite_induct_tac 1);
 qed"DropwhileConc1";
 
@@ -809,7 +809,7 @@
 section "coinductive characterizations of Filter";
 
 
-goal thy "HD`(Filter P`y) = Def x \
+Goal "HD`(Filter P`y) = Def x \
 \         --> y = ((Takewhile (%x. ~P x)`y) @@ (x >> TL`(Dropwhile (%a.~P a)`y)))  \
 \             & Finite (Takewhile (%x. ~ P x)`y)  & P x";
 
@@ -825,7 +825,7 @@
 by (Asm_full_simp_tac 1); 
 qed"divide_Seq_lemma";
 
-goal thy "!! x. (x>>xs) << Filter P`y  \
+Goal "!! x. (x>>xs) << Filter P`y  \
 \   ==> y = ((Takewhile (%a. ~ P a)`y) @@ (x >> TL`(Dropwhile (%a.~P a)`y))) \
 \      & Finite (Takewhile (%a. ~ P a)`y)  & P x";
 by (rtac (divide_Seq_lemma RS mp) 1);
@@ -834,14 +834,14 @@
 qed"divide_Seq";
 
  
-goal thy "~Forall P y --> (? x. HD`(Filter (%a. ~P a)`y) = Def x)";
+Goal "~Forall P y --> (? x. HD`(Filter (%a. ~P a)`y) = Def x)";
 (* Pay attention: is only admissible with chain-finite package to be added to 
         adm test *)
 by (Seq_induct_tac "y" [Forall_def,sforall_def] 1);
 qed"nForall_HDFilter";
 
 
-goal thy "!!y. ~Forall P y  \
+Goal "!!y. ~Forall P y  \
 \  ==> ? x. y= (Takewhile P`y @@ (x >> TL`(Dropwhile P`y))) & \
 \      Finite (Takewhile P`y) & (~ P x)";
 by (dtac (nForall_HDFilter RS mp) 1);
@@ -852,7 +852,7 @@
 qed"divide_Seq2";
 
 
-goal thy  "!! y. ~Forall P y \
+Goal  "!! y. ~Forall P y \
 \  ==> ? x bs rs. y= (bs @@ (x>>rs)) & Finite bs & Forall P bs & (~ P x)";
 by (cut_inst_tac [] divide_Seq2 1);
 (*Auto_tac no longer proves it*)
@@ -867,13 +867,13 @@
 
 section "take_lemma";
 
-goal thy "(!n. seq_take n`x = seq_take n`x') = (x = x')";
+Goal "(!n. seq_take n`x = seq_take n`x') = (x = x')";
 by (rtac iffI 1);
 by (resolve_tac seq.take_lemmas 1);
 by Auto_tac;
 qed"seq_take_lemma";
 
-goal thy 
+Goal 
 "  ! n. ((! k. k < n --> seq_take k`y1 = seq_take k`y2) \
 \   --> seq_take n`(x @@ (t>>y1)) =  seq_take n`(x @@ (t>>y2)))";
 by (Seq_induct_tac "x" [] 1);
@@ -885,7 +885,7 @@
 qed"take_reduction1";
 
 
-goal thy "!! n.[| x=y; s=t;!! k. k<n ==> seq_take k`y1 = seq_take k`y2|] \
+Goal "!! n.[| x=y; s=t;!! k. k<n ==> seq_take k`y1 = seq_take k`y2|] \
 \ ==> seq_take n`(x @@ (s>>y1)) =  seq_take n`(y @@ (t>>y2))";
 
 by (auto_tac (claset() addSIs [take_reduction1 RS spec RS mp],simpset()));
@@ -895,7 +895,7 @@
           take-lemma and take_reduction for << instead of = 
    ------------------------------------------------------------------ *)
 
-goal thy 
+Goal 
 "  ! n. ((! k. k < n --> seq_take k`y1 << seq_take k`y2) \
 \   --> seq_take n`(x @@ (t>>y1)) <<  seq_take n`(x @@ (t>>y2)))";
 by (Seq_induct_tac "x" [] 1);
@@ -907,7 +907,7 @@
 qed"take_reduction_less1";
 
 
-goal thy "!! n.[| x=y; s=t;!! k. k<n ==> seq_take k`y1 << seq_take k`y2|] \
+Goal "!! n.[| x=y; s=t;!! k. k<n ==> seq_take k`y1 << seq_take k`y2|] \
 \ ==> seq_take n`(x @@ (s>>y1)) <<  seq_take n`(y @@ (t>>y2))";
 by (auto_tac (claset() addSIs [take_reduction_less1 RS spec RS mp],simpset()));
 qed"take_reduction_less";
@@ -931,7 +931,7 @@
 qed"take_lemma_less1";
 
 
-goal thy "(!n. seq_take n`x << seq_take n`x') = (x << x')";
+Goal "(!n. seq_take n`x << seq_take n`x') = (x << x')";
 by (rtac iffI 1);
 by (rtac take_lemma_less1 1);
 by Auto_tac;
@@ -942,7 +942,7 @@
           take-lemma proof principles
    ------------------------------------------------------------------ *)
 
-goal thy "!! Q. [|!! s. [| Forall Q s; A s |] ==> (f s) = (g s) ; \
+Goal "!! Q. [|!! s. [| Forall Q s; A s |] ==> (f s) = (g s) ; \
 \           !! s1 s2 y. [| Forall Q s1; Finite s1; ~ Q y; A (s1 @@ y>>s2)|] \
 \                         ==> (f (s1 @@ y>>s2)) = (g (s1 @@ y>>s2)) |] \
 \              ==> A x --> (f x)=(g x)";
@@ -950,7 +950,7 @@
 by (auto_tac (claset() addSDs [divide_Seq3],simpset()));
 qed"take_lemma_principle1";
 
-goal thy "!! Q. [|!! s. [| Forall Q s; A s |] ==> (f s) = (g s) ; \
+Goal "!! Q. [|!! s. [| Forall Q s; A s |] ==> (f s) = (g s) ; \
 \           !! s1 s2 y. [| Forall Q s1; Finite s1; ~ Q y; A (s1 @@ y>>s2)|] \
 \                         ==> ! n. seq_take n`(f (s1 @@ y>>s2)) \
 \                                = seq_take n`(g (s1 @@ y>>s2)) |] \
@@ -970,7 +970,7 @@
          to be imbuilt into the rule, as induction has to be done early and the take lemma 
          has to be used in the trivial direction afterwards for the (Forall Q x) case.  *)
 
-goal thy 
+Goal 
 "!! Q. [|!! s. [| Forall Q s; A s |] ==> (f s) = (g s) ; \
 \        !! s1 s2 y n. [| ! t. A t --> seq_take n`(f t) = seq_take n`(g t);\
 \                         Forall Q s1; Finite s1; ~ Q y; A (s1 @@ y>>s2) |] \
@@ -992,7 +992,7 @@
 qed"take_lemma_induct";
 
 
-goal thy 
+Goal 
 "!! Q. [|!! s. [| Forall Q s; A s |] ==> (f s) = (g s) ; \
 \        !! s1 s2 y n. [| ! t m. m < n --> A t --> seq_take m`(f t) = seq_take m`(g t);\
 \                         Forall Q s1; Finite s1; ~ Q y; A (s1 @@ y>>s2) |] \
@@ -1044,7 +1044,7 @@
 end;
 
 
-goal thy 
+Goal 
 "!! Q. [|!! s h1 h2. [| Forall Q s; A s h1 h2|] ==> (f s h1 h2) = (g s h1 h2) ; \
 \  !! s1 s2 y n. [| ! t h1 h2 m. m < n --> (A t h1 h2) --> seq_take m`(f t h1 h2) = seq_take m`(g t h1 h2);\
 \                         Forall Q s1; Finite s1; ~ Q y; A (s1 @@ y>>s2) h1 h2|] \
@@ -1068,7 +1068,7 @@
 
 
 
-goal thy 
+Goal 
 "!! Q. [|!! s. Forall Q s ==> P ((f s) = (g s)) ; \
 \        !! s1 s2 y n. [| ! t m. m < n --> P (seq_take m`(f t) = seq_take m`(g t));\
 \                         Forall Q s1; Finite s1; ~ Q y|] \
@@ -1097,7 +1097,7 @@
 *)
 
 
-goal thy 
+Goal 
 "!! Q. [| A UU  ==> (f UU) = (g UU) ; \
 \         A nil ==> (f nil) = (g nil) ; \
 \         !! s y n. [| ! t. A t --> seq_take n`(f t) = seq_take n`(g t);\
@@ -1131,27 +1131,27 @@
 
 (* In general: How to do this case without the same adm problems 
    as for the entire proof ? *) 
-goal thy "Forall (%x.~(P x & Q x)) s \
+Goal "Forall (%x.~(P x & Q x)) s \
 \         --> Filter P`(Filter Q`s) =\
 \             Filter (%x. P x & Q x)`s";
 
 by (Seq_induct_tac "s" [Forall_def,sforall_def] 1);
 qed"Filter_lemma1";
 
-goal thy "!! s. Finite s ==>  \
+Goal "!! s. Finite s ==>  \
 \         (Forall (%x. (~P x) | (~ Q x)) s  \
 \         --> Filter P`(Filter Q`s) = nil)";
 by (Seq_Finite_induct_tac 1);
 qed"Filter_lemma2";
 
-goal thy "!! s. Finite s ==>  \
+Goal "!! s. Finite s ==>  \
 \         Forall (%x. (~P x) | (~ Q x)) s  \
 \         --> Filter (%x. P x & Q x)`s = nil";
 by (Seq_Finite_induct_tac 1);
 qed"Filter_lemma3";
 
 
-goal thy "Filter P`(Filter Q`s) = Filter (%x. P x & Q x)`s";
+Goal "Filter P`(Filter Q`s) = Filter (%x. P x & Q x)`s";
 by (res_inst_tac [("A1","%x. True") 
                  ,("Q1","%x.~(P x & Q x)"),("x1","s")]
                  (take_lemma_induct RS mp) 1);
@@ -1170,7 +1170,7 @@
 
 
 
-goal thy "Map f`(x@@y) = (Map f`x) @@ (Map f`y)";
+Goal "Map f`(x@@y) = (Map f`x) @@ (Map f`y)";
 by (res_inst_tac [("A1","%x. True"), ("x1","x")]
     (take_lemma_in_eq_out RS mp) 1);
 by Auto_tac;
--- a/src/HOLCF/IOA/meta_theory/ShortExecutions.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/IOA/meta_theory/ShortExecutions.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -34,17 +34,17 @@
 \      )\
 \    )");
 
-goal thy "oraclebuild P`sch`UU = UU";
+Goal "oraclebuild P`sch`UU = UU";
 by (stac oraclebuild_unfold 1);
 by (Simp_tac 1);
 qed"oraclebuild_UU";
 
-goal thy "oraclebuild P`sch`nil = nil";
+Goal "oraclebuild P`sch`nil = nil";
 by (stac oraclebuild_unfold 1);
 by (Simp_tac 1);
 qed"oraclebuild_nil";
 
-goal thy "oraclebuild P`s`(x>>t) = \
+Goal "oraclebuild P`s`(x>>t) = \
 \         (Takewhile (%a.~ P a)`s)   \
 \          @@ (x>>(oraclebuild P`(TL`(Dropwhile (%a.~ P a)`s))`t))";     
 by (rtac trans 1);
@@ -60,7 +60,7 @@
                    section "Cut rewrite rules";
 (* ---------------------------------------------------------------- *)
 
-goalw thy [Cut_def]
+Goalw [Cut_def]
 "!! s. [| Forall (%a.~ P a) s; Finite s|] \
 \           ==> Cut P s =nil";
 by (subgoal_tac "Filter P`s = nil" 1);
@@ -69,7 +69,7 @@
 by (REPEAT (atac 1));
 qed"Cut_nil";
 
-goalw thy [Cut_def]
+Goalw [Cut_def]
 "!! s. [| Forall (%a.~ P a) s; ~Finite s|] \
 \           ==> Cut P s =UU";
 by (subgoal_tac "Filter P`s= UU" 1);
@@ -78,7 +78,7 @@
 by (REPEAT (atac 1));
 qed"Cut_UU";
 
-goalw thy [Cut_def]
+Goalw [Cut_def]
 "!! s. [| P t;  Forall (%x.~ P x) ss; Finite ss|] \
 \           ==> Cut P (ss @@ (t>> rs)) \
 \                = ss @@ (t >> Cut P rs)";
@@ -93,7 +93,7 @@
 (* ---------------------------------------------------------------- *)
 
 
-goal thy "Filter P`s = Filter P`(Cut P s)";
+Goal "Filter P`s = Filter P`(Cut P s)";
 
 by (res_inst_tac [("A1","%x. True")
                  ,("Q1","%x.~ P x"), ("x1","s")]
@@ -110,7 +110,7 @@
 qed"FilterCut";
 
 
-goal thy "Cut P (Cut P s) = (Cut P s)";
+Goal "Cut P (Cut P s) = (Cut P s)";
 
 by (res_inst_tac [("A1","%x. True")
                  ,("Q1","%x.~ P x"), ("x1","s")]
@@ -128,7 +128,7 @@
 qed"Cut_idemp";
 
 
-goal thy "Map f`(Cut (P o f) s) = Cut P (Map f`s)";
+Goal "Map f`(Cut (P o f) s) = Cut P (Map f`s)";
 
 by (res_inst_tac [("A1","%x. True")
                  ,("Q1","%x.~ P (f x)"), ("x1","s")]
@@ -152,7 +152,7 @@
 qed"MapCut";
 
 
-goal thy "~Finite s --> Cut P s << s";
+Goal "~Finite s --> Cut P s << s";
 by (strip_tac 1);
 by (rtac (take_lemma_less RS iffD1) 1);
 by (strip_tac 1);
@@ -181,7 +181,7 @@
 
 (*
 
-goal thy "Finite s --> (? y. s = Cut P s @@ y)";
+Goal "Finite s --> (? y. s = Cut P s @@ y)";
 by (strip_tac 1);
 by (rtac exI 1);
 by (resolve_tac seq.take_lemmas 1);
@@ -209,7 +209,7 @@
 
 
 
-goal thy "!!ex .is_exec_frag A (s,ex) ==> is_exec_frag A (s,Cut P ex)";
+Goal "!!ex .is_exec_frag A (s,ex) ==> is_exec_frag A (s,Cut P ex)";
 by (case_tac "Finite ex" 1);
 by (cut_inst_tac [("s","ex"),("P","P")] Cut_prefixcl_Finite 1);
 by (assume_tac 1);
@@ -229,7 +229,7 @@
 
 
 
-goalw thy [schedules_def,has_schedule_def]
+Goalw [schedules_def,has_schedule_def]
  "!! sch. [|sch : schedules A ; tr = Filter (%a. a:ext A)`sch|] \
 \   ==> ? sch. sch : schedules A & \
 \              tr = Filter (%a. a:ext A)`sch &\
@@ -268,7 +268,7 @@
                    section "Further Cut lemmas";
 (* ---------------------------------------------------------------- *)
 
-goalw  thy [LastActExtsch_def]
+Goalw [LastActExtsch_def]
   "!! A. [| LastActExtsch A sch; Filter (%x. x:ext A)`sch = nil |] \
 \   ==> sch=nil";
 by (dtac FilternPnilForallP 1);
@@ -278,7 +278,7 @@
 by (Asm_full_simp_tac 1);
 qed"LastActExtimplnil";
 
-goalw  thy [LastActExtsch_def]
+Goalw [LastActExtsch_def]
   "!! A. [| LastActExtsch A sch; Filter (%x. x:ext A)`sch = UU |] \
 \   ==> sch=UU";
 by (dtac FilternPUUForallP 1);
--- a/src/HOLCF/IOA/meta_theory/SimCorrectness.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/IOA/meta_theory/SimCorrectness.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -18,7 +18,7 @@
 (* ---------------------------------------------------------------- *)
 
 
-goal thy "corresp_ex_simC A R  = (LAM ex. (%s. case ex of \
+Goal "corresp_ex_simC A R  = (LAM ex. (%s. case ex of \
 \      nil =>  nil   \
 \    | x##xs => (flift1 (%pr. let a = (fst pr); t = (snd pr); \
 \                                 T' = @t'. ? ex1. (t,t'):R & move A ex1 s a t' \
@@ -33,17 +33,17 @@
 by (simp_tac (simpset() addsimps [flift1_def]) 1);
 qed"corresp_ex_simC_unfold";
 
-goal thy "(corresp_ex_simC A R`UU) s=UU";
+Goal "(corresp_ex_simC A R`UU) s=UU";
 by (stac corresp_ex_simC_unfold 1);
 by (Simp_tac 1);
 qed"corresp_ex_simC_UU";
 
-goal thy "(corresp_ex_simC A R`nil) s = nil";
+Goal "(corresp_ex_simC A R`nil) s = nil";
 by (stac corresp_ex_simC_unfold 1);
 by (Simp_tac 1);
 qed"corresp_ex_simC_nil";
 
-goal thy "(corresp_ex_simC A R`((a,t)>>xs)) s = \
+Goal "(corresp_ex_simC A R`((a,t)>>xs)) s = \
 \          (let T' = @t'. ? ex1. (t,t'):R & move A ex1 s a t' \
 \           in  \
 \            (@cex. move A cex s a T')  \ 
@@ -68,7 +68,7 @@
 
 Delsimps [Let_def];
 
-goalw thy [is_simulation_def]
+Goalw [is_simulation_def]
    "!!f. [|is_simulation R C A; reachable C s; s -a--C-> t; (s,s'):R|] ==>\
 \     let T' = @t'. ? ex1. (t,t'):R & move A ex1 s' a t' in \
 \     (t,T'): R & move A (@ex2. move A ex2 s' a T') s' a T'";
@@ -98,7 +98,7 @@
 
 Addsimps [Let_def];
 
-goal thy
+Goal
    "!!f. [|is_simulation R C A; reachable C s; s-a--C-> t; (s,s'):R|] ==>\
 \   let T' = @t'. ? ex1. (t,t'):R & move A ex1 s' a t' in \
 \    is_exec_frag A (s',@x. move A x s' a T')";
@@ -107,7 +107,7 @@
 by (asm_full_simp_tac (simpset() addsimps [move_def,Let_def]) 1);
 qed"move_subprop1_sim";
 
-goal thy
+Goal
    "!!f. [|is_simulation R C A; reachable C s; s-a--C-> t; (s,s'):R|] ==>\
 \   let T' = @t'. ? ex1. (t,t'):R & move A ex1 s' a t' in \
 \   Finite (@x. move A x s' a T')";
@@ -116,7 +116,7 @@
 by (asm_full_simp_tac (simpset() addsimps [move_def,Let_def]) 1);
 qed"move_subprop2_sim";
 
-goal thy
+Goal
    "!!f. [|is_simulation R C A; reachable C s; s-a--C-> t; (s,s'):R|] ==>\
 \   let T' = @t'. ? ex1. (t,t'):R & move A ex1 s' a t' in \
 \    laststate (s',@x. move A x s' a T') = T'";
@@ -125,7 +125,7 @@
 by (asm_full_simp_tac (simpset() addsimps [move_def,Let_def]) 1);
 qed"move_subprop3_sim";
 
-goal thy
+Goal
    "!!f. [|is_simulation R C A; reachable C s; s-a--C-> t; (s,s'):R|] ==>\
 \   let T' = @t'. ? ex1. (t,t'):R & move A ex1 s' a t' in \
 \     mk_trace A`((@x. move A x s' a T')) = \
@@ -135,7 +135,7 @@
 by (asm_full_simp_tac (simpset() addsimps [move_def,Let_def]) 1);
 qed"move_subprop4_sim";
 
-goal thy
+Goal
    "!!f. [|is_simulation R C A; reachable C s; s-a--C-> t; (s,s'):R|] ==>\
 \   let T' = @t'. ? ex1. (t,t'):R & move A ex1 s' a t' in \
 \     (t,T'):R";
@@ -157,7 +157,7 @@
    ------------------------------------------------------- *)
 
 Delsplits[split_if];
-goal thy 
+Goal 
   "!!f.[|is_simulation R C A; ext C = ext A|] ==>  \     
 \        !s s'. reachable C s & is_exec_frag C (s,ex) & (s,s'): R --> \
 \            mk_trace C`ex = mk_trace A`((corresp_ex_simC A R`ex) s')";
@@ -187,7 +187,7 @@
 (* ----------------------------------------------------------- *)
 
 
-goal thy 
+Goal 
  "!!f.[| is_simulation R C A |] ==>\
 \ !s s'. reachable C s & is_exec_frag C (s,ex) & (s,s'):R  \
 \ --> is_exec_frag A (s',(corresp_ex_simC A R`ex) s')"; 
@@ -239,7 +239,7 @@
      traces_coincide_sim, the second for the start state case. 
      S':= @s'. (s,s'):R & s':starts_of A, where s:starts_of C  *)
 
-goal thy 
+Goal 
 "!!C. [| is_simulation R C A; s:starts_of C |] \
 \ ==> let S' = @s'. (s,s'):R & s':starts_of A in \
 \     (s,S'):R & S':starts_of A";
@@ -258,7 +258,7 @@
 bind_thm("sim_starts2",(rewrite_rule [Let_def] simulation_starts) RS conjunct2);
 
 
-goalw thy [traces_def]
+Goalw [traces_def]
   "!!f. [| ext C = ext A; is_simulation R C A |] \
 \          ==> traces C <= traces A"; 
 
--- a/src/HOLCF/IOA/meta_theory/Simulations.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/IOA/meta_theory/Simulations.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -8,17 +8,17 @@
 
 
 
-goal thy "(A~={}) = (? x. x:A)";
+Goal "(A~={}) = (? x. x:A)";
 by (safe_tac set_cs);
 auto();
 qed"set_non_empty";
 
-goal thy "(A Int B ~= {}) = (? x. x: A & x:B)";
+Goal "(A Int B ~= {}) = (? x. x: A & x:B)";
 by (simp_tac (simpset() addsimps [set_non_empty]) 1);
 qed"Int_non_empty";
 
 
-goalw thy [Image_def]
+Goalw [Image_def]
 "(R^^{x} Int S ~= {}) = (? y. (x,y):R & y:S)";
 by (simp_tac (simpset() addsimps [Int_non_empty]) 1);
 qed"Sim_start_convert";
@@ -26,7 +26,7 @@
 Addsimps [Sim_start_convert];
 
 
-goalw thy [is_ref_map_def,is_simulation_def]
+Goalw [is_ref_map_def,is_simulation_def]
 "!! f. is_ref_map f C A ==> is_simulation {p. (snd p) = f (fst p)} C A";
 (* start states *)
 by (Asm_full_simp_tac 1);
--- a/src/HOLCF/IOA/meta_theory/TL.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/IOA/meta_theory/TL.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -7,24 +7,24 @@
 *)   
 
 
-goal thy "[] <> (.~ P) = (.~ <> [] P)";
+Goal "[] <> (.~ P) = (.~ <> [] P)";
 br ext 1;
 by (simp_tac (simpset() addsimps [Diamond_def,NOT_def,Box_def])1);
 auto();
 qed"simple_try";
 
-goal thy "nil |= [] P";
+Goal "nil |= [] P";
 by (asm_full_simp_tac (simpset() addsimps [satisfies_def,
      Box_def,tsuffix_def,suffix_def,nil_is_Conc])1);
 qed"Boxnil";
 
-goal thy "~(nil |= <> P)";
+Goal "~(nil |= <> P)";
 by (simp_tac (simpset() addsimps [Diamond_def,satisfies_def,NOT_def])1);
 by (cut_inst_tac [] Boxnil 1);
 by (asm_full_simp_tac (simpset() addsimps [satisfies_def])1);
 qed"Diamondnil";
 
-goal thy "(<> F) s = (? s2. tsuffix s2 s & F s2)";
+Goal "(<> F) s = (? s2. tsuffix s2 s & F s2)";
 by (simp_tac (simpset() addsimps [Diamond_def,NOT_def,Box_def])1);
 qed"Diamond_def2";
 
@@ -36,13 +36,13 @@
 (*                 TLA Axiomatization by Merz                       *)
 (* ---------------------------------------------------------------- *)
 
-goal thy "suffix s s";
+Goal "suffix s s";
 by (simp_tac (simpset() addsimps [suffix_def])1);
 by (res_inst_tac [("x","nil")] exI 1);
 auto();
 qed"suffix_refl";
 
-goal thy "s~=UU & s~=nil --> (s |= [] F .--> F)";
+Goal "s~=UU & s~=nil --> (s |= [] F .--> F)";
 by (simp_tac (simpset() addsimps [satisfies_def,IMPLIES_def,Box_def])1);
 by (REPEAT (rtac impI 1));
 by (eres_inst_tac [("x","s")] allE 1);
@@ -50,7 +50,7 @@
 qed"reflT";
 
 
-goal thy "!!x. [| suffix y x ; suffix z y |]  ==> suffix z x";
+Goal "!!x. [| suffix y x ; suffix z y |]  ==> suffix z x";
 by (asm_full_simp_tac (simpset() addsimps [suffix_def])1);
 auto();
 by (res_inst_tac [("x","s1 @@ s1a")] exI 1);
@@ -58,7 +58,7 @@
 by (simp_tac (simpset() addsimps [Conc_assoc]) 1); 
 qed"suffix_trans";
 
-goal thy "s |= [] F .--> [] [] F";
+Goal "s |= [] F .--> [] [] F";
 by (simp_tac (simpset() addsimps [satisfies_def,IMPLIES_def,Box_def,tsuffix_def])1);
 auto();
 bd suffix_trans 1;
@@ -68,13 +68,13 @@
 qed"transT";
 
 
-goal thy "s |= [] (F .--> G) .--> [] F .--> [] G";
+Goal "s |= [] (F .--> G) .--> [] F .--> [] G";
 by (simp_tac (simpset() addsimps [satisfies_def,IMPLIES_def,Box_def])1);
 qed"normalT";
 
 
 (*
-goal thy "s |= <> F .& <> G .--> (<> (F .& <> G) .| <> (G .& <> F))";
+Goal "s |= <> F .& <> G .--> (<> (F .& <> G) .| <> (G .& <> F))";
 by (simp_tac (simpset() addsimps [satisfies_def,IMPLIES_def,AND_def,OR_def,Diamond_def2])1);
 br impI 1;
 be conjE 1;
@@ -84,7 +84,7 @@
 
 br disjI1 1;
 
-goal thy "!!s. [| tsuffix s1 s ; tsuffix s2 s|] ==> tsuffix s2 s1 | tsuffix s1 s2";
+Goal "!!s. [| tsuffix s1 s ; tsuffix s2 s|] ==> tsuffix s2 s1 | tsuffix s1 s2";
 by (asm_full_simp_tac (simpset() addsimps [tsuffix_def,suffix_def])1);
 by (REPEAT (etac conjE 1));
 by (REPEAT (etac exE 1));
@@ -100,15 +100,15 @@
 (*                      TLA Rules by Lamport                        *)
 (* ---------------------------------------------------------------- *)
 
-goal thy "!! P. validT P ==> validT ([] P)";
+Goal "!! P. validT P ==> validT ([] P)";
 by (asm_full_simp_tac (simpset() addsimps [validT_def,satisfies_def,Box_def,tsuffix_def])1);
 qed"STL1a";
 
-goal thy "!! P. valid P ==> validT (Init P)";
+Goal "!! P. valid P ==> validT (Init P)";
 by (asm_full_simp_tac (simpset() addsimps [valid_def,validT_def,satisfies_def,Init_def])1);
 qed"STL1b";
 
-goal thy "!! P. valid P ==> validT ([] (Init P))";
+Goal "!! P. valid P ==> validT ([] (Init P))";
 br STL1a 1;
 be STL1b 1;
 qed"STL1";
@@ -116,7 +116,7 @@
 
 
 (* Note that unlift and HD is not at all used !!! *)
-goal thy "!! P. valid (P .--> Q)  ==> validT ([] (Init P) .--> [] (Init Q))";
+Goal "!! P. valid (P .--> Q)  ==> validT ([] (Init P) .--> [] (Init Q))";
 by (asm_full_simp_tac (simpset() addsimps [valid_def,validT_def,satisfies_def,IMPLIES_def,Box_def,Init_def])1);
 qed"STL4";
 
@@ -128,7 +128,7 @@
 (* ---------------------------------------------------------------- *)
 
 
-goalw thy [tsuffix_def,suffix_def]
+Goalw [tsuffix_def,suffix_def]
 "s~=UU & s~=nil --> tsuffix s2 (TL`s) --> tsuffix s2 s";
 auto();
 by (Seq_case_simp_tac "s" 1);
@@ -139,7 +139,7 @@
 val tsuffix_TL2 = conjI RS tsuffix_TL;
 
 Delsplits[split_if];
-goalw thy [Next_def,satisfies_def,NOT_def,IMPLIES_def,AND_def,Box_def] 
+Goalw [Next_def,satisfies_def,NOT_def,IMPLIES_def,AND_def,Box_def] 
    "s~=UU & s~=nil --> (s |= [] F .--> (F .& (Next ([] F))))";
 auto();
 (* []F .--> F *)
@@ -155,22 +155,22 @@
 Addsplits[split_if];
 
 
-goalw thy [Next_def,satisfies_def,NOT_def,IMPLIES_def] 
+Goalw [Next_def,satisfies_def,NOT_def,IMPLIES_def] 
     "s |= .~ (Next F) .--> (Next (.~ F))";
 by (Asm_full_simp_tac 1);
 qed"LTL2a";
 
-goalw thy [Next_def,satisfies_def,NOT_def,IMPLIES_def] 
+Goalw [Next_def,satisfies_def,NOT_def,IMPLIES_def] 
     "s |= (Next (.~ F)) .--> (.~ (Next F))";
 by (Asm_full_simp_tac 1);
 qed"LTL2b";
 
-goalw thy [Next_def,satisfies_def,NOT_def,IMPLIES_def] 
+Goalw [Next_def,satisfies_def,NOT_def,IMPLIES_def] 
 "ex |= (Next (F .--> G)) .--> (Next F) .--> (Next G)";
 by (Asm_full_simp_tac 1);
 qed"LTL3";
 
-goalw thy [Next_def,satisfies_def,Box_def,NOT_def,IMPLIES_def] 
+Goalw [Next_def,satisfies_def,Box_def,NOT_def,IMPLIES_def] 
  "s |= [] (F .--> Next F) .--> F .--> []F";
 by (Asm_full_simp_tac 1);
 auto();
@@ -182,11 +182,11 @@
 
 
 
-goal thy "!! P. [| validT (P .--> Q); validT P |] ==> validT Q";
+Goal "!! P. [| validT (P .--> Q); validT P |] ==> validT Q";
 by (asm_full_simp_tac (simpset() addsimps [validT_def,satisfies_def,IMPLIES_def])1);
 qed"ModusPonens";
 
 (* works only if validT is defined without restriction to s~=UU, s~=nil *)
-goal thy "!! P. validT P ==> validT (Next P)";
+Goal "!! P. validT P ==> validT (Next P)";
 by (asm_full_simp_tac (simpset() addsimps [validT_def,satisfies_def,Next_def])1);
 (* qed"NextTauto"; *)
--- a/src/HOLCF/IOA/meta_theory/TLS.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/IOA/meta_theory/TLS.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -17,7 +17,7 @@
 (*                                 ex2seqC                          *)
 (* ---------------------------------------------------------------- *)
 
-goal thy "ex2seqC  = (LAM ex. (%s. case ex of \
+Goal "ex2seqC  = (LAM ex. (%s. case ex of \
 \      nil =>  (s,None,s)>>nil   \
 \    | x##xs => (flift1 (%pr. \
 \                (s,Some (fst pr), snd pr)>> (ex2seqC`xs) (snd pr))  \
@@ -30,17 +30,17 @@
 by (simp_tac (simpset() addsimps [flift1_def]) 1);
 qed"ex2seqC_unfold";
 
-goal thy "(ex2seqC `UU) s=UU";
+Goal "(ex2seqC `UU) s=UU";
 by (stac ex2seqC_unfold 1);
 by (Simp_tac 1);
 qed"ex2seqC_UU";
 
-goal thy "(ex2seqC `nil) s = (s,None,s)>>nil";
+Goal "(ex2seqC `nil) s = (s,None,s)>>nil";
 by (stac ex2seqC_unfold 1);
 by (Simp_tac 1);
 qed"ex2seqC_nil";
 
-goal thy "(ex2seqC `((a,t)>>xs)) s = \
+Goal "(ex2seqC `((a,t)>>xs)) s = \
 \          (s,Some a,t)>> ((ex2seqC`xs) t)";
 by (rtac trans 1);
 by (stac ex2seqC_unfold 1);
@@ -51,15 +51,15 @@
 Addsimps [ex2seqC_UU,ex2seqC_nil,ex2seqC_cons];
 
 
-goal thy "ex2seq (s, UU) = UU";
+Goal "ex2seq (s, UU) = UU";
 by (simp_tac (simpset() addsimps [ex2seq_def]) 1);
 qed"ex2seq_UU";
 
-goal thy "ex2seq (s, nil) = (s,None,s)>>nil";
+Goal "ex2seq (s, nil) = (s,None,s)>>nil";
 by (simp_tac (simpset() addsimps [ex2seq_def]) 1);
 qed"ex2seq_nil";
 
-goal thy "ex2seq (s, (a,t)>>ex) = (s,Some a,t) >> ex2seq (t, ex)";
+Goal "ex2seq (s, (a,t)>>ex) = (s,Some a,t) >> ex2seq (t, ex)";
 by (simp_tac (simpset() addsimps [ex2seq_def]) 1);
 qed"ex2seq_cons";
 
@@ -69,17 +69,17 @@
 
 
 (* FIX: Not true for UU, as ex2seq is defined continously !!!!! *)
-goal thy "ex2seq exec ~= UU & ex2seq exec ~= nil";
+Goal "ex2seq exec ~= UU & ex2seq exec ~= nil";
 
 
-goal thy "ex |== [] P .--> P";
+Goal "ex |== [] P .--> P";
 
 
 (* ----------------------------------------------------------- *)
 (*           Interface TL -- TLS                               *)
 (* ---------------------------------------------------------- *)
 
-goalw thy [Init_def,Next_def,temp_sat_def,satisfies_def,IMPLIES_def,AND_def]
+Goalw [Init_def,Next_def,temp_sat_def,satisfies_def,IMPLIES_def,AND_def]
  "!! s. (P s) & s-a--A-> t --> (Q t) \
 \  ==> ex |== (Init (%(s,a,t). P s) .& Init (%(s,a,t). s -a--A-> t) \
 \             .--> (Next (Init (%(s,a,t).Q s))))";
--- a/src/HOLCF/IOA/meta_theory/Traces.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/IOA/meta_theory/Traces.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -26,15 +26,15 @@
 (* ---------------------------------------------------------------- *)
 
 
-goal thy  "filter_act`UU = UU";
+Goal  "filter_act`UU = UU";
 by (simp_tac (simpset() addsimps [filter_act_def]) 1);
 qed"filter_act_UU";
 
-goal thy  "filter_act`nil = nil";
+Goal  "filter_act`nil = nil";
 by (simp_tac (simpset() addsimps [filter_act_def]) 1);
 qed"filter_act_nil";
 
-goal thy "filter_act`(x>>xs) = (fst x) >> filter_act`xs";
+Goal "filter_act`(x>>xs) = (fst x) >> filter_act`xs";
 by (simp_tac (simpset() addsimps [filter_act_def]) 1);
 qed"filter_act_cons";
 
@@ -45,15 +45,15 @@
 (*                             mk_trace                             *)
 (* ---------------------------------------------------------------- *)
 
-goal thy "mk_trace A`UU=UU";
+Goal "mk_trace A`UU=UU";
 by (simp_tac (simpset() addsimps [mk_trace_def]) 1);
 qed"mk_trace_UU";
 
-goal thy "mk_trace A`nil=nil";
+Goal "mk_trace A`nil=nil";
 by (simp_tac (simpset() addsimps [mk_trace_def]) 1);
 qed"mk_trace_nil";
 
-goal thy "mk_trace A`(at >> xs) =    \
+Goal "mk_trace A`(at >> xs) =    \
 \            (if ((fst at):ext A)    \       
 \                 then (fst at) >> (mk_trace A`xs) \   
 \                 else mk_trace A`xs)";
@@ -68,7 +68,7 @@
 (* ---------------------------------------------------------------- *)
 
 
-goal thy "is_exec_fragC A = (LAM ex. (%s. case ex of \
+Goal "is_exec_fragC A = (LAM ex. (%s. case ex of \
 \      nil => TT \
 \    | x##xs => (flift1 \ 
 \            (%p. Def ((s,p):trans_of A) andalso (is_exec_fragC A`xs) (snd p)) \
@@ -81,17 +81,17 @@
 by (simp_tac (simpset() addsimps [flift1_def]) 1);
 qed"is_exec_fragC_unfold";
 
-goal thy "(is_exec_fragC A`UU) s=UU";
+Goal "(is_exec_fragC A`UU) s=UU";
 by (stac is_exec_fragC_unfold 1);
 by (Simp_tac 1);
 qed"is_exec_fragC_UU";
 
-goal thy "(is_exec_fragC A`nil) s = TT";
+Goal "(is_exec_fragC A`nil) s = TT";
 by (stac is_exec_fragC_unfold 1);
 by (Simp_tac 1);
 qed"is_exec_fragC_nil";
 
-goal thy "(is_exec_fragC A`(pr>>xs)) s = \
+Goal "(is_exec_fragC A`(pr>>xs)) s = \
 \                        (Def ((s,pr):trans_of A) \
 \                andalso (is_exec_fragC A`xs)(snd pr))";
 by (rtac trans 1);
@@ -108,15 +108,15 @@
 (*                        is_exec_frag                              *)
 (* ---------------------------------------------------------------- *)
 
-goal thy "is_exec_frag A (s, UU)";
+Goal "is_exec_frag A (s, UU)";
 by (simp_tac (simpset() addsimps [is_exec_frag_def]) 1);
 qed"is_exec_frag_UU";
 
-goal thy "is_exec_frag A (s, nil)";
+Goal "is_exec_frag A (s, nil)";
 by (simp_tac (simpset() addsimps [is_exec_frag_def]) 1);
 qed"is_exec_frag_nil";
 
-goal thy "is_exec_frag A (s, (a,t)>>ex) = \
+Goal "is_exec_frag A (s, (a,t)>>ex) = \
 \                               (((s,a,t):trans_of A) & \
 \                               is_exec_frag A (t, ex))";
 by (simp_tac (simpset() addsimps [is_exec_frag_def]) 1);
@@ -130,15 +130,15 @@
                            section "laststate";
 (* ---------------------------------------------------------------------------- *)
 
-goal thy "laststate (s,UU) = s";
+Goal "laststate (s,UU) = s";
 by (simp_tac (simpset() addsimps [laststate_def]) 1); 
 qed"laststate_UU";
 
-goal thy "laststate (s,nil) = s";
+Goal "laststate (s,nil) = s";
 by (simp_tac (simpset() addsimps [laststate_def]) 1);
 qed"laststate_nil";
 
-goal thy "!! ex. Finite ex ==> laststate (s,at>>ex) = laststate (snd at,ex)";
+Goal "!! ex. Finite ex ==> laststate (s,at>>ex) = laststate (snd at,ex)";
 by (simp_tac (simpset() addsimps [laststate_def]) 1);
 by (case_tac "ex=nil" 1);
 by (Asm_simp_tac 1);
@@ -150,7 +150,7 @@
 
 Addsimps [laststate_UU,laststate_nil,laststate_cons];
 
-goal thy "!!ex. Finite ex ==> (! s. ? u. laststate (s,ex)=u)";
+Goal "!!ex. Finite ex ==> (! s. ? u. laststate (s,ex)=u)";
 by (Seq_Finite_induct_tac 1);
 qed"exists_laststate";
 
@@ -162,7 +162,7 @@
 (* alternative definition of has_trace tailored for the refinement proof, as it does not 
    take the detour of schedules *)
 
-goalw thy  [executions_def,mk_trace_def,has_trace_def,schedules_def,has_schedule_def] 
+Goalw  [executions_def,mk_trace_def,has_trace_def,schedules_def,has_schedule_def] 
 "has_trace A b = (? ex:executions A. b = mk_trace A`(snd ex))";
 
 by (safe_tac set_cs);
@@ -193,7 +193,7 @@
    For executions of parallel automata this assumption is not needed, as in par_def
    this condition is included once more. (see Lemmas 1.1.1c in CompoExecs for example) *)
 
-goal thy 
+Goal 
   "!! A. is_trans_of A ==> \
 \ ! s. is_exec_frag A (s,xs) --> Forall (%a. a:act A) (filter_act`xs)";
 
@@ -204,7 +204,7 @@
 by (REPEAT (asm_full_simp_tac (simpset() addsimps [is_trans_of_def]) 1));
 qed"execfrag_in_sig";
 
-goal thy 
+Goal 
   "!! A.[|  is_trans_of A; x:executions A |] ==> \
 \ Forall (%a. a:act A) (filter_act`(snd x))";
 
@@ -214,7 +214,7 @@
 by Auto_tac;
 qed"exec_in_sig";
 
-goalw thy [schedules_def,has_schedule_def]
+Goalw [schedules_def,has_schedule_def]
   "!! A.[|  is_trans_of A; x:schedules A |] ==> \
 \   Forall (%a. a:act A) x";
 
@@ -225,7 +225,7 @@
 
 is ok but needs ForallQFilterP which has to been proven first (is trivial also)
 
-goalw thy [traces_def,has_trace_def]
+Goalw [traces_def,has_trace_def]
   "!! A.[| x:traces A |] ==> \
 \   Forall (%a. a:act A) x";
  by (safe_tac set_cs );
@@ -240,7 +240,7 @@
 section "executions are prefix closed";
 
 (* only admissible in y, not if done in x !! *)
-goal thy "!x s. is_exec_frag A (s,x) & y<<x  --> is_exec_frag A (s,y)";
+Goal "!x s. is_exec_frag A (s,x) & y<<x  --> is_exec_frag A (s,y)";
 by (pair_induct_tac "y" [is_exec_frag_def] 1);
 by (strip_tac 1);
 by (Seq_case_simp_tac "xa" 1);
@@ -253,7 +253,7 @@
 
 (* second prefix notion for Finite x *)
 
-goal thy "! y s. is_exec_frag A (s,x@@y) --> is_exec_frag A (s,x)";
+Goal "! y s. is_exec_frag A (s,x@@y) --> is_exec_frag A (s,x)";
 by (pair_induct_tac "x" [is_exec_frag_def] 1);
 by (strip_tac 1);
 by (Seq_case_simp_tac "s" 1);
--- a/src/HOLCF/IOA/meta_theory/TrivEx.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/IOA/meta_theory/TrivEx.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -11,7 +11,7 @@
 qed "imp_conj_lemma";
 
 
-goalw  thy [is_abstraction_def] 
+Goalw [is_abstraction_def] 
 "is_abstraction h_abs C_ioa A_ioa";
 by (rtac conjI 1);
 (* ------------- start states ------------ *)
@@ -28,7 +28,7 @@
 qed"h_abs_is_abstraction";
 
 
-goal thy "validIOA C_ioa (<>[] <%(n,a,m). n~=0>)";
+Goal "validIOA C_ioa (<>[] <%(n,a,m). n~=0>)";
 br AbsRuleT1 1;
 br h_abs_is_abstraction 1;
 br MC_result 1;
--- a/src/HOLCF/IOA/meta_theory/TrivEx2.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/IOA/meta_theory/TrivEx2.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -11,7 +11,7 @@
 qed "imp_conj_lemma";
 
 
-goalw  thy [is_abstraction_def] 
+Goalw [is_abstraction_def] 
 "is_abstraction h_abs C_ioa A_ioa";
 by (rtac conjI 1);
 (* ------------- start states ------------ *)
@@ -29,21 +29,21 @@
 
 
 (*
-goalw thy [xt2_def,plift,option_lift]
+Goalw [xt2_def,plift,option_lift]
   "(xt2 (plift afun)) (s,a,t) = (afun a)";
 (* !!!!!!!!!!!!! Occurs check !!!! *)
 by (option.induct_tac "a" 1);
 
 *)
 
-goalw thy [Enabled_def, enabled_def, h_abs_def,A_ioa_def,C_ioa_def,A_trans_def,
+Goalw [Enabled_def, enabled_def, h_abs_def,A_ioa_def,C_ioa_def,A_trans_def,
            C_trans_def,trans_of_def] 
 "!!s. Enabled A_ioa {INC} (h_abs s) ==> Enabled C_ioa {INC} s";
 auto();
 qed"Enabled_implication";
 
 
-goalw thy [is_live_abstraction_def]
+Goalw [is_live_abstraction_def]
 "is_live_abstraction h_abs (C_ioa, WF C_ioa {INC}) (A_ioa, WF A_ioa {INC})";
 auto();
 (* is_abstraction *)
@@ -54,7 +54,7 @@
 qed"h_abs_is_liveabstraction";
 
 
-goalw thy [C_live_ioa_def]
+Goalw [C_live_ioa_def]
 "validLIOA C_live_ioa (<>[] <%(n,a,m). n~=0>)";
 br AbsRuleT2 1;
 br h_abs_is_liveabstraction 1;
@@ -65,7 +65,7 @@
 
 
 (*
-goal thy "validIOA aut_ioa (Box (Init (%(s,a,t). a= Some(Alarm(APonR)))) 
+Goal "validIOA aut_ioa (Box (Init (%(s,a,t). a= Some(Alarm(APonR)))) 
 IMPLIES (Next (Init (%(s,a,t). info_comp(s) = APonR))))";
 
 *)
--- a/src/HOLCF/Lift.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/Lift.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -14,7 +14,7 @@
 
 
 (* flift1 is continuous in its argument itself*)
-goal thy "cont (lift_case UU f)"; 
+Goal "cont (lift_case UU f)"; 
 by (rtac flatdom_strict2cont 1);
 by (Simp_tac 1);
 qed"cont_flift1_arg";
@@ -22,7 +22,7 @@
 (* flift1 is continuous in a variable that occurs only 
    in the Def branch *)
 
-goal thy "!!f. [| !! a. cont (%y. (f y) a) |] ==> \
+Goal "!!f. [| !! a. cont (%y. (f y) a) |] ==> \
 \          cont (%y. lift_case UU (f y))";
 by (rtac cont2cont_CF1L_rev 1);
 by (strip_tac 1);
@@ -34,7 +34,7 @@
 (* flift1 is continuous in a variable that occurs either 
    in the Def branch or in the argument *)
 
-goal thy "!!f. [| !! a. cont (%y. (f y) a); cont g|] ==> \
+Goal "!!f. [| !! a. cont (%y. (f y) a); cont g|] ==> \
 \   cont (%y. lift_case UU (f y) (g y))";
 by (rtac cont2cont_app 1);
 back();
@@ -46,7 +46,7 @@
 
 (* flift2 is continuous in its argument itself *)
 
-goal thy "cont (lift_case UU (%y. Def (f y)))";
+Goal "cont (lift_case UU (%y. Def (f y)))";
 by (rtac flatdom_strict2cont 1);
 by (Simp_tac 1);
 qed"cont_flift2_arg";
@@ -82,25 +82,25 @@
 (* ---------------------------------------------------------- *)
 
 
-goal thy "flift1 f`(Def x) = (f x)";
+Goal "flift1 f`(Def x) = (f x)";
 by (simp_tac (simpset() addsimps [flift1_def]) 1);
 qed"flift1_Def";
 
-goal thy "flift2 f`(Def x) = Def (f x)";
+Goal "flift2 f`(Def x) = Def (f x)";
 by (simp_tac (simpset() addsimps [flift2_def]) 1);
 qed"flift2_Def";
 
-goal thy "flift1 f`UU = UU";
+Goal "flift1 f`UU = UU";
 by (simp_tac (simpset() addsimps [flift1_def]) 1);
 qed"flift1_UU";
 
-goal thy "flift2 f`UU = UU";
+Goal "flift2 f`UU = UU";
 by (simp_tac (simpset() addsimps [flift2_def]) 1);
 qed"flift2_UU";
 
 Addsimps [flift1_Def,flift2_Def,flift1_UU,flift2_UU];
 
-goal thy "!!x. x~=UU ==> (flift2 f)`x~=UU";
+Goal "!!x. x~=UU ==> (flift2 f)`x~=UU";
 by (def_tac 1);
 qed"flift2_nUU";
 
--- a/src/HOLCF/Lift1.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/Lift1.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -13,7 +13,7 @@
 (* less_lift is a partial order on type 'a -> 'b                            *)
 (* ------------------------------------------------------------------------ *)
 
-goalw thy [less_lift_def] "(x::'a lift) << x";
+Goalw [less_lift_def] "(x::'a lift) << x";
 by (fast_tac HOL_cs 1);
 qed"refl_less_lift";
 
--- a/src/HOLCF/Lift2.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/Lift2.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -20,7 +20,7 @@
 (* type ('a)lift is pointed                                                *)
 (* ------------------------------------------------------------------------ *)
 
-goal Lift2.thy  "Undef << x";
+Goal  "Undef << x";
 by (simp_tac (simpset() addsimps [inst_lift_po]) 1);
 qed"minimal_lift";
 
@@ -44,7 +44,7 @@
 
 (* Tailoring notUU_I of Pcpo.ML to Undef *)
 
-goal Lift2.thy "!!x. [| x<<y; ~x=Undef |] ==> ~y=Undef";
+Goal "!!x. [| x<<y; ~x=Undef |] ==> ~y=Undef";
 by (etac contrapos 1);
 by (hyp_subst_tac 1);
 by (rtac antisym_less 1);
@@ -55,7 +55,7 @@
 
 (* Tailoring chain_mono2 of Pcpo.ML to Undef *)
 
-goal Lift2.thy
+Goal
 "!!Y. [|? j.~Y(j)=Undef;chain(Y::nat=>('a)lift)|] \
 \ ==> ? j.!i. j<i-->~Y(i)=Undef";
 by Safe_tac;
@@ -70,7 +70,7 @@
 
 (* Tailoring flat_imp_chfin of Fix.ML to lift *)
 
-goal Lift2.thy
+Goal
         "(! Y. chain(Y::nat=>('a)lift)-->(? n. max_in_chain n Y))";
 by (rewtac max_in_chain_def);  
 by (strip_tac 1);
@@ -109,7 +109,7 @@
 
 (* Main Lemma: cpo_lift *)
 
-goal Lift2.thy  
+Goal  
   "!!Y. chain(Y::nat=>('a)lift) ==> ? x. range(Y) <<|x";
 by (cut_inst_tac [] flat_imp_chfin_poo 1);
 by (Step_tac 1);
--- a/src/HOLCF/Lift3.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/Lift3.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -70,7 +70,7 @@
               section"less_lift";
 (* --------------------------------------------------------*)
 
-goal thy "(x::'a lift) << y = (x=y | x=UU)";
+Goal "(x::'a lift) << y = (x=y | x=UU)";
 by (stac inst_lift_po 1);
 by (Simp_tac 1);
 qed"less_lift";
@@ -80,7 +80,7 @@
                  section"UU and Def";             
 (* ---------------------------------------------------------- *)
 
-goal thy "x=UU | (? y. x=Def y)"; 
+Goal "x=UU | (? y. x=Def y)"; 
 by (lift.induct_tac "x" 1);
 by (Asm_simp_tac 1);
 by (rtac disjI2 1);
@@ -94,7 +94,7 @@
 by (fast_tac (HOL_cs addSEs prems) 1);
 qed"Lift_cases";
 
-goal thy "(x~=UU)=(? y. x=Def y)";
+Goal "(x~=UU)=(? y. x=Def y)";
 by (rtac iffI 1);
 by (rtac Lift_cases 1);
 by (REPEAT (fast_tac (HOL_cs addSIs lift.distinct) 1));
@@ -115,7 +115,7 @@
 by (fast_tac (HOL_cs addSDs [DefE]) 1);
 qed"DefE2";
 
-goal thy "Def x << Def y = (x = y)";
+Goal "Def x << Def y = (x = y)";
 by (stac (hd lift.inject RS sym) 1);
 back();
 by (rtac iffI 1);
@@ -123,7 +123,7 @@
 by (etac (antisym_less_inverse RS conjunct1) 1);
 qed"Def_inject_less_eq";
 
-goal thy "Def x << y = (Def x = y)";
+Goal "Def x << y = (Def x = y)";
 by (simp_tac (simpset() addsimps [less_lift]) 1);
 qed"Def_less_is_eq";
 
@@ -133,19 +133,19 @@
               section"Lift is flat";
 (* ---------------------------------------------------------- *)
 
-goal thy "! x y::'a lift. x << y --> x = UU | x = y";
+Goal "! x y::'a lift. x << y --> x = UU | x = y";
 by (simp_tac (simpset() addsimps [less_lift]) 1);
 qed"ax_flat_lift";
 
 (* Two specific lemmas for the combination of LCF and HOL terms *)
 
-goal thy "!!f.[|cont g; cont f|] ==> cont(%x. ((f x)`(g x)) s)";
+Goal "!!f.[|cont g; cont f|] ==> cont(%x. ((f x)`(g x)) s)";
 by (rtac cont2cont_CF1L 1);
 by (REPEAT (resolve_tac cont_lemmas1 1));
 by Auto_tac;
 qed"cont_fapp_app";
 
-goal thy "!!f.[|cont g; cont f|] ==> cont(%x. ((f x)`(g x)) s t)";
+Goal "!!f.[|cont g; cont f|] ==> cont(%x. ((f x)`(g x)) s t)";
 by (rtac cont2cont_CF1L 1);
 by (etac cont_fapp_app 1);
 by (assume_tac 1);
--- a/src/HOLCF/Porder.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/Porder.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -112,7 +112,7 @@
         ]);
 
 
-goal thy "lub{x} = x";
+Goal "lub{x} = x";
 by (rtac thelubI 1);
 by (simp_tac (simpset() addsimps [is_lub,is_ub]) 1);
 qed "lub_singleton";
--- a/src/HOLCF/Porder0.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/Porder0.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -47,6 +47,6 @@
         (atac 1)
         ]);
 
-goal Porder0.thy "((x::'a::po)=y) = (x << y & y << x)";
+Goal "((x::'a::po)=y) = (x << y & y << x)";
 by (fast_tac (HOL_cs addSEs [antisym_less_inverse] addSIs [antisym_less]) 1);
 qed "po_eq_conv";
--- a/src/HOLCF/Tr.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/HOLCF/Tr.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -95,7 +95,7 @@
 (*  split-tac for If via If2 because the constant has to be a constant *)
 (* ------------------------------------------------------------------- *)
   
-goalw thy [If2_def] 
+Goalw [If2_def] 
   "P (If2 Q x y ) = ((Q=UU --> P UU) & (Q=TT --> P x) & (Q=FF --> P y))";
 by (res_inst_tac [("p","Q")] trE 1);
 by (REPEAT (Asm_full_simp_tac 1));
@@ -111,7 +111,7 @@
 (* ----------------------------------------------------------------- *)
 
 
-goal thy
+Goal
 "!!t.[|t~=UU|]==> ((t andalso s)=FF)=(t=FF | s=FF)";
 by (rtac iffI 1);
 by (res_inst_tac [("p","t")] trE 1);
@@ -120,7 +120,7 @@
 by Auto_tac;
 qed"andalso_or";
 
-goal thy "!!t.[|t~=UU|]==> ((t andalso s)~=FF)=(t~=FF & s~=FF)";
+Goal "!!t.[|t~=UU|]==> ((t andalso s)~=FF)=(t~=FF & s~=FF)";
 by (rtac iffI 1);
 by (res_inst_tac [("p","t")] trE 1);
 by Auto_tac;
@@ -128,23 +128,23 @@
 by Auto_tac;
 qed"andalso_and";
 
-goal thy "(Def x ~=FF)= x";
+Goal "(Def x ~=FF)= x";
 by (simp_tac (simpset() addsimps [FF_def]) 1);
 qed"Def_bool1";
 
-goal thy "(Def x = FF) = (~x)";
+Goal "(Def x = FF) = (~x)";
 by (simp_tac (simpset() addsimps [FF_def]) 1);
 qed"Def_bool2";
 
-goal thy "(Def x = TT) = x";
+Goal "(Def x = TT) = x";
 by (simp_tac (simpset() addsimps [TT_def]) 1);
 qed"Def_bool3";
 
-goal thy "(Def x ~= TT) = (~x)";
+Goal "(Def x ~= TT) = (~x)";
 by (simp_tac (simpset() addsimps [TT_def]) 1);
 qed"Def_bool4";
 
-goal thy 
+Goal 
   "(If Def P then A else B fi)= (if P then A else B)";
 by (res_inst_tac [("p","Def P")]  trE 1);
 by (Asm_full_simp_tac 1);
@@ -163,12 +163,12 @@
    replaced by a more general admissibility test that also checks 
    chain-finiteness, of which these lemmata are specific examples *)
 
-goal thy "x~=FF = (x=TT|x=UU)";
+Goal "x~=FF = (x=TT|x=UU)";
 by (res_inst_tac [("p","x")] trE 1);
 by (TRYALL (Asm_full_simp_tac));
 qed"adm_trick_1";
 
-goal thy "x~=TT = (x=FF|x=UU)";
+Goal "x~=TT = (x=FF|x=UU)";
 by (res_inst_tac [("p","x")] trE 1);
 by (TRYALL (Asm_full_simp_tac));
 qed"adm_trick_2";
@@ -176,12 +176,12 @@
 val adm_tricks = [adm_trick_1,adm_trick_2];
 
 
-goal thy "!!f. cont(f) ==> adm (%x. (f x)~=TT)";
+Goal "!!f. cont(f) ==> adm (%x. (f x)~=TT)";
 by (simp_tac (HOL_basic_ss addsimps adm_tricks) 1);
 by (REPEAT ((resolve_tac (adm_lemmas@cont_lemmas1) 1) ORELSE atac 1));
 qed"adm_nTT";
 
-goal thy "!!f. cont(f) ==> adm (%x. (f x)~=FF)";
+Goal "!!f. cont(f) ==> adm (%x. (f x)~=FF)";
 by (simp_tac (HOL_basic_ss addsimps adm_tricks) 1);
 by (REPEAT ((resolve_tac (adm_lemmas@cont_lemmas1) 1) ORELSE atac 1));
 qed"adm_nFF";
--- a/src/ZF/AC/AC15_WO6.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/AC/AC15_WO6.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -7,7 +7,7 @@
 
 open AC15_WO6;
 
-goal thy "!!x. Ord(x) ==> (UN a<x. F(a)) = (UN a:x. F(a))";
+Goal "!!x. Ord(x) ==> (UN a<x. F(a)) = (UN a:x. F(a))";
 by (fast_tac (claset() addSIs [ltI] addSDs [ltD]) 1);
 qed "OUN_eq_UN";
 
@@ -30,7 +30,7 @@
                 addSDs [HH_subset_x_imp_subset_Diff_UN RS not_emptyI2]) 1);
 val lemma2 = result();
 
-goalw thy [AC15_def, WO6_def] "!!Z. AC15 ==> WO6";
+Goalw [AC15_def, WO6_def] "!!Z. AC15 ==> WO6";
 by (rtac allI 1);
 by (eres_inst_tac [("x","Pow(A)-{0}")] allE 1);
 by (etac impE 1);
--- a/src/ZF/AC/AC16_WO4.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/AC/AC16_WO4.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -11,7 +11,7 @@
 (* The case of finite set                                                 *)
 (* ********************************************************************** *)
 
-goalw thy [Finite_def] "!!A. [| Finite(A); 0<m; m:nat |] ==>  \
+Goalw [Finite_def] "!!A. [| Finite(A); 0<m; m:nat |] ==>  \
 \       EX a f. Ord(a) & domain(f) = a &  \
 \               (UN b<a. f`b) = A & (ALL b<a. f`b lepoll m)";
 by (etac bexE 1);
@@ -34,16 +34,16 @@
 (* well_ord(x,r) ==> well_ord({{y,z}. y:x}, Something(x,z))  **)
 bind_thm ("well_ord_paired", (paired_bij RS bij_is_inj RS well_ord_rvimage));
 
-goal thy "!!A. [| A lepoll B; ~ A lepoll C |] ==> ~ B lepoll C";
+Goal "!!A. [| A lepoll B; ~ A lepoll C |] ==> ~ B lepoll C";
 by (fast_tac (claset() addEs [notE, lepoll_trans]) 1);
 qed "lepoll_trans1";
 
-goalw thy [lepoll_def]
+Goalw [lepoll_def]
         "!!X.[| Y lepoll X; well_ord(X, R) |] ==> EX S. well_ord(Y, S)";
 by (fast_tac (claset() addSEs [well_ord_rvimage]) 1);
 qed "well_ord_lepoll";
 
-goal thy "!!X. [| well_ord(X,R); well_ord(Y,S)  \
+Goal "!!X. [| well_ord(X,R); well_ord(Y,S)  \
 \               |] ==> EX T. well_ord(X Un Y, T)";
 by (eresolve_tac [well_ord_radd RS (Un_lepoll_sum RS well_ord_lepoll)] 1);
 by (assume_tac 1);
@@ -53,7 +53,7 @@
 (* There exists a well ordered set y such that ...                        *)
 (* ********************************************************************** *)
 
-goal thy "EX y R. well_ord(y,R) & x Int y = 0 & ~y lepoll z & ~Finite(y)";
+Goal "EX y R. well_ord(y,R) & x Int y = 0 & ~y lepoll z & ~Finite(y)";
 by (res_inst_tac [("x","{{a,x}. a:nat Un  Hartog(z)}")] exI 1);
 by (resolve_tac [transfer thy Ord_nat RS well_ord_Memrel RS (Ord_Hartog RS
                 well_ord_Memrel RSN (2, well_ord_Un)) RS exE] 1);
@@ -85,7 +85,7 @@
 (* ********************************************************************** *)
 
 (*Proof simplified by LCP*)
-goal thy "!!A. [| ~(EX x:A. f`x=y); f : inj(A, B); y:B |]  \
+Goal "!!A. [| ~(EX x:A. f`x=y); f : inj(A, B); y:B |]  \
 \       ==> (lam a:succ(A). if(a=A, y, f`a)) : inj(succ(A), B)";
 by (res_inst_tac [("d","%z. if(z=y, A, converse(f)`z)")] lam_injective 1);
 by (ALLGOALS
@@ -94,7 +94,7 @@
       setloop (split_tac [expand_if] ORELSE' Step_tac))));
 qed "succ_not_lepoll_lemma";
 
-goalw thy [lepoll_def, eqpoll_def, bij_def, surj_def]
+Goalw [lepoll_def, eqpoll_def, bij_def, surj_def]
         "!!A. [| ~A eqpoll B; A lepoll B |] ==> succ(A) lepoll B";
 by (fast_tac (claset() addSEs [succ_not_lepoll_lemma, inj_is_fun]) 1);
 qed "succ_not_lepoll_imp_eqpoll";
@@ -115,7 +115,7 @@
 (* There is a k-2 element subset of y                                     *)
 (* ********************************************************************** *)
 
-goalw thy [lepoll_def, eqpoll_def]
+Goalw [lepoll_def, eqpoll_def]
         "!!X. [| n:nat; nat lepoll X |] ==> EX Y. Y<=X & n eqpoll Y";
 by (fast_tac (FOL_cs addSDs [Ord_nat RSN (2, OrdmemD) RSN (2, restrict_inj)]
         addSIs [subset_refl]
@@ -125,7 +125,7 @@
 val ordertype_eqpoll =
         ordermap_bij RS (exI RS (eqpoll_def RS def_imp_iff RS iffD2));
 
-goal thy "!!y. [| well_ord(y,R); ~Finite(y); n:nat  \
+Goal "!!y. [| well_ord(y,R); ~Finite(y); n:nat  \
 \       |] ==> EX z. z<=y & n eqpoll z";
 by (etac nat_lepoll_imp_ex_eqpoll_n 1);
 by (resolve_tac [ordertype_eqpoll RS eqpoll_sym RS eqpoll_imp_lepoll
@@ -135,14 +135,14 @@
                         RS lepoll_infinite]) 1);
 qed "ex_subset_eqpoll_n";
 
-goalw thy [lesspoll_def] "!!n. n: nat ==> n lesspoll nat";
+Goalw [lesspoll_def] "!!n. n: nat ==> n lesspoll nat";
 by (fast_tac (claset() addSEs [Ord_nat RSN (2, ltI) RS leI RS le_imp_lepoll,
         eqpoll_sym RS eqpoll_imp_lepoll]
         addIs [Ord_nat RSN (2, nat_succI RS ltI) RS leI
         RS le_imp_lepoll RS lepoll_trans RS succ_lepoll_natE]) 1);
 qed "n_lesspoll_nat";
 
-goal thy "!!y. [| well_ord(y,R); ~Finite(y); k eqpoll a; a <= y; k: nat |]  \
+Goal "!!y. [| well_ord(y,R); ~Finite(y); k eqpoll a; a <= y; k: nat |]  \
 \       ==> y - a eqpoll y";
 by (fast_tac (empty_cs addIs [lepoll_lesspoll_lesspoll]
         addSIs [Card_cardinal, Diff_lesspoll_eqpoll_Card RS eqpoll_trans,
@@ -156,31 +156,31 @@
                 RS lepoll_infinite]) 1);
 qed "Diff_Finite_eqpoll";
 
-goal thy "!!x. [| a<=y; b:y-a; u:x |] ==> cons(b, cons(u, a)) : Pow(x Un y)";
+Goal "!!x. [| a<=y; b:y-a; u:x |] ==> cons(b, cons(u, a)) : Pow(x Un y)";
 by (Fast_tac 1);
 qed "cons_cons_subset";
 
-goal thy "!!x. [| a eqpoll k; a<=y; b:y-a; u:x; x Int y = 0  \
+Goal "!!x. [| a eqpoll k; a<=y; b:y-a; u:x; x Int y = 0  \
 \       |] ==> cons(b, cons(u, a)) eqpoll succ(succ(k))";
 by (fast_tac (claset() addSIs [cons_eqpoll_succ] addEs [equals0D]) 1);
 qed "cons_cons_eqpoll";
 
-goalw thy [s_u_def] "s_u(u, t_n, k, y) <= t_n";
+Goalw [s_u_def] "s_u(u, t_n, k, y) <= t_n";
 by (Fast_tac 1);
 qed "s_u_subset";
 
-goalw thy [s_u_def, succ_def]
+Goalw [s_u_def, succ_def]
         "!!w. [| w:t_n; cons(b,cons(u,a)) <= w; a <= y; b : y-a; k eqpoll a  \
 \       |] ==> w: s_u(u, t_n, succ(k), y)";
 by (fast_tac (claset() addDs [eqpoll_imp_lepoll RS cons_lepoll_cong]
                 addSEs [subset_imp_lepoll RSN (2, lepoll_trans)]) 1);
 qed "s_uI";
 
-goalw thy [s_u_def] "!!v. v : s_u(u, t_n, k, y) ==> u : v";
+Goalw [s_u_def] "!!v. v : s_u(u, t_n, k, y) ==> u : v";
 by (Fast_tac 1);
 qed "in_s_u_imp_u_in";
 
-goal thy
+Goal
         "!!y. [| ALL z:{z: Pow(x Un y) . z eqpoll succ(succ(k))}.  \
 \       EX! w. w:t_n & z <= w; \
 \       k eqpoll a; a <= y; b : y - a; u : x; x Int y = 0 |]  \
@@ -201,7 +201,7 @@
 by (fast_tac (claset() addSEs [s_u_subset RS subsetD, in_s_u_imp_u_in]) 1);
 qed "ex1_superset_a";
 
-goal thy
+Goal
         "!!A. [| succ(k) eqpoll A; k eqpoll B; B <= A; a : A-B; k:nat  \
 \       |] ==> A = cons(a, B)";
 by (rtac equalityI 1);
@@ -218,7 +218,7 @@
         (lepoll_trans RS lepoll_trans) RS succ_lepoll_natE]) 1);
 qed "set_eq_cons";
 
-goal thy
+Goal
         "!!y. [| ALL z:{z: Pow(x Un y) . z eqpoll succ(succ(k))}.  \
 \       EX! w. w:t_n & z <= w; \
 \       ALL v:s_u(u, t_n, succ(k), y). succ(k) eqpoll v Int y;  \
@@ -230,15 +230,15 @@
 by (REPEAT (Fast_tac 1));
 qed "the_eq_cons";
 
-goal thy "!!a. [| cons(x,a) = cons(y,a); x~: a |] ==> x = y ";
+Goal "!!a. [| cons(x,a) = cons(y,a); x~: a |] ==> x = y ";
 by (fast_tac (claset() addSEs [equalityE]) 1);
 qed "cons_eqE";
 
-goal thy "!!A. A = B ==> A Int C = B Int C";
+Goal "!!A. A = B ==> A Int C = B Int C";
 by (Asm_simp_tac 1);
 qed "eq_imp_Int_eq";
 
-goal thy "!!a. [| a=b; a=c; b=d |] ==> c=d";
+Goal "!!a. [| a=b; a=c; b=d |] ==> c=d";
 by (Asm_full_simp_tac 1);
 qed "msubst";
 
@@ -247,7 +247,7 @@
 (*      where a is certain k-2 element subset of y                        *)
 (* ********************************************************************** *)
 
-goal thy
+Goal
         "!!y. [| ALL z:{z: Pow(x Un y) . z eqpoll succ(succ(k))}.  \
 \       EX! w. w:t_n & z <= w; \
 \       ALL v:s_u(u, t_n, succ(k), y). succ(k) eqpoll v Int y;  \
@@ -281,7 +281,7 @@
 (* some arithmetic                                                        *)
 (* ********************************************************************** *)
 
-goal thy 
+Goal 
         "!!k. [| k:nat; m:nat |] ==>  \
 \       ALL A B. A eqpoll k #+ m & k lepoll B & B<=A --> A-B lepoll m";
 by (eres_inst_tac [("n","k")] nat_induct 1);
@@ -300,7 +300,7 @@
 by (Fast_tac 1);
 qed "eqpoll_sum_imp_Diff_lepoll_lemma";
 
-goal thy "!!k. [| A eqpoll succ(k #+ m); B<=A; succ(k) lepoll B;  \
+Goal "!!k. [| A eqpoll succ(k #+ m); B<=A; succ(k) lepoll B;  \
 \       k:nat; m:nat |]  \
 \       ==> A-B lepoll m";
 by (dresolve_tac [add_succ RS ssubst] 1);
@@ -313,7 +313,7 @@
 (* similar properties for eqpoll                                          *)
 (* ********************************************************************** *)
 
-goal thy 
+Goal 
         "!!k. [| k:nat; m:nat |] ==>  \
 \       ALL A B. A eqpoll k #+ m & k eqpoll B & B<=A --> A-B eqpoll m";
 by (eres_inst_tac [("n","k")] nat_induct 1);
@@ -332,7 +332,7 @@
 by (Fast_tac 1);
 qed "eqpoll_sum_imp_Diff_eqpoll_lemma";
 
-goal thy "!!k. [| A eqpoll succ(k #+ m); B<=A; succ(k) eqpoll B;  \
+Goal "!!k. [| A eqpoll succ(k #+ m); B<=A; succ(k) eqpoll B;  \
 \       k:nat; m:nat |]  \
 \       ==> A-B eqpoll m";
 by (dresolve_tac [add_succ RS ssubst] 1);
@@ -345,12 +345,12 @@
 (* back to the second part                                                *)
 (* ********************************************************************** *)
 
-goal thy "!!w. [| x Int y = 0; w <= x Un y |]  \
+Goal "!!w. [| x Int y = 0; w <= x Un y |]  \
 \        ==> w Int (x - {u}) = w - cons(u, w Int y)";
 by (fast_tac (claset() addEs [equals0D]) 1);
 qed "w_Int_eq_w_Diff";
 
-goal thy "!!w. [| w:{v:s_u(u, t_n, succ(l), y). a <= v};  \
+Goal "!!w. [| w:{v:s_u(u, t_n, succ(l), y). a <= v};  \
 \       l eqpoll a; t_n <= {v:Pow(x Un y). v eqpoll succ(succ(l) #+ m)};  \
 \       m:nat; l:nat; x Int y = 0; u : x;  \
 \       ALL v:s_u(u, t_n, succ(l), y). succ(l) eqpoll v Int y  \
@@ -364,12 +364,12 @@
         addSIs [nat_succI, eqpoll_sum_imp_Diff_eqpoll]) 1);
 qed "w_Int_eqpoll_m";
 
-goal thy "!!m. [| 0<m; x eqpoll m; m:nat |] ==> x ~= 0";
+Goal "!!m. [| 0<m; x eqpoll m; m:nat |] ==> x ~= 0";
 by (fast_tac (empty_cs
         addSEs [mem_irrefl, ltE, eqpoll_succ_imp_not_empty, natE]) 1);
 qed "eqpoll_m_not_empty";
 
-goal thy
+Goal
         "!!z. [| z : xa Int (x - {u}); l eqpoll a; a <= y; x Int y = 0; u:x  \
 \       |] ==> cons(z, cons(u, a)) : {v: Pow(x Un y). v eqpoll succ(succ(l))}";
 by (fast_tac (claset() addSIs [cons_eqpoll_succ] addEs [equals0D, eqpoll_sym]) 1);
@@ -380,7 +380,7 @@
 (*      to {v:Pow(x). v eqpoll n-k}                                       *)
 (* ********************************************************************** *)
 
-goal thy 
+Goal 
         "!!x. [| ALL z:{z: Pow(x Un y) . z eqpoll succ(succ(l))}.  \
 \       EX! w. w:t_n & z <= w; \
 \       t_n <= {v:Pow(x Un y). v eqpoll succ(succ(l) #+ m)};  \
@@ -410,13 +410,13 @@
 by (fast_tac (claset() addSEs [s_u_subset RS subsetD, in_s_u_imp_u_in]) 1);
 qed "subset_s_u_lepoll_w";
 
-goal thy "!!k. [| 0<k; k:nat |] ==> EX l:nat. k = succ(l)";
+Goal "!!k. [| 0<k; k:nat |] ==> EX l:nat. k = succ(l)";
 by (etac natE 1);
 by (fast_tac (empty_cs addSEs [ltE, mem_irrefl]) 1);
 by (fast_tac (empty_cs addSIs [refl, bexI]) 1);
 qed "ex_eq_succ";
 
-goal thy
+Goal
  "!!x. [| ALL z:{z: Pow(x Un y) . z eqpoll succ(k)}.  \
 \                  EX! w. w:t_n & z <= w; \
 \         well_ord(y,R); ~Finite(y); u:x; x Int y = 0;  \
@@ -438,12 +438,12 @@
 (* LL can be well ordered                                                 *)
 (* ********************************************************************** *)
 
-goal thy "{x:Pow(X). x lepoll 0} = {0}";
+Goal "{x:Pow(X). x lepoll 0} = {0}";
 by (fast_tac (claset() addSDs [lepoll_0_is_0]
                       addSIs [lepoll_refl]) 1);
 qed "subsets_lepoll_0_eq_unit";
 
-goal thy "!!X. [| well_ord(X, R); ~Finite(X); n:nat |]  \
+Goal "!!X. [| well_ord(X, R); ~Finite(X); n:nat |]  \
 \               ==> EX S. well_ord({Y: Pow(X) . Y eqpoll succ(n)}, S)";
 by (resolve_tac [well_ord_infinite_subsets_eqpoll_X
         RS (eqpoll_def RS def_imp_iff RS iffD1) RS exE] 1
@@ -451,14 +451,14 @@
 by (fast_tac (claset() addSEs [bij_is_inj RS well_ord_rvimage]) 1);
 qed "well_ord_subsets_eqpoll_n";
 
-goal thy "!!n. n:nat ==> {z:Pow(y). z lepoll succ(n)} =  \
+Goal "!!n. n:nat ==> {z:Pow(y). z lepoll succ(n)} =  \
 \       {z:Pow(y). z lepoll n} Un {z:Pow(y). z eqpoll succ(n)}";
 by (fast_tac (claset() addIs [le_refl, leI, le_imp_lepoll]
                 addSDs [lepoll_succ_disj]
                 addSEs [nat_into_Ord, lepoll_trans, eqpoll_imp_lepoll]) 1);
 qed "subsets_lepoll_succ";
 
-goal thy "!!n. n:nat ==>  \
+Goal "!!n. n:nat ==>  \
 \       {z:Pow(y). z lepoll n} Int {z:Pow(y). z eqpoll succ(n)} = 0";
 by (fast_tac (claset() addSEs [eqpoll_sym RS eqpoll_imp_lepoll 
                 RS lepoll_trans RS succ_lepoll_natE]
@@ -469,16 +469,16 @@
 (* unit set is well-ordered by the empty relation                         *)
 (* ********************************************************************** *)
 
-goalw thy [irrefl_def, trans_on_def, part_ord_def, linear_def, tot_ord_def]
+Goalw [irrefl_def, trans_on_def, part_ord_def, linear_def, tot_ord_def]
         "tot_ord({a},0)";
 by (Simp_tac 1);
 qed "tot_ord_unit";
 
-goalw thy [wf_on_def, wf_def] "wf[{a}](0)";
+Goalw [wf_on_def, wf_def] "wf[{a}](0)";
 by (Fast_tac 1);
 qed "wf_on_unit";
 
-goalw thy [well_ord_def] "well_ord({a},0)";
+Goalw [well_ord_def] "well_ord({a},0)";
 by (simp_tac (simpset() addsimps [tot_ord_unit, wf_on_unit]) 1);
 qed "well_ord_unit";
 
@@ -486,7 +486,7 @@
 (* well_ord_subsets_lepoll_n                                              *)
 (* ********************************************************************** *)
 
-goal thy "!!y r. [| well_ord(y,r); ~Finite(y); n:nat |] ==>  \
+Goal "!!y r. [| well_ord(y,r); ~Finite(y); n:nat |] ==>  \
 \       EX R. well_ord({z:Pow(y). z lepoll n}, R)";
 by (etac nat_induct 1);
 by (fast_tac (claset() addSIs [well_ord_unit]
@@ -501,7 +501,7 @@
 by (fast_tac (claset() addSEs [bij_is_inj RS well_ord_rvimage]) 1);
 qed "well_ord_subsets_lepoll_n";
 
-goalw thy [LL_def, MM_def]
+Goalw [LL_def, MM_def]
         "!!x. t_n <= {v:Pow(x Un y). v eqpoll n}  \
 \               ==> LL(t_n, k, y) <= {z:Pow(y). z lepoll n}";
 by (fast_tac (claset() addSEs [RepFunE]
@@ -509,7 +509,7 @@
                 RSN (2, lepoll_trans))]) 1);
 qed "LL_subset";
 
-goal thy "!!x. [| t_n <= {v:Pow(x Un y). v eqpoll n}; \
+Goal "!!x. [| t_n <= {v:Pow(x Un y). v eqpoll n}; \
 \               well_ord(y, R); ~Finite(y); n:nat  \
 \               |] ==> EX S. well_ord(LL(t_n, k, y), S)";
 by (fast_tac (FOL_cs addIs [exI]
@@ -521,7 +521,7 @@
 (* every element of LL is a contained in exactly one element of MM        *)
 (* ********************************************************************** *)
 
-goalw thy [MM_def, LL_def]
+Goalw [MM_def, LL_def]
         "!!x. [| ALL z:{z: Pow(x Un y) . z eqpoll k}. EX! w. w:t_n & z <= w; \
 \       t_n <= {v:Pow(x Un y). v eqpoll n}; \
 \       v:LL(t_n, k, y)  \
@@ -546,7 +546,7 @@
 (* The union of appropriate values is the whole x                         *)
 (* ********************************************************************** *)
 
-goal thy
+Goal
         "!!x. [| ALL z:{z: Pow(x Un y) . z eqpoll succ(k)}.  \
 \       EX! w. w:t_n & z <= w; \
 \       well_ord(y,R); ~Finite(y); u:x; x Int y = 0;  \
@@ -559,15 +559,15 @@
 by (Fast_tac 1);
 qed "exists_in_MM";
 
-goalw thy [LL_def] "!!w. w : MM(t_n, k, y) ==> w Int y : LL(t_n, k, y)";
+Goalw [LL_def] "!!w. w : MM(t_n, k, y) ==> w Int y : LL(t_n, k, y)";
 by (Fast_tac 1);
 qed "Int_in_LL";
 
-goalw thy [MM_def] "MM(t_n, k, y) <= t_n";
+Goalw [MM_def] "MM(t_n, k, y) <= t_n";
 by (Fast_tac 1);
 qed "MM_subset";
 
-goal thy 
+Goal 
         "!!x. [| ALL z:{z: Pow(x Un y) . z eqpoll succ(k)}.  \
 \       EX! w. w:t_n & z <= w; \
 \       well_ord(y,R); ~Finite(y); u:x; x Int y = 0;  \
@@ -585,7 +585,7 @@
 by (REPEAT (fast_tac (claset() addEs [equals0D] addSEs [Int_in_LL]) 1));
 qed "exists_in_LL";
 
-goalw thy [LL_def] 
+Goalw [LL_def] 
         "!!x. [| ALL z:{z: Pow(x Un y) . z eqpoll k}. EX! w. w:t_n & z <= w; \
 \       t_n <= {v:Pow(x Un y). v eqpoll n};  \
 \       v : LL(t_n, k, y) |]  \
@@ -594,7 +594,7 @@
                 unique_superset_in_MM RS the_equality2 RS ssubst]) 1);
 qed "in_LL_eq_Int";
 
-goal thy  
+Goal  
         "!!x. [| ALL z:{z: Pow(x Un y) . z eqpoll k}. EX! w. w:t_n & z <= w; \
 \       t_n <= {v:Pow(x Un y). v eqpoll n};  \
 \       v : LL(t_n, k, y) |]  \
@@ -603,7 +603,7 @@
         (MM_subset RS subsetD)]) 1);
 qed "the_in_MM_subset";
 
-goalw thy [GG_def] 
+Goalw [GG_def] 
         "!!x. [| ALL z:{z: Pow(x Un y) . z eqpoll k}. EX! w. w:t_n & z <= w; \
 \       t_n <= {v:Pow(x Un y). v eqpoll n};  \
 \       v : LL(t_n, k, y) |]  \
@@ -617,7 +617,7 @@
 by (fast_tac (claset() addEs [ssubst]) 1);
 qed "GG_subset";
 
-goal thy  
+Goal  
         "!!x. [| well_ord(LL(t_n, succ(k), y), S);  \
 \       ALL z:{z: Pow(x Un y) . z eqpoll succ(k)}. EX! w. w:t_n & z <= w; \
 \       well_ord(y,R); ~Finite(y); x Int y = 0;  \
@@ -646,18 +646,18 @@
 (* Every element of the family is less than or equipollent to n-k (m)     *)
 (* ********************************************************************** *)
 
-goalw thy [MM_def]
+Goalw [MM_def]
         "!!w. [| w : MM(t_n, k, y); t_n <= {v:Pow(x Un y). v eqpoll n}  \
 \       |] ==> w eqpoll n";
 by (Fast_tac 1);
 qed "in_MM_eqpoll_n";
 
-goalw thy [LL_def, MM_def]
+Goalw [LL_def, MM_def]
         "!!w. w : LL(t_n, k, y) ==> k lepoll w";
 by (Fast_tac 1);
 qed "in_LL_eqpoll_n";
 
-goalw thy [GG_def] 
+Goalw [GG_def] 
         "!!x. [|  \
 \       ALL z:{z: Pow(x Un y) . z eqpoll succ(k)}. EX! w. w:t_n & z <= w; \
 \       t_n <= {v:Pow(x Un y). v eqpoll succ(k #+ m)};  \
@@ -684,7 +684,7 @@
 (* The main theorem AC16(n, k) ==> WO4(n-k)                               *)
 (* ********************************************************************** *)
 
-goalw thy [AC16_def,WO4_def]
+Goalw [AC16_def,WO4_def]
         "!!n k. [| AC16(k #+ m, k); 0 < k; 0 < m; k:nat; m:nat |] ==> WO4(m)";
 by (rtac allI 1);
 by (excluded_middle_tac "Finite(A)" 1);
--- a/src/ZF/AC/AC16_lemmas.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/AC/AC16_lemmas.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -7,11 +7,11 @@
 
 open AC16_lemmas;
 
-goal thy "!!a. a~:A ==> cons(a,A)-{a}=A";
+Goal "!!a. a~:A ==> cons(a,A)-{a}=A";
 by (Fast_tac 1);
 qed "cons_Diff_eq";
 
-goalw thy [lepoll_def] "1 lepoll X <-> (EX x. x:X)";
+Goalw [lepoll_def] "1 lepoll X <-> (EX x. x:X)";
 by (rtac iffI 1);
 by (fast_tac (claset() addIs [inj_is_fun RS apply_type]) 1);
 by (etac exE 1);
@@ -19,7 +19,7 @@
 by (fast_tac (claset() addSIs [lam_injective]) 1);
 qed "nat_1_lepoll_iff";
 
-goal thy "X eqpoll 1 <-> (EX x. X={x})";
+Goal "X eqpoll 1 <-> (EX x. X={x})";
 by (rtac iffI 1);
 by (etac eqpollE 1);
 by (dresolve_tac [nat_1_lepoll_iff RS iffD1] 1);
@@ -27,12 +27,12 @@
 by (fast_tac (claset() addSIs [singleton_eqpoll_1]) 1);
 qed "eqpoll_1_iff_singleton";
 
-goalw thy [succ_def] 
+Goalw [succ_def] 
       "!!x. [| x eqpoll n; y~:x |] ==> cons(y,x) eqpoll succ(n)";
 by (fast_tac (claset() addSEs [cons_eqpoll_cong, mem_irrefl]) 1);
 qed "cons_eqpoll_succ";
 
-goal thy "{Y:Pow(X). Y eqpoll 1} = {{x}. x:X}";
+Goal "{Y:Pow(X). Y eqpoll 1} = {{x}. x:X}";
 by (rtac equalityI 1);
 by (rtac subsetI 1);
 by (etac CollectE 1);
@@ -45,7 +45,7 @@
 by (fast_tac (claset() addSIs [singleton_eqpoll_1]) 1);
 qed "subsets_eqpoll_1_eq";
 
-goalw thy [eqpoll_def, bij_def] "X eqpoll {{x}. x:X}";
+Goalw [eqpoll_def, bij_def] "X eqpoll {{x}. x:X}";
 by (res_inst_tac [("x","lam x:X. {x}")] exI 1);
 by (rtac IntI 1);
 by (rewrite_goals_tac [inj_def, surj_def]);
@@ -56,12 +56,12 @@
 by (fast_tac (claset() addSIs [lam_type]) 1);
 qed "eqpoll_RepFun_sing";
 
-goal thy "{Y:Pow(X). Y eqpoll 1} eqpoll X";
+Goal "{Y:Pow(X). Y eqpoll 1} eqpoll X";
 by (resolve_tac [subsets_eqpoll_1_eq RS ssubst] 1);
 by (resolve_tac [eqpoll_RepFun_sing RS eqpoll_sym] 1);
 qed "subsets_eqpoll_1_eqpoll";
 
-goal thy "!!x. [| InfCard(x); y:Pow(x); y eqpoll succ(z) |]  \
+Goal "!!x. [| InfCard(x); y:Pow(x); y eqpoll succ(z) |]  \
 \               ==> (LEAST i. i:y) : y";
 by (eresolve_tac [eqpoll_sym RS eqpoll_imp_lepoll RS 
                 succ_lepoll_imp_not_empty RS not_emptyE] 1);
@@ -70,7 +70,7 @@
         addEs [Ord_in_Ord]) 1);
 qed "InfCard_Least_in";
 
-goalw thy [lepoll_def] "!!i. [| InfCard(x); n:nat |] ==>  \
+Goalw [lepoll_def] "!!i. [| InfCard(x); n:nat |] ==>  \
 \       {y:Pow(x). y eqpoll succ(succ(n))} lepoll  \
 \       x*{y:Pow(x). y eqpoll succ(n)}";
 by (res_inst_tac [("x","lam y:{y:Pow(x). y eqpoll succ(succ(n))}. \
@@ -104,7 +104,7 @@
 by (fast_tac (claset() addSEs [leE,ltE]) 1);
 qed "set_of_Ord_succ_Union";
 
-goal thy "!!i. j<=i ==> i ~: j";
+Goal "!!i. j<=i ==> i ~: j";
 by (fast_tac (claset() addSEs [mem_irrefl]) 1);
 qed "subset_not_mem";
 
@@ -113,19 +113,19 @@
 by (eresolve_tac prems 1);
 qed "succ_Union_not_mem";
 
-goal thy "Union(cons(succ(Union(z)),z)) = succ(Union(z))";
+Goal "Union(cons(succ(Union(z)),z)) = succ(Union(z))";
 by (Fast_tac 1);
 qed "Union_cons_eq_succ_Union";
 
-goal thy "!!i. [| Ord(i); Ord(j) |] ==> i Un j = i | i Un j = j";
+Goal "!!i. [| Ord(i); Ord(j) |] ==> i Un j = i | i Un j = j";
 by (fast_tac (claset() addSDs [le_imp_subset] addEs [Ord_linear_le]) 1);
 qed "Un_Ord_disj";
 
-goal thy "!!X. x:X ==> Union(X) = x Un Union(X-{x})";
+Goal "!!X. x:X ==> Union(X) = x Un Union(X-{x})";
 by (Fast_tac 1);
 qed "Union_eq_Un";
 
-goal thy "!!n. n:nat ==>  \
+Goal "!!n. n:nat ==>  \
 \       ALL z. (ALL y:z. Ord(y)) & z eqpoll n & z~=0 --> Union(z) : z";
 by (etac nat_induct 1);
 by (fast_tac (claset() addSDs [eqpoll_imp_lepoll RS lepoll_0_is_0]) 1);
@@ -149,13 +149,13 @@
 by (rtac subst_elem 1 THEN (TRYALL assume_tac));
 qed "Union_in_lemma";
 
-goal thy "!!z. [| ALL x:z. Ord(x); z eqpoll n; z~=0; n:nat |]  \
+Goal "!!z. [| ALL x:z. Ord(x); z eqpoll n; z~=0; n:nat |]  \
 \               ==> Union(z) : z";
 by (dtac Union_in_lemma 1);
 by (fast_tac FOL_cs 1);
 qed "Union_in";
 
-goal thy "!!x. [| InfCard(x); z:Pow(x); z eqpoll n; n:nat |]  \
+Goal "!!x. [| InfCard(x); z:Pow(x); z eqpoll n; n:nat |]  \
 \               ==> succ(Union(z)) : x";
 by (resolve_tac [Limit_has_succ RS ltE] 1 THEN (assume_tac 3));
 by (etac InfCard_is_Limit 1);
@@ -170,7 +170,7 @@
 		 (2, InfCard_is_Card RS Card_is_Ord RS Ord_in_Ord)]) 1);
 qed "succ_Union_in_x";
 
-goalw thy [lepoll_def] "!!X. [| InfCard(x); n:nat |] ==>  \
+Goalw [lepoll_def] "!!X. [| InfCard(x); n:nat |] ==>  \
 \       {y:Pow(x). y eqpoll succ(n)} lepoll  \
 \       {y:Pow(x). y eqpoll succ(succ(n))}";
 by (res_inst_tac [("x","lam z:{y:Pow(x). y eqpoll succ(n)}.  \
@@ -188,7 +188,7 @@
 by (fast_tac (claset() addSIs [succ_Union_in_x, nat_succI]) 1);
 qed "succ_lepoll_succ_succ";
 
-goal thy "!!X. [| InfCard(X); n:nat |]  \
+Goal "!!X. [| InfCard(X); n:nat |]  \
 \       ==> {Y:Pow(X). Y eqpoll succ(n)} eqpoll X";
 by (etac nat_induct 1);
 by (rtac subsets_eqpoll_1_eqpoll 1);
@@ -205,18 +205,18 @@
         addSIs [succ_lepoll_succ_succ]) 1);
 qed "subsets_eqpoll_X";
 
-goalw thy [surj_def] "!!f. [| f:surj(A,B); y<=B |]  \
+Goalw [surj_def] "!!f. [| f:surj(A,B); y<=B |]  \
 \       ==> f``(converse(f)``y) = y";
 by (fast_tac (claset() addDs [apply_equality2]
 	              addEs [apply_iff RS iffD2]) 1);
 qed "image_vimage_eq";
 
-goal thy "!!f. [| f:inj(A,B); y<=A |] ==> converse(f)``(f``y) = y";
+Goal "!!f. [| f:inj(A,B); y<=A |] ==> converse(f)``(f``y) = y";
 by (fast_tac (claset() addSEs [inj_is_fun RS apply_Pair]
                 addDs [inj_equality]) 1);
 qed "vimage_image_eq";
 
-goalw thy [eqpoll_def] "!!A B. A eqpoll B  \
+Goalw [eqpoll_def] "!!A B. A eqpoll B  \
 \       ==> {Y:Pow(A). Y eqpoll n} eqpoll {Y:Pow(B). Y eqpoll n}";
 by (etac exE 1);
 by (res_inst_tac [("x","lam X:{Y:Pow(A). EX f. f : bij(Y, n)}. f``X")] exI 1);
@@ -232,22 +232,22 @@
 by (fast_tac (claset() addSEs [bij_is_surj RS image_vimage_eq]) 1);
 qed "subsets_eqpoll";
 
-goalw thy [WO2_def] "!!X. WO2 ==> EX a. Card(a) & X eqpoll a";
+Goalw [WO2_def] "!!X. WO2 ==> EX a. Card(a) & X eqpoll a";
 by (REPEAT (eresolve_tac [allE,exE,conjE] 1));
 by (fast_tac (claset() addSEs [well_ord_Memrel RS well_ord_cardinal_eqpoll RS
                 (eqpoll_sym RSN (2, eqpoll_trans)) RS eqpoll_sym]
                 addSIs [Card_cardinal]) 1);
 qed "WO2_imp_ex_Card";
 
-goal thy "!!X. [| X lepoll Y; ~Finite(X) |] ==> ~Finite(Y)";
+Goal "!!X. [| X lepoll Y; ~Finite(X) |] ==> ~Finite(Y)";
 by (fast_tac (empty_cs addEs [notE, lepoll_Finite] addSIs [notI]) 1); 
 qed "lepoll_infinite";
 
-goalw thy [InfCard_def] "!!X. [| ~Finite(X); Card(X) |] ==> InfCard(X)";
+Goalw [InfCard_def] "!!X. [| ~Finite(X); Card(X) |] ==> InfCard(X)";
 by (fast_tac (claset() addSEs [Card_is_Ord RS nat_le_infinite_Ord]) 1);
 qed "infinite_Card_is_InfCard";
 
-goal thy "!!X n. [| WO2; n:nat; ~Finite(X) |]  \
+Goal "!!X n. [| WO2; n:nat; ~Finite(X) |]  \
 \       ==> {Y:Pow(X). Y eqpoll succ(n)} eqpoll X";
 by (dtac WO2_imp_ex_Card 1);
 by (REPEAT (eresolve_tac [allE,exE,conjE] 1));
@@ -259,12 +259,12 @@
 by (etac eqpoll_sym 1);
 qed "WO2_infinite_subsets_eqpoll_X";
 
-goal thy "!!X. well_ord(X,R) ==> EX a. Card(a) & X eqpoll a";
+Goal "!!X. well_ord(X,R) ==> EX a. Card(a) & X eqpoll a";
 by (fast_tac (claset() addSEs [well_ord_cardinal_eqpoll RS eqpoll_sym]
                 addSIs [Card_cardinal]) 1);
 qed "well_ord_imp_ex_Card";
 
-goal thy "!!X n. [| well_ord(X,R); n:nat; ~Finite(X) |]  \
+Goal "!!X n. [| well_ord(X,R); n:nat; ~Finite(X) |]  \
 \               ==> {Y:Pow(X). Y eqpoll succ(n)} eqpoll X";
 by (dtac well_ord_imp_ex_Card 1);
 by (REPEAT (eresolve_tac [allE,exE,conjE] 1));
--- a/src/ZF/AC/AC17_AC1.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/AC/AC17_AC1.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -11,7 +11,7 @@
 (* more properties of HH                                                   *)
 (* *********************************************************************** *)
 
-goal thy "!!f. [| x - (UN j:LEAST i. HH(lam X:Pow(x)-{0}. {f`X}, x, i) = {x}. \
+Goal "!!f. [| x - (UN j:LEAST i. HH(lam X:Pow(x)-{0}. {f`X}, x, i) = {x}. \
 \       HH(lam X:Pow(x)-{0}. {f`X}, x, j)) = 0;  \
 \       f : Pow(x)-{0} -> x |]  \
 \       ==> EX r. well_ord(x,r)";
@@ -25,7 +25,7 @@
 (* theorems closer to the proof                                            *)
 (* *********************************************************************** *)
 
-goalw thy AC_defs "!!Z. ~AC1 ==>  \
+Goalw AC_defs "!!Z. ~AC1 ==>  \
 \               EX A. ALL f:Pow(A)-{0} -> A. EX u:Pow(A)-{0}. f`u ~: u";
 by (etac swap 1);
 by (rtac allI 1);
@@ -37,7 +37,7 @@
 by (fast_tac (claset() addSIs [restrict_type]) 1);
 qed "not_AC1_imp_ex";
 
-goal thy "!!x. [| ALL f:Pow(x) - {0} -> x. EX u: Pow(x) - {0}. f`u~:u;  \
+Goal "!!x. [| ALL f:Pow(x) - {0} -> x. EX u: Pow(x) - {0}. f`u~:u;  \
 \       EX f: Pow(x)-{0}->x. \
 \       x - (UN a:(LEAST i. HH(lam X:Pow(x)-{0}. {f`X},x,i)={x}).  \
 \       HH(lam X:Pow(x)-{0}. {f`X},x,a)) = 0 |] \
@@ -53,24 +53,24 @@
 by (Fast_tac 1);
 val lemma1 = result();
 
-goal thy "!!x. ~ (EX f: Pow(x)-{0}->x. x - F(f) = 0)  \
+Goal "!!x. ~ (EX f: Pow(x)-{0}->x. x - F(f) = 0)  \
 \       ==> (lam f: Pow(x)-{0}->x. x - F(f))  \
 \               : (Pow(x) -{0} -> x) -> Pow(x) - {0}";
 by (fast_tac (claset() addSIs [lam_type] addSDs [Diff_eq_0_iff RS iffD1]) 1);
 val lemma2 = result();
 
-goal thy "!!f. [| f`Z : Z; Z:Pow(x)-{0} |] ==>  \
+Goal "!!f. [| f`Z : Z; Z:Pow(x)-{0} |] ==>  \
 \       (lam X:Pow(x)-{0}. {f`X})`Z : Pow(Z)-{0}";
 by (Asm_full_simp_tac 1);
 by (fast_tac (claset() addSDs [equals0D]) 1);
 val lemma3 = result();
 
-goal thy "!!z. EX f:F. f`((lam f:F. Q(f))`f) : (lam f:F. Q(f))`f  \
+Goal "!!z. EX f:F. f`((lam f:F. Q(f))`f) : (lam f:F. Q(f))`f  \
 \       ==> EX f:F. f`Q(f) : Q(f)";
 by (Asm_full_simp_tac 1);
 val lemma4 = result();
 
-goalw thy [AC17_def] "!!Z. AC17 ==> AC1";
+Goalw [AC17_def] "!!Z. AC17 ==> AC1";
 by (rtac classical 1);
 by (eresolve_tac [not_AC1_imp_ex RS exE] 1);
 by (excluded_middle_tac
--- a/src/ZF/AC/AC18_AC19.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/AC/AC18_AC19.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -11,7 +11,7 @@
 (* AC1 ==> AC18                                                           *)
 (* ********************************************************************** *)
 
-goal thy "!!f. [| f: (PROD b:{P(a). a:A}. b); ALL a:A. P(a)<=Q(a) |]  \
+Goal "!!f. [| f: (PROD b:{P(a). a:A}. b); ALL a:A. P(a)<=Q(a) |]  \
 \               ==> (lam a:A. f`P(a)):(PROD a:A. Q(a))";
 by (rtac lam_type 1);
 by (dtac apply_type 1);
@@ -20,7 +20,7 @@
 by (etac subsetD 1 THEN (assume_tac 1));
 qed "PROD_subsets";
 
-goal thy "!!X. [| ALL A. 0 ~: A --> (EX f. f : (PROD X:A. X)); A ~= 0 |] ==>  \
+Goal "!!X. [| ALL A. 0 ~: A --> (EX f. f : (PROD X:A. X)); A ~= 0 |] ==>  \
 \  (INT a:A. UN b:B(a). X(a, b)) <= (UN f:PROD a:A. B(a). INT a:A. X(a, f`a))";
 by (rtac subsetI 1);
 by (eres_inst_tac [("x","{{b:B(a). x:X(a,b)}. a:A}")] allE 1);
@@ -55,14 +55,14 @@
 (* AC19 ==> AC1                                                           *)
 (* ********************************************************************** *)
 
-goalw thy [u_def]
+Goalw [u_def]
         "!!A. [| A ~= 0; 0 ~: A |] ==> {u_(a). a:A} ~= 0 & 0 ~: {u_(a). a:A}";
 by (fast_tac (claset() addSIs [not_emptyI, RepFunI]
                 addSEs [not_emptyE, RepFunE]
                 addSDs [sym RS (RepFun_eq_0_iff RS iffD1)]) 1);
 qed "RepRep_conj";
 
-goal thy "!!c. [|c : a; x = c Un {0}; x ~: a |] ==> x - {0} : a";
+Goal "!!c. [|c : a; x = c Un {0}; x ~: a |] ==> x - {0} : a";
 by (hyp_subst_tac 1);
 by (rtac subst_elem 1 THEN (assume_tac 1));
 by (rtac equalityI 1);
@@ -73,14 +73,14 @@
 by (fast_tac (claset() addEs [notE, subst_elem])  1);
 val lemma1_1 = result();
 
-goalw thy [u_def]
+Goalw [u_def]
         "!!a. [| f`(u_(a)) ~: a; f: (PROD B:{u_(a). a:A}. B); a:A |]  \
 \               ==> f`(u_(a))-{0} : a";
 by (fast_tac (claset() addSEs [RepFunI, RepFunE, lemma1_1]
                 addSDs [apply_type]) 1);
 val lemma1_2 = result();
 
-goal thy  "!!A. EX f. f: (PROD B:{u_(a). a:A}. B) ==> EX f. f: (PROD B:A. B)";
+Goal  "!!A. EX f. f: (PROD B:{u_(a). a:A}. B) ==> EX f. f: (PROD B:A. B)";
 by (etac exE 1);
 by (res_inst_tac
         [("x","lam a:A. if(f`(u_(a)) : a, f`(u_(a)), f`(u_(a))-{0})")] exI 1);
@@ -91,21 +91,21 @@
 by (fast_tac (claset() addSEs [lemma1_2]) 1);
 val lemma1 = result();
 
-goalw thy [u_def] "!!a. a~=0 ==> 0: (UN b:u_(a). b)";
+Goalw [u_def] "!!a. a~=0 ==> 0: (UN b:u_(a). b)";
 by (fast_tac (claset() addSEs [not_emptyE] addSIs [UN_I, RepFunI]) 1);
 val lemma2_1 = result();
 
-goal thy "!!A C. [| A~=0; 0~:A |] ==> (INT x:{u_(a). a:A}. UN b:x. b) ~= 0";
+Goal "!!A C. [| A~=0; 0~:A |] ==> (INT x:{u_(a). a:A}. UN b:x. b) ~= 0";
 by (etac not_emptyE 1);
 by (res_inst_tac [("a","0")] not_emptyI 1);
 by (fast_tac (claset() addSIs [INT_I, RepFunI, lemma2_1] addSEs [RepFunE]) 1);
 val lemma2 = result();
 
-goal thy "!!F. (UN f:F. P(f)) ~= 0 ==> F ~= 0";
+Goal "!!F. (UN f:F. P(f)) ~= 0 ==> F ~= 0";
 by (fast_tac (claset() addSEs [not_emptyE]) 1);
 val lemma3 = result();
 
-goalw thy AC_defs "!!Z. AC19 ==> AC1";
+Goalw AC_defs "!!Z. AC19 ==> AC1";
 by (REPEAT (resolve_tac [allI,impI] 1));
 by (excluded_middle_tac "A=0" 1);
 by (fast_tac (claset() addSIs [exI, empty_fun]) 2);
--- a/src/ZF/AC/AC1_AC17.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/AC/AC1_AC17.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -5,13 +5,13 @@
 The proof of AC1 ==> AC17
 *)
 
-goal thy "!!f. f : (PROD X:Pow(A) - {0}. X) ==> f : (Pow(A) - {0} -> A)";
+Goal "!!f. f : (PROD X:Pow(A) - {0}. X) ==> f : (Pow(A) - {0} -> A)";
 by (rtac Pi_type 1 THEN (assume_tac 1));
 by (dtac apply_type 1 THEN (assume_tac 1));
 by (Fast_tac 1);
 val lemma1 = result();
 
-goalw thy AC_defs "!!Z. AC1 ==> AC17";
+Goalw AC_defs "!!Z. AC1 ==> AC17";
 by (rtac allI 1);
 by (rtac ballI 1);
 by (eres_inst_tac [("x","Pow(A)-{0}")] allE 1);
--- a/src/ZF/AC/AC1_WO2.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/AC/AC1_WO2.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -17,7 +17,7 @@
 by (fast_tac (claset() addSIs [prem RS Pi_weaken_type]) 1);
 val lemma1 = uresult() |> standard;
 
-goalw thy [AC1_def, WO2_def, eqpoll_def] "!!Z. AC1 ==> WO2";
+Goalw [AC1_def, WO2_def, eqpoll_def] "!!Z. AC1 ==> WO2";
 by (rtac allI 1);
 by (eres_inst_tac [("x","Pow(A)-{0}")] allE 1);
 by (fast_tac (claset() addSDs [lemma1] addSIs [Ord_Least]) 1);
--- a/src/ZF/AC/AC2_AC6.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/AC/AC2_AC6.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -14,17 +14,17 @@
 (* AC1 ==> AC2                                                            *)
 (* ********************************************************************** *)
 
-goal thy "!!B. [| B:A; f:(PROD X:A. X); 0~:A |]  \
+Goal "!!B. [| B:A; f:(PROD X:A. X); 0~:A |]  \
 \               ==> {f`B} <= B Int {f`C. C:A}";
 by (fast_tac (claset() addSEs [apply_type]) 1);
 val lemma1 = result();
 
-goalw thy [pairwise_disjoint_def]
+Goalw [pairwise_disjoint_def]
         "!!A. [| pairwise_disjoint(A); B:A; C:A; D:B; D:C |] ==> f`B = f`C";
 by (fast_tac (claset() addSEs [equals0D]) 1);
 val lemma2 = result();
 
-goalw thy AC_defs "!!Z. AC1 ==> AC2"; 
+Goalw AC_defs "!!Z. AC1 ==> AC2"; 
 by (rtac allI 1);
 by (rtac impI 1);
 by (REPEAT (eresolve_tac [asm_rl,conjE,allE,exE,impE] 1));
@@ -38,12 +38,12 @@
 (* AC2 ==> AC1                                                            *)
 (* ********************************************************************** *)
 
-goal thy "!!A. 0~:A ==> 0 ~: {B*{B}. B:A}";
+Goal "!!A. 0~:A ==> 0 ~: {B*{B}. B:A}";
 by (fast_tac (claset() addSDs [sym RS (Sigma_empty_iff RS iffD1)]
         addSEs [RepFunE, equals0D]) 1);
 val lemma1 = result();
 
-goal thy "!!A. [| X*{X} Int C = {y}; X:A |]  \
+Goal "!!A. [| X*{X} Int C = {y}; X:A |]  \
 \               ==> (THE y. X*{X} Int C = {y}): X*A";
 by (rtac subst_elem 1);
 by (fast_tac (claset() addSIs [the_equality]
@@ -51,7 +51,7 @@
 by (fast_tac (claset() addSEs [equalityE, make_elim singleton_subsetD]) 1);
 val lemma2 = result();
 
-goal thy "!!A. ALL D:{E*{E}. E:A}. EX y. D Int C = {y}  \
+Goal "!!A. ALL D:{E*{E}. E:A}. EX y. D Int C = {y}  \
 \       ==> (lam x:A. fst(THE z. (x*{x} Int C = {z}))) :  \
 \               (PROD X:A. X) ";
 by (fast_tac (claset() addSEs [lemma2]
@@ -59,7 +59,7 @@
                 addSDs [bspec]) 1);
 val lemma3 = result();
 
-goalw thy (AC_defs@AC_aux_defs) "!!Z. AC2 ==> AC1";
+Goalw (AC_defs@AC_aux_defs) "!!Z. AC2 ==> AC1";
 by (REPEAT (resolve_tac [allI, impI] 1));
 by (REPEAT (eresolve_tac [allE, impE] 1));
 by (fast_tac (claset() addSEs [lemma3]) 2);
@@ -71,11 +71,11 @@
 (* AC1 ==> AC4                                                            *)
 (* ********************************************************************** *)
 
-goal thy "!!R. 0 ~: {R``{x}. x:domain(R)}";
+Goal "!!R. 0 ~: {R``{x}. x:domain(R)}";
 by (fast_tac (claset() addEs [sym RS equals0D]) 1);
 val lemma = result();
 
-goalw thy AC_defs "!!Z. AC1 ==> AC4";
+Goalw AC_defs "!!Z. AC1 ==> AC4";
 by (REPEAT (resolve_tac [allI, impI] 1));
 by (REPEAT (eresolve_tac [allE, lemma RSN (2, impE), exE] 1));
 by (best_tac (claset() addSIs [lam_type] addSEs [apply_type]) 1);
@@ -86,19 +86,19 @@
 (* AC4 ==> AC3                                                            *)
 (* ********************************************************************** *)
 
-goal thy "!!f. f:A->B ==> (UN z:A. {z}*f`z) <= A*Union(B)";
+Goal "!!f. f:A->B ==> (UN z:A. {z}*f`z) <= A*Union(B)";
 by (fast_tac (claset() addSDs [apply_type]) 1);
 val lemma1 = result();
 
-goal thy "!!f. domain(UN z:A. {z}*f(z)) = {a:A. f(a)~=0}";
+Goal "!!f. domain(UN z:A. {z}*f(z)) = {a:A. f(a)~=0}";
 by (fast_tac (claset() addSIs [not_emptyI] addDs [range_type]) 1);
 val lemma2 = result();
 
-goal thy "!!f. x:A ==> (UN z:A. {z}*f(z))``{x} = f(x)";
+Goal "!!f. x:A ==> (UN z:A. {z}*f(z))``{x} = f(x)";
 by (Fast_tac 1);
 val lemma3 = result();
 
-goalw thy AC_defs "!!Z. AC4 ==> AC3";
+Goalw AC_defs "!!Z. AC4 ==> AC3";
 by (REPEAT (resolve_tac [allI,ballI] 1));
 by (REPEAT (eresolve_tac [allE,impE] 1));
 by (etac lemma1 1);
@@ -110,13 +110,13 @@
 (* AC3 ==> AC1                                                            *)
 (* ********************************************************************** *)
 
-goal thy "!!A. b~:A ==> (PROD x:{a:A. id(A)`a~=b}. id(A)`x) = (PROD x:A. x)";
+Goal "!!A. b~:A ==> (PROD x:{a:A. id(A)`a~=b}. id(A)`x) = (PROD x:A. x)";
 by (asm_full_simp_tac (simpset() addsimps [id_def] addcongs [Pi_cong]) 1);
 by (res_inst_tac [("b","A")] subst_context 1);
 by (Fast_tac 1);
 val lemma = result();
 
-goalw thy AC_defs "!!Z. AC3 ==> AC1";
+Goalw AC_defs "!!Z. AC3 ==> AC1";
 by (REPEAT (resolve_tac [allI, impI] 1));
 by (REPEAT (eresolve_tac [allE, ballE] 1));
 by (fast_tac (claset() addSIs [id_type]) 2);
@@ -127,7 +127,7 @@
 (* AC4 ==> AC5                                                            *)
 (* ********************************************************************** *)
 
-goalw thy (range_def::AC_defs) "!!Z. AC4 ==> AC5";
+Goalw (range_def::AC_defs) "!!Z. AC4 ==> AC5";
 by (REPEAT (resolve_tac [allI,ballI] 1));
 by (REPEAT (eresolve_tac [allE,impE] 1));
 by (eresolve_tac [fun_is_rel RS converse_type] 1);
@@ -148,11 +148,11 @@
 (* AC5 ==> AC4, Rubin & Rubin, p. 11                                      *)
 (* ********************************************************************** *)
 
-goal thy "!!A. R <= A*B ==> (lam x:R. fst(x)) : R -> A";
+Goal "!!A. R <= A*B ==> (lam x:R. fst(x)) : R -> A";
 by (fast_tac (claset() addSIs [lam_type, fst_type]) 1);
 val lemma1 = result();
 
-goalw thy [range_def] "!!A. R <= A*B ==> range(lam x:R. fst(x)) = domain(R)";
+Goalw [range_def] "!!A. R <= A*B ==> range(lam x:R. fst(x)) = domain(R)";
 by (rtac equalityI 1);
 by (fast_tac (claset() addSEs [lamE]
                 addEs [subst_elem]
@@ -163,13 +163,13 @@
 by (fast_tac (claset() addSEs [lamI RS subst_elem] addIs [fst_conv RS ssubst]) 1);
 val lemma2 = result();
 
-goal thy "!!A. [| EX f: A->C. P(f,domain(f)); A=B |] ==>  EX f: B->C. P(f,B)";
+Goal "!!A. [| EX f: A->C. P(f,domain(f)); A=B |] ==>  EX f: B->C. P(f,B)";
 by (etac bexE 1);
 by (forward_tac [domain_of_fun] 1);
 by (Fast_tac 1);
 val lemma3 = result();
 
-goal thy "!!g. [| R <= A*B; g: C->R; ALL x:C. (lam z:R. fst(z))` (g`x) = x |] \
+Goal "!!g. [| R <= A*B; g: C->R; ALL x:C. (lam z:R. fst(z))` (g`x) = x |] \
 \               ==> (lam x:C. snd(g`x)): (PROD x:C. R``{x})";
 by (rtac lam_type 1);
 by (dtac apply_type 1 THEN (assume_tac 1));
@@ -180,7 +180,7 @@
 by (Asm_full_simp_tac 1);
 val lemma4 = result();
 
-goalw thy AC_defs "!!Z. AC5 ==> AC4";
+Goalw AC_defs "!!Z. AC5 ==> AC4";
 by (REPEAT (resolve_tac [allI,impI] 1));
 by (REPEAT (eresolve_tac [allE,ballE] 1));
 by (eresolve_tac [lemma1 RSN (2, notE)] 2 THEN (assume_tac 2));
@@ -193,7 +193,7 @@
 (* AC1 <-> AC6                                                            *)
 (* ********************************************************************** *)
 
-goalw thy AC_defs "AC1 <-> AC6";
+Goalw AC_defs "AC1 <-> AC6";
 by (fast_tac (claset() addDs [equals0D] addSEs [not_emptyE]) 1);
 qed "AC1_iff_AC6";
 
--- a/src/ZF/AC/AC7_AC9.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/AC/AC7_AC9.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -17,26 +17,26 @@
 by (Fast_tac 1);
 qed "mem_not_eq_not_mem";
 
-goal thy "!!A. [| 0~:A; B:A |] ==> (nat->Union(A))*B ~= 0";
+Goal "!!A. [| 0~:A; B:A |] ==> (nat->Union(A))*B ~= 0";
 by (fast_tac (claset() addSDs [Sigma_empty_iff RS iffD1]
                 addDs [fun_space_emptyD, mem_not_eq_not_mem]
                 addEs [equals0D]
                 addSIs [equals0I,UnionI]) 1);
 qed "Sigma_fun_space_not0";
 
-goal thy "!!A C. (ALL B:A. B eqpoll C) ==> (ALL B1:A. ALL B2:A. B1 eqpoll B2)";
+Goal "!!A C. (ALL B:A. B eqpoll C) ==> (ALL B1:A. ALL B2:A. B1 eqpoll B2)";
 by (REPEAT (rtac ballI 1));
 by (resolve_tac [bspec RS (bspec RSN (2, eqpoll_sym RSN (2, eqpoll_trans)))] 1
         THEN REPEAT (assume_tac 1));
 qed "all_eqpoll_imp_pair_eqpoll";
 
-goal thy "!!A. [| ALL a:A. if(a=b, P(a), Q(a)) = if(a=b, R(a), S(a)); b:A   \
+Goal "!!A. [| ALL a:A. if(a=b, P(a), Q(a)) = if(a=b, R(a), S(a)); b:A   \
 \       |] ==> P(b)=R(b)";
 by (dtac bspec 1 THEN (assume_tac 1));
 by (Asm_full_simp_tac 1);
 qed "if_eqE1";
 
-goal thy "!!A. ALL a:A. if(a=b, P(a), Q(a)) = if(a=b, R(a), S(a))  \
+Goal "!!A. ALL a:A. if(a=b, P(a), Q(a)) = if(a=b, R(a), S(a))  \
 \       ==> ALL a:A. a~=b --> Q(a)=S(a)";
 by (rtac ballI 1);
 by (rtac impI 1);
@@ -44,13 +44,13 @@
 by (Asm_full_simp_tac 1);
 qed "if_eqE2";
 
-goal thy "!!A. [| (lam x:A. f(x))=(lam x:A. g(x)); a:A |] ==> f(a)=g(a)";
+Goal "!!A. [| (lam x:A. f(x))=(lam x:A. g(x)); a:A |] ==> f(a)=g(a)";
 by (fast_tac (claset() addDs [subsetD]
                 addSIs [lamI]
                 addEs [equalityE, lamE]) 1);
 qed "lam_eqE";
 
-goalw thy [inj_def]
+Goalw [inj_def]
         "!!A. C:A ==> (lam g:(nat->Union(A))*C.  \
 \               (lam n:nat. if(n=0, snd(g), fst(g)`(n #- 1))))  \
 \               : inj((nat->Union(A))*C, (nat->Union(A)) ) ";
@@ -70,7 +70,7 @@
 by (asm_full_simp_tac (simpset() addsimps [succ_not_0 RS if_not_P]) 1);
 val lemma = result();
 
-goal thy "!!A. [| C:A; 0~:A |] ==> (nat->Union(A)) * C eqpoll (nat->Union(A))";
+Goal "!!A. [| C:A; 0~:A |] ==> (nat->Union(A)) * C eqpoll (nat->Union(A))";
 by (rtac eqpollI 1);
 by (fast_tac (claset() addSEs [prod_lepoll_self, not_sym RS not_emptyE,
                 subst_elem] addEs [swap]) 2);
@@ -83,7 +83,7 @@
 (* AC6 ==> AC7                                                            *)
 (* ********************************************************************** *)
 
-goalw thy AC_defs "!!Z. AC6 ==> AC7";
+Goalw AC_defs "!!Z. AC6 ==> AC7";
 by (Fast_tac 1);
 qed "AC6_AC7";
 
@@ -93,27 +93,27 @@
 (* the proof.                                                             *)
 (* ********************************************************************** *)
 
-goal thy "!!y. y: (PROD B:A. Y*B) ==> (lam B:A. snd(y`B)): (PROD B:A. B)";
+Goal "!!y. y: (PROD B:A. Y*B) ==> (lam B:A. snd(y`B)): (PROD B:A. B)";
 by (fast_tac (claset() addSIs [lam_type, snd_type, apply_type]) 1);
 val lemma1_1 = result();
 
-goal thy "!!A. y: (PROD B:{Y*C. C:A}. B)  \
+Goal "!!A. y: (PROD B:{Y*C. C:A}. B)  \
 \               ==> (lam B:A. y`(Y*B)): (PROD B:A. Y*B)";
 by (fast_tac (claset() addSIs [lam_type, apply_type]) 1);
 val lemma1_2 = result();
 
-goal thy "!!A. (PROD B:{(nat->Union(A))*C. C:A}. B) ~= 0  \
+Goal "!!A. (PROD B:{(nat->Union(A))*C. C:A}. B) ~= 0  \
 \               ==> (PROD B:A. B) ~= 0";
 by (fast_tac (claset() addSIs [equals0I,lemma1_1, lemma1_2]
                 addSEs [equals0D]) 1);
 val lemma1 = result();
 
-goal thy "!!A. 0 ~: A ==> 0 ~: {(nat -> Union(A)) * C. C:A}";
+Goal "!!A. 0 ~: A ==> 0 ~: {(nat -> Union(A)) * C. C:A}";
 by (fast_tac (claset() addEs [RepFunE,
                 Sigma_fun_space_not0 RS not_sym RS notE]) 1);
 val lemma2 = result();
 
-goalw thy AC_defs "!!Z. AC7 ==> AC6";
+Goalw AC_defs "!!Z. AC7 ==> AC6";
 by (rtac allI 1);
 by (rtac impI 1);
 by (excluded_middle_tac "A=0" 1);
@@ -131,7 +131,7 @@
 (* AC1 ==> AC8                                                            *)
 (* ********************************************************************** *)
 
-goalw thy [eqpoll_def]
+Goalw [eqpoll_def]
         "!!A. ALL B:A. EX B1 B2. B=<B1,B2> & B1 eqpoll B2  \
 \       ==> 0 ~: { bij(fst(B),snd(B)). B:A }";
 by (rtac notI 1);
@@ -142,13 +142,13 @@
 by (Asm_full_simp_tac 1);
 val lemma1 = result();
 
-goal thy "!!A. [| f: (PROD X:RepFun(A,p). X); D:A |]  \
+Goal "!!A. [| f: (PROD X:RepFun(A,p). X); D:A |]  \
 \               ==> (lam x:A. f`p(x))`D : p(D)";
 by (resolve_tac [beta RS ssubst] 1 THEN (assume_tac 1));
 by (fast_tac (claset() addSEs [apply_type]) 1);
 val lemma2 = result();
 
-goalw thy AC_defs "!!Z. AC1 ==> AC8";
+Goalw AC_defs "!!Z. AC1 ==> AC8";
 by (rtac allI 1);
 by (etac allE 1);
 by (rtac impI 1);
@@ -164,16 +164,16 @@
 (*    AC8 ==> AC1 and AC1 ==> AC9                                         *)
 (* ********************************************************************** *)
 
-goal thy "!!A. ALL B1:A. ALL B2:A. B1 eqpoll B2 ==>  \
+Goal "!!A. ALL B1:A. ALL B2:A. B1 eqpoll B2 ==>  \
 \               ALL B:A*A. EX B1 B2. B=<B1,B2> & B1 eqpoll B2";
 by (Fast_tac 1);
 val lemma1 = result();
 
-goal thy "!!f. f:bij(fst(<a,b>),snd(<a,b>)) ==> f:bij(a,b)";
+Goal "!!f. f:bij(fst(<a,b>),snd(<a,b>)) ==> f:bij(a,b)";
 by (Asm_full_simp_tac 1);
 val lemma2 = result();
 
-goalw thy AC_defs "!!Z. AC8 ==> AC9";
+Goalw AC_defs "!!Z. AC8 ==> AC9";
 by (rtac allI 1);
 by (rtac impI 1);
 by (etac allE 1);
@@ -196,7 +196,7 @@
         ((prod_commute_eqpoll RS eqpoll_imp_lepoll) RSN (2,lepoll_trans));
 
 
-goal thy "!!A. [| 0~:A; A~=0 |]  \
+Goal "!!A. [| 0~:A; A~=0 |]  \
 \       ==> ALL B1: ({((nat->Union(A))*B)*nat. B:A}  \
 \               Un {cons(0,((nat->Union(A))*B)*nat). B:A}).  \
 \       ALL B2: ({((nat->Union(A))*B)*nat. B:A}  \
@@ -209,7 +209,7 @@
                         (2, prod_eqpoll_cong), Sigma_fun_space_eqpoll]) 1);
 val lemma1 = result();
 
-goal thy "!!A. ALL B1:{(F*B)*N. B:A} Un {cons(0,(F*B)*N). B:A}.  \
+Goal "!!A. ALL B1:{(F*B)*N. B:A} Un {cons(0,(F*B)*N). B:A}.  \
 \       ALL B2:{(F*B)*N. B:A}  \
 \       Un {cons(0,(F*B)*N). B:A}. f`<B1,B2> : bij(B1, B2)  \
 \       ==> (lam B:A. snd(fst((f`<cons(0,(F*B)*N),(F*B)*N>)`0))) :  \
@@ -221,7 +221,7 @@
 by (fast_tac (claset() addSIs [fun_weaken_type, bij_is_fun]) 1);
 val lemma2 = result();
 
-goalw thy AC_defs "!!Z. AC9 ==> AC1";
+Goalw AC_defs "!!Z. AC9 ==> AC1";
 by (rtac allI 1);
 by (rtac impI 1);
 by (etac allE 1);
--- a/src/ZF/AC/AC_Equiv.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/AC/AC_Equiv.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -38,7 +38,7 @@
 (*             lemmas concerning FOL and pure ZF theory                   *)
 (* ********************************************************************** *)
 
-goal thy "!!X. (A->X)=0 ==> X=0";
+Goal "!!X. (A->X)=0 ==> X=0";
 by (fast_tac (claset() addSIs [equals0I] addEs [lam_type RSN (2, equals0D)]) 1);
 qed "fun_space_emptyD";
 
@@ -95,12 +95,12 @@
 by (asm_simp_tac (simpset() addsimps [Card_is_Ord RS nat_le_infinite_Ord]) 1);
 qed "Inf_Card_is_InfCard";
 
-goal thy "(THE z. {x}={z}) = x";
+Goal "(THE z. {x}={z}) = x";
 by (fast_tac (claset() addSIs [the_equality]
                 addSEs [singleton_eq_iff RS iffD1 RS sym]) 1);
 qed "the_element";
 
-goal thy "(lam x:A. {x}) : bij(A, {{x}. x:A})";
+Goal "(lam x:A. {x}) : bij(A, {{x}. x:A})";
 by (res_inst_tac [("d","%z. THE x. z={x}")] lam_bijective 1);
 by (TRYALL (eresolve_tac [RepFunI, RepFunE]));
 by (REPEAT (asm_full_simp_tac (simpset() addsimps [the_element]) 1));
@@ -118,15 +118,15 @@
         addSIs [major RS CollectD1 RS Pi_type, major RS CollectD2]) 1);
 qed "inj_strengthen_type";
 
-goal thy "A*B=0 <-> A=0 | B=0";
+Goal "A*B=0 <-> A=0 | B=0";
 by (fast_tac (claset() addSIs [equals0I] addEs [equals0D]) 1);
 qed "Sigma_empty_iff";
 
-goalw thy [Finite_def] "!!n. n:nat ==> Finite(n)";
+Goalw [Finite_def] "!!n. n:nat ==> Finite(n)";
 by (fast_tac (claset() addSIs [eqpoll_refl]) 1);
 qed "nat_into_Finite";
 
-goalw thy [Finite_def] "~Finite(nat)";
+Goalw [Finite_def] "~Finite(nat)";
 by (fast_tac (claset() addSDs [eqpoll_imp_lepoll]
                 addIs [Ord_nat RSN (2, ltI) RS lt_not_lepoll RS notE]) 1);
 qed "nat_not_Finite";
@@ -137,7 +137,7 @@
 (* Another elimination rule for EX!                                       *)
 (* ********************************************************************** *)
 
-goal thy "!!x. [| EX! x. P(x); P(x); P(y) |] ==> x=y";
+Goal "!!x. [| EX! x. P(x); P(x); P(y) |] ==> x=y";
 by (etac ex1E 1);
 by (res_inst_tac [("b","xa")] (sym RSN (2, trans)) 1);
 by (Fast_tac 1);
@@ -148,7 +148,7 @@
 (* image of a surjection                                                  *)
 (* ********************************************************************** *)
 
-goalw thy [surj_def] "!!f. f : surj(A, B) ==> f``A = B";
+Goalw [surj_def] "!!f. f : surj(A, B) ==> f``A = B";
 by (etac CollectE 1);
 by (resolve_tac [subset_refl RSN (2, image_fun) RS ssubst] 1 
     THEN (assume_tac 1));
@@ -156,11 +156,11 @@
 qed "surj_image_eq";
 
 
-goal thy "!!y. succ(x) lepoll y ==> y ~= 0";
+Goal "!!y. succ(x) lepoll y ==> y ~= 0";
 by (fast_tac (claset() addSDs [lepoll_0_is_0]) 1);
 qed "succ_lepoll_imp_not_empty";
 
-goal thy "!!x. x eqpoll succ(n) ==> x ~= 0";
+Goal "!!x. x eqpoll succ(n) ==> x ~= 0";
 by (fast_tac (claset() addSEs [eqpoll_sym RS eqpoll_0_is_0 RS succ_neq_0]) 1);
 qed "eqpoll_succ_imp_not_empty";
 
--- a/src/ZF/AC/Cardinal_aux.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/AC/Cardinal_aux.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -26,7 +26,7 @@
 by (fast_tac (claset() addSDs [lepoll_imp_ex_le_eqpoll] addSEs [leE]) 1);
 qed "lesspoll_imp_ex_lt_eqpoll";
 
-goalw thy [InfCard_def] "!!i. [| ~Finite(i); Ord(i) |] ==> InfCard(|i|)";
+Goalw [InfCard_def] "!!i. [| ~Finite(i); Ord(i) |] ==> InfCard(|i|)";
 by (rtac conjI 1);
 by (rtac Card_cardinal 1);
 by (resolve_tac [Card_nat RS (Card_def RS def_imp_iff RS iffD1 RS ssubst)] 1);
@@ -39,12 +39,12 @@
         asm_rl, lepoll_refl] MRS (prod_lepoll_mono RSN (2, lepoll_trans))
         |> standard;
 
-goal thy "!!A. [| A lepoll B; 2 lepoll A |] ==> A+B lepoll A*B";
+Goal "!!A. [| A lepoll B; 2 lepoll A |] ==> A+B lepoll A*B";
 by (REPEAT (ares_tac [[sum_lepoll_mono, sum_lepoll_prod]
                 MRS lepoll_trans, lepoll_refl] 1));
 qed "lepoll_imp_sum_lepoll_prod";
 
-goal thy "!!A. [| A eqpoll i; B eqpoll i; ~Finite(i); Ord(i) |]  \
+Goal "!!A. [| A eqpoll i; B eqpoll i; ~Finite(i); Ord(i) |]  \
 \               ==> A Un B eqpoll i";
 by (rtac eqpollI 1);
 by (eresolve_tac [subset_imp_lepoll RSN (2, eqpoll_sym RS eqpoll_imp_lepoll
@@ -78,22 +78,22 @@
 by (fast_tac (claset() addSIs [the_equality] addEs [equalityE]) 1);
 qed "paired_bij_lemma";
 
-goal thy "(lam y:{{y,z}. y:x}. if(y-{z}=0, z, THE w. y-{z}={w}))  \
+Goal "(lam y:{{y,z}. y:x}. if(y-{z}=0, z, THE w. y-{z}={w}))  \
 \               : bij({{y,z}. y:x}, x)";
 by (res_inst_tac [("d","%a. {a,z}")] lam_bijective 1);
 by (TRYALL (fast_tac (claset() addSEs [RepFunE] addSIs [RepFunI] 
                 addss (simpset() addsimps [paired_bij_lemma]))));
 qed "paired_bij";
 
-goalw thy [eqpoll_def] "{{y,z}. y:x} eqpoll x";
+Goalw [eqpoll_def] "{{y,z}. y:x} eqpoll x";
 by (fast_tac (claset() addSIs [paired_bij]) 1);
 qed "paired_eqpoll";
 
-goal thy "!!A. EX B. B eqpoll A & B Int C = 0";
+Goal "!!A. EX B. B eqpoll A & B Int C = 0";
 by (fast_tac (claset() addSIs [paired_eqpoll, equals0I] addEs [mem_asym]) 1);
 qed "ex_eqpoll_disjoint";
 
-goal thy "!!A. [| A lepoll i; B lepoll i; ~Finite(i); Ord(i) |]  \
+Goal "!!A. [| A lepoll i; B lepoll i; ~Finite(i); Ord(i) |]  \
 \               ==> A Un B lepoll i";
 by (res_inst_tac [("A1","i"), ("C1","i")] (ex_eqpoll_disjoint RS exE) 1);
 by (etac conjE 1);
@@ -105,7 +105,7 @@
         THEN (REPEAT (assume_tac 1)));
 qed "Un_lepoll_Inf_Ord";
 
-goal thy "!!P. [| P(i); i:j; Ord(j) |] ==> (LEAST i. P(i)) : j";
+Goal "!!P. [| P(i); i:j; Ord(j) |] ==> (LEAST i. P(i)) : j";
 by (eresolve_tac [Least_le RS leE] 1);
 by (etac Ord_in_Ord 1 THEN (assume_tac 1));
 by (etac ltE 1);
@@ -113,7 +113,7 @@
 by (etac subst_elem 1 THEN (assume_tac 1));
 qed "Least_in_Ord";
 
-goal thy "!!x. [| well_ord(x,r); y<=x; y lepoll succ(n); n:nat |]  \
+Goal "!!x. [| well_ord(x,r); y<=x; y lepoll succ(n); n:nat |]  \
 \       ==> y-{THE b. first(b,y,r)} lepoll n";
 by (res_inst_tac [("Q","y=0")] (excluded_middle RS disjE) 1);
 by (fast_tac (claset() addSIs [Diff_sing_lepoll, the_first_in]) 1);
@@ -122,18 +122,18 @@
 by (Fast_tac 1);
 qed "Diff_first_lepoll";
 
-goal thy "(UN x:X. P(x)) <= (UN x:X. P(x)-Q(x)) Un (UN x:X. Q(x))";
+Goal "(UN x:X. P(x)) <= (UN x:X. P(x)-Q(x)) Un (UN x:X. Q(x))";
 by (Fast_tac 1);
 qed "UN_subset_split";
 
-goalw thy [lepoll_def] "!!a. Ord(a) ==> (UN x:a. {P(x)}) lepoll a";
+Goalw [lepoll_def] "!!a. Ord(a) ==> (UN x:a. {P(x)}) lepoll a";
 by (res_inst_tac [("x","lam z:(UN x:a. {P(x)}). (LEAST i. P(i)=z)")] exI 1);
 by (res_inst_tac [("d","%z. P(z)")] lam_injective 1);
 by (fast_tac (claset() addSIs [Least_in_Ord]) 1);
 by (fast_tac (claset() addIs [LeastI] addSEs [Ord_in_Ord]) 1);
 qed "UN_sing_lepoll";
 
-goal thy "!!a T. [| well_ord(T, R); ~Finite(a); Ord(a); n:nat |] ==>  \
+Goal "!!a T. [| well_ord(T, R); ~Finite(a); Ord(a); n:nat |] ==>  \
 \       ALL f. (ALL b:a. f`b lepoll n & f`b <= T) --> (UN b:a. f`b) lepoll a";
 by (etac nat_induct 1);
 by (rtac allI 1);
@@ -155,13 +155,13 @@
 by (etac UN_sing_lepoll 1);
 qed "UN_fun_lepoll_lemma";
 
-goal thy "!!a f. [| ALL b:a. f`b lepoll n & f`b <= T; well_ord(T, R);  \
+Goal "!!a f. [| ALL b:a. f`b lepoll n & f`b <= T; well_ord(T, R);  \
 \       ~Finite(a); Ord(a); n:nat |] ==> (UN b:a. f`b) lepoll a";
 by (eresolve_tac [UN_fun_lepoll_lemma RS allE] 1 THEN (REPEAT (assume_tac 1)));
 by (Fast_tac 1);
 qed "UN_fun_lepoll";
 
-goal thy "!!a f. [| ALL b:a. F(b) lepoll n & F(b) <= T; well_ord(T, R);  \
+Goal "!!a f. [| ALL b:a. F(b) lepoll n & F(b) <= T; well_ord(T, R);  \
 \       ~Finite(a); Ord(a); n:nat |] ==> (UN b:a. F(b)) lepoll a";
 by (rtac impE 1 THEN (assume_tac 3));
 by (res_inst_tac [("f","lam b:a. F(b)")] (UN_fun_lepoll) 2 
@@ -170,7 +170,7 @@
 by (Asm_full_simp_tac 1);
 qed "UN_lepoll";
 
-goal thy "!!a. Ord(a) ==> (UN b:a. F(b)) = (UN b:a. F(b) - (UN c:b. F(c)))";
+Goal "!!a. Ord(a) ==> (UN b:a. F(b)) = (UN b:a. F(b) - (UN c:b. F(c)))";
 by (rtac equalityI 1);
 by (Fast_tac 2);
 by (rtac subsetI 1);
@@ -185,7 +185,7 @@
 by (eresolve_tac [Ord_Least RSN (2, ltI)] 1);
 qed "UN_eq_UN_Diffs";
 
-goalw thy [eqpoll_def] "!!A B. A Int B = 0 ==> A Un B eqpoll A + B";
+Goalw [eqpoll_def] "!!A B. A Int B = 0 ==> A Un B eqpoll A + B";
 by (res_inst_tac [("x","lam a:A Un B. if(a:A,Inl(a),Inr(a))")] exI 1);
 by (res_inst_tac [("d","%z. case(%x. x, %x. x, z)")] lam_bijective 1);
 by (fast_tac (claset() addSIs [if_type, InlI, InrI]) 1);
@@ -195,7 +195,7 @@
 by (fast_tac (claset() addDs [equals0D]) 1);
 qed "disj_Un_eqpoll_sum";
 
-goalw thy [lepoll_def, eqpoll_def]
+Goalw [lepoll_def, eqpoll_def]
         "!!X a. a lepoll X ==> EX Y. Y<=X & a eqpoll Y";
 by (etac exE 1);
 by (forward_tac [subset_refl RSN (2, restrict_bij)] 1);
@@ -207,7 +207,7 @@
 (* Diff_lesspoll_eqpoll_Card                                              *)
 (* ********************************************************************** *)
 
-goal thy "!!A B. [| A eqpoll a; ~Finite(a); Card(a); B lesspoll a;  \
+Goal "!!A B. [| A eqpoll a; ~Finite(a); Card(a); B lesspoll a;  \
 \               A-B lesspoll a |] ==> P";
 by (REPEAT (eresolve_tac [lesspoll_imp_ex_lt_eqpoll RS exE,
         Card_is_Ord, conjE] 1));
@@ -237,7 +237,7 @@
 by (fast_tac (claset() addSDs [lesspoll_def RS def_imp_iff RS iffD1]) 1);
 qed "Diff_lesspoll_eqpoll_Card_lemma";
 
-goal thy "!!A. [| A eqpoll a; ~Finite(a); Card(a); B lesspoll a |]  \
+Goal "!!A. [| A eqpoll a; ~Finite(a); Card(a); B lesspoll a |]  \
 \       ==> A - B eqpoll a";
 by (rtac swap 1 THEN (Fast_tac 1));
 by (rtac Diff_lesspoll_eqpoll_Card_lemma 1 THEN (REPEAT (assume_tac 1)));
--- a/src/ZF/AC/DC.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/AC/DC.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -33,12 +33,12 @@
 (*                                                                        *)
 (* ********************************************************************** *)
 
-goal thy "{<z1,z2>:XX*XX. domain(z2)=succ(domain(z1))  \
+Goal "{<z1,z2>:XX*XX. domain(z2)=succ(domain(z1))  \
 \       & restrict(z2, domain(z1)) = z1} <= XX*XX";
 by (Fast_tac 1);
 val lemma1_1 = result();
 
-goal thy "!!X. ALL Y:Pow(X). Y lesspoll nat --> (EX x:X. <Y, x> : R)  \
+Goal "!!X. ALL Y:Pow(X). Y lesspoll nat --> (EX x:X. <Y, x> : R)  \
 \       ==> {<z1,z2>: (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R}) *  \
 \               (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R}).  \
 \               domain(z2)=succ(domain(z1))  \
@@ -58,7 +58,7 @@
                 singleton_0]) 1);
 val lemma1_2 = result();
 
-goal thy "!!X. ALL Y:Pow(X). Y lesspoll nat --> (EX x:X. <Y, x> : R)  \
+Goal "!!X. ALL Y:Pow(X). Y lesspoll nat --> (EX x:X. <Y, x> : R)  \
 \       ==> range({<z1,z2>: (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R}) *  \
 \               (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R}).  \
 \               domain(z2)=succ(domain(z1))  \
@@ -92,7 +92,7 @@
         addsimps [domain_cons, domain_of_fun, succ_def, restrict_cons_eq]) 1);
 val lemma1_3 = result();
 
-goal thy "!!X. [| XX = (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R});  \
+Goal "!!X. [| XX = (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R});  \
 \       RR = {<z1,z2>:XX*XX. domain(z2)=succ(domain(z1))  \
 \       & restrict(z2, domain(z1)) = z1};  \
 \       ALL Y:Pow(X). Y lesspoll nat --> (EX x:X. <Y, x> : R)  \
@@ -100,7 +100,7 @@
 by (fast_tac (claset() addSIs [lemma1_1] addSEs [lemma1_2, lemma1_3]) 1);
 val lemma1 = result();
 
-goal thy
+Goal
 "!!X.[| XX = (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R});  \
 \       ALL n:nat. <f`n, f`succ(n)> :  \
 \       {<z1,z2>:XX*XX. domain(z2)=succ(domain(z1))  \
@@ -136,7 +136,7 @@
 by (deepen_tac (claset() addDs [domain_of_fun RS sym RS trans]) 0 1);
 val lemma2 = result();
 
-goal thy 
+Goal 
 "!!X.[| XX = (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R});  \
 \       ALL n:nat. <f`n, f`succ(n)> :  \
 \       {<z1,z2>:XX*XX. domain(z2)=succ(domain(z1))  \
@@ -162,12 +162,12 @@
         addSEs [bexE, nat_into_Ord RSN (2, OrdmemD) RS subsetD]) 1);
 val lemma3_1 = result();
 
-goal thy "!!n. ALL x:n. f`succ(n)`x = f`succ(x)`x   \
+Goal "!!n. ALL x:n. f`succ(n)`x = f`succ(x)`x   \
 \       ==> {f`succ(x)`x. x:n} = {f`succ(n)`x. x:n}";
 by (Asm_full_simp_tac 1);
 val lemma3_2 = result();
 
-goal thy "!!X. [| XX = (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R});  \
+Goal "!!X. [| XX = (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R});  \
 \       ALL n:nat. <f`n, f`succ(n)> :  \
 \       {<z1,z2>:XX*XX. domain(z2)=succ(domain(z1))  \
 \       & restrict(z2, domain(z1)) = z1};  \
@@ -186,12 +186,12 @@
                             (2, image_fun RS sym)]) 1);
 val lemma3 = result();
 
-goal thy "!!f. [| f:A->B; B<=C |] ==> f:A->C";
+Goal "!!f. [| f:A->B; B<=C |] ==> f:A->C";
 by (rtac Pi_type 1 THEN (assume_tac 1));
 by (fast_tac (claset() addSEs [apply_type]) 1);
 qed "fun_type_gen";
 
-goalw thy [DC_def, DC0_def] "!!Z. DC0 ==> DC(nat)";
+Goalw [DC_def, DC0_def] "!!Z. DC0 ==> DC(nat)";
 by (REPEAT (resolve_tac [allI, impI] 1));
 by (REPEAT (eresolve_tac [conjE, allE] 1));
 by (eresolve_tac [[refl, refl] MRS lemma1 RSN (2, impE)] 1
@@ -241,13 +241,13 @@
                                                                           
 ************************************************************************* *)
 
-goalw thy [lesspoll_def, Finite_def]
+Goalw [lesspoll_def, Finite_def]
         "!!A. A lesspoll nat ==> Finite(A)";
 by (fast_tac (claset() addSDs [ltD, lepoll_imp_ex_le_eqpoll]
         addSIs [Ord_nat]) 1);
 qed "lesspoll_nat_is_Finite";
 
-goal thy "!!n. n:nat ==> ALL A. (A eqpoll n & A <= X) --> A : Fin(X)";
+Goal "!!n. n:nat ==> ALL A. (A eqpoll n & A <= X) --> A : Fin(X)";
 by (etac nat_induct 1);
 by (rtac allI 1);
 by (fast_tac (claset() addSIs [Fin.emptyI]
@@ -266,7 +266,7 @@
 by (asm_full_simp_tac (simpset() addsimps [cons_Diff]) 1);
 qed "Finite_Fin_lemma";
 
-goalw thy [Finite_def] "!!A. [| Finite(A); A <= X |] ==> A : Fin(X)";
+Goalw [Finite_def] "!!A. [| Finite(A); A <= X |] ==> A : Fin(X)";
 by (etac bexE 1);
 by (dtac Finite_Fin_lemma 1);
 by (etac allE 1);
@@ -275,14 +275,14 @@
 by (Fast_tac 1);
 qed "Finite_Fin";
 
-goal thy "!!x. x: X  \
+Goal "!!x. x: X  \
 \ ==> {<0,x>}: (UN n:nat. {f:succ(n)->X. ALL k:n. <f`k, f`succ(k)> : R})";
 by (rtac (nat_0I RS UN_I) 1);
 by (fast_tac (claset() addSIs [singleton_fun RS Pi_type]
         addss (simpset() addsimps [singleton_0 RS sym])) 1);
 qed "singleton_in_funs";
 
-goal thy
+Goal
         "!!X. [| XX = (UN n:nat.  \
 \               {f:succ(n)->domain(R). ALL k:n. <f`k, f`succ(k)> : R});  \
 \       RR = {<z1,z2>:Fin(XX)*XX. (domain(z2)=succ(UN f:z1. domain(f))  \
@@ -307,31 +307,31 @@
                                      cons_fun_type2, empty_fun]) 1);
 val lemma4 = result();
 
-goal thy "!!f. [| f:nat->X; n:nat |] ==>  \
+Goal "!!f. [| f:nat->X; n:nat |] ==>  \
 \       (UN x:f``succ(n). P(x)) =  P(f`n) Un (UN x:f``n. P(x))";
 by (asm_full_simp_tac (simpset()
         addsimps [Ord_nat RSN (2, OrdmemD) RSN (2, image_fun),
         [nat_succI, Ord_nat] MRS OrdmemD RSN (2, image_fun)]) 1);
 qed "UN_image_succ_eq";
 
-goal thy "!!f. [| (UN x:f``n. P(x)) = y; P(f`n) = succ(y);  \
+Goal "!!f. [| (UN x:f``n. P(x)) = y; P(f`n) = succ(y);  \
 \       f:nat -> X; n:nat |] ==> (UN x:f``succ(n). P(x)) = succ(y)";
 by (asm_full_simp_tac (simpset() addsimps [UN_image_succ_eq]) 1);
 by (Fast_tac 1);
 qed "UN_image_succ_eq_succ";
 
-goal thy "!!f. [| f:succ(n) -> D;  n:nat;  \
+Goal "!!f. [| f:succ(n) -> D;  n:nat;  \
 \       domain(f)=succ(x); x=y |] ==> f`y : D";
 by (fast_tac (claset() addEs [apply_type]
         addSDs [[sym, domain_of_fun] MRS trans]) 1);
 qed "apply_domain_type";
 
-goal thy "!!f. [| f : nat -> X; n:nat |] ==> f``succ(n) = cons(f`n, f``n)";
+Goal "!!f. [| f : nat -> X; n:nat |] ==> f``succ(n) = cons(f`n, f``n)";
 by (asm_full_simp_tac (simpset()
         addsimps [nat_succI, Ord_nat RSN (2, OrdmemD), image_fun]) 1);
 qed "image_fun_succ";
 
-goal thy "!!f. [| domain(f`n) = succ(u); f : nat -> (UN n:nat.  \
+Goal "!!f. [| domain(f`n) = succ(u); f : nat -> (UN n:nat.  \
 \       {f:succ(n)->domain(R). ALL k:n. <f`k, f`succ(k)> : R});  \
 \       u=k; n:nat  \
 \       |] ==> f`n : succ(k) -> domain(R)";
@@ -339,7 +339,7 @@
 by (fast_tac (claset() addEs [UN_E, domain_eq_imp_fun_type]) 1);
 qed "f_n_type";
 
-goal thy "!!f. [| f : nat -> (UN n:nat.  \
+Goal "!!f. [| f : nat -> (UN n:nat.  \
 \       {f:succ(n)->domain(R). ALL k:n. <f`k, f`succ(k)> : R});  \
 \       domain(f`n) = succ(k); n:nat  \
 \       |] ==> ALL i:k. <f`n`i, f`n`succ(i)> : R";
@@ -350,7 +350,7 @@
 by (Asm_full_simp_tac 1);
 qed "f_n_pairs_in_R";
 
-goalw thy [restrict_def]
+Goalw [restrict_def]
         "!!f. [| restrict(f, domain(x))=x; f:n->X; domain(x) <= n  \
 \       |] ==> restrict(cons(<n, y>, f), domain(x)) = x";
 by (eresolve_tac [sym RS trans RS sym] 1);
@@ -360,7 +360,7 @@
 by (asm_full_simp_tac (simpset() addsimps [subsetD RS cons_val_k]) 1);
 qed "restrict_cons_eq_restrict";
 
-goal thy "!!f. [| ALL x:f``n. restrict(f`n, domain(x))=x;  \
+Goal "!!f. [| ALL x:f``n. restrict(f`n, domain(x))=x;  \
 \       f : nat -> (UN n:nat.  \
 \       {f:succ(n)->domain(R). ALL k:n. <f`k, f`succ(k)> : R});  \
 \       n:nat; domain(f`n) = succ(n);  \
@@ -375,7 +375,7 @@
 by (fast_tac (claset() addSEs [restrict_cons_eq_restrict]) 1);
 qed "all_in_image_restrict_eq";
 
-goal thy
+Goal
 "!!X.[| ALL b<nat. <f``b, f`b> :  \
 \       {<z1,z2>:Fin(XX)*XX. (domain(z2)=succ(UN f:z1. domain(f))  & \
 \                            (ALL f:z1. restrict(z2, domain(f)) = f)) |  \
@@ -436,7 +436,7 @@
 qed "simplify_recursion";
 
 
-goal thy "!!X. [| XX = (UN n:nat.  \
+Goal "!!X. [| XX = (UN n:nat.  \
 \               {f:succ(n)->domain(R). ALL k:n. <f`k, f`succ(k)> : R});  \
 \       ALL b<nat. <f``b, f`b> :  \
 \       {<z1,z2>:Fin(XX)*XX. (domain(z2)=succ(UN f:z1. domain(f))  \
@@ -455,7 +455,7 @@
 by (fast_tac (FOL_cs addSEs [conjE, f_n_pairs_in_R, trans, subst_context]) 1);
 val lemma2 = result();
 
-goal thy 
+Goal 
 "!!n. [| XX = (UN n:nat.  \
 \                {f:succ(n)->domain(R). ALL k:n. <f`k, f`succ(k)> : R});  \
 \       ALL b<nat. <f``b, f`b> :  \
@@ -478,7 +478,7 @@
 val lemma3 = result();
 
 
-goalw thy [DC_def, DC0_def] "!!Z. DC(nat) ==> DC0";
+Goalw [DC_def, DC0_def] "!!Z. DC(nat) ==> DC0";
 by (REPEAT (resolve_tac [allI, impI] 1));
 by (REPEAT (eresolve_tac [asm_rl, conjE, ex_in_domain RS exE, allE] 1));
 by (eresolve_tac [[refl, refl] MRS lemma4 RSN (2, impE)] 1
@@ -501,7 +501,7 @@
 (* ALL K. Card(K) --> DC(K) ==> WO3                                       *)
 (* ********************************************************************** *)
 
-goalw thy [lesspoll_def]
+Goalw [lesspoll_def]
         "!!A. [| ~ A lesspoll B; C lesspoll B |] ==> A - C ~= 0";
 by (fast_tac (claset() addSDs [Diff_eq_0_iff RS iffD1 RS subset_imp_lepoll]
         addSIs [eqpollI] addEs [notE] addSEs [eqpollE, lepoll_trans]) 1);
@@ -518,11 +518,11 @@
 by (REPEAT (fast_tac (claset() addDs [not_eq, not_eq RS not_sym]) 1));
 qed "fun_Ord_inj";
 
-goal thy "!!a. [| f:X->Y; A<=X; a:A |] ==> f`a : f``A";
+Goal "!!a. [| f:X->Y; A<=X; a:A |] ==> f`a : f``A";
 by (fast_tac (claset() addSEs [image_fun RS ssubst]) 1);
 qed "value_in_image";
 
-goalw thy [DC_def, WO3_def]
+Goalw [DC_def, WO3_def]
         "!!Z. ALL K. Card(K) --> DC(K) ==> WO3";
 by (rtac allI 1);
 by (excluded_middle_tac "A lesspoll Hartog(A)" 1);
@@ -551,7 +551,7 @@
 (* WO1 ==> ALL K. Card(K) --> DC(K)                                       *)
 (* ********************************************************************** *)
 
-goal thy
+Goal
         "!!a. [| Ord(a); b:a |] ==> (lam x:a. P(x))``b = (lam x:b. P(x))``b";
 by (rtac images_eq 1);
 by (REPEAT (fast_tac (claset() addSEs [RepFunI, OrdmemD]
@@ -562,23 +562,23 @@
 by (Asm_full_simp_tac 1);
 qed "lam_images_eq";
 
-goalw thy [lesspoll_def] "!!K. [| Card(K); b:K |] ==> b lesspoll K";
+Goalw [lesspoll_def] "!!K. [| Card(K); b:K |] ==> b lesspoll K";
 by (asm_full_simp_tac (simpset() addsimps [Card_iff_initial]) 1);
 by (fast_tac (claset() addSIs [le_imp_lepoll, ltI, leI]) 1);
 qed "in_Card_imp_lesspoll";
 
-goal thy "(lam b:a. P(b)) : a -> {P(b). b:a}";
+Goal "(lam b:a. P(b)) : a -> {P(b). b:a}";
 by (fast_tac (claset() addSIs [lam_type, RepFunI]) 1);
 qed "lam_type_RepFun";
 
-goal thy "!!Z. [| ALL Y:Pow(X). Y lesspoll a --> (EX x:X. <Y, x> : R);  \
+Goal "!!Z. [| ALL Y:Pow(X). Y lesspoll a --> (EX x:X. <Y, x> : R);  \
 \       b:a; Z:Pow(X); Z lesspoll a |]  \
 \       ==> {x:X. <Z,x> : R} ~= 0";
 by (fast_tac (FOL_cs addEs [bexE, equals0D]
         addSDs [bspec] addIs [CollectI]) 1);
 val lemma_ = result();
 
-goal thy "!!K. [| Card(K); well_ord(X,Q);  \
+Goal "!!K. [| Card(K); well_ord(X,Q);  \
 \       ALL Y:Pow(X). Y lesspoll K --> (EX x:X. <Y, x> : R); b:K |]  \
 \       ==> ff(b, X, Q, R) : {x:X. <(lam c:b. ff(c, X, Q, R))``b, x> : R}";
 by (res_inst_tac [("P","b:K")] impE 1 THEN TRYALL assume_tac);
@@ -597,7 +597,7 @@
                 MRS lepoll_lesspoll_lesspoll]) 1);
 val lemma = result();
 
-goalw thy [DC_def, WO1_def]
+Goalw [DC_def, WO1_def]
         "!!Z. WO1 ==> ALL K. Card(K) --> DC(K)";
 by (REPEAT (resolve_tac [allI,impI] 1));
 by (REPEAT (eresolve_tac [allE,exE,conjE] 1));
--- a/src/ZF/AC/DC_lemmas.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/AC/DC_lemmas.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -14,7 +14,7 @@
 by (REPEAT (eresolve_tac [RepFunE, LeastI, prem RS Ord_in_Ord] 1));
 qed "RepFun_lepoll";
 
-goalw thy [lesspoll_def] "!!n. n:nat ==> n lesspoll nat";
+Goalw [lesspoll_def] "!!n. n:nat ==> n lesspoll nat";
 by (rtac conjI 1);
 by (eresolve_tac [Ord_nat RSN (2, OrdmemD) RS subset_imp_lepoll] 1);
 by (rtac notI 1);
@@ -25,7 +25,7 @@
         THEN (assume_tac 1));
 qed "n_lesspoll_nat";
 
-goalw thy [lepoll_def]
+Goalw [lepoll_def]
         "!!f. [| f:X->Y; Ord(X) |] ==> f``X lepoll X";
 by (res_inst_tac [("x","lam x:f``X. LEAST y. f`y = x")] exI 1);
 by (res_inst_tac [("d","%z. f`z")] lam_injective 1);
@@ -46,35 +46,35 @@
 by (fast_tac (claset() addSEs [mem_irrefl]) 1);
 qed "mem_not_eq";
 
-goalw thy [succ_def] "!!g. g:n->X ==> cons(<n,x>, g) : succ(n) -> cons(x, X)";
+Goalw [succ_def] "!!g. g:n->X ==> cons(<n,x>, g) : succ(n) -> cons(x, X)";
 by (fast_tac (claset() addSIs [fun_extend] addSEs [mem_irrefl]) 1);
 qed "cons_fun_type";
 
-goal thy "!!g. [| g:n->X; x:X |] ==> cons(<n,x>, g) : succ(n) -> X";
+Goal "!!g. [| g:n->X; x:X |] ==> cons(<n,x>, g) : succ(n) -> X";
 by (etac (cons_absorb RS subst) 1 THEN etac cons_fun_type 1);
 qed "cons_fun_type2";
 
-goal thy "!!n. n: nat ==> cons(<n,x>, g)``n = g``n";
+Goal "!!n. n: nat ==> cons(<n,x>, g)``n = g``n";
 by (fast_tac (claset() addSEs [mem_irrefl]) 1);
 qed "cons_image_n";
 
-goal thy "!!n. g:n->X ==> cons(<n,x>, g)`n = x";
+Goal "!!n. g:n->X ==> cons(<n,x>, g)`n = x";
 by (fast_tac (claset() addSIs [apply_equality] addSEs [cons_fun_type]) 1);
 qed "cons_val_n";
 
-goal thy "!!k. k : n ==> cons(<n,x>, g)``k = g``k";
+Goal "!!k. k : n ==> cons(<n,x>, g)``k = g``k";
 by (fast_tac (claset() addEs [mem_asym]) 1);
 qed "cons_image_k";
 
-goal thy "!!k. [| k:n; g:n->X |] ==> cons(<n,x>, g)`k = g`k";
+Goal "!!k. [| k:n; g:n->X |] ==> cons(<n,x>, g)`k = g`k";
 by (fast_tac (claset() addSIs [apply_equality, consI2] addSEs [cons_fun_type, apply_Pair]) 1);
 qed "cons_val_k";
 
-goal thy "!!f. domain(f)=x ==> domain(cons(<x,y>, f)) = succ(x)";
+Goal "!!f. domain(f)=x ==> domain(cons(<x,y>, f)) = succ(x)";
 by (asm_full_simp_tac (simpset() addsimps [domain_cons, succ_def]) 1);
 qed "domain_cons_eq_succ";
 
-goalw thy [restrict_def] "!!g. g:n->X ==> restrict(cons(<n,x>, g), n)=g";
+Goalw [restrict_def] "!!g. g:n->X ==> restrict(cons(<n,x>, g), n)=g";
 by (rtac fun_extension 1);
 by (rtac lam_type 1);
 by (eresolve_tac [cons_fun_type RS apply_type] 1);
@@ -83,25 +83,25 @@
 by (asm_full_simp_tac (simpset() addsimps [cons_val_k]) 1);
 qed "restrict_cons_eq";
 
-goal thy "!!k. [| Ord(k); i:k |] ==> succ(i) : succ(k)";
+Goal "!!k. [| Ord(k); i:k |] ==> succ(i) : succ(k)";
 by (resolve_tac [Ord_linear RS disjE] 1 THEN (assume_tac 3));
 by (REPEAT (fast_tac (claset() addSIs [Ord_succ]
         addEs [Ord_in_Ord, mem_irrefl, mem_asym]
         addSDs [succ_inject]) 1));
 qed "succ_in_succ";
 
-goalw thy [restrict_def]
+Goalw [restrict_def]
         "!!f. [| restrict(f, domain(g)) = g; x: domain(g) |] ==> f`x = g`x";
 by (etac subst 1);
 by (Asm_full_simp_tac 1);
 qed "restrict_eq_imp_val_eq";
 
-goal thy "!!f. [| domain(f)=A; f:B->C |] ==> f:A->C";
+Goal "!!f. [| domain(f)=A; f:B->C |] ==> f:A->C";
 by (forward_tac [domain_of_fun] 1);
 by (Fast_tac 1);
 qed "domain_eq_imp_fun_type";
 
-goal thy "!!R. [| R <= A * B; R ~= 0 |] ==> EX x. x:domain(R)";
+Goal "!!R. [| R <= A * B; R ~= 0 |] ==> EX x. x:domain(R)";
 by (fast_tac (claset() addSEs [not_emptyE]) 1);
 qed "ex_in_domain";
 
--- a/src/ZF/AC/HH.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/AC/HH.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -14,25 +14,25 @@
 (* Lemmas useful in each of the three proofs                              *)
 (* ********************************************************************** *)
 
-goal thy "HH(f,x,a) =  \
+Goal "HH(f,x,a) =  \
 \       (let z = x - (UN b:a. HH(f,x,b))  \
 \       in  if(f`z:Pow(z)-{0}, f`z, {x}))";
 by (resolve_tac [HH_def RS def_transrec RS trans] 1);
 by (Simp_tac 1);
 qed "HH_def_satisfies_eq";
 
-goal thy "HH(f,x,a) : Pow(x)-{0} | HH(f,x,a)={x}";
+Goal "HH(f,x,a) : Pow(x)-{0} | HH(f,x,a)={x}";
 by (resolve_tac [HH_def_satisfies_eq RS ssubst] 1);
 by (simp_tac (simpset() addsimps [Let_def, Diff_subset RS PowI] 
                     setloop split_tac [expand_if]) 1);
 by (Fast_tac 1);
 qed "HH_values";
 
-goal thy "!!A. B<=A ==> X-(UN a:A. P(a)) = X-(UN a:A-B. P(a))-(UN b:B. P(b))";
+Goal "!!A. B<=A ==> X-(UN a:A. P(a)) = X-(UN a:A-B. P(a))-(UN b:B. P(b))";
 by (Fast_tac 1);
 qed "subset_imp_Diff_eq";
 
-goal thy "!!c. [| c:a-b; b<a |] ==> c=b | b<c & c<a";
+Goal "!!c. [| c:a-b; b<a |] ==> c=b | b<c & c<a";
 by (etac ltE 1);
 by (dtac Ord_linear 1);
 by (fast_tac (claset() addSIs [ltI] addIs [Ord_in_Ord]) 2);
@@ -44,14 +44,14 @@
 by (fast_tac (claset() addSDs [prem] addSEs [mem_irrefl]) 1);
 qed "Diff_UN_eq_self";
 
-goal thy "!!a. x - (UN b:a. HH(f,x,b)) = x - (UN b:a1. HH(f,x,b))  \
+Goal "!!a. x - (UN b:a. HH(f,x,b)) = x - (UN b:a1. HH(f,x,b))  \
 \               ==> HH(f,x,a) = HH(f,x,a1)";
 by (resolve_tac [HH_def_satisfies_eq RS
                 (HH_def_satisfies_eq RS sym RSN (3, trans RS trans))] 1);
 by (etac subst_context 1);
 qed "HH_eq";
 
-goal thy "!!a. [| HH(f,x,b)={x}; b<a |] ==> HH(f,x,a)={x}";
+Goal "!!a. [| HH(f,x,b)={x}; b<a |] ==> HH(f,x,a)={x}";
 by (res_inst_tac [("P","b<a")] impE 1 THEN REPEAT (assume_tac 2));
 by (eresolve_tac [lt_Ord2 RS trans_induct] 1);
 by (rtac impI 1);
@@ -64,14 +64,14 @@
 by (fast_tac (claset() addEs [ltE]) 1);
 qed "HH_is_x_gt_too";
 
-goal thy "!!a. [| HH(f,x,a) : Pow(x)-{0}; b<a |] ==> HH(f,x,b) : Pow(x)-{0}";
+Goal "!!a. [| HH(f,x,a) : Pow(x)-{0}; b<a |] ==> HH(f,x,b) : Pow(x)-{0}";
 by (resolve_tac [HH_values RS disjE] 1 THEN (assume_tac 1));
 by (dtac HH_is_x_gt_too 1 THEN (assume_tac 1));
 by (dtac subst 1 THEN (assume_tac 1));
 by (fast_tac (claset() addSEs [mem_irrefl]) 1);
 qed "HH_subset_x_lt_too";
 
-goal thy "!!a. HH(f,x,a) : Pow(x)-{0}   \
+Goal "!!a. HH(f,x,a) : Pow(x)-{0}   \
 \               ==> HH(f,x,a) : Pow(x - (UN b:a. HH(f,x,b)))-{0}";
 by (dresolve_tac [HH_def_satisfies_eq RS subst] 1);
 by (resolve_tac [HH_def_satisfies_eq RS ssubst] 1);
@@ -81,14 +81,14 @@
 by (fast_tac (subset_cs addSEs [mem_irrefl]) 1);
 qed "HH_subset_x_imp_subset_Diff_UN";
 
-goal thy "!!x. [| HH(f,x,v)=HH(f,x,w); HH(f,x,v): Pow(x)-{0}; v:w |] ==> P";
+Goal "!!x. [| HH(f,x,v)=HH(f,x,w); HH(f,x,v): Pow(x)-{0}; v:w |] ==> P";
 by (forw_inst_tac [("P","%y. y: Pow(x)-{0}")] subst 1 THEN (assume_tac 1));
 by (dres_inst_tac [("a","w")] HH_subset_x_imp_subset_Diff_UN 1);
 by (dtac subst_elem 1 THEN (assume_tac 1));
 by (fast_tac (claset() addSIs [singleton_iff RS iffD2, equals0I]) 1);
 qed "HH_eq_arg_lt";
 
-goal thy "!!x. [| HH(f,x,v)=HH(f,x,w); HH(f,x,w): Pow(x)-{0};  \
+Goal "!!x. [| HH(f,x,v)=HH(f,x,w); HH(f,x,w): Pow(x)-{0};  \
 \               Ord(v); Ord(w) |] ==> v=w";
 by (res_inst_tac [("j","w")] Ord_linear_lt 1 THEN TRYALL assume_tac);
 by (resolve_tac [sym RS (ltD RSN (3, HH_eq_arg_lt))] 2
@@ -97,7 +97,7 @@
 by (fast_tac (FOL_cs addDs [ltD] addSEs [HH_eq_arg_lt]) 1);
 qed "HH_eq_imp_arg_eq";
 
-goalw thy [lepoll_def, inj_def]
+Goalw [lepoll_def, inj_def]
         "!!i. [| HH(f, x, i) : Pow(x)-{0}; Ord(i) |] ==> i lepoll Pow(x)-{0}";
 by (res_inst_tac [("x","lam j:i. HH(f, x, j)")] exI 1);
 by (Asm_simp_tac 1);
@@ -105,17 +105,17 @@
                 addSIs [lam_type, ballI, ltI] addIs [bexI]) 1);
 qed "HH_subset_x_imp_lepoll";
 
-goal thy "HH(f, x, Hartog(Pow(x)-{0})) = {x}";
+Goal "HH(f, x, Hartog(Pow(x)-{0})) = {x}";
 by (resolve_tac [HH_values RS disjE] 1 THEN (assume_tac 2));
 by (fast_tac (FOL_cs addSDs [HH_subset_x_imp_lepoll]
                 addSIs [Ord_Hartog] addSEs [HartogE]) 1);
 qed "HH_Hartog_is_x";
 
-goal thy "HH(f, x, LEAST i. HH(f, x, i) = {x}) = {x}";
+Goal "HH(f, x, LEAST i. HH(f, x, i) = {x}) = {x}";
 by (fast_tac (claset() addSIs [Ord_Hartog, HH_Hartog_is_x, LeastI]) 1);
 qed "HH_Least_eq_x";
 
-goal thy "!!a. a:(LEAST i. HH(f,x,i)={x}) ==> HH(f,x,a) : Pow(x)-{0}";
+Goal "!!a. a:(LEAST i. HH(f,x,i)={x}) ==> HH(f,x,a) : Pow(x)-{0}";
 by (resolve_tac [HH_values RS disjE] 1 THEN (assume_tac 1));
 by (rtac less_LeastE 1);
 by (eresolve_tac [Ord_Least RSN (2, ltI)] 2);
@@ -126,7 +126,7 @@
 (* Lemmas used in the proofs of AC1 ==> WO2 and AC17 ==> AC1              *)
 (* ********************************************************************** *)
 
-goalw thy [inj_def]
+Goalw [inj_def]
         "(lam a:(LEAST i. HH(f,x,i)={x}). HH(f,x,a)) :  \
 \               inj(LEAST i. HH(f,x,i)={x}, Pow(x)-{0})";
 by (Asm_full_simp_tac 1);
@@ -134,14 +134,14 @@
                 addSEs [HH_eq_imp_arg_eq, Ord_Least RS Ord_in_Ord]) 1);
 qed "lam_Least_HH_inj_Pow";
 
-goal thy "!!x. ALL a:(LEAST i. HH(f,x,i)={x}). EX z:x. HH(f,x,a) = {z}  \
+Goal "!!x. ALL a:(LEAST i. HH(f,x,i)={x}). EX z:x. HH(f,x,a) = {z}  \
 \               ==> (lam a:(LEAST i. HH(f,x,i)={x}). HH(f,x,a))  \
 \                       : inj(LEAST i. HH(f,x,i)={x}, {{y}. y:x})";
 by (resolve_tac [lam_Least_HH_inj_Pow RS inj_strengthen_type] 1);
 by (Asm_full_simp_tac 1);
 qed "lam_Least_HH_inj";
 
-goalw thy [surj_def]
+Goalw [surj_def]
         "!!x. [| x - (UN a:A. F(a)) = 0;  \
 \               ALL a:A. EX z:x. F(a) = {z} |]  \
 \               ==> (lam a:A. F(a)) : surj(A, {{y}. y:x})";
@@ -151,12 +151,12 @@
 by (deepen_tac (claset() addSIs [bexI] addSEs [equalityE]) 4 1);
 qed "lam_surj_sing";
 
-goal thy "!!x. y:Pow(x)-{0} ==> x ~= 0";
+Goal "!!x. y:Pow(x)-{0} ==> x ~= 0";
 by (fast_tac (claset() addSIs [equals0I, singletonI RS subst_elem]
                 addSDs [equals0D]) 1);
 qed "not_emptyI2";
 
-goal thy "!!f. f`(x - (UN j:i. HH(f,x,j))): Pow(x - (UN j:i. HH(f,x,j)))-{0}  \
+Goal "!!f. f`(x - (UN j:i. HH(f,x,j))): Pow(x - (UN j:i. HH(f,x,j)))-{0}  \
 \       ==> HH(f, x, i) : Pow(x) - {0}";
 by (resolve_tac [HH_def_satisfies_eq RS ssubst] 1);
 by (asm_full_simp_tac (simpset() addsimps [Let_def, Diff_subset RS PowI,
@@ -174,20 +174,20 @@
                 addSEs [mem_irrefl]) 1);
 qed "f_subsets_imp_UN_HH_eq_x";
 
-goal thy "HH(f,x,i)=f`(x - (UN j:i. HH(f,x,j))) | HH(f,x,i)={x}";
+Goal "HH(f,x,i)=f`(x - (UN j:i. HH(f,x,j))) | HH(f,x,i)={x}";
 by (resolve_tac [HH_def_satisfies_eq RS ssubst] 1);
 by (simp_tac (simpset() addsimps [Let_def, Diff_subset RS PowI]
               setloop split_tac [expand_if]) 1);
 qed "HH_values2";
 
-goal thy
+Goal
      "!!f. HH(f,x,i): Pow(x)-{0} ==> HH(f,x,i)=f`(x - (UN j:i. HH(f,x,j)))";
 by (resolve_tac [HH_values2 RS disjE] 1 THEN (assume_tac 1));
 by (fast_tac (claset() addSEs [equalityE, mem_irrefl]
         addSDs [singleton_subsetD]) 1);
 qed "HH_subset_imp_eq";
 
-goal thy "!!f. [| f : (Pow(x)-{0}) -> {{z}. z:x};  \
+Goal "!!f. [| f : (Pow(x)-{0}) -> {{z}. z:x};  \
 \       a:(LEAST i. HH(f,x,i)={x}) |] ==> EX z:x. HH(f,x,a) = {z}";
 by (dtac less_Least_subset_x 1);
 by (forward_tac [HH_subset_imp_eq] 1);
@@ -198,7 +198,7 @@
 by (fast_tac (claset() addss (simpset())) 1);
 qed "f_sing_imp_HH_sing";
 
-goalw thy [bij_def] 
+Goalw [bij_def] 
         "!!f. [| x - (UN j: (LEAST i. HH(f,x,i)={x}). HH(f,x,j)) = 0;  \
 \       f : (Pow(x)-{0}) -> {{z}. z:x} |]  \
 \       ==> (lam a:(LEAST i. HH(f,x,i)={x}). HH(f,x,a))  \
@@ -207,7 +207,7 @@
                               f_sing_imp_HH_sing]) 1);
 qed "f_sing_lam_bij";
 
-goal thy "!!f. f: (PROD X: Pow(x)-{0}. F(X))  \
+Goal "!!f. f: (PROD X: Pow(x)-{0}. F(X))  \
 \       ==> (lam X:Pow(x)-{0}. {f`X}) : (PROD X: Pow(x)-{0}. {{z}. z:F(X)})";
 by (fast_tac (FOL_cs addSIs [lam_type, RepFun_eqI, singleton_eq_iff RS iffD2]
                      addDs [apply_type]) 1);
--- a/src/ZF/AC/Hartog.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/AC/Hartog.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -7,12 +7,12 @@
 
 open Hartog;
 
-goal thy "!!X. ALL a. Ord(a) --> a:X ==> P";
+Goal "!!X. ALL a. Ord(a) --> a:X ==> P";
 by (res_inst_tac [("X1","{y:X. Ord(y)}")] (ON_class RS revcut_rl) 1);
 by (Fast_tac 1);
 qed "Ords_in_set";
 
-goalw thy [lepoll_def] "!!X. [| Ord(a); a lepoll X |] ==>  \
+Goalw [lepoll_def] "!!X. [| Ord(a); a lepoll X |] ==>  \
 \               EX Y. Y<=X & (EX R. well_ord(Y,R) & ordertype(Y,R)=a)";
 by (etac exE 1);
 by (REPEAT (resolve_tac [exI, conjI] 1));
@@ -28,7 +28,7 @@
         THEN (REPEAT (assume_tac 1)));
 qed "Ord_lepoll_imp_ex_well_ord";
 
-goal thy "!!X. [| Ord(a); a lepoll X |] ==>  \
+Goal "!!X. [| Ord(a); a lepoll X |] ==>  \
 \               EX Y. Y<=X & (EX R. R<=X*X & ordertype(Y,R)=a)";
 by (dtac Ord_lepoll_imp_ex_well_ord 1 THEN (assume_tac 1));
 by Safe_tac;
@@ -37,7 +37,7 @@
 by (Fast_tac 1);
 qed "Ord_lepoll_imp_eq_ordertype";
 
-goal thy "!!X. ALL a. Ord(a) --> a lepoll X ==>  \
+Goal "!!X. ALL a. Ord(a) --> a lepoll X ==>  \
 \       ALL a. Ord(a) -->  \
 \       a:{a. Z:Pow(X)*Pow(X*X), EX Y R. Z=<Y,R> & ordertype(Y,R)=a}";
 by (REPEAT (resolve_tac [allI,impI] 1));
@@ -47,18 +47,18 @@
 by (fast_tac (claset() addSIs [ReplaceI] addEs [sym]) 1);
 qed "Ords_lepoll_set_lemma";
 
-goal thy "!!X. ALL a. Ord(a) --> a lepoll X ==> P";
+Goal "!!X. ALL a. Ord(a) --> a lepoll X ==> P";
 by (eresolve_tac [Ords_lepoll_set_lemma RS Ords_in_set] 1);
 qed "Ords_lepoll_set";
 
-goal thy "EX a. Ord(a) & ~a lepoll X";
+Goal "EX a. Ord(a) & ~a lepoll X";
 by (rtac swap 1);
 by (Fast_tac 1);
 by (rtac Ords_lepoll_set 1);
 by (Fast_tac 1);
 qed "ex_Ord_not_lepoll";
 
-goalw thy [Hartog_def] "~ Hartog(A) lepoll A";
+Goalw [Hartog_def] "~ Hartog(A) lepoll A";
 by (resolve_tac [ex_Ord_not_lepoll RS exE] 1);
 by (rtac LeastI 1);
 by (REPEAT (Fast_tac 1));
@@ -66,19 +66,19 @@
 
 val HartogE = HartogI RS notE;
 
-goalw thy [Hartog_def] "Ord(Hartog(A))";
+Goalw [Hartog_def] "Ord(Hartog(A))";
 by (rtac Ord_Least 1);
 qed "Ord_Hartog";
 
-goalw thy [Hartog_def] "!!i. [| i < Hartog(A); ~ i lepoll A |] ==> P";
+Goalw [Hartog_def] "!!i. [| i < Hartog(A); ~ i lepoll A |] ==> P";
 by (fast_tac (claset() addEs [less_LeastE]) 1);
 qed "less_HartogE1";
 
-goal thy "!!i. [| i < Hartog(A); i eqpoll Hartog(A) |] ==> P";
+Goal "!!i. [| i < Hartog(A); i eqpoll Hartog(A) |] ==> P";
 by (fast_tac (claset() addEs [less_HartogE1, eqpoll_sym RS eqpoll_imp_lepoll
                 RS lepoll_trans RS HartogE]) 1);
 qed "less_HartogE";
 
-goal thy "Card(Hartog(A))";
+Goal "Card(Hartog(A))";
 by (fast_tac (claset() addSIs [CardI, Ord_Hartog] addEs [less_HartogE]) 1);
 qed "Card_Hartog";
--- a/src/ZF/AC/WO1_AC.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/AC/WO1_AC.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -31,7 +31,7 @@
 (* WO1 ==> AC1                                                            *)
 (* ********************************************************************** *)
 
-goalw thy [AC1_def, WO1_def] "!!Z. WO1 ==> AC1";
+Goalw [AC1_def, WO1_def] "!!Z. WO1 ==> AC1";
 by (fast_tac (claset() addSEs [ex_choice_fun]) 1);
 qed "WO1_AC1";
 
@@ -39,7 +39,7 @@
 (* WO1 ==> AC10(n) (n >= 1)                                               *)
 (* ********************************************************************** *)
 
-goalw thy [WO1_def] "!!A. [| WO1; ALL B:A. EX C:D(B). P(C,B) |]  \
+Goalw [WO1_def] "!!A. [| WO1; ALL B:A. EX C:D(B). P(C,B) |]  \
 \               ==> EX f. ALL B:A. P(f`B,B)";
 by (eres_inst_tac [("x","Union({{C:D(B). P(C,B)}. B:A})")] allE 1);
 by (etac exE 1);
@@ -52,7 +52,7 @@
                 addSEs [CollectD2]) 1);
 val lemma1 = result();
 
-goalw thy [WO1_def] "!!A. [| ~Finite(B); WO1 |] ==> |B| + |B| eqpoll  B";
+Goalw [WO1_def] "!!A. [| ~Finite(B); WO1 |] ==> |B| + |B| eqpoll  B";
 by (rtac eqpoll_trans 1);
 by (fast_tac (claset() addSEs [well_ord_cardinal_eqpoll]) 2);
 by (resolve_tac [eqpoll_sym RS eqpoll_trans] 1);
@@ -68,17 +68,17 @@
 by (fast_tac (claset() addSEs [well_ord_cardinal_eqpoll]) 1);
 val lemma2_1 = result();
 
-goal thy "!!f. f : bij(D+D, B) ==> {{f`Inl(i), f`Inr(i)}. i:D} : Pow(Pow(B))";
+Goal "!!f. f : bij(D+D, B) ==> {{f`Inl(i), f`Inr(i)}. i:D} : Pow(Pow(B))";
 by (fast_tac (claset() addSIs [InlI, InrI]
                 addSEs [RepFunE, bij_is_fun RS apply_type]) 1);
 val lemma2_2 = result();
 
-goal thy "!!f. [| f:inj(A,B); f`a = f`b; a:A; b:A |] ==> a=b";
+Goal "!!f. [| f:inj(A,B); f`a = f`b; a:A; b:A |] ==> a=b";
 by (rtac inj_equality 1);
 by (TRYALL (fast_tac (claset() addSEs [inj_is_fun RS apply_Pair] addEs [subst])));
 val lemma = result();
 
-goalw thy AC_aux_defs
+Goalw AC_aux_defs
         "!!f. f : bij(D+D, B) ==>  \
 \               pairwise_disjoint({{f`Inl(i), f`Inr(i)}. i:D})";
 by (fast_tac (claset() addSEs [RepFunE, not_emptyE] 
@@ -99,12 +99,12 @@
         addSEs [RepFunE, not_emptyE, mem_irrefl]) 1);
 val lemma2_4 = result();
 
-goalw thy [bij_def, surj_def]
+Goalw [bij_def, surj_def]
         "!!f. f : bij(D+D, B) ==> Union({{f`Inl(i), f`Inr(i)}. i:D})=B";
 by (fast_tac (claset() addSEs [inj_is_fun RS apply_type]) 1);
 val lemma2_5 = result();
 
-goal thy "!!A. [| WO1; ~Finite(B); 1 le n  |]  \
+Goal "!!A. [| WO1; ~Finite(B); 1 le n  |]  \
 \       ==> EX C:Pow(Pow(B)). pairwise_disjoint(C) &  \
 \               sets_of_size_between(C, 2, succ(n)) &  \
 \               Union(C)=B";
@@ -114,6 +114,6 @@
                 addSEs [lemma2_2, lemma2_3, lemma2_4, lemma2_5]) 1);
 val lemma2 = result();
 
-goalw thy AC_defs "!!n. [| WO1; 1 le n |] ==> AC10(n)";
+Goalw AC_defs "!!n. [| WO1; 1 le n |] ==> AC10(n)";
 by (fast_tac (claset() addSIs [lemma1] addSEs [lemma2]) 1);
 qed "WO1_AC10";
--- a/src/ZF/AC/WO1_WO6.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/AC/WO1_WO6.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -13,39 +13,39 @@
 
 (* ********************************************************************** *)
 
-goalw thy WO_defs "!!Z. WO2 ==> WO3";
+Goalw WO_defs "!!Z. WO2 ==> WO3";
 by (Fast_tac 1);
 qed "WO2_WO3";
 
 (* ********************************************************************** *)
 
-goalw thy (eqpoll_def::WO_defs) "!!Z. WO3 ==> WO1";
+Goalw (eqpoll_def::WO_defs) "!!Z. WO3 ==> WO1";
 by (fast_tac (claset() addSEs [bij_is_inj RS well_ord_rvimage, 
 			      well_ord_Memrel RS well_ord_subset]) 1);
 qed "WO3_WO1";
 
 (* ********************************************************************** *)
 
-goalw thy (eqpoll_def::WO_defs) "!!Z. WO1 ==> WO2";
+Goalw (eqpoll_def::WO_defs) "!!Z. WO1 ==> WO2";
 by (fast_tac (claset() addSIs [Ord_ordertype, ordermap_bij]) 1);
 qed "WO1_WO2";
 
 (* ********************************************************************** *)
 
-goal thy "!!f. f: A->B ==> (lam x:A. {f`x}): A -> {{b}. b:B}";
+Goal "!!f. f: A->B ==> (lam x:A. {f`x}): A -> {{b}. b:B}";
 by (fast_tac (claset() addSIs [lam_type, apply_type]) 1);
 qed "lam_sets";
 
-goalw thy [surj_def] "!!f. f:surj(A,B) ==> (UN a:A. {f`a}) = B";
+Goalw [surj_def] "!!f. f:surj(A,B) ==> (UN a:A. {f`a}) = B";
 by (fast_tac (claset() addSEs [apply_type]) 1);
 qed "surj_imp_eq_";
 
-goal thy "!!f. [| f:surj(A,B); Ord(A) |] ==> (UN a<A. {f`a}) = B";
+Goal "!!f. [| f:surj(A,B); Ord(A) |] ==> (UN a<A. {f`a}) = B";
 by (fast_tac (claset() addSDs [surj_imp_eq_]
                 addSIs [ltI] addSEs [ltE]) 1);
 qed "surj_imp_eq";
 
-goalw thy WO_defs "!!Z. WO1 ==> WO4(1)";
+Goalw WO_defs "!!Z. WO1 ==> WO4(1)";
 by (rtac allI 1);
 by (eres_inst_tac [("x","A")] allE 1);
 by (etac exE 1);
@@ -61,20 +61,20 @@
 
 (* ********************************************************************** *)
 
-goalw thy WO_defs "!!Z. [| m:nat; n:nat; m le n; WO4(m) |] ==> WO4(n)";
+Goalw WO_defs "!!Z. [| m:nat; n:nat; m le n; WO4(m) |] ==> WO4(n)";
 by (fast_tac (claset() addIs [nat_le_imp_lepoll RSN (2, lepoll_trans)]) 1);
 qed "WO4_mono";
 
 (* ********************************************************************** *)
 
-goalw thy WO_defs "!!Z. [| m:nat; 1 le m; WO4(m) |] ==> WO5";
+Goalw WO_defs "!!Z. [| m:nat; 1 le m; WO4(m) |] ==> WO5";
     (*ZF_cs is essential: default claset's too slow*)
 by (fast_tac ZF_cs 1);
 qed "WO4_WO5";
 
 (* ********************************************************************** *)
 
-goalw thy WO_defs "!!Z. WO5 ==> WO6";
+Goalw WO_defs "!!Z. WO5 ==> WO6";
     (*ZF_cs is essential: default claset's too slow*)
 by (fast_tac ZF_cs 1);
 qed "WO5_WO6";
--- a/src/ZF/AC/WO1_WO7.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/AC/WO1_WO7.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -10,7 +10,7 @@
 (* It is easy to see, that WO7 is equivallent to (**)                     *)
 (* ********************************************************************** *)
 
-goalw thy [WO7_def] "WO7 <-> (ALL X. ~Finite(X) -->  \
+Goalw [WO7_def] "WO7 <-> (ALL X. ~Finite(X) -->  \
 \                       (EX R. well_ord(X,R) & ~well_ord(X,converse(R))))";
 by (fast_tac (claset() addSEs [Finite_well_ord_converse]) 1);
 qed "WO7_iff_LEMMA";
@@ -19,7 +19,7 @@
 (* It is also easy to show that LEMMA implies WO1.                        *)
 (* ********************************************************************** *)
 
-goalw thy [WO1_def] "!!Z. ALL X. ~Finite(X) -->  \
+Goalw [WO1_def] "!!Z. ALL X. ~Finite(X) -->  \
 \               (EX R. well_ord(X,R) & ~well_ord(X,converse(R))) ==> WO1";
 by (rtac allI 1);
 by (etac allE 1);
@@ -43,7 +43,7 @@
 (* gives the conclusion.                                                  *)
 (* ********************************************************************** *)
 
-goalw thy [wf_on_def, wf_def] 
+Goalw [wf_on_def, wf_def] 
     "!!a. [| Ord(a); ~Finite(a) |] ==> ~wf[a](converse(Memrel(a)))";
 by (dresolve_tac [nat_le_infinite_Ord RS le_imp_subset] 1 
     THEN (assume_tac 1));
@@ -57,12 +57,12 @@
                             nat_succI RSN (2, subsetD)]) 1);
 qed "converse_Memrel_not_wf_on";
 
-goalw thy [well_ord_def] 
+Goalw [well_ord_def] 
     "!!a. [| Ord(a); ~Finite(a) |] ==> ~well_ord(a,converse(Memrel(a)))";
 by (fast_tac (claset() addSDs [converse_Memrel_not_wf_on]) 1);
 qed "converse_Memrel_not_well_ord";
 
-goal thy "!!A. [| well_ord(A,r); well_ord(A,converse(r)) |]  \
+Goal "!!A. [| well_ord(A,r); well_ord(A,converse(r)) |]  \
 \       ==> well_ord(ordertype(A,r), converse(Memrel(ordertype(A, r))))";
 by (rtac ([ordertype_ord_iso RS ord_iso_sym RS ord_iso_rvimage_eq, 
                 Memrel_type RS (subset_Int_iff RS iffD1)] 
@@ -74,7 +74,7 @@
         THEN (assume_tac 1));
 qed "well_ord_converse_Memrel";
 
-goalw thy [WO1_def] "!!Z. WO1 ==> ALL X. ~Finite(X) -->  \
+Goalw [WO1_def] "!!Z. WO1 ==> ALL X. ~Finite(X) -->  \
 \                       (EX R. well_ord(X,R) & ~well_ord(X,converse(R)))";
 by (REPEAT (resolve_tac [allI,impI] 1));
 by (REPEAT (eresolve_tac [allE,exE] 1));
--- a/src/ZF/AC/WO1_WO8.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/AC/WO1_WO8.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -9,7 +9,7 @@
 (* Trivial implication "WO1 ==> WO8"                                      *)
 (* ********************************************************************** *)
 
-goalw thy WO_defs "!!Z. WO1 ==> WO8";
+Goalw WO_defs "!!Z. WO1 ==> WO8";
 by (Fast_tac 1);
 qed "WO1_WO8";
 
@@ -17,7 +17,7 @@
 (* The proof of "WO8 ==> WO1" - faithful image of Rubin & Rubin's proof   *)
 (* ********************************************************************** *)
 
-goalw thy WO_defs "!!Z. WO8 ==> WO1";
+Goalw WO_defs "!!Z. WO8 ==> WO1";
 by (rtac allI 1);
 by (eres_inst_tac [("x","{{x}. x:A}")] allE 1);
 by (etac impE 1);
--- a/src/ZF/AC/WO2_AC16.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/AC/WO2_AC16.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -18,7 +18,7 @@
 (* case of limit ordinal                                                  *)
 (* ********************************************************************** *)
 
-goal thy "!!Z. [| ALL y<x. ALL z<a. z<y | (EX Y: F(y). f(z)<=Y)  \
+Goal "!!Z. [| ALL y<x. ALL z<a. z<y | (EX Y: F(y). f(z)<=Y)  \
 \               --> (EX! Y. Y:F(y) & f(z)<=Y);  \
 \               ALL i j. i le j --> F(i) <= F(j); j le i; i<x; z<a;  \
 \               V: F(i); f(z)<=V; W:F(j); f(z)<=W |]  \
@@ -31,7 +31,7 @@
 by (etac ex1_two_eq 1 THEN (REPEAT (ares_tac [conjI] 1)));
 val lemma3_1 = result();
 
-goal thy "!!Z. [| ALL y<x. ALL z<a. z<y | (EX Y: F(y). f(z)<=Y)  \
+Goal "!!Z. [| ALL y<x. ALL z<a. z<y | (EX Y: F(y). f(z)<=Y)  \
 \               --> (EX! Y. Y:F(y) & f(z)<=Y);  \
 \               ALL i j. i le j --> F(i) <= F(j); i<x; j<x; z<a;  \
 \               V: F(i); f(z)<=V; W:F(j); f(z)<=W |]  \
@@ -42,7 +42,7 @@
 by (etac lemma3_1 1 THEN (REPEAT (assume_tac 1)));
 val lemma3 = result();
 
-goal thy "!!a. [| ALL y<x. F(y) <= X &  \
+Goal "!!a. [| ALL y<x. F(y) <= X &  \
 \               (ALL x<a. x < y | (EX Y:F(y). fa(x) <= Y) -->  \
 \                       (EX! Y. Y : F(y) & fa(x) <= Y)); x < a |]  \
 \               ==> ALL y<x. ALL z<a. z < y | (EX Y:F(y). fa(z) <= Y) -->  \
@@ -52,7 +52,7 @@
 by (fast_tac (FOL_cs addSEs [oallE]) 1);
 val lemma4 = result();
 
-goal thy "!!a. [| ALL y<x. F(y) <= X &  \
+Goal "!!a. [| ALL y<x. F(y) <= X &  \
 \               (ALL x<a. x < y | (EX Y:F(y). fa(x) <= Y) -->  \
 \                       (EX! Y. Y : F(y) & fa(x) <= Y)); \
 \               x < a; Limit(x); ALL i j. i le j --> F(i) <= F(j) |]  \
@@ -101,7 +101,7 @@
 (* dbl_Diff_eqpoll_Card                                                   *)
 (* ********************************************************************** *)
 
-goal thy "!!A. [| A eqpoll a; Card(a); ~Finite(a); B lesspoll a;  \
+Goal "!!A. [| A eqpoll a; Card(a); ~Finite(a); B lesspoll a;  \
 \       C lesspoll a |] ==> A - B - C eqpoll a";
 by (rtac Diff_lesspoll_eqpoll_Card 1 THEN (REPEAT (assume_tac 1)));
 by (rtac Diff_lesspoll_eqpoll_Card 1 THEN (REPEAT (assume_tac 1)));
@@ -111,7 +111,7 @@
 (* Case of finite ordinals                                                *)
 (* ********************************************************************** *)
 
-goalw thy [lesspoll_def]
+Goalw [lesspoll_def]
         "!!X. [| Finite(X); ~Finite(a); Ord(a) |] ==> X lesspoll a";
 by (rtac conjI 1);
 by (dresolve_tac [nat_le_infinite_Ord RS le_imp_lepoll] 1
@@ -123,11 +123,11 @@
         subset_imp_lepoll RSN (2, eqpoll_imp_lepoll RS lepoll_trans)]) 1);
 qed "Finite_lesspoll_infinite_Ord";
 
-goal thy "!!x. x:X ==> Union(X) = Union(X-{x}) Un x";
+Goal "!!x. x:X ==> Union(X) = Union(X-{x}) Un x";
 by (Fast_tac 1);
 qed "Union_eq_Un_Diff";
 
-goal thy "!!n. n:nat ==> ALL X. X eqpoll n --> (ALL x:X. Finite(x))  \
+Goal "!!n. n:nat ==> ALL X. X eqpoll n --> (ALL x:X. Finite(x))  \
 \       --> Finite(Union(X))";
 by (etac nat_induct 1);
 by (fast_tac (claset() addSDs [eqpoll_imp_lepoll RS lepoll_0_is_0]
@@ -141,19 +141,19 @@
 by (fast_tac (claset() addSIs [Diff_sing_eqpoll]) 1);
 qed "Finite_Union_lemma";
 
-goal thy "!!X. [| ALL x:X. Finite(x); Finite(X) |] ==> Finite(Union(X))";
+Goal "!!X. [| ALL x:X. Finite(x); Finite(X) |] ==> Finite(Union(X))";
 by (eresolve_tac [Finite_def RS def_imp_iff RS iffD1 RS bexE] 1);
 by (dtac Finite_Union_lemma 1);
 by (Fast_tac 1);
 qed "Finite_Union";
 
-goalw thy [Finite_def] "!!x. [| x lepoll n; n:nat |] ==> Finite(x)";
+Goalw [Finite_def] "!!x. [| x lepoll n; n:nat |] ==> Finite(x)";
 by (fast_tac (claset()
         addEs [nat_into_Ord RSN (2, lepoll_imp_ex_le_eqpoll) RS exE,
         Ord_nat RSN (2, ltI) RSN (2, lt_trans1) RS ltD]) 1);
 qed "lepoll_nat_num_imp_Finite";
 
-goal thy "!!X. [| ALL x:X. x lepoll n & x <= T; well_ord(T, R); X lepoll b;  \
+Goal "!!X. [| ALL x:X. x lepoll n & x <= T; well_ord(T, R); X lepoll b;  \
 \       b<a; ~Finite(a); Card(a); n:nat |]  \
 \       ==> Union(X) lesspoll a";
 by (excluded_middle_tac "Finite(X)" 1);
@@ -181,24 +181,24 @@
 (* recfunAC16_lepoll_index                                                *)
 (* ********************************************************************** *)
 
-goal thy "A Un {a} = cons(a, A)";
+Goal "A Un {a} = cons(a, A)";
 by (Fast_tac 1);
 qed "Un_sing_eq_cons";
 
-goal thy "!!A. A lepoll B ==> A Un {a} lepoll succ(B)";
+Goal "!!A. A lepoll B ==> A Un {a} lepoll succ(B)";
 by (asm_simp_tac (simpset() addsimps [Un_sing_eq_cons, succ_def]) 1);
 by (eresolve_tac [mem_not_refl RSN (2, cons_lepoll_cong)] 1);
 qed "Un_lepoll_succ";
 
-goal thy "!!a. Ord(a) ==> F(a) - (UN b<succ(a). F(b)) = 0";
+Goal "!!a. Ord(a) ==> F(a) - (UN b<succ(a). F(b)) = 0";
 by (fast_tac (claset() addSIs [OUN_I, le_refl]) 1);
 qed "Diff_UN_succ_empty";
 
-goal thy "!!a. Ord(a) ==> F(a) Un X - (UN b<succ(a). F(b)) <= X";
+Goal "!!a. Ord(a) ==> F(a) Un X - (UN b<succ(a). F(b)) <= X";
 by (fast_tac (claset() addSIs [OUN_I, le_refl]) 1);
 qed "Diff_UN_succ_subset";
 
-goal thy "!!x. Ord(x) ==>  \
+Goal "!!x. Ord(x) ==>  \
 \       recfunAC16(f, g, x, a) - (UN i<x. recfunAC16(f, g, i, a)) lepoll 1";
 by (etac Ord_cases 1);
 by (asm_simp_tac (simpset() addsimps [recfunAC16_0,
@@ -213,12 +213,12 @@
         (singleton_eqpoll_1 RS eqpoll_imp_lepoll RSN (2, lepoll_trans))]) 1);
 qed "recfunAC16_Diff_lepoll_1";
 
-goal thy "!!z. [| z : F(x); Ord(x) |]  \
+Goal "!!z. [| z : F(x); Ord(x) |]  \
 \       ==> z:F(LEAST i. z:F(i)) - (UN j<(LEAST i. z:F(i)). F(j))";
 by (fast_tac (claset() addEs [less_LeastE] addSEs [OUN_E, LeastI]) 1);
 qed "in_Least_Diff";
 
-goal thy "!!w. [| (LEAST i. w:F(i)) = (LEAST i. z:F(i));  \
+Goal "!!w. [| (LEAST i. w:F(i)) = (LEAST i. z:F(i));  \
 \       w:(UN i<a. F(i)); z:(UN i<a. F(i)) |]  \
 \       ==> EX b<a. w:(F(b) - (UN c<b. F(c))) & z:(F(b) - (UN c<b. F(c)))";
 by (REPEAT (etac OUN_E 1));
@@ -231,11 +231,11 @@
         THEN (REPEAT (assume_tac 1)));
 qed "Least_eq_imp_ex";
 
-goal thy "!!A. [| A lepoll 1; a:A; b:A |] ==> a=b";
+Goal "!!A. [| A lepoll 1; a:A; b:A |] ==> a=b";
 by (fast_tac (claset() addSDs [lepoll_1_is_sing]) 1);
 qed "two_in_lepoll_1";
 
-goal thy "!!a. [| ALL i<a. F(i)-(UN j<i. F(j)) lepoll 1; Limit(a) |]  \
+Goal "!!a. [| ALL i<a. F(i)-(UN j<i. F(j)) lepoll 1; Limit(a) |]  \
 \       ==> (UN x<a. F(x)) lepoll a";
 by (resolve_tac [lepoll_def RS (def_imp_iff RS iffD2)] 1);
 by (res_inst_tac [("x","lam z: (UN x<a. F(x)). LEAST i. z:F(i)")] exI 1);
@@ -254,7 +254,7 @@
 by (fast_tac (claset() addSEs [two_in_lepoll_1]) 1);
 qed "UN_lepoll_index";
 
-goal thy "!!y. Ord(y) ==> recfunAC16(f, fa, y, a) lepoll y";
+Goal "!!y. Ord(y) ==> recfunAC16(f, fa, y, a) lepoll y";
 by (etac trans_induct 1);
 by (etac Ord_cases 1);
 by (asm_simp_tac (simpset() addsimps [recfunAC16_0, lepoll_refl]) 1);
@@ -268,7 +268,7 @@
         addSIs [UN_lepoll_index]) 1);
 qed "recfunAC16_lepoll_index";
 
-goal thy "!!y. [| recfunAC16(f,g,y,a) <= {X: Pow(A). X eqpoll n};  \
+Goal "!!y. [| recfunAC16(f,g,y,a) <= {X: Pow(A). X eqpoll n};  \
 \               A eqpoll a; y<a; ~Finite(a); Card(a); n:nat |]  \
 \               ==> Union(recfunAC16(f,g,y,a)) lesspoll a";
 by (eresolve_tac [eqpoll_def RS def_imp_iff RS iffD1 RS exE] 1);
@@ -279,7 +279,7 @@
 by (fast_tac (claset() addSEs [eqpoll_imp_lepoll]) 1);
 qed "Union_recfunAC16_lesspoll";
 
-goal thy
+Goal
   "!!a. [| recfunAC16(f, fa, y, a) <= {X: Pow(A) . X eqpoll succ(k #+ m)};  \
 \       Card(a); ~ Finite(a); A eqpoll a;  \
 \       k : nat; m : nat; y<a;  \
@@ -302,7 +302,7 @@
                              nat_sum_eqpoll_sum RSN (3, 
                              eqpoll_trans RS eqpoll_trans))) |> standard;
 
-goal thy "!!x. [| x : Pow(A - B - fa`i); x eqpoll m;  \
+Goal "!!x. [| x : Pow(A - B - fa`i); x eqpoll m;  \
 \       fa : bij(a, {x: Pow(A) . x eqpoll k}); i<a; k:nat; m:nat |]  \
 \       ==> fa ` i Un x : {x: Pow(A) . x eqpoll k #+ m}";
 by (rtac CollectI 1);
@@ -319,7 +319,7 @@
 (* Lemmas simplifying assumptions                                         *)
 (* ********************************************************************** *)
 
-goal thy "!!j. [| ALL y<succ(j). F(y)<=X & (ALL x<a. x<y | P(x,y)  \
+Goal "!!j. [| ALL y<succ(j). F(y)<=X & (ALL x<a. x<y | P(x,y)  \
 \       --> Q(x,y)); succ(j)<a |]  \
 \       ==> F(j)<=X & (ALL x<a. x<j | P(x,j) --> Q(x,j))";
 by (dtac ospec 1);
@@ -327,7 +327,7 @@
         THEN (REPEAT (assume_tac 1)));
 val lemma6 = result();
 
-goal thy "!!j. [| F(j)<=X; (ALL x<a. x<j | P(x,j) --> Q(x,j)); succ(j)<a |]  \
+Goal "!!j. [| F(j)<=X; (ALL x<a. x<j | P(x,j) --> Q(x,j)); succ(j)<a |]  \
 \       ==> P(j,j) --> F(j) <= X & (ALL x<a. x le j | P(x,j) --> Q(x,j))";
 by (fast_tac (claset() addSEs [leE]) 1);
 val lemma7 = result();
@@ -337,7 +337,7 @@
 (* ordinal there is a set satisfying certain properties                   *)
 (* ********************************************************************** *)
 
-goal thy "!!A. [| A eqpoll a; ~ Finite(a); Ord(a); m:nat |]  \
+Goal "!!A. [| A eqpoll a; ~ Finite(a); Ord(a); m:nat |]  \
 \       ==> EX X:Pow(A). X eqpoll m";
 by (eresolve_tac [Ord_nat RSN (2, ltI) RS 
                 (nat_le_infinite_Ord RSN (2, lt_trans2)) RS 
@@ -348,11 +348,11 @@
 by (fast_tac (claset() addSEs [eqpoll_sym]) 1);
 qed "ex_subset_eqpoll";
 
-goal thy "!!A. [| A <= B Un C; A Int C = 0 |] ==> A <= B";
+Goal "!!A. [| A <= B Un C; A Int C = 0 |] ==> A <= B";
 by (fast_tac (claset() addDs [equals0D]) 1);
 qed "subset_Un_disjoint";
 
-goal thy "!!F. [| X:Pow(A - Union(B) -C); T:B; F<=T |] ==> F Int X = 0";
+Goal "!!F. [| X:Pow(A - Union(B) -C); T:B; F<=T |] ==> F Int X = 0";
 by (fast_tac (claset() addSIs [equals0I]) 1);
 qed "Int_empty";
 
@@ -360,11 +360,11 @@
 (* equipollent subset (and finite) is the whole set                       *)
 (* ********************************************************************** *)
 
-goal thy "!!A. [| A <= B; a : A; A - {a} = B - {a} |] ==> A = B";
+Goal "!!A. [| A <= B; a : A; A - {a} = B - {a} |] ==> A = B";
 by (fast_tac (claset() addSEs [equalityE]) 1);
 qed "Diffs_eq_imp_eq";
 
-goal thy "!!A. m:nat ==> ALL A B. A <= B & m lepoll A & B lepoll m --> A=B";
+Goal "!!A. m:nat ==> ALL A B. A <= B & m lepoll A & B lepoll m --> A=B";
 by (etac nat_induct 1);
 by (fast_tac (claset() addSDs [lepoll_0_is_0]) 1);
 by (REPEAT (resolve_tac [allI, impI] 1));
@@ -382,11 +382,11 @@
         THEN REPEAT (assume_tac 1));
 qed "subset_imp_eq_lemma";
 
-goal thy "!!A. [| A <= B; m lepoll A; B lepoll m; m:nat |] ==> A=B";
+Goal "!!A. [| A <= B; m lepoll A; B lepoll m; m:nat |] ==> A=B";
 by (fast_tac (FOL_cs addSDs [subset_imp_eq_lemma]) 1);
 qed "subset_imp_eq";
 
-goal thy "!!f. [| f:bij(a, {Y:X. Y eqpoll succ(k)}); k:nat; f`b<=f`y; b<a;  \
+Goal "!!f. [| f:bij(a, {Y:X. Y eqpoll succ(k)}); k:nat; f`b<=f`y; b<a;  \
 \       y<a |] ==> b=y";
 by (dtac subset_imp_eq 1);
 by (etac nat_succI 3);
@@ -398,7 +398,7 @@
 by (fast_tac (claset() addSDs [ltD]) 1);
 qed "bij_imp_arg_eq";
 
-goal thy
+Goal
   "!!a. [| recfunAC16(f, fa, y, a) <= {X: Pow(A) . X eqpoll succ(k #+ m)};  \
 \       Card(a); ~ Finite(a); A eqpoll a;  \
 \       k : nat; m : nat; y<a;  \
@@ -434,7 +434,7 @@
 (* ordinal there is a number of the set satisfying certain properties     *)
 (* ********************************************************************** *)
 
-goal thy
+Goal
   "!!a. [| recfunAC16(f, fa, y, a) <= {X: Pow(A) . X eqpoll succ(k #+ m)};  \
 \       Card(a); ~ Finite(a); A eqpoll a;  \
 \       k : nat; m : nat; y<a;  \
@@ -452,7 +452,7 @@
         THEN REPEAT (ares_tac [Card_is_Ord] 1));
 qed "ex_next_Ord";
 
-goal thy "!!Z. [| EX! Y. Y:Z & P(Y); ~P(W) |] ==> EX! Y. Y: (Z Un {W}) & P(Y)";
+Goal "!!Z. [| EX! Y. Y:Z & P(Y); ~P(W) |] ==> EX! Y. Y: (Z Un {W}) & P(Y)";
 by (Fast_tac 1);
 qed "ex1_in_Un_sing";
 
@@ -460,7 +460,7 @@
 (* Lemma simplifying assumptions                                          *)
 (* ********************************************************************** *)
 
-goal thy "!!j. [| ALL x<a. x<j | (EX xa:F(j). P(x, xa))  \
+Goal "!!j. [| ALL x<a. x<j | (EX xa:F(j). P(x, xa))  \
 \       --> (EX! Y. Y:F(j) & P(x, Y)); F(j) <= X;  \
 \       L : X; P(j, L) & (ALL x<a. P(x, L) --> (ALL xa:F(j). ~P(x, xa))) |]  \
 \       ==> F(j) Un {L} <= X &  \
@@ -489,7 +489,7 @@
 (* lemma main_induct                                                      *)
 (* ********************************************************************** *)
 
-goal thy
+Goal
         "!!a. [| b < a; f : bij(a, {Y: Pow(A) . Y eqpoll succ(k #+ m)});  \
 \       fa : bij(a, {Y: Pow(A) . Y eqpoll succ(k)});  \
 \       ~Finite(a); Card(a); A eqpoll a; k : nat; m : nat |]  \
@@ -574,7 +574,7 @@
 (* The target theorem                                                     *)
 (* ********************************************************************** *)
 
-goalw thy [AC16_def]
+Goalw [AC16_def]
         "!!n k. [| WO2; 0<m; k:nat; m:nat |] ==> AC16(k #+ m,k)";
 by (rtac allI 1);
 by (rtac impI 1);
--- a/src/ZF/AC/WO6_WO1.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/AC/WO6_WO1.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -32,25 +32,25 @@
 (* some properties of relation uu(beta, gamma, delta) - p. 2              *)
 (* ********************************************************************** *)
 
-goalw thy [uu_def] "domain(uu(f,b,g,d)) <= f`b";
+Goalw [uu_def] "domain(uu(f,b,g,d)) <= f`b";
 by (Blast_tac 1);
 qed "domain_uu_subset";
 
-goal thy "!! a. ALL b<a. f`b lepoll m ==> \
+Goal "!! a. ALL b<a. f`b lepoll m ==> \
 \               ALL b<a. ALL g<a. ALL d<a. domain(uu(f,b,g,d)) lepoll m";
 by (fast_tac (claset() addSEs
         [domain_uu_subset RS subset_imp_lepoll RS lepoll_trans]) 1);
 qed "quant_domain_uu_lepoll_m";
 
-goalw thy [uu_def] "uu(f,b,g,d) <= f`b * f`g";
+Goalw [uu_def] "uu(f,b,g,d) <= f`b * f`g";
 by (Blast_tac 1);
 qed "uu_subset1";
 
-goalw thy [uu_def] "uu(f,b,g,d) <= f`d";
+Goalw [uu_def] "uu(f,b,g,d) <= f`d";
 by (Blast_tac 1);
 qed "uu_subset2";
 
-goal thy "!! a. [| ALL b<a. f`b lepoll m;  d<a |] ==> uu(f,b,g,d) lepoll m";
+Goal "!! a. [| ALL b<a. f`b lepoll m;  d<a |] ==> uu(f,b,g,d) lepoll m";
 by (fast_tac (claset()
         addSEs [uu_subset2 RS subset_imp_lepoll RS lepoll_trans]) 1);
 qed "uu_lepoll_m";
@@ -58,7 +58,7 @@
 (* ********************************************************************** *)
 (* Two cases for lemma ii                                                 *)
 (* ********************************************************************** *)
-goalw thy [lesspoll_def] 
+Goalw [lesspoll_def] 
   "!! a f u. ALL b<a. ALL g<a. ALL d<a. u(f,b,g,d) lepoll m ==>  \
 \            (ALL b<a. f`b ~= 0 --> (EX g<a. EX d<a. u(f,b,g,d) ~= 0 &  \
 \                                       u(f,b,g,d) lesspoll m)) |  \
@@ -71,7 +71,7 @@
 (* ********************************************************************** *)
 (* Lemmas used in both cases                                              *)
 (* ********************************************************************** *)
-goal thy "!!a C. Ord(a) ==> (UN b<a++a. C(b)) = (UN b<a. C(b) Un C(a++b))";
+Goal "!!a C. Ord(a) ==> (UN b<a++a. C(b)) = (UN b<a. C(b) Un C(a++b))";
 by (fast_tac (claset() addSIs [equalityI] addIs [ltI] 
                     addSDs [lt_oadd_disj]
                     addSEs [lt_oadd1, oadd_lt_mono2]) 1);
@@ -82,7 +82,7 @@
 (* Case 1 : lemmas                                                        *)
 (* ********************************************************************** *)
 
-goalw thy [vv1_def] "vv1(f,m,b) <= f`b";
+Goalw [vv1_def] "vv1(f,m,b) <= f`b";
 by (rtac (LetI RS LetI) 1);
 by (split_tac [expand_if] 1);
 by (simp_tac (simpset() addsimps [domain_uu_subset]) 1);
@@ -91,7 +91,7 @@
 (* ********************************************************************** *)
 (* Case 1 : Union of images is the whole "y"                              *)
 (* ********************************************************************** *)
-goalw thy [gg1_def]
+Goalw [gg1_def]
   "!! a f y. [| Ord(a);  m:nat |] ==>   \
 \            (UN b<a++a. gg1(f,a,m)`b) = (UN b<a. f`b)";
 by (asm_simp_tac
@@ -101,14 +101,14 @@
                            vv1_subset RS Diff_partition, ww1_def]) 1);
 qed "UN_gg1_eq";
 
-goal thy "domain(gg1(f,a,m)) = a++a";
+Goal "domain(gg1(f,a,m)) = a++a";
 by (simp_tac (simpset() addsimps [lam_funtype RS domain_of_fun, gg1_def]) 1);
 qed "domain_gg1";
 
 (* ********************************************************************** *)
 (* every value of defined function is less than or equipollent to m       *)
 (* ********************************************************************** *)
-goal thy "!!a b. [| P(a, b);  Ord(a);  Ord(b);  \
+Goal "!!a b. [| P(a, b);  Ord(a);  Ord(b);  \
 \               Least_a = (LEAST a. EX x. Ord(x) & P(a, x)) |]  \
 \               ==> P(Least_a, LEAST b. P(Least_a, b))";
 by (etac ssubst 1);
@@ -122,7 +122,7 @@
         [("P","%g d. domain(uu(f,b,g,d)) ~= 0 &  \
 \               domain(uu(f,b,g,d)) lepoll m")] nested_LeastI);
 
-goalw thy [gg1_def]
+Goalw [gg1_def]
     "!!a. [| Ord(a);  m:nat;  \
 \            ALL b<a. f`b ~=0 -->                                       \
 \            (EX g<a. EX d<a. domain(uu(f,b,g,d)) ~= 0  &               \
@@ -159,7 +159,7 @@
 (* Case 2 : vv2_subset                                                    *)
 (* ********************************************************************** *)
 
-goalw thy [uu_def] "!!f. [| b<a;  g<a;  f`b~=0;  f`g~=0;        \
+Goalw [uu_def] "!!f. [| b<a;  g<a;  f`b~=0;  f`g~=0;        \
 \                           y*y <= y;  (UN b<a. f`b)=y          \
 \                        |] ==> EX d<a. uu(f,b,g,d) ~= 0";
 by (fast_tac (claset() addSIs [not_emptyI] 
@@ -167,7 +167,7 @@
                     addSEs [not_emptyE]) 1);
 qed "ex_d_uu_not_empty";
 
-goal thy "!!f. [| b<a; g<a; f`b~=0; f`g~=0;  \
+Goal "!!f. [| b<a; g<a; f`b~=0; f`g~=0;  \
 \                       y*y<=y; (UN b<a. f`b)=y |]  \
 \               ==> uu(f,b,g,LEAST d. (uu(f,b,g,d) ~= 0)) ~= 0";
 by (dtac ex_d_uu_not_empty 1 THEN REPEAT (assume_tac 1));
@@ -179,7 +179,7 @@
                 sym RSN (2, subst_elem) RS domainI RS not_emptyI] 1));
 qed "not_empty_rel_imp_domain";
 
-goal thy "!!f. [| b<a; g<a; f`b~=0; f`g~=0;  \
+Goal "!!f. [| b<a; g<a; f`b~=0; f`g~=0;  \
 \                       y*y <= y; (UN b<a. f`b)=y |]  \
 \               ==> (LEAST d. uu(f,b,g,d) ~= 0) < a";
 by (eresolve_tac [ex_d_uu_not_empty RS oexE] 1
@@ -193,7 +193,7 @@
 qed "subset_Diff_sing";
 
 (*Could this be proved more directly?*)
-goal thy "!!A B. [| A lepoll m; m lepoll B; B <= A; m:nat |] ==> A=B";
+Goal "!!A B. [| A lepoll m; m lepoll B; B <= A; m:nat |] ==> A=B";
 by (etac natE 1);
 by (fast_tac (claset() addSDs [lepoll_0_is_0] addSIs [equalityI]) 1);
 by (hyp_subst_tac 1);
@@ -207,7 +207,7 @@
         THEN REPEAT (assume_tac 1));
 qed "supset_lepoll_imp_eq";
 
-goal thy
+Goal
  "!!a. [| ALL g<a. ALL d<a. domain(uu(f, b, g, d))~=0 -->               \
 \         domain(uu(f, b, g, d)) eqpoll succ(m);                        \
 \         ALL b<a. f`b lepoll succ(m);  y*y <= y;                       \
@@ -225,7 +225,7 @@
 by (REPEAT (fast_tac (claset() addSIs [domain_uu_subset, nat_succI]) 1));
 qed "uu_Least_is_fun";
 
-goalw thy [vv2_def]
+Goalw [vv2_def]
     "!!a. [| ALL g<a. ALL d<a. domain(uu(f, b, g, d))~=0 -->            \
 \            domain(uu(f, b, g, d)) eqpoll succ(m);                     \
 \            ALL b<a. f`b lepoll succ(m); y*y <= y;                     \
@@ -240,7 +240,7 @@
 (* ********************************************************************** *)
 (* Case 2 : Union of images is the whole "y"                              *)
 (* ********************************************************************** *)
-goalw thy [gg2_def]
+Goalw [gg2_def]
     "!!a. [| ALL g<a. ALL d<a. domain(uu(f,b,g,d)) ~= 0 -->             \
 \            domain(uu(f,b,g,d)) eqpoll succ(m);                        \
 \            ALL b<a. f`b lepoll succ(m); y*y<=y;                       \
@@ -254,7 +254,7 @@
                            vv2_subset RS Diff_partition]) 1);
 qed "UN_gg2_eq";
 
-goal thy "domain(gg2(f,a,b,s)) = a++a";
+Goal "domain(gg2(f,a,b,s)) = a++a";
 by (simp_tac (simpset() addsimps [lam_funtype RS domain_of_fun, gg2_def]) 1);
 qed "domain_gg2";
 
@@ -262,7 +262,7 @@
 (* every value of defined function is less than or equipollent to m       *)
 (* ********************************************************************** *)
 
-goalw thy [vv2_def]
+Goalw [vv2_def]
     "!!m. [| m:nat; m~=0 |] ==> vv2(f,b,g,s) lepoll m";
 by (asm_simp_tac (simpset() addsimps [empty_lepollI]
                               setloop split_tac [expand_if]) 1);
@@ -273,7 +273,7 @@
                 nat_into_Ord, nat_1I]) 1);
 qed "vv2_lepoll";
 
-goalw thy [ww2_def]
+Goalw [ww2_def]
     "!!m. [| ALL b<a. f`b lepoll succ(m);  g<a;  m:nat;  vv2(f,b,g,d) <= f`g  \
 \         |] ==> ww2(f,b,g,d) lepoll m";
 by (excluded_middle_tac "f`g = 0" 1);
@@ -284,7 +284,7 @@
 by (asm_simp_tac (simpset() addsimps [vv2_def, expand_if, not_emptyI]) 1);
 qed "ww2_lepoll";
 
-goalw thy [gg2_def]
+Goalw [gg2_def]
     "!!a. [| ALL g<a. ALL d<a. domain(uu(f,b,g,d)) ~= 0 -->             \
 \            domain(uu(f,b,g,d)) eqpoll succ(m);                        \
 \            ALL b<a. f`b lepoll succ(m);  y*y <= y;                    \
@@ -299,7 +299,7 @@
 (* ********************************************************************** *)
 (* lemma ii                                                               *)
 (* ********************************************************************** *)
-goalw thy [NN_def]
+Goalw [NN_def]
         "!!y. [| succ(m) : NN(y); y*y <= y; m:nat; m~=0 |] ==> m : NN(y)";
 by (REPEAT (eresolve_tac [CollectE, exE, conjE] 1));
 by (resolve_tac [quant_domain_uu_lepoll_m RS cases RS disjE] 1
@@ -331,19 +331,19 @@
 (* the quantifier ALL looks inelegant but makes the proofs shorter  *)
 (* (used only in the following two lemmas)                          *)
 
-goal thy "ALL n:nat. rec(n, x, %k r. r Un r*r) <=  \
+Goal "ALL n:nat. rec(n, x, %k r. r Un r*r) <=  \
 \                    rec(succ(n), x, %k r. r Un r*r)";
 by (fast_tac (claset() addIs [rec_succ RS ssubst]) 1);
 qed "z_n_subset_z_succ_n";
 
-goal thy "!!n. [| ALL n:nat. f(n)<=f(succ(n)); n le m; n : nat; m: nat |]  \
+Goal "!!n. [| ALL n:nat. f(n)<=f(succ(n)); n le m; n : nat; m: nat |]  \
 \              ==> f(n)<=f(m)";
 by (eres_inst_tac [("P","n le m")] rev_mp 1);
 by (res_inst_tac [("P","%z. n le z --> f(n) <= f(z)")] nat_induct 1);
 by (REPEAT (fast_tac le_cs 1));
 qed "le_subsets";
 
-goal thy "!!n m. [| n le m; m:nat |] ==>  \
+Goal "!!n m. [| n le m; m:nat |] ==>  \
 \       rec(n, x, %k r. r Un r*r) <= rec(m, x, %k r. r Un r*r)";
 by (resolve_tac [z_n_subset_z_succ_n RS le_subsets] 1 
     THEN (TRYALL assume_tac));
@@ -351,7 +351,7 @@
     THEN (assume_tac 1));
 qed "le_imp_rec_subset";
 
-goal thy "EX y. x Un y*y <= y";
+Goal "EX y. x Un y*y <= y";
 by (res_inst_tac [("x","UN n:nat. rec(n, x, %k r. r Un r*r)")] exI 1);
 by Safe_tac;
 by (rtac (nat_0I RS UN_I) 1);
@@ -378,7 +378,7 @@
 (*      WO6 ==> NN(y) ~= 0                                                *)
 (* ********************************************************************** *)
 
-goalw thy [WO6_def, NN_def] "!!y. WO6 ==> NN(y) ~= 0";
+Goalw [WO6_def, NN_def] "!!y. WO6 ==> NN(y) ~= 0";
 by (fast_tac (ZF_cs addEs [equals0D]) 1);
 qed "WO6_imp_NN_not_empty";
 
@@ -386,18 +386,18 @@
 (*      1 : NN(y) ==> y can be well-ordered                               *)
 (* ********************************************************************** *)
 
-goal thy "!!f. [| (UN b<a. f`b)=y; x:y; ALL b<a. f`b lepoll 1; Ord(a) |]  \
+Goal "!!f. [| (UN b<a. f`b)=y; x:y; ALL b<a. f`b lepoll 1; Ord(a) |]  \
 \               ==> EX c<a. f`c = {x}";
 by (fast_tac (claset() addSEs [lepoll_1_is_sing]) 1);
 val lemma1 = result();
 
-goal thy "!!f. [| (UN b<a. f`b)=y; x:y; ALL b<a. f`b lepoll 1; Ord(a) |]  \
+Goal "!!f. [| (UN b<a. f`b)=y; x:y; ALL b<a. f`b lepoll 1; Ord(a) |]  \
 \               ==> f` (LEAST i. f`i = {x}) = {x}";
 by (dtac lemma1 1 THEN REPEAT (assume_tac 1));
 by (fast_tac (claset() addSEs [lt_Ord] addIs [LeastI]) 1);
 val lemma2 = result();
 
-goalw thy [NN_def] "!!y. 1 : NN(y) ==> EX a f. Ord(a) & f:inj(y, a)";
+Goalw [NN_def] "!!y. 1 : NN(y) ==> EX a f. Ord(a) & f:inj(y, a)";
 by (etac CollectE 1);
 by (REPEAT (eresolve_tac [exE, conjE] 1));
 by (res_inst_tac [("x","a")] exI 1);
@@ -410,7 +410,7 @@
 by (fast_tac (claset() addSIs [the_equality]) 1);
 qed "NN_imp_ex_inj";
 
-goal thy "!!y. [| y*y <= y; 1 : NN(y) |] ==> EX r. well_ord(y, r)";
+Goal "!!y. [| y*y <= y; 1 : NN(y) |] ==> EX r. well_ord(y, r)";
 by (dtac NN_imp_ex_inj 1);
 by (fast_tac (claset() addSEs [well_ord_Memrel RSN (2,  well_ord_rvimage)]) 1);
 qed "y_well_ord";
@@ -439,11 +439,11 @@
 by (REPEAT (ares_tac prems 1));
 qed "rev_induct";
 
-goalw thy [NN_def] "!!n. n:NN(y) ==> n:nat";
+Goalw [NN_def] "!!n. n:NN(y) ==> n:nat";
 by (etac CollectD1 1);
 qed "NN_into_nat";
 
-goal thy "!!n. [| n:NN(y); y*y <= y; n~=0 |] ==> 1:NN(y)";
+Goal "!!n. [| n:NN(y); y*y <= y; n~=0 |] ==> 1:NN(y)";
 by (rtac rev_induct 1 THEN REPEAT (ares_tac [NN_into_nat] 1));
 by (rtac lemma_ii 1 THEN REPEAT (assume_tac 1));
 val lemma3 = result();
@@ -453,12 +453,12 @@
 (* ********************************************************************** *)
 
 (* another helpful lemma *)
-goalw thy [NN_def] "!!y. 0:NN(y) ==> y=0";
+Goalw [NN_def] "!!y. 0:NN(y) ==> y=0";
 by (fast_tac (claset() addSIs [equalityI] 
                     addSDs [lepoll_0_is_0] addEs [subst]) 1);
 qed "NN_y_0";
 
-goalw thy [WO1_def] "!!Z. WO6 ==> WO1";
+Goalw [WO1_def] "!!Z. WO6 ==> WO1";
 by (rtac allI 1);
 by (excluded_middle_tac "A=0" 1);
 by (fast_tac (claset() addSIs [well_ord_Memrel, nat_0I RS nat_into_Ord]) 2);
--- a/src/ZF/AC/WO_AC.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/AC/WO_AC.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -7,16 +7,16 @@
 
 open WO_AC;
 
-goal thy "!!A. [| well_ord(Union(A),r); 0~:A; B:A |]  \
+Goal "!!A. [| well_ord(Union(A),r); 0~:A; B:A |]  \
 \               ==> (THE b. first(b,B,r)) : B";
 by (fast_tac (claset() addSEs [well_ord_imp_ex1_first RS theI RS
                 (first_def RS def_imp_iff RS iffD1 RS conjunct1)]) 1);
 qed "first_in_B";
 
-goal thy "!!A. [| well_ord(Union(A), R); 0~:A |] ==> EX f. f:(PROD X:A. X)";
+Goal "!!A. [| well_ord(Union(A), R); 0~:A |] ==> EX f. f:(PROD X:A. X)";
 by (fast_tac (claset() addSEs [first_in_B] addSIs [lam_type]) 1);
 qed "ex_choice_fun";
 
-goal thy "!!A. well_ord(A, R) ==> EX f. f:(PROD X: Pow(A)-{0}. X)";
+Goal "!!A. well_ord(A, R) ==> EX f. f:(PROD X: Pow(A)-{0}. X)";
 by (fast_tac (claset() addSEs [well_ord_subset RS ex_choice_fun]) 1);
 qed "ex_choice_fun_Pow";
--- a/src/ZF/AC/recfunAC16.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/AC/recfunAC16.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -11,11 +11,11 @@
 (* Basic properties of recfunAC16                                         *)
 (* ********************************************************************** *)
 
-goalw thy [recfunAC16_def] "recfunAC16(f,fa,0,a) = 0";
+Goalw [recfunAC16_def] "recfunAC16(f,fa,0,a) = 0";
 by (rtac transrec2_0 1);
 qed "recfunAC16_0";
 
-goalw thy [recfunAC16_def] "recfunAC16(f,fa,succ(i),a) =  \
+Goalw [recfunAC16_def] "recfunAC16(f,fa,succ(i),a) =  \
 \       if (EX y:recfunAC16(f,fa,i,a). fa ` i <= y, recfunAC16(f,fa,i,a), \
 \       recfunAC16(f,fa,i,a) Un {f ` (LEAST j. fa ` i <= f ` j &  \
 \       (ALL b<a. (fa`b <= f`j  \
@@ -23,7 +23,7 @@
 by (rtac transrec2_succ 1);
 qed "recfunAC16_succ";
 
-goalw thy [recfunAC16_def] "!!i. Limit(i)  \
+Goalw [recfunAC16_def] "!!i. Limit(i)  \
 \       ==> recfunAC16(f,fa,i,a) = (UN j<i. recfunAC16(f,fa,j,a))";
 by (etac transrec2_Limit 1);
 qed "recfunAC16_Limit";
@@ -57,7 +57,7 @@
 (* Monotonicity of recfunAC16                                             *)
 (* ********************************************************************** *)
 
-goalw thy [recfunAC16_def]
+Goalw [recfunAC16_def]
         "!!i. i le j ==> recfunAC16(f, g, i, a) <= recfunAC16(f, g, j, a)";
 by (rtac transrec2_mono 1);
 by (REPEAT (fast_tac (claset() addIs [expand_if RS iffD2]) 1));
--- a/src/ZF/Coind/ECR.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/Coind/ECR.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -8,7 +8,7 @@
 
 (* Specialised co-induction rule *)
 
-goal ECR.thy
+Goal
  "!!x.[| x:ExVar; e:Exp; t:Ty; ve:ValEnv; te:TyEnv; \
 \    <te, e_fn(x, e), t>:ElabRel; ve_dom(ve) = te_dom(te); \
 \    {<ve_app(ve,y),te_app(te,y)>.y:ve_dom(ve)}: \
@@ -46,7 +46,7 @@
 
 (* Properties of the pointwise extension to environments *)
 
-goalw ECR.thy [hastyenv_def]
+Goalw [hastyenv_def]
   "!!ve.[| ve:ValEnv; te:TyEnv; hastyenv(ve,te); <v,t>:HasTyRel |] ==> \
 \   hastyenv(ve_owr(ve,x,v),te_owr(te,x,t))";
 by Safe_tac;
@@ -67,7 +67,7 @@
 by (asm_simp_tac (simpset() addsimps [ve_app_owr1,te_app_owr1]) 1);
 qed "hastyenv_owr";
 
-goalw ECR.thy  [isofenv_def,hastyenv_def]
+Goalw  [isofenv_def,hastyenv_def]
   "!!ve.[| ve:ValEnv; te:TyEnv; isofenv(ve,te) |] ==> hastyenv(ve,te)";
 by Safe_tac;
 by (dtac bspec 1);
--- a/src/ZF/Coind/MT.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/Coind/MT.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -11,21 +11,21 @@
 (* The Consistency theorem                                      *)
 (* ############################################################ *)
 
-goal MT.thy 
+Goal 
   "!!t. [| c:Const; hastyenv(ve,te);<te,e_const(c),t>:ElabRel |] ==>         \
 \       <v_const(c), t> : HasTyRel";
 by (Fast_tac 1);
 qed "consistency_const";
 
 
-goalw  MT.thy [hastyenv_def]
+Goalw [hastyenv_def]
   "!!t. [| x:ve_dom(ve); hastyenv(ve,te); <te,e_var(x),t>:ElabRel |] ==>     \
 \  <ve_app(ve,x),t>:HasTyRel";
 by (Fast_tac 1);
 qed "consistency_var";
 
 
-goalw MT.thy [hastyenv_def]
+Goalw [hastyenv_def]
   "!!t. [| ve:ValEnv; x:ExVar; e:Exp; hastyenv(ve,te);       \
 \          <te,e_fn(x,e),t>:ElabRel  \
 \       |] ==> <v_clos(x, e, ve), t> : HasTyRel";
@@ -143,7 +143,7 @@
 
 fun mt_cases_tac t = ((rtac t 1) THEN (TRYALL assume_tac)); 
 
-goal MT.thy 
+Goal 
   "!!e. <ve,e,v>:EvalRel ==>         \
 \       (ALL t te. hastyenv(ve,te) --> <te,e,t>:ElabRel --> <v,t>:HasTyRel)";
 by (etac EvalRel.induct 1);
--- a/src/ZF/Coind/Map.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/Coind/Map.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -9,22 +9,22 @@
 (** Some sample proofs of inclusions for the final coalgebra "U" (by lcp) **)
 
 
-goalw Map.thy [TMap_def]
+Goalw [TMap_def]
      "{0,1} <= {1} Un TMap(I, {0,1})";
 by (Blast_tac 1);
 result();
 
-goalw Map.thy [TMap_def]
+Goalw [TMap_def]
      "{0} Un TMap(I,1) <= {1} Un TMap(I, {0} Un TMap(I,1))";
 by (Blast_tac 1);
 result();
 
-goalw Map.thy [TMap_def]
+Goalw [TMap_def]
      "{0,1} Un TMap(I,2) <= {1} Un TMap(I, {0,1} Un TMap(I,2))";
 by (Blast_tac 1);
 result();
 
-goalw Map.thy [TMap_def]
+Goalw [TMap_def]
      "{0,1} Un TMap(I,TMap(I,2)) Un TMap(I,2) <= \
 \     {1} Un TMap(I, {0,1} Un TMap(I,TMap(I,2)) Un TMap(I,2))";
 (*twice as fast as Blast_tac alone*)
@@ -36,19 +36,19 @@
 (* Lemmas                                                       *)
 (* ############################################################ *)
 
-goal Map.thy "!!A. a:A ==> Sigma(A,B)``{a} = B(a)";
+Goal "!!A. a:A ==> Sigma(A,B)``{a} = B(a)";
 by (Fast_tac 1);
 qed "qbeta";
 
-goal Map.thy "!!A. a~:A ==> Sigma(A,B)``{a} = 0";
+Goal "!!A. a~:A ==> Sigma(A,B)``{a} = 0";
 by (Fast_tac 1);
 qed "qbeta_emp";
 
-goal Map.thy "!!A. a ~: A ==> Sigma(A,B)``{a}=0";
+Goal "!!A. a ~: A ==> Sigma(A,B)``{a}=0";
 by (Fast_tac 1);
 qed "image_Sigma1";
 
-goal Map.thy "0``A = 0";
+Goal "0``A = 0";
 by (Fast_tac 1);
 qed "image_02";
 
@@ -58,7 +58,7 @@
 
 (* Lemmas *)
 
-goalw Map.thy [quniv_def]
+Goalw [quniv_def]
     "!!A. A <= univ(X) ==> Pow(A * Union(quniv(X))) <= quniv(X)";
 by (rtac Pow_mono 1);
 by (rtac ([Sigma_mono, product_univ] MRS subset_trans) 1);
@@ -82,7 +82,7 @@
 (* Monotonicity                                                 *)
 (* ############################################################ *)
 
-goalw Map.thy [PMap_def,TMap_def] "!!A. B<=C ==> PMap(A,B)<=PMap(A,C)";
+Goalw [PMap_def,TMap_def] "!!A. B<=C ==> PMap(A,B)<=PMap(A,C)";
 by (Fast_tac 1);
 qed "map_mono";
 
@@ -94,7 +94,7 @@
 
 (** map_emp **)
 
-goalw Map.thy [map_emp_def,PMap_def,TMap_def] "map_emp:PMap(A,B)";
+Goalw [map_emp_def,PMap_def,TMap_def] "map_emp:PMap(A,B)";
 by Safe_tac;
 by (rtac image_02 1);
 qed "pmap_empI";
@@ -102,7 +102,7 @@
 (** map_owr **)
 
 
-goalw Map.thy [map_owr_def,PMap_def,TMap_def] 
+Goalw [map_owr_def,PMap_def,TMap_def] 
   "!! A.[| m:PMap(A,B); a:A; b:B |]  ==> map_owr(m,a,b):PMap(A,B)";
 by Safe_tac;
 by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [if_iff])));
@@ -123,7 +123,7 @@
 
 (** map_app **)
 
-goalw Map.thy [TMap_def,map_app_def] 
+Goalw [TMap_def,map_app_def] 
   "!!m.[| m:TMap(A,B); a:domain(m) |] ==> map_app(m,a) ~=0";
 by (etac domainE 1);
 by (dtac imageI 1);
@@ -131,12 +131,12 @@
 by (etac not_emptyI 1);
 qed "tmap_app_notempty";
 
-goalw Map.thy [TMap_def,map_app_def,domain_def] 
+Goalw [TMap_def,map_app_def,domain_def] 
   "!!m.[| m:TMap(A,B); a:domain(m) |] ==> map_app(m,a):B";
 by (Fast_tac 1);
 qed "tmap_appI";
 
-goalw Map.thy [PMap_def]
+Goalw [PMap_def]
   "!!m.[| m:PMap(A,B); a:domain(m) |] ==> map_app(m,a):B";
 by (forward_tac [tmap_app_notempty] 1); 
 by (assume_tac 1);
@@ -147,12 +147,12 @@
 
 (** domain **)
 
-goalw Map.thy [TMap_def]
+Goalw [TMap_def]
   "!!m.[| m:TMap(A,B); a:domain(m) |] ==> a:A";
 by (Fast_tac 1);
 qed "tmap_domainD";
 
-goalw Map.thy [PMap_def,TMap_def]
+Goalw [PMap_def,TMap_def]
   "!!m.[| m:PMap(A,B); a:domain(m) |] ==> a:A";
 by (Fast_tac 1);
 qed "pmap_domainD";
@@ -165,22 +165,22 @@
 
 (* Lemmas *)
 
-goal Map.thy  "domain(UN x:A. B(x)) = (UN x:A. domain(B(x)))";
+Goal  "domain(UN x:A. B(x)) = (UN x:A. domain(B(x)))";
 by (Fast_tac 1);
 qed "domain_UN";
 
-goal Map.thy  "domain(Sigma(A,B)) = {x:A. EX y. y:B(x)}";
+Goal  "domain(Sigma(A,B)) = {x:A. EX y. y:B(x)}";
 by (simp_tac (simpset() addsimps [domain_UN,domain_0,domain_cons]) 1);
 by (Fast_tac 1);
 qed "domain_Sigma";
 
 (* Theorems *)
 
-goalw Map.thy [map_emp_def] "domain(map_emp) = 0";
+Goalw [map_emp_def] "domain(map_emp) = 0";
 by (Fast_tac 1);
 qed "map_domain_emp";
 
-goalw Map.thy [map_owr_def] 
+Goalw [map_owr_def] 
   "!!a. b ~= 0 ==> domain(map_owr(f,a,b)) = {a} Un domain(f)";
 by (simp_tac (simpset() addsimps [domain_Sigma]) 1);
 by (rtac equalityI 1);
@@ -197,14 +197,14 @@
 
 (** Application **)
 
-goalw Map.thy [map_app_def,map_owr_def] 
+Goalw [map_app_def,map_owr_def] 
   "map_app(map_owr(f,a,b),a) = b";
 by (stac qbeta 1);
 by (Fast_tac 1);
 by (Simp_tac 1);
 qed "map_app_owr1";
 
-goalw Map.thy [map_app_def,map_owr_def] 
+Goalw [map_app_def,map_owr_def] 
   "!!a. c ~= a ==> map_app(map_owr(f,a,b),c)= map_app(f,c)";
 by (rtac (excluded_middle RS disjE) 1);
 by (stac qbeta_emp 1);
--- a/src/ZF/Coind/Types.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/Coind/Types.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -9,40 +9,40 @@
 val te_owrE = 
   TyEnv.mk_cases TyEnv.con_defs "te_owr(te,f,t):TyEnv";
 
-goalw Types.thy TyEnv.con_defs "rank(te) < rank(te_owr(te,x,t))";
+Goalw TyEnv.con_defs "rank(te) < rank(te_owr(te,x,t))";
 by (simp_tac rank_ss 1);
 qed "rank_te_owr1";
 
-goal Types.thy "te_rec(te_emp,c_te_emp,f_te_owr) = c_te_emp";
+Goal "te_rec(te_emp,c_te_emp,f_te_owr) = c_te_emp";
 by (rtac (te_rec_def RS def_Vrec RS trans) 1);
 by (simp_tac (simpset() addsimps (rank_te_owr1::TyEnv.case_eqns)) 1);
 qed "te_rec_emp";
 
-goal Types.thy 
+Goal 
   " te_rec(te_owr(te,x,t),c_te_emp,f_te_owr) = \
 \   f_te_owr(te,x,t,te_rec(te,c_te_emp,f_te_owr))";
 by (rtac (te_rec_def RS def_Vrec RS trans) 1);
 by (simp_tac (rank_ss addsimps (rank_te_owr1::TyEnv.case_eqns)) 1);
 qed "te_rec_owr";
 
-goalw Types.thy [te_dom_def] "te_dom(te_emp) = 0";
+Goalw [te_dom_def] "te_dom(te_emp) = 0";
 by (simp_tac (simpset() addsimps [te_rec_emp]) 1);
 qed "te_dom_emp";
 
-goalw Types.thy [te_dom_def] "te_dom(te_owr(te,x,v)) = te_dom(te) Un {x}";
+Goalw [te_dom_def] "te_dom(te_owr(te,x,v)) = te_dom(te) Un {x}";
 by (simp_tac (simpset() addsimps [te_rec_owr]) 1);
 qed "te_dom_owr";
 
-goalw Types.thy [te_app_def] "te_app(te_owr(te,x,t),x) = t";
+Goalw [te_app_def] "te_app(te_owr(te,x,t),x) = t";
 by (simp_tac (simpset() addsimps [te_rec_owr]) 1);
 qed "te_app_owr1";
 
-goalw Types.thy [te_app_def]
+Goalw [te_app_def]
   "!!x y. x ~= y ==> te_app(te_owr(te,x,t),y) = te_app(te,y)";
 by (asm_simp_tac (simpset() addsimps [te_rec_owr,(not_sym RS if_not_P)]) 1);
 qed "te_app_owr2";
 
-goal Types.thy
+Goal
     "!!te. [| te:TyEnv; x:ExVar; x:te_dom(te) |] ==> te_app(te,x):Ty";
 by (res_inst_tac [("P","x:te_dom(te)")] impE 1);
 by (assume_tac 2);
--- a/src/ZF/Coind/Values.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/Coind/Values.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -49,7 +49,7 @@
 by (Fast_tac 1);
 qed "set_notemptyE";
 
-goalw Values.thy (sum_def::QPair_def::QInl_def::QInr_def::Val_ValEnv.con_defs)
+Goalw (sum_def::QPair_def::QInl_def::QInr_def::Val_ValEnv.con_defs)
   "v_clos(x,e,ve) ~= 0";
 by (rtac not_emptyI 1);
 by (rtac UnI2 1);
@@ -59,7 +59,7 @@
 by (Fast_tac 1);
 qed "v_closNE";
 
-goalw Values.thy (sum_def::QPair_def::QInl_def::QInr_def::Val_ValEnv.con_defs)
+Goalw (sum_def::QPair_def::QInl_def::QInr_def::Val_ValEnv.con_defs)
   "!!c. c:Const ==> v_const(c) ~= 0";
 by (dtac constNEE 1);
 by (etac not_emptyE 1);
@@ -75,7 +75,7 @@
 
 (* Proving that the empty set is not a value *)
 
-goal Values.thy "!!v. v:Val ==> v ~= 0";
+Goal "!!v. v:Val ==> v ~= 0";
 by (etac ValE 1);
 by (ALLGOALS hyp_subst_tac);
 by (etac v_constNE 1);
@@ -94,14 +94,14 @@
 by (rtac Un_commute 1);
 qed "ve_dom_owr";
 
-goalw Values.thy [ve_app_def,ve_owr_def]
+Goalw [ve_app_def,ve_owr_def]
 "!!ve. ve:ValEnv ==> ve_app(ve_owr(ve,x,v),x) = v"; 
 by (etac ValEnvE 1);
 by (asm_full_simp_tac (simpset() addsimps Val_ValEnv.case_eqns) 1);
 by (rtac map_app_owr1 1);
 qed "ve_app_owr1";
 
-goalw Values.thy [ve_app_def,ve_owr_def]
+Goalw [ve_app_def,ve_owr_def]
  "!!ve. ve:ValEnv ==> x ~= y ==> ve_app(ve_owr(ve,x,v),y) = ve_app(ve,y)";
 by (etac ValEnvE 1);
 by (asm_full_simp_tac (simpset() addsimps Val_ValEnv.case_eqns) 1);
@@ -111,7 +111,7 @@
 
 (* Introduction rules for operators on value environments *)
 
-goalw Values.thy [ve_app_def,ve_owr_def,ve_dom_def]
+Goalw [ve_app_def,ve_owr_def,ve_dom_def]
   "!!ve.[| ve:ValEnv; x:ve_dom(ve) |] ==> ve_app(ve,x):Val";
 by (etac ValEnvE 1);
 by (hyp_subst_tac 1);
@@ -121,7 +121,7 @@
 by (assume_tac 1);
 qed "ve_appI";
 
-goalw Values.thy [ve_dom_def]
+Goalw [ve_dom_def]
   "!!ve.[| ve:ValEnv; x:ve_dom(ve) |] ==> x:ExVar";
 by (etac ValEnvE 1);
 by (hyp_subst_tac 1);
@@ -131,12 +131,12 @@
 by (assume_tac 1);
 qed "ve_domI";
 
-goalw Values.thy [ve_emp_def] "ve_emp:ValEnv";
+Goalw [ve_emp_def] "ve_emp:ValEnv";
 by (rtac Val_ValEnv.ve_mk_I 1);
 by (rtac pmap_empI 1);
 qed "ve_empI";
 
-goalw Values.thy [ve_owr_def] 
+Goalw [ve_owr_def] 
   "!!ve.[|ve:ValEnv; x:ExVar; v:Val |] ==> ve_owr(ve,x,v):ValEnv";
 by (etac ValEnvE 1);
 by (hyp_subst_tac 1);
--- a/src/ZF/IMP/Equiv.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/IMP/Equiv.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -44,7 +44,7 @@
 val bexp2 = bexp_iff RS iffD2;
 
 
-goal Equiv.thy "!!c. <c,sigma> -c-> sigma' ==> <sigma,sigma'> : C(c)";
+Goal "!!c. <c,sigma> -c-> sigma' ==> <sigma,sigma'> : C(c)";
 by (etac evalc.induct 1);
 by (ALLGOALS (asm_simp_tac (simpset() addsimps [bexp1])));
 (* skip *)
@@ -70,7 +70,7 @@
 AddEs  [C_type,C_type_fst];
 
 
-goal Equiv.thy "!!c. c : com ==> ALL x:C(c). <c,fst(x)> -c-> snd(x)";
+Goal "!!c. c : com ==> ALL x:C(c). <c,fst(x)> -c-> snd(x)";
 by (etac com.induct 1);
 (* skip *)
 by (fast_tac (claset() addss (simpset())) 1);
@@ -93,7 +93,7 @@
 
 (**** Proof of Equivalence ****)
 
-goal Equiv.thy
+Goal
     "ALL c:com. C(c) = {io:(loc->nat)*(loc->nat). <c,fst(io)> -c-> snd(io)}";
 by (fast_tac (claset() addIs [C_subset RS subsetD]
 		       addEs [com2 RS bspec]
--- a/src/ZF/Resid/Confluence.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/Resid/Confluence.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -11,7 +11,7 @@
 (*        strip lemmas                                                       *)
 (* ------------------------------------------------------------------------- *)
 
-goalw Confluence.thy [confluence_def,strip_def] 
+Goalw [confluence_def,strip_def] 
     "!!u.[|confluence(Spar_red1)|]==> strip";
 by (resolve_tac [impI RS allI RS allI] 1);
 by (etac Spar_red.induct 1);
@@ -20,7 +20,7 @@
 qed "strip_lemma_r";
 
 
-goalw Confluence.thy [confluence_def,strip_def] 
+Goalw [confluence_def,strip_def] 
     "!!u. strip==> confluence(Spar_red)";
 by (resolve_tac [impI RS allI RS allI] 1);
 by (etac Spar_red.induct 1);
@@ -38,7 +38,7 @@
 (* ------------------------------------------------------------------------- *)
 
 
-goalw Confluence.thy [confluence_def] "confluence(Spar_red1)";
+Goalw [confluence_def] "confluence(Spar_red1)";
 by (Clarify_tac 1);
 by (forward_tac [simulation] 1);
 by (forw_inst_tac [("n","z")] simulation 1);
@@ -51,7 +51,7 @@
 bind_thm ("confluence_parallel_reduction",
 	  parallel_moves RS strip_lemma_r RS strip_lemma_l);
 
-goalw Confluence.thy [confluence_def] 
+Goalw [confluence_def] 
     "!!u.[|confluence(Spar_red)|]==> confluence(Sred)";
 by (blast_tac (claset() addIs [par_red_red, red_par_red]) 1);
 val lemma1 = result();
--- a/src/ZF/Resid/Conversion.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/Resid/Conversion.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -9,7 +9,7 @@
 
 AddIs (Sconv.intrs @ Sconv1.intrs);
 
-goal Conversion.thy  
+Goal  
     "!!u. m<--->n ==> n<--->m";
 by (etac Sconv.induct 1);
 by (etac Sconv1.induct 1);
@@ -20,7 +20,7 @@
 (*      Church_Rosser Theorem                                                *)
 (* ------------------------------------------------------------------------- *)
 
-goal Conversion.thy  
+Goal  
     "!!u. m<--->n ==> EX p.(m --->p) & (n ---> p)";
 by (etac Sconv.induct 1);
 by (etac Sconv1.induct 1);
--- a/src/ZF/Resid/Cube.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/Resid/Cube.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -17,7 +17,7 @@
 (* ------------------------------------------------------------------------- *)
 
 (* Having more assumptions than needed -- removed below  *)
-goal Cube.thy 
+Goal 
     "!!u. v<==u ==> \
 \   regular(u)-->(ALL w. w~v-->w~u-->  \
 \             w |> u = (w|>v) |> (u|>v))";
@@ -31,7 +31,7 @@
 by (asm_full_simp_tac prism_ss 1);
 qed_spec_mp "prism_l";
 
-goal Cube.thy 
+Goal 
     "!!u.[|v <== u; regular(u); w~v|]==> \
 \       w |> u = (w|>v) |> (u|>v)";
 by (rtac prism_l 1);
@@ -45,7 +45,7 @@
 (*    Levy's Cube Lemma                                                      *)
 (* ------------------------------------------------------------------------- *)
 
-goal Cube.thy 
+Goal 
     "!!u.[|u~v; regular(v); regular(u); w~u|]==>  \
 \          (w|>u) |> (v|>u) = (w|>v) |> (u|>v)";
 by (stac preservation 1 
@@ -68,7 +68,7 @@
 (*           paving theorem                                                  *)
 (* ------------------------------------------------------------------------- *)
 
-goal Cube.thy 
+Goal 
     "!!u.[|w~u; w~v; regular(u); regular(v)|]==> \
 \          EX uv vu. (w|>u) |> vu = (w|>v) |> uv & (w|>u)~vu &\
 \            regular(vu) & (w|>v)~uv & regular(uv) ";
--- a/src/ZF/Resid/Redex.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/Resid/Redex.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -11,17 +11,17 @@
 (*      redex_rec conversions                                                *)
 (* ------------------------------------------------------------------------- *)
 
-goal Redex.thy  "redex_rec(Var(n),b,c,d) = b(n)";
+Goal  "redex_rec(Var(n),b,c,d) = b(n)";
 by (rtac (redex_rec_def RS def_Vrec RS trans) 1);
 by (simp_tac (rank_ss addsimps redexes.con_defs)  1);
 qed "redex_rec_Var";
 
-goal Redex.thy  "redex_rec(Fun(t),b,c,d) = c(t,redex_rec(t,b,c,d))";
+Goal  "redex_rec(Fun(t),b,c,d) = c(t,redex_rec(t,b,c,d))";
 by (rtac (redex_rec_def RS def_Vrec RS trans) 1);
 by (simp_tac (rank_ss addsimps redexes.con_defs)  1);
 qed "redex_rec_Fun";
 
-goal Redex.thy  "redex_rec(App(m,f,a),b,c,d) =  \
+Goal  "redex_rec(App(m,f,a),b,c,d) =  \
 \                      d(m,f,a,redex_rec(f,b,c,d),redex_rec(a,b,c,d))";
 by (rtac (redex_rec_def RS def_Vrec RS trans) 1);
 by (simp_tac (rank_ss addsimps redexes.con_defs)  1);
--- a/src/ZF/Resid/Reduction.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/Resid/Reduction.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -37,33 +37,33 @@
 (*     Lemmas for reduction                                                  *)
 (* ------------------------------------------------------------------------- *)
 
-goal Reduction.thy  "!!u. m--->n ==> Fun(m) ---> Fun(n)";
+Goal  "!!u. m--->n ==> Fun(m) ---> Fun(n)";
 by (etac Sred.induct 1);
 by (resolve_tac [Sred.trans] 3);
 by (ALLGOALS (Asm_simp_tac ));
 qed "red_Fun";
 
-goal Reduction.thy  
+Goal  
     "!!u.[|n:lambda; m ---> m'|]==> Apl(m,n)--->Apl(m',n)";
 by (etac Sred.induct 1);
 by (resolve_tac [Sred.trans] 3);
 by (ALLGOALS (Asm_simp_tac ));
 qed "red_Apll";
 
-goal Reduction.thy  
+Goal  
     "!!u.[|n:lambda; m ---> m'|]==> Apl(n,m)--->Apl(n,m')";
 by (etac Sred.induct 1);
 by (resolve_tac [Sred.trans] 3);
 by (ALLGOALS (Asm_simp_tac ));
 qed "red_Aplr";
 
-goal Reduction.thy  
+Goal  
     "!!u.[|m ---> m'; n--->n'|]==> Apl(m,n)--->Apl(m',n')";
 by (res_inst_tac [("n","Apl(m',n)")] Sred.trans 1);
 by (ALLGOALS (asm_simp_tac (simpset() addsimps [red_Apll,red_Aplr]) ));
 qed "red_Apl";
 
-goal Reduction.thy  
+Goal  
     "!!u.[|m:lambda; m':lambda; n:lambda; n':lambda; m ---> m'; n--->n'|]==> \
 \              Apl(Fun(m),n)---> n'/m'";
 by (res_inst_tac [("n","Apl(Fun(m'),n')")] Sred.trans 1);
@@ -76,23 +76,23 @@
 (* ------------------------------------------------------------------------- *)
 
 
-goal Reduction.thy "!!u. m:lambda==> m =1=> m";
+Goal "!!u. m:lambda==> m =1=> m";
 by (eresolve_tac [lambda.induct] 1);
 by (ALLGOALS (Asm_simp_tac ));
 qed "refl_par_red1";
 
-goal Reduction.thy "!!u. m-1->n ==> m=1=>n";
+Goal "!!u. m-1->n ==> m=1=>n";
 by (etac Sred1.induct 1);
 by (ALLGOALS (asm_simp_tac (simpset() addsimps [refl_par_red1]) ));
 qed "red1_par_red1";
 
-goal Reduction.thy "!!u. m--->n ==> m===>n";
+Goal "!!u. m--->n ==> m===>n";
 by (etac Sred.induct 1);
 by (resolve_tac [Spar_red.trans] 3);
 by (ALLGOALS (asm_simp_tac (simpset() addsimps [refl_par_red1,red1_par_red1]) ));
 qed "red_par_red";
 
-goal Reduction.thy "!!u. m===>n ==> m--->n";
+Goal "!!u. m===>n ==> m--->n";
 by (etac Spar_red.induct 1);
 by (etac Spar_red1.induct 1);
 by (resolve_tac [Sred.trans] 5);
@@ -104,7 +104,7 @@
 (*      Simulation                                                           *)
 (* ------------------------------------------------------------------------- *)
 
-goal Reduction.thy  
+Goal  
     "!!u. m=1=>n ==> EX v. m|>v = n & m~v & regular(v)";
 by (etac Spar_red1.induct 1);
 by Safe_tac;
@@ -118,14 +118,14 @@
 (*           commuting of unmark and subst                                   *)
 (* ------------------------------------------------------------------------- *)
 
-goal Reduction.thy  
+Goal  
     "!!u. u:redexes ==> \
 \           ALL k:nat. unmark(lift_rec(u,k)) = lift_rec(unmark(u),k)";
 by (eresolve_tac [redexes.induct] 1);
 by (ALLGOALS (asm_full_simp_tac (addsplit (simpset()))));
 qed "unmmark_lift_rec";
 
-goal Reduction.thy  
+Goal  
     "!!u. v:redexes ==> ALL k:nat. ALL u:redexes.  \
 \         unmark(subst_rec(u,v,k)) = subst_rec(unmark(u),unmark(v),k)";
 by (eresolve_tac [redexes.induct] 1);
@@ -138,7 +138,7 @@
 (*        Completeness                                                       *)
 (* ------------------------------------------------------------------------- *)
 
-goal Reduction.thy  
+Goal  
     "!!u. u~v ==> regular(v) --> unmark(u) =1=> unmark(u|>v)";
 by (etac Scomp.induct 1);
 by (ALLGOALS (asm_full_simp_tac (reducL_ss addsimps [unmmark_subst_rec]) ));
@@ -146,7 +146,7 @@
 by (asm_full_simp_tac reducL_ss 1);
 qed_spec_mp "completeness_l";
 
-goal Reduction.thy  
+Goal  
     "!!u.[|u:lambda; u~v; regular(v)|]==> u =1=> unmark(u|>v)";
 by (dtac completeness_l 1);
 by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [lambda_unmark]) ));
--- a/src/ZF/Resid/Residuals.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/Resid/Residuals.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -43,13 +43,13 @@
 (*       residuals is a  partial function                                    *)
 (* ------------------------------------------------------------------------- *)
 
-goal Residuals.thy 
+Goal 
     "!!u. residuals(u,v,w) ==> ALL w1. residuals(u,v,w1) --> w1 = w";
 by (etac Sres.induct 1);
 by (ALLGOALS Fast_tac);
 qed_spec_mp "residuals_function";
 
-goal Residuals.thy 
+Goal 
     "!!u. u~v ==> regular(v) --> (EX w. residuals(u,v,w))";
 by (etac Scomp.induct 1);
 by (ALLGOALS Fast_tac);
@@ -73,29 +73,29 @@
 val resfunc_cs = (claset() addIs  [the_equality,residuals_function] 
                           addSEs [comp_resfuncE]);
 
-goalw Residuals.thy [res_func_def]
+Goalw [res_func_def]
     "!!n. n:nat ==> Var(n) |> Var(n) = Var(n)";
 by (fast_tac resfunc_cs 1);
 qed "res_Var";
 
-goalw Residuals.thy [res_func_def]
+Goalw [res_func_def]
     "!!n.[|s~t; regular(t)|]==> Fun(s) |> Fun(t) = Fun(s |> t)";
 by (fast_tac resfunc_cs 1);
 qed "res_Fun";
 
-goalw Residuals.thy [res_func_def]
+Goalw [res_func_def]
     "!!n.[|s~u; regular(u); t~v; regular(v); b:bool|]==> \
 \           App(b,s,t) |> App(0,u,v) = App(b, s |> u, t |> v)";
 by (fast_tac resfunc_cs 1);
 qed "res_App";
 
-goalw Residuals.thy [res_func_def]
+Goalw [res_func_def]
     "!!n.[|s~u; regular(u); t~v; regular(v); b:bool|]==> \
 \           App(b,Fun(s),t) |> App(1,Fun(u),v) = (t |> v)/ (s |> u)";
 by (fast_tac resfunc_cs 1);
 qed "res_redex";
 
-goal Residuals.thy
+Goal
     "!!n.[|s~t; regular(t)|]==> regular(t) --> s |> t:redexes";
 by (etac Scomp.induct 1);
 by (ALLGOALS (asm_full_simp_tac 
@@ -118,19 +118,19 @@
 (*     Commutation theorem                                                   *)
 (* ------------------------------------------------------------------------- *)
 
-goal Residuals.thy 
+Goal 
     "!!u. u<==v ==> u~v";
 by (etac Ssub.induct 1);
 by (ALLGOALS Asm_simp_tac);
 qed "sub_comp";
 
-goal Residuals.thy 
+Goal 
     "!!u. u<==v  ==> regular(v) --> regular(u)";
 by (etac Ssub.induct 1);
 by (ALLGOALS (asm_simp_tac res1L_ss));
 qed "sub_preserve_reg";
 
-goal Residuals.thy 
+Goal 
     "!!u.[|u~v; k:nat|]==> regular(v)--> (ALL n:nat.  \
 \        lift_rec(u,n) |> lift_rec(v,n) = lift_rec(u |> v,n))";
 by (etac Scomp.induct 1);
@@ -140,7 +140,7 @@
 by (Asm_full_simp_tac 1);
 qed "residuals_lift_rec";
 
-goal Residuals.thy 
+Goal 
     "!!u. u1~u2 ==>  ALL v1 v2. v1~v2 --> regular(v2) --> regular(u2) -->\
 \   (ALL n:nat. subst_rec(v1,u1,n) |> subst_rec(v2,u2,n) = \
 \              subst_rec(v1 |> v2, u1 |> u2,n))";
@@ -153,7 +153,7 @@
 qed "residuals_subst_rec";
 
 
-goal Residuals.thy 
+Goal 
     "!!u.[|u1~u2; v1~v2; regular(u2); regular(v2)|]==>\
 \       (v1/u1) |> (v2/u2) = (v1 |> v2)/(u1 |> u2)";
 by (asm_simp_tac (simpset() addsimps ([residuals_subst_rec])) 1);
@@ -163,7 +163,7 @@
 (*     Residuals are comp and regular                                        *)
 (* ------------------------------------------------------------------------- *)
 
-goal Residuals.thy 
+Goal 
     "!!u. u~v ==> ALL w. u~w --> v~w --> regular(w) --> (u|>w) ~ (v|>w)";
 by (etac Scomp.induct 1);
 by (ALLGOALS (asm_simp_tac res1L_ss));
@@ -175,7 +175,7 @@
 qed_spec_mp "residuals_preserve_comp";
 
 
-goal Residuals.thy 
+Goal 
     "!!u. u~v ==> regular(u) --> regular(v) --> regular(u|>v)";
 by (etac Scomp.induct 1);
 by Safe_tac;
@@ -187,13 +187,13 @@
 (*     Preservation lemma                                                    *)
 (* ------------------------------------------------------------------------- *)
 
-goal Residuals.thy 
+Goal 
     "!!u. u~v ==> v ~ (u un v)";
 by (etac Scomp.induct 1);
 by (ALLGOALS Asm_full_simp_tac);
 qed "union_preserve_comp";
 
-goal Residuals.thy 
+Goal 
     "!!u. u ~ v ==> regular(v) --> u|>v = (u un v)|>v";
 by (etac Scomp.induct 1);
 by Safe_tac;
--- a/src/ZF/Resid/SubUnion.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/Resid/SubUnion.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -21,19 +21,19 @@
 (*    Equality rules for union                                               *)
 (* ------------------------------------------------------------------------- *)
 
-goalw SubUnion.thy [union_def]
+Goalw [union_def]
     "!!u. n:nat==>Var(n) un Var(n)=Var(n)";
 by (Asm_simp_tac 1);
 by (simp_tac (rank_ss addsimps redexes.con_defs)  1);
 qed "union_Var";
 
-goalw SubUnion.thy [union_def]
+Goalw [union_def]
     "!!u.[|u:redexes; v:redexes|]==>Fun(u) un Fun(v)=Fun(u un v)";
 by (Asm_simp_tac 1);
 by (simp_tac (rank_ss addsimps redexes.con_defs)  1);
 qed "union_Fun";
  
-goalw SubUnion.thy [union_def]
+Goalw [union_def]
  "!!u.[|b1:bool; b2:bool; u1:redexes; v1:redexes; u2:redexes; v2:redexes|]==> \
 \     App(b1,u1,v1) un App(b2,u2,v2)=App(b1 or b2,u1 un u2,v1 un v2)";
 by (Asm_simp_tac 1);
@@ -65,24 +65,24 @@
 (*    comp proofs                                                            *)
 (* ------------------------------------------------------------------------- *)
 
-goal SubUnion.thy "!!u. u:redexes ==> u ~ u";
+Goal "!!u. u:redexes ==> u ~ u";
 by (eresolve_tac [redexes.induct] 1);
 by (ALLGOALS Fast_tac);
 qed "comp_refl";
 
-goal SubUnion.thy 
+Goal 
     "!!u. u ~ v ==> v ~ u";
 by (etac Scomp.induct 1);
 by (ALLGOALS Fast_tac);
 qed "comp_sym";
 
-goal SubUnion.thy 
+Goal 
     "u ~ v <-> v ~ u";
 by (fast_tac (claset() addIs [comp_sym]) 1);
 qed "comp_sym_iff";
 
 
-goal SubUnion.thy 
+Goal 
     "!!u. u ~ v ==> ALL w. v ~ w-->u ~ w";
 by (etac Scomp.induct 1);
 by (ALLGOALS Fast_tac);
@@ -92,21 +92,21 @@
 (*   union proofs                                                            *)
 (* ------------------------------------------------------------------------- *)
 
-goal SubUnion.thy 
+Goal 
     "!!u. u ~ v ==> u <== (u un v)";
 by (etac Scomp.induct 1);
 by (etac boolE 3);
 by (ALLGOALS Asm_full_simp_tac);
 qed "union_l";
 
-goal SubUnion.thy 
+Goal 
     "!!u. u ~ v ==> v <== (u un v)";
 by (etac Scomp.induct 1);
 by (eres_inst_tac [("c","b2")] boolE 3);
 by (ALLGOALS Asm_full_simp_tac);
 qed "union_r";
 
-goal SubUnion.thy 
+Goal 
     "!!u. u ~ v ==> u un v = v un u";
 by (etac Scomp.induct 1);
 by (ALLGOALS(asm_simp_tac (simpset() addsimps [or_commute])));
@@ -116,7 +116,7 @@
 (*      regular proofs                                                       *)
 (* ------------------------------------------------------------------------- *)
 
-goal SubUnion.thy 
+Goal 
     "!!u. u ~ v ==> regular(u)-->regular(v)-->regular(u un v)";
 by (etac Scomp.induct 1);
 by (ALLGOALS(asm_full_simp_tac
--- a/src/ZF/Resid/Substitution.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/Resid/Substitution.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -42,28 +42,28 @@
 (* ------------------------------------------------------------------------- *)
 (*     lift_rec equality rules                                               *)
 (* ------------------------------------------------------------------------- *)
-goalw Substitution.thy [lift_rec_def] 
+Goalw [lift_rec_def] 
     "!!n.[|n:nat; i:nat|]==> lift_rec(Var(i),n) =if(i<n,Var(i),Var(succ(i)))";
 by (Asm_full_simp_tac 1);
 qed "lift_rec_Var";
 
-goalw Substitution.thy [lift_rec_def] 
+Goalw [lift_rec_def] 
     "!!n.[|n:nat; i:nat; k:nat; k le i|]==> lift_rec(Var(i),k) = Var(succ(i))";
 by (Asm_full_simp_tac 1);
 qed "lift_rec_le";
 
-goalw Substitution.thy [lift_rec_def] 
+Goalw [lift_rec_def] 
     "!!n.[|i:nat; k:nat; i<k |]==> lift_rec(Var(i),k) = Var(i)";
 by (Asm_full_simp_tac 1);
 qed "lift_rec_gt";
 
-goalw Substitution.thy [lift_rec_def] 
+Goalw [lift_rec_def] 
     "!!n.[|n:nat; k:nat|]==>   \
 \        lift_rec(Fun(t),k) = Fun(lift_rec(t,succ(k)))";
 by (Asm_full_simp_tac 1);
 qed "lift_rec_Fun";
 
-goalw Substitution.thy [lift_rec_def] 
+Goalw [lift_rec_def] 
     "!!n.[|n:nat; k:nat|]==>   \
 \        lift_rec(App(b,f,a),k) = App(b,lift_rec(f,k),lift_rec(a,k))";
 by (Asm_full_simp_tac 1);
@@ -73,36 +73,36 @@
 (*    substitution quality rules                                             *)
 (* ------------------------------------------------------------------------- *)
 
-goalw Substitution.thy [subst_rec_def] 
+Goalw [subst_rec_def] 
     "!!n.[|i:nat; k:nat; u:redexes|]==>  \
 \        subst_rec(u,Var(i),k) = if(k<i,Var(i#-1),if(k=i,u,Var(i)))";
 by (asm_full_simp_tac (simpset() addsimps [gt_not_eq,leI]) 1);
 qed "subst_Var";
 
-goalw Substitution.thy [subst_rec_def] 
+Goalw [subst_rec_def] 
     "!!n.[|n:nat; u:redexes|]==> subst_rec(u,Var(n),n) = u";
 by (asm_full_simp_tac (simpset()) 1);
 qed "subst_eq";
 
-goalw Substitution.thy [subst_rec_def] 
+Goalw [subst_rec_def] 
     "!!n.[|n:nat; u:redexes; p:nat; p<n|]==>  \
 \        subst_rec(u,Var(n),p) = Var(n#-1)";
 by (asm_full_simp_tac (simpset()) 1);
 qed "subst_gt";
 
-goalw Substitution.thy [subst_rec_def] 
+Goalw [subst_rec_def] 
     "!!n.[|n:nat; u:redexes; p:nat; n<p|]==>  \
 \        subst_rec(u,Var(n),p) = Var(n)";
 by (asm_full_simp_tac (simpset() addsimps [gt_not_eq,leI]) 1);
 qed "subst_lt";
 
-goalw Substitution.thy [subst_rec_def] 
+Goalw [subst_rec_def] 
     "!!n.[|p:nat; u:redexes|]==>  \
 \        subst_rec(u,Fun(t),p) = Fun(subst_rec(lift(u),t,succ(p))) ";
 by (asm_full_simp_tac (simpset()) 1);
 qed "subst_Fun";
 
-goalw Substitution.thy [subst_rec_def] 
+Goalw [subst_rec_def] 
     "!!n.[|p:nat; u:redexes|]==>  \
 \        subst_rec(u,App(b,f,a),p) = App(b,subst_rec(u,f,p),subst_rec(u,a,p))";
 by (asm_full_simp_tac (simpset()) 1);
@@ -112,7 +112,7 @@
                 addsimps [lift_rec_Var,subst_Var]);
 
 
-goal Substitution.thy  
+Goal  
     "!!n. u:redexes ==> ALL k:nat. lift_rec(u,k):redexes";
 by (eresolve_tac [redexes.induct] 1);
 by (ALLGOALS(asm_full_simp_tac 
@@ -120,7 +120,7 @@
 qed "lift_rec_type_a";
 val lift_rec_type = lift_rec_type_a RS bspec;
 
-goalw Substitution.thy [] 
+Goalw [] 
     "!!n. v:redexes ==>  ALL n:nat. ALL u:redexes. subst_rec(u,v,n):redexes";
 by (eresolve_tac [redexes.induct] 1);
 by (ALLGOALS(asm_full_simp_tac 
@@ -139,7 +139,7 @@
 (*    lift and  substitution proofs                                          *)
 (* ------------------------------------------------------------------------- *)
 
-goalw Substitution.thy [] 
+Goalw [] 
     "!!n. u:redexes ==> ALL n:nat. ALL i:nat. i le n -->   \
 \       (lift_rec(lift_rec(u,i),succ(n)) = lift_rec(lift_rec(u,n),i))";
 by ((eresolve_tac [redexes.induct] 1)
@@ -152,13 +152,13 @@
 qed "lift_lift_rec";
 
 
-goalw Substitution.thy [] 
+Goalw [] 
     "!!n.[|u:redexes; n:nat|]==>  \
 \      lift_rec(lift(u),succ(n)) = lift(lift_rec(u,n))";
 by (asm_simp_tac (simpset() addsimps [lift_lift_rec]) 1);
 qed "lift_lift";
 
-goal Substitution.thy 
+Goal 
     "!!n. v:redexes ==>  \
 \      ALL n:nat. ALL m:nat. ALL u:redexes. n le m-->\
 \         lift_rec(subst_rec(u,v,n),m) = \
@@ -178,14 +178,14 @@
 by (asm_full_simp_tac (simpset() addsimps [leI]) 1);
 qed "lift_rec_subst_rec";
 
-goalw Substitution.thy [] 
+Goalw [] 
     "!!n.[|v:redexes; u:redexes; n:nat|]==>  \
 \        lift_rec(u/v,n) = lift_rec(u,n)/lift_rec(v,succ(n))";
 by (asm_full_simp_tac (simpset() addsimps [lift_rec_subst_rec]) 1);
 qed "lift_subst";
 
 
-goalw Substitution.thy [] 
+Goalw [] 
     "!!n. v:redexes ==>  \
 \      ALL n:nat. ALL m:nat. ALL u:redexes. m le n-->\
 \         lift_rec(subst_rec(u,v,n),m) = \
@@ -206,7 +206,7 @@
 qed "lift_rec_subst_rec_lt";
 
 
-goalw Substitution.thy [] 
+Goalw [] 
     "!!n. u:redexes ==>  \
 \       ALL n:nat. ALL v:redexes. subst_rec(v,lift_rec(u,n),n) =  u";
 by ((eresolve_tac [redexes.induct] 1)
@@ -218,7 +218,7 @@
 by (asm_full_simp_tac (simpset()) 1);
 qed "subst_rec_lift_rec";
 
-goal Substitution.thy  
+Goal  
     "!!n. v:redexes ==>  \
 \       ALL m:nat. ALL n:nat. ALL u:redexes. ALL w:redexes. m le  n --> \
 \    subst_rec(subst_rec(w,u,n),subst_rec(lift_rec(w,m),v,succ(n)),m)=\
@@ -249,7 +249,7 @@
 qed "subst_rec_subst_rec";
 
 
-goalw Substitution.thy [] 
+Goalw [] 
     "!!n.[|v:redexes; u:redexes; w:redexes; n:nat|]==>  \
 \       subst_rec(w,u,n)/subst_rec(lift(w),v,succ(n)) = subst_rec(w,u/v,n)";
 by (asm_simp_tac (simpset() addsimps [subst_rec_subst_rec]) 1);
@@ -261,13 +261,13 @@
 (* ------------------------------------------------------------------------- *)
 
 
-goal Substitution.thy
+Goal
     "!!n.[|n:nat; u ~ v|]==> ALL m:nat. lift_rec(u,m) ~ lift_rec(v,m)";
 by (etac Scomp.induct 1);
 by (ALLGOALS(asm_simp_tac (simpset() addsimps [comp_refl])));
 qed "lift_rec_preserve_comp";
 
-goal Substitution.thy
+Goal
     "!!n. u2 ~ v2 ==> ALL m:nat. ALL u1:redexes. ALL v1:redexes.\
 \            u1 ~ v1--> subst_rec(u1,u2,m) ~ subst_rec(v1,v2,m)";
 by (etac Scomp.induct 1);
@@ -275,13 +275,13 @@
             ([lift_rec_preserve_comp,comp_refl]))));
 qed "subst_rec_preserve_comp";
 
-goal Substitution.thy
+Goal
     "!!n. regular(u) ==> ALL m:nat. regular(lift_rec(u,m))";
 by (eresolve_tac [Sreg.induct] 1);
 by (ALLGOALS(asm_full_simp_tac (addsplit (simpset()))));
 qed "lift_rec_preserve_reg";
 
-goal Substitution.thy
+Goal
     "!!n. regular(v) ==>  \
 \       ALL m:nat. ALL u:redexes. regular(u)-->regular(subst_rec(u,v,m))";
 by (eresolve_tac [Sreg.induct] 1);
--- a/src/ZF/Resid/Terms.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/Resid/Terms.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -11,17 +11,17 @@
 (*       unmark simplifications                                              *)
 (* ------------------------------------------------------------------------- *)
 
-goalw Terms.thy [unmark_def] 
+Goalw [unmark_def] 
     "unmark(Var(n)) = Var(n)";
 by (Asm_simp_tac 1);
 qed "unmark_Var";
 
-goalw Terms.thy [unmark_def] 
+Goalw [unmark_def] 
     "unmark(Fun(t)) = Fun(unmark(t))";
 by (Asm_simp_tac 1);
 qed "unmark_Fun";
 
-goalw Terms.thy [unmark_def] 
+Goalw [unmark_def] 
     "unmark(App(b,n,m)) = App(0,unmark(n),unmark(m))";
 by (Asm_simp_tac 1);
 qed "unmark_App";
@@ -41,13 +41,13 @@
 (*        unmark lemmas                                                      *)
 (* ------------------------------------------------------------------------- *)
 
-goalw Terms.thy [unmark_def] 
+Goalw [unmark_def] 
     "!!u. u:redexes ==> unmark(u):lambda";
 by (eresolve_tac [redexes.induct] 1);
 by (ALLGOALS Asm_simp_tac);
 qed "unmark_type";
 
-goal Terms.thy  
+Goal  
     "!!u. u:lambda ==> unmark(u) = u";
 by (eresolve_tac [lambda.induct] 1);
 by (ALLGOALS Asm_simp_tac);
@@ -58,14 +58,14 @@
 (*         lift and subst preserve lambda                                    *)
 (* ------------------------------------------------------------------------- *)
 
-goal Terms.thy  
+Goal  
     "!!u.[|v:lambda|]==> ALL k:nat. lift_rec(v,k):lambda";
 by (eresolve_tac [lambda.induct] 1);
 by (ALLGOALS(asm_full_simp_tac (addsplit (simpset()))));
 qed "liftL_typea";
 val liftL_type =liftL_typea RS bspec ;
 
-goal Terms.thy  
+Goal  
     "!!n.[|v:lambda|]==>  ALL n:nat. ALL u:lambda. subst_rec(u,v,n):lambda";
 by (eresolve_tac [lambda.induct] 1);
 by (ALLGOALS(asm_full_simp_tac ((addsplit (simpset())) addsimps [liftL_type])));
--- a/src/ZF/ex/Acc.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/ex/Acc.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -20,7 +20,7 @@
 by (fast_tac (claset() addIs (prems@acc.intrs)) 1);
 qed "accI";
 
-goal Acc.thy "!!a b r. [| b: acc(r);  <a,b>: r |] ==> a: acc(r)";
+Goal "!!a b r. [| b: acc(r);  <a,b>: r |] ==> a: acc(r)";
 by (etac acc.elim 1);
 by (Fast_tac 1);
 qed "acc_downward";
@@ -37,7 +37,7 @@
 by (etac (Collect_subset RS Pow_mono RS subsetD) 1);
 qed "acc_induct";
 
-goal Acc.thy "wf[acc(r)](r)";
+Goal "wf[acc(r)](r)";
 by (rtac wf_onI2 1);
 by (etac acc_induct 1);
 by (Fast_tac 1);
@@ -55,6 +55,6 @@
 by (Fast_tac 1);
 qed "acc_wfD";
 
-goal Acc.thy "wf(r) <-> field(r) <= acc(r)";
+Goal "wf(r) <-> field(r) <= acc(r)";
 by (EVERY1 [rtac iffI, etac acc_wfD, etac acc_wfI]);
 qed "wf_acc_iff";
--- a/src/ZF/ex/BT.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/ex/BT.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -19,13 +19,13 @@
 
 (**  Lemmas to justify using "bt" in other recursive type definitions **)
 
-goalw BT.thy bt.defs "!!A B. A<=B ==> bt(A) <= bt(B)";
+Goalw bt.defs "!!A B. A<=B ==> bt(A) <= bt(B)";
 by (rtac lfp_mono 1);
 by (REPEAT (rtac bt.bnd_mono 1));
 by (REPEAT (ares_tac (univ_mono::basic_monos) 1));
 qed "bt_mono";
 
-goalw BT.thy (bt.defs@bt.con_defs) "bt(univ(A)) <= univ(A)";
+Goalw (bt.defs@bt.con_defs) "bt(univ(A)) <= univ(A)";
 by (rtac lfp_lowerbound 1);
 by (rtac (A_subset_univ RS univ_mono) 2);
 by (fast_tac (claset() addSIs [zero_in_univ, Inl_in_univ, Inr_in_univ,
@@ -37,20 +37,20 @@
 
 (** bt_rec -- by Vset recursion **)
 
-goalw BT.thy bt.con_defs "rank(l) < rank(Br(a,l,r))";
+Goalw bt.con_defs "rank(l) < rank(Br(a,l,r))";
 by (simp_tac rank_ss 1);
 qed "rank_Br1";
 
-goalw BT.thy bt.con_defs "rank(r) < rank(Br(a,l,r))";
+Goalw bt.con_defs "rank(r) < rank(Br(a,l,r))";
 by (simp_tac rank_ss 1);
 qed "rank_Br2";
 
-goal BT.thy "bt_rec(Lf,c,h) = c";
+Goal "bt_rec(Lf,c,h) = c";
 by (rtac (bt_rec_def RS def_Vrec RS trans) 1);
 by (Simp_tac 1);
 qed "bt_rec_Lf";
 
-goal BT.thy
+Goal
     "bt_rec(Br(a,l,r), c, h) = h(a, l, r, bt_rec(l,c,h), bt_rec(r,c,h))";
 by (rtac (bt_rec_def RS def_Vrec RS trans) 1);
 by (simp_tac (rank_ss addsimps (bt.case_eqns @ [rank_Br1, rank_Br2])) 1);
@@ -110,7 +110,7 @@
 val [bt_reflect_Lf, bt_reflect_Br] = bt_recs bt_reflect_def;
 Addsimps [bt_reflect_Lf, bt_reflect_Br];
 
-goalw BT.thy [bt_reflect_def] "!!xs. xs: bt(A) ==> bt_reflect(xs) : bt(A)";
+Goalw [bt_reflect_def] "!!xs. xs: bt(A) ==> bt_reflect(xs) : bt(A)";
 by (REPEAT (ares_tac (bt.intrs @ [bt_rec_type]) 1));
 qed "bt_reflect_type";
 
--- a/src/ZF/ex/Bin.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/ex/Bin.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -19,19 +19,19 @@
 
 (** bin_rec -- by Vset recursion **)
 
-goal Bin.thy "bin_rec(Plus,a,b,h) = a";
+Goal "bin_rec(Plus,a,b,h) = a";
 by (rtac (bin_rec_def RS def_Vrec RS trans) 1);
 by (rewrite_goals_tac bin.con_defs);
 by (simp_tac rank_ss 1);
 qed "bin_rec_Plus";
 
-goal Bin.thy "bin_rec(Minus,a,b,h) = b";
+Goal "bin_rec(Minus,a,b,h) = b";
 by (rtac (bin_rec_def RS def_Vrec RS trans) 1);
 by (rewrite_goals_tac bin.con_defs);
 by (simp_tac rank_ss 1);
 qed "bin_rec_Minus";
 
-goal Bin.thy "bin_rec(Bcons(w,x),a,b,h) = h(w, x, bin_rec(w,a,b,h))";
+Goal "bin_rec(Bcons(w,x),a,b,h) = h(w, x, bin_rec(w,a,b,h))";
 by (rtac (bin_rec_def RS def_Vrec RS trans) 1);
 by (rewrite_goals_tac bin.con_defs);
 by (simp_tac rank_ss 1);
@@ -72,23 +72,23 @@
 fun bin_recs def = map standard
         ([def] RL [def_bin_rec_Plus, def_bin_rec_Minus, def_bin_rec_Bcons]);
 
-goalw Bin.thy [norm_Bcons_def] "norm_Bcons(Plus,0) = Plus";
+Goalw [norm_Bcons_def] "norm_Bcons(Plus,0) = Plus";
 by (Asm_simp_tac 1);
 qed "norm_Bcons_Plus_0";
 
-goalw Bin.thy [norm_Bcons_def] "norm_Bcons(Plus,1) = Bcons(Plus,1)";
+Goalw [norm_Bcons_def] "norm_Bcons(Plus,1) = Bcons(Plus,1)";
 by (Asm_simp_tac 1);
 qed "norm_Bcons_Plus_1";
 
-goalw Bin.thy [norm_Bcons_def] "norm_Bcons(Minus,0) = Bcons(Minus,0)";
+Goalw [norm_Bcons_def] "norm_Bcons(Minus,0) = Bcons(Minus,0)";
 by (Asm_simp_tac 1);
 qed "norm_Bcons_Minus_0";
 
-goalw Bin.thy [norm_Bcons_def] "norm_Bcons(Minus,1) = Minus";
+Goalw [norm_Bcons_def] "norm_Bcons(Minus,1) = Minus";
 by (Asm_simp_tac 1);
 qed "norm_Bcons_Minus_1";
 
-goalw Bin.thy [norm_Bcons_def]
+Goalw [norm_Bcons_def]
     "norm_Bcons(Bcons(w,x),b) = Bcons(Bcons(w,x),b)";
 by (asm_simp_tac (simpset() addsimps bin.case_eqns) 1);
 qed "norm_Bcons_Bcons";
@@ -101,41 +101,41 @@
 
 val bin_typechecks0 = bin_rec_type :: bin.intrs;
 
-goalw Bin.thy [integ_of_bin_def]
+Goalw [integ_of_bin_def]
     "!!w. w: bin ==> integ_of_bin(w) : integ";
 by (typechk_tac (bin_typechecks0@integ_typechecks@
                  nat_typechecks@[bool_into_nat]));
 qed "integ_of_bin_type";
 
-goalw Bin.thy [norm_Bcons_def]
+Goalw [norm_Bcons_def]
     "!!w. [| w: bin; b: bool |] ==> norm_Bcons(w,b) : bin";
 by (etac bin.elim 1);
 by (ALLGOALS (asm_simp_tac (simpset() addsimps bin.case_eqns)));
 by (typechk_tac (bin_typechecks0@bool_typechecks));
 qed "norm_Bcons_type";
 
-goalw Bin.thy [bin_succ_def]
+Goalw [bin_succ_def]
     "!!w. w: bin ==> bin_succ(w) : bin";
 by (typechk_tac ([norm_Bcons_type]@bin_typechecks0@bool_typechecks));
 qed "bin_succ_type";
 
-goalw Bin.thy [bin_pred_def]
+Goalw [bin_pred_def]
     "!!w. w: bin ==> bin_pred(w) : bin";
 by (typechk_tac ([norm_Bcons_type]@bin_typechecks0@bool_typechecks));
 qed "bin_pred_type";
 
-goalw Bin.thy [bin_minus_def]
+Goalw [bin_minus_def]
     "!!w. w: bin ==> bin_minus(w) : bin";
 by (typechk_tac ([bin_pred_type]@bin_typechecks0@bool_typechecks));
 qed "bin_minus_type";
 
-goalw Bin.thy [bin_add_def]
+Goalw [bin_add_def]
     "!!v w. [| v: bin; w: bin |] ==> bin_add(v,w) : bin";
 by (typechk_tac ([norm_Bcons_type, bin_succ_type, bin_pred_type]@
                  bin_typechecks0@ bool_typechecks@ZF_typechecks));
 qed "bin_add_type";
 
-goalw Bin.thy [bin_mult_def]
+Goalw [bin_mult_def]
     "!!v w. [| v: bin; w: bin |] ==> bin_mult(v,w) : bin";
 by (typechk_tac ([norm_Bcons_type, bin_minus_type, bin_add_type]@
                  bin_typechecks0@ bool_typechecks));
@@ -177,7 +177,7 @@
 
 
 (*norm_Bcons preserves the integer value of its argument*)
-goal Bin.thy
+Goal
     "!!w. [| w: bin; b: bool |] ==>     \
 \         integ_of_bin(norm_Bcons(w,b)) = integ_of_bin(Bcons(w,b))";
 by (etac bin.elim 1);
@@ -186,7 +186,7 @@
 by (ALLGOALS (asm_simp_tac (simpset() addsimps (norm_Bcons_simps))));
 qed "integ_of_bin_norm_Bcons";
 
-goal Bin.thy
+Goal
     "!!w. w: bin ==> integ_of_bin(bin_succ(w)) = $#1 $+ integ_of_bin(w)";
 by (etac bin.induct 1);
 by (Simp_tac 1);
@@ -196,7 +196,7 @@
     (asm_simp_tac (simpset() addsimps integ_of_bin_norm_Bcons::zadd_ac)));
 qed "integ_of_bin_succ";
 
-goal Bin.thy
+Goal
     "!!w. w: bin ==> integ_of_bin(bin_pred(w)) = $~ ($#1) $+ integ_of_bin(w)";
 by (etac bin.induct 1);
 by (Simp_tac 1);
@@ -214,7 +214,7 @@
 Addsimps (bin_recs bin_minus_def @
 	  [integ_of_bin_succ, integ_of_bin_pred]);
 
-goal Bin.thy
+Goal
     "!!w. w: bin ==> integ_of_bin(bin_minus(w)) = $~ integ_of_bin(w)";
 by (etac bin.induct 1);
 by (Simp_tac 1);
@@ -227,23 +227,23 @@
 
 (*** bin_add: binary addition ***)
 
-goalw Bin.thy [bin_add_def] "!!w. w: bin ==> bin_add(Plus,w) = w";
+Goalw [bin_add_def] "!!w. w: bin ==> bin_add(Plus,w) = w";
 by (Asm_simp_tac 1);
 qed "bin_add_Plus";
 
-goalw Bin.thy [bin_add_def] "!!w. w: bin ==> bin_add(Minus,w) = bin_pred(w)";
+Goalw [bin_add_def] "!!w. w: bin ==> bin_add(Minus,w) = bin_pred(w)";
 by (Asm_simp_tac 1);
 qed "bin_add_Minus";
 
-goalw Bin.thy [bin_add_def] "bin_add(Bcons(v,x),Plus) = Bcons(v,x)";
+Goalw [bin_add_def] "bin_add(Bcons(v,x),Plus) = Bcons(v,x)";
 by (Simp_tac 1);
 qed "bin_add_Bcons_Plus";
 
-goalw Bin.thy [bin_add_def] "bin_add(Bcons(v,x),Minus) = bin_pred(Bcons(v,x))";
+Goalw [bin_add_def] "bin_add(Bcons(v,x),Minus) = bin_pred(Bcons(v,x))";
 by (Simp_tac 1);
 qed "bin_add_Bcons_Minus";
 
-goalw Bin.thy [bin_add_def]
+Goalw [bin_add_def]
     "!!w y. [| w: bin;  y: bool |] ==> \
 \           bin_add(Bcons(v,x), Bcons(w,y)) = \
 \           norm_Bcons(bin_add(v, cond(x and y, bin_succ(w), w)), x xor y)";
@@ -257,7 +257,7 @@
 
 Addsimps [bool_subset_nat RS subsetD];
 
-goal Bin.thy
+Goal
     "!!v. v: bin ==> \
 \         ALL w: bin. integ_of_bin(bin_add(v,w)) = \
 \                     integ_of_bin(v) $+ integ_of_bin(w)";
@@ -299,19 +299,19 @@
 val [bin_succ_Plus, bin_succ_Minus, _] = bin_recs bin_succ_def;
 val [bin_pred_Plus, bin_pred_Minus, _] = bin_recs bin_pred_def;
 
-goal Bin.thy "bin_succ(Bcons(w,1)) = Bcons(bin_succ(w), 0)";
+Goal "bin_succ(Bcons(w,1)) = Bcons(bin_succ(w), 0)";
 by (Simp_tac 1);
 qed "bin_succ_Bcons1";
 
-goal Bin.thy "bin_succ(Bcons(w,0)) = norm_Bcons(w,1)";
+Goal "bin_succ(Bcons(w,0)) = norm_Bcons(w,1)";
 by (Simp_tac 1);
 qed "bin_succ_Bcons0";
 
-goal Bin.thy "bin_pred(Bcons(w,1)) = norm_Bcons(w,0)";
+Goal "bin_pred(Bcons(w,1)) = norm_Bcons(w,0)";
 by (Simp_tac 1);
 qed "bin_pred_Bcons1";
 
-goal Bin.thy "bin_pred(Bcons(w,0)) = Bcons(bin_pred(w), 1)";
+Goal "bin_pred(Bcons(w,0)) = Bcons(bin_pred(w), 1)";
 by (Simp_tac 1);
 qed "bin_pred_Bcons0";
 
@@ -319,29 +319,29 @@
 
 val [bin_minus_Plus, bin_minus_Minus, _] = bin_recs bin_minus_def;
 
-goal Bin.thy "bin_minus(Bcons(w,1)) = bin_pred(Bcons(bin_minus(w), 0))";
+Goal "bin_minus(Bcons(w,1)) = bin_pred(Bcons(bin_minus(w), 0))";
 by (Simp_tac 1);
 qed "bin_minus_Bcons1";
 
-goal Bin.thy "bin_minus(Bcons(w,0)) = Bcons(bin_minus(w), 0)";
+Goal "bin_minus(Bcons(w,0)) = Bcons(bin_minus(w), 0)";
 by (Simp_tac 1);
 qed "bin_minus_Bcons0";
 
 (** extra rules for bin_add **)
 
-goal Bin.thy 
+Goal 
     "!!w. w: bin ==> bin_add(Bcons(v,1), Bcons(w,1)) = \
 \                    norm_Bcons(bin_add(v, bin_succ(w)), 0)";
 by (Asm_simp_tac 1);
 qed "bin_add_Bcons_Bcons11";
 
-goal Bin.thy 
+Goal 
     "!!w. w: bin ==> bin_add(Bcons(v,1), Bcons(w,0)) =  \
 \                    norm_Bcons(bin_add(v,w), 1)";
 by (Asm_simp_tac 1);
 qed "bin_add_Bcons_Bcons10";
 
-goal Bin.thy 
+Goal 
     "!!w y. [| w: bin;  y: bool |] ==> \
 \           bin_add(Bcons(v,0), Bcons(w,y)) = norm_Bcons(bin_add(v,w), y)";
 by (Asm_simp_tac 1);
@@ -351,12 +351,12 @@
 
 val [bin_mult_Plus, bin_mult_Minus, _] = bin_recs bin_mult_def;
 
-goal Bin.thy
+Goal
     "bin_mult(Bcons(v,1), w) = bin_add(norm_Bcons(bin_mult(v,w),0), w)";
 by (Simp_tac 1);
 qed "bin_mult_Bcons1";
 
-goal Bin.thy "bin_mult(Bcons(v,0), w) = norm_Bcons(bin_mult(v,w),0)";
+Goal "bin_mult(Bcons(v,0), w) = norm_Bcons(bin_mult(v,w),0)";
 by (Simp_tac 1);
 qed "bin_mult_Bcons0";
 
@@ -387,63 +387,63 @@
 set proof_timing;
 (*All runtimes below are on a SPARCserver 10*)
 
-goal Bin.thy "#13  $+  #19 = #32";
+Goal "#13  $+  #19 = #32";
 by (simp_tac bin_comp_ss 1);    (*0.4 secs*)
 result();
 
 bin_add(binary_of_int 13, binary_of_int 19);
 
-goal Bin.thy "#1234  $+  #5678 = #6912";
+Goal "#1234  $+  #5678 = #6912";
 by (simp_tac bin_comp_ss 1);    (*1.3 secs*)
 result();
 
 bin_add(binary_of_int 1234, binary_of_int 5678);
 
-goal Bin.thy "#1359  $+  #~2468 = #~1109";
+Goal "#1359  $+  #~2468 = #~1109";
 by (simp_tac bin_comp_ss 1);    (*1.2 secs*)
 result();
 
 bin_add(binary_of_int 1359, binary_of_int ~2468);
 
-goal Bin.thy "#93746  $+  #~46375 = #47371";
+Goal "#93746  $+  #~46375 = #47371";
 by (simp_tac bin_comp_ss 1);    (*1.9 secs*)
 result();
 
 bin_add(binary_of_int 93746, binary_of_int ~46375);
 
-goal Bin.thy "$~ #65745 = #~65745";
+Goal "$~ #65745 = #~65745";
 by (simp_tac bin_comp_ss 1);    (*0.4 secs*)
 result();
 
 bin_minus(binary_of_int 65745);
 
 (* negation of ~54321 *)
-goal Bin.thy "$~ #~54321 = #54321";
+Goal "$~ #~54321 = #54321";
 by (simp_tac bin_comp_ss 1);    (*0.5 secs*)
 result();
 
 bin_minus(binary_of_int ~54321);
 
-goal Bin.thy "#13  $*  #19 = #247";
+Goal "#13  $*  #19 = #247";
 by (simp_tac bin_comp_ss 1);    (*0.7 secs*)
 result();
 
 bin_mult(binary_of_int 13, binary_of_int 19);
 
-goal Bin.thy "#~84  $*  #51 = #~4284";
+Goal "#~84  $*  #51 = #~4284";
 by (simp_tac bin_comp_ss 1);    (*1.3 secs*)
 result();
 
 bin_mult(binary_of_int ~84, binary_of_int 51);
 
 (*The worst case for 8-bit operands *)
-goal Bin.thy "#255  $*  #255 = #65025";
+Goal "#255  $*  #255 = #65025";
 by (simp_tac bin_comp_ss 1);    (*4.3 secs*)
 result();
 
 bin_mult(binary_of_int 255, binary_of_int 255);
 
-goal Bin.thy "#1359  $*  #~2468 = #~3354012";
+Goal "#1359  $*  #~2468 = #~3354012";
 by (simp_tac bin_comp_ss 1);    (*6.1 secs*)
 result();
 
--- a/src/ZF/ex/Brouwer.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/ex/Brouwer.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -12,7 +12,7 @@
 
 (** The Brouwer ordinals **)
 
-goal Brouwer.thy "brouwer = {0} + brouwer + (nat -> brouwer)";
+Goal "brouwer = {0} + brouwer + (nat -> brouwer)";
 let open brouwer;  val rew = rewrite_rule con_defs in  
 by (fast_tac (claset() addSIs (map rew intrs) addEs [rew elim]) 1)
 end;
@@ -35,7 +35,7 @@
 
 (** The Martin-Löf wellordering type **)
 
-goal Brouwer.thy "Well(A,B) = (SUM x:A. B(x) -> Well(A,B))";
+Goal "Well(A,B) = (SUM x:A. B(x) -> Well(A,B))";
 let open Well;  val rew = rewrite_rule con_defs in  
 by (fast_tac (claset() addSIs (map rew intrs) addEs [rew elim]) 1)
 end;
@@ -56,7 +56,7 @@
 
 (*In fact it's isomorphic to nat, but we need a recursion operator for
   Well to prove this.*)
-goal Brouwer.thy "Well(bool, %x. x) = 1 + (1 -> Well(bool, %x. x))";
+Goal "Well(bool, %x. x) = 1 + (1 -> Well(bool, %x. x))";
 by (resolve_tac [Well_unfold RS trans] 1);
 by (simp_tac (simpset() addsimps [Sigma_bool, Pi_empty1, succ_def]) 1);
 qed "Well_bool_unfold";
--- a/src/ZF/ex/CoUnit.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/ex/CoUnit.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -19,7 +19,7 @@
 val Con_iff = counit.mk_free "Con(x)=Con(y) <-> x=y";
 
 (*Should be a singleton, not everything!*)
-goal CoUnit.thy "counit = quniv(0)";
+Goal "counit = quniv(0)";
 by (rtac (counit.dom_subset RS equalityI) 1);
 by (rtac subsetI 1);
 by (etac counit.coinduct 1);
@@ -38,12 +38,12 @@
 (*Proving freeness results*)
 val Con2_iff = counit2.mk_free "Con2(x,y)=Con2(x',y') <-> x=x' & y=y'";
 
-goalw CoUnit.thy counit2.con_defs "bnd_mono(univ(0), %x. Con2(x,x))";
+Goalw counit2.con_defs "bnd_mono(univ(0), %x. Con2(x,x))";
 by (rtac bnd_monoI 1);
 by (REPEAT (ares_tac [subset_refl, QPair_subset_univ, QPair_mono] 1));
 qed "Con2_bnd_mono";
 
-goal CoUnit.thy "lfp(univ(0), %x. Con2(x,x)) : counit2";
+Goal "lfp(univ(0), %x. Con2(x,x)) : counit2";
 by (rtac (singletonI RS counit2.coinduct) 1);
 by (rtac (qunivI RS singleton_subsetI) 1);
 by (rtac ([lfp_subset, empty_subsetI RS univ_mono] MRS subset_trans) 1);
@@ -51,7 +51,7 @@
 qed "lfp_Con2_in_counit2";
 
 (*Lemma for proving finality.  Borrowed from ex/llist_eq.ML!*)
-goal CoUnit.thy
+Goal
     "!!i. Ord(i) ==> ALL x y. x: counit2 & y: counit2 --> x Int Vset(i) <= y";
 by (etac trans_induct 1);
 by (safe_tac subset_cs);
@@ -66,12 +66,12 @@
 val counit2_Int_Vset_subset = standard
         (counit2_Int_Vset_subset_lemma RS spec RS spec RS mp);
 
-goal CoUnit.thy "!!x y. [| x: counit2;  y: counit2 |] ==> x=y";
+Goal "!!x y. [| x: counit2;  y: counit2 |] ==> x=y";
 by (rtac equalityI 1);
 by (REPEAT (ares_tac [conjI, counit2_Int_Vset_subset RS Int_Vset_subset] 1));
 qed "counit2_implies_equal";
 
-goal CoUnit.thy "counit2 = {lfp(univ(0), %x. Con2(x,x))}";
+Goal "counit2 = {lfp(univ(0), %x. Con2(x,x))}";
 by (rtac equalityI 1);
 by (rtac (lfp_Con2_in_counit2 RS singleton_subsetI) 2);
 by (rtac subsetI 1);
--- a/src/ZF/ex/Comb.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/ex/Comb.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -19,7 +19,7 @@
 
 (*** Transitive closure preserves the Church-Rosser property ***)
 
-goalw Comb.thy [diamond_def]
+Goalw [diamond_def]
     "!!x y r. [| diamond(r);  <x,y>:r^+ |] ==> \
 \    ALL y'. <x,y'>:r --> (EX z. <y',z>: r^+ & <y,z>: r)";
 by (etac trancl_induct 1);
@@ -52,7 +52,7 @@
 val contract_combD1 = contract.dom_subset RS subsetD RS SigmaD1;
 val contract_combD2 = contract.dom_subset RS subsetD RS SigmaD2;
 
-goal Comb.thy "field(contract) = comb";
+Goal "field(contract) = comb";
 by (blast_tac (claset() addIs [contract.K] addSEs [contract_combE2]) 1);
 qed "field_contract_eq";
 
@@ -70,11 +70,11 @@
                      contract.Ap2 RS rtrancl_into_rtrancl2];
 
 (*Example only: not used*)
-goalw Comb.thy [I_def] "!!p. p:comb ==> I#p ---> p";
+Goalw [I_def] "!!p. p:comb ==> I#p ---> p";
 by (blast_tac (claset() addIs reduction_rls) 1);
 qed "reduce_I";
 
-goalw Comb.thy [I_def] "I: comb";
+Goalw [I_def] "I: comb";
 by (Blast_tac 1);
 qed "comb_I";
 
@@ -88,15 +88,15 @@
 AddIs  [contract.Ap1, contract.Ap2];
 AddSEs [K_contractE, S_contractE, Ap_contractE];
 
-goalw Comb.thy [I_def] "!!r. I -1-> r ==> P";
+Goalw [I_def] "!!r. I -1-> r ==> P";
 by (Blast_tac 1);
 qed "I_contract_E";
 
-goal Comb.thy "!!p r. K#p -1-> r ==> (EX q. r = K#q & p -1-> q)";
+Goal "!!p r. K#p -1-> r ==> (EX q. r = K#q & p -1-> q)";
 by (Blast_tac 1);
 qed "K1_contractD";
 
-goal Comb.thy "!!p r. [| p ---> q;  r: comb |] ==> p#r ---> q#r";
+Goal "!!p r. [| p ---> q;  r: comb |] ==> p#r ---> q#r";
 by (forward_tac [rtrancl_type RS subsetD RS SigmaD1] 1);
 by (dtac (field_contract_eq RS equalityD1 RS subsetD) 1);
 by (etac rtrancl_induct 1);
@@ -105,7 +105,7 @@
 by (blast_tac (claset() addIs (contract_combD2::reduction_rls)) 1);
 qed "Ap_reduce1";
 
-goal Comb.thy "!!p r. [| p ---> q;  r: comb |] ==> r#p ---> r#q";
+Goal "!!p r. [| p ---> q;  r: comb |] ==> r#p ---> r#q";
 by (forward_tac [rtrancl_type RS subsetD RS SigmaD1] 1);
 by (dtac (field_contract_eq RS equalityD1 RS subsetD) 1);
 by (etac rtrancl_induct 1);
@@ -116,19 +116,19 @@
 
 (** Counterexample to the diamond property for -1-> **)
 
-goal Comb.thy "K#I#(I#I) -1-> I";
+Goal "K#I#(I#I) -1-> I";
 by (REPEAT (ares_tac [contract.K, comb_I, comb_Ap] 1));
 qed "KIII_contract1";
 
-goalw Comb.thy [I_def] "K#I#(I#I) -1-> K#I#((K#I)#(K#I))";
+Goalw [I_def] "K#I#(I#I) -1-> K#I#((K#I)#(K#I))";
 by (DEPTH_SOLVE (resolve_tac (comb.intrs @ contract.intrs) 1));
 qed "KIII_contract2";
 
-goal Comb.thy "K#I#((K#I)#(K#I)) -1-> I";
+Goal "K#I#((K#I)#(K#I)) -1-> I";
 by (REPEAT (ares_tac (comb.intrs @ [contract.K, comb_I]) 1));
 qed "KIII_contract3";
 
-goalw Comb.thy [diamond_def] "~ diamond(contract)";
+Goalw [diamond_def] "~ diamond(contract)";
 by (blast_tac (claset() addIs [KIII_contract1,KIII_contract2,KIII_contract3]
                        addSEs [I_contract_E]) 1);
 qed "not_diamond_contract";
@@ -142,7 +142,7 @@
 val parcontract_combD1 = parcontract.dom_subset RS subsetD RS SigmaD1;
 val parcontract_combD2 = parcontract.dom_subset RS subsetD RS SigmaD2;
 
-goal Comb.thy "field(parcontract) = comb";
+Goal "field(parcontract) = comb";
 by (blast_tac (claset() addIs [parcontract.K] 
                       addSEs [parcontract_combE2]) 1);
 qed "field_parcontract_eq";
@@ -157,21 +157,21 @@
 
 (*** Basic properties of parallel contraction ***)
 
-goal Comb.thy "!!p r. K#p =1=> r ==> (EX p'. r = K#p' & p =1=> p')";
+Goal "!!p r. K#p =1=> r ==> (EX p'. r = K#p' & p =1=> p')";
 by (Blast_tac 1);
 qed "K1_parcontractD";
 
-goal Comb.thy "!!p r. S#p =1=> r ==> (EX p'. r = S#p' & p =1=> p')";
+Goal "!!p r. S#p =1=> r ==> (EX p'. r = S#p' & p =1=> p')";
 by (Blast_tac 1);
 qed "S1_parcontractD";
 
-goal Comb.thy
+Goal
  "!!p q r. S#p#q =1=> r ==> (EX p' q'. r = S#p'#q' & p =1=> p' & q =1=> q')";
 by (blast_tac (claset() addSDs [S1_parcontractD]) 1);
 qed "S2_parcontractD";
 
 (*Church-Rosser property for parallel contraction*)
-goalw Comb.thy [diamond_def] "diamond(parcontract)";
+Goalw [diamond_def] "diamond(parcontract)";
 by (rtac (impI RS allI RS allI) 1);
 by (etac parcontract.induct 1);
 by (ALLGOALS 
@@ -181,12 +181,12 @@
 
 (*** Equivalence of p--->q and p===>q ***)
 
-goal Comb.thy "!!p q. p-1->q ==> p=1=>q";
+Goal "!!p q. p-1->q ==> p=1=>q";
 by (etac contract.induct 1);
 by (ALLGOALS Blast_tac);
 qed "contract_imp_parcontract";
 
-goal Comb.thy "!!p q. p--->q ==> p===>q";
+Goal "!!p q. p--->q ==> p===>q";
 by (forward_tac [rtrancl_type RS subsetD RS SigmaD1] 1);
 by (dtac (field_contract_eq RS equalityD1 RS subsetD) 1);
 by (etac rtrancl_induct 1);
@@ -196,7 +196,7 @@
 qed "reduce_imp_parreduce";
 
 
-goal Comb.thy "!!p q. p=1=>q ==> p--->q";
+Goal "!!p q. p=1=>q ==> p--->q";
 by (etac parcontract.induct 1);
 by (blast_tac (claset() addIs reduction_rls) 1);
 by (blast_tac (claset() addIs reduction_rls) 1);
@@ -207,7 +207,7 @@
 		    parcontract_combD2]) 1);
 qed "parcontract_imp_reduce";
 
-goal Comb.thy "!!p q. p===>q ==> p--->q";
+Goal "!!p q. p===>q ==> p--->q";
 by (forward_tac [trancl_type RS subsetD RS SigmaD1] 1);
 by (dtac (field_parcontract_eq RS equalityD1 RS subsetD) 1);
 by (etac trancl_induct 1);
@@ -216,7 +216,7 @@
 by (etac parcontract_imp_reduce 1);
 qed "parreduce_imp_reduce";
 
-goal Comb.thy "p===>q <-> p--->q";
+Goal "p===>q <-> p--->q";
 by (REPEAT (ares_tac [iffI, parreduce_imp_reduce, reduce_imp_parreduce] 1));
 qed "parreduce_iff_reduce";
 
--- a/src/ZF/ex/Data.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/ex/Data.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -9,7 +9,7 @@
 
 open Data;
 
-goal Data.thy "data(A,B) = ({0} + A) + (A*B + A*B*data(A,B))";
+Goal "data(A,B) = ({0} + A) + (A*B + A*B*data(A,B))";
 let open data;  val rew = rewrite_rule con_defs in  
 by (fast_tac (claset() addSIs (map rew intrs) addEs [rew elim]) 1)
 end;
@@ -17,13 +17,13 @@
 
 (**  Lemmas to justify using "data" in other recursive type definitions **)
 
-goalw Data.thy data.defs "!!A B. [| A<=C; B<=D |] ==> data(A,B) <= data(C,D)";
+Goalw data.defs "!!A B. [| A<=C; B<=D |] ==> data(A,B) <= data(C,D)";
 by (rtac lfp_mono 1);
 by (REPEAT (rtac data.bnd_mono 1));
 by (REPEAT (ares_tac (univ_mono::Un_mono::basic_monos) 1));
 qed "data_mono";
 
-goalw Data.thy (data.defs@data.con_defs) "data(univ(A),univ(A)) <= univ(A)";
+Goalw (data.defs@data.con_defs) "data(univ(A),univ(A)) <= univ(A)";
 by (rtac lfp_lowerbound 1);
 by (rtac ([A_subset_univ, Un_upper1] MRS subset_trans RS univ_mono) 2);
 by (fast_tac (claset() addSIs [zero_in_univ, Inl_in_univ, Inr_in_univ,
--- a/src/ZF/ex/Enum.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/ex/Enum.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -10,7 +10,7 @@
 
 open Enum;
 
-goal Enum.thy "C00 ~= C01";
+Goal "C00 ~= C01";
 by (simp_tac (simpset() addsimps enum.free_iffs) 1);
 result();
 
--- a/src/ZF/ex/Integ.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/ex/Integ.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -31,20 +31,20 @@
 
 (** Natural deduction for intrel **)
 
-goalw Integ.thy [intrel_def]
+Goalw [intrel_def]
     "<<x1,y1>,<x2,y2>>: intrel <-> \
 \    x1: nat & y1: nat & x2: nat & y2: nat & x1#+y2 = x2#+y1";
 by (Fast_tac 1);
 qed "intrel_iff";
 
-goalw Integ.thy [intrel_def]
+Goalw [intrel_def]
     "!!x1 x2. [| x1#+y2 = x2#+y1; x1: nat; y1: nat; x2: nat; y2: nat |] ==> \
 \             <<x1,y1>,<x2,y2>>: intrel";
 by (fast_tac (claset() addIs prems) 1);
 qed "intrelI";
 
 (*intrelE is hard to derive because fast_tac tries hyp_subst_tac so soon*)
-goalw Integ.thy [intrel_def]
+Goalw [intrel_def]
   "p: intrel --> (EX x1 y1 x2 y2. \
 \                  p = <<x1,y1>,<x2,y2>> & x1#+y2 = x2#+y1 & \
 \                  x1: nat & y1: nat & x2: nat & y2: nat)";
@@ -63,7 +63,7 @@
 AddSIs [intrelI];
 AddSEs [intrelE];
 
-goalw Integ.thy [equiv_def, refl_def, sym_def, trans_def]
+Goalw [equiv_def, refl_def, sym_def, trans_def]
     "equiv(nat*nat, intrel)";
 by (fast_tac (claset() addSEs [sym, integ_trans_lemma]) 1);
 qed "equiv_intrel";
@@ -77,12 +77,12 @@
 
 (** znat: the injection from nat to integ **)
 
-goalw Integ.thy [integ_def,quotient_def,znat_def]
+Goalw [integ_def,quotient_def,znat_def]
     "!!m. m : nat ==> $#m : integ";
 by (fast_tac (claset() addSIs [nat_0I]) 1);
 qed "znat_type";
 
-goalw Integ.thy [znat_def]
+Goalw [znat_def]
     "!!m n. [| $#m = $#n;  n: nat |] ==> m=n";
 by (dtac eq_intrelD 1);
 by (typechk_tac [nat_0I, SigmaI]);
@@ -92,7 +92,7 @@
 
 (**** zminus: unary negation on integ ****)
 
-goalw Integ.thy [congruent_def]
+Goalw [congruent_def]
     "congruent(intrel, %<x,y>. intrel``{<y,x>})";
 by Safe_tac;
 by (asm_full_simp_tac (simpset() addsimps add_ac) 1);
@@ -101,13 +101,13 @@
 (*Resolve th against the corresponding facts for zminus*)
 val zminus_ize = RSLIST [equiv_intrel, zminus_congruent];
 
-goalw Integ.thy [integ_def,zminus_def]
+Goalw [integ_def,zminus_def]
     "!!z. z : integ ==> $~z : integ";
 by (typechk_tac [split_type, SigmaI, zminus_ize UN_equiv_class_type,
                  quotientI]);
 qed "zminus_type";
 
-goalw Integ.thy [integ_def,zminus_def]
+Goalw [integ_def,zminus_def]
     "!!z w. [| $~z = $~w;  z: integ;  w: integ |] ==> z=w";
 by (etac (zminus_ize UN_equiv_class_inject) 1);
 by Safe_tac;
@@ -116,17 +116,17 @@
                        setloop dtac eq_intrelD) 1);
 qed "zminus_inject";
 
-goalw Integ.thy [zminus_def]
+Goalw [zminus_def]
     "!!x y.[| x: nat;  y: nat |] ==> $~ (intrel``{<x,y>}) = intrel `` {<y,x>}";
 by (asm_simp_tac (simpset() addsimps [zminus_ize UN_equiv_class, SigmaI]) 1);
 qed "zminus";
 
-goalw Integ.thy [integ_def] "!!z. z : integ ==> $~ ($~ z) = z";
+Goalw [integ_def] "!!z. z : integ ==> $~ ($~ z) = z";
 by (REPEAT (eresolve_tac [quotientE,SigmaE,ssubst] 1));
 by (asm_simp_tac (simpset() addsimps [zminus]) 1);
 qed "zminus_zminus";
 
-goalw Integ.thy [integ_def, znat_def] "$~ ($#0) = $#0";
+Goalw [integ_def, znat_def] "$~ ($#0) = $#0";
 by (simp_tac (simpset() addsimps [zminus]) 1);
 qed "zminus_0";
 
@@ -134,7 +134,7 @@
 (**** znegative: the test for negative integers ****)
 
 (*No natural number is negative!*)
-goalw Integ.thy [znegative_def, znat_def]  "~ znegative($# n)";
+Goalw [znegative_def, znat_def]  "~ znegative($# n)";
 by Safe_tac;
 by (dres_inst_tac [("psi", "?lhs=?rhs")] asm_rl 1);
 by (dres_inst_tac [("psi", "?lhs<?rhs")] asm_rl 1);
@@ -142,7 +142,7 @@
               (simpset() addsimps [add_le_self2 RS le_imp_not_lt])) 1);
 qed "not_znegative_znat";
 
-goalw Integ.thy [znegative_def, znat_def]
+Goalw [znegative_def, znat_def]
     "!!n. n: nat ==> znegative($~ $# succ(n))";
 by (asm_simp_tac (simpset() addsimps [zminus]) 1);
 by (REPEAT 
@@ -153,7 +153,7 @@
 
 (**** zmagnitude: magnitide of an integer, as a natural number ****)
 
-goalw Integ.thy [congruent_def]
+Goalw [congruent_def]
     "congruent(intrel, %<x,y>. (y#-x) #+ (x#-y))";
 by Safe_tac;
 by (ALLGOALS Asm_simp_tac);
@@ -176,25 +176,25 @@
 (*Resolve th against the corresponding facts for zmagnitude*)
 val zmagnitude_ize = RSLIST [equiv_intrel, zmagnitude_congruent];
 
-goalw Integ.thy [integ_def,zmagnitude_def]
+Goalw [integ_def,zmagnitude_def]
     "!!z. z : integ ==> zmagnitude(z) : nat";
 by (typechk_tac [split_type, zmagnitude_ize UN_equiv_class_type,
                  add_type, diff_type]);
 qed "zmagnitude_type";
 
-goalw Integ.thy [zmagnitude_def]
+Goalw [zmagnitude_def]
     "!!x y. [| x: nat;  y: nat |] ==> \
 \           zmagnitude (intrel``{<x,y>}) = (y #- x) #+ (x #- y)";
 by (asm_simp_tac
     (simpset() addsimps [zmagnitude_ize UN_equiv_class, SigmaI]) 1);
 qed "zmagnitude";
 
-goalw Integ.thy [znat_def]
+Goalw [znat_def]
     "!!n. n: nat ==> zmagnitude($# n) = n";
 by (asm_simp_tac (simpset() addsimps [zmagnitude]) 1);
 qed "zmagnitude_znat";
 
-goalw Integ.thy [znat_def]
+Goalw [znat_def]
     "!!n. n: nat ==> zmagnitude($~ $# n) = n";
 by (asm_simp_tac (simpset() addsimps [zmagnitude, zminus]) 1);
 qed "zmagnitude_zminus_znat";
@@ -204,7 +204,7 @@
 
 (** Congruence property for addition **)
 
-goalw Integ.thy [congruent2_def]
+Goalw [congruent2_def]
     "congruent2(intrel, %z1 z2.                      \
 \         let <x1,y1>=z1; <x2,y2>=z2                 \
 \                           in intrel``{<x1#+x2, y1#+y2>})";
@@ -222,14 +222,14 @@
 (*Resolve th against the corresponding facts for zadd*)
 val zadd_ize = RSLIST [equiv_intrel, zadd_congruent2];
 
-goalw Integ.thy [integ_def,zadd_def]
+Goalw [integ_def,zadd_def]
     "!!z w. [| z: integ;  w: integ |] ==> z $+ w : integ";
 by (rtac (zadd_ize UN_equiv_class_type2) 1);
 by (simp_tac (simpset() addsimps [Let_def]) 3);
 by (REPEAT (ares_tac [split_type, add_type, quotientI, SigmaI] 1));
 qed "zadd_type";
 
-goalw Integ.thy [zadd_def]
+Goalw [zadd_def]
   "!!x1 y1. [| x1: nat; y1: nat;  x2: nat; y2: nat |] ==>       \
 \           (intrel``{<x1,y1>}) $+ (intrel``{<x2,y2>}) =        \
 \           intrel `` {<x1#+x2, y1#+y2>}";
@@ -237,24 +237,24 @@
 by (simp_tac (simpset() addsimps [Let_def]) 1);
 qed "zadd";
 
-goalw Integ.thy [integ_def,znat_def] "!!z. z : integ ==> $#0 $+ z = z";
+Goalw [integ_def,znat_def] "!!z. z : integ ==> $#0 $+ z = z";
 by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
 by (asm_simp_tac (simpset() addsimps [zadd]) 1);
 qed "zadd_0";
 
-goalw Integ.thy [integ_def]
+Goalw [integ_def]
     "!!z w. [| z: integ;  w: integ |] ==> $~ (z $+ w) = $~ z $+ $~ w";
 by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
 by (asm_simp_tac (simpset() addsimps [zminus,zadd]) 1);
 qed "zminus_zadd_distrib";
 
-goalw Integ.thy [integ_def]
+Goalw [integ_def]
     "!!z w. [| z: integ;  w: integ |] ==> z $+ w = w $+ z";
 by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
 by (asm_simp_tac (simpset() addsimps (add_ac @ [zadd])) 1);
 qed "zadd_commute";
 
-goalw Integ.thy [integ_def]
+Goalw [integ_def]
     "!!z1 z2 z3. [| z1: integ;  z2: integ;  z3: integ |] ==> \
 \                (z1 $+ z2) $+ z3 = z1 $+ (z2 $+ z3)";
 by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
@@ -271,22 +271,22 @@
 (*Integer addition is an AC operator*)
 val zadd_ac = [zadd_assoc, zadd_commute, zadd_left_commute];
 
-goalw Integ.thy [znat_def]
+Goalw [znat_def]
     "!!m n. [| m: nat;  n: nat |] ==> $# (m #+ n) = ($#m) $+ ($#n)";
 by (asm_simp_tac (simpset() addsimps [zadd]) 1);
 qed "znat_add";
 
-goalw Integ.thy [integ_def,znat_def] "!!z. z : integ ==> z $+ ($~ z) = $#0";
+Goalw [integ_def,znat_def] "!!z. z : integ ==> z $+ ($~ z) = $#0";
 by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
 by (asm_simp_tac (simpset() addsimps [zminus, zadd, add_commute]) 1);
 qed "zadd_zminus_inverse";
 
-goal Integ.thy "!!z. z : integ ==> ($~ z) $+ z = $#0";
+Goal "!!z. z : integ ==> ($~ z) $+ z = $#0";
 by (asm_simp_tac
     (simpset() addsimps [zadd_commute, zminus_type, zadd_zminus_inverse]) 1);
 qed "zadd_zminus_inverse2";
 
-goal Integ.thy "!!z. z:integ ==> z $+ $#0 = z";
+Goal "!!z. z:integ ==> z $+ $#0 = z";
 by (rtac (zadd_commute RS trans) 1);
 by (REPEAT (ares_tac [znat_type, nat_0I, zadd_0] 1));
 qed "zadd_0_right";
@@ -300,7 +300,7 @@
 
 (** Congruence property for multiplication **)
 
-goal Integ.thy 
+Goal 
     "congruent2(intrel, %p1 p2.                 \
 \               split(%x1 y1. split(%x2 y2.     \
 \                   intrel``{<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}, p2), p1))";
@@ -320,50 +320,50 @@
 (*Resolve th against the corresponding facts for zmult*)
 val zmult_ize = RSLIST [equiv_intrel, zmult_congruent2];
 
-goalw Integ.thy [integ_def,zmult_def]
+Goalw [integ_def,zmult_def]
     "!!z w. [| z: integ;  w: integ |] ==> z $* w : integ";
 by (REPEAT (ares_tac [zmult_ize UN_equiv_class_type2,
                       split_type, add_type, mult_type, 
                       quotientI, SigmaI] 1));
 qed "zmult_type";
 
-goalw Integ.thy [zmult_def]
+Goalw [zmult_def]
      "!!x1 x2. [| x1: nat; y1: nat;  x2: nat; y2: nat |] ==>    \
 \              (intrel``{<x1,y1>}) $* (intrel``{<x2,y2>}) =     \
 \              intrel `` {<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}";
 by (asm_simp_tac (simpset() addsimps [zmult_ize UN_equiv_class2, SigmaI]) 1);
 qed "zmult";
 
-goalw Integ.thy [integ_def,znat_def] "!!z. z : integ ==> $#0 $* z = $#0";
+Goalw [integ_def,znat_def] "!!z. z : integ ==> $#0 $* z = $#0";
 by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
 by (asm_simp_tac (simpset() addsimps [zmult]) 1);
 qed "zmult_0";
 
-goalw Integ.thy [integ_def,znat_def]
+Goalw [integ_def,znat_def]
     "!!z. z : integ ==> $#1 $* z = z";
 by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
 by (asm_simp_tac (simpset() addsimps [zmult, add_0_right]) 1);
 qed "zmult_1";
 
-goalw Integ.thy [integ_def]
+Goalw [integ_def]
     "!!z w. [| z: integ;  w: integ |] ==> ($~ z) $* w = $~ (z $* w)";
 by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
 by (asm_simp_tac (simpset() addsimps ([zminus, zmult] @ add_ac)) 1);
 qed "zmult_zminus";
 
-goalw Integ.thy [integ_def]
+Goalw [integ_def]
     "!!z w. [| z: integ;  w: integ |] ==> ($~ z) $* ($~ w) = (z $* w)";
 by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
 by (asm_simp_tac (simpset() addsimps ([zminus, zmult] @ add_ac)) 1);
 qed "zmult_zminus_zminus";
 
-goalw Integ.thy [integ_def]
+Goalw [integ_def]
     "!!z w. [| z: integ;  w: integ |] ==> z $* w = w $* z";
 by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
 by (asm_simp_tac (simpset() addsimps ([zmult] @ add_ac @ mult_ac)) 1);
 qed "zmult_commute";
 
-goalw Integ.thy [integ_def]
+Goalw [integ_def]
     "!!z1 z2 z3. [| z1: integ;  z2: integ;  z3: integ |] ==> \
 \                (z1 $* z2) $* z3 = z1 $* (z2 $* z3)";
 by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
@@ -382,7 +382,7 @@
 (*Integer multiplication is an AC operator*)
 val zmult_ac = [zmult_assoc, zmult_commute, zmult_left_commute];
 
-goalw Integ.thy [integ_def]
+Goalw [integ_def]
     "!!z1 z2 z3. [| z1: integ;  z2: integ;  w: integ |] ==> \
 \                (z1 $+ z2) $* w = (z1 $* w) $+ (z2 $* w)";
 by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
--- a/src/ZF/ex/LList.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/ex/LList.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -21,7 +21,7 @@
 val LCons_iff      = llist.mk_free "LCons(a,l)=LCons(a',l') <-> a=a' & l=l'";
 val LNil_LCons_iff = llist.mk_free "~ LNil=LCons(a,l)";
 
-goal LList.thy "llist(A) = {0} <+> (A <*> llist(A))";
+Goal "llist(A) = {0} <+> (A <*> llist(A))";
 let open llist;  val rew = rewrite_rule con_defs in  
 by (fast_tac (claset() addSIs (subsetI ::map rew intrs) addEs [rew elim]) 1)
 end;
@@ -29,7 +29,7 @@
 
 (*** Lemmas to justify using "llist" in other recursive type definitions ***)
 
-goalw LList.thy llist.defs "!!A B. A<=B ==> llist(A) <= llist(B)";
+Goalw llist.defs "!!A B. A<=B ==> llist(A) <= llist(B)";
 by (rtac gfp_mono 1);
 by (REPEAT (rtac llist.bnd_mono 1));
 by (REPEAT (ares_tac (quniv_mono::basic_monos) 1));
@@ -43,7 +43,7 @@
 AddSDs [qunivD];
 AddSEs [Ord_in_Ord];
 
-goal LList.thy
+Goal
    "!!i. Ord(i) ==> ALL l: llist(quniv(A)). l Int Vset(i) <= univ(eclose(A))";
 by (etac trans_induct 1);
 by (rtac ballI 1);
@@ -55,7 +55,7 @@
 by (deepen_tac (claset() addIs [Ord_trans, Int_lower1 RS subset_trans]) 2 1);
 qed "llist_quniv_lemma";
 
-goal LList.thy "llist(quniv(A)) <= quniv(A)";
+Goal "llist(quniv(A)) <= quniv(A)";
 by (rtac (qunivI RS subsetI) 1);
 by (rtac Int_Vset_subset 1);
 by (REPEAT (ares_tac [llist_quniv_lemma RS bspec] 1));
@@ -71,7 +71,7 @@
 AddSEs [Ord_in_Ord, Pair_inject];
 
 (*Lemma for proving finality.  Unfold the lazy list; use induction hypothesis*)
-goal LList.thy
+Goal
    "!!i. Ord(i) ==> ALL l l'. <l,l'> : lleq(A) --> l Int Vset(i) <= l'";
 by (etac trans_induct 1);
 by (REPEAT (resolve_tac [allI, impI] 1));
@@ -94,7 +94,7 @@
 by (ALLGOALS Fast_tac);
 qed "lleq_symmetric";
 
-goal LList.thy "!!l l'. <l,l'> : lleq(A) ==> l=l'";
+Goal "!!l l'. <l,l'> : lleq(A) ==> l=l'";
 by (rtac equalityI 1);
 by (REPEAT (ares_tac [lleq_Int_Vset_subset RS Int_Vset_subset] 1
      ORELSE etac lleq_symmetric 1));
@@ -116,7 +116,7 @@
 
 (*** lconst -- defined directly using lfp, but equivalent to a LList_corec ***)
 
-goalw LList.thy llist.con_defs "bnd_mono(univ(a), %l. LCons(a,l))";
+Goalw llist.con_defs "bnd_mono(univ(a), %l. LCons(a,l))";
 by (rtac bnd_monoI 1);
 by (REPEAT (ares_tac [subset_refl, QInr_mono, QPair_mono] 2));
 by (REPEAT (ares_tac [subset_refl, A_subset_univ, 
@@ -131,12 +131,12 @@
 
 bind_thm ("member_subset_Union_eclose", (arg_into_eclose RS Union_upper));
 
-goal LList.thy "!!a A. a : A ==> lconst(a) : quniv(A)";
+Goal "!!a A. a : A ==> lconst(a) : quniv(A)";
 by (rtac (lconst_subset RS subset_trans RS qunivI) 1);
 by (etac (arg_into_eclose RS eclose_subset RS univ_mono) 1);
 qed "lconst_in_quniv";
 
-goal LList.thy "!!a A. a:A ==> lconst(a): llist(A)";
+Goal "!!a A. a:A ==> lconst(a): llist(A)";
 by (rtac (singletonI RS llist.coinduct) 1);
 by (etac (lconst_in_quniv RS singleton_subsetI) 1);
 by (fast_tac (claset() addSIs [lconst]) 1);
@@ -155,7 +155,7 @@
 
 (*Reasoning borrowed from lleq.ML; a similar proof works for all
   "productive" functions -- cf Coquand's "Infinite Objects in Type Theory".*)
-goal LList.thy
+Goal
    "!!i. Ord(i) ==> ALL l: llist(bool). flip(l) Int Vset(i) <= \
 \                   univ(eclose(bool))";
 by (etac trans_induct 1);
@@ -168,7 +168,7 @@
 by (deepen_tac (claset() addIs [Ord_trans, Int_lower1 RS subset_trans]) 2 1);
 qed "flip_llist_quniv_lemma";
 
-goal LList.thy "!!l. l: llist(bool) ==> flip(l) : quniv(bool)";
+Goal "!!l. l: llist(bool) ==> flip(l) : quniv(bool)";
 by (rtac (flip_llist_quniv_lemma RS bspec RS Int_Vset_subset RS qunivI) 1);
 by (REPEAT (assume_tac 1));
 qed "flip_in_quniv";
--- a/src/ZF/ex/Limit.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/ex/Limit.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -315,7 +315,7 @@
 brr(cpo_refl::prems) 1;
 qed "chain_const";
 
-goalw Limit.thy [islub_def,isub_def] (* islub_const *)
+Goalw [islub_def,isub_def] (* islub_const *)
     "!!x D. [|x:set(D); cpo(D)|] ==> islub(D,(lam n:nat. x),x)";
 by (Asm_simp_tac 1);
 by (Blast_tac 1);
@@ -1717,7 +1717,7 @@
 by (asm_full_simp_tac(simpset() addsimps prems) 1); (* Surprise, surprise. *)
 qed "succ_le_pos";
 
-goal Limit.thy  (* lemma_le_exists *)
+Goal  (* lemma_le_exists *)
     "!!z. [|n:nat; m:nat|] ==> m le n --> (EX k:nat. n = m #+ k)";
 by (res_inst_tac[("n","m")]nat_induct 1);
 by (assume_tac 1);
--- a/src/ZF/ex/ListN.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/ex/ListN.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -11,31 +11,31 @@
 
 open ListN;
 
-goal ListN.thy "!!l. l:list(A) ==> <length(l),l> : listn(A)";
+Goal "!!l. l:list(A) ==> <length(l),l> : listn(A)";
 by (etac list.induct 1);
 by (ALLGOALS Asm_simp_tac);
 by (REPEAT (ares_tac listn.intrs 1));
 qed "list_into_listn";
 
-goal ListN.thy "<n,l> : listn(A) <-> l:list(A) & length(l)=n";
+Goal "<n,l> : listn(A) <-> l:list(A) & length(l)=n";
 by (rtac iffI 1);
 by (etac listn.induct 1);
 by (safe_tac (claset() addSIs (list_typechecks @
                             [length_Nil, length_Cons, list_into_listn])));
 qed "listn_iff";
 
-goal ListN.thy "listn(A)``{n} = {l:list(A). length(l)=n}";
+Goal "listn(A)``{n} = {l:list(A). length(l)=n}";
 by (rtac equality_iffI 1);
 by (simp_tac (simpset() addsimps [listn_iff,separation,image_singleton_iff]) 1);
 qed "listn_image_eq";
 
-goalw ListN.thy listn.defs "!!A B. A<=B ==> listn(A) <= listn(B)";
+Goalw listn.defs "!!A B. A<=B ==> listn(A) <= listn(B)";
 by (rtac lfp_mono 1);
 by (REPEAT (rtac listn.bnd_mono 1));
 by (REPEAT (ares_tac ([univ_mono,Sigma_mono,list_mono] @ basic_monos) 1));
 qed "listn_mono";
 
-goal ListN.thy
+Goal
     "!!n l. [| <n,l> : listn(A);  <n',l'> : listn(A) |] ==> \
 \           <n#+n', l@l'> : listn(A)";
 by (etac listn.induct 1);
--- a/src/ZF/ex/Mutil.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/ex/Mutil.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -11,33 +11,33 @@
 
 (** Basic properties of evnodd **)
 
-goalw thy [evnodd_def] "<i,j>: evnodd(A,b) <-> <i,j>: A & (i#+j) mod 2 = b";
+Goalw [evnodd_def] "<i,j>: evnodd(A,b) <-> <i,j>: A & (i#+j) mod 2 = b";
 by (Blast_tac 1);
 qed "evnodd_iff";
 
-goalw thy [evnodd_def] "evnodd(A, b) <= A";
+Goalw [evnodd_def] "evnodd(A, b) <= A";
 by (Blast_tac 1);
 qed "evnodd_subset";
 
 (* Finite(X) ==> Finite(evnodd(X,b)) *)
 bind_thm("Finite_evnodd", evnodd_subset RS subset_imp_lepoll RS lepoll_Finite);
 
-goalw thy [evnodd_def] "evnodd(A Un B, b) = evnodd(A,b) Un evnodd(B,b)";
+Goalw [evnodd_def] "evnodd(A Un B, b) = evnodd(A,b) Un evnodd(B,b)";
 by (simp_tac (simpset() addsimps [Collect_Un]) 1);
 qed "evnodd_Un";
 
-goalw thy [evnodd_def] "evnodd(A - B, b) = evnodd(A,b) - evnodd(B,b)";
+Goalw [evnodd_def] "evnodd(A - B, b) = evnodd(A,b) - evnodd(B,b)";
 by (simp_tac (simpset() addsimps [Collect_Diff]) 1);
 qed "evnodd_Diff";
 
-goalw thy [evnodd_def]
+Goalw [evnodd_def]
     "evnodd(cons(<i,j>,C), b) = \
 \    if((i#+j) mod 2 = b, cons(<i,j>, evnodd(C,b)), evnodd(C,b))";
 by (asm_simp_tac (simpset() addsimps [evnodd_def, Collect_cons] 
                         setloop split_tac [expand_if]) 1);
 qed "evnodd_cons";
 
-goalw thy [evnodd_def] "evnodd(0, b) = 0";
+Goalw [evnodd_def] "evnodd(0, b) = 0";
 by (simp_tac (simpset() addsimps [evnodd_def]) 1);
 qed "evnodd_0";
 
@@ -45,11 +45,11 @@
 
 (*** Dominoes ***)
 
-goal thy "!!d. d:domino ==> Finite(d)";
+Goal "!!d. d:domino ==> Finite(d)";
 by (blast_tac (claset() addSIs [Finite_cons, Finite_0] addEs [domino.elim]) 1);
 qed "domino_Finite";
 
-goal thy "!!d. [| d:domino; b<2 |] ==> EX i' j'. evnodd(d,b) = {<i',j'>}";
+Goal "!!d. [| d:domino; b<2 |] ==> EX i' j'. evnodd(d,b) = {<i',j'>}";
 by (eresolve_tac [domino.elim] 1);
 by (res_inst_tac [("k1", "i#+j")] (mod2_cases RS disjE) 2);
 by (res_inst_tac [("k1", "i#+j")] (mod2_cases RS disjE) 1);
@@ -65,7 +65,7 @@
 
 (** The union of two disjoint tilings is a tiling **)
 
-goal thy "!!t. t: tiling(A) ==> \
+Goal "!!t. t: tiling(A) ==> \
 \              u: tiling(A) --> t Int u = 0 --> t Un u : tiling(A)";
 by (etac tiling.induct 1);
 by (simp_tac (simpset() addsimps tiling.intrs) 1);
@@ -74,13 +74,13 @@
 by (blast_tac (claset() addIs tiling.intrs) 1);
 qed_spec_mp "tiling_UnI";
 
-goal thy "!!t. t:tiling(domino) ==> Finite(t)";
+Goal "!!t. t:tiling(domino) ==> Finite(t)";
 by (eresolve_tac [tiling.induct] 1);
 by (rtac Finite_0 1);
 by (blast_tac (claset() addSIs [Finite_Un] addIs [domino_Finite]) 1);
 qed "tiling_domino_Finite";
 
-goal thy "!!t. t: tiling(domino) ==> |evnodd(t,0)| = |evnodd(t,1)|";
+Goal "!!t. t: tiling(domino) ==> |evnodd(t,0)| = |evnodd(t,1)|";
 by (eresolve_tac [tiling.induct] 1);
 by (simp_tac (simpset() addsimps [evnodd_def]) 1);
 by (res_inst_tac [("b1","0")] (domino_singleton RS exE) 1);
@@ -95,7 +95,7 @@
 by (blast_tac (claset() addSDs [evnodd_subset RS subsetD] addEs [equalityE]) 1);
 qed "tiling_domino_0_1";
 
-goal thy "!!i n. [| i: nat;  n: nat |] ==> {i} * (n #+ n) : tiling(domino)";
+Goal "!!i n. [| i: nat;  n: nat |] ==> {i} * (n #+ n) : tiling(domino)";
 by (nat_ind_tac "n" [] 1);
 by (simp_tac (simpset() addsimps tiling.intrs) 1);
 by (asm_simp_tac (simpset() addsimps [Un_assoc RS sym, Sigma_succ2]) 1);
@@ -108,7 +108,7 @@
 by (blast_tac (claset() addEs [mem_irrefl, mem_asym]) 1);
 qed "dominoes_tile_row";
 
-goal thy "!!m n. [| m: nat;  n: nat |] ==> m * (n #+ n) : tiling(domino)";
+Goal "!!m n. [| m: nat;  n: nat |] ==> m * (n #+ n) : tiling(domino)";
 by (nat_ind_tac "m" [] 1);
 by (simp_tac (simpset() addsimps tiling.intrs) 1);
 by (asm_simp_tac (simpset() addsimps [Sigma_succ1]) 1);
@@ -117,7 +117,7 @@
 qed "dominoes_tile_matrix";
 
 
-goal thy "!!m n. [| m: nat;  n: nat;  \
+Goal "!!m n. [| m: nat;  n: nat;  \
 \                   t = (succ(m)#+succ(m))*(succ(n)#+succ(n));  \
 \                   t' = t - {<0,0>} - {<succ(m#+m), succ(n#+n)>} |] ==> \
 \                t' ~: tiling(domino)";
--- a/src/ZF/ex/Ntree.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/ex/Ntree.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -13,7 +13,7 @@
 
 (** ntree **)
 
-goal Ntree.thy "ntree(A) = A * (UN n: nat. n -> ntree(A))";
+Goal "ntree(A) = A * (UN n: nat. n -> ntree(A))";
 let open ntree;  val rew = rewrite_rule con_defs in  
 by (fast_tac (claset() addSIs (map rew intrs) addEs [rew elim]) 1)
 end;
@@ -48,14 +48,14 @@
 
 (**  Lemmas to justify using "Ntree" in other recursive type definitions **)
 
-goalw Ntree.thy ntree.defs "!!A B. A<=B ==> ntree(A) <= ntree(B)";
+Goalw ntree.defs "!!A B. A<=B ==> ntree(A) <= ntree(B)";
 by (rtac lfp_mono 1);
 by (REPEAT (rtac ntree.bnd_mono 1));
 by (REPEAT (ares_tac (univ_mono::basic_monos) 1));
 qed "ntree_mono";
 
 (*Easily provable by induction also*)
-goalw Ntree.thy (ntree.defs@ntree.con_defs) "ntree(univ(A)) <= univ(A)";
+Goalw (ntree.defs@ntree.con_defs) "ntree(univ(A)) <= univ(A)";
 by (rtac lfp_lowerbound 1);
 by (rtac (A_subset_univ RS univ_mono) 2);
 by Safe_tac;
@@ -67,7 +67,7 @@
 
 (** maptree **)
 
-goal Ntree.thy "maptree(A) = A * (maptree(A) -||> maptree(A))";
+Goal "maptree(A) = A * (maptree(A) -||> maptree(A))";
 let open maptree;  val rew = rewrite_rule con_defs in  
 by (fast_tac (claset() addSIs (map rew intrs) addEs [rew elim]) 1)
 end;
@@ -91,7 +91,7 @@
 
 (** maptree2 **)
 
-goal Ntree.thy "maptree2(A,B) = A * (B -||> maptree2(A,B))";
+Goal "maptree2(A,B) = A * (B -||> maptree2(A,B))";
 let open maptree2;  val rew = rewrite_rule con_defs in  
 by (fast_tac (claset() addSIs (map rew intrs) addEs [rew elim]) 1)
 end;
--- a/src/ZF/ex/Primes.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/ex/Primes.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -14,7 +14,7 @@
 (** Divides Relation                           **)
 (************************************************)
 
-goalw thy [dvd_def] "!!m. m dvd n ==> m:nat & n:nat & (EX k:nat. n = m#*k)";
+Goalw [dvd_def] "!!m. m dvd n ==> m:nat & n:nat & (EX k:nat. n = m#*k)";
 by (assume_tac 1);
 qed "dvdD";
 
@@ -22,23 +22,23 @@
 bind_thm ("dvd_imp_nat2", dvdD RS conjunct2 RS conjunct1);
 
 
-goalw thy [dvd_def] "!!m. m:nat ==> m dvd 0";
+Goalw [dvd_def] "!!m. m:nat ==> m dvd 0";
 by (fast_tac (claset() addIs [nat_0I, mult_0_right RS sym]) 1);
 qed "dvd_0_right";
 
-goalw thy [dvd_def] "!!m. 0 dvd m ==> m = 0";
+Goalw [dvd_def] "!!m. 0 dvd m ==> m = 0";
 by (fast_tac (claset() addss (simpset())) 1);
 qed "dvd_0_left";
 
-goalw thy [dvd_def] "!!m. m:nat ==> m dvd m";
+Goalw [dvd_def] "!!m. m:nat ==> m dvd m";
 by (fast_tac (claset() addIs [nat_1I, mult_1_right RS sym]) 1);
 qed "dvd_refl";
 
-goalw thy [dvd_def] "!!m n p. [| m dvd n; n dvd p |] ==> m dvd p";
+Goalw [dvd_def] "!!m n p. [| m dvd n; n dvd p |] ==> m dvd p";
 by (fast_tac (claset() addIs [mult_assoc, mult_type] ) 1);
 qed "dvd_trans";
 
-goalw thy [dvd_def] "!!m n. [| m dvd n; n dvd m |] ==> m=n";
+Goalw [dvd_def] "!!m n. [| m dvd n; n dvd m |] ==> m=n";
 by (fast_tac (claset() addDs [mult_eq_self_implies_10]
                     addss (simpset() addsimps [mult_assoc, mult_eq_1_iff])) 1);
 qed "dvd_anti_sym";
@@ -50,35 +50,35 @@
 
 (* GCD by Euclid's Algorithm *)
 
-goalw thy [egcd_def] "!!m. m:nat ==> egcd(m,0) = m";
+Goalw [egcd_def] "!!m. m:nat ==> egcd(m,0) = m";
 by (stac transrec 1);
 by (Asm_simp_tac 1);
 qed "egcd_0";
 
-goalw thy [egcd_def]
+Goalw [egcd_def]
     "!!m. [| 0<n; m:nat; n:nat |] ==> egcd(m,n) = egcd(n, m mod n)";
 by (res_inst_tac [("P", "%z. ?left(z) = ?right")] (transrec RS ssubst) 1);
 by (asm_simp_tac (simpset() addsimps [ltD RS mem_imp_not_eq RS not_sym,
                                      mod_less_divisor RS ltD]) 1);
 qed "egcd_lt_0";
 
-goal thy "!!m. m:nat ==> egcd(m,0) dvd m";
+Goal "!!m. m:nat ==> egcd(m,0) dvd m";
 by (asm_simp_tac (simpset() addsimps [egcd_0,dvd_refl]) 1);
 qed "egcd_0_dvd_m";
 
-goal thy "!!m. m:nat ==> egcd(m,0) dvd 0";
+Goal "!!m. m:nat ==> egcd(m,0) dvd 0";
 by (asm_simp_tac (simpset() addsimps [egcd_0,dvd_0_right]) 1);
 qed "egcd_0_dvd_0";
 
-goalw thy [dvd_def] "!!k. [| k dvd a; k dvd b |] ==> k dvd (a #+ b)";
+Goalw [dvd_def] "!!k. [| k dvd a; k dvd b |] ==> k dvd (a #+ b)";
 by (fast_tac (claset() addIs [add_mult_distrib_left RS sym, add_type]) 1);
 qed "dvd_add";
 
-goalw thy [dvd_def] "!!k. [| k dvd a; q:nat |] ==> k dvd (q #* a)";
+Goalw [dvd_def] "!!k. [| k dvd a; q:nat |] ==> k dvd (q #* a)";
 by (fast_tac (claset() addIs [mult_left_commute, mult_type]) 1);
 qed "dvd_mult";
 
-goal thy "!!k. [| k dvd b; k dvd (a mod b); 0 < b; a:nat |] ==> k dvd a";
+Goal "!!k. [| k dvd b; k dvd (a mod b); 0 < b; a:nat |] ==> k dvd a";
 by (deepen_tac 
     (claset() addIs [mod_div_equality RS subst]
            addDs [dvdD]
@@ -88,7 +88,7 @@
 
 (* egcd type *)
 
-goal thy "!!b. b:nat ==> ALL a:nat. egcd(a,b):nat";
+Goal "!!b. b:nat ==> ALL a:nat. egcd(a,b):nat";
 by (etac complete_induct 1);
 by (rtac ballI 1);
 by (excluded_middle_tac "x=0" 1);
@@ -105,7 +105,7 @@
 
 (* Property 1: egcd(a,b) divides a and b *)
 
-goal thy "!!b. b:nat ==> ALL a: nat. (egcd(a,b) dvd a) & (egcd(a,b) dvd b)";
+Goal "!!b. b:nat ==> ALL a: nat. (egcd(a,b) dvd a) & (egcd(a,b) dvd b)";
 by (res_inst_tac [("i","b")] complete_induct 1);
 by (assume_tac 1);
 by (rtac ballI 1);
@@ -124,7 +124,7 @@
 
 (* if f divides a and b then f divides egcd(a,b) *)
 
-goalw thy [dvd_def] "!!a. [| f dvd a; f dvd b; 0<b |] ==> f dvd (a mod b)";
+Goalw [dvd_def] "!!a. [| f dvd a; f dvd b; 0<b |] ==> f dvd (a mod b)";
 by (safe_tac (claset() addSIs [mult_type, mod_type]));
 ren "m n" 1;
 by (rtac (zero_lt_mult_iff RS iffD1 RS conjE) 1);
@@ -141,7 +141,7 @@
 
 (* Property 2: for all a,b,f naturals, 
                if f divides a and f divides b then f divides egcd(a,b)*)
-goal thy "!!b. [| b:nat; f:nat |] ==>    \
+Goal "!!b. [| b:nat; f:nat |] ==>    \
 \              ALL a. (f dvd a) & (f dvd b) --> f dvd egcd(a,b)";
 by (etac complete_induct 1);
 by (rtac allI 1);
@@ -162,14 +162,14 @@
 
 (* GCD PROOF : GCD exists and egcd fits the definition *)
 
-goalw thy [gcd_def] "!!b. [| a: nat; b:nat |] ==> gcd(egcd(a,b), a, b)";
+Goalw [gcd_def] "!!b. [| a: nat; b:nat |] ==> gcd(egcd(a,b), a, b)";
 by (asm_simp_tac (simpset() addsimps [egcd_prop1]) 1);
 by (fast_tac (claset() addIs [egcd_prop2 RS spec RS mp, dvd_imp_nat1]) 1);
 qed "gcd";
 
 (* GCD is unique *)
 
-goalw thy [gcd_def] "!!a. gcd(m,a,b) & gcd(n,a,b) --> m=n";
+Goalw [gcd_def] "!!a. gcd(m,a,b) & gcd(n,a,b) --> m=n";
 by (fast_tac (claset() addIs [dvd_anti_sym]) 1);
 qed "gcd_unique";
 
--- a/src/ZF/ex/Primrec.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/ex/Primrec.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -24,32 +24,32 @@
 
 simpset_ref() := simpset() setSolver (type_auto_tac pr_typechecks);
 
-goalw Primrec.thy [SC_def]
+Goalw [SC_def]
     "!!x l. [| x:nat;  l: list(nat) |] ==> SC ` (Cons(x,l)) = succ(x)";
 by (Asm_simp_tac 1);
 qed "SC";
 
-goalw Primrec.thy [CONST_def]
+Goalw [CONST_def]
     "!!l. [| l: list(nat) |] ==> CONST(k) ` l = k";
 by (Asm_simp_tac 1);
 qed "CONST";
 
-goalw Primrec.thy [PROJ_def]
+Goalw [PROJ_def]
     "!!l. [| x: nat;  l: list(nat) |] ==> PROJ(0) ` (Cons(x,l)) = x";
 by (Asm_simp_tac 1);
 qed "PROJ_0";
 
-goalw Primrec.thy [COMP_def]
+Goalw [COMP_def]
     "!!l. [| l: list(nat) |] ==> COMP(g,[f]) ` l = g` [f`l]";
 by (Asm_simp_tac 1);
 qed "COMP_1";
 
-goalw Primrec.thy [PREC_def]
+Goalw [PREC_def]
     "!!l. l: list(nat) ==> PREC(f,g) ` (Cons(0,l)) = f`l";
 by (Asm_simp_tac 1);
 qed "PREC_0";
 
-goalw Primrec.thy [PREC_def]
+Goalw [PREC_def]
     "!!l. [| x:nat;  l: list(nat) |] ==>  \
 \         PREC(f,g) ` (Cons(succ(x),l)) = \
 \         g ` Cons(PREC(f,g)`(Cons(x,l)), Cons(x,l))";
@@ -64,7 +64,7 @@
 simpset_ref() := simpset() setSolver (type_auto_tac ([primrec_into_fun] @ 
 					      pr_typechecks @ primrec.intrs));
 
-goalw Primrec.thy [ACK_def] "!!i. i:nat ==> ACK(i): primrec";
+Goalw [ACK_def] "!!i. i:nat ==> ACK(i): primrec";
 by (etac nat_induct 1);
 by (ALLGOALS Asm_simp_tac);
 qed "ACK_in_primrec";
@@ -79,26 +79,26 @@
     REPEAT
       (SOMEGOAL (test_assume_tac ORELSE' match_tac (rls @ ack_typechecks)));
 
-goal Primrec.thy "!!i j. [| i:nat;  j:nat |] ==>  ack(i,j): nat";
+Goal "!!i j. [| i:nat;  j:nat |] ==>  ack(i,j): nat";
 by (tc_tac []);
 qed "ack_type";
 
 (** Ackermann's function cases **)
 
 (*PROPERTY A 1*)
-goalw Primrec.thy [ACK_def] "!!j. j:nat ==> ack(0,j) = succ(j)";
+Goalw [ACK_def] "!!j. j:nat ==> ack(0,j) = succ(j)";
 by (asm_simp_tac (simpset() addsimps [SC]) 1);
 qed "ack_0";
 
 (*PROPERTY A 2*)
-goalw Primrec.thy [ACK_def] "ack(succ(i), 0) = ack(i,1)";
+Goalw [ACK_def] "ack(succ(i), 0) = ack(i,1)";
 by (asm_simp_tac (simpset() addsimps [CONST,PREC_0]) 1);
 qed "ack_succ_0";
 
 (*PROPERTY A 3*)
 (*Could be proved in Primrec0, like the previous two cases, but using
   primrec_into_fun makes type-checking easier!*)
-goalw Primrec.thy [ACK_def]
+Goalw [ACK_def]
     "!!i j. [| i:nat;  j:nat |] ==> \
 \           ack(succ(i), succ(j)) = ack(i, ack(succ(i), j))";
 by (asm_simp_tac (simpset() addsimps [CONST,PREC_succ,COMP_1,PROJ_0]) 1);
@@ -107,7 +107,7 @@
 Addsimps [ack_0, ack_succ_0, ack_succ_succ, ack_type, nat_into_Ord];
 
 (*PROPERTY A 4*)
-goal Primrec.thy "!!i. i:nat ==> ALL j:nat. j < ack(i,j)";
+Goal "!!i. i:nat ==> ALL j:nat. j < ack(i,j)";
 by (etac nat_induct 1);
 by (Asm_simp_tac 1);
 by (rtac ballI 1);
@@ -120,13 +120,13 @@
 bind_thm ("lt_ack2", (lt_ack2_lemma RS bspec));
 
 (*PROPERTY A 5-, the single-step lemma*)
-goal Primrec.thy "!!i j. [| i:nat; j:nat |] ==> ack(i,j) < ack(i, succ(j))";
+Goal "!!i j. [| i:nat; j:nat |] ==> ack(i,j) < ack(i, succ(j))";
 by (etac nat_induct 1);
 by (ALLGOALS (asm_simp_tac (simpset() addsimps [lt_ack2])));
 qed "ack_lt_ack_succ2";
 
 (*PROPERTY A 5, monotonicity for < *)
-goal Primrec.thy "!!i j k. [| j<k; i:nat; k:nat |] ==> ack(i,j) < ack(i,k)";
+Goal "!!i j k. [| j<k; i:nat; k:nat |] ==> ack(i,j) < ack(i,k)";
 by (forward_tac [lt_nat_in_nat] 1 THEN assume_tac 1);
 by (etac succ_lt_induct 1);
 by (assume_tac 1);
@@ -135,14 +135,14 @@
 qed "ack_lt_mono2";
 
 (*PROPERTY A 5', monotonicity for le *)
-goal Primrec.thy
+Goal
     "!!i j k. [| j le k;  i: nat;  k:nat |] ==> ack(i,j) le ack(i,k)";
 by (res_inst_tac [("f", "%j. ack(i,j)")] Ord_lt_mono_imp_le_mono 1);
 by (REPEAT (ares_tac [ack_lt_mono2, ack_type RS nat_into_Ord] 1));
 qed "ack_le_mono2";
 
 (*PROPERTY A 6*)
-goal Primrec.thy
+Goal
     "!!i j. [| i:nat;  j:nat |] ==> ack(i, succ(j)) le ack(succ(i), j)";
 by (nat_ind_tac "j" [] 1);
 by (ALLGOALS Asm_simp_tac);
@@ -152,14 +152,14 @@
 qed "ack2_le_ack1";
 
 (*PROPERTY A 7-, the single-step lemma*)
-goal Primrec.thy "!!i j. [| i:nat; j:nat |] ==> ack(i,j) < ack(succ(i),j)";
+Goal "!!i j. [| i:nat; j:nat |] ==> ack(i,j) < ack(succ(i),j)";
 by (rtac (ack_lt_mono2 RS lt_trans2) 1);
 by (rtac ack2_le_ack1 4);
 by (REPEAT (ares_tac ([nat_le_refl, ack_type] @ pr_typechecks) 1));
 qed "ack_lt_ack_succ1";
 
 (*PROPERTY A 7, monotonicity for < *)
-goal Primrec.thy "!!i j k. [| i<j; j:nat; k:nat |] ==> ack(i,k) < ack(j,k)";
+Goal "!!i j k. [| i<j; j:nat; k:nat |] ==> ack(i,k) < ack(j,k)";
 by (forward_tac [lt_nat_in_nat] 1 THEN assume_tac 1);
 by (etac succ_lt_induct 1);
 by (assume_tac 1);
@@ -168,26 +168,26 @@
 qed "ack_lt_mono1";
 
 (*PROPERTY A 7', monotonicity for le *)
-goal Primrec.thy
+Goal
     "!!i j k. [| i le j; j:nat; k:nat |] ==> ack(i,k) le ack(j,k)";
 by (res_inst_tac [("f", "%j. ack(j,k)")] Ord_lt_mono_imp_le_mono 1);
 by (REPEAT (ares_tac [ack_lt_mono1, ack_type RS nat_into_Ord] 1));
 qed "ack_le_mono1";
 
 (*PROPERTY A 8*)
-goal Primrec.thy "!!j. j:nat ==> ack(1,j) = succ(succ(j))";
+Goal "!!j. j:nat ==> ack(1,j) = succ(succ(j))";
 by (etac nat_induct 1);
 by (ALLGOALS Asm_simp_tac);
 qed "ack_1";
 
 (*PROPERTY A 9*)
-goal Primrec.thy "!!j. j:nat ==> ack(succ(1),j) = succ(succ(succ(j#+j)))";
+Goal "!!j. j:nat ==> ack(succ(1),j) = succ(succ(succ(j#+j)))";
 by (etac nat_induct 1);
 by (ALLGOALS (asm_simp_tac (simpset() addsimps [ack_1, add_succ_right])));
 qed "ack_2";
 
 (*PROPERTY A 10*)
-goal Primrec.thy
+Goal
     "!!i1 i2 j. [| i1:nat; i2:nat; j:nat |] ==> \
 \               ack(i1, ack(i2,j)) < ack(succ(succ(i1#+i2)), j)";
 by (rtac (ack2_le_ack1 RSN (2,lt_trans2)) 1);
@@ -198,7 +198,7 @@
 qed "ack_nest_bound";
 
 (*PROPERTY A 11*)
-goal Primrec.thy
+Goal
     "!!i1 i2 j. [| i1:nat; i2:nat; j:nat |] ==> \
 \          ack(i1,j) #+ ack(i2,j) < ack(succ(succ(succ(succ(i1#+i2)))), j)";
 by (res_inst_tac [("j", "ack(succ(1), ack(i1 #+ i2, j))")] lt_trans 1);
@@ -212,7 +212,7 @@
 
 (*PROPERTY A 12.  Article uses existential quantifier but the ALF proof
   used k#+4.  Quantified version must be nested EX k'. ALL i,j... *)
-goal Primrec.thy
+Goal
     "!!i j k. [| i < ack(k,j);  j:nat;  k:nat |] ==> \
 \             i#+j < ack(succ(succ(succ(succ(k)))), j)";
 by (res_inst_tac [("j", "ack(k,j) #+ ack(0,j)")] lt_trans 1);
@@ -225,7 +225,7 @@
 
 Addsimps [list_add_type, nat_into_Ord];
 
-goalw Primrec.thy [SC_def]
+Goalw [SC_def]
     "!!l. l: list(nat) ==> SC ` l < ack(1, list_add(l))";
 by (etac list.elim 1);
 by (asm_simp_tac (simpset() addsimps [succ_iff]) 1);
@@ -233,19 +233,19 @@
 qed "SC_case";
 
 (*PROPERTY A 4'? Extra lemma needed for CONST case, constant functions*)
-goal Primrec.thy "!!j. [| i:nat; j:nat |] ==> i < ack(i,j)";
+Goal "!!j. [| i:nat; j:nat |] ==> i < ack(i,j)";
 by (etac nat_induct 1);
 by (asm_simp_tac (simpset() addsimps [nat_0_le]) 1);
 by (etac ([succ_leI, ack_lt_ack_succ1] MRS lt_trans1) 1);
 by (tc_tac []);
 qed "lt_ack1";
 
-goalw Primrec.thy [CONST_def]
+Goalw [CONST_def]
     "!!l. [| l: list(nat);  k: nat |] ==> CONST(k) ` l < ack(k, list_add(l))";
 by (asm_simp_tac (simpset() addsimps [lt_ack1]) 1);
 qed "CONST_case";
 
-goalw Primrec.thy [PROJ_def]
+Goalw [PROJ_def]
     "!!l. l: list(nat) ==> ALL i:nat. PROJ(i) ` l < ack(0, list_add(l))";
 by (Asm_simp_tac 1);
 by (etac list.induct 1);
@@ -263,7 +263,7 @@
 
 (** COMP case **)
 
-goal Primrec.thy
+Goal
  "!!fs. fs : list({f: primrec .                                 \
 \                  EX kf:nat. ALL l:list(nat).                  \
 \                             f`l < ack(kf, list_add(l))})      \
@@ -282,7 +282,7 @@
 by (tc_tac []);
 qed "COMP_map_lemma";
 
-goalw Primrec.thy [COMP_def]
+Goalw [COMP_def]
  "!!g. [| kg: nat;                                 \
 \         ALL l:list(nat). g`l < ack(kg, list_add(l));          \
 \         fs : list({f: primrec .                               \
@@ -302,7 +302,7 @@
 
 (** PREC case **)
 
-goalw Primrec.thy [PREC_def]
+Goalw [PREC_def]
  "!!f g. [| ALL l:list(nat). f`l #+ list_add(l) < ack(kf, list_add(l)); \
 \           ALL l:list(nat). g`l #+ list_add(l) < ack(kg, list_add(l)); \
 \           f: primrec;  kf: nat;                                       \
@@ -333,7 +333,7 @@
 by (tc_tac []);
 qed "PREC_case_lemma";
 
-goal Primrec.thy
+Goal
  "!!f g. [| f: primrec;  kf: nat;                               \
 \           g: primrec;  kg: nat;                               \
 \           ALL l:list(nat). f`l < ack(kf, list_add(l));        \
@@ -349,7 +349,7 @@
               rtac (ack_add_bound2 RS ballI) THEN' etac bspec])));
 qed "PREC_case";
 
-goal Primrec.thy
+Goal
     "!!f. f:primrec ==> EX k:nat. ALL l:list(nat). f`l < ack(k, list_add(l))";
 by (etac primrec.induct 1);
 by Safe_tac;
@@ -358,7 +358,7 @@
                        bexI, ballI] @ nat_typechecks) 1));
 qed "ack_bounds_primrec";
 
-goal Primrec.thy
+Goal
     "~ (lam l:list(nat). list_case(0, %x xs. ack(x,x), l)) : primrec";
 by (rtac notI 1);
 by (etac (ack_bounds_primrec RS bexE) 1);
--- a/src/ZF/ex/PropLog.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/ex/PropLog.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -15,19 +15,19 @@
 
 (** conversion rules **)
 
-goal PropLog.thy "prop_rec(Fls,b,c,d) = b";
+Goal "prop_rec(Fls,b,c,d) = b";
 by (rtac (prop_rec_def RS def_Vrec RS trans) 1);
 by (rewrite_goals_tac prop.con_defs);
 by (Simp_tac 1);
 qed "prop_rec_Fls";
 
-goal PropLog.thy "prop_rec(#v,b,c,d) = c(v)";
+Goal "prop_rec(#v,b,c,d) = c(v)";
 by (rtac (prop_rec_def RS def_Vrec RS trans) 1);
 by (rewrite_goals_tac prop.con_defs);
 by (Simp_tac 1);
 qed "prop_rec_Var";
 
-goal PropLog.thy "prop_rec(p=>q,b,c,d) = \
+Goal "prop_rec(p=>q,b,c,d) = \
 \      d(p, q, prop_rec(p,b,c,d), prop_rec(q,b,c,d))";
 by (rtac (prop_rec_def RS def_Vrec RS trans) 1);
 by (rewrite_goals_tac prop.con_defs);
@@ -40,31 +40,31 @@
 
 (** The function is_true **)
 
-goalw PropLog.thy [is_true_def] "is_true(Fls,t) <-> False";
+Goalw [is_true_def] "is_true(Fls,t) <-> False";
 by (simp_tac (simpset() addsimps [one_not_0 RS not_sym]) 1);
 qed "is_true_Fls";
 
-goalw PropLog.thy [is_true_def] "is_true(#v,t) <-> v:t";
+Goalw [is_true_def] "is_true(#v,t) <-> v:t";
 by (simp_tac (simpset() addsimps [one_not_0 RS not_sym] 
               setloop (split_tac [expand_if])) 1);
 qed "is_true_Var";
 
-goalw PropLog.thy [is_true_def]
+Goalw [is_true_def]
     "is_true(p=>q,t) <-> (is_true(p,t)-->is_true(q,t))";
 by (simp_tac (simpset() setloop (split_tac [expand_if])) 1);
 qed "is_true_Imp";
 
 (** The function hyps **)
 
-goalw PropLog.thy [hyps_def] "hyps(Fls,t) = 0";
+Goalw [hyps_def] "hyps(Fls,t) = 0";
 by (Simp_tac 1);
 qed "hyps_Fls";
 
-goalw PropLog.thy [hyps_def] "hyps(#v,t) = {if(v:t, #v, #v=>Fls)}";
+Goalw [hyps_def] "hyps(#v,t) = {if(v:t, #v, #v=>Fls)}";
 by (Simp_tac 1);
 qed "hyps_Var";
 
-goalw PropLog.thy [hyps_def] "hyps(p=>q,t) = hyps(p,t) Un hyps(q,t)";
+Goalw [hyps_def] "hyps(p=>q,t) = hyps(p,t) Un hyps(q,t)";
 by (Simp_tac 1);
 qed "hyps_Imp";
 
@@ -74,7 +74,7 @@
 
 (*** Proof theory of propositional logic ***)
 
-goalw PropLog.thy thms.defs "!!G H. G<=H ==> thms(G) <= thms(H)";
+Goalw thms.defs "!!G H. G<=H ==> thms(G) <= thms(H)";
 by (rtac lfp_mono 1);
 by (REPEAT (rtac thms.bnd_mono 1));
 by (REPEAT (ares_tac (univ_mono::basic_monos) 1));
@@ -85,13 +85,13 @@
 val ImpE = prop.mk_cases prop.con_defs "p=>q : prop";
 
 (*Stronger Modus Ponens rule: no typechecking!*)
-goal PropLog.thy "!!p q H. [| H |- p=>q;  H |- p |] ==> H |- q";
+Goal "!!p q H. [| H |- p=>q;  H |- p |] ==> H |- q";
 by (rtac thms.MP 1);
 by (REPEAT (eresolve_tac [asm_rl, thms_in_pl, thms_in_pl RS ImpE] 1));
 qed "thms_MP";
 
 (*Rule is called I for Identity Combinator, not for Introduction*)
-goal PropLog.thy "!!p H. p:prop ==> H |- p=>p";
+Goal "!!p H. p:prop ==> H |- p=>p";
 by (rtac (thms.S RS thms_MP RS thms_MP) 1);
 by (rtac thms.K 5);
 by (rtac thms.K 4);
@@ -109,13 +109,13 @@
 val weaken_left_Un1  = Un_upper1 RS weaken_left;
 val weaken_left_Un2  = Un_upper2 RS weaken_left;
 
-goal PropLog.thy "!!H p q. [| H |- q;  p:prop |] ==> H |- p=>q";
+Goal "!!H p q. [| H |- q;  p:prop |] ==> H |- p=>q";
 by (rtac (thms.K RS thms_MP) 1);
 by (REPEAT (ares_tac [thms_in_pl] 1));
 qed "weaken_right";
 
 (*The deduction theorem*)
-goal PropLog.thy "!!p q H. [| cons(p,H) |- q;  p:prop |] ==>  H |- p=>q";
+Goal "!!p q H. [| cons(p,H) |- q;  p:prop |] ==>  H |- p=>q";
 by (etac thms.induct 1);
 by (fast_tac (claset() addIs [thms_I, thms.H RS weaken_right]) 1);
 by (fast_tac (claset() addIs [thms.K RS weaken_right]) 1);
@@ -126,12 +126,12 @@
 
 
 (*The cut rule*)
-goal PropLog.thy "!!H p q. [| H|-p;  cons(p,H) |- q |] ==>  H |- q";
+Goal "!!H p q. [| H|-p;  cons(p,H) |- q |] ==>  H |- q";
 by (rtac (deduction RS thms_MP) 1);
 by (REPEAT (ares_tac [thms_in_pl] 1));
 qed "cut";
 
-goal PropLog.thy "!!H p. [| H |- Fls; p:prop |] ==> H |- p";
+Goal "!!H p. [| H |- Fls; p:prop |] ==> H |- p";
 by (rtac (thms.DN RS thms_MP) 1);
 by (rtac weaken_right 2);
 by (REPEAT (ares_tac (prop.intrs@[consI1]) 1));
@@ -141,7 +141,7 @@
 bind_thm ("thms_notE", (thms_MP RS thms_FlsE));
 
 (*Soundness of the rules wrt truth-table semantics*)
-goalw PropLog.thy [logcon_def] "!!H. H |- p ==> H |= p";
+Goalw [logcon_def] "!!H. H |- p ==> H |= p";
 by (etac thms.induct 1);
 by (fast_tac (claset() addSDs [is_true_Imp RS iffD1 RS mp]) 5);
 by (ALLGOALS Asm_simp_tac);
@@ -289,12 +289,12 @@
 qed "completeness_0";
 
 (*A semantic analogue of the Deduction Theorem*)
-goalw PropLog.thy [logcon_def] "!!H p q. [| cons(p,H) |= q |] ==> H |= p=>q";
+Goalw [logcon_def] "!!H p q. [| cons(p,H) |= q |] ==> H |= p=>q";
 by (Simp_tac 1);
 by (Fast_tac 1);
 qed "logcon_Imp";
 
-goal PropLog.thy "!!H. H: Fin(prop) ==> ALL p:prop. H |= p --> H |- p";
+Goal "!!H. H: Fin(prop) ==> ALL p:prop. H |= p --> H |- p";
 by (etac Fin_induct 1);
 by (safe_tac (claset() addSIs [completeness_0]));
 by (rtac (weaken_left_cons RS thms_MP) 1);
--- a/src/ZF/ex/Ramsey.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/ex/Ramsey.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -22,16 +22,16 @@
 
 (*** Cliques and Independent sets ***)
 
-goalw Ramsey.thy [Clique_def] "Clique(0,V,E)";
+Goalw [Clique_def] "Clique(0,V,E)";
 by (Fast_tac 1);
 qed "Clique0";
 
-goalw Ramsey.thy [Clique_def]
+Goalw [Clique_def]
     "!!C V E. [| Clique(C,V',E);  V'<=V |] ==> Clique(C,V,E)";
 by (Fast_tac 1);
 qed "Clique_superset";
 
-goalw Ramsey.thy [Indept_def] "Indept(0,V,E)";
+Goalw [Indept_def] "Indept(0,V,E)";
 by (Fast_tac 1);
 qed "Indept0";
 
@@ -42,21 +42,21 @@
 
 (*** Atleast ***)
 
-goalw Ramsey.thy [Atleast_def, inj_def, Pi_def, function_def] "Atleast(0,A)";
+Goalw [Atleast_def, inj_def, Pi_def, function_def] "Atleast(0,A)";
 by (Fast_tac 1);
 qed "Atleast0";
 
-goalw Ramsey.thy [Atleast_def]
+Goalw [Atleast_def]
     "!!m A. Atleast(succ(m),A) ==> EX x:A. Atleast(m, A-{x})";
 by (fast_tac (claset() addEs [inj_is_fun RS apply_type, inj_succ_restrict]) 1);
 qed "Atleast_succD";
 
-goalw Ramsey.thy [Atleast_def]
+Goalw [Atleast_def]
     "!!n A. [| Atleast(n,A);  A<=B |] ==> Atleast(n,B)";
 by (fast_tac (claset() addEs [inj_weaken_type]) 1);
 qed "Atleast_superset";
 
-goalw Ramsey.thy [Atleast_def,succ_def]
+Goalw [Atleast_def,succ_def]
     "!!m. [| Atleast(m,B);  b~: B |] ==> Atleast(succ(m), cons(b,B))";
 by (etac exE 1);
 by (rtac exI 1);
@@ -65,7 +65,7 @@
 by (assume_tac 1);
 qed "Atleast_succI";
 
-goal Ramsey.thy
+Goal
     "!!m. [| Atleast(m, B-{x});  x: B |] ==> Atleast(succ(m), B)";
 by (etac (Atleast_succI RS Atleast_superset) 1);
 by (Fast_tac 1);
@@ -119,11 +119,11 @@
 
 (** Base cases of induction; they now admit ANY Ramsey number **)
 
-goalw Ramsey.thy [Ramsey_def] "Ramsey(n,0,j)";
+Goalw [Ramsey_def] "Ramsey(n,0,j)";
 by (fast_tac (claset() addIs [Clique0,Atleast0]) 1);
 qed "Ramsey0j";
 
-goalw Ramsey.thy [Ramsey_def] "Ramsey(n,i,0)";
+Goalw [Ramsey_def] "Ramsey(n,i,0)";
 by (fast_tac (claset() addIs [Indept0,Atleast0]) 1);
 qed "Ramseyi0";
 
@@ -201,7 +201,7 @@
 qed "ramsey_lemma";
 
 (*Final statement in a tidy form, without succ(...) *)
-goal Ramsey.thy "!!i j. [| i: nat;  j: nat |] ==> EX n:nat. Ramsey(n,i,j)";
+Goal "!!i j. [| i: nat;  j: nat |] ==> EX n:nat. Ramsey(n,i,j)";
 by (best_tac (claset() addDs [ramsey_lemma] addSIs [nat_succI]) 1);
 qed "ramsey";
 
--- a/src/ZF/ex/Rmap.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/ex/Rmap.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -8,7 +8,7 @@
 
 open Rmap;
 
-goalw Rmap.thy rmap.defs "!!r s. r<=s ==> rmap(r) <= rmap(s)";
+Goalw rmap.defs "!!r s. r<=s ==> rmap(r) <= rmap(s)";
 by (rtac lfp_mono 1);
 by (REPEAT (rtac rmap.bnd_mono 1));
 by (REPEAT (ares_tac ([Sigma_mono, list_mono, domain_mono, range_mono] @ 
@@ -21,20 +21,20 @@
 AddIs  rmap.intrs;
 AddSEs [Nil_rmap_case, Cons_rmap_case];
 
-goal Rmap.thy "!!r. r <= A*B ==> rmap(r) <= list(A)*list(B)";
+Goal "!!r. r <= A*B ==> rmap(r) <= list(A)*list(B)";
 by (rtac (rmap.dom_subset RS subset_trans) 1);
 by (REPEAT (ares_tac [domain_rel_subset, range_rel_subset,
                       Sigma_mono, list_mono] 1));
 qed "rmap_rel_type";
 
-goal Rmap.thy
+Goal
     "!!r. A <= domain(r) ==> list(A) <= domain(rmap(r))";
 by (rtac subsetI 1);
 by (etac list.induct 1);
 by (ALLGOALS Fast_tac);
 qed "rmap_total";
 
-goalw Rmap.thy [function_def] "!!r. function(r) ==> function(rmap(r))";
+Goalw [function_def] "!!r. function(r) ==> function(rmap(r))";
 by (rtac (impI RS allI RS allI) 1);
 by (etac rmap.induct 1);
 by (ALLGOALS Fast_tac);
@@ -43,16 +43,16 @@
 
 (** If f is a function then rmap(f) behaves as expected. **)
 
-goal Rmap.thy "!!f. f: A->B ==> rmap(f): list(A)->list(B)";
+Goal "!!f. f: A->B ==> rmap(f): list(A)->list(B)";
 by (asm_full_simp_tac 
     (simpset() addsimps [Pi_iff, rmap_rel_type, rmap_functional, rmap_total]) 1);
 qed "rmap_fun_type";
 
-goalw Rmap.thy [apply_def] "rmap(f)`Nil = Nil";
+Goalw [apply_def] "rmap(f)`Nil = Nil";
 by (fast_tac (claset() addIs [the_equality]) 1);
 qed "rmap_Nil";
 
-goal Rmap.thy "!!f. [| f: A->B;  x: A;  xs: list(A) |] ==> \
+Goal "!!f. [| f: A->B;  x: A;  xs: list(A) |] ==> \
 \                   rmap(f) ` Cons(x,xs) = Cons(f`x, rmap(f)`xs)";
 by (rtac apply_equality 1);
 by (REPEAT (ares_tac ([apply_Pair, rmap_fun_type] @ rmap.intrs) 1));
--- a/src/ZF/ex/TF.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/ex/TF.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -19,22 +19,22 @@
 
 val [_, tree_def, forest_def] = tree_forest.defs;
 
-goalw TF.thy [tree_def] "tree(A) <= tree_forest(A)";
+Goalw [tree_def] "tree(A) <= tree_forest(A)";
 by (rtac Part_subset 1);
 qed "tree_subset_TF";
 
-goalw TF.thy [forest_def] "forest(A) <= tree_forest(A)";
+Goalw [forest_def] "forest(A) <= tree_forest(A)";
 by (rtac Part_subset 1);
 qed "forest_subset_TF";
 
-goal TF.thy "tree(A) Un forest(A) = tree_forest(A)";
+Goal "tree(A) Un forest(A) = tree_forest(A)";
 by (safe_tac (subset_cs addSIs [equalityI, tree_subset_TF, forest_subset_TF]));
 by (fast_tac (claset() addSIs tree_forest.intrs addEs [tree_forest.elim]) 1);
 qed "TF_equals_Un";
 
 (** NOT useful, but interesting... **)
 
-goalw TF.thy [tree_def, forest_def] 
+Goalw [tree_def, forest_def] 
     "tree_forest(A) = (A*forest(A)) + ({0} + tree(A)*forest(A))";
 let open tree_forest;  
     val rew = rewrite_rule (con_defs @ tl defs) in  
@@ -45,13 +45,13 @@
 val tree_forest_unfold' = rewrite_rule [tree_def, forest_def] 
                           tree_forest_unfold;
 
-goalw TF.thy [tree_def, forest_def]
+Goalw [tree_def, forest_def]
     "tree(A) = {Inl(x). x: A*forest(A)}";
 by (rtac (Part_Inl RS subst) 1);
 by (rtac (tree_forest_unfold' RS subst_context) 1);
 qed "tree_unfold";
 
-goalw TF.thy [tree_def, forest_def]
+Goalw [tree_def, forest_def]
     "forest(A) = {Inr(x). x: {0} + tree(A)*forest(A)}";
 by (rtac (Part_Inr RS subst) 1);
 by (rtac (tree_forest_unfold' RS subst_context) 1);
@@ -62,19 +62,19 @@
 
 (** conversion rules **)
 
-goal TF.thy "TF_rec(Tcons(a,f), b, c, d) = b(a, f, TF_rec(f,b,c,d))";
+Goal "TF_rec(Tcons(a,f), b, c, d) = b(a, f, TF_rec(f,b,c,d))";
 by (rtac (TF_rec_def RS def_Vrec RS trans) 1);
 by (rewrite_goals_tac tree_forest.con_defs);
 by (simp_tac rank_ss 1);
 qed "TF_rec_Tcons";
 
-goal TF.thy "TF_rec(Fnil, b, c, d) = c";
+Goal "TF_rec(Fnil, b, c, d) = c";
 by (rtac (TF_rec_def RS def_Vrec RS trans) 1);
 by (rewrite_goals_tac tree_forest.con_defs);
 by (Simp_tac 1);
 qed "TF_rec_Fnil";
 
-goal TF.thy "TF_rec(Fcons(t,f), b, c, d) = \
+Goal "TF_rec(Fcons(t,f), b, c, d) = \
 \      d(t, f, TF_rec(t, b, c, d), TF_rec(f, b, c, d))";
 by (rtac (TF_rec_def RS def_Vrec RS trans) 1);
 by (rewrite_goals_tac tree_forest.con_defs);
@@ -142,7 +142,7 @@
         TF_recs list_of_TF_def;
 Addsimps [list_of_TF_Tcons, list_of_TF_Fnil, list_of_TF_Fcons];
 
-goalw TF.thy [list_of_TF_def]
+Goalw [list_of_TF_def]
     "!!z A. z: tree_forest(A) ==> list_of_TF(z) : list(tree(A))";
 by (REPEAT (ares_tac ([TF_rec_type,TconsI] @ list_typechecks) 1));
 qed "list_of_TF_type";
@@ -150,7 +150,7 @@
 val [TF_of_list_Nil,TF_of_list_Cons] = list_recs TF_of_list_def;
 Addsimps [TF_of_list_Nil,TF_of_list_Cons];
 
-goalw TF.thy [TF_of_list_def] 
+Goalw [TF_of_list_def] 
     "!!l A. l: list(tree(A)) ==> TF_of_list(l) : forest(A)";
 by (REPEAT (ares_tac [list_rec_type, FnilI, FconsI] 1));
 qed "TF_of_list_type";
@@ -175,7 +175,7 @@
 val [TF_size_Tcons, TF_size_Fnil, TF_size_Fcons] = TF_recs TF_size_def;
 Addsimps [TF_size_Tcons, TF_size_Fnil, TF_size_Fcons];
 
-goalw TF.thy [TF_size_def]
+Goalw [TF_size_def]
     "!!z A. z: tree_forest(A) ==> TF_size(z) : nat";
 by (REPEAT (ares_tac [TF_rec_type, add_type, nat_0I, nat_succI] 1));
 qed "TF_size_type";
@@ -187,7 +187,7 @@
         TF_recs TF_preorder_def;
 Addsimps [TF_preorder_Tcons, TF_preorder_Fnil, TF_preorder_Fcons];
 
-goalw TF.thy [TF_preorder_def]
+Goalw [TF_preorder_def]
     "!!z A. z: tree_forest(A) ==> TF_preorder(z) : list(A)";
 by (REPEAT (ares_tac ([TF_rec_type, app_type] @ list.intrs) 1));
 qed "TF_preorder_type";
@@ -216,12 +216,12 @@
 by (REPEAT (ares_tac (TrueI::prems) 1));
 qed "forest_induct";
 
-goal TF.thy "!!f A. f: forest(A) ==> TF_of_list(list_of_TF(f)) = f";
+Goal "!!f A. f: forest(A) ==> TF_of_list(list_of_TF(f)) = f";
 by (etac forest_induct 1);
 by (ALLGOALS Asm_simp_tac);
 qed "forest_iso";
 
-goal TF.thy
+Goal
     "!!ts. ts: list(tree(A)) ==> list_of_TF(TF_of_list(ts)) = ts";
 by (etac list.induct 1);
 by (ALLGOALS Asm_simp_tac);
@@ -229,12 +229,12 @@
 
 (** theorems about TF_map **)
 
-goal TF.thy "!!z A. z: tree_forest(A) ==> TF_map(%u. u, z) = z";
+Goal "!!z A. z: tree_forest(A) ==> TF_map(%u. u, z) = z";
 by (etac tree_forest.induct 1);
 by (ALLGOALS Asm_simp_tac);
 qed "TF_map_ident";
 
-goal TF.thy
+Goal
  "!!z A. z: tree_forest(A) ==> TF_map(h, TF_map(j,z)) = TF_map(%u. h(j(u)), z)";
 by (etac tree_forest.induct 1);
 by (ALLGOALS Asm_simp_tac);
@@ -242,13 +242,13 @@
 
 (** theorems about TF_size **)
 
-goal TF.thy
+Goal
     "!!z A. z: tree_forest(A) ==> TF_size(TF_map(h,z)) = TF_size(z)";
 by (etac tree_forest.induct 1);
 by (ALLGOALS Asm_simp_tac);
 qed "TF_size_TF_map";
 
-goal TF.thy
+Goal
     "!!z A. z: tree_forest(A) ==> TF_size(z) = length(TF_preorder(z))";
 by (etac tree_forest.induct 1);
 by (ALLGOALS (asm_simp_tac (simpset() addsimps [length_app])));
@@ -256,7 +256,7 @@
 
 (** theorems about TF_preorder **)
 
-goal TF.thy "!!z A. z: tree_forest(A) ==> \
+Goal "!!z A. z: tree_forest(A) ==> \
 \                      TF_preorder(TF_map(h,z)) = map(h, TF_preorder(z))";
 by (etac tree_forest.induct 1);
 by (ALLGOALS (asm_simp_tac (simpset() addsimps [map_app_distrib])));
--- a/src/ZF/ex/Term.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/ex/Term.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -9,7 +9,7 @@
 
 open Term;
 
-goal Term.thy "term(A) = A * list(term(A))";
+Goal "term(A) = A * list(term(A))";
 let open term;  val rew = rewrite_rule con_defs in  
 by (fast_tac (claset() addSIs (map rew intrs) addEs [rew elim]) 1)
 end;
@@ -41,14 +41,14 @@
 
 (**  Lemmas to justify using "term" in other recursive type definitions **)
 
-goalw Term.thy term.defs "!!A B. A<=B ==> term(A) <= term(B)";
+Goalw term.defs "!!A B. A<=B ==> term(A) <= term(B)";
 by (rtac lfp_mono 1);
 by (REPEAT (rtac term.bnd_mono 1));
 by (REPEAT (ares_tac (univ_mono::basic_monos) 1));
 qed "term_mono";
 
 (*Easily provable by induction also*)
-goalw Term.thy (term.defs@term.con_defs) "term(univ(A)) <= univ(A)";
+Goalw (term.defs@term.con_defs) "term(univ(A)) <= univ(A)";
 by (rtac lfp_lowerbound 1);
 by (rtac (A_subset_univ RS univ_mono) 2);
 by Safe_tac;
@@ -58,7 +58,7 @@
 val term_subset_univ = 
     term_mono RS (term_univ RSN (2,subset_trans)) |> standard;
 
-goal Term.thy "!!t A B. [| t: term(A);  A <= univ(B) |] ==> t: univ(B)";
+Goal "!!t A B. [| t: term(A);  A <= univ(B) |] ==> t: univ(B)";
 by (REPEAT (ares_tac [term_subset_univ RS subsetD] 1));
 qed "term_into_univ";
 
@@ -140,7 +140,7 @@
 
 bind_thm ("term_size", (term_size_def RS def_term_rec));
 
-goalw Term.thy [term_size_def] "!!t A. t: term(A) ==> term_size(t) : nat";
+Goalw [term_size_def] "!!t A. t: term(A) ==> term_size(t) : nat";
 by (REPEAT (ares_tac [term_rec_simple_type, list_add_type, nat_succI] 1));
 qed "term_size_type";
 
@@ -149,7 +149,7 @@
 
 bind_thm ("reflect", (reflect_def RS def_term_rec));
 
-goalw Term.thy [reflect_def] "!!t A. t: term(A) ==> reflect(t) : term(A)";
+Goalw [reflect_def] "!!t A. t: term(A) ==> reflect(t) : term(A)";
 by (REPEAT (ares_tac [term_rec_simple_type, rev_type, term.Apply_I] 1));
 qed "reflect_type";
 
@@ -157,7 +157,7 @@
 
 bind_thm ("preorder", (preorder_def RS def_term_rec));
 
-goalw Term.thy [preorder_def]
+Goalw [preorder_def]
     "!!t A. t: term(A) ==> preorder(t) : list(A)";
 by (REPEAT (ares_tac [term_rec_simple_type, list.Cons_I, flat_type] 1));
 qed "preorder_type";
@@ -177,18 +177,18 @@
 
 (** theorems about term_map **)
 
-goal Term.thy "!!t A. t: term(A) ==> term_map(%u. u, t) = t";
+Goal "!!t A. t: term(A) ==> term_map(%u. u, t) = t";
 by (etac term_induct_eqn 1);
 by (asm_simp_tac (simpset() addsimps [map_ident]) 1);
 qed "term_map_ident";
 
-goal Term.thy
+Goal
   "!!t A. t: term(A) ==> term_map(f, term_map(g,t)) = term_map(%u. f(g(u)), t)";
 by (etac term_induct_eqn 1);
 by (asm_simp_tac (simpset() addsimps [map_compose]) 1);
 qed "term_map_compose";
 
-goal Term.thy
+Goal
     "!!t A. t: term(A) ==> term_map(f, reflect(t)) = reflect(term_map(f,t))";
 by (etac term_induct_eqn 1);
 by (asm_simp_tac (simpset() addsimps [rev_map_distrib RS sym, map_compose]) 1);
@@ -197,19 +197,19 @@
 
 (** theorems about term_size **)
 
-goal Term.thy
+Goal
     "!!t A. t: term(A) ==> term_size(term_map(f,t)) = term_size(t)";
 by (etac term_induct_eqn 1);
 by (asm_simp_tac (simpset() addsimps [map_compose]) 1);
 qed "term_size_term_map";
 
-goal Term.thy "!!t A. t: term(A) ==> term_size(reflect(t)) = term_size(t)";
+Goal "!!t A. t: term(A) ==> term_size(reflect(t)) = term_size(t)";
 by (etac term_induct_eqn 1);
 by (asm_simp_tac (simpset() addsimps [rev_map_distrib RS sym, map_compose,
                                     list_add_rev]) 1);
 qed "term_size_reflect";
 
-goal Term.thy "!!t A. t: term(A) ==> term_size(t) = length(preorder(t))";
+Goal "!!t A. t: term(A) ==> term_size(t) = length(preorder(t))";
 by (etac term_induct_eqn 1);
 by (asm_simp_tac (simpset() addsimps [length_flat, map_compose]) 1);
 qed "term_size_length";
@@ -217,7 +217,7 @@
 
 (** theorems about reflect **)
 
-goal Term.thy "!!t A. t: term(A) ==> reflect(reflect(t)) = t";
+Goal "!!t A. t: term(A) ==> reflect(reflect(t)) = t";
 by (etac term_induct_eqn 1);
 by (asm_simp_tac (simpset() addsimps [rev_map_distrib, map_compose,
                                     map_ident, rev_rev_ident]) 1);
@@ -226,7 +226,7 @@
 
 (** theorems about preorder **)
 
-goal Term.thy
+Goal
     "!!t A. t: term(A) ==> preorder(term_map(f,t)) = map(f, preorder(t))";
 by (etac term_induct_eqn 1);
 by (asm_simp_tac (simpset() addsimps [map_compose, map_flat]) 1);
--- a/src/ZF/ex/misc.ML	Mon Jun 22 17:12:27 1998 +0200
+++ b/src/ZF/ex/misc.ML	Mon Jun 22 17:13:09 1998 +0200
@@ -25,7 +25,7 @@
  W. Bledsoe.  A Maximal Method for Set Variables in Automatic Theorem-proving.
  In: J. Hayes and D. Michie and L. Mikulich, eds.  Machine Intelligence 9.
  Ellis Horwood, 53-100 (1979). *)
-goal thy "(ALL F. {x}: F --> {y}:F) --> (ALL A. x:A --> y:A)";
+Goal "(ALL F. {x}: F --> {y}:F) --> (ALL A. x:A --> y:A)";
 by (Best_tac 1);
 result();
 
@@ -69,7 +69,7 @@
 
 (** A characterization of functions, suggested by Tobias Nipkow **)
 
-goalw thy [Pi_def, function_def]
+Goalw [Pi_def, function_def]
     "r: domain(r)->B  <->  r <= domain(r)*B & (ALL X. r `` (r -`` X) <= X)";
 by (Best_tac 1);
 result();