expandshort
authorpaulson
Thu, 28 Oct 1999 14:55:23 +0200
changeset 7961 422ac6888c7f
parent 7960 d5c91c131070
child 7962 d139b03fcac2
expandshort
src/HOL/BCV/DFAandWTI.ML
src/HOL/BCV/DFAimpl.ML
src/HOL/BCV/Fixpoint.ML
src/HOL/BCV/Machine.ML
src/HOL/BCV/Orders.ML
src/HOL/BCV/Orders0.ML
src/HOL/BCV/Plus.ML
src/HOL/BCV/SemiLattice.ML
src/HOL/BCV/Types.ML
src/HOL/BCV/Types0.ML
src/HOL/Map.ML
--- a/src/HOL/BCV/DFAandWTI.ML	Thu Oct 28 14:00:25 1999 +0200
+++ b/src/HOL/BCV/DFAandWTI.ML	Thu Oct 28 14:55:23 1999 +0200
@@ -8,39 +8,39 @@
  "[| wti_is_fix_step step wti succs n L; \
 \    sos : listsn n (option L); p < n |] ==> \
 \ wti p sos = stable step succs p sos";
-by(Asm_simp_tac 1);
+by (Asm_simp_tac 1);
 qed "wti_is_fix_stepD";
 
 Goalw [is_dfa_def]
  "[| is_dfa dfa step succs n L; sos : listsn n (option L) |] ==> \
 \ dfa sos = (? tos : listsn (size sos) (option L). \
 \             sos <= tos & (!p < size tos. stable step succs p tos))";
-by(Asm_simp_tac 1);
+by (Asm_simp_tac 1);
 qed "is_dfaD";
 
 Goalw [welltyping_def]
  "[| wti_is_fix_step step wti succs n L; is_dfa dfa step succs n L; \
 \    sos : listsn n (option L) |] ==> \
 \ dfa sos = (? tos : listsn n (option L). sos <= tos & welltyping wti tos)";
-bd is_dfaD 1;
-ba 1;
-by(Asm_full_simp_tac 1);
-by(blast_tac (claset() addDs [wti_is_fix_stepD]) 1);
+by (dtac is_dfaD 1);
+by (assume_tac 1);
+by (Asm_full_simp_tac 1);
+by (blast_tac (claset() addDs [wti_is_fix_stepD]) 1);
 qed "dfa_iff_welltyping";
 
 Goal
  "x:M ==> replicate n x : listsn n M";
-by(induct_tac "n" 1);
-by(Simp_tac 1);
-by(asm_simp_tac (simpset() addsimps [in_listsn_Suc_iff,Bex_def]) 1);
+by (induct_tac "n" 1);
+by (Simp_tac 1);
+by (asm_simp_tac (simpset() addsimps [in_listsn_Suc_iff,Bex_def]) 1);
 qed "replicate_in_listsn";
 Addsimps [replicate_in_listsn];
 
 Goal "!ys. replicate n x <= ys = (length ys = n & (!y:set ys. x <= y))";
-by(induct_tac "n" 1);
- by(Force_tac 1);
-by(asm_full_simp_tac (simpset() addsimps [Cons_le_iff,length_Suc_conv]) 1);
-by(Force_tac 1);
+by (induct_tac "n" 1);
+ by (Force_tac 1);
+by (asm_full_simp_tac (simpset() addsimps [Cons_le_iff,length_Suc_conv]) 1);
+by (Force_tac 1);
 qed_spec_mp "replicate_le_conv";
 AddIffs [replicate_le_conv];
 
@@ -49,31 +49,31 @@
 \    0 < n; init : (option L) |] ==> \
 \ dfa (init # replicate (n-1) None) = \
 \ (? tos : listsn n (option L). init <= tos!0 & welltyping wti tos)";
-by(subgoal_tac "? m. n = m+1" 1);
- by(res_inst_tac [("x","n-1")] exI 2);
- by(arith_tac 2);
-by(Clarify_tac 1);
-by(Asm_full_simp_tac 1);
-bd dfa_iff_welltyping 1;
-  ba 1;
- be trans 2;
- by(asm_simp_tac (simpset() addsimps [in_listsn_Suc_iff,Bex_def]) 1);
-by(asm_full_simp_tac (simpset() addsimps [Cons_le_iff]) 1);
-br iffI 1;
- by(Clarify_tac 1);
- br bexI 1;
-  br conjI 1;
-   ba 2;
-  ba 2;
- by(Asm_simp_tac 1);
-by(Clarify_tac 1);
-by(asm_full_simp_tac (simpset() addsimps [in_listsn_Suc_iff,Bex_def]) 1);
-by(Clarify_tac 1);
-by(asm_full_simp_tac (simpset() addsimps []) 1);
-br exI 1;
-br conjI 1;
- br conjI 2;
-  ba 3;
- by(Blast_tac 1);
-by(Force_tac 1);
+by (subgoal_tac "? m. n = m+1" 1);
+ by (res_inst_tac [("x","n-1")] exI 2);
+ by (arith_tac 2);
+by (Clarify_tac 1);
+by (Asm_full_simp_tac 1);
+by (dtac dfa_iff_welltyping 1);
+  by (assume_tac 1);
+ by (etac trans 2);
+ by (asm_simp_tac (simpset() addsimps [in_listsn_Suc_iff,Bex_def]) 1);
+by (asm_full_simp_tac (simpset() addsimps [Cons_le_iff]) 1);
+by (rtac iffI 1);
+ by (Clarify_tac 1);
+ by (rtac bexI 1);
+  by (rtac conjI 1);
+   by (assume_tac 2);
+  by (assume_tac 2);
+ by (Asm_simp_tac 1);
+by (Clarify_tac 1);
+by (asm_full_simp_tac (simpset() addsimps [in_listsn_Suc_iff,Bex_def]) 1);
+by (Clarify_tac 1);
+by (asm_full_simp_tac (simpset() addsimps []) 1);
+by (rtac exI 1);
+by (rtac conjI 1);
+ by (rtac conjI 2);
+  by (assume_tac 3);
+ by (Blast_tac 1);
+by (Force_tac 1);
 qed "dfa_iff_welltyping0";
--- a/src/HOL/BCV/DFAimpl.ML	Thu Oct 28 14:00:25 1999 +0200
+++ b/src/HOL/BCV/DFAimpl.ML	Thu Oct 28 14:55:23 1999 +0200
@@ -7,40 +7,40 @@
 (** merges **)
 
 Goal "!sos. size(merges t ps sos) = size sos";
-by(induct_tac "ps" 1);
-by(Auto_tac);
+by (induct_tac "ps" 1);
+by (Auto_tac);
 qed_spec_mp "length_merges";
 Addsimps [length_merges];
 
 Goal
  "!xs. xs : listsn n (option A) --> x : A --> (!p : set ps. p<n) --> \
 \      semilat A --> merges x ps xs : listsn n (option A)";
-by(induct_tac "ps" 1);
-by(Auto_tac);
+by (induct_tac "ps" 1);
+by (Auto_tac);
 qed_spec_mp "merges_preserves_type";
 Addsimps [merges_preserves_type];
 (*AddSIs [merges_preserves_type];*)
 
 Goal "semilat A ==> !xs. xs : listsn n (option A) --> x:A --> (!p:set ps. p<n) \
 \ --> xs <= merges x ps xs";
-by(induct_tac "ps" 1);
- by(Simp_tac 1);
-by(Simp_tac 1);
-by(Clarify_tac 1);
-br order_trans 1;
- be list_update_incr 1;
-   by(Blast_tac 2);
-  ba 2;
- be (Some_in_option RS iffD2) 1;
-by(blast_tac (claset() addSIs [listsnE_set]
+by (induct_tac "ps" 1);
+ by (Simp_tac 1);
+by (Simp_tac 1);
+by (Clarify_tac 1);
+by (rtac order_trans 1);
+ by (etac list_update_incr 1);
+   by (Blast_tac 2);
+  by (assume_tac 2);
+ by (etac (Some_in_option RS iffD2) 1);
+by (blast_tac (claset() addSIs [listsnE_set]
           addIs [semilat_plus,listsnE_length RS nth_in]) 1);
 qed_spec_mp "merges_incr";
 
 (* is redundant but useful *)
 Goal "[| xs!i = Some x; xs : listsn n (option A); i < n |] ==> x:A";
-bd listsnE_nth_in 1;
-ba 1;
-by(Asm_full_simp_tac 1);
+by (dtac listsnE_nth_in 1);
+by (assume_tac 1);
+by (Asm_full_simp_tac 1);
 qed "listsn_optionE_in";
 (*Addsimps [listsn_optionE_in];*)
 
@@ -48,29 +48,29 @@
  "[| semilat L |] ==> \
 \ (!xs. xs : listsn n (option L) --> x:L --> (!p:set ps. p<n) --> \
 \      (merges x ps xs = xs) = (!p:set ps. Some x <= xs!p))";
-by(induct_tac "ps" 1);
- by(Asm_simp_tac 1);
-by(Clarsimp_tac 1);
-by(rename_tac "p ps xs" 1);
-br iffI 1;
- br context_conjI 1;
-  by(subgoal_tac "xs[p := Some x + xs!p] <= xs" 1);
-   by(EVERY[etac subst 2, rtac merges_incr 2]);
-      ba 2;
-     by(Force_tac 2);
-    ba 2;
-   ba 2;
-  by(exhaust_tac "xs!p" 1);
-   by(asm_full_simp_tac (simpset() addsimps [list_update_le_conv]) 1);
-  by(asm_full_simp_tac (simpset() addsimps
+by (induct_tac "ps" 1);
+ by (Asm_simp_tac 1);
+by (Clarsimp_tac 1);
+by (rename_tac "p ps xs" 1);
+by (rtac iffI 1);
+ by (rtac context_conjI 1);
+  by (subgoal_tac "xs[p := Some x + xs!p] <= xs" 1);
+   by (EVERY[etac subst 2, rtac merges_incr 2]);
+      by (assume_tac 2);
+     by (Force_tac 2);
+    by (assume_tac 2);
+   by (assume_tac 2);
+  by (exhaust_tac "xs!p" 1);
+   by (asm_full_simp_tac (simpset() addsimps [list_update_le_conv]) 1);
+  by (asm_full_simp_tac (simpset() addsimps
       [list_update_le_conv,listsn_optionE_in]) 1);
- by(Clarify_tac 1);
- by(rotate_tac ~3 1);
- by(asm_full_simp_tac (simpset() addsimps
+ by (Clarify_tac 1);
+ by (rotate_tac ~3 1);
+ by (asm_full_simp_tac (simpset() addsimps
       [le_iff_plus_unchanged RS iffD1,listsn_optionE_in,
        list_update_same_conv RS iffD2]) 1);
-by(Clarify_tac 1);
-by(asm_simp_tac (simpset() addsimps
+by (Clarify_tac 1);
+by (asm_simp_tac (simpset() addsimps
       [le_iff_plus_unchanged RS iffD1,listsn_optionE_in,
        list_update_same_conv RS iffD2]) 1);
 qed_spec_mp "merges_same_conv";
@@ -80,11 +80,11 @@
  "xs <= ys --> xs : listsn n (option L) --> ys : listsn n (option L) --> \
 \ p < n --> ys!p = Some y --> x <= y --> x : L --> semilat L --> \
 \ xs[p := Some x + xs!p] <= ys";
-by(simp_tac (simpset()  addsimps [nth_list_update]
+by (simp_tac (simpset()  addsimps [nth_list_update]
                         addsplits [option.split]) 1);
-by(Clarify_tac 1);
-by(rotate_tac 3 1);
-by(force_tac (claset() addEs [listsn_optionE_in] addIs [semilat_lub],
+by (Clarify_tac 1);
+by (rotate_tac 3 1);
+by (force_tac (claset() addEs [listsn_optionE_in] addIs [semilat_lub],
               simpset()) 1);
 qed_spec_mp "list_update_le_listI";
 
@@ -94,9 +94,9 @@
 \    !p. p:set ps --> p<n |] ==> \
 \ set qs <= set ps  --> \
 \ (!sos. sos : listsn n (option L) & sos <= tos --> merges t qs sos <= tos)";
-by(induct_tac "qs" 1);
- by(Asm_simp_tac 1);
-by(force_tac (claset(),simpset() addsimps [list_update_le_listI]) 1);
+by (induct_tac "qs" 1);
+ by (Asm_simp_tac 1);
+by (force_tac (claset(),simpset() addsimps [list_update_le_listI]) 1);
 val lemma = result();
 
 Goal
@@ -105,7 +105,7 @@
 \    !p. p:set ps --> p<n; \
 \    sos <= tos; sos : listsn n (option L); tos : listsn n (option L) |] \
 \ ==> merges t ps sos <= tos";
-by(blast_tac (claset() addDs [lemma]) 1);
+by (blast_tac (claset() addDs [lemma]) 1);
 qed "merges_pres_le_ub";
 
 
@@ -115,9 +115,9 @@
  "[| is_next next; next step succs sos = None; succs_bounded succs n; \
 \    sos : listsn n S |] ==> \
 \ ? p<n. ? s. sos!p = Some s & step p s = None";
-by(subgoal_tac "n=size sos" 1);
-by(Blast_tac 1);
-by(Asm_simp_tac 1);
+by (subgoal_tac "n=size sos" 1);
+by (Blast_tac 1);
+by (Asm_simp_tac 1);
 qed "next_None";
 
 Goalw [is_next_def]
@@ -125,9 +125,9 @@
 \ next step succs sos = Some sos --> \
 \ (!p<n. !s. sos!p = Some s --> (? t. \
 \         step p s = Some(t) & merges t (succs p) sos = sos))";
-by(subgoal_tac "n=size sos" 1);
-by(Blast_tac 1);
-by(Asm_simp_tac 1);
+by (subgoal_tac "n=size sos" 1);
+by (Blast_tac 1);
+by (Asm_simp_tac 1);
 qed "next_Some1";
 
 Goalw [is_next_def]
@@ -135,9 +135,9 @@
 \    succs_bounded succs n; sos : listsn n S |] ==> \
 \ ? p<n. ? s. sos!p = Some s & (? t. \
 \     step p s = Some(t) & merges t (succs p) sos = sos')";
-by(subgoal_tac "n=size sos" 1);
-by(Blast_tac 1);
-by(Asm_simp_tac 1);
+by (subgoal_tac "n=size sos" 1);
+by (Blast_tac 1);
+by (Asm_simp_tac 1);
 qed "next_Some2";
 
 Goal
@@ -145,26 +145,26 @@
 \ (next step succs sos = Some sos) = \
 \ (!p<n. !s. sos!p = Some s --> (? t. \
 \         step p s = Some(t) & merges t (succs p) sos = sos))";
-br iffI 1;
- by(asm_simp_tac (simpset() addsimps [next_Some1]) 1);
-by(exhaust_tac "next step succs sos" 1);
- by(best_tac (claset() addSDs [next_None] addss simpset()) 1);
-by(rename_tac "sos'" 1);
-by(case_tac "sos' = sos" 1);
- by(Blast_tac 1);
-by(best_tac (claset() addSDs [next_Some2] addss simpset()) 1);
+by (rtac iffI 1);
+ by (asm_simp_tac (simpset() addsimps [next_Some1]) 1);
+by (exhaust_tac "next step succs sos" 1);
+ by (best_tac (claset() addSDs [next_None] addss simpset()) 1);
+by (rename_tac "sos'" 1);
+by (case_tac "sos' = sos" 1);
+ by (Blast_tac 1);
+by (best_tac (claset() addSDs [next_Some2] addss simpset()) 1);
 qed "next_Some1_eq";
 
 Addsimps [next_Some1_eq];
 
 Goalw [step_pres_type_def]
  "[| step_pres_type step n L; s:L; p<n; step p s = Some(t) |] ==> t:L";
-by(Blast_tac 1);
+by (Blast_tac 1);
 qed "step_pres_typeD";
 
 Goalw [succs_bounded_def]
  "[| succs_bounded succs n; p < n; q : set(succs p) |] ==> q < n";
-by(Blast_tac 1);
+by (Blast_tac 1);
 qed "succs_boundedD";
 
 Goal
@@ -172,21 +172,21 @@
 \    step_pres_type step n A; succs_bounded succs n;\
 \    sos : listsn n (option A) |] ==> \
 \ next step succs sos : option (listsn n (option A))";
-by(exhaust_tac "next step succs sos" 1);
- by(ALLGOALS Asm_simp_tac);
-by(rename_tac "sos'" 1);
-by(case_tac "sos' = sos" 1);
- by(Asm_simp_tac 1);
-by(blast_tac (claset() addIs [step_pres_typeD,succs_boundedD,listsn_optionE_in] addSIs [merges_preserves_type] addDs [next_Some2]) 1);
+by (exhaust_tac "next step succs sos" 1);
+ by (ALLGOALS Asm_simp_tac);
+by (rename_tac "sos'" 1);
+by (case_tac "sos' = sos" 1);
+ by (Asm_simp_tac 1);
+by (blast_tac (claset() addIs [step_pres_typeD,succs_boundedD,listsn_optionE_in] addSIs [merges_preserves_type] addDs [next_Some2]) 1);
 qed_spec_mp "next_preserves_type";
 
 Goal
  "[| is_next next; semilat A; \
 \    step_pres_type step n A; succs_bounded succs n; \
 \    next step succs xs = Some ys; xs : listsn n (option A) |] ==> xs <= ys";
-by(case_tac "ys = xs" 1);
- by(Asm_full_simp_tac 1);
-by(blast_tac (claset() addSIs [merges_incr] addIs [listsn_optionE_in]
+by (case_tac "ys = xs" 1);
+ by (Asm_full_simp_tac 1);
+by (blast_tac (claset() addSIs [merges_incr] addIs [listsn_optionE_in]
        addDs [step_pres_typeD,succs_boundedD,next_Some2]) 1);
 qed_spec_mp "next_incr";
 
@@ -196,13 +196,13 @@
 
 Goalw [is_next_def]
  "is_next (%step succs sos. itnext (size sos) step succs sos)";
-by(Clarify_tac 1);
-by(etac thin_rl 1);
-by(res_inst_tac [("n","length sos")] nat_induct 1);
- by(Asm_full_simp_tac 1);
-by(asm_full_simp_tac (simpset() addsimps [less_Suc_eq,Let_def]
+by (Clarify_tac 1);
+by (etac thin_rl 1);
+by (res_inst_tac [("n","length sos")] nat_induct 1);
+ by (Asm_full_simp_tac 1);
+by (asm_full_simp_tac (simpset() addsimps [less_Suc_eq,Let_def]
                                 addsplits [option.split])1);
-by(Blast_tac 1);
+by (Blast_tac 1);
 qed "is_next_itnext";
 
 Unify.trace_bound := !(fst lemma);
@@ -213,37 +213,37 @@
 Goalw [step_mono_None_def]
  "[| step_mono_None step n L; s : L; p < n; s <= t; step p s = None |] ==> \
 \ step p t = None";
-by(Blast_tac 1);
+by (Blast_tac 1);
 qed "step_mono_NoneD";
 
 Goalw [step_mono_def]
  "[| step_mono step n L; s : L; p < n; s <= t; step p s = Some(u) |] ==> \
 \ !v. step p t = Some(v) --> u <= v";
-by(Blast_tac 1);
+by (Blast_tac 1);
 qed "step_monoD";
 
 Goalw [stable_def]
 "[| is_next next; semilat L; sos : listsn n (option L); \
 \   step_pres_type step n L; succs_bounded succs n |] \
 \ ==> (next step succs sos = Some sos) = (!p<n. stable step succs p sos)";
-by(Asm_simp_tac 1);
-br iffI 1;
- by(Clarify_tac 1);
- by(etac allE 1 THEN mp_tac 1);
- by(etac allE 1 THEN mp_tac 1);
- by(Clarify_tac 1);
+by (Asm_simp_tac 1);
+by (rtac iffI 1);
+ by (Clarify_tac 1);
+ by (etac allE 1 THEN mp_tac 1);
+ by (etac allE 1 THEN mp_tac 1);
+ by (Clarify_tac 1);
  bd(merges_same_conv RS iffD1)1;
-     ba 4;
-    ba 1;
-   by(blast_tac (claset() addIs [step_pres_typeD,listsn_optionE_in]) 1);
-  by(blast_tac (claset() addIs [succs_boundedD,listsn_optionE_in]) 1);
- by(Blast_tac 1);
-by(Clarify_tac 1);
-by(etac allE 1 THEN mp_tac 1);
-by(Asm_full_simp_tac 1);
-by(Clarify_tac 1);
-by(Asm_simp_tac 1);
-by(blast_tac (claset() addSIs [merges_same_conv RS iffD2]
+     by (assume_tac 4);
+    by (assume_tac 1);
+   by (blast_tac (claset() addIs [step_pres_typeD,listsn_optionE_in]) 1);
+  by (blast_tac (claset() addIs [succs_boundedD,listsn_optionE_in]) 1);
+ by (Blast_tac 1);
+by (Clarify_tac 1);
+by (etac allE 1 THEN mp_tac 1);
+by (Asm_full_simp_tac 1);
+by (Clarify_tac 1);
+by (Asm_simp_tac 1);
+by (blast_tac (claset() addSIs [merges_same_conv RS iffD2]
      addIs [step_pres_typeD,succs_boundedD,listsn_optionE_in]) 1);
 qed "fixpoint_next_iff_stable";
 
@@ -253,60 +253,60 @@
 \    tos:listsn n (option L); next step succs tos = Some tos; \
 \    sos:listsn n (option L); sos <= tos |] \
 \ ==> ? sos'. next step succs sos = Some sos' & sos' <= tos";
-by(subgoal_tac
+by (subgoal_tac
    "!p<n. !s. sos!p = Some s --> (? u. \
 \             step p s = Some(u) & (!q:set(succs p). Some u<=tos!q))" 1);
- by(exhaust_tac "next step succs sos" 1);
-  bd next_None 1;
-     ba 1;
-    ba 1;
-   ba 1;
-  by(Force_tac 1);
- by(rename_tac "sos'" 1);
- by(case_tac "sos' = sos" 1);
-  by(Blast_tac 1);
- bd next_Some2 1;
-    by(EVERY1[atac, atac, atac, atac]);
- by(Clarify_tac 1);
- by(etac allE 1 THEN mp_tac 1);
- by(etac allE 1 THEN mp_tac 1);
- by(Clarify_tac 1);
- by(EVERY1[rtac exI, rtac conjI, atac]);
- br merges_pres_le_ub 1;
-       ba 1;
-      by(blast_tac (claset() addIs [step_pres_typeD,listsn_optionE_in]) 1);
-     by(Asm_full_simp_tac 1);
-    by(blast_tac (claset() addIs [succs_boundedD,listsn_optionE_in]) 1);
-   by(REPEAT(atac 1));
-by(Clarify_tac 1);
-by(exhaust_tac "tos!p" 1);
- by(force_tac (claset() addDs [le_listD],simpset()) 1);
-by(rename_tac "t" 1);
-by(subgoal_tac "s <= t" 1);
- by(force_tac (claset() addDs [le_listD],simpset()) 2);
-by(exhaust_tac "step p s" 1);
- bd step_mono_NoneD 1;
-     ba 4;
-    by(blast_tac (claset() addIs [listsn_optionE_in]) 1);
-   ba 1;
-  ba 1;
- by(Force_tac 1);
-bd step_monoD 1;
-    ba 4;
-   by(blast_tac (claset() addIs [listsn_optionE_in]) 1);
-  ba 1;
- ba 1;
-by(Asm_full_simp_tac 1);
-by(etac allE 1 THEN mp_tac 1);
-by(etac allE 1 THEN mp_tac 1);
-by(Clarify_tac 1);
-by(Asm_full_simp_tac 1);
-by(forward_tac[merges_same_conv RS iffD1]1);
-    ba 4;
-   ba 1;
-  by(blast_tac (claset() addIs [step_pres_typeD,listsn_optionE_in]) 1);
- by(blast_tac (claset() addIs [succs_boundedD,listsn_optionE_in]) 1);
-by(blast_tac (claset() addIs [order_trans]) 1);
+ by (exhaust_tac "next step succs sos" 1);
+  by (dtac next_None 1);
+     by (assume_tac 1);
+    by (assume_tac 1);
+   by (assume_tac 1);
+  by (Force_tac 1);
+ by (rename_tac "sos'" 1);
+ by (case_tac "sos' = sos" 1);
+  by (Blast_tac 1);
+ by (dtac next_Some2 1);
+    by (EVERY1[atac, atac, atac, atac]);
+ by (Clarify_tac 1);
+ by (etac allE 1 THEN mp_tac 1);
+ by (etac allE 1 THEN mp_tac 1);
+ by (Clarify_tac 1);
+ by (EVERY1[rtac exI, rtac conjI, atac]);
+ by (rtac merges_pres_le_ub 1);
+       by (assume_tac 1);
+      by (blast_tac (claset() addIs [step_pres_typeD,listsn_optionE_in]) 1);
+     by (Asm_full_simp_tac 1);
+    by (blast_tac (claset() addIs [succs_boundedD,listsn_optionE_in]) 1);
+   by (REPEAT(atac 1));
+by (Clarify_tac 1);
+by (exhaust_tac "tos!p" 1);
+ by (force_tac (claset() addDs [le_listD],simpset()) 1);
+by (rename_tac "t" 1);
+by (subgoal_tac "s <= t" 1);
+ by (force_tac (claset() addDs [le_listD],simpset()) 2);
+by (exhaust_tac "step p s" 1);
+ by (dtac step_mono_NoneD 1);
+     by (assume_tac 4);
+    by (blast_tac (claset() addIs [listsn_optionE_in]) 1);
+   by (assume_tac 1);
+  by (assume_tac 1);
+ by (Force_tac 1);
+by (dtac step_monoD 1);
+    by (assume_tac 4);
+   by (blast_tac (claset() addIs [listsn_optionE_in]) 1);
+  by (assume_tac 1);
+ by (assume_tac 1);
+by (Asm_full_simp_tac 1);
+by (etac allE 1 THEN mp_tac 1);
+by (etac allE 1 THEN mp_tac 1);
+by (Clarify_tac 1);
+by (Asm_full_simp_tac 1);
+by (forward_tac[merges_same_conv RS iffD1]1);
+    by (assume_tac 4);
+   by (assume_tac 1);
+  by (blast_tac (claset() addIs [step_pres_typeD,listsn_optionE_in]) 1);
+ by (blast_tac (claset() addIs [succs_boundedD,listsn_optionE_in]) 1);
+by (blast_tac (claset() addIs [order_trans]) 1);
 val lemma = result();
 
 Goalw [is_dfa_def]
@@ -314,14 +314,14 @@
 \    step_pres_type step n L; succs_bounded succs n; \
 \    step_mono_None step n L; step_mono step n L |] ==> \
 \ is_dfa (%sos. fix(next step succs, sos)) step succs n L";
-by(Clarify_tac 1);
-by(stac fix_iff_has_fixpoint 1);
-     by(etac (acc_option RS acc_listsn) 1);
-    by(blast_tac (claset() addIs [lemma]) 1);
-   by(blast_tac (claset() addIs [next_preserves_type]) 1);
-  by(blast_tac (claset() addIs [next_incr]) 1);
- ba 1;
-by(asm_simp_tac (simpset() delsimps [next_Some1_eq] addsimps [fixpoint_next_iff_stable]) 1);
+by (Clarify_tac 1);
+by (stac fix_iff_has_fixpoint 1);
+     by (etac (acc_option RS acc_listsn) 1);
+    by (blast_tac (claset() addIs [lemma]) 1);
+   by (blast_tac (claset() addIs [next_preserves_type]) 1);
+  by (blast_tac (claset() addIs [next_incr]) 1);
+ by (assume_tac 1);
+by (asm_simp_tac (simpset() delsimps [next_Some1_eq] addsimps [fixpoint_next_iff_stable]) 1);
 qed "is_dfa_fix_next";
 
 Goal
@@ -332,5 +332,5 @@
 \    sos : listsn n (option L) |] ==> \
 \ fix(next step succs, sos) = \
 \ (? tos:listsn n (option L). sos<=tos & welltyping wti tos)";
-by(blast_tac (claset() addSDs [dfa_iff_welltyping] addSIs [is_dfa_fix_next]) 1);
+by (blast_tac (claset() addSDs [dfa_iff_welltyping] addSIs [is_dfa_fix_next]) 1);
 qed "fix_next_iff_welltyping";
--- a/src/HOL/BCV/Fixpoint.ML	Thu Oct 28 14:00:25 1999 +0200
+++ b/src/HOL/BCV/Fixpoint.ML	Thu Oct 28 14:55:23 1999 +0200
@@ -5,27 +5,27 @@
 *)
 
 goalw_cterm [] (cterm_of (sign_of thy) (HOLogic.mk_Trueprop (hd fix.tcs)));
-by(asm_full_simp_tac (simpset() addsimps [wf_iff_no_infinite_down_chain]
+by (asm_full_simp_tac (simpset() addsimps [wf_iff_no_infinite_down_chain]
                                 addsplits [split_split]) 1);
-by(Clarify_tac 1);
-br ccontr 1;
-by(Asm_full_simp_tac 1);
-by(res_inst_tac[("x","0")] allE 1);
-ba 1;
-by(Clarify_tac 1);
-by(rename_tac "next s0" 1);
-by(subgoal_tac "!i. fst(f i) = next" 1);
-by(eres_inst_tac[("x","%i. snd(f i)")] allE 1);
-by(Asm_full_simp_tac 1);
-by(Clarify_tac 1);
-by(eres_inst_tac[("x","i")] allE 1);
-by(eres_inst_tac[("x","i")] allE 1);
-by(Force_tac 1);
-br allI 1;
-by(induct_tac "i" 1);
-by(Asm_simp_tac 1);
-by(eres_inst_tac[("x","n")] allE 1);
-by(Auto_tac);
+by (Clarify_tac 1);
+by (rtac ccontr 1);
+by (Asm_full_simp_tac 1);
+by (res_inst_tac[("x","0")] allE 1);
+by (assume_tac 1);
+by (Clarify_tac 1);
+by (rename_tac "next s0" 1);
+by (subgoal_tac "!i. fst(f i) = next" 1);
+by (eres_inst_tac[("x","%i. snd(f i)")] allE 1);
+by (Asm_full_simp_tac 1);
+by (Clarify_tac 1);
+by (eres_inst_tac[("x","i")] allE 1);
+by (eres_inst_tac[("x","i")] allE 1);
+by (Force_tac 1);
+by (rtac allI 1);
+by (induct_tac "i" 1);
+by (Asm_simp_tac 1);
+by (eres_inst_tac[("x","n")] allE 1);
+by (Auto_tac);
 val fix_tc = result();
 
 Goalw [wf_iff_no_infinite_down_chain RS eq_reflection]
@@ -33,21 +33,21 @@
 \    !a:A. next a : option A; s:A |] ==> \
 \ fix(next,s) = (case next s of None => False \
 \                | Some t => if t=s then True else fix(next,t))";
-by(stac (fix_tc RS (hd fix.rules)) 1);
-by(Simp_tac 1);
-by(Clarify_tac 1);
-by(subgoal_tac "!i. f i : A" 1);
- by(Blast_tac 1);
-br allI 1;
-by(induct_tac "i" 1);
- by(Asm_simp_tac 1);
-by(Auto_tac);
+by (stac (fix_tc RS (hd fix.rules)) 1);
+by (Simp_tac 1);
+by (Clarify_tac 1);
+by (subgoal_tac "!i. f i : A" 1);
+ by (Blast_tac 1);
+by (rtac allI 1);
+by (induct_tac "i" 1);
+ by (Asm_simp_tac 1);
+by (Auto_tac);
 qed "fix_unfold";
 
 (* Thm: next has fixpoint above s iff fix(next,s) *)
 
 Goal "[| x = Some y; x : option A |] ==> y : A";
-by(Blast_tac 1);
+by (Blast_tac 1);
 val lemma = result();
 
 Goal
@@ -55,38 +55,38 @@
 \   !t:L. !s:L. next t = Some t & s <= t --> (? u. next s = Some u & u <= t); \
 \   !a:L. next a : option L; !a:L. !b. next a = Some b --> a <= b |] \
 \ ==> s:L --> fix(next,s) = (? t:L. s <= t & next t = Some t)";
-by(subgoal_tac "wf{(y,x::'a::order). x : L & next x = Some y & y ~= x}" 1);
- by(full_simp_tac (simpset() addsimps [acc_def]) 2);
- be wf_subset 2;
- by(simp_tac (simpset() addsimps [order_less_le]) 2);
- by(blast_tac (claset() addDs [lemma]) 2);
-by(res_inst_tac [("a","s")] wf_induct 1);
- ba 1;
-by(Clarify_tac 1);
-by(Full_simp_tac 1);
-by(stac fix_unfold 1);
-   ba 1;
-  ba 1;
- ba 1;
-by(split_tac [option.split] 1);
-br conjI 1;
- by(thin_tac "!a:L. !b. next a = Some b --> a <= b" 1);
- by(Force_tac 1);
-by(Clarify_tac 1);
-by(split_tac [split_if] 1);
-br conjI 1;
- by(Blast_tac 1);
-by(Clarify_tac 1);
-by(EVERY[etac allE 1, etac impE 1, etac impE 2, etac trans 3]);
-  by(Blast_tac 1);
- be lemma 1;
- by(Blast_tac 1);
-br iffI 1;
- by(blast_tac (claset() addIs [order_trans]) 1);
-by(Clarify_tac 1);
-by(EVERY[rtac bexI 1, atac 2, rtac conjI 1, atac 2]);
-by(thin_tac "!a:L. !b. next a = Some b --> a <= b" 1);
-by(Force_tac 1);
+by (subgoal_tac "wf{(y,x::'a::order). x : L & next x = Some y & y ~= x}" 1);
+ by (full_simp_tac (simpset() addsimps [acc_def]) 2);
+ by (etac wf_subset 2);
+ by (simp_tac (simpset() addsimps [order_less_le]) 2);
+ by (blast_tac (claset() addDs [lemma]) 2);
+by (res_inst_tac [("a","s")] wf_induct 1);
+ by (assume_tac 1);
+by (Clarify_tac 1);
+by (Full_simp_tac 1);
+by (stac fix_unfold 1);
+   by (assume_tac 1);
+  by (assume_tac 1);
+ by (assume_tac 1);
+by (split_tac [option.split] 1);
+by (rtac conjI 1);
+ by (thin_tac "!a:L. !b. next a = Some b --> a <= b" 1);
+ by (Force_tac 1);
+by (Clarify_tac 1);
+by (split_tac [split_if] 1);
+by (rtac conjI 1);
+ by (Blast_tac 1);
+by (Clarify_tac 1);
+by (EVERY[etac allE 1, etac impE 1, etac impE 2, etac trans 3]);
+  by (Blast_tac 1);
+ by (etac lemma 1);
+ by (Blast_tac 1);
+by (rtac iffI 1);
+ by (blast_tac (claset() addIs [order_trans]) 1);
+by (Clarify_tac 1);
+by (EVERY[rtac bexI 1, atac 2, rtac conjI 1, atac 2]);
+by (thin_tac "!a:L. !b. next a = Some b --> a <= b" 1);
+by (Force_tac 1);
 qed_spec_mp "fix_iff_has_fixpoint";
 
 
@@ -103,40 +103,40 @@
 \                   next t = None | (? v. next t = Some v & u <= v); \
 \   !a:L. next a : option L; !a:L. !b. next a = Some b --> a <= b |] \
 \ ==> s:L --> fix(next,s) = (? t:L. s <= t & next t = Some t)";
-by(res_inst_tac [("a","s")] wf_induct 1);
- ba 1;
-by(Clarify_tac 1);
-by(Full_simp_tac 1);
-by(stac fix_unfold 1);
-   ba 1;
-  ba 1;
- ba 1;
-by(split_tac [option.split] 1);
-br conjI 1;
- by(blast_tac (claset() addDs [sym RS trans]) 1);
-by(Clarify_tac 1);
-by(split_tac [split_if] 1);
-br conjI 1;
- by(Blast_tac 1);
-by(Clarify_tac 1);
-by(EVERY[etac allE 1, etac impE 1, etac impE 2, etac trans 3]);
-  by(Blast_tac 1);
- be lemma 1;
- by(Blast_tac 1);
-br iffI 1;
- by(subgoal_tac "next a ~= None" 1);
-  by(Clarify_tac 1);
-  by(EVERY[etac ballE 1, etac ballE 1, etac allE 1, etac impE 1,etac disjE 2]);
-      br conjI 1;
-       ba 1;
-      ba 1;
-     by(blast_tac (claset() addDs [sym RS trans]) 1);
-    by(blast_tac (claset() addIs [order_trans]) 1);
-   by(blast_tac (claset() addIs [order_trans]) 1);
-  by(Blast_tac 1);
- br notI 1;
- by(Clarify_tac 1);
- by(blast_tac (claset() addDs [sym RS trans,lemma]) 1);
-by(blast_tac (claset() addDs [sym RS trans]) 1);
+by (res_inst_tac [("a","s")] wf_induct 1);
+ by (assume_tac 1);
+by (Clarify_tac 1);
+by (Full_simp_tac 1);
+by (stac fix_unfold 1);
+   by (assume_tac 1);
+  by (assume_tac 1);
+ by (assume_tac 1);
+by (split_tac [option.split] 1);
+by (rtac conjI 1);
+ by (blast_tac (claset() addDs [sym RS trans]) 1);
+by (Clarify_tac 1);
+by (split_tac [split_if] 1);
+by (rtac conjI 1);
+ by (Blast_tac 1);
+by (Clarify_tac 1);
+by (EVERY[etac allE 1, etac impE 1, etac impE 2, etac trans 3]);
+  by (Blast_tac 1);
+ by (etac lemma 1);
+ by (Blast_tac 1);
+by (rtac iffI 1);
+ by (subgoal_tac "next a ~= None" 1);
+  by (Clarify_tac 1);
+  by (EVERY[etac ballE 1, etac ballE 1, etac allE 1, etac impE 1,etac disjE 2]);
+      by (rtac conjI 1);
+       by (assume_tac 1);
+      by (assume_tac 1);
+     by (blast_tac (claset() addDs [sym RS trans]) 1);
+    by (blast_tac (claset() addIs [order_trans]) 1);
+   by (blast_tac (claset() addIs [order_trans]) 1);
+  by (Blast_tac 1);
+ by (rtac notI 1);
+ by (Clarify_tac 1);
+ by (blast_tac (claset() addDs [sym RS trans,lemma]) 1);
+by (blast_tac (claset() addDs [sym RS trans]) 1);
 qed_spec_mp "fix_iff_has_fixpoint";
 *)
--- a/src/HOL/BCV/Machine.ML	Thu Oct 28 14:00:25 1999 +0200
+++ b/src/HOL/BCV/Machine.ML	Thu Oct 28 14:55:23 1999 +0200
@@ -12,29 +12,29 @@
 (*** Machine specific ***)
 
 Goal "!p. p<size(bs) --> maxreg(bs!p) < maxregs(bs)";
-by(induct_tac "bs" 1);
-by(auto_tac (claset(), simpset() addsimps [nth_Cons,less_max_iff_disj]
+by (induct_tac "bs" 1);
+by (auto_tac (claset(), simpset() addsimps [nth_Cons,less_max_iff_disj]
                                  addsplits [nat.split]));
 qed_spec_mp "maxreg_less_maxregs";
 
 Goalw [le_typ] "(t <= Ctyp) = (t = Ctyp)";
-by(Simp_tac 1);
+by (Simp_tac 1);
 qed "le_Ctyp_conv";
 
 Goalw [le_typ] "(t ~= Top) = (t = Atyp | t = Btyp | t = Ctyp)";
-by(induct_tac "t" 1);
-by(ALLGOALS Simp_tac);
+by (induct_tac "t" 1);
+by (ALLGOALS Simp_tac);
 qed "le_Top_conv";
 
 Goalw [step_pres_type_def,listsn_def,exec_def,stype_def]
  "step_pres_type (%p. exec (bs!p)) (length bs) (stype bs)";
-by(force_tac (claset(), simpset() addsplits [instr.split_asm,split_if_asm]) 1);
+by (force_tac (claset(), simpset() addsplits [instr.split_asm,split_if_asm]) 1);
 qed "exec_pres_type";
 
 Goalw [wti_is_fix_step_def,stable_def,wt_instr_def,exec_def,succs_def]
  "wti_is_fix_step (%p. exec (bs!p)) (%u. wt_instr (bs ! u) u) \
 \     (%p. succs (bs ! p) p) (length bs) (stype bs)";
-by(force_tac (claset() addDs [maxreg_less_maxregs,
+by (force_tac (claset() addDs [maxreg_less_maxregs,
                               le_Top_conv RS iffD1,le_Ctyp_conv RS iffD1],
               simpset() addsplits [option.split,instr.split]) 1);
 qed "wt_instr_is_fix_exec";
@@ -42,62 +42,62 @@
 
 Goalw [step_mono_None_def,exec_def,succs_def,le_list]
  "step_mono_None (%p. exec (bs!p)) (length bs) (stype bs)";
-by(Clarify_tac 1);
-by(forward_tac [maxreg_less_maxregs] 1);
-by(split_asm_tac [instr.split_asm] 1);
+by (Clarify_tac 1);
+by (ftac maxreg_less_maxregs 1);
+by (split_asm_tac [instr.split_asm] 1);
 
-by(ALLGOALS (asm_full_simp_tac (simpset() addsplits [split_if_asm])));
+by (ALLGOALS (asm_full_simp_tac (simpset() addsplits [split_if_asm])));
 
-by(rotate_tac 1 1);
-by(Asm_full_simp_tac 1);
-by(blast_tac (claset() addIs [order_trans]) 1);
+by (rotate_tac 1 1);
+by (Asm_full_simp_tac 1);
+by (blast_tac (claset() addIs [order_trans]) 1);
 
-by(rotate_tac 1 1);
-by(Asm_full_simp_tac 1);
-by(blast_tac (claset() addIs [order_trans]) 1);
+by (rotate_tac 1 1);
+by (Asm_full_simp_tac 1);
+by (blast_tac (claset() addIs [order_trans]) 1);
 
-by(rotate_tac 1 1);
-by(Asm_full_simp_tac 1);
-by(blast_tac (claset() addIs [order_trans]) 1);
+by (rotate_tac 1 1);
+by (Asm_full_simp_tac 1);
+by (blast_tac (claset() addIs [order_trans]) 1);
 
-by(rotate_tac 1 1);
-by(Asm_full_simp_tac 1);
-by(subgoal_tac "s!nat1 <= t!nat1" 1);
-by(Blast_tac 2);
-by(subgoal_tac "s!nat2 <= t!nat2" 1);
-by(Blast_tac 2);
-be thin_rl 1;
-by(asm_full_simp_tac (simpset() addsimps [le_typ])1);
-by(exhaust_tac "t!nat1" 1);
-by(ALLGOALS Asm_full_simp_tac);
-by(exhaust_tac "t!nat2" 1);
-by(ALLGOALS Asm_full_simp_tac);
-by(exhaust_tac "t!nat2" 1);
-by(ALLGOALS Asm_full_simp_tac);
-by(exhaust_tac "s!nat1" 1);
-by(ALLGOALS Asm_full_simp_tac);
-by(exhaust_tac "t!nat2" 1);
-by(ALLGOALS Asm_full_simp_tac);
+by (rotate_tac 1 1);
+by (Asm_full_simp_tac 1);
+by (subgoal_tac "s!nat1 <= t!nat1" 1);
+by (Blast_tac 2);
+by (subgoal_tac "s!nat2 <= t!nat2" 1);
+by (Blast_tac 2);
+by (etac thin_rl 1);
+by (asm_full_simp_tac (simpset() addsimps [le_typ])1);
+by (exhaust_tac "t!nat1" 1);
+by (ALLGOALS Asm_full_simp_tac);
+by (exhaust_tac "t!nat2" 1);
+by (ALLGOALS Asm_full_simp_tac);
+by (exhaust_tac "t!nat2" 1);
+by (ALLGOALS Asm_full_simp_tac);
+by (exhaust_tac "s!nat1" 1);
+by (ALLGOALS Asm_full_simp_tac);
+by (exhaust_tac "t!nat2" 1);
+by (ALLGOALS Asm_full_simp_tac);
 
 qed "exec_mono_None";
 
 Goalw [step_mono_def,exec_def,succs_def]
  "step_mono (%p. exec (bs!p)) (length bs) (stype bs)";
-by(Clarify_tac 1);
-by(forward_tac [maxreg_less_maxregs] 1);
-by(split_asm_tac [instr.split_asm] 1);
-by(ALLGOALS
+by (Clarify_tac 1);
+by (ftac maxreg_less_maxregs 1);
+by (split_asm_tac [instr.split_asm] 1);
+by (ALLGOALS
    (fast_tac (claset() addIs [list_update_le_cong,le_listD]
                addss (simpset() addsplits [split_if_asm]))));
 
 qed_spec_mp "exec_mono_Some";
 
 Goalw [stype_def] "semilat(stype bs)";
-by(Blast_tac 1);
+by (Blast_tac 1);
 qed "lat_stype";
 
 Goalw [stype_def] "acc(stype bs)";
-by(Simp_tac 1);
+by (Simp_tac 1);
 qed "acc_stype";
 
 Delsimps [stype_def];
@@ -109,8 +109,8 @@
 \ fix(next (%p. exec (bs!p)) (%p. succs (bs!p) p), sos) = \
 \ (? tos : listsn (size bs) (option(stype bs)). \
 \      sos <= tos & welltyping (%p. wt_instr (bs!p) p) tos)";
-by(simp_tac (simpset() delsimps [not_None_eq]) 1);
-by(REPEAT(ares_tac [fix_next_iff_welltyping,exec_pres_type,
+by (simp_tac (simpset() delsimps [not_None_eq]) 1);
+by (REPEAT(ares_tac [fix_next_iff_welltyping,exec_pres_type,
                     exec_mono_None,exec_mono_Some,
                     wt_instr_is_fix_exec,lat_stype,acc_stype] 1));
 
--- a/src/HOL/BCV/Orders.ML	Thu Oct 28 14:00:25 1999 +0200
+++ b/src/HOL/BCV/Orders.ML	Thu Oct 28 14:55:23 1999 +0200
@@ -10,36 +10,36 @@
 section "option";
 
 Goalw [option_def] "None : option A";
-by(Simp_tac 1);
+by (Simp_tac 1);
 qed "None_in_option";
 AddIffs [None_in_option];
 
 Goalw [option_def] "(Some x : option A) = (x:A)";
-by(Auto_tac);
+by (Auto_tac);
 qed "Some_in_option";
 AddIffs [Some_in_option];
 
 Goalw [less_option,le_option]
  "Some x < opt = (? y. opt = Some y & x < (y::'a::order))";
-by(simp_tac (simpset() addsimps [order_less_le] addsplits [option.split]) 1);
+by (simp_tac (simpset() addsimps [order_less_le] addsplits [option.split]) 1);
 qed_spec_mp "Some_less_conv";
 AddIffs [Some_less_conv];
 
 Goalw [less_option,le_option] "None < x = (? a. x = Some a)";
-by(asm_simp_tac (simpset() addsplits [option.split]) 1);
+by (asm_simp_tac (simpset() addsplits [option.split]) 1);
 qed_spec_mp "None_less_iff";
 AddIffs [None_less_iff];
 
 
 Goalw [acc_def] "acc A ==> acc(option A)";
-by(asm_full_simp_tac (simpset() addsimps [wf_eq_minimal]) 1);
-by(Clarify_tac 1);
-by(case_tac "? a. Some a : Q" 1);
- by(eres_inst_tac [("x","{a . Some a : Q}")] allE 1);
- by(Blast_tac 1);
-by(exhaust_tac "x" 1);
- by(Fast_tac 1);
-by(Blast_tac 1);
+by (asm_full_simp_tac (simpset() addsimps [wf_eq_minimal]) 1);
+by (Clarify_tac 1);
+by (case_tac "? a. Some a : Q" 1);
+ by (eres_inst_tac [("x","{a . Some a : Q}")] allE 1);
+ by (Blast_tac 1);
+by (exhaust_tac "x" 1);
+ by (Fast_tac 1);
+by (Blast_tac 1);
 qed "acc_option";
 Addsimps [acc_option];
 AddSIs [acc_option];
@@ -49,129 +49,129 @@
 section "list";
 
 Goalw [le_list] "~ [] <= x#xs";
-by(Simp_tac 1);
+by (Simp_tac 1);
 qed "Nil_notle_Cons";
 
 Goalw [le_list] "~ x#xs <= []";
-by(Simp_tac 1);
+by (Simp_tac 1);
 qed "Cons_notle_Nil";
 
 AddIffs [Nil_notle_Cons,Cons_notle_Nil];
 
 Goalw [le_list] "x#xs <= y#ys = (x <= y & xs <= ys)";
-br iffI 1;
- by(Asm_full_simp_tac 1);
- by(Clarify_tac 1);
- br conjI 1;
-  by(eres_inst_tac [("x","0")] allE 1);
-  by(Asm_full_simp_tac 1);
- by(Clarify_tac 1);
- by(eres_inst_tac [("x","Suc i")] allE 1);
- by(Asm_full_simp_tac 1);
-by(Asm_full_simp_tac 1);
-by(Clarify_tac 1);
-by(exhaust_tac "i" 1);
-by(ALLGOALS Asm_simp_tac);
+by (rtac iffI 1);
+ by (Asm_full_simp_tac 1);
+ by (Clarify_tac 1);
+ by (rtac conjI 1);
+  by (eres_inst_tac [("x","0")] allE 1);
+  by (Asm_full_simp_tac 1);
+ by (Clarify_tac 1);
+ by (eres_inst_tac [("x","Suc i")] allE 1);
+ by (Asm_full_simp_tac 1);
+by (Asm_full_simp_tac 1);
+by (Clarify_tac 1);
+by (exhaust_tac "i" 1);
+by (ALLGOALS Asm_simp_tac);
 qed "Cons_le_Cons";
 AddIffs [Cons_le_Cons];
 
 Goal "(x#xs <= ys) = (? z zs. ys = z#zs & x <= z & xs <= zs)";
-by(exhaust_tac "ys" 1);
-by(ALLGOALS Asm_simp_tac);
+by (exhaust_tac "ys" 1);
+by (ALLGOALS Asm_simp_tac);
 qed "Cons_le_iff";
 
 Goal "(xs <= y#ys) = (? z zs. xs = z#zs & z <= y & zs <= ys)";
-by(exhaust_tac "xs" 1);
-by(ALLGOALS Asm_simp_tac);
+by (exhaust_tac "xs" 1);
+by (ALLGOALS Asm_simp_tac);
 qed "le_Cons_iff";
 
 Goalw [less_list]
  "x#xs < y#ys = (x < (y::'a::order) & xs <= ys  |  x = y & xs < ys)";
-by(simp_tac (simpset() addsimps [order_less_le]) 1);
-by(Blast_tac 1);
+by (simp_tac (simpset() addsimps [order_less_le]) 1);
+by (Blast_tac 1);
 qed "Cons_less_Cons";
 AddIffs [Cons_less_Cons];
 
 Goalw [le_list]
  "[| i<size xs; xs <= ys; x <= y |] ==> xs[i:=x] <= ys[i:=y]";
-by(asm_full_simp_tac (simpset() addsimps [nth_list_update]) 1);
+by (asm_full_simp_tac (simpset() addsimps [nth_list_update]) 1);
 qed "list_update_le_cong";
 
 Goal "!i < length xs. (xs[i := x] <= xs) = ((x::'a::order) <= xs!i)";
-by(induct_tac "xs" 1);
- by(Simp_tac 1);
-by(asm_simp_tac (simpset() addsplits [nat.split]) 1);
+by (induct_tac "xs" 1);
+ by (Simp_tac 1);
+by (asm_simp_tac (simpset() addsplits [nat.split]) 1);
 qed_spec_mp "list_update_le_conv";
 
 Goalw [listsn_def] "xs : listsn n A ==> length xs = n";
-by(Blast_tac 1);
+by (Blast_tac 1);
 qed "listsnE_length";
 Addsimps [listsnE_length];
 
 Goalw [listsn_def] "xs : listsn n A ==> set xs <= A";
-by(Blast_tac 1);
+by (Blast_tac 1);
 qed "listsnE_set";
 Addsimps [listsnE_set];
 
 Goalw [listsn_def] "listsn 0 A = {[]}";
-by(Auto_tac); 
+by (Auto_tac); 
 qed "listsn_0";
 Addsimps [listsn_0];
 
 Goalw [listsn_def]
  "(xs : listsn (Suc n) A) = (? y:A. ? ys:listsn n A. xs = y#ys)";
-by(exhaust_tac "xs" 1);
-by(Auto_tac);
+by (exhaust_tac "xs" 1);
+by (Auto_tac);
 qed "in_listsn_Suc_iff";
 
 
 Goal "? a. a:A ==> ? xs. xs : listsn n A";
-by(induct_tac "n" 1);
- by(Simp_tac 1);
-by(asm_simp_tac (simpset() addsimps [in_listsn_Suc_iff]) 1);
-by(Blast_tac 1);
+by (induct_tac "n" 1);
+ by (Simp_tac 1);
+by (asm_simp_tac (simpset() addsimps [in_listsn_Suc_iff]) 1);
+by (Blast_tac 1);
 qed "listsn_not_empty";
 
 
 Goal "!i n. length xs = n --> set xs <= A --> i < n --> (xs!i) : A";
-by(induct_tac "xs" 1);
- by(Simp_tac 1);
-by(asm_full_simp_tac (simpset() addsimps [nth_Cons] addsplits [nat.split]) 1);
+by (induct_tac "xs" 1);
+ by (Simp_tac 1);
+by (asm_full_simp_tac (simpset() addsimps [nth_Cons] addsplits [nat.split]) 1);
 qed_spec_mp "nth_in";
 Addsimps [nth_in];
 
 Goal "[| xs : listsn n A; i < n |] ==> (xs!i) : A";
-by(blast_tac (HOL_cs addIs [nth_in,listsnE_length,listsnE_set]) 1);
+by (blast_tac (HOL_cs addIs [nth_in,listsnE_length,listsnE_set]) 1);
 qed "listsnE_nth_in";
 
 Goalw [listsn_def]
  "[| xs : listsn n A; x:A; i < n |] ==> xs[i := x] : listsn n A";
-by(Asm_full_simp_tac 1);
-by(blast_tac (claset() addDs [set_list_update_subset RS subsetD]) 1);
+by (Asm_full_simp_tac 1);
+by (blast_tac (claset() addDs [set_list_update_subset RS subsetD]) 1);
 qed "list_update_in_listsn";
 Addsimps [list_update_in_listsn];
 AddSIs [list_update_in_listsn];
 
 
 Goalw [acc_def] "acc(A) ==> acc(listsn n A)";
-by(asm_full_simp_tac (simpset() addsimps [wf_eq_minimal]) 1);
-by(induct_tac "n" 1);
- by(simp_tac (simpset() addcongs [conj_cong]) 1);
-by(simp_tac (simpset() addsimps [in_listsn_Suc_iff]) 1);
-by(Clarify_tac 1);
-by(rename_tac "M m" 1);
-by(case_tac "? x xs. x:A & xs:listsn n A & x#xs : M" 1);
- be thin_rl 2;
- be thin_rl 2;
- by(Blast_tac 2);
-by(eres_inst_tac [("x","{a. a:A & (? xs. xs:listsn n A & a#xs:M)}")] allE 1);
-be impE 1;
- by(Blast_tac 1);
-by(thin_tac "? x xs. ?P x xs" 1);
-by(Clarify_tac 1);
-by(rename_tac "maxA xs" 1);
-by(eres_inst_tac [("x","{xs. xs:listsn n A & maxA#xs : M}")] allE 1);
-by(Blast_tac 1);
+by (asm_full_simp_tac (simpset() addsimps [wf_eq_minimal]) 1);
+by (induct_tac "n" 1);
+ by (simp_tac (simpset() addcongs [conj_cong]) 1);
+by (simp_tac (simpset() addsimps [in_listsn_Suc_iff]) 1);
+by (Clarify_tac 1);
+by (rename_tac "M m" 1);
+by (case_tac "? x xs. x:A & xs:listsn n A & x#xs : M" 1);
+ by (etac thin_rl 2);
+ by (etac thin_rl 2);
+ by (Blast_tac 2);
+by (eres_inst_tac [("x","{a. a:A & (? xs. xs:listsn n A & a#xs:M)}")] allE 1);
+by (etac impE 1);
+ by (Blast_tac 1);
+by (thin_tac "? x xs. ?P x xs" 1);
+by (Clarify_tac 1);
+by (rename_tac "maxA xs" 1);
+by (eres_inst_tac [("x","{xs. xs:listsn n A & maxA#xs : M}")] allE 1);
+by (Blast_tac 1);
 qed "acc_listsn";
 Addsimps [acc_listsn];
 AddSIs [acc_listsn];
--- a/src/HOL/BCV/Orders0.ML	Thu Oct 28 14:00:25 1999 +0200
+++ b/src/HOL/BCV/Orders0.ML	Thu Oct 28 14:55:23 1999 +0200
@@ -8,30 +8,30 @@
 section "option";
 
 Goalw [le_option] "(o1::('a::order)option) <= o1";
-by(simp_tac (simpset() addsplits [option.split]) 1);
+by (simp_tac (simpset() addsplits [option.split]) 1);
 qed "le_option_refl";
 
 Goalw [le_option] "(o1::('a::order)option) <= o2 --> o2<=o3 --> o1<=o3";
-by(simp_tac (simpset() addsplits [option.split]) 1);
-by(blast_tac (claset() addIs [order_trans]) 1);
+by (simp_tac (simpset() addsplits [option.split]) 1);
+by (blast_tac (claset() addIs [order_trans]) 1);
 qed_spec_mp "le_option_trans";
 
 Goalw [le_option] "(o1::('a::order)option) <= o2 --> o2<=o1 --> o1=o2";
-by(simp_tac (simpset() addsplits [option.split]) 1);
-by(blast_tac (claset() addIs [order_antisym]) 1);
+by (simp_tac (simpset() addsplits [option.split]) 1);
+by (blast_tac (claset() addIs [order_antisym]) 1);
 qed_spec_mp "le_option_antisym";
 
 Goalw [less_option] "(o1::('a::ord)option) < o2 = (o1<=o2 & o1 ~= o2)";
-br refl 1;
+by (rtac refl 1);
 qed "less_le_option";
 
 Goalw [le_option] "Some x <= opt = (? y. opt = Some y & x <= y)";
-by(simp_tac (simpset() addsplits [option.split]) 1);
+by (simp_tac (simpset() addsplits [option.split]) 1);
 qed_spec_mp "Some_le_conv";
 AddIffs [Some_le_conv];
 
 Goalw [le_option] "None <= opt";
-by(simp_tac (simpset() addsplits [option.split]) 1);
+by (simp_tac (simpset() addsplits [option.split]) 1);
 qed "None_le";
 AddIffs [None_le];
 
@@ -40,90 +40,90 @@
 section "list";
 
 Goalw [le_list] "[| xs <= ys; p < size xs |] ==> xs!p <= ys!p";
-by(Blast_tac 1);
+by (Blast_tac 1);
 qed "le_listD";
 
 Goalw [le_list] "([] <= ys) = (ys = [])";
-by(Simp_tac 1);
+by (Simp_tac 1);
 qed "Nil_le_conv";
 AddIffs [Nil_le_conv];
 
 Goalw [le_list] "(xs::('a::order)list) <= xs";
-by(induct_tac "xs" 1);
-by(Auto_tac);
+by (induct_tac "xs" 1);
+by (Auto_tac);
 qed "le_list_refl";
 
 Goalw [le_list]
  "!ys zs.(xs::('a::order)list) <= ys --> ys <= zs --> xs <= zs";
-by(induct_tac "xs" 1);
- by(Simp_tac 1);
-br allI 1;
-by(exhaust_tac "ys" 1);
- by(hyp_subst_tac 1);
- by(Simp_tac 1);
-br allI 1;
-by(exhaust_tac "zs" 1);
- by(hyp_subst_tac 1);
- by(Simp_tac 1);
-by(hyp_subst_tac 1);
-by(simp_tac (simpset() addsimps [nth_Cons] addsplits [nat.split]) 1);
-by(Clarify_tac 1);
-br conjI 1;
- by(REPEAT(EVERY1[etac allE, etac conjE, etac impE, rtac refl]));
- by(blast_tac (claset() addIs [order_trans]) 1);
-by(Clarify_tac 1);
-by(EVERY[etac allE 1, etac allE 1, etac impE 1, etac impE 2]);
-  br conjI 1;
-  ba 1;
-  by(Blast_tac 1);
- br conjI 1;
- ba 1;
- by(Blast_tac 1);
-by(Asm_full_simp_tac 1);
+by (induct_tac "xs" 1);
+ by (Simp_tac 1);
+by (rtac allI 1);
+by (exhaust_tac "ys" 1);
+ by (hyp_subst_tac 1);
+ by (Simp_tac 1);
+by (rtac allI 1);
+by (exhaust_tac "zs" 1);
+ by (hyp_subst_tac 1);
+ by (Simp_tac 1);
+by (hyp_subst_tac 1);
+by (simp_tac (simpset() addsimps [nth_Cons] addsplits [nat.split]) 1);
+by (Clarify_tac 1);
+by (rtac conjI 1);
+ by (REPEAT(EVERY1[etac allE, etac conjE, etac impE, rtac refl]));
+ by (blast_tac (claset() addIs [order_trans]) 1);
+by (Clarify_tac 1);
+by (EVERY[etac allE 1, etac allE 1, etac impE 1, etac impE 2]);
+  by (rtac conjI 1);
+  by (assume_tac 1);
+  by (Blast_tac 1);
+ by (rtac conjI 1);
+ by (assume_tac 1);
+ by (Blast_tac 1);
+by (Asm_full_simp_tac 1);
 qed_spec_mp "le_list_trans";
 
 Goalw [le_list]
  "!ys. (xs::('a::order)list) <= ys --> ys <= xs --> xs = ys";
-by(induct_tac "xs" 1);
- by(Simp_tac 1);
-br allI 1;
-by(exhaust_tac "ys" 1);
- by(hyp_subst_tac 1);
- by(Simp_tac 1);
-by(hyp_subst_tac 1);
-by(simp_tac (simpset() addsimps [nth_Cons] addsplits [nat.split]) 1);
-by(Clarify_tac 1);
-br conjI 1;
- by(blast_tac (claset() addIs [order_antisym]) 1);
-by(Asm_full_simp_tac 1);
+by (induct_tac "xs" 1);
+ by (Simp_tac 1);
+by (rtac allI 1);
+by (exhaust_tac "ys" 1);
+ by (hyp_subst_tac 1);
+ by (Simp_tac 1);
+by (hyp_subst_tac 1);
+by (simp_tac (simpset() addsimps [nth_Cons] addsplits [nat.split]) 1);
+by (Clarify_tac 1);
+by (rtac conjI 1);
+ by (blast_tac (claset() addIs [order_antisym]) 1);
+by (Asm_full_simp_tac 1);
 qed_spec_mp "le_list_antisym";
 
 Goalw [less_list] "(xs::('a::ord)list) < ys = (xs<=ys & xs ~= ys)";
-br refl 1;
+by (rtac refl 1);
 qed "less_le_list";
 
 (** product **)
 section "product";
 
 Goalw [le_prod] "(p1::('a::order * 'b::order)) <= p1";
-by(Simp_tac 1);
+by (Simp_tac 1);
 qed "le_prod_refl";
 
 Goalw [le_prod]
  "[| (p1::('a::order * 'b::order)) <= p2; p2<=p3 |] ==> p1<=p3";
-by(blast_tac (claset() addIs [order_trans]) 1);
+by (blast_tac (claset() addIs [order_trans]) 1);
 qed "le_prod_trans";
 
 Goalw [le_prod]
  "[| (p1::('a::order * 'b::order)) <= p2; p2 <= p1 |] ==> p1 = p2";
-by(blast_tac (claset() addIs [order_antisym,trans,surjective_pairing,sym]) 1);
+by (blast_tac (claset() addIs [order_antisym,trans,surjective_pairing,sym]) 1);
 qed_spec_mp "le_prod_antisym";
 
 Goalw [less_prod] "(p1::('a::order * 'b::order)) < p2 = (p1<=p2 & p1 ~= p2)";
-br refl 1;
+by (rtac refl 1);
 qed "less_le_prod";
 
 Goalw [le_prod] "((a,b) <= (c,d)) = (a <= c & b <= d)";
-by(Simp_tac 1);
+by (Simp_tac 1);
 qed "le_prod_iff";
 AddIffs [le_prod_iff];
--- a/src/HOL/BCV/Plus.ML	Thu Oct 28 14:00:25 1999 +0200
+++ b/src/HOL/BCV/Plus.ML	Thu Oct 28 14:55:23 1999 +0200
@@ -7,42 +7,42 @@
 (** option **)
 
 Goalw [plus_option] "x+None = x";
-by(simp_tac (simpset() addsplits [option.split]) 1);
+by (simp_tac (simpset() addsplits [option.split]) 1);
 qed "plus_None";
 Addsimps [plus_None];
 
 Goalw [plus_option] "None+x = x";
-by(simp_tac (simpset() addsplits [option.split]) 1);
+by (simp_tac (simpset() addsplits [option.split]) 1);
 qed "None_plus";
 Addsimps [None_plus];
 
 Goalw [plus_option] "Some x + Some y = Some(x+y)";
-by(Simp_tac 1);
+by (Simp_tac 1);
 qed "Some_plus_Some";
 Addsimps [Some_plus_Some];
 
 Goalw [plus_option] "? y. Some x + opt = Some y";
-by(simp_tac (simpset() addsplits [option.split]) 1);
+by (simp_tac (simpset() addsplits [option.split]) 1);
 qed "plus_option_Some_Some";
 
 (** list **)
 
 Goal "list_plus xs [] = xs";
-by(induct_tac "xs" 1);
- by(Simp_tac 1);
-by(asm_simp_tac (simpset() addsplits [list.split]) 1);
+by (induct_tac "xs" 1);
+ by (Simp_tac 1);
+by (asm_simp_tac (simpset() addsplits [list.split]) 1);
 qed "list_plus_Nil2";
 Addsimps [list_plus_Nil2];
 
 Goal "!ys. length(list_plus xs ys) = max(length xs) (length ys)";
-by(induct_tac "xs" 1);
- by(Simp_tac 1);
-by(asm_simp_tac (simpset() addsplits [list.split]) 1);
+by (induct_tac "xs" 1);
+ by (Simp_tac 1);
+by (asm_simp_tac (simpset() addsplits [list.split]) 1);
 qed_spec_mp "length_list_plus";
 Addsimps [length_list_plus];
 
 Goalw [plus_list]
  "length(ts+us) = max (length ts) (length us)";
-by(Simp_tac 1);
+by (Simp_tac 1);
 qed "length_plus_list";
 Addsimps [length_plus_list];
--- a/src/HOL/BCV/SemiLattice.ML	Thu Oct 28 14:00:25 1999 +0200
+++ b/src/HOL/BCV/SemiLattice.ML	Thu Oct 28 14:55:23 1999 +0200
@@ -5,58 +5,58 @@
 *)
 
 Goalw [semilat_def] "[| x:A; y:A; semilat A |] ==> x <= x+y";
-by(Blast_tac 1);
+by (Blast_tac 1);
 qed "semilat_ub1";
 
 Goalw [semilat_def] "[| x:A; y:A; semilat A |] ==> y <= x+y";
-by(Blast_tac 1);
+by (Blast_tac 1);
 qed "semilat_ub2";
 
 Goalw [semilat_def] "[| x:A; y:A; z:A; x <= z; y <= z; semilat A |] ==> x+y <= z";
-by(Blast_tac 1);
+by (Blast_tac 1);
 qed "semilat_lub";
 
 Goalw [semilat_def] "[| x:A; y:A; semilat A |] ==> x+y : A";
-by(Blast_tac 1);
+by (Blast_tac 1);
 qed "semilat_plus";
 
 Addsimps [semilat_ub1,semilat_ub2,semilat_lub,semilat_plus];
 
 Goal "[| x:A; y:A; semilat A |] ==> (x <= y) = (x+y = y)";
-br iffI 1;
- by(REPEAT(ares_tac [order_antisym,semilat_lub,order_refl,semilat_ub2] 1));
-be subst 1;
-by(Asm_simp_tac 1);
+by (rtac iffI 1);
+ by (REPEAT(ares_tac [order_antisym,semilat_lub,order_refl,semilat_ub2] 1));
+by (etac subst 1);
+by (Asm_simp_tac 1);
 qed "le_iff_plus_unchanged";
 
 Goal "[| x:A; y:A; semilat A |] ==> (x <= y) = (y = x+y)";
-br iffI 1;
- by(REPEAT(ares_tac [order_antisym,semilat_lub,order_refl,semilat_ub2] 1));
-be ssubst 1;
-by(Asm_simp_tac 1);
+by (rtac iffI 1);
+ by (REPEAT(ares_tac [order_antisym,semilat_lub,order_refl,semilat_ub2] 1));
+by (etac ssubst 1);
+by (Asm_simp_tac 1);
 qed "le_iff_plus_unchanged2";
 
 Goal "[| x:A; y:A; z:A; semilat A; x <= y | x <= z |] ==> x <= y+z";
-by(blast_tac (claset() addIs [order_trans,semilat_ub1,semilat_ub2]) 1);
+by (blast_tac (claset() addIs [order_trans,semilat_ub1,semilat_ub2]) 1);
 qed "plus_mono";
 
 Goal "[| x:A; semilat A |] ==> x+x = x";
-by(REPEAT(ares_tac [order_refl RSN (4,le_iff_plus_unchanged RS iffD1)] 1));
+by (REPEAT(ares_tac [order_refl RSN (4,le_iff_plus_unchanged RS iffD1)] 1));
 qed "semilat_idemp";
 Addsimps [semilat_idemp];
 
 Goal "[| x:A; y:A; z:A; semilat A |] ==> (x+y)+z = x+(y+z)";
-by(blast_tac (HOL_cs addSIs [order_refl,order_antisym,semilat_lub,semilat_plus]
+by (blast_tac (HOL_cs addSIs [order_refl,order_antisym,semilat_lub,semilat_plus]
                      addIs [plus_mono]) 1);
 qed "semilat_assoc";
 
 Goal "[| x:A; y:A; semilat A |] ==> x+(x+y) = x+y";
-by(asm_simp_tac (simpset() addsimps [semilat_assoc RS sym]) 1);
+by (asm_simp_tac (simpset() addsimps [semilat_assoc RS sym]) 1);
 qed "semilat_assoc_idemp";
 Addsimps [semilat_assoc_idemp];
 
 Goal "[| x:A; y:A; z:A; semilat A |] ==> (x+y <= z) = (x <= z & y <= z)";
-by(blast_tac (claset() addIs [semilat_ub1,semilat_ub2,semilat_lub,order_trans]) 1);
+by (blast_tac (claset() addIs [semilat_ub1,semilat_ub2,semilat_lub,order_trans]) 1);
 qed "plus_le_conv";
 Addsimps [plus_le_conv];
 
@@ -64,92 +64,92 @@
 (** option **)
 
 Goal "semilat A ==> semilat (option A)";
-by(simp_tac (simpset() addsimps [semilat_def,le_option,plus_option]
+by (simp_tac (simpset() addsimps [semilat_def,le_option,plus_option]
                        addsplits [option.split]) 1);
-by(blast_tac (claset() addIs [semilat_ub1,semilat_ub2,semilat_lub,semilat_plus]) 1);
+by (blast_tac (claset() addIs [semilat_ub1,semilat_ub2,semilat_lub,semilat_plus]) 1);
 qed "semilat_option";
 Addsimps [semilat_option];
 AddSIs [semilat_option];
 
 (** list **)
 Goalw [plus_list] "[] + [] = []";
-by(Simp_tac 1);
+by (Simp_tac 1);
 qed "plus_Nil_Nil";
 Addsimps [plus_Nil_Nil];
 
 Goalw [plus_list] "(x#xs) + (y#ys) = (x+y)#(xs+ys)";
-by(Simp_tac 1);
+by (Simp_tac 1);
 qed "plus_Cons_Cons";
 Addsimps [plus_Cons_Cons];
 
 Goal
  "!xs ys i. length xs = n--> length ys = n--> i<n--> (xs+ys)!i = xs!i + ys!i";
-by(induct_tac "n" 1);
- by(Simp_tac 1);
-by(Clarify_tac 1);
-by(exhaust_tac "xs" 1);
- by(Asm_full_simp_tac 1);
-by(exhaust_tac "ys" 1);
- by(Asm_full_simp_tac 1);
-by(asm_full_simp_tac (simpset() addsimps [nth_Cons] addsplits [nat.split]) 1);
+by (induct_tac "n" 1);
+ by (Simp_tac 1);
+by (Clarify_tac 1);
+by (exhaust_tac "xs" 1);
+ by (Asm_full_simp_tac 1);
+by (exhaust_tac "ys" 1);
+ by (Asm_full_simp_tac 1);
+by (asm_full_simp_tac (simpset() addsimps [nth_Cons] addsplits [nat.split]) 1);
 qed_spec_mp "nth_plus";
 Addsimps [nth_plus];
 
 Goalw [le_list]
  "semilat A ==> !xs:listsn n A. !ys:listsn n A. xs <= xs+ys";
-by(asm_simp_tac (simpset() addsimps [le_max_iff_disj]) 1);
-by(Clarify_tac 1);
-by(blast_tac (HOL_cs addIs [nth_in,nth_plus,listsnE_length,listsnE_set,semilat_ub1]) 1);
+by (asm_simp_tac (simpset() addsimps [le_max_iff_disj]) 1);
+by (Clarify_tac 1);
+by (blast_tac (HOL_cs addIs [nth_in,nth_plus,listsnE_length,listsnE_set,semilat_ub1]) 1);
 qed_spec_mp "plus_list_ub1";
 
 Goalw [le_list]
  "semilat A ==> !xs:listsn n A. !ys:listsn n A. ys <= xs+ys";
-by(asm_simp_tac (simpset() addsimps [le_max_iff_disj]) 1);
-by(Clarify_tac 1);
-by(blast_tac (HOL_cs addIs [nth_in,nth_plus,listsnE_length,listsnE_set,semilat_ub2]) 1);
+by (asm_simp_tac (simpset() addsimps [le_max_iff_disj]) 1);
+by (Clarify_tac 1);
+by (blast_tac (HOL_cs addIs [nth_in,nth_plus,listsnE_length,listsnE_set,semilat_ub2]) 1);
 qed_spec_mp "plus_list_ub2";
 
 Goalw [le_list]
  "semilat A ==> !xs:listsn n A. !ys:listsn n A. !zs:listsn n A. \
 \              xs <= zs & ys <= zs --> xs+ys <= zs";
-by(asm_simp_tac (simpset() addsimps [le_max_iff_disj] addcongs [conj_cong]) 1);
-by(Clarify_tac 1);
-by(blast_tac (HOL_cs addIs [nth_in,nth_plus,listsnE_length,listsnE_set,semilat_lub]) 1);
+by (asm_simp_tac (simpset() addsimps [le_max_iff_disj] addcongs [conj_cong]) 1);
+by (Clarify_tac 1);
+by (blast_tac (HOL_cs addIs [nth_in,nth_plus,listsnE_length,listsnE_set,semilat_lub]) 1);
 qed_spec_mp "plus_list_lub";
 
 Goalw [listsn_def]
  "semilat A ==> !xs:listsn n A. !ys:listsn n A. xs+ys : listsn n A";
-by(induct_tac "n" 1);
- by(Simp_tac 1);
-by(Asm_full_simp_tac 1);
-by(Clarify_tac 1);
-by(exhaust_tac "xs" 1);
- by(Asm_full_simp_tac 1);
-by(exhaust_tac "ys" 1);
- by(Asm_full_simp_tac 1);
-by(Asm_full_simp_tac 1);
-by(blast_tac (claset() addIs [semilat_plus]) 1);
+by (induct_tac "n" 1);
+ by (Simp_tac 1);
+by (Asm_full_simp_tac 1);
+by (Clarify_tac 1);
+by (exhaust_tac "xs" 1);
+ by (Asm_full_simp_tac 1);
+by (exhaust_tac "ys" 1);
+ by (Asm_full_simp_tac 1);
+by (Asm_full_simp_tac 1);
+by (blast_tac (claset() addIs [semilat_plus]) 1);
 qed_spec_mp "plus_list_closed";
 
 Goal "semilat A ==> semilat (listsn n A)";
-by(simp_tac (HOL_basic_ss addsimps [semilat_def]) 1);
-by(asm_simp_tac (HOL_ss addsimps [plus_list_ub1,plus_list_ub2,plus_list_lub,plus_list_closed]) 1);
+by (simp_tac (HOL_basic_ss addsimps [semilat_def]) 1);
+by (asm_simp_tac (HOL_ss addsimps [plus_list_ub1,plus_list_ub2,plus_list_lub,plus_list_closed]) 1);
 qed "semilat_listsn";
 Addsimps [semilat_listsn];
 AddSIs [semilat_listsn];
 
 Goalw [le_list]
  "!i xs. xs : listsn n A --> x:A --> semilat A --> i<n --> xs <= xs[i := x + xs!i]";
-by(Simp_tac 1);
-by(induct_tac "n" 1);
- by(Simp_tac 1);
-by(asm_simp_tac (simpset() addsimps [in_listsn_Suc_iff]) 1);
-by(Clarify_tac 1);
-by(asm_simp_tac (simpset() addsimps [nth_Cons] addsplits [nat.split]) 1);
+by (Simp_tac 1);
+by (induct_tac "n" 1);
+ by (Simp_tac 1);
+by (asm_simp_tac (simpset() addsimps [in_listsn_Suc_iff]) 1);
+by (Clarify_tac 1);
+by (asm_simp_tac (simpset() addsimps [nth_Cons] addsplits [nat.split]) 1);
 qed_spec_mp "list_update_incr";
 
 (* product *)
 
 Goalw [semilat_def,plus_prod] "[| semilat A; semilat B |] ==> semilat (A Times B)";
-by(Asm_simp_tac 1);
+by (Asm_simp_tac 1);
 qed "semilat_Times";
--- a/src/HOL/BCV/Types.ML	Thu Oct 28 14:00:25 1999 +0200
+++ b/src/HOL/BCV/Types.ML	Thu Oct 28 14:55:23 1999 +0200
@@ -5,27 +5,27 @@
 *)
 
 Goalw [semilat_def,le_typ,plus_typ] "semilat (UNIV::typ set)";
-by(Auto_tac);
+by (Auto_tac);
 qed "semilat_typ";
 AddIffs [semilat_typ];
 
 Goal "{(x,y::'a::order). y<x}^+ = {(x,y::'a::order). y<x}";
-by(Auto_tac);
- be trancl_induct 1;
- by(Blast_tac 1);
- by(blast_tac (claset() addIs [order_less_trans]) 1);
-by(blast_tac (claset() addIs [r_into_trancl]) 1);
+by (Auto_tac);
+ by (etac trancl_induct 1);
+ by (Blast_tac 1);
+ by (blast_tac (claset() addIs [order_less_trans]) 1);
+by (blast_tac (claset() addIs [r_into_trancl]) 1);
 qed "trancl_order1_conv";
 Addsimps [trancl_order1_conv];
 
 Goalw [acyclic_def] "acyclic{(x,y::'a::order). y<x}";
-by(Simp_tac 1);
+by (Simp_tac 1);
 qed "acyclic_order1";
 AddIffs [acyclic_order1];
 
 Goalw [acc_def] "acc(UNIV::typ set)";
-br finite_acyclic_wf 1;
- by(fast_tac (claset() addIs [finite_SigmaI RSN (2,finite_subset)]) 1);
-by(blast_tac (claset() addIs [acyclic_subset]) 1);
+by (rtac finite_acyclic_wf 1);
+ by (fast_tac (claset() addIs [finite_SigmaI RSN (2,finite_subset)]) 1);
+by (blast_tac (claset() addIs [acyclic_subset]) 1);
 qed "acc_UNIV_typ";
 AddIffs [acc_UNIV_typ];
--- a/src/HOL/BCV/Types0.ML	Thu Oct 28 14:00:25 1999 +0200
+++ b/src/HOL/BCV/Types0.ML	Thu Oct 28 14:55:23 1999 +0200
@@ -5,29 +5,29 @@
 *)
 
 Goalw [le_typ] "(t::typ) <= t";
-by(Blast_tac 1);
+by (Blast_tac 1);
 qed "le_typ_refl";
 
 Goalw [le_typ] "[| (t1::typ) <= t2; t2<=t3 |] ==> t1<=t3";
-by(Blast_tac 1);
+by (Blast_tac 1);
 qed "le_typ_trans";
 
 Goalw [le_typ] "[| (t1::typ) <= t2; t2<=t1 |] ==> t1=t2";
-by(Blast_tac 1);
+by (Blast_tac 1);
 qed "le_typ_antisym";
 
 Goalw [less_typ] "(t1::typ) < t2 = (t1<=t2 & t1 ~= t2)";
-br refl 1;
+by (rtac refl 1);
 qed "less_le_typ";
 
 Goal "(UNIV::typ set) = {Atyp,Btyp,Ctyp,Top}";
-by(Auto_tac);
-by(rename_tac "t" 1);
-by(exhaust_tac "t" 1);
-by(Auto_tac);
+by (Auto_tac);
+by (rename_tac "t" 1);
+by (exhaust_tac "t" 1);
+by (Auto_tac);
 qed "UNIV_typ_enum";
 
 Goal "finite(UNIV::typ set)";
-by(simp_tac (simpset() addsimps [UNIV_typ_enum]) 1);
+by (simp_tac (simpset() addsimps [UNIV_typ_enum]) 1);
 qed "finite_UNIV_typ";
 AddIffs [finite_UNIV_typ];
--- a/src/HOL/Map.ML	Thu Oct 28 14:00:25 1999 +0200
+++ b/src/HOL/Map.ML	Thu Oct 28 14:55:23 1999 +0200
@@ -54,8 +54,8 @@
 qed_spec_mp "map_of_SomeD";
 
 Goal "[| map_of xs k = Some z; P z |] ==> map_of [(k,z):xs . P z] k = Some z";
-br mp 1;
-ba  2;
+by (rtac mp 1);
+by (assume_tac 2);
 by (etac thin_rl 1);
 by (induct_tac "xs" 1);
 by  Auto_tac;