tuned version by Stephan Merz (unbatchified etc.);
authorwenzelm
Thu, 03 Aug 2000 19:29:03 +0200
changeset 9517 f58863b1406a
parent 9516 72b5d28aae58
child 9518 0c8422ed066f
tuned version by Stephan Merz (unbatchified etc.);
src/HOL/TLA/Action.ML
src/HOL/TLA/Action.thy
src/HOL/TLA/Buffer/Buffer.ML
src/HOL/TLA/Inc/Inc.ML
src/HOL/TLA/Inc/Inc.thy
src/HOL/TLA/IntLemmas.ML
src/HOL/TLA/Intensional.ML
src/HOL/TLA/Intensional.thy
src/HOL/TLA/Memory/MIlive.ML
src/HOL/TLA/Memory/MIsafe.ML
src/HOL/TLA/Memory/MemClerk.ML
src/HOL/TLA/Memory/MemClerk.thy
src/HOL/TLA/Memory/Memory.ML
src/HOL/TLA/Memory/Memory.thy
src/HOL/TLA/Memory/MemoryImplementation.ML
src/HOL/TLA/Memory/MemoryImplementation.thy
src/HOL/TLA/Memory/MemoryParameters.ML
src/HOL/TLA/Memory/MemoryParameters.thy
src/HOL/TLA/Memory/ProcedureInterface.ML
src/HOL/TLA/Memory/ProcedureInterface.thy
src/HOL/TLA/Memory/RPC.ML
src/HOL/TLA/Memory/RPC.thy
src/HOL/TLA/Memory/RPCParameters.ML
src/HOL/TLA/Memory/RPCParameters.thy
src/HOL/TLA/TLA.ML
src/HOL/TLA/TLA.thy
--- a/src/HOL/TLA/Action.ML	Thu Aug 03 19:28:37 2000 +0200
+++ b/src/HOL/TLA/Action.ML	Thu Aug 03 19:29:03 2000 +0200
@@ -9,11 +9,13 @@
 (* The following assertion specializes "intI" for any world type 
    which is a pair, not just for "state * state".
 *)
-qed_goal "actionI" Action.thy "(!!s t. (s,t) |= A) ==> |- A"
-  (fn [prem] => [REPEAT (resolve_tac [prem,intI,prod_induct] 1)]);
+val [prem] = goal thy "(!!s t. (s,t) |= A) ==> |- A";
+by (REPEAT (resolve_tac [prem,intI,prod_induct] 1));
+qed "actionI";
 
-qed_goal "actionD" Action.thy "|- A ==> (s,t) |= A"
-  (fn [prem] => [rtac (prem RS intD) 1]);
+Goal "|- A ==> (s,t) |= A";
+by (etac intD 1);
+qed "actionD";
 
 local
   fun prover s = prove_goal Action.thy s 
@@ -57,166 +59,154 @@
 
 (* ===================== Update simpset and classical prover ============================= *)
 
-(***
-(* Make the simplifier use action_use rather than int_use
-   when action simplifications are added.
-*)
-
-let
-  val ss = simpset_ref()
-  fun try_rewrite th = 
-      (action_rewrite th) handle _ => (action_use th) handle _ => th
-in
-  ss := !ss setmksimps ((mksimps mksimps_pairs) o try_rewrite)
-end;
-***)
-
 AddSIs [actionI];
 AddDs  [actionD];
 
 (* =========================== square / angle brackets =========================== *)
 
-qed_goalw "idle_squareI" Action.thy [square_def]
-   "!!s t. (s,t) |= unchanged v ==> (s,t) |= [A]_v"
-   (fn _ => [ Asm_full_simp_tac 1 ]);
+Goalw [square_def] "(s,t) |= unchanged v ==> (s,t) |= [A]_v";
+by (Asm_full_simp_tac 1);
+qed "idle_squareI";
+
+Goalw [square_def] "(s,t) |= A ==> (s,t) |= [A]_v";
+by (Asm_simp_tac 1);
+qed "busy_squareI";
 
-qed_goalw "busy_squareI" Action.thy [square_def]
-   "!!s t. (s,t) |= A ==> (s,t) |= [A]_v"
-   (fn _ => [ Asm_simp_tac 1 ]);
-
-qed_goal "squareE" Action.thy
-  "[| (s,t) |= [A]_v; A (s,t) ==> B (s,t); v t = v s ==> B (s,t) |] ==> B (s,t)"
-  (fn prems => [cut_facts_tac prems 1,
-                rewrite_goals_tac (square_def::action_rews),
-                etac disjE 1,
-                REPEAT (eresolve_tac prems 1)]);
+val prems = goal thy
+  "[| (s,t) |= [A]_v; A (s,t) ==> B (s,t); v t = v s ==> B (s,t) |] ==> B (s,t)";
+by (cut_facts_tac prems 1);
+by (rewrite_goals_tac (square_def::action_rews));
+by (etac disjE 1);
+by (REPEAT (eresolve_tac prems 1));
+qed "squareE";
 
-qed_goalw "squareCI" Action.thy (square_def::action_rews)
-  "[| v t ~= v s ==> A (s,t) |] ==> (s,t) |= [A]_v"
-  (fn prems => [rtac disjCI 1,
-                eresolve_tac prems 1]);
+val prems = goalw thy (square_def::action_rews)
+  "[| v t ~= v s ==> A (s,t) |] ==> (s,t) |= [A]_v";
+by (rtac disjCI 1);
+by (eresolve_tac prems 1);
+qed "squareCI";
 
-qed_goalw "angleI" Action.thy [angle_def]
-  "!!s t. [| A (s,t); v t ~= v s |] ==> (s,t) |= <A>_v"
-  (fn _ => [ Asm_simp_tac 1 ]);
+goalw thy [angle_def]
+  "!!s t. [| A (s,t); v t ~= v s |] ==> (s,t) |= <A>_v";
+by (Asm_simp_tac 1);
+qed "angleI";
 
-qed_goalw "angleE" Action.thy (angle_def::action_rews)
-  "[| (s,t) |= <A>_v; [| A (s,t); v t ~= v s |] ==> R |] ==> R"
-  (fn prems => [cut_facts_tac prems 1,
-                etac conjE 1,
-                REPEAT (ares_tac prems 1)]);
+val prems = goalw thy (angle_def::action_rews)
+  "[| (s,t) |= <A>_v; [| A (s,t); v t ~= v s |] ==> R |] ==> R";
+by (cut_facts_tac prems 1);
+by (etac conjE 1);
+by (REPEAT (ares_tac prems 1));
+qed "angleE";
 
 AddIs [angleI, squareCI];
 AddEs [angleE, squareE];
 
-qed_goal "square_simulation" Action.thy
+goal thy
    "!!f. [| |- unchanged f & ~B --> unchanged g;   \
 \           |- A & ~unchanged g --> B              \
-\        |] ==> |- [A]_f --> [B]_g"
-   (fn _ => [Clarsimp_tac 1,
-             etac squareE 1,
-             auto_tac (claset(), simpset() addsimps [square_def])
-            ]);
+\        |] ==> |- [A]_f --> [B]_g";
+by (Clarsimp_tac 1);
+by (etac squareE 1);
+by (auto_tac (claset(), simpset() addsimps [square_def]));
+qed "square_simulation";
 
-qed_goalw "not_square" Action.thy [square_def,angle_def]
-   "|- (~ [A]_v) = <~A>_v"
-   (fn _ => [ Auto_tac ]);
+goalw thy [square_def,angle_def]
+   "|- (~ [A]_v) = <~A>_v";
+by Auto_tac;
+qed "not_square";
 
-qed_goalw "not_angle" Action.thy [square_def,angle_def]
-   "|- (~ <A>_v) = [~A]_v"
-   (fn _ => [ Auto_tac ]);
+goalw thy [square_def,angle_def]
+   "|- (~ <A>_v) = [~A]_v";
+by Auto_tac;
+qed "not_angle";
 
 (* ============================== Facts about ENABLED ============================== *)
 
-qed_goal "enabledI" Action.thy
-  "|- A --> $Enabled A"
-  (fn _ => [ auto_tac (claset(), simpset() addsimps [enabled_def]) ]);
+goal thy "|- A --> $Enabled A";
+by (auto_tac (claset(), simpset() addsimps [enabled_def]));
+qed "enabledI";
 
-qed_goalw "enabledE" Action.thy [enabled_def]
-  "[| s |= Enabled A; !!u. A (s,u) ==> Q |] ==> Q"
-  (fn prems => [cut_facts_tac prems 1,
-                etac exE 1,
-                resolve_tac prems 1, atac 1
-               ]);
+val prems = goalw thy [enabled_def]
+  "[| s |= Enabled A; !!u. A (s,u) ==> Q |] ==> Q";
+by (cut_facts_tac prems 1);
+by (etac exE 1);
+by (resolve_tac prems 1);
+by (atac 1);
+qed "enabledE";
 
-qed_goal "notEnabledD" Action.thy
-  "|- ~$Enabled G --> ~ G"
-  (fn _ => [ auto_tac (claset(), simpset() addsimps [enabled_def]) ]);
+goal thy "|- ~$Enabled G --> ~ G";
+by (auto_tac (claset(), simpset() addsimps [enabled_def]));
+qed "notEnabledD";
 
 (* Monotonicity *)
-qed_goal "enabled_mono" Action.thy
-  "[| s |= Enabled F; |- F --> G |] ==> s |= Enabled G"
-  (fn [min,maj] => [rtac (min RS enabledE) 1,
-                    rtac (action_use enabledI) 1,
-                    etac (action_use maj) 1
-                   ]);
+val [min,maj] = goal thy
+  "[| s |= Enabled F; |- F --> G |] ==> s |= Enabled G";
+by (rtac (min RS enabledE) 1);
+by (rtac (action_use enabledI) 1);
+by (etac (action_use maj) 1);
+qed "enabled_mono";
 
 (* stronger variant *)
-qed_goal "enabled_mono2" Action.thy
-   "[| s |= Enabled F; !!t. F (s,t) ==> G (s,t) |] ==> s |= Enabled G"
-   (fn [min,maj] => [rtac (min RS enabledE) 1,
-		     rtac (action_use enabledI) 1,
-		     etac maj 1
-		    ]);
+val [min,maj] = goal thy
+  "[| s |= Enabled F; !!t. F (s,t) ==> G (s,t) |] ==> s |= Enabled G";
+by (rtac (min RS enabledE) 1);
+by (rtac (action_use enabledI) 1);
+by (etac maj 1);
+qed "enabled_mono2";
 
-qed_goal "enabled_disj1" Action.thy
-  "|- Enabled F --> Enabled (F | G)"
-  (fn _ => [ auto_tac (claset() addSEs [enabled_mono], simpset()) ]);
+goal thy "|- Enabled F --> Enabled (F | G)";
+by (auto_tac (claset() addSEs [enabled_mono], simpset()));
+qed "enabled_disj1";
 
-qed_goal "enabled_disj2" Action.thy
-  "|- Enabled G --> Enabled (F | G)"
-  (fn _ => [ auto_tac (claset() addSEs [enabled_mono], simpset()) ]);
+goal thy "|- Enabled G --> Enabled (F | G)";
+by (auto_tac (claset() addSEs [enabled_mono], simpset()));
+qed "enabled_disj2";
 
-qed_goal "enabled_conj1" Action.thy
-  "|- Enabled (F & G) --> Enabled F"
-  (fn _ => [ auto_tac (claset() addSEs [enabled_mono], simpset()) ]);
+goal thy "|- Enabled (F & G) --> Enabled F";
+by (auto_tac (claset() addSEs [enabled_mono], simpset()));
+qed "enabled_conj1";
 
-qed_goal "enabled_conj2" Action.thy
-  "|- Enabled (F & G) --> Enabled G"
-  (fn _ => [ auto_tac (claset() addSEs [enabled_mono], simpset()) ]);
+goal thy "|- Enabled (F & G) --> Enabled G";
+by (auto_tac (claset() addSEs [enabled_mono], simpset()));
+qed "enabled_conj2";
 
-qed_goal "enabled_conjE" Action.thy
-  "[| s |= Enabled (F & G); [| s |= Enabled F; s |= Enabled G |] ==> Q |] ==> Q"
-  (fn prems => [cut_facts_tac prems 1, resolve_tac prems 1,
-                etac (action_use enabled_conj1) 1, 
-		etac (action_use enabled_conj2) 1
-	       ]);
-
-qed_goal "enabled_disjD" Action.thy
-  "|- Enabled (F | G) --> Enabled F | Enabled G"
-  (fn _ => [ auto_tac (claset(), simpset() addsimps [enabled_def]) ]);
+val prems = goal thy
+  "[| s |= Enabled (F & G); [| s |= Enabled F; s |= Enabled G |] ==> Q |] ==> Q";
+by (cut_facts_tac prems 1);
+by (resolve_tac prems 1);
+by (etac (action_use enabled_conj1) 1);
+by (etac (action_use enabled_conj2) 1);
+qed "enabled_conjE";
 
-qed_goal "enabled_disj" Action.thy
-  "|- Enabled (F | G) = (Enabled F | Enabled G)"
-  (fn _ => [Clarsimp_tac 1,
-	    rtac iffI 1,
-            etac (action_use enabled_disjD) 1,
-            REPEAT (eresolve_tac (disjE::map action_use [enabled_disj1,enabled_disj2]) 1)
-           ]);
+goal thy "|- Enabled (F | G) --> Enabled F | Enabled G";
+by (auto_tac (claset(), simpset() addsimps [enabled_def]));
+qed "enabled_disjD";
 
-qed_goal "enabled_ex" Action.thy
-  "|- Enabled (? x. F x) = (? x. Enabled (F x))"
-  (fn _ => [ force_tac (claset(), simpset() addsimps [enabled_def]) 1 ]);
+goal thy "|- Enabled (F | G) = (Enabled F | Enabled G)";
+by (Clarsimp_tac 1);
+by (rtac iffI 1);
+by (etac (action_use enabled_disjD) 1);
+by (REPEAT (eresolve_tac (disjE::map action_use [enabled_disj1,enabled_disj2]) 1));
+qed "enabled_disj";
+
+goal thy "|- Enabled (EX x. F x) = (EX x. Enabled (F x))";
+by (force_tac (claset(), simpset() addsimps [enabled_def]) 1);
+qed "enabled_ex";
 
 
-(* A rule that combines enabledI and baseE, but generates fewer possible instantiations *)
-qed_goal "base_enabled" Action.thy
-  "[| basevars vs; ? c. ! u. vs u = c --> A(s,u) |] ==> s |= Enabled A"
-  (fn prems => [cut_facts_tac prems 1,
-		etac exE 1, etac baseE 1, 
-                rtac (action_use enabledI) 1,
-                etac allE 1, etac mp 1, atac 1
-               ]);
-                
-(*** old version immediately generates schematic variable
-qed_goal "base_enabled" Action.thy
-  "[| basevars vs; !!u. vs u = c s ==> A (s,u) |] ==> s |= Enabled A"
-  (fn prems => [cut_facts_tac prems 1,
-		etac baseE 1, rtac (action_use enabledI) 1,
-		REPEAT (ares_tac prems 1)]);
-***)
+(* A rule that combines enabledI and baseE, but generates fewer instantiations *)
+val prems = goal thy
+  "[| basevars vs; EX c. ! u. vs u = c --> A(s,u) |] ==> s |= Enabled A";
+by (cut_facts_tac prems 1);
+by (etac exE 1);
+by (etac baseE 1);
+by (rtac (action_use enabledI) 1);
+by (etac allE 1);
+by (etac mp 1);
+by (atac 1);
+qed "base_enabled";
 
-(* ================================ action_simp_tac ================================== *)
+(* ======================= action_simp_tac ============================== *)
 
 (* A dumb simplification-based tactic with just a little first-order logic:
    should plug in only "very safe" rules that can be applied blindly.
@@ -244,16 +234,6 @@
    - Solve for the unknowns using standard HOL reasoning.
    The following tactic combines these steps except the final one.
 *)
-(*** old version
-fun enabled_tac base_vars i =
-    EVERY [(* apply actionI (plus rewriting) if the goal is of the form $(Enabled A),
-	      do nothing if it is of the form s |= Enabled A *)
-	   TRY ((resolve_tac [actionI,intI] i) 
-                THEN (SELECT_GOAL (rewrite_goals_tac action_rews) i)),
-	   clarify_tac (claset() addSIs [base_vars RS base_enabled]) i,
-	   (SELECT_GOAL (rewrite_goals_tac action_rews) i)
-	  ];
-***)
 
 fun enabled_tac base_vars =
     clarsimp_tac (claset() addSIs [base_vars RS base_enabled], simpset());
--- a/src/HOL/TLA/Action.thy	Thu Aug 03 19:28:37 2000 +0200
+++ b/src/HOL/TLA/Action.thy	Thu Aug 03 19:29:03 2000 +0200
@@ -50,9 +50,9 @@
 translations
   "ACT A"            =>   "(A::state*state => _)"
   "_before"          ==   "before"
-  "_after"           =>   "_prime"
+  "_after"           ==   "after"
+  "_prime"           =>   "_after"
   "_unchanged"       ==   "unch"
-  "_prime"           ==   "after"
   "_SqAct"           ==   "SqAct"
   "_AnAct"           ==   "AnAct"
   "_Enabled"         ==   "enabled"
@@ -63,7 +63,7 @@
 
 rules
   unl_before    "(ACT $v) (s,t) == v s"
-  unl_after     "(ACT v`) (s,t) == v t"
+  unl_after     "(ACT v$) (s,t) == v t"
 
   unchanged_def "(s,t) |= unchanged v == (v t = v s)"
   square_def    "ACT [A]_v == ACT (A | unchanged v)"
--- a/src/HOL/TLA/Buffer/Buffer.ML	Thu Aug 03 19:28:37 2000 +0200
+++ b/src/HOL/TLA/Buffer/Buffer.ML	Thu Aug 03 19:29:03 2000 +0200
@@ -26,7 +26,8 @@
 (* Enabling condition for dequeue -- NOT NEEDED *)
 Goalw [temp_rewrite Deq_visible]
    "!!q. basevars (ic,q,oc) ==> |- Enabled (<Deq ic q oc>_(ic,q,oc)) = (q ~= #[])";
-by (force_tac (claset() addSEs [base_enabled,enabledE], simpset() addsimps [Deq_def]) 1);
+by (force_tac (claset() addSEs [base_enabled,enabledE], 
+               simpset() addsimps [Deq_def]) 1);
 qed "Deq_enabled";
 
 (* For the left-to-right implication, we don't need the base variable stuff *)
--- a/src/HOL/TLA/Inc/Inc.ML	Thu Aug 03 19:28:37 2000 +0200
+++ b/src/HOL/TLA/Inc/Inc.ML	Thu Aug 03 19:29:03 2000 +0200
@@ -14,37 +14,46 @@
 
 (*** Invariant proof for Psi: "manual" proof proves individual lemmas ***)
 
-qed_goal "PsiInv_Init" Inc.thy "|- InitPsi --> PsiInv"
- (fn _ => [ auto_tac (Inc_css addsimps2 InitPsi_def::PsiInv_defs) ]);
+Goal "|- InitPsi --> PsiInv";
+by (auto_tac (Inc_css addsimps2 InitPsi_def::PsiInv_defs));
+qed "PsiInv_Init";
 
-qed_goal "PsiInv_alpha1" Inc.thy "|- alpha1 & $PsiInv --> PsiInv`"
-  (fn _ => [ auto_tac (Inc_css addsimps2 alpha1_def::PsiInv_defs) ]);
+Goal "|- alpha1 & $PsiInv --> PsiInv$";
+by (auto_tac (Inc_css addsimps2 alpha1_def::PsiInv_defs));
+qed "PsiInv_alpha1";
 
-qed_goal "PsiInv_alpha2" Inc.thy "|- alpha2 & $PsiInv --> PsiInv`"
-  (fn _ => [ auto_tac (Inc_css addsimps2 alpha2_def::PsiInv_defs) ]);
+Goal "|- alpha2 & $PsiInv --> PsiInv$";
+by (auto_tac (Inc_css addsimps2 alpha2_def::PsiInv_defs));
+qed "PsiInv_alpha2";
 
-qed_goal "PsiInv_beta1" Inc.thy "|- beta1 & $PsiInv --> PsiInv`"
-  (fn _ => [ auto_tac (Inc_css addsimps2 beta1_def::PsiInv_defs) ]);
+Goal "|- beta1 & $PsiInv --> PsiInv$";
+by (auto_tac (Inc_css addsimps2 beta1_def::PsiInv_defs));
+qed "PsiInv_beta1";
 
-qed_goal "PsiInv_beta2" Inc.thy "|- beta2 & $PsiInv --> PsiInv`"
-  (fn _ => [ auto_tac (Inc_css addsimps2 beta2_def::PsiInv_defs) ]);
+Goal "|- beta2 & $PsiInv --> PsiInv$";
+by (auto_tac (Inc_css addsimps2 beta2_def::PsiInv_defs));
+qed "PsiInv_beta2";
 
-qed_goal "PsiInv_gamma1" Inc.thy "|- gamma1 & $PsiInv --> PsiInv`"
-  (fn _ => [ auto_tac (Inc_css addsimps2 gamma1_def::PsiInv_defs) ]);
+Goal "|- gamma1 & $PsiInv --> PsiInv$";
+by (auto_tac (Inc_css addsimps2 gamma1_def::PsiInv_defs));
+qed "PsiInv_gamma1";
 
-qed_goal "PsiInv_gamma2" Inc.thy "|- gamma2 & $PsiInv --> PsiInv`"
-  (fn _ => [ auto_tac (Inc_css addsimps2 gamma2_def::PsiInv_defs) ]);
+Goal "|- gamma2 & $PsiInv --> PsiInv$";
+by (auto_tac (Inc_css addsimps2 gamma2_def::PsiInv_defs));
+qed "PsiInv_gamma2";
 
-qed_goal "PsiInv_stutter" Inc.thy "|- unchanged (x,y,sem,pc1,pc2) & $PsiInv --> PsiInv`"
-  (fn _ => [ auto_tac (Inc_css addsimps2 PsiInv_defs) ]);
+Goal "|- unchanged (x,y,sem,pc1,pc2) & $PsiInv --> PsiInv$";
+by (auto_tac (Inc_css addsimps2 PsiInv_defs));
+qed "PsiInv_stutter";
 
-qed_goal "PsiInv" Inc.thy "|- Psi --> []PsiInv" (K [
-	    inv_tac (Inc_css addsimps2 [Psi_def]) 1,
-	    force_tac (Inc_css addsimps2 [PsiInv_Init, Init_def]) 1,
-	    auto_tac (Inc_css addIs2
-		        [PsiInv_alpha1,PsiInv_alpha2,PsiInv_beta1,
-			 PsiInv_beta2,PsiInv_gamma1,PsiInv_gamma2,PsiInv_stutter]
-                        addsimps2 [square_def,N1_def, N2_def]) ]);
+Goal "|- Psi --> []PsiInv";
+by (inv_tac (Inc_css addsimps2 [Psi_def]) 1);
+ by (force_tac (Inc_css addsimps2 [PsiInv_Init, Init_def]) 1);
+by (auto_tac (Inc_css
+              addIs2 [PsiInv_alpha1,PsiInv_alpha2,PsiInv_beta1,
+                      PsiInv_beta2,PsiInv_gamma1,PsiInv_gamma2,PsiInv_stutter]
+              addsimps2 [square_def,N1_def, N2_def]));
+qed "PsiInv";
 
 (* Automatic proof works too, but it make take a while on a slow machine.
    More realistic examples require user guidance anyway.
@@ -56,13 +65,14 @@
 
 (**** Step simulation ****)
 
-qed_goal "Init_sim" Inc.thy "|- Psi --> Init InitPhi"
-  (fn _ => [ auto_tac (Inc_css addsimps2 [InitPhi_def,Psi_def,InitPsi_def,Init_def]) ]);
+Goal "|- Psi --> Init InitPhi";
+by (auto_tac (Inc_css addsimps2 [InitPhi_def,Psi_def,InitPsi_def,Init_def]));
+qed "Init_sim";
 
-qed_goal "Step_sim" Inc.thy "|- Psi --> [][M1 | M2]_(x,y)"
-  (fn _ => [auto_tac (Inc_css addsimps2 [square_def,M1_def,M2_def] @ Psi_defs
-                              addSEs2 [STL4E])
-           ]);
+Goal "|- Psi --> [][M1 | M2]_(x,y)";
+by (auto_tac (Inc_css addsimps2 [square_def,M1_def,M2_def] @ Psi_defs
+                      addSEs2 [STL4E]));
+qed "Step_sim";
 
 (**** Proof of fairness ****)
 
@@ -82,166 +92,152 @@
    the auxiliary lemmas are very similar.
 *)
 
-qed_goal "Stuck_at_b" Inc.thy
-  "|- [][(N1 | N2) & ~ beta1]_(x,y,sem,pc1,pc2) --> stable(pc1 = #b)"
-  (fn _ => [ auto_tac (Inc_css addSEs2 [Stable,squareE] addsimps2 Psi_defs) ]);
+Goal "|- [][(N1 | N2) & ~ beta1]_(x,y,sem,pc1,pc2) --> stable(pc1 = #b)";
+by (auto_tac (Inc_css addSEs2 [Stable,squareE] addsimps2 Psi_defs));
+qed "Stuck_at_b";
 
-qed_goal "N1_enabled_at_g" Inc.thy
-  "|- pc1 = #g --> Enabled (<N1>_(x,y,sem,pc1,pc2))"
-  (fn _ => [Clarsimp_tac 1,
-	    res_inst_tac [("F","gamma1")] enabled_mono 1,
-	    enabled_tac Inc_base 1,
-            force_tac (Inc_css addsimps2 [gamma1_def]) 1,
-	    force_tac (Inc_css addsimps2 [angle_def,gamma1_def,N1_def]) 1
-	   ]);
+Goal "|- pc1 = #g --> Enabled (<N1>_(x,y,sem,pc1,pc2))";
+by (Clarsimp_tac 1);
+by (res_inst_tac [("F","gamma1")] enabled_mono 1);
+by (enabled_tac Inc_base 1);
+ by (force_tac (Inc_css addsimps2 [gamma1_def]) 1);
+by (force_tac (Inc_css addsimps2 [angle_def,gamma1_def,N1_def]) 1);
+qed "N1_enabled_at_g";
 
-qed_goal "g1_leadsto_a1" Inc.thy
-  "|- [][(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2) & SF(N1)_(x,y,sem,pc1,pc2) & []#True \
-\     --> (pc1 = #g ~> pc1 = #a)"
-  (fn _ => [rtac SF1 1,
-	    action_simp_tac (simpset() addsimps Psi_defs) [] [squareE] 1,
-	    action_simp_tac (simpset() addsimps angle_def::Psi_defs) [] [] 1,
-	    (* reduce |- []A --> <>Enabled B  to  |- A --> Enabled B *)
-	    auto_tac (Inc_css addSIs2 [InitDmd_gen, N1_enabled_at_g]
-		              addSDs2 [STL2_gen]
-		              addsimps2 [Init_def])
-	   ]);
+Goal "|- [][(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2) & SF(N1)_(x,y,sem,pc1,pc2) & []#True \
+\        --> (pc1 = #g ~> pc1 = #a)";
+by (rtac SF1 1);
+by (action_simp_tac (simpset() addsimps Psi_defs) [] [squareE] 1);
+by (action_simp_tac (simpset() addsimps angle_def::Psi_defs) [] [] 1);
+(* reduce |- []A --> <>Enabled B  to  |- A --> Enabled B *)
+by (auto_tac (Inc_css addSIs2 [InitDmd_gen, N1_enabled_at_g]
+	              addSDs2 [STL2_gen]
+                      addsimps2 [Init_def]));
+qed "g1_leadsto_a1";
 
 (* symmetrical for N2, and similar for beta2 *)
-qed_goal "N2_enabled_at_g" Inc.thy
-  "|- pc2 = #g --> Enabled (<N2>_(x,y,sem,pc1,pc2))"
-  (fn _ => [Clarsimp_tac 1,
-	    res_inst_tac [("F","gamma2")] enabled_mono 1,
-	    enabled_tac Inc_base 1,
-            force_tac (Inc_css addsimps2 [gamma2_def]) 1,
-	    force_tac (Inc_css addsimps2 [angle_def,gamma2_def,N2_def]) 1
-	   ]);
+Goal "|- pc2 = #g --> Enabled (<N2>_(x,y,sem,pc1,pc2))";
+by (Clarsimp_tac 1);
+by (res_inst_tac [("F","gamma2")] enabled_mono 1);
+by (enabled_tac Inc_base 1);
+ by (force_tac (Inc_css addsimps2 [gamma2_def]) 1);
+by (force_tac (Inc_css addsimps2 [angle_def,gamma2_def,N2_def]) 1);
+qed "N2_enabled_at_g";
 
-qed_goal "g2_leadsto_a2" Inc.thy
-  "|- [][(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2) & SF(N2)_(x,y,sem,pc1,pc2) & []#True \
-\     --> (pc2 = #g ~> pc2 = #a)"
-  (fn _ => [rtac SF1 1,
-	    action_simp_tac (simpset() addsimps Psi_defs) [] [squareE] 1,
-	    action_simp_tac (simpset() addsimps angle_def::Psi_defs) [] [] 1,
-	    auto_tac (Inc_css addSIs2 [InitDmd_gen, N2_enabled_at_g]
-		              addSDs2 [STL2_gen]
-		              addsimps2 [Init_def])
-	   ]);
+Goal "|- [][(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2) & SF(N2)_(x,y,sem,pc1,pc2) & []#True \
+\        --> (pc2 = #g ~> pc2 = #a)";
+by (rtac SF1 1);
+by (action_simp_tac (simpset() addsimps Psi_defs) [] [squareE] 1);
+by (action_simp_tac (simpset() addsimps angle_def::Psi_defs) [] [] 1);
+by (auto_tac (Inc_css addSIs2 [InitDmd_gen, N2_enabled_at_g]
+	              addSDs2 [STL2_gen]
+                      addsimps2 [Init_def]));
+qed "g2_leadsto_a2";
 
-qed_goal "N2_enabled_at_b" Inc.thy
-  "|- pc2 = #b --> Enabled (<N2>_(x,y,sem,pc1,pc2))"
-  (fn _ => [Clarsimp_tac 1,
-	    res_inst_tac [("F","beta2")] enabled_mono 1,
-	    enabled_tac Inc_base 1,
-            force_tac (Inc_css addsimps2 [beta2_def]) 1,
-	    force_tac (Inc_css addsimps2 [angle_def,beta2_def,N2_def]) 1
-	   ]);
+Goal "|- pc2 = #b --> Enabled (<N2>_(x,y,sem,pc1,pc2))";
+by (Clarsimp_tac 1);
+by (res_inst_tac [("F","beta2")] enabled_mono 1);
+by (enabled_tac Inc_base 1);
+ by (force_tac (Inc_css addsimps2 [beta2_def]) 1);
+by (force_tac (Inc_css addsimps2 [angle_def,beta2_def,N2_def]) 1);
+qed "N2_enabled_at_b";
 
-qed_goal "b2_leadsto_g2" Inc.thy
-  "|- [][(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2) & SF(N2)_(x,y,sem,pc1,pc2) & []#True \
-\     --> (pc2 = #b ~> pc2 = #g)"
-  (fn _ => [rtac SF1 1,
-	    action_simp_tac (simpset() addsimps Psi_defs) [] [squareE] 1,
-	    action_simp_tac (simpset() addsimps angle_def::Psi_defs) [] [] 1,
-	    auto_tac (Inc_css addSIs2 [InitDmd_gen, N2_enabled_at_b]
-		              addSDs2 [STL2_gen]
-		              addsimps2 [Init_def])
-	   ]);
+Goal "|- [][(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2) & SF(N2)_(x,y,sem,pc1,pc2) & []#True \
+\        --> (pc2 = #b ~> pc2 = #g)";
+by (rtac SF1 1);
+by (action_simp_tac (simpset() addsimps Psi_defs) [] [squareE] 1);
+by (action_simp_tac (simpset() addsimps angle_def::Psi_defs) [] [] 1);
+by (auto_tac (Inc_css addSIs2 [InitDmd_gen, N2_enabled_at_b]
+                      addSDs2 [STL2_gen]
+                      addsimps2 [Init_def]));
+qed "b2_leadsto_g2";
 
 (* Combine above lemmas: the second component will eventually reach pc2 = a *)
-qed_goal "N2_leadsto_a" Inc.thy
-  "|- [][(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2) & SF(N2)_(x,y,sem,pc1,pc2) & []#True \
-\     --> (pc2 = #a | pc2 = #b | pc2 = #g ~> pc2 = #a)"
-  (fn _ => [auto_tac (Inc_css addSIs2 [LatticeDisjunctionIntro]),
-	    rtac (temp_use LatticeReflexivity) 1,
-	    rtac (temp_use LatticeTransitivity) 1,
-	    auto_tac (Inc_css addSIs2 [b2_leadsto_g2,g2_leadsto_a2])
-	   ]);
+Goal "|- [][(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2) & SF(N2)_(x,y,sem,pc1,pc2) & []#True \
+\        --> (pc2 = #a | pc2 = #b | pc2 = #g ~> pc2 = #a)";
+by (auto_tac (Inc_css addSIs2 [LatticeDisjunctionIntro]));
+by (rtac (temp_use LatticeReflexivity) 1);
+by (rtac (temp_use LatticeTransitivity) 1);
+by (auto_tac (Inc_css addSIs2 [b2_leadsto_g2,g2_leadsto_a2]));
+qed "N2_leadsto_a";
 
-(* Get rid of complete disjunction on the left-hand side of ~> above. *)
-qed_goal "N2_live" Inc.thy
-  "|- [][(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2) & SF(N2)_(x,y,sem,pc1,pc2) \
-\     --> <>(pc2 = #a)"
-  (fn _ => [auto_tac (Inc_css addsimps2 Init_defs
-                              addSIs2 [(temp_use N2_leadsto_a) 
-                                       RSN(2, (temp_use leadsto_init))]),
-	    case_tac "pc2 (st1 sigma)" 1,
-	    Auto_tac
-	   ]);
+(* Get rid of disjunction on the left-hand side of ~> above. *)
+Goal "|- [][(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2) & SF(N2)_(x,y,sem,pc1,pc2) \
+\        --> <>(pc2 = #a)";
+by (auto_tac (Inc_css addsimps2 Init_defs
+                      addSIs2 [(temp_use N2_leadsto_a) 
+                               RSN(2, (temp_use leadsto_init))]));
+by (case_tac "pc2 (st1 sigma)" 1);
+by Auto_tac;
+qed "N2_live";
 
 (* Now prove that the first component will eventually reach pc1 = b from pc1 = a *)
 
-qed_goal "N1_enabled_at_both_a" Inc.thy
-  "|- pc2 = #a & (PsiInv & pc1 = #a) --> Enabled (<N1>_(x,y,sem,pc1,pc2))"
-  (fn _ => [Clarsimp_tac 1,
-	    res_inst_tac [("F","alpha1")] enabled_mono 1,
-	    enabled_tac Inc_base 1,
-            force_tac (Inc_css addsimps2 (alpha1_def::PsiInv_defs)) 1,
-	    force_tac (Inc_css addsimps2 [angle_def,alpha1_def,N1_def]) 1
-	   ]);
+Goal "|- pc2 = #a & (PsiInv & pc1 = #a) --> Enabled (<N1>_(x,y,sem,pc1,pc2))";
+by (Clarsimp_tac 1);
+by (res_inst_tac [("F","alpha1")] enabled_mono 1);
+by (enabled_tac Inc_base 1);
+ by (force_tac (Inc_css addsimps2 (alpha1_def::PsiInv_defs)) 1);
+by (force_tac (Inc_css addsimps2 [angle_def,alpha1_def,N1_def]) 1);
+qed "N1_enabled_at_both_a";
 
-qed_goal "a1_leadsto_b1" Inc.thy
-  "|- []($PsiInv & [(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2))            \
-\           & SF(N1)_(x,y,sem,pc1,pc2) & [] SF(N2)_(x,y,sem,pc1,pc2)  \
-\     --> (pc1 = #a ~> pc1 = #b)"
-  (fn _ => [rtac SF1 1,
-            action_simp_tac (simpset() addsimps Psi_defs) [] [squareE] 1,
-            action_simp_tac (simpset() addsimps angle_def::Psi_defs) [] [] 1,
-	    clarsimp_tac (Inc_css addSIs2 [N1_enabled_at_both_a RS (temp_use DmdImpl)]) 1,
-	    auto_tac (Inc_css addSIs2 [BoxDmd2_simple, N2_live]
-		              addsimps2 split_box_conj::more_temp_simps)
-	   ]);
+Goal "|- []($PsiInv & [(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2))      \
+\        & SF(N1)_(x,y,sem,pc1,pc2) & [] SF(N2)_(x,y,sem,pc1,pc2)  \
+\        --> (pc1 = #a ~> pc1 = #b)";
+by (rtac SF1 1);
+by (action_simp_tac (simpset() addsimps Psi_defs) [] [squareE] 1);
+by (action_simp_tac (simpset() addsimps angle_def::Psi_defs) [] [] 1);
+by (clarsimp_tac (Inc_css addSIs2 [N1_enabled_at_both_a RS (temp_use DmdImpl)]) 1);
+by (auto_tac (Inc_css addSIs2 [BoxDmd2_simple, N2_live]
+	              addsimps2 split_box_conj::more_temp_simps));
+qed "a1_leadsto_b1";
 
 (* Combine the leadsto properties for N1: it will arrive at pc1 = b *)
 
-qed_goal "N1_leadsto_b" Inc.thy
-  "|- []($PsiInv & [(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2))             \
-\            & SF(N1)_(x,y,sem,pc1,pc2) & [] SF(N2)_(x,y,sem,pc1,pc2)  \
-\     --> (pc1 = #b | pc1 = #g | pc1 = #a ~> pc1 = #b)"
-  (fn _ => [auto_tac (Inc_css addSIs2 [LatticeDisjunctionIntro]),
-	    rtac (temp_use LatticeReflexivity) 1,
-	    rtac (temp_use LatticeTransitivity) 1,
-	    auto_tac (Inc_css addSIs2 [a1_leadsto_b1,g1_leadsto_a1]
-		              addsimps2 [split_box_conj])
-	   ]);
+Goal "|- []($PsiInv & [(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2))             \
+\        & SF(N1)_(x,y,sem,pc1,pc2) & [] SF(N2)_(x,y,sem,pc1,pc2)  \
+\        --> (pc1 = #b | pc1 = #g | pc1 = #a ~> pc1 = #b)";
+by (auto_tac (Inc_css addSIs2 [LatticeDisjunctionIntro]));
+by (rtac (temp_use LatticeReflexivity) 1);
+by (rtac (temp_use LatticeTransitivity) 1);
+by (auto_tac (Inc_css addSIs2 [a1_leadsto_b1,g1_leadsto_a1]
+	              addsimps2 [split_box_conj]));
+qed "N1_leadsto_b";
 
-qed_goal "N1_live" Inc.thy
-  "|- []($PsiInv & [(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2))             \
-\            & SF(N1)_(x,y,sem,pc1,pc2) & [] SF(N2)_(x,y,sem,pc1,pc2)  \
-\     --> <>(pc1 = #b)"
-  (fn _ => [auto_tac (Inc_css addsimps2 Init_defs
-                              addSIs2 [(temp_use N1_leadsto_b) 
-                                       RSN(2, temp_use leadsto_init)]),
-	    case_tac "pc1 (st1 sigma)" 1,
-	    Auto_tac
-	   ]);
+Goal "|- []($PsiInv & [(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2))             \
+\        & SF(N1)_(x,y,sem,pc1,pc2) & [] SF(N2)_(x,y,sem,pc1,pc2)  \
+\        --> <>(pc1 = #b)";
+by (auto_tac (Inc_css addsimps2 Init_defs
+                      addSIs2 [(temp_use N1_leadsto_b) 
+                               RSN(2, temp_use leadsto_init)]));
+by (case_tac "pc1 (st1 sigma)" 1);
+by Auto_tac;
+qed "N1_live";
 
-qed_goal "N1_enabled_at_b" Inc.thy
-  "|- pc1 = #b --> Enabled (<N1>_(x,y,sem,pc1,pc2))"
-  (fn _ => [Clarsimp_tac 1,
-	    res_inst_tac [("F","beta1")] enabled_mono 1,
-	    enabled_tac Inc_base 1,
-            force_tac (Inc_css addsimps2 [beta1_def]) 1,
-	    force_tac (Inc_css addsimps2 [angle_def,beta1_def,N1_def]) 1
-	   ]);
+Goal "|- pc1 = #b --> Enabled (<N1>_(x,y,sem,pc1,pc2))";
+by (Clarsimp_tac 1);
+by (res_inst_tac [("F","beta1")] enabled_mono 1);
+by (enabled_tac Inc_base 1);
+ by (force_tac (Inc_css addsimps2 [beta1_def]) 1);
+by (force_tac (Inc_css addsimps2 [angle_def,beta1_def,N1_def]) 1);
+qed "N1_enabled_at_b";
 
 (* Now assemble the bits and pieces to prove that Psi is fair. *)
 
-qed_goal "Fair_M1_lemma" Inc.thy
-  "|- []($PsiInv & [(N1 | N2)]_(x,y,sem,pc1,pc2))   \
-\     & SF(N1)_(x,y,sem,pc1,pc2) & []SF(N2)_(x,y,sem,pc1,pc2)  \
-\     --> SF(M1)_(x,y)"
-  (fn _ => [ res_inst_tac [("B","beta1"),("P","PRED pc1 = #b")] SF2 1,
-               (* action premises *)
-             force_tac (Inc_css addsimps2 [angle_def,M1_def,beta1_def]) 1,
-             force_tac (Inc_css addsimps2 angle_def::Psi_defs) 1,
-             force_tac (Inc_css addSEs2 [N1_enabled_at_b]) 1,
-               (* temporal premise: use previous lemmas and simple TL *)
-             force_tac (Inc_css addSIs2 [DmdStable, N1_live,Stuck_at_b] 
-                                addEs2 [STL4E] addsimps2 [square_def]) 1
-            ]);
+Goal "|- []($PsiInv & [(N1 | N2)]_(x,y,sem,pc1,pc2))   \
+\        & SF(N1)_(x,y,sem,pc1,pc2) & []SF(N2)_(x,y,sem,pc1,pc2)  \
+\        --> SF(M1)_(x,y)";
+by (res_inst_tac [("B","beta1"),("P","PRED pc1 = #b")] SF2 1);
+   (* action premises *)
+by (force_tac (Inc_css addsimps2 [angle_def,M1_def,beta1_def]) 1);
+by (force_tac (Inc_css addsimps2 angle_def::Psi_defs) 1);
+by (force_tac (Inc_css addSEs2 [N1_enabled_at_b]) 1);
+   (* temporal premise: use previous lemmas and simple TL *)
+by (force_tac (Inc_css addSIs2 [DmdStable, N1_live,Stuck_at_b] 
+                       addEs2 [STL4E] addsimps2 [square_def]) 1);
+qed "Fair_M1_lemma";
 
-qed_goal "Fair_M1" Inc.thy "|- Psi --> WF(M1)_(x,y)"
-  (fn _ => [auto_tac (Inc_css addSIs2 [SFImplWF, Fair_M1_lemma, PsiInv]
-		              addsimps2 [Psi_def,split_box_conj]@more_temp_simps)
-	   ]);
+Goal "|- Psi --> WF(M1)_(x,y)";
+by (auto_tac (Inc_css addSIs2 [SFImplWF, Fair_M1_lemma, PsiInv]
+		      addsimps2 [Psi_def,split_box_conj]@more_temp_simps));
+qed "Fair_M1";
--- a/src/HOL/TLA/Inc/Inc.thy	Thu Aug 03 19:28:37 2000 +0200
+++ b/src/HOL/TLA/Inc/Inc.thy	Thu Aug 03 19:29:03 2000 +0200
@@ -1,5 +1,5 @@
 (* 
-    File:        TLA/ex/inc/Inc.thy
+    File:        TLA/Inc/Inc.thy
     Author:      Stephan Merz
     Copyright:   1997 University of Munich
 
@@ -34,17 +34,17 @@
 
   (* definitions for high-level program *)
   InitPhi_def   "InitPhi == PRED x = # 0 & y = # 0"
-  M1_def        "M1      == ACT  x` = Suc<$x> & y` = $y"
-  M2_def        "M2      == ACT  y` = Suc<$y> & x` = $x"
+  M1_def        "M1      == ACT  x$ = Suc<$x> & y$ = $y"
+  M2_def        "M2      == ACT  y$ = Suc<$y> & x$ = $x"
   Phi_def       "Phi     == TEMP Init InitPhi & [][M1 | M2]_(x,y)
                                  & WF(M1)_(x,y) & WF(M2)_(x,y)"
 
   (* definitions for low-level program *)
   InitPsi_def   "InitPsi == PRED pc1 = #a & pc2 = #a
                                  & x = # 0 & y = # 0 & sem = # 1"
-  alpha1_def    "alpha1  == ACT  $pc1 = #a & pc1$ = #b & $sem = Suc<sem`> 
+  alpha1_def    "alpha1  == ACT  $pc1 = #a & pc1$ = #b & $sem = Suc<sem$> 
                                  & unchanged(x,y,pc2)"
-  alpha2_def    "alpha2  == ACT  $pc2 = #a & pc2$ = #b & $sem = Suc<sem`>
+  alpha2_def    "alpha2  == ACT  $pc2 = #a & pc2$ = #b & $sem = Suc<sem$>
                                  & unchanged(x,y,pc1)"
   beta1_def     "beta1   == ACT  $pc1 = #b & pc1$ = #g & x$ = Suc<$x>
                                  & unchanged(y,sem,pc2)"
--- a/src/HOL/TLA/IntLemmas.ML	Thu Aug 03 19:28:37 2000 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,361 +0,0 @@
-(* 
-    File:	 IntLemmas.ML
-    Author:      Stephan Merz
-    Copyright:   1998 University of Munich
-
-Lemmas and tactics for "intensional" logics. 
-
-Mostly a lifting of standard HOL lemmas. They are not required in standard
-reasoning about intensional logics, which starts by unlifting proof goals
-to the HOL level.
-*)
-
-
-qed_goal "substW" Intensional.thy
-  "[| |- x = y; w |= P(x) |] ==> w |= P(y)"
-  (fn [prem1,prem2] => [rtac (rewrite_rule ([prem1] RL [inteq_reflection]) prem2) 1]);
-                        
-
-(* Lift HOL rules to intensional reasoning *)
-
-qed_goal "reflW" Intensional.thy "|- x = x"
-  (fn _ => [Simp_tac 1]);
-
-qed_goal "symW" Intensional.thy "|- s = t  ==>  |- t = s"
-  (fn prems => [ cut_facts_tac prems 1,
-                 rtac intI 1, dtac intD 1,
-                 rewrite_goals_tac intensional_rews,
-                 etac sym 1 ]);
-
-qed_goal "not_symW" Intensional.thy "|- s ~= t  ==>  |- t ~= s"
-  (fn prems => [ cut_facts_tac prems 1,
-                 rtac intI 1, dtac intD 1,
-                 rewrite_goals_tac intensional_rews,
-                 etac not_sym 1 ]);
-
-qed_goal "transW" Intensional.thy 
-  "[| |- r = s; |- s = t |] ==> |- r = t"
-  (fn prems => [ cut_facts_tac prems 1,
-                 rtac intI 1, REPEAT (dtac intD 1),
-                 rewrite_goals_tac intensional_rews,
-                 etac trans 1,
-                 atac 1 ]);
-
-qed_goal "box_equalsW" Intensional.thy 
-   "[| |- a = b; |- a = c; |- b = d |] ==> |- c = d"
-   (fn prems => [ (rtac transW 1),
-                  (rtac transW 1),
-                  (rtac symW 1),
-                  (REPEAT (resolve_tac prems 1)) ]);
-
-
-(* NB: Antecedent is a standard HOL (non-intensional) formula. *)
-qed_goal "fun_congW" Intensional.thy 
-   "f = g ==> |- f<x> = g<x>"
-   (fn prems => [ cut_facts_tac prems 1,
-                  rtac intI 1,
-                  rewrite_goals_tac intensional_rews,
-                  etac fun_cong 1 ]);
-
-qed_goal "fun_cong2W" Intensional.thy 
-   "f = g ==> |- f<x,y> = g<x,y>"
-   (fn prems => [ cut_facts_tac prems 1,
-                  rtac intI 1,
-                  Asm_full_simp_tac 1 ]);
-
-qed_goal "fun_cong3W" Intensional.thy 
-   "f = g ==> |- f<x,y,z> = g<x,y,z>"
-   (fn prems => [ cut_facts_tac prems 1,
-                  rtac intI 1,
-                  Asm_full_simp_tac 1 ]);
-
-
-qed_goal "arg_congW" Intensional.thy "|- x = y ==> |- f<x> = f<y>"
-   (fn prems => [ cut_facts_tac prems 1,
-                  rtac intI 1,
-                  dtac intD 1,
-                  rewrite_goals_tac intensional_rews,
-                  etac arg_cong 1 ]);
-
-qed_goal "arg_cong2W" Intensional.thy 
-   "[| |- u = v; |- x = y |] ==> |- f<u,x> = f<v,y>"
-   (fn prems => [ cut_facts_tac prems 1,
-                  rtac intI 1,
-                  REPEAT (dtac intD 1),
-                  rewrite_goals_tac intensional_rews,
-                  REPEAT (etac subst 1),
-                  rtac refl 1 ]);
-
-qed_goal "arg_cong3W" Intensional.thy 
-   "[| |- r = s; |- u = v; |- x = y |] ==> |- f<r,u,x> = f<s,v,y>"
-   (fn prems => [ cut_facts_tac prems 1,
-                  rtac intI 1,
-                  REPEAT (dtac intD 1),
-                  rewrite_goals_tac intensional_rews,
-                  REPEAT (etac subst 1),
-                  rtac refl 1 ]);
-
-qed_goal "congW" Intensional.thy 
-   "[| f = g; |- x = y |] ==> |- f<x> = g<y>"
-   (fn prems => [ rtac box_equalsW 1,
-                  rtac reflW 3,
-                  rtac arg_congW 1,
-                  resolve_tac prems 1,
-                  rtac fun_congW 1,
-                  rtac sym 1,
-                  resolve_tac prems 1 ]);
-
-qed_goal "cong2W" Intensional.thy 
-   "[| f = g; |- u = v; |- x = y |] ==> |- f<u,x> = g<v,y>"
-   (fn prems => [ rtac box_equalsW 1,
-                  rtac reflW 3,
-                  rtac arg_cong2W 1,
-                  REPEAT (resolve_tac prems 1),
-                  rtac fun_cong2W 1,
-                  rtac sym 1,
-                  resolve_tac prems 1 ]);
-
-qed_goal "cong3W" Intensional.thy 
-   "[| f = g; |- r = s; |- u = v; |- x = y |] ==> |- f<r,u,x> = g<s,v,y>"
-   (fn prems => [ rtac box_equalsW 1,
-                  rtac reflW 3,
-                  rtac arg_cong3W 1,
-                  REPEAT (resolve_tac prems 1),
-                  rtac fun_cong3W 1,
-                  rtac sym 1,
-                  resolve_tac prems 1 ]);
-
-
-(** Lifted equivalence **)
-
-(* Note the object-level implication in the hypothesis. Meta-level implication
-   would be incorrect! *)
-qed_goal "iffIW" Intensional.thy 
-  "[| |- A --> B; |- B --> A |] ==> |- A = B"
-  (fn prems => [ cut_facts_tac prems 1,
-                 rewrite_goals_tac (Valid_def::intensional_rews),
-                 Blast_tac 1 ]);
-
-qed_goal "iffD2W" Intensional.thy 
-  "[| |- P = Q; w |= Q |] ==> w |= P"
- (fn prems => [ cut_facts_tac prems 1,
-	        rewrite_goals_tac (Valid_def::intensional_rews),
-                Blast_tac 1 ]);
-
-val iffD1W = symW RS iffD2W;
-
-(** #True **)
-
-qed_goal "eqTrueIW" Intensional.thy "|- P ==> |- P = #True"
-  (fn prems => [cut_facts_tac prems 1,
-                rtac intI 1,
-                dtac intD 1,
-		Asm_full_simp_tac 1]);
-
-qed_goal "eqTrueEW" Intensional.thy "|- P = #True ==> |- P"
-  (fn prems => [cut_facts_tac prems 1,
-                rtac intI 1,
-                dtac intD 1,
-		Asm_full_simp_tac 1]);
-
-(** #False **)
-
-qed_goal "FalseEW" Intensional.thy "|- #False ==> |- P"
-  (fn prems => [cut_facts_tac prems 1,
-                rtac intI 1,
-                dtac intD 1,
-                rewrite_goals_tac intensional_rews,
-                etac FalseE 1]);
-
-qed_goal "False_neq_TrueW" Intensional.thy 
- "|- #False = #True ==> |- P"
- (fn [prem] => [rtac (prem RS eqTrueEW RS FalseEW) 1]);
-
-
-(** Negation **)
-
-(* Again use object-level implication *)
-qed_goal "notIW" Intensional.thy "|- P --> #False ==> |- ~P"
-  (fn prems => [cut_facts_tac prems 1,
-		rewrite_goals_tac (Valid_def::intensional_rews),
-		Blast_tac 1]);
-
-qed_goal "notEWV" Intensional.thy 
-  "[| |- ~P; |- P |] ==> |- R"
-  (fn prems => [cut_facts_tac prems 1,
-		rtac intI 1,
-                REPEAT (dtac intD 1),
-                rewrite_goals_tac intensional_rews,
-                etac notE 1, atac 1]);
-
-(* The following rule is stronger: It is enough to detect an 
-   inconsistency at *some* world to conclude R. Note also that P and R
-   are allowed to be (intensional) formulas of different types! *)
-
-qed_goal "notEW" Intensional.thy 
-   "[| w |= ~P; w |= P |] ==> |- R"
-  (fn prems => [cut_facts_tac prems 1,
-                rtac intI 1,
-                rewrite_goals_tac intensional_rews,
-                etac notE 1, atac 1]);
-
-(** Implication **)
-
-qed_goal "impIW" Intensional.thy "(!!w. (w |= A) ==> (w |= B)) ==> |- A --> B"
-  (fn [prem] => [ rtac intI 1,
-                 rewrite_goals_tac intensional_rews,
-                 rtac impI 1,
-                 etac prem 1 ]);
-
-
-qed_goal "mpW" Intensional.thy "[| |- A --> B; w |= A |] ==> w |= B"
-   (fn prems => [ cut_facts_tac prems 1,
-                  dtac intD 1,
-                  rewrite_goals_tac intensional_rews,
-                  etac mp 1,
-                  atac 1 ]);
-
-qed_goal "impEW" Intensional.thy 
-  "[| |- A --> B; w |= A; w |= B ==> w |= C |] ==> w |= C"
-  (fn prems => [ (REPEAT (resolve_tac (prems@[mpW]) 1)) ]);
-
-qed_goal "rev_mpW" Intensional.thy "[| w |= P; |- P --> Q |] ==> w |= Q"
-  (fn prems => [ (REPEAT (resolve_tac (prems@[mpW]) 1)) ]);
-
-qed_goalw "contraposW" Intensional.thy intensional_rews
-  "[| w |= ~Q; |- P --> Q |] ==> w |= ~P"
-  (fn [major,minor] => [rtac (major RS contrapos) 1,
-                        etac rev_mpW 1,
-                        rtac minor 1]);
-
-qed_goal "iffEW" Intensional.thy
-    "[| |- P = Q; [| |- P --> Q; |- Q --> P |] ==> R |] ==> R"
- (fn [p1,p2] => [REPEAT(ares_tac([p1 RS iffD2W, p1 RS iffD1W, p2, impIW])1)]);
-
-
-(** Conjunction **)
-
-qed_goalw "conjIW" Intensional.thy intensional_rews "[| w |= P; w |= Q |] ==> w |= P & Q"
-  (fn prems => [REPEAT (resolve_tac ([conjI]@prems) 1)]);
-
-qed_goal "conjunct1W" Intensional.thy "(w |= P & Q) ==> w |= P"
-  (fn prems => [cut_facts_tac prems 1,
-                rewrite_goals_tac intensional_rews,
-                etac conjunct1 1]);
-
-qed_goal "conjunct2W" Intensional.thy "(w |= P & Q) ==> w |= Q"
-  (fn prems => [cut_facts_tac prems 1,
-                rewrite_goals_tac intensional_rews,
-                etac conjunct2 1]);
-
-qed_goal "conjEW" Intensional.thy 
-  "[| w |= P & Q; [| w |= P; w |= Q |] ==> w |= R |] ==> w |= R"
-  (fn prems => [cut_facts_tac prems 1, resolve_tac prems 1,
-	        etac conjunct1W 1, etac conjunct2W 1]);
-
-
-(** Disjunction **)
-
-qed_goalw "disjI1W" Intensional.thy intensional_rews "w |= P ==> w |= P | Q"
-  (fn [prem] => [REPEAT (resolve_tac [disjI1,prem] 1)]);
-
-qed_goalw "disjI2W" Intensional.thy intensional_rews "w |= Q ==> w |= P | Q"
-  (fn [prem] => [REPEAT (resolve_tac [disjI2,prem] 1)]);
-
-qed_goal "disjEW" Intensional.thy 
-         "[| w |= P | Q; |- P --> R; |- Q --> R |] ==> w |= R"
-  (fn prems => [cut_facts_tac prems 1,
-                REPEAT (dtac intD 1),
-                rewrite_goals_tac intensional_rews,
-		Blast_tac 1]);
-
-(** Classical propositional logic **)
-
-qed_goalw "classicalW" Intensional.thy (Valid_def::intensional_rews)
-  "!!P. |- ~P --> P  ==>  |- P"
-  (fn prems => [Blast_tac 1]);
-
-qed_goal "notnotDW" Intensional.thy "!!P. |- ~~P  ==>  |- P"
-  (fn prems => [rtac intI 1,
-                dtac intD 1,
-                rewrite_goals_tac intensional_rews,
-                etac notnotD 1]);
-
-qed_goal "disjCIW" Intensional.thy "!!P Q. (w |= ~Q --> P) ==> (w |= P|Q)"
-  (fn prems => [rewrite_goals_tac intensional_rews,
-                Blast_tac 1]);
-
-qed_goal "impCEW" Intensional.thy 
-   "[| |- P --> Q; (w |= ~P) ==> (w |= R); (w |= Q) ==> (w |= R) |] ==> w |= R"
-  (fn [a1,a2,a3] => 
-    [rtac (excluded_middle RS disjE) 1,
-     etac (rewrite_rule intensional_rews a2) 1,
-     rtac a3 1,
-     etac (a1 RS mpW) 1]);
-
-qed_goalw "iffCEW" Intensional.thy intensional_rews
-   "[| |- P = Q;      \
-\      [| (w |= P); (w |= Q) |] ==> (w |= R);   \
-\      [| (w |= ~P); (w |= ~Q) |] ==> (w |= R)  \
-\   |] ==> w |= R"
-   (fn [a1,a2,a3] =>
-      [rtac iffCE 1,
-       etac a2 2, atac 2,
-       etac a3 2, atac 2,
-       rtac (int_unlift a1) 1]);
-
-qed_goal "case_split_thmW" Intensional.thy 
-   "!!P. [| |- P --> Q; |- ~P --> Q |] ==> |- Q"
-  (fn _ => [rewrite_goals_tac (Valid_def::intensional_rews),
-	    Blast_tac 1]);
-
-fun case_tacW a = res_inst_tac [("P",a)] case_split_thmW;
-
-
-(** Rigid quantifiers **)
-
-qed_goal "allIW" Intensional.thy "(!!x. |- P x) ==> |- ! x. P(x)"
-  (fn [prem] => [rtac intI 1,
-                 rewrite_goals_tac intensional_rews,
-                 rtac allI 1,
-                 rtac (prem RS intD) 1]);
-
-qed_goal "specW" Intensional.thy "|- ! x. P x ==> |- P x"
-  (fn prems => [cut_facts_tac prems 1,
-                rtac intI 1,
-                dtac intD 1,
-                rewrite_goals_tac intensional_rews,
-                etac spec 1]);
-
-
-qed_goal "allEW" Intensional.thy 
-         "[| |- ! x. P x;  |- P x ==> |- R |] ==> |- R"
- (fn major::prems=>
-  [ (REPEAT (resolve_tac (prems @ [major RS specW]) 1)) ]);
-
-qed_goal "all_dupEW" Intensional.thy 
-    "[| |- ! x. P x;  [| |- P x; |- ! x. P x |] ==> |- R |] ==> |- R"
- (fn prems =>
-  [ (REPEAT (resolve_tac (prems @ (prems RL [specW])) 1)) ]);
-
-
-qed_goal "exIW" Intensional.thy "|- P x ==> |- ? x. P x"
-  (fn [prem] => [rtac intI 1,
-                 rewrite_goals_tac intensional_rews,
-                 rtac exI 1,
-                 rtac (prem RS intD) 1]);
-
-qed_goal "exEW" Intensional.thy 
-  "[| w |= ? x. P x; !!x. |- P x --> Q |] ==> w |= Q"
-  (fn [major,minor] => [rtac exE 1,
-                        rtac (rewrite_rule intensional_rews major) 1,
-                        etac rev_mpW 1,
-                        rtac minor 1]);
-
-(** Classical quantifier reasoning **)
-
-qed_goal "exCIW" Intensional.thy 
-  "!!P. w |= (! x. ~P x) --> P a ==> w |= ? x. P x"
-  (fn prems => [rewrite_goals_tac intensional_rews,
-                Blast_tac 1]);
-
--- a/src/HOL/TLA/Intensional.ML	Thu Aug 03 19:28:37 2000 +0200
+++ b/src/HOL/TLA/Intensional.ML	Thu Aug 03 19:29:03 2000 +0200
@@ -8,16 +8,19 @@
 
 val intensional_rews = [unl_con,unl_lift,unl_lift2,unl_lift3,unl_Rall,unl_Rex,unl_Rex1];
 
-qed_goalw "inteq_reflection" Intensional.thy  [Valid_def,unl_lift2]
-  "|- x=y  ==>  (x==y)"
-  (fn [prem] => [rtac eq_reflection 1, rtac ext 1, rtac (prem RS spec) 1 ]);
+Goalw [Valid_def,unl_lift2] "|- x=y  ==>  (x==y)";
+by (rtac eq_reflection 1);
+by (rtac ext 1);
+by (etac spec 1);
+qed "inteq_reflection";
 
-qed_goalw "intI" Intensional.thy [Valid_def] "(!!w. w |= A) ==> |- A"
-  (fn [prem] => [REPEAT (resolve_tac [allI,prem] 1)]);
+val [prem] = goalw thy [Valid_def] "(!!w. w |= A) ==> |- A";
+by (REPEAT (resolve_tac [allI,prem] 1));
+qed "intI";
 
-qed_goalw "intD" Intensional.thy [Valid_def] "|- A ==> w |= A"
-  (fn [prem] => [rtac (prem RS spec) 1]);
-
+Goalw [Valid_def] "|- A ==> w |= A";
+by (etac spec 1);
+qed "intD";
 
 (** Lift usual HOL simplifications to "intensional" level. **)
 local
@@ -45,11 +48,12 @@
    "|- (P | P) = P", "|- (P | ~P) = #True", "|- (~P | P) = #True",
    "|- (! x. P) = P", "|- (? x. P) = P", 
    "|- (~Q --> ~P) = (P --> Q)",
-   "|- (P|Q --> R) = ((P-->R)&(Q-->R))" ];
+   "|- (P|Q --> R) = ((P-->R)&(Q-->R))" ]
 end;
 
-qed_goal "TrueW" Intensional.thy "|- #True"
-  (fn _ => [simp_tac (simpset() addsimps [Valid_def,unl_con]) 1]);
+Goal "|- #True";
+by (simp_tac (simpset() addsimps [Valid_def,unl_con]) 1);
+qed "TrueW";
 
 Addsimps (TrueW::intensional_rews);
 Addsimps int_simps;
@@ -109,33 +113,13 @@
               ((flatten (int_unlift th)) handle _ => th)
     | _ => th;
 
-(***
-(* Make the simplifier accept "intensional" goals by either turning them into
-   a meta-equality or by unlifting them.
-*)
-
-let 
-  val ss = simpset_ref()
-  fun try_rewrite th = (int_rewrite th) handle _ => (int_use th) handle _ => th
-in 
-  ss := !ss setmksimps ((mksimps mksimps_pairs) o try_rewrite)
-end;
-***)
-
 (* ========================================================================= *)
 
-qed_goal "Not_Rall" Intensional.thy
-   "|- (~(! x. F x)) = (? x. ~F x)"
-   (fn _ => [simp_tac (simpset() addsimps [Valid_def]) 1]);
-
-qed_goal "Not_Rex" Intensional.thy
-   "|- (~ (? x. F x)) = (! x. ~ F x)"
-   (fn _ => [simp_tac (simpset() addsimps [Valid_def]) 1]);
+Goalw [Valid_def] "|- (~(! x. F x)) = (? x. ~F x)";
+by (Simp_tac 1);
+qed "Not_Rall";
 
-(* IntLemmas.ML contains a collection of further lemmas about "intensional" logic.
-   These are not loaded by default because they are not required for the
-   standard proof procedures that first unlift proof goals to the HOL level.
+Goalw [Valid_def] "|- (~ (? x. F x)) = (! x. ~ F x)";
+by (Simp_tac 1);
+qed "Not_Rex";
 
-use "IntLemmas.ML";
-
-*)
--- a/src/HOL/TLA/Intensional.thy	Thu Aug 03 19:28:37 2000 +0200
+++ b/src/HOL/TLA/Intensional.thy	Thu Aug 03 19:29:03 2000 +0200
@@ -82,17 +82,17 @@
   (** TODO: syntax for lifted collection / comprehension **)
   "_liftPair"   :: [lift,liftargs] => lift                   ("(1'(_,/ _'))")
   (* infix syntax for list operations *)
-  "_liftCons" :: [lift, lift] => lift                    ("(_ #/ _)" [65,66] 65)
-  "_liftApp"  :: [lift, lift] => lift                    ("(_ @/ _)" [65,66] 65)
-  "_liftList" :: liftargs => lift                        ("[(_)]")
+  "_liftCons" :: [lift, lift] => lift                  ("(_ #/ _)" [65,66] 65)
+  "_liftApp"  :: [lift, lift] => lift                  ("(_ @/ _)" [65,66] 65)
+  "_liftList" :: liftargs => lift                      ("[(_)]")
 
   (* Rigid quantification (syntax level) *)
-  "_RAll"  :: [idts, lift] => lift                     ("(3! _./ _)" [0, 10] 10)
-  "_REx"   :: [idts, lift] => lift                     ("(3? _./ _)" [0, 10] 10)
-  "_REx1"  :: [idts, lift] => lift                     ("(3?! _./ _)" [0, 10] 10)
-  "_ARAll" :: [idts, lift] => lift                     ("(3ALL _./ _)" [0, 10] 10)
-  "_AREx"  :: [idts, lift] => lift                     ("(3EX _./ _)" [0, 10] 10)
-  "_AREx1" :: [idts, lift] => lift                     ("(3EX! _./ _)" [0, 10] 10)
+  "_ARAll"  :: [idts, lift] => lift                    ("(3! _./ _)" [0, 10] 10)
+  "_AREx"   :: [idts, lift] => lift                    ("(3? _./ _)" [0, 10] 10)
+  "_AREx1"  :: [idts, lift] => lift                    ("(3?! _./ _)" [0, 10] 10)
+  "_RAll" :: [idts, lift] => lift                      ("(3ALL _./ _)" [0, 10] 10)
+  "_REx"  :: [idts, lift] => lift                      ("(3EX _./ _)" [0, 10] 10)
+  "_REx1" :: [idts, lift] => lift                      ("(3EX! _./ _)" [0, 10] 10)
 
 translations
   "_const"        == "const"
@@ -142,9 +142,9 @@
   "w |= A | B"    <= "_liftOr A B w"
   "w |= A --> B"  <= "_liftImp A B w"
   "w |= u = v"    <= "_liftEqu u v w"
-  "w |= ! x. A"   <= "_RAll x A w"
-  "w |= ? x. A"   <= "_REx x A w"
-  "w |= ?! x. A"  <= "_REx1 x A w"
+  "w |= ALL x. A"   <= "_RAll x A w"
+  "w |= EX x. A"   <= "_REx x A w"
+  "w |= EX! x. A"  <= "_REx1 x A w"
 
 syntax (symbols)
   "_Valid"      :: lift => bool                        ("(\\<turnstile> _)" 5)
@@ -172,7 +172,7 @@
   unl_lift2   "LIFT f<x, y> w == f (x w) (y w)"
   unl_lift3   "LIFT f<x, y, z> w == f (x w) (y w) (z w)"
 
-  unl_Rall    "w |= ! x. A x  ==  ! x. (w |= A x)" 
-  unl_Rex     "w |= ? x. A x  ==  ? x. (w |= A x)"
-  unl_Rex1    "w |= ?! x. A x  ==  ?! x. (w |= A x)"
+  unl_Rall    "w |= ALL x. A x  ==  ALL x. (w |= A x)" 
+  unl_Rex     "w |= EX x. A x   ==  EX x. (w |= A x)"
+  unl_Rex1    "w |= EX! x. A x  ==  EX! x. (w |= A x)"
 end
--- a/src/HOL/TLA/Memory/MIlive.ML	Thu Aug 03 19:28:37 2000 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,365 +0,0 @@
-(* 
-    File:        MIlive.ML
-    Author:      Stephan Merz
-    Copyright:   1997 University of Munich
-
-    RPC-Memory example: Lower-level lemmas for the liveness proof
-*)
-
-(* Liveness assertions for the different implementation states, based on the
-   fairness conditions. Prove subgoals of WF1 / SF1 rules as separate lemmas
-   for readability. Reuse action proofs from safety part.
-*)
-
-(* ------------------------------ State S1 ------------------------------ *)
-
-qed_goal "S1_successors" MemoryImplementation.thy
-   "|- $S1 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)  \
-\      --> (S1 rmhist p)` | (S2 rmhist p)`"
-   (fn _ => [split_idle_tac [] 1,
-	     auto_tac (MI_css addSDs2 [Step1_2_1])
-	    ]);
-
-(* Show that the implementation can satisfy the high-level fairness requirements
-   by entering the state S1 infinitely often.
-*)
-
-qed_goal "S1_RNextdisabled" MemoryImplementation.thy
-   "|- S1 rmhist p --> \
-\      ~Enabled (<RNext memCh mm (resbar rmhist) p>_(rtrner memCh!p, resbar rmhist!p))"
-   (fn _ => [action_simp_tac (simpset() addsimps [angle_def,S_def,S1_def])
-	                     [notI] [enabledE,temp_elim Memoryidle] 1,
-	     Force_tac 1
-	    ]);
-
-qed_goal "S1_Returndisabled" MemoryImplementation.thy
-   "|- S1 rmhist p --> \
-\      ~Enabled (<MemReturn memCh (resbar rmhist) p>_(rtrner memCh!p, resbar rmhist!p))"
-   (fn _ => [action_simp_tac (simpset() addsimps [angle_def,MemReturn_def,Return_def,S_def,S1_def])
-	                     [notI] [enabledE] 1
-	    ]);
-
-qed_goal "RNext_fair" MemoryImplementation.thy
-   "|- []<>S1 rmhist p   \
-\      --> WF(RNext memCh mm (resbar rmhist) p)_(rtrner memCh!p, resbar rmhist!p)"
-   (fn _ => [auto_tac (MI_css addsimps2 [WF_alt]
-			      addSIs2 [S1_RNextdisabled] addSEs2 [STL4E,DmdImplE])
-	    ]);
-
-qed_goal "Return_fair" MemoryImplementation.thy
-   "|- []<>S1 rmhist p   \
-\      --> WF(MemReturn memCh (resbar rmhist) p)_(rtrner memCh!p, resbar rmhist!p)"
-   (fn _ => [auto_tac (MI_css addsimps2 [WF_alt]
-			      addSIs2 [S1_Returndisabled] addSEs2 [STL4E,DmdImplE])
-	    ]);
-
-(* ------------------------------ State S2 ------------------------------ *)
-
-qed_goal "S2_successors" MemoryImplementation.thy
-   "|- $S2 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)   \
-\      --> (S2 rmhist p)` | (S3 rmhist p)`"
-   (fn _ => [split_idle_tac [] 1,
-	     auto_tac (MI_css addSDs2 [Step1_2_2])
-	    ]);
-
-qed_goal "S2MClkFwd_successors" MemoryImplementation.thy
-   "|- ($S2 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p))    \
-\      & <MClkFwd memCh crCh cst p>_(c p) \
-\      --> (S3 rmhist p)`"
-   (fn _ => [ auto_tac (MI_css addsimps2 [angle_def] addSDs2 [Step1_2_2]) ]);
-
-qed_goal "S2MClkFwd_enabled" MemoryImplementation.thy
-   "|- $S2 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)    \
-\      --> $Enabled (<MClkFwd memCh crCh cst p>_(c p))"
-   (fn _ => [auto_tac (MI_css addsimps2 [c_def] addSIs2 [MClkFwd_ch_enabled,MClkFwd_enabled]),
-             cut_facts_tac [MI_base] 1,
-             blast_tac (claset() addDs [base_pair]) 1,
-             ALLGOALS (asm_full_simp_tac (simpset() addsimps [S_def,S2_def]))
-	    ]);
-
-qed_goal "S2_live" MemoryImplementation.thy
-   "|- [](ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)) & WF(MClkFwd memCh crCh cst p)_(c p) \
-\      --> (S2 rmhist p ~> S3 rmhist p)"
-   (fn _ => [REPEAT (resolve_tac [WF1,S2_successors,
-				  S2MClkFwd_successors,S2MClkFwd_enabled] 1)
-	    ]);
-
-
-(* ------------------------------ State S3 ------------------------------ *)
-
-qed_goal "S3_successors" MemoryImplementation.thy
-   "|- $S3 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)   \
-\      --> (S3 rmhist p)` | (S4 rmhist p | S6 rmhist p)`"
-   (fn _ => [split_idle_tac [] 1,
-	     auto_tac (MI_css addSDs2 [Step1_2_3])
-	    ]);
-
-qed_goal "S3RPC_successors" MemoryImplementation.thy
-   "|- ($S3 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p))   \
-\      & <RPCNext crCh rmCh rst p>_(r p) \
-\      --> (S4 rmhist p | S6 rmhist p)`"
-   (fn _ => [ auto_tac (MI_css addsimps2 [angle_def] addSDs2 [Step1_2_3]) ]);
-
-qed_goal "S3RPC_enabled" MemoryImplementation.thy
-   "|- $S3 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)   \
-\      --> $Enabled (<RPCNext crCh rmCh rst p>_(r p))"
-   (fn _ => [auto_tac (MI_css addsimps2 [r_def]
-		              addSIs2 [RPCFail_Next_enabled,RPCFail_enabled]),
-	     cut_facts_tac [MI_base] 1,
-	     blast_tac (claset() addDs [base_pair]) 1,
-             ALLGOALS (asm_full_simp_tac (simpset() addsimps [S_def,S3_def]))
-	    ]);
-
-qed_goal "S3_live" MemoryImplementation.thy
-   "|- [](ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)) & WF(RPCNext crCh rmCh rst p)_(r p) \
-\   --> (S3 rmhist p ~> S4 rmhist p | S6 rmhist p)"
-   (fn _ => [REPEAT (resolve_tac [WF1,S3_successors,S3RPC_successors,S3RPC_enabled] 1)]);
-
-(* ------------- State S4 -------------------------------------------------- *)
-
-qed_goal "S4_successors" MemoryImplementation.thy
-   "|- $S4 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) \
-\                   & (!l. $MemInv mm l)  \
-\      --> (S4 rmhist p)` | (S5 rmhist p)`"
-   (fn _ => [split_idle_tac [] 1,
-	     auto_tac (MI_css addSDs2 [Step1_2_4])
-	    ]);
-
-(* ------------- State S4a: S4 /\ (ires p = NotAResult) ------------------------------ *)
-
-qed_goal "S4a_successors" MemoryImplementation.thy
-   "|- $(S4 rmhist p & ires!p = #NotAResult) \
-\      & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) & (!l. $MemInv mm l) \
-\      --> (S4 rmhist p & ires!p = #NotAResult)`  \
-\        | ((S4 rmhist p & ires!p ~= #NotAResult) | S5 rmhist p)`"
-   (fn _ => [split_idle_tac [m_def] 1,
-	     auto_tac (MI_css addSDs2 [Step1_2_4])
-	    ]);
-
-qed_goal "S4aRNext_successors" MemoryImplementation.thy
-   "|- ($(S4 rmhist p & ires!p = #NotAResult)  \
-\       & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) & (!l. $MemInv mm l))  \
-\      & <RNext rmCh mm ires p>_(m p) \
-\      --> ((S4 rmhist p & ires!p ~= #NotAResult) | S5 rmhist p)`"
-   (fn _ => [auto_tac (MI_css addsimps2 [angle_def]
-		              addSDs2 [Step1_2_4, ReadResult, WriteResult])
-	    ]);
-
-qed_goal "S4aRNext_enabled" MemoryImplementation.thy
-   "|- $(S4 rmhist p & ires!p = #NotAResult) \
-\      & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) & (!l. $MemInv mm l)  \
-\   --> $Enabled (<RNext rmCh mm ires p>_(m p))"
-   (fn _ => [auto_tac (MI_css addsimps2 [m_def] addSIs2 [RNext_enabled]),
-	     cut_facts_tac [MI_base] 1,
-	     blast_tac (claset() addDs [base_pair]) 1,
-	     asm_full_simp_tac (simpset() addsimps [S_def,S4_def]) 1
-	    ]);
-
-qed_goal "S4a_live" MemoryImplementation.thy
-  "|- [](ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) & (!l. $MemInv mm l)) \
-\     & WF(RNext rmCh mm ires p)_(m p) \
-\     --> (S4 rmhist p & ires!p = #NotAResult  \
-\          ~> (S4 rmhist p & ires!p ~= #NotAResult) | S5 rmhist p)"
-   (K [REPEAT (resolve_tac [WF1, S4a_successors, S4aRNext_successors, S4aRNext_enabled] 1)]);
-
-(* ------------- State S4b: S4 /\ (ires p # NotAResult) ------------------------------ *)
-
-qed_goal "S4b_successors" MemoryImplementation.thy
-   "|- $(S4 rmhist p & ires!p ~= #NotAResult)  \
-\      & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) & (!l. $MemInv mm l) \
-\      --> (S4 rmhist p & ires!p ~= #NotAResult)` | (S5 rmhist p)`"
-   (fn _ => [split_idle_tac [m_def] 1,
-	     auto_tac (MI_css addSDs2 [WriteResult,Step1_2_4,ReadResult])
-	    ]);
-
-qed_goal "S4bReturn_successors" MemoryImplementation.thy
-   "|- ($(S4 rmhist p & ires!p ~= #NotAResult)  \
-\       & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) & (!l. $MemInv mm l))   \
-\      & <MemReturn rmCh ires p>_(m p) \
-\      --> (S5 rmhist p)`"
-   (fn _ => [force_tac (MI_css addsimps2 [angle_def] addSDs2 [Step1_2_4]
-                               addDs2 [ReturnNotReadWrite]) 1
-	    ]);
-
-qed_goal "S4bReturn_enabled" MemoryImplementation.thy
-   "|- $(S4 rmhist p & ires!p ~= #NotAResult)  \
-\      & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) & (!l. $MemInv mm l)  \
-\      --> $Enabled (<MemReturn rmCh ires p>_(m p))"
-   (fn _ => [auto_tac (MI_css addsimps2 [m_def] addSIs2 [MemReturn_enabled]),
-	     cut_facts_tac [MI_base] 1,
-             blast_tac (claset() addDs [base_pair]) 1,
-	     asm_full_simp_tac (simpset() addsimps [S_def,S4_def]) 1
-	    ]);
-
-qed_goal "S4b_live" MemoryImplementation.thy
-  "|- [](ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) & (!l. $MemInv mm l)) \
-\     & WF(MemReturn rmCh ires p)_(m p) \
-\     --> (S4 rmhist p & ires!p ~= #NotAResult ~> S5 rmhist p)"
-   (K [REPEAT (resolve_tac [WF1, S4b_successors,S4bReturn_successors, S4bReturn_enabled] 1)]);
-
-(* ------------------------------ State S5 ------------------------------ *)
-
-qed_goal "S5_successors" MemoryImplementation.thy
-   "|- $S5 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) \
-\      --> (S5 rmhist p)` | (S6 rmhist p)`"
-   (fn _ => [split_idle_tac [] 1,
-	     auto_tac (MI_css addSDs2 [Step1_2_5])
-	    ]);
-
-qed_goal "S5RPC_successors" MemoryImplementation.thy
-   "|- ($S5 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)) \
-\     & <RPCNext crCh rmCh rst p>_(r p) \
-\     --> (S6 rmhist p)`"
-   (fn _ => [ auto_tac (MI_css addsimps2 [angle_def] addSDs2 [Step1_2_5]) ]);
-
-qed_goal "S5RPC_enabled" MemoryImplementation.thy
-   "|- $S5 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) \
-\      --> $Enabled (<RPCNext crCh rmCh rst p>_(r p))"
-   (fn _ => [auto_tac (MI_css addsimps2 [r_def]
-		              addSIs2 [RPCFail_Next_enabled, RPCFail_enabled]),
-	     cut_facts_tac [MI_base] 1,
-	     blast_tac (claset() addDs [base_pair]) 1,
-	     ALLGOALS (asm_full_simp_tac (simpset() addsimps [S_def,S5_def]))
-	    ]);
-
-qed_goal "S5_live" MemoryImplementation.thy
-   "|- [](ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p))   \
-\      & WF(RPCNext crCh rmCh rst p)_(r p) \
-\      --> (S5 rmhist p ~> S6 rmhist p)"
-   (fn _ => [REPEAT (resolve_tac [WF1,S5_successors,S5RPC_successors,S5RPC_enabled] 1)]);
-
-
-(* ------------------------------ State S6 ------------------------------ *)
-
-qed_goal "S6_successors" MemoryImplementation.thy
-   "|- $S6 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) \
-\      --> (S1 rmhist p)` | (S3 rmhist p)` | (S6 rmhist p)`"
-   (fn _ => [split_idle_tac [] 1,
-	     auto_tac (MI_css addSDs2 [Step1_2_6])
-	    ]);
-
-qed_goal "S6MClkReply_successors" MemoryImplementation.thy
-   "|- ($S6 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)) \
-\      & <MClkReply memCh crCh cst p>_(c p) \
-\      --> (S1 rmhist p)`"
-   (fn _ => [auto_tac (MI_css addsimps2 [angle_def] addSDs2 [Step1_2_6, MClkReplyNotRetry])
-	    ]);
-
-qed_goal "MClkReplyS6" MemoryImplementation.thy
-   "|- $ImpInv rmhist p & <MClkReply memCh crCh cst p>_(c p) --> $S6 rmhist p"
-   (fn _ => [action_simp_tac
-	        (simpset() addsimps
-		    [angle_def,MClkReply_def,Return_def,
-		     ImpInv_def,S_def,S1_def,S2_def,S3_def,S4_def,S5_def])
-		[] [] 1
-	    ]);
-
-qed_goal "S6MClkReply_enabled" MemoryImplementation.thy
-   "|- S6 rmhist p --> Enabled (<MClkReply memCh crCh cst p>_(c p))"
-   (fn _ => [auto_tac (MI_css addsimps2 [c_def] addSIs2 [MClkReply_enabled]),
-	     cut_facts_tac [MI_base] 1,
-	     blast_tac (claset() addDs [base_pair]) 1,
-	     ALLGOALS (action_simp_tac (simpset() addsimps [S_def,S6_def]) [] [])
-	    ]);
-
-qed_goal "S6_live" MemoryImplementation.thy
-   "|- [](ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) & $(ImpInv rmhist p)) \
-\      & SF(MClkReply memCh crCh cst p)_(c p) & []<>(S6 rmhist p)  \
-\      --> []<>(S1 rmhist p)"
-   (fn _ => [Clarsimp_tac 1,
-	     subgoal_tac "sigma |= []<>(<MClkReply memCh crCh cst p>_(c p))" 1,
-             etac InfiniteEnsures 1, atac 1,
-	     action_simp_tac (simpset()) []
-	                     (map temp_elim [MClkReplyS6,S6MClkReply_successors]) 1,
-	     auto_tac (MI_css addsimps2 [SF_def]),
-	     etac swap 1,
-	     auto_tac (MI_css addSIs2 [S6MClkReply_enabled] addSEs2 [STL4E, DmdImplE])
-	    ]);
-
-(* ------------------------------ complex leadsto properties ------------------------------ *)
-
-qed_goal "S5S6LeadstoS6" MemoryImplementation.thy
-   "!!sigma. sigma |= S5 rmhist p ~> S6 rmhist p \
-\      ==> sigma |= (S5 rmhist p | S6 rmhist p) ~> S6 rmhist p"
-   (fn _ => [auto_tac (MI_css addSIs2 [LatticeDisjunctionIntro, LatticeReflexivity])
-	    ]);
-
-qed_goal "S4bS5S6LeadstoS6" MemoryImplementation.thy
-   "!!sigma. [| sigma |= S4 rmhist p & ires!p ~= #NotAResult ~> S5 rmhist p;  \
-\               sigma |= S5 rmhist p ~> S6 rmhist p |]  \
-\      ==> sigma |= (S4 rmhist p & ires!p ~= #NotAResult) | S5 rmhist p | S6 rmhist p \
-\                   ~> S6 rmhist p"
-   (fn _ => [auto_tac (MI_css addSIs2 [LatticeDisjunctionIntro,S5S6LeadstoS6]
-		              addIs2 [LatticeTransitivity])
-            ]);
-
-qed_goal "S4S5S6LeadstoS6" MemoryImplementation.thy
-   "!!sigma. [| sigma |= S4 rmhist p & ires!p = #NotAResult \
-\                        ~> (S4 rmhist p & ires!p ~= #NotAResult) | S5 rmhist p; \
-\               sigma |= S4 rmhist p & ires!p ~= #NotAResult ~> S5 rmhist p;  \
-\               sigma |= S5 rmhist p ~> S6 rmhist p |]  \
-\      ==> sigma |= S4 rmhist p | S5 rmhist p | S6 rmhist p ~> S6 rmhist p"
-   (fn _ => [subgoal_tac "sigma |= (S4 rmhist p & ires!p = #NotAResult) | (S4 rmhist p & ires!p ~= #NotAResult) | S5 rmhist p | S6 rmhist p ~> S6 rmhist p" 1,
-	     eres_inst_tac [("G", "PRED ((S4 rmhist p & ires!p = #NotAResult) | (S4 rmhist p & ires!p ~= #NotAResult) | S5 rmhist p | S6 rmhist p)")] (temp_use LatticeTransitivity) 1,
-	     force_tac (MI_css addsimps2 Init_defs addSIs2 [ImplLeadsto_gen, necT]) 1,
-	     rtac (temp_use LatticeDisjunctionIntro) 1,
-	     etac (temp_use LatticeTransitivity) 1,
-	     etac (temp_use LatticeTriangle2) 1, atac 1,
-	     auto_tac (MI_css addSIs2 [S4bS5S6LeadstoS6])
-	    ]);
-
-qed_goal "S3S4S5S6LeadstoS6" MemoryImplementation.thy
-   "!!sigma. [| sigma |= S3 rmhist p ~> S4 rmhist p | S6 rmhist p;   \
-\               sigma |= S4 rmhist p & ires!p = #NotAResult \
-\                         ~> (S4 rmhist p & ires!p ~= #NotAResult) | S5 rmhist p; \
-\               sigma |= S4 rmhist p & ires!p ~= #NotAResult ~> S5 rmhist p;  \
-\               sigma |= S5 rmhist p ~> S6 rmhist p |]  \
-\      ==> sigma |= S3 rmhist p | S4 rmhist p | S5 rmhist p | S6 rmhist p ~> S6 rmhist p"
-   (fn _ => [rtac (temp_use LatticeDisjunctionIntro) 1,
-	     etac (temp_use LatticeTriangle2) 1,
-	     rtac (S4S5S6LeadstoS6 RS (temp_use LatticeTransitivity)) 1,
-	     auto_tac (MI_css addSIs2 [S4S5S6LeadstoS6,necT]
-			      addIs2 [ImplLeadsto_gen] addsimps2 Init_defs)
-	    ]);
-
-qed_goal "S2S3S4S5S6LeadstoS6" MemoryImplementation.thy
-   "!!sigma. [| sigma |= S2 rmhist p ~> S3 rmhist p; \
-\               sigma |= S3 rmhist p ~> S4 rmhist p | S6 rmhist p;   \
-\               sigma |= S4 rmhist p & ires!p = #NotAResult \
-\                         ~> S4 rmhist p & ires!p ~= #NotAResult | S5 rmhist p; \
-\               sigma |= S4 rmhist p & ires!p ~= #NotAResult ~> S5 rmhist p;  \
-\               sigma |= S5 rmhist p ~> S6 rmhist p |]  \
-\      ==> sigma |= S2 rmhist p | S3 rmhist p | S4 rmhist p | S5 rmhist p | S6 rmhist p \
-\                   ~> S6 rmhist p"
-   (fn _ => [rtac (temp_use LatticeDisjunctionIntro) 1,
-	     rtac (temp_use LatticeTransitivity) 1, atac 2,
-	     rtac (S3S4S5S6LeadstoS6 RS (temp_use LatticeTransitivity)) 1,
-	     auto_tac (MI_css addSIs2 [S3S4S5S6LeadstoS6,necT]
-			      addIs2 [ImplLeadsto_gen] addsimps2 Init_defs)
-	    ]);
-
-qed_goal "NotS1LeadstoS6" MemoryImplementation.thy
-   "!!sigma. [| sigma |= []ImpInv rmhist p; \
-\        sigma |= S2 rmhist p ~> S3 rmhist p; \
-\        sigma |= S3 rmhist p ~> S4 rmhist p | S6 rmhist p; \
-\        sigma |= S4 rmhist p & ires!p = #NotAResult \
-\                 ~> S4 rmhist p & ires!p ~= #NotAResult | S5 rmhist p; \
-\        sigma |= S4 rmhist p & ires!p ~= #NotAResult ~> S5 rmhist p;  \
-\        sigma |= S5 rmhist p ~> S6 rmhist p |] \
-\        ==> sigma |= ~S1 rmhist p ~> S6 rmhist p"
-   (fn _ => [rtac (S2S3S4S5S6LeadstoS6 RS (temp_use LatticeTransitivity)) 1,
-             TRYALL atac,
-             etac (temp_use INV_leadsto) 1,
-             rtac (temp_use ImplLeadsto_gen) 1,
-             rtac (temp_use necT) 1,
-	     auto_tac (MI_css addsimps2 ImpInv_def::Init_defs addSIs2 [necT])
-	    ]);
-
-qed_goal "S1Infinite" MemoryImplementation.thy
-   "!!sigma. [| sigma |= ~S1 rmhist p ~> S6 rmhist p; \
-\               sigma |= []<>S6 rmhist p --> []<>S1 rmhist p |] \
-\            ==> sigma |= []<>S1 rmhist p"
-   (fn _ => [rtac classical 1,
-	     asm_full_simp_tac (simpset() addsimps [temp_use NotBox, NotDmd]) 1,
-	     auto_tac (MI_css addSEs2 [mp,leadsto_infinite] addSDs2 [DBImplBD])
-	    ]);
--- a/src/HOL/TLA/Memory/MIsafe.ML	Thu Aug 03 19:28:37 2000 +0200
+++ b/src/HOL/TLA/Memory/MIsafe.ML	Thu Aug 03 19:29:03 2000 +0200
@@ -10,17 +10,17 @@
 
 (* RPCFailure notin MemVals U {OK,BadArg} *)
 
-qed_goalw "MVOKBAnotRF" MemoryImplementation.thy [MVOKBA_def]
-   "!!x. MVOKBA x ==> x ~= RPCFailure"
-   (fn _ => [ Auto_tac ]);
+Goalw [MVOKBA_def] "MVOKBA x ==> x ~= RPCFailure";
+by Auto_tac;
+qed "MVOKBAnotRF";
 
 (* NotAResult notin MemVals U {OK,BadArg,RPCFailure} *)
 
-qed_goalw "MVOKBARFnotNR" MemoryImplementation.thy [MVOKBARF_def]
-   "!!x. MVOKBARF x ==> x ~= NotAResult"
-   (fn _ => [ Auto_tac ]);
+Goalw [MVOKBARF_def] "MVOKBARF x ==> x ~= NotAResult";
+by Auto_tac;
+qed "MVOKBARFnotNR";
 
-(* ========================= Si's are mutually exclusive ==================================== *)
+(* ================ Si's are mutually exclusive ================================ *)
 (* Si and Sj are mutually exclusive for i # j. This helps to simplify the big
    conditional in the definition of resbar when doing the step-simulation proof.
    We prove a weaker result, which suffices for our purposes: 
@@ -28,42 +28,41 @@
 *)
 
 (* --- not used ---
-qed_goal "S1_excl" MemoryImplementation.thy 
-     "|- S1 rmhist p --> S1 rmhist p & ~S2 rmhist p & ~S3 rmhist p & \
-\                        ~S4 rmhist p & ~S5 rmhist p & ~S6 rmhist p"
-   (fn _ => [ auto_tac (MI_css addsimps2 [S_def, S1_def, S2_def,
-                                          S3_def, S4_def, S5_def, S6_def])
-            ]);
+Goal "|- S1 rmhist p --> S1 rmhist p & ~S2 rmhist p & ~S3 rmhist p & \
+\                        ~S4 rmhist p & ~S5 rmhist p & ~S6 rmhist p";
+by (auto_tac (MI_css addsimps2 [S_def, S1_def, S2_def,
+                                S3_def, S4_def, S5_def, S6_def]));
+qed "S1_excl";
 *)
 
-qed_goal "S2_excl" MemoryImplementation.thy 
-     "|- S2 rmhist p --> S2 rmhist p & ~S1 rmhist p"
-   (fn _ => [ auto_tac (MI_css addsimps2 [S_def, S1_def, S2_def]) ]);
+Goal "|- S2 rmhist p --> S2 rmhist p & ~S1 rmhist p";
+by (auto_tac (MI_css addsimps2 [S_def, S1_def, S2_def]));
+qed "S2_excl";
 
-qed_goal "S3_excl" MemoryImplementation.thy 
-     "|- S3 rmhist p --> S3 rmhist p & ~S1 rmhist p & ~S2 rmhist p"
-   (fn _ => [ auto_tac (MI_css addsimps2 [S_def, S1_def, S2_def, S3_def]) ]);
+Goal "|- S3 rmhist p --> S3 rmhist p & ~S1 rmhist p & ~S2 rmhist p";
+by (auto_tac (MI_css addsimps2 [S_def, S1_def, S2_def, S3_def]));
+qed "S3_excl";
 
-qed_goal "S4_excl" MemoryImplementation.thy 
-     "|- S4 rmhist p --> S4 rmhist p & ~S1 rmhist p & ~S2 rmhist p & ~S3 rmhist p"
-   (fn _ => [ auto_tac (MI_css addsimps2 [S_def,S1_def,S2_def,S3_def,S4_def]) ]);
+Goal "|- S4 rmhist p --> S4 rmhist p & ~S1 rmhist p & ~S2 rmhist p & ~S3 rmhist p";
+by (auto_tac (MI_css addsimps2 [S_def,S1_def,S2_def,S3_def,S4_def]));
+qed "S4_excl";
 
-qed_goal "S5_excl" MemoryImplementation.thy 
-     "|- S5 rmhist p --> S5 rmhist p & ~S1 rmhist p & ~S2 rmhist p \
-\                        & ~S3 rmhist p & ~S4 rmhist p"
-   (fn _ => [ auto_tac (MI_css addsimps2 [S_def,S1_def,S2_def,S3_def,S4_def,S5_def]) ]);
+Goal "|- S5 rmhist p --> S5 rmhist p & ~S1 rmhist p & ~S2 rmhist p \
+\                        & ~S3 rmhist p & ~S4 rmhist p";
+by (auto_tac (MI_css addsimps2 [S_def,S1_def,S2_def,S3_def,S4_def,S5_def]));
+qed "S5_excl";
 
-qed_goal "S6_excl" MemoryImplementation.thy 
-     "|- S6 rmhist p --> S6 rmhist p & ~S1 rmhist p & ~S2 rmhist p  \
-\                        & ~S3 rmhist p & ~S4 rmhist p & ~S5 rmhist p"
-   (fn _ => [ auto_tac (MI_css addsimps2 [S_def,S1_def,S2_def,S3_def,S4_def,S5_def,S6_def]) ]);
+Goal "|- S6 rmhist p --> S6 rmhist p & ~S1 rmhist p & ~S2 rmhist p  \
+\                        & ~S3 rmhist p & ~S4 rmhist p & ~S5 rmhist p";
+by (auto_tac (MI_css addsimps2 [S_def,S1_def,S2_def,S3_def,S4_def,S5_def,S6_def]));
+qed "S6_excl";
 
 
 (* ==================== Lemmas about the environment ============================== *)
 
-qed_goal "Envbusy" MemoryImplementation.thy
-   "|- $(Calling memCh p) --> ~ENext p"
-   (fn _ => [ auto_tac (MI_css addsimps2 [ENext_def,Call_def]) ]);
+Goal "|- $(Calling memCh p) --> ~ENext p";
+by (auto_tac (MI_css addsimps2 [ENext_def,Call_def]));
+qed "Envbusy";
 
 (* ==================== Lemmas about the implementation's states ==================== *)
 
@@ -74,287 +73,268 @@
 
 (* ------------------------------ State S1 ---------------------------------------- *) 
 
-qed_goal "S1Env" MemoryImplementation.thy
-   "|- ENext p & $(S1 rmhist p) & unchanged (c p, r p, m p, rmhist!p) --> (S2 rmhist p)$"
-   (fn _ => [force_tac (MI_css
-		        addsimps2 [ENext_def,Call_def,c_def,r_def,m_def,
-			   	   caller_def,rtrner_def,MVNROKBA_def,
-                                   S_def,S1_def,S2_def,Calling_def]) 1
-	    ]);
+Goal "|- ENext p & $(S1 rmhist p) & unchanged (c p, r p, m p, rmhist!p) \
+\        --> (S2 rmhist p)$";
+by (force_tac (MI_css addsimps2 [ENext_def,Call_def,c_def,r_def,m_def,
+                                 caller_def,rtrner_def,MVNROKBA_def,
+                                 S_def,S1_def,S2_def,Calling_def]) 1);
+qed "S1Env";
 
-qed_goal "S1ClerkUnch" MemoryImplementation.thy 
-   "|- [MClkNext memCh crCh cst p]_(c p) & $(S1 rmhist p) --> unchanged (c p)"
-   (fn _ => [auto_tac (MI_fast_css addSDs2 [MClkidle] addsimps2 [S_def,S1_def]) ]);
+Goal "|- [MClkNext memCh crCh cst p]_(c p) & $(S1 rmhist p) --> unchanged (c p)";
+by (auto_tac (MI_fast_css addSDs2 [MClkidle] addsimps2 [S_def,S1_def]));
+qed "S1ClerkUnch";
 
-qed_goal "S1RPCUnch" MemoryImplementation.thy
-   "|- [RPCNext crCh rmCh rst p]_(r p) & $(S1 rmhist p) --> unchanged (r p)"
-   (fn _ => [auto_tac (MI_fast_css addSDs2 [RPCidle] addsimps2 [S_def,S1_def]) ]);
+Goal "|- [RPCNext crCh rmCh rst p]_(r p) & $(S1 rmhist p) --> unchanged (r p)";
+by (auto_tac (MI_fast_css addSDs2 [RPCidle] addsimps2 [S_def,S1_def]));
+qed "S1RPCUnch";
 
-qed_goal "S1MemUnch" MemoryImplementation.thy
-   "|- [RNext rmCh mm ires p]_(m p) & $(S1 rmhist p) --> unchanged (m p)"
-   (fn _ => [auto_tac (MI_fast_css addSDs2 [Memoryidle] addsimps2 [S_def,S1_def]) ]);
+Goal "|- [RNext rmCh mm ires p]_(m p) & $(S1 rmhist p) --> unchanged (m p)";
+by (auto_tac (MI_fast_css addSDs2 [Memoryidle] addsimps2 [S_def,S1_def]));
+qed "S1MemUnch";
 
-qed_goal "S1Hist" MemoryImplementation.thy
-   "|- [HNext rmhist p]_(c p,r p,m p,rmhist!p) & $(S1 rmhist p) --> unchanged (rmhist!p)"
-   (fn _ => [action_simp_tac (simpset() addsimps [HNext_def, S_def, S1_def, MemReturn_def, 
-                                                  RPCFail_def,MClkReply_def,Return_def])
-                             [] [squareE] 1
-	    ]);
+Goal "|- [HNext rmhist p]_(c p,r p,m p,rmhist!p) & $(S1 rmhist p)\
+\        --> unchanged (rmhist!p)";
+by (action_simp_tac (simpset() addsimps [HNext_def, S_def, S1_def, MemReturn_def, 
+                                         RPCFail_def,MClkReply_def,Return_def])
+                    [] [squareE] 1);
+qed "S1Hist";
 
 (* ------------------------------ State S2 ---------------------------------------- *)
 
-qed_goal "S2EnvUnch" MemoryImplementation.thy
-   "|- [ENext p]_(e p) & $(S2 rmhist p) --> unchanged (e p)"
-   (fn _ => [auto_tac (MI_css addSDs2 [Envbusy] addsimps2 [S_def,S2_def]) ]);
+Goal "|- [ENext p]_(e p) & $(S2 rmhist p) --> unchanged (e p)";
+by (auto_tac (MI_css addSDs2 [Envbusy] addsimps2 [S_def,S2_def]));
+qed "S2EnvUnch";
 
-qed_goal "S2Clerk" MemoryImplementation.thy
-   "|- MClkNext memCh crCh cst p & $(S2 rmhist p) --> MClkFwd memCh crCh cst p"
-   (fn _ => [auto_tac (MI_css addsimps2 [MClkNext_def,MClkRetry_def,MClkReply_def,
-					 S_def,S2_def])
-	    ]);
+Goal "|- MClkNext memCh crCh cst p & $(S2 rmhist p) --> MClkFwd memCh crCh cst p";
+by (auto_tac (MI_css addsimps2 [MClkNext_def,MClkRetry_def,MClkReply_def,
+                                S_def,S2_def]));
+qed "S2Clerk";
 
-qed_goal "S2Forward" MemoryImplementation.thy
-   "|- $(S2 rmhist p) & MClkFwd memCh crCh cst p & unchanged (e p, r p, m p, rmhist!p) \
-\      --> (S3 rmhist p)$"
-   (fn _ => [action_simp_tac (simpset() addsimps
+Goal "|- $(S2 rmhist p) & MClkFwd memCh crCh cst p\
+\        & unchanged (e p, r p, m p, rmhist!p) \
+\        --> (S3 rmhist p)$";
+by (action_simp_tac 
+         (simpset() addsimps
                 [MClkFwd_def,Call_def,e_def,r_def,m_def,caller_def,rtrner_def,
                  S_def,S2_def,S3_def,Calling_def])
-               [] [] 1
-	     ]);
+         [] [] 1);
+qed "S2Forward";
 
-qed_goal "S2RPCUnch" MemoryImplementation.thy
-   "|- [RPCNext crCh rmCh rst p]_(r p) & $(S2 rmhist p) --> unchanged (r p)"
-   (fn _ => [auto_tac (MI_css addsimps2 [S_def,S2_def] addSDs2 [RPCidle]) ]);
+Goal "|- [RPCNext crCh rmCh rst p]_(r p) & $(S2 rmhist p) --> unchanged (r p)";
+by (auto_tac (MI_css addsimps2 [S_def,S2_def] addSDs2 [RPCidle]));
+qed "S2RPCUnch";
 
-qed_goal "S2MemUnch" MemoryImplementation.thy
-   "|- [RNext rmCh mm ires p]_(m p) & $(S2 rmhist p) --> unchanged (m p)"
-   (fn _ => [auto_tac (MI_css addsimps2 [S_def,S2_def] addSDs2 [Memoryidle]) ]);
+Goal "|- [RNext rmCh mm ires p]_(m p) & $(S2 rmhist p) --> unchanged (m p)";
+by (auto_tac (MI_css addsimps2 [S_def,S2_def] addSDs2 [Memoryidle]));
+qed "S2MemUnch";
 
-qed_goal "S2Hist" MemoryImplementation.thy
-   "|- [HNext rmhist p]_(c p,r p,m p,rmhist!p) & $(S2 rmhist p) --> unchanged (rmhist!p)"
-   (fn _ => [auto_tac (MI_fast_css
-		       addsimps2 [HNext_def,MemReturn_def,
-				  RPCFail_def,MClkReply_def,Return_def,S_def,S2_def])
-	    ]);
+Goal "|- [HNext rmhist p]_(c p,r p,m p,rmhist!p) & $(S2 rmhist p)\
+\        --> unchanged (rmhist!p)";
+by (auto_tac (MI_fast_css
+		addsimps2 [HNext_def,MemReturn_def,
+		           RPCFail_def,MClkReply_def,Return_def,S_def,S2_def]));
+qed "S2Hist";
 
 (* ------------------------------ State S3 ---------------------------------------- *)
 
-qed_goal "S3EnvUnch" MemoryImplementation.thy
-   "|- [ENext p]_(e p) & $(S3 rmhist p) --> unchanged (e p)"
-   (fn _ => [auto_tac (MI_css addSDs2 [Envbusy] addsimps2 [S_def,S3_def]) ]);
+Goal "|- [ENext p]_(e p) & $(S3 rmhist p) --> unchanged (e p)";
+by (auto_tac (MI_css addSDs2 [Envbusy] addsimps2 [S_def,S3_def]));
+qed "S3EnvUnch";
 
-qed_goal "S3ClerkUnch" MemoryImplementation.thy 
-   "|- [MClkNext memCh crCh cst p]_(c p) & $(S3 rmhist p) --> unchanged (c p)"
-   (fn _ => [auto_tac (MI_css addSDs2 [MClkbusy] addsimps2 [square_def,S_def,S3_def]) ]);
+Goal "|- [MClkNext memCh crCh cst p]_(c p) & $(S3 rmhist p) --> unchanged (c p)";
+by (auto_tac (MI_css addSDs2 [MClkbusy] addsimps2 [square_def,S_def,S3_def]));
+qed "S3ClerkUnch";
 
-qed_goal "S3LegalRcvArg" MemoryImplementation.thy
-   "|- S3 rmhist p --> IsLegalRcvArg<arg<crCh!p>>"
-   (fn _ => [auto_tac (MI_css addsimps2 [IsLegalRcvArg_def,MClkRelayArg_def,S_def,S3_def]) ]);
+Goal "|- S3 rmhist p --> IsLegalRcvArg<arg<crCh!p>>";
+by (auto_tac (MI_css addsimps2 [IsLegalRcvArg_def,MClkRelayArg_def,S_def,S3_def]));
+qed "S3LegalRcvArg";
 
-qed_goal "S3RPC" MemoryImplementation.thy
-   "|- RPCNext crCh rmCh rst p & $(S3 rmhist p) \
-\      --> RPCFwd crCh rmCh rst p | RPCFail crCh rmCh rst p"
-   (fn _ => [Clarsimp_tac 1,
-             forward_tac [action_use S3LegalRcvArg] 1,
-	     auto_tac (MI_css addsimps2 [RPCNext_def,RPCReject_def,RPCReply_def,S_def,S3_def])
-	    ]);
+Goal "|- RPCNext crCh rmCh rst p & $(S3 rmhist p) \
+\        --> RPCFwd crCh rmCh rst p | RPCFail crCh rmCh rst p";
+by (Clarsimp_tac 1);
+by (forward_tac [action_use S3LegalRcvArg] 1);
+by (auto_tac (MI_css addsimps2 [RPCNext_def,RPCReject_def,RPCReply_def,S_def,S3_def]));
+qed "S3RPC";
 
-qed_goal "S3Forward" MemoryImplementation.thy
-   "|- RPCFwd crCh rmCh rst p & HNext rmhist p & $(S3 rmhist p) & unchanged (e p, c p, m p) \
-\      --> (S4 rmhist p)$ & unchanged (rmhist!p)"
-   (fn _ => [action_simp_tac 
-               (simpset() addsimps [RPCFwd_def,HNext_def,MemReturn_def,RPCFail_def,MClkReply_def,
-				   Return_def,Call_def,e_def,c_def,m_def,caller_def,rtrner_def, 
-				   S_def,S3_def,S4_def,Calling_def])
-	       [] [] 1
-	    ]);
+Goal "|- RPCFwd crCh rmCh rst p & HNext rmhist p & $(S3 rmhist p)\
+\        & unchanged (e p, c p, m p) \
+\        --> (S4 rmhist p)$ & unchanged (rmhist!p)";
+by (action_simp_tac 
+      (simpset() addsimps [RPCFwd_def,HNext_def,MemReturn_def,RPCFail_def,MClkReply_def,
+	                   Return_def,Call_def,e_def,c_def,m_def,caller_def,rtrner_def, 
+                           S_def,S3_def,S4_def,Calling_def])
+      [] [] 1);
+qed "S3Forward";
 
-qed_goal "S3Fail" MemoryImplementation.thy
-   "|- RPCFail crCh rmCh rst p & $(S3 rmhist p) & HNext rmhist p & unchanged (e p, c p, m p) \
-\      --> (S6 rmhist p)$"
-   (fn _ => [action_simp_tac 
-               (simpset() addsimps [HNext_def,RPCFail_def,Return_def,e_def,c_def,m_def,
-				   caller_def,rtrner_def,MVOKBARF_def,
-				   S_def,S3_def,S6_def,Calling_def])
-               [] [] 1
-	    ]);
+Goal "|- RPCFail crCh rmCh rst p & $(S3 rmhist p) & HNext rmhist p\
+\        & unchanged (e p, c p, m p) \
+\        --> (S6 rmhist p)$";
+by (action_simp_tac 
+      (simpset() addsimps [HNext_def,RPCFail_def,Return_def,e_def,c_def,m_def,
+			   caller_def,rtrner_def,MVOKBARF_def,
+			   S_def,S3_def,S6_def,Calling_def])
+      [] [] 1);
+qed "S3Fail";
 
-qed_goal "S3MemUnch" MemoryImplementation.thy
-   "|- [RNext rmCh mm ires p]_(m p) & $(S3 rmhist p) --> unchanged (m p)"
-   (fn _ => [auto_tac (MI_css addsimps2 [S_def,S3_def] addSDs2 [Memoryidle]) ]);
+Goal "|- [RNext rmCh mm ires p]_(m p) & $(S3 rmhist p) --> unchanged (m p)";
+by (auto_tac (MI_css addsimps2 [S_def,S3_def] addSDs2 [Memoryidle]));
+qed "S3MemUnch";
 
-qed_goal "S3Hist" MemoryImplementation.thy
-   "|- HNext rmhist p & $(S3 rmhist p) & unchanged (r p) --> unchanged (rmhist!p)"
-   (fn _ => [auto_tac (MI_css
-		       addsimps2 [HNext_def,MemReturn_def,RPCFail_def,MClkReply_def,
-				  Return_def,r_def,rtrner_def,S_def,S3_def,Calling_def])
-	    ]);
-
+Goal "|- HNext rmhist p & $(S3 rmhist p) & unchanged (r p) --> unchanged (rmhist!p)";
+by (auto_tac (MI_css addsimps2 [HNext_def,MemReturn_def,RPCFail_def,MClkReply_def,
+			        Return_def,r_def,rtrner_def,S_def,S3_def,Calling_def]));
+qed "S3Hist";
 
 (* ------------------------------ State S4 ---------------------------------------- *)
 
-qed_goal "S4EnvUnch" MemoryImplementation.thy
-   "|- [ENext p]_(e p) & $(S4 rmhist p) --> unchanged (e p)"
-   (fn _ => [auto_tac (MI_css addsimps2 [S_def,S4_def] addSDs2 [Envbusy]) ]);
+Goal "|- [ENext p]_(e p) & $(S4 rmhist p) --> unchanged (e p)";
+by (auto_tac (MI_css addsimps2 [S_def,S4_def] addSDs2 [Envbusy]));
+qed "S4EnvUnch";
 
-qed_goal "S4ClerkUnch" MemoryImplementation.thy
-   "|- [MClkNext memCh crCh cst p]_(c p) & $(S4 rmhist p) --> unchanged (c p)"
-   (fn _ => [auto_tac (MI_css addsimps2 [S_def,S4_def] addSDs2 [MClkbusy]) ]);
+Goal "|- [MClkNext memCh crCh cst p]_(c p) & $(S4 rmhist p) --> unchanged (c p)";
+by (auto_tac (MI_css addsimps2 [S_def,S4_def] addSDs2 [MClkbusy]));
+qed "S4ClerkUnch";
 
-qed_goal "S4RPCUnch" MemoryImplementation.thy
-   "|- [RPCNext crCh rmCh rst p]_(r p) & $(S4 rmhist p) --> unchanged (r p)"
-   (fn _ => [auto_tac (MI_fast_css addsimps2 [S_def,S4_def] addSDs2 [RPCbusy]) ]);
+Goal "|- [RPCNext crCh rmCh rst p]_(r p) & $(S4 rmhist p) --> unchanged (r p)";
+by (auto_tac (MI_fast_css addsimps2 [S_def,S4_def] addSDs2 [RPCbusy]));
+qed "S4RPCUnch";
 
-qed_goal "S4ReadInner" MemoryImplementation.thy
-   "|- ReadInner rmCh mm ires p l & $(S4 rmhist p) & unchanged (e p, c p, r p) \
-\           & HNext rmhist p & $(MemInv mm l) \
-\      --> (S4 rmhist p)$ & unchanged (rmhist!p)"
-   (fn _ => [action_simp_tac 
-               (simpset() addsimps [ReadInner_def,GoodRead_def, BadRead_def,HNext_def,
-				   MemReturn_def, RPCFail_def,MClkReply_def,Return_def,
-				   e_def,c_def,r_def,rtrner_def,caller_def,MVNROKBA_def,
-				   S_def,S4_def,RdRequest_def,Calling_def,MemInv_def])
-               [] [] 1
-	    ]);
+Goal "|- ReadInner rmCh mm ires p l & $(S4 rmhist p) & unchanged (e p, c p, r p) \
+\        & HNext rmhist p & $(MemInv mm l) \
+\        --> (S4 rmhist p)$ & unchanged (rmhist!p)";
+by (action_simp_tac 
+      (simpset() addsimps [ReadInner_def,GoodRead_def, BadRead_def,HNext_def,
+			   MemReturn_def, RPCFail_def,MClkReply_def,Return_def,
+			   e_def,c_def,r_def,rtrner_def,caller_def,MVNROKBA_def,
+			   S_def,S4_def,RdRequest_def,Calling_def,MemInv_def])
+      [] [] 1);
+qed "S4ReadInner";
 
-qed_goal "S4Read" MemoryImplementation.thy
-   "|- Read rmCh mm ires p & $(S4 rmhist p) & unchanged (e p, c p, r p) \
-\           & HNext rmhist p & (!l. $MemInv mm l) \
-\      --> (S4 rmhist p)$ & unchanged (rmhist!p)"
-   (fn _ => [auto_tac (MI_css addsimps2 [Read_def] addSDs2 [S4ReadInner]) ]);
+Goal "|- Read rmCh mm ires p & $(S4 rmhist p) & unchanged (e p, c p, r p) \
+\        & HNext rmhist p & (!l. $MemInv mm l) \
+\        --> (S4 rmhist p)$ & unchanged (rmhist!p)";
+by (auto_tac (MI_css addsimps2 [Read_def] addSDs2 [S4ReadInner]));
+qed "S4Read";
 
-qed_goal "S4WriteInner" MemoryImplementation.thy
-   "|- WriteInner rmCh mm ires p l v & $(S4 rmhist p) & unchanged (e p, c p, r p) \
-\           & HNext rmhist p \
-\      --> (S4 rmhist p)$ & unchanged (rmhist!p)"
-   (fn _ => [action_simp_tac 
-               (simpset() addsimps [WriteInner_def,GoodWrite_def, BadWrite_def,HNext_def,
-				   MemReturn_def,RPCFail_def,MClkReply_def,Return_def,
-				   e_def,c_def,r_def,rtrner_def,caller_def,MVNROKBA_def, 
-				   S_def,S4_def,WrRequest_def,Calling_def])
-               [] [] 1
-	    ]);
+Goal "|- WriteInner rmCh mm ires p l v & $(S4 rmhist p) & unchanged (e p, c p, r p) \
+\        & HNext rmhist p \
+\        --> (S4 rmhist p)$ & unchanged (rmhist!p)";
+by (action_simp_tac 
+      (simpset() addsimps [WriteInner_def,GoodWrite_def, BadWrite_def,HNext_def,
+			   MemReturn_def,RPCFail_def,MClkReply_def,Return_def,
+			   e_def,c_def,r_def,rtrner_def,caller_def,MVNROKBA_def, 
+			   S_def,S4_def,WrRequest_def,Calling_def])
+      [] [] 1);
+qed "S4WriteInner";
 
-qed_goal "S4Write" MemoryImplementation.thy
-   "|- Write rmCh mm ires p l & $(S4 rmhist p) & unchanged (e p, c p, r p) & (HNext rmhist p) \
-\      --> (S4 rmhist p)$ & unchanged (rmhist!p)"
-   (fn _ => [ auto_tac (MI_css addsimps2 [Write_def] addSDs2 [S4WriteInner]) ]);
+Goal "|- Write rmCh mm ires p l & $(S4 rmhist p) & unchanged (e p, c p, r p)\
+\        & (HNext rmhist p) \
+\        --> (S4 rmhist p)$ & unchanged (rmhist!p)";
+by (auto_tac (MI_css addsimps2 [Write_def] addSDs2 [S4WriteInner]));
+qed "S4Write";
 
-qed_goal "WriteS4" MemoryImplementation.thy
-   "|- $ImpInv rmhist p & Write rmCh mm ires p l --> $S4 rmhist p"
-   (fn _ => [auto_tac (MI_css
-		       addsimps2 [Write_def,WriteInner_def,ImpInv_def,WrRequest_def,
-				  S_def,S1_def,S2_def,S3_def,S4_def,S5_def,S6_def])
-            ]);
+Goal "|- $ImpInv rmhist p & Write rmCh mm ires p l --> $S4 rmhist p";
+by (auto_tac (MI_css addsimps2 [Write_def,WriteInner_def,ImpInv_def,WrRequest_def,
+			        S_def,S1_def,S2_def,S3_def,S4_def,S5_def,S6_def]));
+qed "WriteS4";
 
-qed_goal "S4Return" MemoryImplementation.thy
-   "|- MemReturn rmCh ires p & $S4 rmhist p & unchanged (e p, c p, r p) & HNext rmhist p \
-\      --> (S5 rmhist p)$"
-   (fn _ => [auto_tac (MI_css
-		       addsimps2 [HNext_def,MemReturn_def,Return_def,e_def,c_def,r_def,
-				  rtrner_def,caller_def,MVNROKBA_def,MVOKBA_def,
-		                  S_def,S4_def,S5_def,Calling_def])
-	    ]);
+Goal "|- MemReturn rmCh ires p & $S4 rmhist p & unchanged (e p, c p, r p)\
+\        & HNext rmhist p \
+\        --> (S5 rmhist p)$";
+by (auto_tac (MI_css addsimps2 [HNext_def,MemReturn_def,Return_def,e_def,c_def,r_def,
+				rtrner_def,caller_def,MVNROKBA_def,MVOKBA_def,
+		                S_def,S4_def,S5_def,Calling_def]));
+qed "S4Return";
 
-qed_goal "S4Hist" MemoryImplementation.thy
-   "|- HNext rmhist p & $S4 rmhist p & (m p)$ = $(m p) --> (rmhist!p)$ = $(rmhist!p)"
-   (fn _ => [auto_tac (MI_css
-		       addsimps2 [HNext_def,MemReturn_def,RPCFail_def,MClkReply_def,
-				  Return_def,m_def,rtrner_def,S_def,S4_def,Calling_def])
-	    ]);
+Goal "|- HNext rmhist p & $S4 rmhist p & (m p)$ = $(m p) --> (rmhist!p)$ = $(rmhist!p)";
+by (auto_tac (MI_css addsimps2 [HNext_def,MemReturn_def,RPCFail_def,MClkReply_def,
+				Return_def,m_def,rtrner_def,S_def,S4_def,Calling_def]));
+qed "S4Hist";
 
 (* ------------------------------ State S5 ---------------------------------------- *)
 
-qed_goal "S5EnvUnch" MemoryImplementation.thy
-   "|- [ENext p]_(e p) & $(S5 rmhist p) --> unchanged (e p)"
-   (fn _ => [auto_tac (MI_css addsimps2 [S_def,S5_def] addSDs2 [Envbusy]) ]);
+Goal "|- [ENext p]_(e p) & $(S5 rmhist p) --> unchanged (e p)";
+by (auto_tac (MI_css addsimps2 [S_def,S5_def] addSDs2 [Envbusy]));
+qed "S5EnvUnch";
 
-qed_goal "S5ClerkUnch" MemoryImplementation.thy
-   "|- [MClkNext memCh crCh cst p]_(c p) & $(S5 rmhist p) --> unchanged (c p)"
-   (fn _ => [auto_tac (MI_css addsimps2 [S_def,S5_def] addSDs2 [MClkbusy]) ]);
+Goal "|- [MClkNext memCh crCh cst p]_(c p) & $(S5 rmhist p) --> unchanged (c p)";
+by (auto_tac (MI_css addsimps2 [S_def,S5_def] addSDs2 [MClkbusy]));
+qed "S5ClerkUnch";
 
-qed_goal "S5RPC" MemoryImplementation.thy
-   "|- RPCNext crCh rmCh rst p & $(S5 rmhist p)   \
-\      --> RPCReply crCh rmCh rst p | RPCFail crCh rmCh rst p"
-   (fn _ => [auto_tac (MI_css
-		       addsimps2 [RPCNext_def,RPCReject_def,RPCFwd_def,S_def,S5_def])
-	    ]);
+Goal "|- RPCNext crCh rmCh rst p & $(S5 rmhist p)   \
+\        --> RPCReply crCh rmCh rst p | RPCFail crCh rmCh rst p";
+by (auto_tac (MI_css addsimps2 [RPCNext_def,RPCReject_def,RPCFwd_def,S_def,S5_def]));
+qed "S5RPC";
 
-qed_goal "S5Reply" MemoryImplementation.thy
-   "|- RPCReply crCh rmCh rst p & $(S5 rmhist p) & unchanged (e p, c p, m p,rmhist!p) \
-\      --> (S6 rmhist p)$"
-   (fn _ => [action_simp_tac 
-               (simpset()
-		addsimps [RPCReply_def,Return_def,e_def,c_def,m_def,
-			  MVOKBA_def,MVOKBARF_def,caller_def,rtrner_def,
-			  S_def,S5_def,S6_def,Calling_def])
-               [] [] 1
-	    ]);
+Goal "|- RPCReply crCh rmCh rst p & $(S5 rmhist p) & unchanged (e p, c p, m p,rmhist!p)\
+\      --> (S6 rmhist p)$";
+by (action_simp_tac 
+      (simpset() addsimps [RPCReply_def,Return_def,e_def,c_def,m_def,
+			   MVOKBA_def,MVOKBARF_def,caller_def,rtrner_def,
+			   S_def,S5_def,S6_def,Calling_def])
+      [] [] 1);
+qed "S5Reply";
 
-qed_goal "S5Fail" MemoryImplementation.thy
-   "|- RPCFail crCh rmCh rst p & $(S5 rmhist p) & unchanged (e p, c p, m p,rmhist!p) \
-\      --> (S6 rmhist p)$"
-   (fn _ => [action_simp_tac
-	       (simpset()
-		addsimps [RPCFail_def,Return_def,e_def,c_def,m_def,
-			  MVOKBARF_def,caller_def,rtrner_def,
-			  S_def,S5_def,S6_def,Calling_def])
-               [] [] 1
-	    ]);
+Goal "|- RPCFail crCh rmCh rst p & $(S5 rmhist p) & unchanged (e p, c p, m p,rmhist!p) \
+\        --> (S6 rmhist p)$";
+by (action_simp_tac
+      (simpset() addsimps [RPCFail_def,Return_def,e_def,c_def,m_def,
+			   MVOKBARF_def,caller_def,rtrner_def,
+			   S_def,S5_def,S6_def,Calling_def])
+      [] [] 1);
+qed "S5Fail";
 
-qed_goal "S5MemUnch" MemoryImplementation.thy
-   "|- [RNext rmCh mm ires p]_(m p) & $(S5 rmhist p) --> unchanged (m p)"
-   (fn _ => [auto_tac (MI_css addsimps2 [S_def,S5_def] addSDs2 [Memoryidle]) ]);
+Goal "|- [RNext rmCh mm ires p]_(m p) & $(S5 rmhist p) --> unchanged (m p)";
+by (auto_tac (MI_css addsimps2 [S_def,S5_def] addSDs2 [Memoryidle]));
+qed "S5MemUnch";
 
-qed_goal "S5Hist" MemoryImplementation.thy
-   "|- [HNext rmhist p]_(c p, r p, m p, rmhist!p) & $(S5 rmhist p) --> (rmhist!p)$ = $(rmhist!p)"
-   (fn _ => [auto_tac (MI_fast_css
-		       addsimps2 [HNext_def,MemReturn_def,
-				  RPCFail_def,MClkReply_def,Return_def,S_def,S5_def])
-	    ]);
+Goal "|- [HNext rmhist p]_(c p, r p, m p, rmhist!p) & $(S5 rmhist p)\
+\        --> (rmhist!p)$ = $(rmhist!p)";
+by (auto_tac (MI_fast_css
+	      addsimps2 [HNext_def,MemReturn_def,
+		         RPCFail_def,MClkReply_def,Return_def,S_def,S5_def]));
+qed "S5Hist";
 
 (* ------------------------------ State S6 ---------------------------------------- *)
 
-qed_goal "S6EnvUnch" MemoryImplementation.thy
-   "|- [ENext p]_(e p) & $(S6 rmhist p) --> unchanged (e p)"
-   (fn _ => [auto_tac (MI_css addsimps2 [S_def,S6_def] addSDs2 [Envbusy]) ]);
+Goal "|- [ENext p]_(e p) & $(S6 rmhist p) --> unchanged (e p)";
+by (auto_tac (MI_css addsimps2 [S_def,S6_def] addSDs2 [Envbusy]));
+qed "S6EnvUnch";
 
-qed_goal "S6Clerk" MemoryImplementation.thy
-   "|- MClkNext memCh crCh cst p & $(S6 rmhist p) \
-\      --> MClkRetry memCh crCh cst p | MClkReply memCh crCh cst p"
-   (fn _ => [ auto_tac (MI_css addsimps2 [MClkNext_def,MClkFwd_def,S_def,S6_def]) ]);
+Goal "|- MClkNext memCh crCh cst p & $(S6 rmhist p) \
+\        --> MClkRetry memCh crCh cst p | MClkReply memCh crCh cst p";
+by (auto_tac (MI_css addsimps2 [MClkNext_def,MClkFwd_def,S_def,S6_def]));
+qed "S6Clerk";
 
-qed_goal "S6Retry" MemoryImplementation.thy
-   "|- MClkRetry memCh crCh cst p & HNext rmhist p & $S6 rmhist p & unchanged (e p,r p,m p) \
-\      --> (S3 rmhist p)$ & unchanged (rmhist!p)"
-   (fn _ => [action_simp_tac
-	        (simpset() addsimps [HNext_def,MClkReply_def,MClkRetry_def,Call_def,
-				    Return_def,e_def,r_def,m_def,caller_def,rtrner_def,
-		                    S_def,S6_def,S3_def,Calling_def])
-                [] [] 1]);
+Goal "|- MClkRetry memCh crCh cst p & HNext rmhist p & $S6 rmhist p\
+\        & unchanged (e p,r p,m p) \
+\        --> (S3 rmhist p)$ & unchanged (rmhist!p)";
+by (action_simp_tac
+      (simpset() addsimps [HNext_def,MClkReply_def,MClkRetry_def,Call_def,
+	                   Return_def,e_def,r_def,m_def,caller_def,rtrner_def,
+                           S_def,S6_def,S3_def,Calling_def])
+      [] [] 1);
+qed "S6Retry";
 
-qed_goal "S6Reply" MemoryImplementation.thy
-   "|- MClkReply memCh crCh cst p & HNext rmhist p & $S6 rmhist p & unchanged (e p,r p,m p) \
-\      --> (S1 rmhist p)$"
-   (fn _ => [action_simp_tac (simpset()
-			      addsimps [HNext_def,MemReturn_def,RPCFail_def,Return_def,
-					MClkReply_def,e_def,r_def,m_def,caller_def,rtrner_def,
-					S_def,S6_def,S1_def,Calling_def])
-	                     [] [] 1
-	    ]);
+Goal "|- MClkReply memCh crCh cst p & HNext rmhist p & $S6 rmhist p\
+\        & unchanged (e p,r p,m p) \
+\        --> (S1 rmhist p)$";
+by (action_simp_tac 
+      (simpset() addsimps [HNext_def,MemReturn_def,RPCFail_def,Return_def,
+		           MClkReply_def,e_def,r_def,m_def,caller_def,rtrner_def,
+                           S_def,S6_def,S1_def,Calling_def])
+      [] [] 1);
+qed "S6Reply";
 
-qed_goal "S6RPCUnch" MemoryImplementation.thy
-   "|- [RPCNext crCh rmCh rst p]_(r p) & $S6 rmhist p --> unchanged (r p)"
-   (fn _ => [auto_tac (MI_css addsimps2 [S_def,S6_def] addSDs2 [RPCidle]) ]);
-
-qed_goal "S6MemUnch" MemoryImplementation.thy
-   "|- [RNext rmCh mm ires p]_(m p) & $(S6 rmhist p) --> unchanged (m p)"
-   (fn _ => [auto_tac (MI_css addsimps2 [S_def,S6_def] addSDs2 [Memoryidle]) ]);
+Goal "|- [RPCNext crCh rmCh rst p]_(r p) & $S6 rmhist p --> unchanged (r p)";
+by (auto_tac (MI_css addsimps2 [S_def,S6_def] addSDs2 [RPCidle]));
+qed "S6RPCUnch";
 
-qed_goal "S6Hist" MemoryImplementation.thy
-   "|- HNext rmhist p & $S6 rmhist p & (c p)$ = $(c p) --> (rmhist!p)$ = $(rmhist!p)"
-   (fn _ => [auto_tac (MI_css
-		       addsimps2 [HNext_def,MClkReply_def,Return_def,c_def,rtrner_def,
-		                  S_def,S6_def,Calling_def])
-	    ]);
+Goal "|- [RNext rmCh mm ires p]_(m p) & $(S6 rmhist p) --> unchanged (m p)";
+by (auto_tac (MI_css addsimps2 [S_def,S6_def] addSDs2 [Memoryidle]));
+qed "S6MemUnch";
 
+Goal "|- HNext rmhist p & $S6 rmhist p & (c p)$ = $(c p) --> (rmhist!p)$ = $(rmhist!p)";
+by (auto_tac (MI_css addsimps2 [HNext_def,MClkReply_def,Return_def,c_def,rtrner_def,
+		                S_def,S6_def,Calling_def]));
+qed "S6Hist";
--- a/src/HOL/TLA/Memory/MemClerk.ML	Thu Aug 03 19:28:37 2000 +0200
+++ b/src/HOL/TLA/Memory/MemClerk.ML	Thu Aug 03 19:29:03 2000 +0200
@@ -16,47 +16,42 @@
 (* The Clerk engages in an action for process p only if there is an outstanding,
    unanswered call for that process.
 *)
+Goal "|- ~$Calling send p & $(cst!p) = #clkA --> ~MClkNext send rcv cst p";
+by (auto_tac (mem_css addsimps2 (Return_def::MC_action_defs)));
+qed "MClkidle";
 
-qed_goal "MClkidle" MemClerk.thy
-   "|- ~$Calling send p & $(cst!p) = #clkA --> ~MClkNext send rcv cst p"
-   (fn _ => [ auto_tac (mem_css addsimps2 (Return_def::MC_action_defs)) ]);
-
-qed_goal "MClkbusy" MemClerk.thy
-   "|- $Calling rcv p --> ~MClkNext send rcv cst p"
-   (fn _ => [ auto_tac (mem_css addsimps2 (MC_action_defs @ [Call_def])) ]);
+Goal "|- $Calling rcv p --> ~MClkNext send rcv cst p";
+by (auto_tac (mem_css addsimps2 (Call_def::MC_action_defs)));
+qed "MClkbusy";
 
 (* Enabledness of actions *)
 
-qed_goal "MClkFwd_enabled" MemClerk.thy
-   "!!p. basevars (rtrner send!p, caller rcv!p, cst!p) ==> \
-\        |- Calling send p & ~Calling rcv p & cst!p = #clkA  \
-\           --> Enabled (MClkFwd send rcv cst p)"
-   (fn _ => [action_simp_tac (simpset() addsimps [MClkFwd_def,Call_def,caller_def,rtrner_def])
-                             [exI] [base_enabled,Pair_inject] 1]);
+Goal "!!p. basevars (rtrner send!p, caller rcv!p, cst!p) ==> \
+\     |- Calling send p & ~Calling rcv p & cst!p = #clkA  \
+\        --> Enabled (MClkFwd send rcv cst p)";
+by (action_simp_tac (simpset() addsimps [MClkFwd_def,Call_def,caller_def,rtrner_def])
+                    [exI] [base_enabled,Pair_inject] 1);
+qed "MClkFwd_enabled";
 
-qed_goal "MClkFwd_ch_enabled" MemClerk.thy
-   "|- Enabled (MClkFwd send rcv cst p)  -->  \
-\      Enabled (<MClkFwd send rcv cst p>_(cst!p, rtrner send!p, caller rcv!p))"
-   (fn _ => [auto_tac (mem_css addSEs2 [enabled_mono]
-	                       addsimps2 [angle_def,MClkFwd_def])
-  	    ]);
+Goal "|- Enabled (MClkFwd send rcv cst p)  -->  \
+\        Enabled (<MClkFwd send rcv cst p>_(cst!p, rtrner send!p, caller rcv!p))";
+by (auto_tac (mem_css addSEs2 [enabled_mono] addsimps2 [angle_def,MClkFwd_def]));
+qed "MClkFwd_ch_enabled";
 
-qed_goal "MClkReply_change" MemClerk.thy
-   "|- MClkReply send rcv cst p --> <MClkReply send rcv cst p>_(cst!p, rtrner send!p, caller rcv!p)"
-   (fn _ => [auto_tac (mem_css addsimps2 [angle_def,MClkReply_def]
-			       addEs2 [Return_changed])
-            ]);
+Goal "|- MClkReply send rcv cst p --> \
+\        <MClkReply send rcv cst p>_(cst!p, rtrner send!p, caller rcv!p)";
+by (auto_tac (mem_css addsimps2 [angle_def,MClkReply_def]
+	              addEs2 [Return_changed]));
+qed "MClkReply_change";
 
-qed_goal "MClkReply_enabled" MemClerk.thy
-   "!!p. basevars (rtrner send!p, caller rcv!p, cst!p) ==> \
-\        |- Calling send p & ~Calling rcv p & cst!p = #clkB  \
-\           --> Enabled (<MClkReply send rcv cst p>_(cst!p, rtrner send!p, caller rcv!p))"
-   (fn _ => [action_simp_tac (simpset()) [MClkReply_change RSN (2,enabled_mono)] [] 1,
-	     action_simp_tac (simpset() addsimps [MClkReply_def,Return_def,caller_def,rtrner_def])
-                             [exI] [base_enabled,Pair_inject] 1
-	    ]);
+Goal "!!p. basevars (rtrner send!p, caller rcv!p, cst!p) ==> \
+\     |- Calling send p & ~Calling rcv p & cst!p = #clkB  \
+\        --> Enabled (<MClkReply send rcv cst p>_(cst!p, rtrner send!p, caller rcv!p))";
+by (action_simp_tac (simpset()) [MClkReply_change RSN (2,enabled_mono)] [] 1);
+by (action_simp_tac (simpset() addsimps [MClkReply_def,Return_def,caller_def,rtrner_def])
+                    [exI] [base_enabled,Pair_inject] 1);
+qed "MClkReply_enabled";
 
-qed_goal "MClkReplyNotRetry" MemClerk.thy
-   "|- MClkReply send rcv cst p --> ~MClkRetry send rcv cst p"
-   (fn _ => [ auto_tac (mem_css addsimps2 [MClkReply_def,MClkRetry_def]) ]);
-
+Goal "|- MClkReply send rcv cst p --> ~MClkRetry send rcv cst p";
+by (auto_tac (mem_css addsimps2 [MClkReply_def,MClkRetry_def]));
+qed "MClkReplyNotRetry";
--- a/src/HOL/TLA/Memory/MemClerk.thy	Thu Aug 03 19:28:37 2000 +0200
+++ b/src/HOL/TLA/Memory/MemClerk.thy	Thu Aug 03 19:29:03 2000 +0200
@@ -62,7 +62,7 @@
                          & SF(MClkReply send rcv cst p)_(cst!p, rtrner send!p, caller rcv!p)"
 
   MClkISpec     :: "mClkSndChType => mClkRcvChType => mClkStType => temporal"
-      "MClkISpec send rcv cst == TEMP (!p. MClkIPSpec send rcv cst p)"
+      "MClkISpec send rcv cst == TEMP (ALL p. MClkIPSpec send rcv cst p)"
 
 end
 
--- a/src/HOL/TLA/Memory/Memory.ML	Thu Aug 03 19:28:37 2000 +0200
+++ b/src/HOL/TLA/Memory/Memory.ML	Thu Aug 03 19:29:03 2000 +0200
@@ -22,112 +22,104 @@
 (* -------------------- Proofs ---------------------------------------------- *)
 
 (* The reliable memory is an implementation of the unreliable one *)
-qed_goal "ReliableImplementsUnReliable" Memory.thy 
-   "|- IRSpec ch mm rs --> IUSpec ch mm rs"
-   (K [force_tac (temp_css addsimps2 ([UNext_def,UPSpec_def,IUSpec_def] @ RM_temp_defs)
-			   addSEs2 [STL4E,squareE]) 1]);
+Goal "|- IRSpec ch mm rs --> IUSpec ch mm rs";
+by (force_tac (temp_css addsimps2 ([UNext_def,UPSpec_def,IUSpec_def] @ RM_temp_defs)
+			addSEs2 [STL4E,squareE]) 1);
+qed "ReliableImplementsUnReliable";
 
 (* The memory spec implies the memory invariant *)
-qed_goal "MemoryInvariant" Memory.thy 
-   "|- MSpec ch mm rs l --> [](MemInv mm l)"
-   (fn _ => [ auto_inv_tac (simpset() addsimps RM_temp_defs @ RM_action_defs) 1 ]);
+Goal "|- MSpec ch mm rs l --> [](MemInv mm l)";
+by (auto_inv_tac (simpset() addsimps RM_temp_defs @ RM_action_defs) 1);
+qed "MemoryInvariant";
 
 (* The invariant is trivial for non-locations *)
-qed_goal "NonMemLocInvariant" Memory.thy
-   "|- #l ~: #MemLoc --> [](MemInv mm l)"
-   (K [ auto_tac (temp_css addsimps2 [MemInv_def] addSIs2 [necT]) ]);
+Goal "|- #l ~: #MemLoc --> [](MemInv mm l)";
+by (auto_tac (temp_css addsimps2 [MemInv_def] addSIs2 [necT]));
+qed "NonMemLocInvariant";
 
-qed_goal "MemoryInvariantAll" Memory.thy
-   "|- (!l. #l : #MemLoc --> MSpec ch mm rs l) --> (!l. [](MemInv mm l))"
-    (K [step_tac temp_cs 1,
-	case_tac "l : MemLoc" 1,
-	auto_tac (temp_css addSEs2 [MemoryInvariant,NonMemLocInvariant]) ]);
+Goal "|- (ALL l. #l : #MemLoc --> MSpec ch mm rs l) --> (ALL l. [](MemInv mm l))";
+by (step_tac temp_cs 1);
+by (case_tac "l : MemLoc" 1);
+by (auto_tac (temp_css addSEs2 [MemoryInvariant,NonMemLocInvariant]));
+qed "MemoryInvariantAll";
 
 (* The memory engages in an action for process p only if there is an 
    unanswered call from p.
    We need this only for the reliable memory.
 *)
 
-qed_goal "Memoryidle" Memory.thy
-   "|- ~$(Calling ch p) --> ~ RNext ch mm rs p"
-   (K [ auto_tac (mem_css addsimps2 (Return_def::RM_action_defs)) ]);
+Goal "|- ~$(Calling ch p) --> ~ RNext ch mm rs p";
+by (auto_tac (mem_css addsimps2 (Return_def::RM_action_defs)));
+qed "Memoryidle";
 
 (* Enabledness conditions *)
 
-qed_goal "MemReturn_change" Memory.thy
-   "|- MemReturn ch rs p --> <MemReturn ch rs p>_(rtrner ch ! p, rs!p)"
-   (K [ force_tac (mem_css addsimps2 [MemReturn_def,angle_def]) 1]);
+Goal "|- MemReturn ch rs p --> <MemReturn ch rs p>_(rtrner ch ! p, rs!p)";
+by (force_tac (mem_css addsimps2 [MemReturn_def,angle_def]) 1);
+qed "MemReturn_change";
 
-qed_goal "MemReturn_enabled" Memory.thy
-   "!!p. basevars (rtrner ch ! p, rs!p) ==> \
-\        |- Calling ch p & (rs!p ~= #NotAResult) \
-\           --> Enabled (<MemReturn ch rs p>_(rtrner ch ! p, rs!p))"
-  (K [action_simp_tac (simpset()) [MemReturn_change RSN (2,enabled_mono)] [] 1,
-      action_simp_tac (simpset() addsimps [MemReturn_def,Return_def,rtrner_def])
-                      [exI] [base_enabled,Pair_inject] 1
-     ]);
+Goal "!!p. basevars (rtrner ch ! p, rs!p) ==> \
+\     |- Calling ch p & (rs!p ~= #NotAResult) \
+\        --> Enabled (<MemReturn ch rs p>_(rtrner ch ! p, rs!p))";
+by (action_simp_tac (simpset()) [MemReturn_change RSN (2,enabled_mono)] [] 1);
+by (action_simp_tac (simpset() addsimps [MemReturn_def,Return_def,rtrner_def])
+                    [exI] [base_enabled,Pair_inject] 1);
+qed "MemReturn_enabled";
 
-qed_goal "ReadInner_enabled" Memory.thy
- "!!p. basevars (rtrner ch ! p, rs!p) ==> \
-\      |- Calling ch p & (arg<ch!p> = #(read l)) --> Enabled (ReadInner ch mm rs p l)"
-   (fn _ => [case_tac "l : MemLoc" 1,
-             ALLGOALS
-	        (force_tac (mem_css addsimps2 [ReadInner_def,GoodRead_def,
-                                               BadRead_def,RdRequest_def]
-                            addSIs2 [exI] addSEs2 [base_enabled]))
-            ]);
+Goal "!!p. basevars (rtrner ch ! p, rs!p) ==> \
+\     |- Calling ch p & (arg<ch!p> = #(read l)) --> Enabled (ReadInner ch mm rs p l)";
+by (case_tac "l : MemLoc" 1);
+by (ALLGOALS
+     (force_tac (mem_css addsimps2 [ReadInner_def,GoodRead_def,
+                                    BadRead_def,RdRequest_def]
+                         addSIs2 [exI] addSEs2 [base_enabled])));
+qed "ReadInner_enabled";
 
-qed_goal "WriteInner_enabled" Memory.thy
-   "!!p. basevars (mm!l, rtrner ch ! p, rs!p) ==> \
-\        |- Calling ch p & (arg<ch!p> = #(write l v)) \
-\           --> Enabled (WriteInner ch mm rs p l v)"
-   (fn _ => [case_tac "l:MemLoc & v:MemVal" 1,
-             ALLGOALS 
-	        (force_tac (mem_css addsimps2 [WriteInner_def,GoodWrite_def,
-                                               BadWrite_def,WrRequest_def]
-                            addSIs2 [exI] addSEs2 [base_enabled]))
-            ]);
+Goal "!!p. basevars (mm!l, rtrner ch ! p, rs!p) ==> \
+\     |- Calling ch p & (arg<ch!p> = #(write l v)) \
+\        --> Enabled (WriteInner ch mm rs p l v)";
+by (case_tac "l:MemLoc & v:MemVal" 1);
+by (ALLGOALS 
+     (force_tac (mem_css addsimps2 [WriteInner_def,GoodWrite_def,
+                                    BadWrite_def,WrRequest_def]
+                         addSIs2 [exI] addSEs2 [base_enabled])));
+qed "WriteInner_enabled";
 
-qed_goal "ReadResult" Memory.thy
-   "|- Read ch mm rs p & (!l. $(MemInv mm l)) --> (rs!p)` ~= #NotAResult"
-   (fn _ => [force_tac (mem_css addsimps2 
-                            [Read_def,ReadInner_def,GoodRead_def,BadRead_def,MemInv_def]) 1]);
+Goal "|- Read ch mm rs p & (!l. $(MemInv mm l)) --> (rs!p)` ~= #NotAResult";
+by (force_tac (mem_css addsimps2 
+                       [Read_def,ReadInner_def,GoodRead_def,BadRead_def,MemInv_def]) 1);
+qed "ReadResult";
 
-qed_goal "WriteResult" Memory.thy
-   "|- Write ch mm rs p l --> (rs!p)` ~= #NotAResult"
-   (fn _ => [auto_tac (mem_css addsimps2 ([Write_def,WriteInner_def,GoodWrite_def,BadWrite_def]))
-	    ]);
+Goal "|- Write ch mm rs p l --> (rs!p)` ~= #NotAResult";
+by (auto_tac (mem_css addsimps2 ([Write_def,WriteInner_def,GoodWrite_def,BadWrite_def])));
+qed "WriteResult";
 
-qed_goal "ReturnNotReadWrite" Memory.thy
-   "|- (!l. $MemInv mm l) & MemReturn ch rs p \
-\      --> ~ Read ch mm rs p & (!l. ~ Write ch mm rs p l)"
-   (fn _ => [auto_tac
-	       (mem_css addsimps2 [MemReturn_def] addSDs2 [WriteResult, ReadResult])
-	    ]);
+Goal "|- (ALL l. $MemInv mm l) & MemReturn ch rs p \
+\        --> ~ Read ch mm rs p & (ALL l. ~ Write ch mm rs p l)";
+by (auto_tac (mem_css addsimps2 [MemReturn_def] addSDs2 [WriteResult, ReadResult]));
+qed "ReturnNotReadWrite";
 
-qed_goal "RWRNext_enabled" Memory.thy
-   "|- (rs!p = #NotAResult) & (!l. MemInv mm l)  \
-\          & Enabled (Read ch mm rs p | (? l. Write ch mm rs p l)) \
-\      --> Enabled (<RNext ch mm rs p>_(rtrner ch ! p, rs!p))"
-   (K [force_tac (mem_css addsimps2 [RNext_def,angle_def]
-	     addSEs2 [enabled_mono2]
-	     addDs2 [ReadResult, WriteResult]) 1]);
+Goal "|- (rs!p = #NotAResult) & (!l. MemInv mm l)  \
+\        & Enabled (Read ch mm rs p | (? l. Write ch mm rs p l)) \
+\        --> Enabled (<RNext ch mm rs p>_(rtrner ch ! p, rs!p))";
+by (force_tac (mem_css addsimps2 [RNext_def,angle_def]
+	               addSEs2 [enabled_mono2]
+	               addDs2 [ReadResult, WriteResult]) 1);
+qed "RWRNext_enabled";
 
 
 (* Combine previous lemmas: the memory can make a visible step if there is an
    outstanding call for which no result has been produced.
 *)
-qed_goal "RNext_enabled" Memory.thy
-"!!p. !l. basevars (mm!l, rtrner ch!p, rs!p) ==> \
+Goal "!!p. !l. basevars (mm!l, rtrner ch!p, rs!p) ==> \
 \     |- (rs!p = #NotAResult) & Calling ch p & (!l. MemInv mm l)  \
-\        --> Enabled (<RNext ch mm rs p>_(rtrner ch ! p, rs!p))" (K [
-	     auto_tac (mem_css addsimps2 [enabled_disj]
-		                  addSIs2 [RWRNext_enabled]),
-             case_tac "arg(ch w p)" 1,
- 	      action_simp_tac (simpset()addsimps[Read_def,enabled_ex])
-	                      [ReadInner_enabled,exI] [] 1,
-              force_tac (mem_css addDs2 [base_pair]) 1,
-	     etac swap 1,
-	     action_simp_tac (simpset() addsimps [Write_def,enabled_ex])
-	                     [WriteInner_enabled,exI] [] 1]);
-
+\        --> Enabled (<RNext ch mm rs p>_(rtrner ch ! p, rs!p))";
+by (auto_tac (mem_css addsimps2 [enabled_disj] addSIs2 [RWRNext_enabled]));
+by (case_tac "arg(ch w p)" 1);
+ by (action_simp_tac (simpset()addsimps[Read_def,enabled_ex])
+                     [ReadInner_enabled,exI] [] 1);
+ by (force_tac (mem_css addDs2 [base_pair]) 1);
+by (etac swap 1);
+by (action_simp_tac (simpset() addsimps [Write_def,enabled_ex])
+	            [WriteInner_enabled,exI] [] 1);
+qed "RNext_enabled";
--- a/src/HOL/TLA/Memory/Memory.thy	Thu Aug 03 19:28:37 2000 +0200
+++ b/src/HOL/TLA/Memory/Memory.thy	Thu Aug 03 19:29:03 2000 +0200
@@ -74,7 +74,7 @@
                          & (GoodRead mm rs p l  |  BadRead mm rs p l)
                          & unchanged (rtrner ch ! p)"
   (* the read action with l quantified *)
-  Read_def          "Read ch mm rs p == ACT (? l. ReadInner ch mm rs p l)"
+  Read_def          "Read ch mm rs p == ACT (EX l. ReadInner ch mm rs p l)"
 
   (* similar definitions for the write action *)
   GoodWrite_def     "GoodWrite mm rs p l v == ACT
@@ -87,7 +87,7 @@
                         $(WrRequest ch p l v)
                         & (GoodWrite mm rs p l v  |  BadWrite mm rs p l v)
                         & unchanged (rtrner ch ! p)"
-  Write_def         "Write ch mm rs p l == ACT (? v. WriteInner ch mm rs p l v)"
+  Write_def         "Write ch mm rs p l == ACT (EX v. WriteInner ch mm rs p l v)"
 
   (* the return action *)
   MemReturn_def     "MemReturn ch rs p == ACT
@@ -103,7 +103,7 @@
   (* next-state relations for reliable / unreliable memory *)
   RNext_def         "RNext ch mm rs p == ACT 
                        (  Read ch mm rs p
-                        | (? l. Write ch mm rs p l)
+                        | (EX l. Write ch mm rs p l)
                         | MemReturn ch rs p)"
   UNext_def         "UNext ch mm rs p == ACT
                         (RNext ch mm rs p | MemFail ch rs p)"
@@ -120,13 +120,13 @@
                         & WF(MemReturn ch rs p)_(rtrner ch ! p, rs!p)"
   MSpec_def         "MSpec ch mm rs l == TEMP
                         Init(MInit mm l)
-                        & [][ ? p. Write ch mm rs p l ]_(mm!l)"
+                        & [][ EX p. Write ch mm rs p l ]_(mm!l)"
   IRSpec_def        "IRSpec ch mm rs == TEMP
-                        (! p. RPSpec ch mm rs p)
-                        & (! l. #l : #MemLoc --> MSpec ch mm rs l)"
+                        (ALL p. RPSpec ch mm rs p)
+                        & (ALL l. #l : #MemLoc --> MSpec ch mm rs l)"
   IUSpec_def        "IUSpec ch mm rs == TEMP
-                        (! p. UPSpec ch mm rs p)
-                        & (! l. #l : #MemLoc --> MSpec ch mm rs l)"
+                        (ALL p. UPSpec ch mm rs p)
+                        & (ALL l. #l : #MemLoc --> MSpec ch mm rs l)"
 
   RSpec_def         "RSpec ch rs == TEMP (EEX mm. IRSpec ch mm rs)"
   USpec_def         "USpec ch == TEMP (EEX mm rs. IUSpec ch mm rs)"
--- a/src/HOL/TLA/Memory/MemoryImplementation.ML	Thu Aug 03 19:28:37 2000 +0200
+++ b/src/HOL/TLA/Memory/MemoryImplementation.ML	Thu Aug 03 19:29:03 2000 +0200
@@ -36,30 +36,30 @@
 (****************************** The history variable ******************************)
 section "History variable";
 
-qed_goal "HistoryLemma" MemoryImplementation.thy
-   "|- Init(!p. ImpInit p) & [](!p. ImpNext p)  \
-\      --> (EEX rmhist. Init(! p. HInit rmhist p) \
-\                     & [](!p. [HNext rmhist p]_(c p, r p, m p, rmhist!p)))"
-   (fn _ => [Clarsimp_tac 1,
-             rtac historyI 1, TRYALL atac, rtac MI_base 1,
-             action_simp_tac (simpset() addsimps [HInit_def]) [] [] 1,
-             etac fun_cong 1,
-             action_simp_tac (simpset() addsimps [HNext_def]) [busy_squareI] [] 1,
-             etac fun_cong 1
-            ]);
+Goal "|- Init(ALL p. ImpInit p) & [](ALL p. ImpNext p)  \
+\        --> (EEX rmhist. Init(ALL p. HInit rmhist p) \
+\                         & [](ALL p. [HNext rmhist p]_(c p, r p, m p, rmhist!p)))";
+by (Clarsimp_tac 1);
+by (rtac historyI 1); 
+by (TRYALL atac); 
+by (rtac MI_base 1);
+by (action_simp_tac (simpset() addsimps [HInit_def]) [] [] 1);
+by (etac fun_cong 1);
+by (action_simp_tac (simpset() addsimps [HNext_def]) [busy_squareI] [] 1);
+by (etac fun_cong 1);
+qed "HistoryLemma";
 
-qed_goal "History" MemoryImplementation.thy
-   "|- Implementation --> (EEX rmhist. Hist rmhist)"
-   (fn _ => [Clarsimp_tac 1,
-             rtac ((temp_use HistoryLemma) RS eex_mono) 1,
-             force_tac (MI_css 
-                        addsimps2 [Hist_def,HistP_def,Init_def,all_box,split_box_conj]) 3,
-             auto_tac (MI_css
-                       addsimps2 [Implementation_def,MClkISpec_def,RPCISpec_def,IRSpec_def,
-                                  MClkIPSpec_def,RPCIPSpec_def,RPSpec_def,
-                                  ImpInit_def,Init_def,ImpNext_def,
-                                  c_def,r_def,m_def,all_box,split_box_conj])
-            ]);
+Goal "|- Implementation --> (EEX rmhist. Hist rmhist)";
+by (Clarsimp_tac 1);
+by (rtac ((temp_use HistoryLemma) RS eex_mono) 1);
+by (force_tac (MI_css 
+               addsimps2 [Hist_def,HistP_def,Init_def,all_box,split_box_conj]) 3);
+by (auto_tac (MI_css
+              addsimps2 [Implementation_def,MClkISpec_def,RPCISpec_def,IRSpec_def,
+                         MClkIPSpec_def,RPCIPSpec_def,RPSpec_def,
+                         ImpInit_def,Init_def,ImpNext_def,
+                         c_def,r_def,m_def,all_box,split_box_conj]));
+qed "History";
 
 (******************************** The safety part *********************************)
 
@@ -74,82 +74,76 @@
 (* ========== Step 1.1 ================================================= *)
 (* The implementation's initial condition implies the state predicate S1 *)
 
-qed_goal "Step1_1" MemoryImplementation.thy
-   "|- ImpInit p & HInit rmhist p --> S1 rmhist p"
-   (fn _ => [auto_tac (MI_fast_css
-		       addsimps2 [MVNROKBA_def,MClkInit_def,RPCInit_def,PInit_def,
-			          HInit_def,ImpInit_def,S_def,S1_def])
-	    ]);
+Goal "|- ImpInit p & HInit rmhist p --> S1 rmhist p";
+by (auto_tac (MI_fast_css
+              addsimps2 [MVNROKBA_def,MClkInit_def,RPCInit_def,PInit_def,
+		         HInit_def,ImpInit_def,S_def,S1_def]));
+qed "Step1_1";
 
 (* ========== Step 1.2 ================================================== *)
 (* Figure 16 is a predicate-action diagram for the implementation. *)
 
-qed_goal "Step1_2_1" MemoryImplementation.thy
-   "|- [HNext rmhist p]_(c p,r p,m p, rmhist!p) & ImpNext p  \
-\             & ~unchanged (e p, c p, r p, m p, rmhist!p)  & $S1 rmhist p \
-\      --> (S2 rmhist p)$ & ENext p & unchanged (c p, r p, m p)"
-   (fn _ => [action_simp_tac (simpset() addsimps [ImpNext_def]) []
-                             (map temp_elim [S1ClerkUnch,S1RPCUnch,S1MemUnch,S1Hist]) 1,
-             auto_tac (MI_fast_css addSIs2 [S1Env])
-	    ]);
+Goal "|- [HNext rmhist p]_(c p,r p,m p, rmhist!p) & ImpNext p  \
+\        & ~unchanged (e p, c p, r p, m p, rmhist!p)  & $S1 rmhist p \
+\        --> (S2 rmhist p)$ & ENext p & unchanged (c p, r p, m p)";
+by (action_simp_tac (simpset() addsimps [ImpNext_def]) []
+                    (map temp_elim [S1ClerkUnch,S1RPCUnch,S1MemUnch,S1Hist]) 1);
+by (auto_tac (MI_fast_css addSIs2 [S1Env]));
+qed "Step1_2_1";
 
-qed_goal "Step1_2_2" MemoryImplementation.thy
-   "|- [HNext rmhist p]_(c p,r p,m p, rmhist!p) & ImpNext p  \
-\             & ~unchanged (e p, c p, r p, m p, rmhist!p) & $S2 rmhist p \
-\      --> (S3 rmhist p)$ & MClkFwd memCh crCh cst p & unchanged (e p, r p, m p, rmhist!p)"
-   (fn _ => [action_simp_tac (simpset() addsimps [ImpNext_def]) []
-                             (map temp_elim [S2EnvUnch,S2RPCUnch,S2MemUnch,S2Hist]) 1,
-	     auto_tac (MI_fast_css addSIs2 [S2Clerk,S2Forward])
-	    ]);
+Goal "|- [HNext rmhist p]_(c p,r p,m p, rmhist!p) & ImpNext p  \
+\        & ~unchanged (e p, c p, r p, m p, rmhist!p) & $S2 rmhist p \
+\        --> (S3 rmhist p)$ & MClkFwd memCh crCh cst p\
+\            & unchanged (e p, r p, m p, rmhist!p)";
+by (action_simp_tac (simpset() addsimps [ImpNext_def]) []
+                    (map temp_elim [S2EnvUnch,S2RPCUnch,S2MemUnch,S2Hist]) 1);
+by (auto_tac (MI_fast_css addSIs2 [S2Clerk,S2Forward]));
+qed "Step1_2_2";
 
-qed_goal "Step1_2_3" MemoryImplementation.thy
-   "|- [HNext rmhist p]_(c p,r p,m p, rmhist!p) & ImpNext p  \
-\             & ~unchanged (e p, c p, r p, m p, rmhist!p) & $S3 rmhist p \
-\      --> ((S4 rmhist p)$ & RPCFwd crCh rmCh rst p & unchanged (e p, c p, m p, rmhist!p)) \
-\        | ((S6 rmhist p)$ & RPCFail crCh rmCh rst p & unchanged (e p, c p, m p))"
-   (fn _ => [action_simp_tac (simpset() addsimps [ImpNext_def]) []
-	          (map temp_elim [S3EnvUnch,S3ClerkUnch,S3MemUnch]) 1,
-             action_simp_tac (simpset()) [] 
-                  (squareE::map temp_elim [S3RPC,S3Forward,S3Fail]) 1,
-             auto_tac (MI_css addDs2 [S3Hist])
-	    ]);
+Goal "|- [HNext rmhist p]_(c p,r p,m p, rmhist!p) & ImpNext p  \
+\        & ~unchanged (e p, c p, r p, m p, rmhist!p) & $S3 rmhist p \
+\        --> ((S4 rmhist p)$ & RPCFwd crCh rmCh rst p & unchanged (e p, c p, m p, rmhist!p)) \
+\            | ((S6 rmhist p)$ & RPCFail crCh rmCh rst p & unchanged (e p, c p, m p))";
+by (action_simp_tac (simpset() addsimps [ImpNext_def]) []
+	            (map temp_elim [S3EnvUnch,S3ClerkUnch,S3MemUnch]) 1);
+by (action_simp_tac (simpset()) [] 
+                    (squareE::map temp_elim [S3RPC,S3Forward,S3Fail]) 1);
+by (auto_tac (MI_css addDs2 [S3Hist]));
+qed "Step1_2_3";
 
-qed_goal "Step1_2_4" MemoryImplementation.thy
-   "|- [HNext rmhist p]_(c p,r p,m p, rmhist!p) & ImpNext p  \
+Goal "|- [HNext rmhist p]_(c p,r p,m p, rmhist!p) & ImpNext p  \
 \             & ~unchanged (e p, c p, r p, m p, rmhist!p) \
 \             & $S4 rmhist p & (!l. $(MemInv mm l))     \
-\      --> ((S4 rmhist p)$ & Read rmCh mm ires p & unchanged (e p, c p, r p, rmhist!p)) \
-\        | ((S4 rmhist p)$ & (? l. Write rmCh mm ires p l) & unchanged (e p, c p, r p, rmhist!p)) \
-\        | ((S5 rmhist p)$ & MemReturn rmCh ires p & unchanged (e p, c p, r p))"
-   (fn _ => [action_simp_tac (simpset() addsimps [ImpNext_def]) []
-                             (map temp_elim [S4EnvUnch,S4ClerkUnch,S4RPCUnch]) 1,
-             action_simp_tac (simpset() addsimps [RNext_def]) []
-                             (squareE::map temp_elim [S4Read,S4Write,S4Return]) 1,
-             auto_tac (MI_css addDs2 [S4Hist])
-            ]);
+\        --> ((S4 rmhist p)$ & Read rmCh mm ires p & unchanged (e p, c p, r p, rmhist!p)) \
+\            | ((S4 rmhist p)$ & (? l. Write rmCh mm ires p l) & unchanged (e p, c p, r p, rmhist!p)) \
+\            | ((S5 rmhist p)$ & MemReturn rmCh ires p & unchanged (e p, c p, r p))";
+by (action_simp_tac (simpset() addsimps [ImpNext_def]) []
+                    (map temp_elim [S4EnvUnch,S4ClerkUnch,S4RPCUnch]) 1);
+by (action_simp_tac (simpset() addsimps [RNext_def]) []
+                    (squareE::map temp_elim [S4Read,S4Write,S4Return]) 1);
+by (auto_tac (MI_css addDs2 [S4Hist]));
+qed "Step1_2_4";
 
-qed_goal "Step1_2_5" MemoryImplementation.thy
-   "|- [HNext rmhist p]_(c p,r p,m p, rmhist!p) & ImpNext p  \
+Goal "|- [HNext rmhist p]_(c p,r p,m p, rmhist!p) & ImpNext p  \
 \             & ~unchanged (e p, c p, r p, m p, rmhist!p) & $S5 rmhist p \
-\      --> ((S6 rmhist p)$ & RPCReply crCh rmCh rst p & unchanged (e p, c p, m p)) \
-\        | ((S6 rmhist p)$ & RPCFail crCh rmCh rst p & unchanged (e p, c p, m p))"
-   (fn _ => [action_simp_tac (simpset() addsimps [ImpNext_def]) []
-                             (map temp_elim [S5EnvUnch,S5ClerkUnch,S5MemUnch,S5Hist]) 1,
-	     action_simp_tac (simpset()) [] [squareE, temp_elim S5RPC] 1,
-	     auto_tac (MI_fast_css addSDs2 [S5Reply,S5Fail])
-	    ]);
+\        --> ((S6 rmhist p)$ & RPCReply crCh rmCh rst p & unchanged (e p, c p, m p)) \
+\            | ((S6 rmhist p)$ & RPCFail crCh rmCh rst p & unchanged (e p, c p, m p))";
+by (action_simp_tac (simpset() addsimps [ImpNext_def]) []
+                    (map temp_elim [S5EnvUnch,S5ClerkUnch,S5MemUnch,S5Hist]) 1);
+by (action_simp_tac (simpset()) [] [squareE, temp_elim S5RPC] 1);
+by (auto_tac (MI_fast_css addSDs2 [S5Reply,S5Fail]));
+qed "Step1_2_5";
 
-qed_goal "Step1_2_6" MemoryImplementation.thy
-   "|- [HNext rmhist p]_(c p,r p,m p, rmhist!p) & ImpNext p  \
+Goal "|- [HNext rmhist p]_(c p,r p,m p, rmhist!p) & ImpNext p  \
 \             & ~unchanged (e p, c p, r p, m p, rmhist!p) & $S6 rmhist p \
-\      --> ((S1 rmhist p)$ & MClkReply memCh crCh cst p & unchanged (e p, r p, m p))\
-\        | ((S3 rmhist p)$ & MClkRetry memCh crCh cst p & unchanged (e p,r p,m p,rmhist!p))"
-   (fn _ => [action_simp_tac (simpset() addsimps [ImpNext_def]) []
-                             (map temp_elim [S6EnvUnch,S6RPCUnch,S6MemUnch]) 1,
-             action_simp_tac (simpset()) []
-                             (squareE::map temp_elim [S6Clerk,S6Retry,S6Reply]) 1,
-             auto_tac (MI_css addDs2 [S6Hist])
-            ]);
+\        --> ((S1 rmhist p)$ & MClkReply memCh crCh cst p & unchanged (e p, r p, m p))\
+\            | ((S3 rmhist p)$ & MClkRetry memCh crCh cst p & unchanged (e p,r p,m p,rmhist!p))";
+by (action_simp_tac (simpset() addsimps [ImpNext_def]) []
+                    (map temp_elim [S6EnvUnch,S6RPCUnch,S6MemUnch]) 1);
+by (action_simp_tac (simpset()) []
+                    (squareE::map temp_elim [S6Clerk,S6Retry,S6Reply]) 1);
+by (auto_tac (MI_css addDs2 [S6Hist]));
+qed "Step1_2_6";
 
 (* --------------------------------------------------------------------------
    Step 1.3: S1 implies the barred initial condition.
@@ -157,11 +151,10 @@
 
 section "Initialization (Step 1.3)";
 
-qed_goal "Step1_3" MemoryImplementation.thy 
-   "|- S1 rmhist p --> PInit (resbar rmhist) p"
-   (fn _ => [action_simp_tac (simpset() addsimps [resbar_def,PInit_def,S_def,S1_def])
-                             [] [] 1
-            ]);
+Goal "|- S1 rmhist p --> PInit (resbar rmhist) p";
+by (action_simp_tac (simpset() addsimps [resbar_def,PInit_def,S_def,S1_def])
+                    [] [] 1);
+qed "Step1_3";
 
 (* ----------------------------------------------------------------------
    Step 1.4: Implementation's next-state relation simulates specification's
@@ -170,171 +163,161 @@
 
 section "Step simulation (Step 1.4)";
 
-qed_goal "Step1_4_1" MemoryImplementation.thy
-   "|- ENext p & $S1 rmhist p & (S2 rmhist p)$ & unchanged (c p, r p, m p) \
-\      --> unchanged (rtrner memCh!p, resbar rmhist!p)"
-  (fn _ => [ auto_tac (MI_fast_css addsimps2 [c_def,r_def,m_def,resbar_def]) ]);
+Goal "|- ENext p & $S1 rmhist p & (S2 rmhist p)$ & unchanged (c p, r p, m p) \
+\        --> unchanged (rtrner memCh!p, resbar rmhist!p)";
+by (auto_tac (MI_fast_css addsimps2 [c_def,r_def,m_def,resbar_def]));
+qed "Step1_4_1";
 
-qed_goal "Step1_4_2" MemoryImplementation.thy
-   "|- MClkFwd memCh crCh cst p & $S2 rmhist p & (S3 rmhist p)$  \
-\                & unchanged (e p, r p, m p, rmhist!p) \
-\      --> unchanged (rtrner memCh!p, resbar rmhist!p)"
-  (fn _ => [action_simp_tac
-                (simpset() addsimps [MClkFwd_def, e_def, r_def, m_def, resbar_def,
-                                     S_def, S2_def, S3_def]) [] [] 1
-           ]);
+Goal "|- MClkFwd memCh crCh cst p & $S2 rmhist p & (S3 rmhist p)$  \
+\        & unchanged (e p, r p, m p, rmhist!p) \
+\        --> unchanged (rtrner memCh!p, resbar rmhist!p)";
+by (action_simp_tac
+      (simpset() addsimps [MClkFwd_def, e_def, r_def, m_def, resbar_def,
+                           S_def, S2_def, S3_def]) [] [] 1);
+qed "Step1_4_2";
 
-qed_goal "Step1_4_3a" MemoryImplementation.thy
-   "|- RPCFwd crCh rmCh rst p & $S3 rmhist p & (S4 rmhist p)$    \
-\                  & unchanged (e p, c p, m p, rmhist!p) \
-\      --> unchanged (rtrner memCh!p, resbar rmhist!p)"
-  (fn _ => [Clarsimp_tac 1,
-            REPEAT (dresolve_tac (map temp_use [S3_excl,S4_excl]) 1),
-            action_simp_tac 
-                 (simpset() addsimps [e_def,c_def,m_def,resbar_def,S_def, S3_def]) [] [] 1
-           ]);
+Goal "|- RPCFwd crCh rmCh rst p & $S3 rmhist p & (S4 rmhist p)$    \
+\        & unchanged (e p, c p, m p, rmhist!p) \
+\        --> unchanged (rtrner memCh!p, resbar rmhist!p)";
+by (Clarsimp_tac 1);
+by (REPEAT (dresolve_tac (map temp_use [S3_excl,S4_excl]) 1));
+by (action_simp_tac 
+      (simpset() addsimps [e_def,c_def,m_def,resbar_def,S_def, S3_def]) [] [] 1);
+qed "Step1_4_3a";
 
-qed_goal "Step1_4_3b" MemoryImplementation.thy
-   "|- RPCFail crCh rmCh rst p & $S3 rmhist p & (S6 rmhist p)$ & unchanged (e p, c p, m p) \
-\      --> MemFail memCh (resbar rmhist) p"
-  (fn _ => [Clarsimp_tac 1,
-            dtac (temp_use S6_excl) 1,
-            auto_tac (MI_css addsimps2 [RPCFail_def,MemFail_def,e_def,c_def,m_def,
-		                        resbar_def]),
-            force_tac (MI_css addsimps2 [S3_def,S_def]) 1,
-            auto_tac (MI_css addsimps2 [Return_def])
-           ]);
-
+Goal "|- RPCFail crCh rmCh rst p & $S3 rmhist p & (S6 rmhist p)$\
+\        & unchanged (e p, c p, m p) \
+\        --> MemFail memCh (resbar rmhist) p";
+by (Clarsimp_tac 1);
+by (dtac (temp_use S6_excl) 1);
+by (auto_tac (MI_css addsimps2 [RPCFail_def,MemFail_def,e_def,c_def,m_def,
+	                        resbar_def]));
+by (force_tac (MI_css addsimps2 [S3_def,S_def]) 1);
+by (auto_tac (MI_css addsimps2 [Return_def]));
+qed "Step1_4_3b";
 
-qed_goal "Step1_4_4a1" MemoryImplementation.thy
-   "|- $S4 rmhist p & (S4 rmhist p)$ & ReadInner rmCh mm ires p l \
-\             & unchanged (e p, c p, r p, rmhist!p) & $MemInv mm l \
-\      --> ReadInner memCh mm (resbar rmhist) p l"
-  (fn _ => [Clarsimp_tac 1,
-            REPEAT (dtac (temp_use S4_excl) 1),
-            action_simp_tac 
-               (simpset() addsimps [ReadInner_def,GoodRead_def,BadRead_def,e_def,c_def,m_def]) 
-               [] [] 1,
-            auto_tac (MI_css addsimps2 [resbar_def]),
-	    ALLGOALS (action_simp_tac 
-                        (simpset() addsimps [RPCRelayArg_def,MClkRelayArg_def,
-		                            S_def,S4_def,RdRequest_def,MemInv_def])
-		      [] [impE,MemValNotAResultE])
-           ]);
+Goal "|- $S4 rmhist p & (S4 rmhist p)$ & ReadInner rmCh mm ires p l \
+\        & unchanged (e p, c p, r p, rmhist!p) & $MemInv mm l \
+\        --> ReadInner memCh mm (resbar rmhist) p l";
+by (Clarsimp_tac 1);
+by (REPEAT (dtac (temp_use S4_excl) 1));
+by (action_simp_tac 
+      (simpset() addsimps [ReadInner_def,GoodRead_def,BadRead_def,e_def,c_def,m_def]) 
+      [] [] 1);
+by (auto_tac (MI_css addsimps2 [resbar_def]));
+by (ALLGOALS (action_simp_tac 
+                (simpset() addsimps [RPCRelayArg_def,MClkRelayArg_def,
+	                             S_def,S4_def,RdRequest_def,MemInv_def])
+		[] [impE,MemValNotAResultE]));
+qed "Step1_4_4a1";
 
-qed_goal "Step1_4_4a" MemoryImplementation.thy
-   "|- Read rmCh mm ires p & $S4 rmhist p & (S4 rmhist p)$ \
-\           & unchanged (e p, c p, r p, rmhist!p) & (!l. $(MemInv mm l)) \
-\      --> Read memCh mm (resbar rmhist) p"
-  (fn _ => [ force_tac (MI_css addsimps2 [Read_def] addSEs2 [Step1_4_4a1]) 1 ]);
+Goal "|- Read rmCh mm ires p & $S4 rmhist p & (S4 rmhist p)$ \
+\        & unchanged (e p, c p, r p, rmhist!p) & (!l. $(MemInv mm l)) \
+\        --> Read memCh mm (resbar rmhist) p";
+by (force_tac (MI_css addsimps2 [Read_def] addSEs2 [Step1_4_4a1]) 1);
+qed "Step1_4_4a";
 
-qed_goal "Step1_4_4b1" MemoryImplementation.thy
-   "|- $S4 rmhist p & (S4 rmhist p)$ & WriteInner rmCh mm ires p l v   \
-\                   & unchanged (e p, c p, r p, rmhist!p) \
-\      --> WriteInner memCh mm (resbar rmhist) p l v"
-  (fn _ => [Clarsimp_tac 1,
-            REPEAT (dtac (temp_use S4_excl) 1),
-            action_simp_tac 
-               (simpset() addsimps [WriteInner_def, GoodWrite_def, BadWrite_def,
-			           e_def, c_def, m_def])
-               [] [] 1,
-	    auto_tac (MI_css addsimps2 [resbar_def]),
-	    ALLGOALS (action_simp_tac
-                        (simpset() addsimps [RPCRelayArg_def,MClkRelayArg_def,
-		                            S_def,S4_def,WrRequest_def])
-		      [] [])
-           ]);
+Goal "|- $S4 rmhist p & (S4 rmhist p)$ & WriteInner rmCh mm ires p l v   \
+\        & unchanged (e p, c p, r p, rmhist!p) \
+\        --> WriteInner memCh mm (resbar rmhist) p l v";
+by (Clarsimp_tac 1);
+by (REPEAT (dtac (temp_use S4_excl) 1));
+by (action_simp_tac 
+      (simpset() addsimps [WriteInner_def, GoodWrite_def, BadWrite_def,
+		           e_def, c_def, m_def])
+      [] [] 1);
+by (auto_tac (MI_css addsimps2 [resbar_def]));
+by (ALLGOALS (action_simp_tac
+                (simpset() addsimps [RPCRelayArg_def,MClkRelayArg_def,
+	                             S_def,S4_def,WrRequest_def])
+		[] []));
+qed "Step1_4_4b1";
 
-qed_goal "Step1_4_4b" MemoryImplementation.thy
-   "|- Write rmCh mm ires p l & $S4 rmhist p & (S4 rmhist p)$   \
-\                 & unchanged (e p, c p, r p, rmhist!p) \
-\      --> Write memCh mm (resbar rmhist) p l"
-  (fn _ => [ force_tac (MI_css addsimps2 [Write_def] addSEs2 [Step1_4_4b1]) 1 ]);
+Goal "|- Write rmCh mm ires p l & $S4 rmhist p & (S4 rmhist p)$   \
+\        & unchanged (e p, c p, r p, rmhist!p) \
+\        --> Write memCh mm (resbar rmhist) p l";
+by (force_tac (MI_css addsimps2 [Write_def] addSEs2 [Step1_4_4b1]) 1);
+qed "Step1_4_4b";
 
-qed_goal "Step1_4_4c" MemoryImplementation.thy
-   "|- MemReturn rmCh ires p & $S4 rmhist p & (S5 rmhist p)$ & unchanged (e p, c p, r p) \
-\      --> unchanged (rtrner memCh!p, resbar rmhist!p)"
-  (fn _ => [action_simp_tac
-	       (simpset() addsimps [e_def,c_def,r_def,resbar_def]) [] [] 1,
-	    REPEAT (dresolve_tac [temp_use S4_excl, temp_use S5_excl] 1),
-	    auto_tac (MI_fast_css addsimps2 [MemReturn_def,Return_def])
-           ]);
+Goal "|- MemReturn rmCh ires p & $S4 rmhist p & (S5 rmhist p)$\
+\        & unchanged (e p, c p, r p) \
+\        --> unchanged (rtrner memCh!p, resbar rmhist!p)";
+by (action_simp_tac
+      (simpset() addsimps [e_def,c_def,r_def,resbar_def]) [] [] 1);
+by (REPEAT (dresolve_tac [temp_use S4_excl, temp_use S5_excl] 1));
+by (auto_tac (MI_fast_css addsimps2 [MemReturn_def,Return_def]));
+qed "Step1_4_4c";
 
-qed_goal "Step1_4_5a" MemoryImplementation.thy
-   "|- RPCReply crCh rmCh rst p & $S5 rmhist p & (S6 rmhist p)$ & unchanged (e p, c p, m p) \
-\      --> unchanged (rtrner memCh!p, resbar rmhist!p)"
-  (fn _ => [Clarsimp_tac 1,
-            REPEAT (dresolve_tac [temp_use S5_excl, temp_use S6_excl] 1),
-            auto_tac (MI_css addsimps2 [e_def,c_def,m_def, resbar_def]),
-	    auto_tac (MI_css addsimps2 [RPCReply_def,Return_def,S5_def,S_def] 
-                             addSDs2 [MVOKBAnotRF])
-           ]);
+Goal "|- RPCReply crCh rmCh rst p & $S5 rmhist p & (S6 rmhist p)$\
+\        & unchanged (e p, c p, m p) \
+\        --> unchanged (rtrner memCh!p, resbar rmhist!p)";
+by (Clarsimp_tac 1);
+by (REPEAT (dresolve_tac [temp_use S5_excl, temp_use S6_excl] 1));
+by (auto_tac (MI_css addsimps2 [e_def,c_def,m_def, resbar_def]));
+by (auto_tac (MI_css addsimps2 [RPCReply_def,Return_def,S5_def,S_def] 
+                     addSDs2 [MVOKBAnotRF]));
+qed "Step1_4_5a";
 
-qed_goal "Step1_4_5b" MemoryImplementation.thy
-   "|- RPCFail crCh rmCh rst p & $S5 rmhist p & (S6 rmhist p)$ & unchanged (e p, c p, m p) \
-\      --> MemFail memCh (resbar rmhist) p"
-  (fn _ => [Clarsimp_tac 1,
-            dtac (temp_use S6_excl) 1,
-            auto_tac (MI_css addsimps2 [e_def, c_def, m_def, RPCFail_def, Return_def,
-		 		        MemFail_def, resbar_def]),
-	    auto_tac (MI_css addsimps2 [S5_def,S_def])
-           ]);
+Goal "|- RPCFail crCh rmCh rst p & $S5 rmhist p & (S6 rmhist p)$\
+\        & unchanged (e p, c p, m p) \
+\        --> MemFail memCh (resbar rmhist) p";
+by (Clarsimp_tac 1);
+by (dtac (temp_use S6_excl) 1);
+by (auto_tac (MI_css addsimps2 [e_def, c_def, m_def, RPCFail_def, Return_def,
+		 		MemFail_def, resbar_def]));
+by (auto_tac (MI_css addsimps2 [S5_def,S_def]));
+qed "Step1_4_5b";
 
-qed_goal "Step1_4_6a" MemoryImplementation.thy
-   "|- MClkReply memCh crCh cst p & $S6 rmhist p & (S1 rmhist p)$ & unchanged (e p, r p, m p) \
-\      --> MemReturn memCh (resbar rmhist) p"
-  (fn _ => [Clarsimp_tac 1,
-            dtac (temp_use S6_excl) 1,
-            action_simp_tac
-	      (simpset() addsimps [e_def, r_def, m_def, MClkReply_def, MemReturn_def,
-				  Return_def, resbar_def]) [] [] 1,
-	    ALLGOALS Asm_full_simp_tac,  (* simplify if-then-else *)
-	    ALLGOALS (action_simp_tac
-    	              (simpset() addsimps [MClkReplyVal_def,S6_def,S_def])
-		      [] [MVOKBARFnotNR])
-           ]);
+Goal "|- MClkReply memCh crCh cst p & $S6 rmhist p & (S1 rmhist p)$\
+\        & unchanged (e p, r p, m p) \
+\        --> MemReturn memCh (resbar rmhist) p";
+by (Clarsimp_tac 1);
+by (dtac (temp_use S6_excl) 1);
+by (action_simp_tac
+      (simpset() addsimps [e_def,r_def,m_def,MClkReply_def,MemReturn_def,
+		           Return_def,resbar_def]) [] [] 1);
+by (ALLGOALS Asm_full_simp_tac);  (* simplify if-then-else *)
+by (ALLGOALS (action_simp_tac
+                (simpset() addsimps [MClkReplyVal_def,S6_def,S_def])
+	        [] [MVOKBARFnotNR]));
+qed "Step1_4_6a";
 
-qed_goal "Step1_4_6b" MemoryImplementation.thy
-   "|- MClkRetry memCh crCh cst p & $S6 rmhist p & (S3 rmhist p)$   \
-\                & unchanged (e p, r p, m p, rmhist!p) \
-\      --> MemFail memCh (resbar rmhist) p"
-  (fn _ => [Clarsimp_tac 1,
-            dtac (temp_use S3_excl) 1,
-            action_simp_tac
-	       (simpset() addsimps [e_def, r_def, m_def, MClkRetry_def, MemFail_def, resbar_def])
-	       [] [] 1,
-	    auto_tac (MI_css addsimps2 [S6_def,S_def])
-           ]);
+Goal "|- MClkRetry memCh crCh cst p & $S6 rmhist p & (S3 rmhist p)$   \
+\        & unchanged (e p, r p, m p, rmhist!p) \
+\        --> MemFail memCh (resbar rmhist) p";
+by (Clarsimp_tac 1);
+by (dtac (temp_use S3_excl) 1);
+by (action_simp_tac
+      (simpset() addsimps [e_def, r_def, m_def, MClkRetry_def, MemFail_def, resbar_def])
+      [] [] 1);
+by (auto_tac (MI_css addsimps2 [S6_def,S_def]));
+qed "Step1_4_6b";
 
-qed_goal "S_lemma" MemoryImplementation.thy
-   "|- unchanged (e p, c p, r p, m p, rmhist!p) \
-\      --> unchanged (S rmhist ec cc rc cs rs hs1 hs2 p)"
-   (fn _ => [auto_tac (MI_css addsimps2 [e_def,c_def,r_def,m_def,caller_def,rtrner_def,
-					 S_def,Calling_def])
-            ]);
+Goal "|- unchanged (e p, c p, r p, m p, rmhist!p) \
+\        --> unchanged (S rmhist ec cc rc cs rs hs1 hs2 p)";
+by (auto_tac (MI_css addsimps2 [e_def,c_def,r_def,m_def,caller_def,rtrner_def,
+			        S_def,Calling_def]));
+qed "S_lemma";
 
-qed_goal "Step1_4_7H" MemoryImplementation.thy
-   "|- unchanged (e p, c p, r p, m p, rmhist!p) \
-\      --> unchanged (rtrner memCh!p, S1 rmhist p, S2 rmhist p, S3 rmhist p, \
-\                     S4 rmhist p, S5 rmhist p, S6 rmhist p)"
-   (fn _ => [Clarsimp_tac 1,
-             rtac conjI 1,
-             force_tac (MI_css addsimps2 [c_def]) 1,
-             force_tac (MI_css addsimps2 [S1_def,S2_def,S3_def,S4_def,S5_def,S6_def]
-                               addSIs2 [S_lemma]) 1
-            ]);
+Goal "|- unchanged (e p, c p, r p, m p, rmhist!p) \
+\        --> unchanged (rtrner memCh!p, S1 rmhist p, S2 rmhist p, S3 rmhist p, \
+\                       S4 rmhist p, S5 rmhist p, S6 rmhist p)";
+by (Clarsimp_tac 1);
+by (rtac conjI 1);
+by (force_tac (MI_css addsimps2 [c_def]) 1);
+by (force_tac (MI_css addsimps2 [S1_def,S2_def,S3_def,S4_def,S5_def,S6_def]
+                      addSIs2 [S_lemma]) 1);
+qed "Step1_4_7H";
 
-qed_goal "Step1_4_7" MemoryImplementation.thy
-   "|- unchanged (e p, c p, r p, m p, rmhist!p) \
-\      --> unchanged (rtrner memCh!p, resbar rmhist!p, S1 rmhist p, S2 rmhist p, S3 rmhist p, \
-\                     S4 rmhist p, S5 rmhist p, S6 rmhist p)"
-  (fn _ => [rtac actionI 1,
-            rewrite_goals_tac action_rews,
-            rtac impI 1,
-            forward_tac [temp_use Step1_4_7H] 1,
-	    auto_tac (MI_css addsimps2 [e_def,c_def,r_def,m_def,rtrner_def,resbar_def])
-           ]);
-
+Goal "|- unchanged (e p, c p, r p, m p, rmhist!p) \
+\        --> unchanged (rtrner memCh!p, resbar rmhist!p, S1 rmhist p, S2 rmhist p, \
+\                       S3 rmhist p, S4 rmhist p, S5 rmhist p, S6 rmhist p)";
+by (rtac actionI 1);
+by (rewrite_goals_tac action_rews);
+by (rtac impI 1);
+by (forward_tac [temp_use Step1_4_7H] 1);
+by (auto_tac (MI_css addsimps2 [e_def,c_def,r_def,m_def,rtrner_def,resbar_def]));
+qed "Step1_4_7";
 
 (* Frequently needed abbreviation: distinguish between idling and non-idling
    steps of the implementation, and try to solve the idling case by simplification
@@ -354,74 +337,66 @@
 
 (* Steps that leave all variables unchanged are safe, so I may assume
    that some variable changes in the proof that a step is safe. *)
-qed_goal "unchanged_safe" MemoryImplementation.thy
-   "|- (~unchanged (e p, c p, r p, m p, rmhist!p) \
-\        --> [UNext memCh mm (resbar rmhist) p]_(rtrner memCh!p, resbar rmhist!p)) \
-\      --> [UNext memCh mm (resbar rmhist) p]_(rtrner memCh!p, resbar rmhist!p)"
-   (fn _ => [split_idle_tac [square_def] 1,
-             Force_tac 1
-            ]);
+Goal "|- (~unchanged (e p, c p, r p, m p, rmhist!p) \
+\            --> [UNext memCh mm (resbar rmhist) p]_(rtrner memCh!p, resbar rmhist!p)) \
+\        --> [UNext memCh mm (resbar rmhist) p]_(rtrner memCh!p, resbar rmhist!p)";
+by (split_idle_tac [square_def] 1);
+by (Force_tac 1);
+qed "unchanged_safe";
 (* turn into (unsafe, looping!) introduction rule *)
 bind_thm("unchanged_safeI", impI RS (action_use unchanged_safe));
 
-qed_goal "S1safe" MemoryImplementation.thy
-   "|- $S1 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)   \
-\      --> [UNext memCh mm (resbar rmhist) p]_(rtrner memCh!p, resbar rmhist!p)"
-   (fn _ => [Clarsimp_tac 1, 
-             rtac unchanged_safeI 1,
-             rtac idle_squareI 1,
-	     auto_tac (MI_css addSDs2 [Step1_2_1,Step1_4_1])
-	    ]);
+Goal "|- $S1 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)   \
+\        --> [UNext memCh mm (resbar rmhist) p]_(rtrner memCh!p, resbar rmhist!p)";
+by (Clarsimp_tac 1);
+by (rtac unchanged_safeI 1);
+by (rtac idle_squareI 1);
+by (auto_tac (MI_css addSDs2 [Step1_2_1,Step1_4_1]));
+qed "S1safe";
 
-qed_goal "S2safe" MemoryImplementation.thy
-   "|- $S2 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)   \
-\      --> [UNext memCh mm (resbar rmhist) p]_(rtrner memCh!p, resbar rmhist!p)"
-   (fn _ => [Clarsimp_tac 1, 
-             rtac unchanged_safeI 1,
-             rtac idle_squareI 1,
-	     auto_tac (MI_css addSDs2 [Step1_2_2,Step1_4_2])
-	    ]);
+Goal "|- $S2 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)   \
+\        --> [UNext memCh mm (resbar rmhist) p]_(rtrner memCh!p, resbar rmhist!p)";
+by (Clarsimp_tac 1);
+by (rtac unchanged_safeI 1);
+by (rtac idle_squareI 1);
+by (auto_tac (MI_css addSDs2 [Step1_2_2,Step1_4_2]));
+qed "S2safe";
 
-qed_goal "S3safe" MemoryImplementation.thy
-   "|- $S3 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)   \
-\      --> [UNext memCh mm (resbar rmhist) p]_(rtrner memCh!p, resbar rmhist!p)"
-   (fn _ => [Clarsimp_tac 1,
-	     rtac unchanged_safeI 1,
-             auto_tac (MI_css addSDs2 [Step1_2_3]),
-	     auto_tac (MI_css addsimps2 [square_def,UNext_def]
-		              addSDs2 [Step1_4_3a,Step1_4_3b])
-	    ]);
+Goal "|- $S3 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)   \
+\        --> [UNext memCh mm (resbar rmhist) p]_(rtrner memCh!p, resbar rmhist!p)";
+by (Clarsimp_tac 1);
+by (rtac unchanged_safeI 1);
+by (auto_tac (MI_css addSDs2 [Step1_2_3]));
+by (auto_tac (MI_css addsimps2 [square_def,UNext_def] addSDs2 [Step1_4_3a,Step1_4_3b]));
+qed "S3safe";
 
-qed_goal "S4safe" MemoryImplementation.thy
-   "|- $S4 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)  \
-\                   & (!l. $(MemInv mm l)) \
-\      --> [UNext memCh mm (resbar rmhist) p]_(rtrner memCh!p, resbar rmhist!p)"
-   (fn _ => [Clarsimp_tac 1,
-	     rtac unchanged_safeI 1,
-             auto_tac (MI_css addSDs2 [Step1_2_4]),
-	     auto_tac (MI_css addsimps2 [square_def,UNext_def,RNext_def]
-                              addSDs2 [Step1_4_4a,Step1_4_4b,Step1_4_4c])
-	    ]);
+Goal "|- $S4 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)  \
+\        & (!l. $(MemInv mm l)) \
+\        --> [UNext memCh mm (resbar rmhist) p]_(rtrner memCh!p, resbar rmhist!p)";
+by (Clarsimp_tac 1);
+by (rtac unchanged_safeI 1);
+by (auto_tac (MI_css addSDs2 [Step1_2_4]));
+by (auto_tac (MI_css addsimps2 [square_def,UNext_def,RNext_def]
+                     addSDs2 [Step1_4_4a,Step1_4_4b,Step1_4_4c]));
+qed "S4safe";
 
-qed_goal "S5safe" MemoryImplementation.thy
-   "|- $S5 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)  \
-\      --> [UNext memCh mm (resbar rmhist) p]_(rtrner memCh!p, resbar rmhist!p)"
-   (fn _ => [Clarsimp_tac 1,
-	     rtac unchanged_safeI 1,
-             auto_tac (MI_css addSDs2 [Step1_2_5]),
-	     auto_tac (MI_css addsimps2 [square_def,UNext_def]
-		              addSDs2 [Step1_4_5a,Step1_4_5b])
-	    ]);
+Goal "|- $S5 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)  \
+\        --> [UNext memCh mm (resbar rmhist) p]_(rtrner memCh!p, resbar rmhist!p)";
+by (Clarsimp_tac 1);
+by (rtac unchanged_safeI 1);
+by (auto_tac (MI_css addSDs2 [Step1_2_5]));
+by (auto_tac (MI_css addsimps2 [square_def,UNext_def]
+	             addSDs2 [Step1_4_5a,Step1_4_5b]));
+qed "S5safe";
 
-qed_goal "S6safe" MemoryImplementation.thy
-   "|- $S6 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)   \
-\      --> [UNext memCh mm (resbar rmhist) p]_(rtrner memCh!p, resbar rmhist!p)"
-   (fn _ => [Clarsimp_tac 1,
-	     rtac unchanged_safeI 1,
-             auto_tac (MI_css addSDs2 [Step1_2_6]),
-	     auto_tac (MI_css addsimps2 [square_def,UNext_def,RNext_def]
-		              addSDs2 [Step1_4_6a,Step1_4_6b])
-	    ]);
+Goal "|- $S6 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)   \
+\        --> [UNext memCh mm (resbar rmhist) p]_(rtrner memCh!p, resbar rmhist!p)";
+by (Clarsimp_tac 1);
+by (rtac unchanged_safeI 1);
+by (auto_tac (MI_css addSDs2 [Step1_2_6]));
+by (auto_tac (MI_css addsimps2 [square_def,UNext_def,RNext_def]
+	             addSDs2 [Step1_4_6a,Step1_4_6b]));
+qed "S6safe";
 
 (* ----------------------------------------------------------------------
    Step 1.5: Temporal refinement proof, based on previous steps.
@@ -429,7 +404,343 @@
 
 section "The liveness part";
 
-use "MIlive.ML";
+(* Liveness assertions for the different implementation states, based on the
+   fairness conditions. Prove subgoals of WF1 / SF1 rules as separate lemmas
+   for readability. Reuse action proofs from safety part.
+*)
+
+(* ------------------------------ State S1 ------------------------------ *)
+
+Goal "|- $S1 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)  \
+\        --> (S1 rmhist p)$ | (S2 rmhist p)$";
+by (split_idle_tac [] 1);
+by (auto_tac (MI_css addSDs2 [Step1_2_1]));
+qed "S1_successors";
+
+(* Show that the implementation can satisfy the high-level fairness requirements
+   by entering the state S1 infinitely often.
+*)
+
+Goal "|- S1 rmhist p --> \
+\        ~Enabled (<RNext memCh mm (resbar rmhist) p>_(rtrner memCh!p, resbar rmhist!p))";
+by (action_simp_tac (simpset() addsimps [angle_def,S_def,S1_def])
+	            [notI] [enabledE,temp_elim Memoryidle] 1);
+by (Force_tac 1);
+qed "S1_RNextdisabled";
+
+Goal "|- S1 rmhist p --> \
+\        ~Enabled (<MemReturn memCh (resbar rmhist) p>_(rtrner memCh!p, resbar rmhist!p))";
+by (action_simp_tac 
+      (simpset() addsimps [angle_def,MemReturn_def,Return_def,S_def,S1_def])
+      [notI] [enabledE] 1);
+qed "S1_Returndisabled";
+
+Goal "|- []<>S1 rmhist p   \
+\        --> WF(RNext memCh mm (resbar rmhist) p)_(rtrner memCh!p, resbar rmhist!p)";
+by (auto_tac (MI_css addsimps2 [WF_alt]
+		     addSIs2 [S1_RNextdisabled] addSEs2 [STL4E,DmdImplE]));
+qed "RNext_fair";
+
+Goal "|- []<>S1 rmhist p   \
+\        --> WF(MemReturn memCh (resbar rmhist) p)_(rtrner memCh!p, resbar rmhist!p)";
+by (auto_tac (MI_css addsimps2 [WF_alt]
+		     addSIs2 [S1_Returndisabled] addSEs2 [STL4E,DmdImplE]));
+qed "Return_fair";
+
+(* ------------------------------ State S2 ------------------------------ *)
+
+Goal "|- $S2 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)   \
+\        --> (S2 rmhist p)$ | (S3 rmhist p)$";
+by (split_idle_tac [] 1);
+by (auto_tac (MI_css addSDs2 [Step1_2_2]));
+qed "S2_successors";
+
+Goal "|- ($S2 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p))\
+\        & <MClkFwd memCh crCh cst p>_(c p) \
+\        --> (S3 rmhist p)$";
+by (auto_tac (MI_css addsimps2 [angle_def] addSDs2 [Step1_2_2]));
+qed "S2MClkFwd_successors";
+
+Goal "|- $S2 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)\
+\        --> $Enabled (<MClkFwd memCh crCh cst p>_(c p))";
+by (auto_tac (MI_css addsimps2 [c_def] addSIs2 [MClkFwd_ch_enabled,MClkFwd_enabled]));
+by (cut_facts_tac [MI_base] 1);
+by (blast_tac (claset() addDs [base_pair]) 1);
+by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [S_def,S2_def])));
+qed "S2MClkFwd_enabled";
+
+Goal "|- [](ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p))\
+\        & WF(MClkFwd memCh crCh cst p)_(c p) \
+\        --> (S2 rmhist p ~> S3 rmhist p)";
+by (REPEAT (resolve_tac [WF1,S2_successors,
+                         S2MClkFwd_successors,S2MClkFwd_enabled] 1));
+qed "S2_live";
+
+(* ------------------------------ State S3 ------------------------------ *)
+
+Goal "|- $S3 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)\
+\        --> (S3 rmhist p)$ | (S4 rmhist p | S6 rmhist p)$";
+by (split_idle_tac [] 1);
+by (auto_tac (MI_css addSDs2 [Step1_2_3]));
+qed "S3_successors";
+
+Goal "|- ($S3 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p))\
+\        & <RPCNext crCh rmCh rst p>_(r p) \
+\        --> (S4 rmhist p | S6 rmhist p)$";
+by (auto_tac (MI_css addsimps2 [angle_def] addSDs2 [Step1_2_3]));
+qed "S3RPC_successors";
+
+Goal "|- $S3 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)\
+\        --> $Enabled (<RPCNext crCh rmCh rst p>_(r p))";
+by (auto_tac (MI_css addsimps2 [r_def]
+		     addSIs2 [RPCFail_Next_enabled,RPCFail_enabled]));
+by (cut_facts_tac [MI_base] 1);
+by (blast_tac (claset() addDs [base_pair]) 1);
+by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [S_def,S3_def])));
+qed "S3RPC_enabled";
+
+Goal "|- [](ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p))\
+\        & WF(RPCNext crCh rmCh rst p)_(r p) \
+\        --> (S3 rmhist p ~> S4 rmhist p | S6 rmhist p)";
+by (REPEAT (resolve_tac [WF1,S3_successors,S3RPC_successors,S3RPC_enabled] 1));
+qed "S3_live";
+
+(* ------------- State S4 -------------------------------------------------- *)
+
+Goal"|- $S4 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) \
+\       & (ALL l. $MemInv mm l)  \
+\       --> (S4 rmhist p)$ | (S5 rmhist p)$";
+by (split_idle_tac [] 1);
+by (auto_tac (MI_css addSDs2 [Step1_2_4]));
+qed "S4_successors";
+
+(* --------- State S4a: S4 /\ (ires p = NotAResult) ------------------------ *)
+
+Goal "|- $(S4 rmhist p & ires!p = #NotAResult) \
+\        & ImpNext p & [HNext rmhist p]_(c p,r p,m p,rmhist!p) & (ALL l. $MemInv mm l)\
+\        --> (S4 rmhist p & ires!p = #NotAResult)$  \
+\            | ((S4 rmhist p & ires!p ~= #NotAResult) | S5 rmhist p)$";
+by (split_idle_tac [m_def] 1);
+by (auto_tac (MI_css addSDs2 [Step1_2_4]));
+qed "S4a_successors";
+
+Goal "|- ($(S4 rmhist p & ires!p = #NotAResult)  \
+\        & ImpNext p & [HNext rmhist p]_(c p,r p,m p,rmhist!p) & (ALL l. $MemInv mm l))\
+\        & <RNext rmCh mm ires p>_(m p) \
+\        --> ((S4 rmhist p & ires!p ~= #NotAResult) | S5 rmhist p)$";
+by (auto_tac (MI_css addsimps2 [angle_def]
+		     addSDs2 [Step1_2_4, ReadResult, WriteResult]));
+qed "S4aRNext_successors";
+
+Goal "|- $(S4 rmhist p & ires!p = #NotAResult) \
+\        & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) & (ALL l. $MemInv mm l)\
+\        --> $Enabled (<RNext rmCh mm ires p>_(m p))";
+by (auto_tac (MI_css addsimps2 [m_def] addSIs2 [RNext_enabled]));
+by (cut_facts_tac [MI_base] 1);
+by (blast_tac (claset() addDs [base_pair]) 1);
+by (asm_full_simp_tac (simpset() addsimps [S_def,S4_def]) 1);
+qed "S4aRNext_enabled";
+
+Goal "|- [](ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)\
+\        & (ALL l. $MemInv mm l)) & WF(RNext rmCh mm ires p)_(m p) \
+\        --> (S4 rmhist p & ires!p = #NotAResult  \
+\             ~> (S4 rmhist p & ires!p ~= #NotAResult) | S5 rmhist p)";
+by (REPEAT (resolve_tac [WF1, S4a_successors, S4aRNext_successors, S4aRNext_enabled] 1));
+qed "S4a_live";
+
+(* ---------- State S4b: S4 /\ (ires p # NotAResult) --------------------------- *)
+
+Goal "|- $(S4 rmhist p & ires!p ~= #NotAResult)  \
+\        & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) & (ALL l. $MemInv mm l)\
+\        --> (S4 rmhist p & ires!p ~= #NotAResult)$ | (S5 rmhist p)$";
+by (split_idle_tac [m_def] 1);
+by (auto_tac (MI_css addSDs2 [WriteResult,Step1_2_4,ReadResult]));
+qed "S4b_successors";
+
+Goal "|- ($(S4 rmhist p & ires!p ~= #NotAResult)  \
+\        & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) \
+\        & (ALL l. $MemInv mm l)) & <MemReturn rmCh ires p>_(m p) \
+\        --> (S5 rmhist p)$";
+by (force_tac (MI_css addsimps2 [angle_def] addSDs2 [Step1_2_4]
+                      addDs2 [ReturnNotReadWrite]) 1);
+qed "S4bReturn_successors";
+
+Goal "|- $(S4 rmhist p & ires!p ~= #NotAResult)  \
+\        & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)\
+\        & (ALL l. $MemInv mm l)  \
+\        --> $Enabled (<MemReturn rmCh ires p>_(m p))";
+by (auto_tac (MI_css addsimps2 [m_def] addSIs2 [MemReturn_enabled]));
+by (cut_facts_tac [MI_base] 1);
+by (blast_tac (claset() addDs [base_pair]) 1);
+by (asm_full_simp_tac (simpset() addsimps [S_def,S4_def]) 1);
+qed "S4bReturn_enabled";
+
+Goal "|- [](ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) & (!l. $MemInv mm l)) \
+\        & WF(MemReturn rmCh ires p)_(m p) \
+\        --> (S4 rmhist p & ires!p ~= #NotAResult ~> S5 rmhist p)";
+by (REPEAT (resolve_tac [WF1, S4b_successors,S4bReturn_successors, S4bReturn_enabled] 1));
+qed "S4b_live";
+
+(* ------------------------------ State S5 ------------------------------ *)
+
+Goal "|- $S5 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) \
+\        --> (S5 rmhist p)$ | (S6 rmhist p)$";
+by (split_idle_tac [] 1);
+by (auto_tac (MI_css addSDs2 [Step1_2_5]));
+qed "S5_successors";
+
+Goal "|- ($S5 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)) \
+\        & <RPCNext crCh rmCh rst p>_(r p) \
+\        --> (S6 rmhist p)$";
+by (auto_tac (MI_css addsimps2 [angle_def] addSDs2 [Step1_2_5]));
+qed "S5RPC_successors";
+
+Goal "|- $S5 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) \
+\        --> $Enabled (<RPCNext crCh rmCh rst p>_(r p))";
+by (auto_tac (MI_css addsimps2 [r_def]
+		     addSIs2 [RPCFail_Next_enabled, RPCFail_enabled]));
+by (cut_facts_tac [MI_base] 1);
+by (blast_tac (claset() addDs [base_pair]) 1);
+by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [S_def,S5_def])));
+qed "S5RPC_enabled";
+
+Goal "|- [](ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p))\
+\        & WF(RPCNext crCh rmCh rst p)_(r p) \
+\        --> (S5 rmhist p ~> S6 rmhist p)";
+by (REPEAT (resolve_tac [WF1,S5_successors,S5RPC_successors,S5RPC_enabled] 1));
+qed "S5_live";
+
+(* ------------------------------ State S6 ------------------------------ *)
+
+Goal "|- $S6 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) \
+\        --> (S1 rmhist p)$ | (S3 rmhist p)$ | (S6 rmhist p)$";
+by (split_idle_tac [] 1);
+by (auto_tac (MI_css addSDs2 [Step1_2_6]));
+qed "S6_successors";
+
+Goal "|- ($S6 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)) \
+\        & <MClkReply memCh crCh cst p>_(c p) \
+\        --> (S1 rmhist p)$";
+by (auto_tac (MI_css addsimps2 [angle_def] addSDs2 [Step1_2_6, MClkReplyNotRetry]));
+qed "S6MClkReply_successors";
+
+Goal "|- $ImpInv rmhist p & <MClkReply memCh crCh cst p>_(c p) --> $S6 rmhist p";
+by (action_simp_tac
+      (simpset() addsimps [angle_def,MClkReply_def,Return_def,
+		     ImpInv_def,S_def,S1_def,S2_def,S3_def,S4_def,S5_def])
+      [] [] 1);
+qed "MClkReplyS6";
+
+Goal "|- S6 rmhist p --> Enabled (<MClkReply memCh crCh cst p>_(c p))";
+by (auto_tac (MI_css addsimps2 [c_def] addSIs2 [MClkReply_enabled]));
+by (cut_facts_tac [MI_base] 1);
+by (blast_tac (claset() addDs [base_pair]) 1);
+by (ALLGOALS (action_simp_tac (simpset() addsimps [S_def,S6_def]) [] []));
+qed "S6MClkReply_enabled";
+
+Goal "|- [](ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) & $(ImpInv rmhist p))\
+\        & SF(MClkReply memCh crCh cst p)_(c p) & []<>(S6 rmhist p)  \
+\        --> []<>(S1 rmhist p)";
+by (Clarsimp_tac 1);
+by (subgoal_tac "sigma |= []<>(<MClkReply memCh crCh cst p>_(c p))" 1);
+by (etac InfiniteEnsures 1); 
+by (atac 1);
+by (action_simp_tac (simpset()) []
+	            (map temp_elim [MClkReplyS6,S6MClkReply_successors]) 1);
+by (auto_tac (MI_css addsimps2 [SF_def]));
+by (etac swap 1);
+by (auto_tac (MI_css addSIs2 [S6MClkReply_enabled] addSEs2 [STL4E, DmdImplE]));
+qed "S6_live";
+
+(* --------------- aggregate leadsto properties----------------------------- *)
+
+Goal "sigma |= S5 rmhist p ~> S6 rmhist p \
+\     ==> sigma |= (S5 rmhist p | S6 rmhist p) ~> S6 rmhist p";
+by (auto_tac (MI_css addSIs2 [LatticeDisjunctionIntro, LatticeReflexivity]));
+qed "S5S6LeadstoS6";
+
+Goal "[| sigma |= S4 rmhist p & ires!p ~= #NotAResult ~> S5 rmhist p;\
+\        sigma |= S5 rmhist p ~> S6 rmhist p |]  \
+\     ==> sigma |= (S4 rmhist p & ires!p ~= #NotAResult) | S5 rmhist p | S6 rmhist p \
+\                   ~> S6 rmhist p";
+by (auto_tac (MI_css addSIs2 [LatticeDisjunctionIntro,S5S6LeadstoS6]
+		     addIs2 [LatticeTransitivity]));
+qed "S4bS5S6LeadstoS6";
+
+Goal "[| sigma |= S4 rmhist p & ires!p = #NotAResult \
+\                 ~> (S4 rmhist p & ires!p ~= #NotAResult) | S5 rmhist p; \
+\        sigma |= S4 rmhist p & ires!p ~= #NotAResult ~> S5 rmhist p;  \
+\        sigma |= S5 rmhist p ~> S6 rmhist p |]  \
+\     ==> sigma |= S4 rmhist p | S5 rmhist p | S6 rmhist p ~> S6 rmhist p";
+by (subgoal_tac 
+     "sigma |= (S4 rmhist p & ires!p = #NotAResult)\
+\            | (S4 rmhist p & ires!p ~= #NotAResult)\
+\            | S5 rmhist p | S6 rmhist p ~> S6 rmhist p" 1);
+ by (eres_inst_tac 
+      [("G", "PRED ((S4 rmhist p & ires!p = #NotAResult)\
+\                | (S4 rmhist p & ires!p ~= #NotAResult)\
+\                | S5 rmhist p | S6 rmhist p)")] 
+      (temp_use LatticeTransitivity) 1);
+ by (force_tac (MI_css addsimps2 Init_defs addSIs2 [ImplLeadsto_gen, necT]) 1);
+by (rtac (temp_use LatticeDisjunctionIntro) 1);
+by (etac (temp_use LatticeTransitivity) 1);
+by (etac (temp_use LatticeTriangle2) 1); 
+by (atac 1);
+by (auto_tac (MI_css addSIs2 [S4bS5S6LeadstoS6]));
+qed "S4S5S6LeadstoS6";
+
+Goal "[| sigma |= S3 rmhist p ~> S4 rmhist p | S6 rmhist p;   \
+\        sigma |= S4 rmhist p & ires!p = #NotAResult \
+\                 ~> (S4 rmhist p & ires!p ~= #NotAResult) | S5 rmhist p; \
+\        sigma |= S4 rmhist p & ires!p ~= #NotAResult ~> S5 rmhist p;  \
+\        sigma |= S5 rmhist p ~> S6 rmhist p |]  \
+\     ==> sigma |= S3 rmhist p | S4 rmhist p | S5 rmhist p | S6 rmhist p ~> S6 rmhist p";
+by (rtac (temp_use LatticeDisjunctionIntro) 1);
+by (etac (temp_use LatticeTriangle2) 1);
+by (rtac (S4S5S6LeadstoS6 RS (temp_use LatticeTransitivity)) 1);
+by (auto_tac (MI_css addSIs2 [S4S5S6LeadstoS6,necT]
+	             addIs2 [ImplLeadsto_gen] addsimps2 Init_defs));
+qed "S3S4S5S6LeadstoS6";
+
+Goal "[| sigma |= S2 rmhist p ~> S3 rmhist p; \
+\        sigma |= S3 rmhist p ~> S4 rmhist p | S6 rmhist p;   \
+\        sigma |= S4 rmhist p & ires!p = #NotAResult \
+\                 ~> S4 rmhist p & ires!p ~= #NotAResult | S5 rmhist p; \
+\        sigma |= S4 rmhist p & ires!p ~= #NotAResult ~> S5 rmhist p;  \
+\        sigma |= S5 rmhist p ~> S6 rmhist p |]  \
+\     ==> sigma |= S2 rmhist p | S3 rmhist p | S4 rmhist p | S5 rmhist p | S6 rmhist p \
+\                  ~> S6 rmhist p";
+by (rtac (temp_use LatticeDisjunctionIntro) 1);
+by (rtac (temp_use LatticeTransitivity) 1); 
+by (atac 2);
+by (rtac (S3S4S5S6LeadstoS6 RS (temp_use LatticeTransitivity)) 1);
+by (auto_tac (MI_css addSIs2 [S3S4S5S6LeadstoS6,necT]
+	             addIs2 [ImplLeadsto_gen] addsimps2 Init_defs));
+qed "S2S3S4S5S6LeadstoS6";
+
+Goal "[| sigma |= []ImpInv rmhist p; \
+\        sigma |= S2 rmhist p ~> S3 rmhist p; \
+\        sigma |= S3 rmhist p ~> S4 rmhist p | S6 rmhist p; \
+\        sigma |= S4 rmhist p & ires!p = #NotAResult \
+\                 ~> S4 rmhist p & ires!p ~= #NotAResult | S5 rmhist p; \
+\        sigma |= S4 rmhist p & ires!p ~= #NotAResult ~> S5 rmhist p;  \
+\        sigma |= S5 rmhist p ~> S6 rmhist p |] \
+\     ==> sigma |= ~S1 rmhist p ~> S6 rmhist p";
+by (rtac (S2S3S4S5S6LeadstoS6 RS (temp_use LatticeTransitivity)) 1);
+by (TRYALL atac);
+by (etac (temp_use INV_leadsto) 1);
+by (rtac (temp_use ImplLeadsto_gen) 1);
+by (rtac (temp_use necT) 1);
+by (auto_tac (MI_css addsimps2 ImpInv_def::Init_defs addSIs2 [necT]));
+qed "NotS1LeadstoS6";
+
+Goal "[| sigma |= ~S1 rmhist p ~> S6 rmhist p; \
+\        sigma |= []<>S6 rmhist p --> []<>S1 rmhist p |] \
+\     ==> sigma |= []<>S1 rmhist p";
+by (rtac classical 1);
+by (asm_full_simp_tac (simpset() addsimps [temp_use NotBox, NotDmd]) 1);
+by (auto_tac (MI_css addSEs2 [mp,leadsto_infinite] addSDs2 [DBImplBD]));
+qed "S1Infinite";
 
 section "Refinement proof (step 1.5)";
 
@@ -437,148 +748,134 @@
    a. memory invariant
    b. "implementation invariant": always in states S1,...,S6
 *)
-qed_goal "Step1_5_1a" MemoryImplementation.thy 
-   "|- IPImp p --> (!l. []$MemInv mm l)"
-   (fn _ => [auto_tac (MI_css addsimps2 [IPImp_def,box_stp_act]
-			      addSIs2 [MemoryInvariantAll])
-	    ]);
+Goal "|- IPImp p --> (ALL l. []$MemInv mm l)";
+by (auto_tac (MI_css addsimps2 [IPImp_def,box_stp_act]
+	             addSIs2 [MemoryInvariantAll]));
+qed "Step1_5_1a";
 
-qed_goal "Step1_5_1b" MemoryImplementation.thy
-   "|- Init(ImpInit p & HInit rmhist p) & [](ImpNext p) \
-\      & [][HNext rmhist p]_(c p, r p, m p, rmhist!p) & [](!l. $MemInv mm l) \
-\      --> []ImpInv rmhist p"
-   (fn _ => [inv_tac MI_css 1,
-	     auto_tac (MI_css addsimps2 [Init_def, ImpInv_def, box_stp_act]
-                              addSDs2 [Step1_1]
-		              addDs2 [S1_successors,S2_successors,S3_successors,
-			              S4_successors,S5_successors,S6_successors])
-            ]);
+Goal "|- Init(ImpInit p & HInit rmhist p) & [](ImpNext p) \
+\        & [][HNext rmhist p]_(c p, r p, m p, rmhist!p) & [](ALL l. $MemInv mm l) \
+\        --> []ImpInv rmhist p";
+by (inv_tac MI_css 1);
+by (auto_tac (MI_css addsimps2 [Init_def, ImpInv_def, box_stp_act]
+                     addSDs2 [Step1_1]
+	             addDs2 [S1_successors,S2_successors,S3_successors,
+		             S4_successors,S5_successors,S6_successors]));
+qed "Step1_5_1b";
 
 (*** Initialization ***)
-qed_goal "Step1_5_2a" MemoryImplementation.thy
-   "|- Init(ImpInit p & HInit rmhist p) --> Init(PInit (resbar rmhist) p)"
-   (fn _ => [auto_tac (MI_css addsimps2 [Init_def]
-                              addSIs2 [Step1_1,Step1_3])
-            ]);
+Goal "|- Init(ImpInit p & HInit rmhist p) --> Init(PInit (resbar rmhist) p)";
+by (auto_tac (MI_css addsimps2 [Init_def] addSIs2 [Step1_1,Step1_3]));
+qed "Step1_5_2a";
 
 (*** step simulation ***)
-qed_goal "Step1_5_2b" MemoryImplementation.thy
-   "|- [](ImpNext p & [HNext rmhist p]_(c p, r p, m p, rmhist!p)   \
-\                   & $ImpInv rmhist p & (!l. $MemInv mm l))       \
-\      --> [][UNext memCh mm (resbar rmhist) p]_(rtrner memCh!p, resbar rmhist!p)"
-   (fn _ => [auto_tac (MI_css 
-                          addsimps2 [ImpInv_def] addSEs2 [STL4E]
-                          addSDs2 [S1safe,S2safe,S3safe,S4safe,S5safe,S6safe])
-            ]);
-
+Goal "|- [](ImpNext p & [HNext rmhist p]_(c p, r p, m p, rmhist!p)   \
+\        & $ImpInv rmhist p & (!l. $MemInv mm l)) \
+\        --> [][UNext memCh mm (resbar rmhist) p]_(rtrner memCh!p, resbar rmhist!p)";
+by (auto_tac (MI_css addsimps2 [ImpInv_def] addSEs2 [STL4E]
+                     addSDs2 [S1safe,S2safe,S3safe,S4safe,S5safe,S6safe]));
+qed "Step1_5_2b";
 
 (*** Liveness ***)
-qed_goal "GoodImpl" MemoryImplementation.thy
-   "|- IPImp p & HistP rmhist p  \
-\      -->   Init(ImpInit p & HInit rmhist p)   \
-\          & [](ImpNext p & [HNext rmhist p]_(c p, r p, m p, rmhist!p)) \
-\          & [](!l. $MemInv mm l) & []($ImpInv rmhist p) \
-\          & ImpLive p"
-   (fn _ => [Clarsimp_tac 1,
-	     subgoal_tac
-	       "sigma |= Init(ImpInit p & HInit rmhist p) \
-\                        & [](ImpNext p) \
-\                        & [][HNext rmhist p]_(c p, r p, m p, rmhist!p) \
-\                        & [](!l. $MemInv mm l)" 1,
-	     auto_tac (MI_css addsimps2 [split_box_conj,box_stp_act] addSDs2 [Step1_5_1b]),
-	     force_tac (MI_css addsimps2 [IPImp_def,MClkIPSpec_def,RPCIPSpec_def,RPSpec_def,
-					  ImpLive_def,c_def,r_def,m_def]) 1,
-	     force_tac (MI_css addsimps2 [IPImp_def,MClkIPSpec_def,RPCIPSpec_def,RPSpec_def,
-					  HistP_def,Init_def,ImpInit_def]) 1,
-	     force_tac (MI_css addsimps2 [IPImp_def,MClkIPSpec_def,RPCIPSpec_def,RPSpec_def,
-					  ImpNext_def,c_def,r_def,m_def,split_box_conj]) 1,
-	     force_tac (MI_css addsimps2 [HistP_def]) 1,
-             force_tac (MI_css addsimps2 [temp_use allT] addSDs2 [Step1_5_1a]) 1
-	    ]);
+Goal "|- IPImp p & HistP rmhist p  \
+\        -->   Init(ImpInit p & HInit rmhist p)   \
+\            & [](ImpNext p & [HNext rmhist p]_(c p, r p, m p, rmhist!p)) \
+\            & [](ALL l. $MemInv mm l) & []($ImpInv rmhist p) \
+\            & ImpLive p";
+by (Clarsimp_tac 1);
+by (subgoal_tac
+      "sigma |= Init(ImpInit p & HInit rmhist p) \
+\             & [](ImpNext p) \
+\             & [][HNext rmhist p]_(c p, r p, m p, rmhist!p) \
+\             & [](ALL l. $MemInv mm l)" 1);
+by (auto_tac (MI_css addsimps2 [split_box_conj,box_stp_act] addSDs2 [Step1_5_1b]));
+by (force_tac (MI_css addsimps2 [IPImp_def,MClkIPSpec_def,RPCIPSpec_def,RPSpec_def,
+				 ImpLive_def,c_def,r_def,m_def]) 1);
+by (force_tac (MI_css addsimps2 [IPImp_def,MClkIPSpec_def,RPCIPSpec_def,RPSpec_def,
+	                         HistP_def,Init_def,ImpInit_def]) 1);
+by (force_tac (MI_css addsimps2 [IPImp_def,MClkIPSpec_def,RPCIPSpec_def,RPSpec_def,
+                                 ImpNext_def,c_def,r_def,m_def,split_box_conj]) 1);
+by (force_tac (MI_css addsimps2 [HistP_def]) 1);
+by (force_tac (MI_css addsimps2 [temp_use allT] addSDs2 [Step1_5_1a]) 1);
+qed "GoodImpl";
 
 (* The implementation is infinitely often in state S1... *)
-qed_goal "Step1_5_3a" MemoryImplementation.thy
-   "|- [](ImpNext p & [HNext rmhist p]_(c p, r p, m p, rmhist!p)) \
-\      & [](!l. $MemInv mm l)  \
-\      & []($ImpInv rmhist p) & ImpLive p  \
-\      --> []<>S1 rmhist p"
-   (fn _ => [clarsimp_tac (MI_css addsimps2 [ImpLive_def]) 1,
-             rtac S1Infinite 1,
-	     force_tac (MI_css
-			  addsimps2 [split_box_conj,box_stp_act]
-			  addSIs2 [NotS1LeadstoS6,S2_live,S3_live,S4a_live,S4b_live,S5_live]) 1,
-             auto_tac (MI_css addsimps2 [split_box_conj] addSIs2 [S6_live])
-            ]);
+Goal "|- [](ImpNext p & [HNext rmhist p]_(c p, r p, m p, rmhist!p)) \
+\        & [](ALL l. $MemInv mm l)  \
+\        & []($ImpInv rmhist p) & ImpLive p  \
+\        --> []<>S1 rmhist p";
+by (clarsimp_tac (MI_css addsimps2 [ImpLive_def]) 1);
+by (rtac S1Infinite 1);
+by (force_tac
+      (MI_css addsimps2 [split_box_conj,box_stp_act]
+              addSIs2 [NotS1LeadstoS6,S2_live,S3_live,S4a_live,S4b_live,S5_live]) 1);
+by (auto_tac (MI_css addsimps2 [split_box_conj] addSIs2 [S6_live]));
+qed "Step1_5_3a";
 
-(* ... which implies that it satisfies the fairness requirements of the specification *)
-qed_goal "Step1_5_3b" MemoryImplementation.thy
-   "|- [](ImpNext p & [HNext rmhist p]_(c p, r p, m p, rmhist!p)) \
-\      & [](!l. $MemInv mm l) & []($ImpInv rmhist p) & ImpLive p  \
-\      --> WF(RNext memCh mm (resbar rmhist) p)_(rtrner memCh!p, resbar rmhist!p)"
-   (fn _ => [ auto_tac (MI_css addSIs2 [RNext_fair,Step1_5_3a]) ]);
+(* ... and therefore satisfies the fairness requirements of the specification *)
+Goal "|- [](ImpNext p & [HNext rmhist p]_(c p, r p, m p, rmhist!p)) \
+\        & [](ALL l. $MemInv mm l) & []($ImpInv rmhist p) & ImpLive p  \
+\        --> WF(RNext memCh mm (resbar rmhist) p)_(rtrner memCh!p, resbar rmhist!p)";
+by (auto_tac (MI_css addSIs2 [RNext_fair,Step1_5_3a]));
+qed "Step1_5_3b";
 
-qed_goal "Step1_5_3c" MemoryImplementation.thy
-   "|- [](ImpNext p & [HNext rmhist p]_(c p, r p, m p, rmhist!p)) \
-\      & [](!l. $MemInv mm l) & []($ImpInv rmhist p) & ImpLive p  \
-\      --> WF(MemReturn memCh (resbar rmhist) p)_(rtrner memCh!p, resbar rmhist!p)"
-   (fn _ => [ auto_tac (MI_css addSIs2 [Return_fair,Step1_5_3a]) ]);
-
+Goal "|- [](ImpNext p & [HNext rmhist p]_(c p, r p, m p, rmhist!p)) \
+\        & [](ALL l. $MemInv mm l) & []($ImpInv rmhist p) & ImpLive p  \
+\        --> WF(MemReturn memCh (resbar rmhist) p)_(rtrner memCh!p, resbar rmhist!p)";
+by (auto_tac (MI_css addSIs2 [Return_fair,Step1_5_3a]));
+qed "Step1_5_3c";
 
 (* QED step of step 1 *)
-qed_goal "Step1" MemoryImplementation.thy
-   "|- IPImp p & HistP rmhist p --> UPSpec memCh mm (resbar rmhist) p"
-   (fn _ => [auto_tac
-               (MI_css addsimps2 [UPSpec_def,split_box_conj]
-		       addSDs2 [GoodImpl]
-                       addSIs2 [Step1_5_2a,Step1_5_2b,Step1_5_3b,Step1_5_3c])
-            ]);
-
+Goal "|- IPImp p & HistP rmhist p --> UPSpec memCh mm (resbar rmhist) p";
+by (auto_tac (MI_css addsimps2 [UPSpec_def,split_box_conj]
+		     addSDs2 [GoodImpl]
+                     addSIs2 [Step1_5_2a,Step1_5_2b,Step1_5_3b,Step1_5_3c]));
+qed "Step1";
 
 (* ------------------------------ Step 2 ------------------------------ *)
 section "Step 2";
 
-qed_goal "Step2_2a" MemoryImplementation.thy
-   "|- Write rmCh mm ires p l & ImpNext p & [HNext rmhist p]_(c p, r p, m p, rmhist!p) \
-\      & $ImpInv rmhist p  \
-\      --> (S4 rmhist p)` & unchanged (e p, c p, r p, rmhist!p)"
-   (fn _ => [Clarsimp_tac 1,
-             dtac (action_use WriteS4) 1, atac 1,
-             split_idle_tac [] 1,
-             auto_tac (MI_css addsimps2 [ImpNext_def] 
-                              addSDs2 [S4EnvUnch,S4ClerkUnch,S4RPCUnch]),
-             auto_tac (MI_css addsimps2 [square_def] addDs2 [S4Write])
-            ]);
+Goal "|- Write rmCh mm ires p l & ImpNext p\
+\        & [HNext rmhist p]_(c p, r p, m p, rmhist!p) \
+\        & $ImpInv rmhist p  \
+\        --> (S4 rmhist p)$ & unchanged (e p, c p, r p, rmhist!p)";
+by (Clarsimp_tac 1);
+by (dtac (action_use WriteS4) 1); 
+by (atac 1);
+by (split_idle_tac [] 1);
+by (auto_tac (MI_css addsimps2 [ImpNext_def] 
+                     addSDs2 [S4EnvUnch,S4ClerkUnch,S4RPCUnch]));
+by (auto_tac (MI_css addsimps2 [square_def] addDs2 [S4Write]));
+qed "Step2_2a";
 
-qed_goal "Step2_2" MemoryImplementation.thy
-   "|-   (!p. ImpNext p) \
-\      & (!p. [HNext rmhist p]_(c p, r p, m p, rmhist!p)) \
-\      & (!p. $ImpInv rmhist p) \
-\      & [? q. Write rmCh mm ires q l]_(mm!l) \
-\      --> [? q. Write memCh mm (resbar rmhist) q l]_(mm!l)"
-   (fn _ => [auto_tac (MI_css addSIs2 [squareCI] addSEs2 [squareE]),
-             REPEAT (ares_tac [exI, action_use Step1_4_4b] 1),
-             force_tac (MI_css addSIs2 [WriteS4]) 1,
-             auto_tac (MI_css addSDs2 [Step2_2a])
-            ]);
+Goal "|-   (ALL p. ImpNext p) \
+\        & (ALL p. [HNext rmhist p]_(c p, r p, m p, rmhist!p)) \
+\        & (ALL p. $ImpInv rmhist p) \
+\        & [EX q. Write rmCh mm ires q l]_(mm!l) \
+\        --> [EX q. Write memCh mm (resbar rmhist) q l]_(mm!l)";
+by (auto_tac (MI_css addSIs2 [squareCI] addSEs2 [squareE]));
+by (REPEAT (ares_tac [exI, action_use Step1_4_4b] 1));
+by (force_tac (MI_css addSIs2 [WriteS4]) 1);
+by (auto_tac (MI_css addSDs2 [Step2_2a]));
+qed "Step2_2";
 
-qed_goal "Step2_lemma" MemoryImplementation.thy
-   "|-  [](  (!p. ImpNext p) \
-\          & (!p. [HNext rmhist p]_(c p, r p, m p, rmhist!p)) \
-\          & (!p. $ImpInv rmhist p) \
-\          & [? q. Write rmCh mm ires q l]_(mm!l)) \
-\       --> [][? q. Write memCh mm (resbar rmhist) q l]_(mm!l)"
-   (fn _ => [ force_tac (MI_css addSEs2 [STL4E] addSDs2 [Step2_2]) 1 ]);
+Goal "|- [](  (ALL p. ImpNext p) \
+\           & (ALL p. [HNext rmhist p]_(c p, r p, m p, rmhist!p)) \
+\           & (ALL p. $ImpInv rmhist p) \
+\           & [EX q. Write rmCh mm ires q l]_(mm!l)) \
+\        --> [][EX q. Write memCh mm (resbar rmhist) q l]_(mm!l)";
+by (force_tac (MI_css addSEs2 [STL4E] addSDs2 [Step2_2]) 1);
+qed "Step2_lemma";
 
-qed_goal "Step2" MemoryImplementation.thy
-   "|- #l : #MemLoc & (!p. IPImp p & HistP rmhist p)  \
-\      --> MSpec memCh mm (resbar rmhist) l"
-   (fn _ => [auto_tac (MI_css addsimps2 [MSpec_def]),
-	     force_tac (MI_css addsimps2 [IPImp_def,MSpec_def]) 1,
-	     auto_tac (MI_css addSIs2 [Step2_lemma]
-		              addsimps2 [split_box_conj,all_box]),
-	     force_tac (MI_css addsimps2 [IPImp_def,MSpec_def]) 4,
-             auto_tac (MI_css addsimps2 [split_box_conj] addSEs2 [allE] addSDs2 [GoodImpl])
-	    ]);
+Goal "|- #l : #MemLoc & (ALL p. IPImp p & HistP rmhist p)  \
+\        --> MSpec memCh mm (resbar rmhist) l";
+by (auto_tac (MI_css addsimps2 [MSpec_def]));
+by (force_tac (MI_css addsimps2 [IPImp_def,MSpec_def]) 1);
+by (auto_tac (MI_css addSIs2 [Step2_lemma]
+	             addsimps2 [split_box_conj,all_box]));
+by (force_tac (MI_css addsimps2 [IPImp_def,MSpec_def]) 4);
+by (auto_tac (MI_css addsimps2 [split_box_conj] addSEs2 [allE] addSDs2 [GoodImpl]));
+qed "Step2";
 
 (* ----------------------------- Main theorem --------------------------------- *)
 section "Memory implementation";
@@ -590,21 +887,17 @@
 (* Implementation of internal specification by combination of implementation
    and history variable with explicit refinement mapping
 *)
-qed_goal "Impl_IUSpec" MemoryImplementation.thy
-   "|- Implementation & Hist rmhist --> IUSpec memCh mm (resbar rmhist)"
-   (fn _ => [auto_tac (MI_css addsimps2 [IUSpec_def,Implementation_def,IPImp_def,
-					 MClkISpec_def,RPCISpec_def,IRSpec_def,Hist_def]
-		              addSIs2 [Step1,Step2])
-	    ]);
+Goal "|- Implementation & Hist rmhist --> IUSpec memCh mm (resbar rmhist)";
+by (auto_tac (MI_css addsimps2 [IUSpec_def,Implementation_def,IPImp_def,
+			        MClkISpec_def,RPCISpec_def,IRSpec_def,Hist_def]
+		     addSIs2 [Step1,Step2]));
+qed "Impl_IUSpec";
 
 (* The main theorem: introduce hiding and eliminate history variable. *)
-qed_goal "Implementation" MemoryImplementation.thy
-   "|- Implementation --> USpec memCh"
-   (fn _ => [Clarsimp_tac 1,
-             forward_tac [temp_use History] 1,
-             auto_tac (MI_css addsimps2 [USpec_def] 
-                              addIs2 [eexI, Impl_IUSpec, MI_base]
-                              addSEs2 [eexE])
-            ]);
-
-
+Goal "|- Implementation --> USpec memCh";
+by (Clarsimp_tac 1);
+by (forward_tac [temp_use History] 1);
+by (auto_tac (MI_css addsimps2 [USpec_def] 
+                     addIs2 [eexI, Impl_IUSpec, MI_base]
+                     addSEs2 [eexE]));
+qed "Implementation";
--- a/src/HOL/TLA/Memory/MemoryImplementation.thy	Thu Aug 03 19:28:37 2000 +0200
+++ b/src/HOL/TLA/Memory/MemoryImplementation.thy	Thu Aug 03 19:29:03 2000 +0200
@@ -33,9 +33,6 @@
   rst           :: "rpcStType"
   cst           :: "mClkStType"
   ires          :: "resType"
-(* the history variable : not defined as a constant
-  rmhist        :: "histType"
-*)
 
 constdefs
   (* auxiliary predicates *)
@@ -78,7 +75,7 @@
                            & [][HNext rmhist p]_(c p,r p,m p, rmhist!p)"
 
   Hist          :: "histType => temporal"
-      "Hist rmhist == TEMP (!p. HistP rmhist p)"
+      "Hist rmhist == TEMP (ALL p. HistP rmhist p)"
 
   (* the implementation *)
   IPImp          :: "PrIds => temporal"
@@ -86,7 +83,7 @@
 	               & MClkIPSpec memCh crCh cst p
   	               & RPCIPSpec crCh rmCh rst p
 	               & RPSpec rmCh mm ires p
-		       & (! l. #l : #MemLoc --> MSpec rmCh mm ires l))"
+		       & (ALL l. #l : #MemLoc --> MSpec rmCh mm ires l))"
 
   ImpInit        :: "PrIds => stpred"
       "ImpInit p == PRED (  ~Calling memCh p
@@ -108,7 +105,7 @@
 			& WF(MemReturn rmCh ires p)_(m p)"
 
   Implementation :: "temporal"
-      "Implementation == TEMP ( (!p. Init (~Calling memCh p) & [][ENext p]_(e p))
+      "Implementation == TEMP ( (ALL p. Init (~Calling memCh p) & [][ENext p]_(e p))
                                & MClkISpec memCh crCh cst
                                & RPCISpec crCh rmCh rst
                                & IRSpec rmCh mm ires)"
--- a/src/HOL/TLA/Memory/MemoryParameters.ML	Thu Aug 03 19:28:37 2000 +0200
+++ b/src/HOL/TLA/Memory/MemoryParameters.ML	Thu Aug 03 19:29:03 2000 +0200
@@ -6,24 +6,13 @@
     RPC-Memory example: memory parameters (ML file)
 *)
 
-(*
-val MP_simps = [BadArgNoMemVal,MemFailNoMemVal,InitValMemVal,NotAResultNotVal,
-                  NotAResultNotOK, NotAResultNotBA, NotAResultNotMF]
-               @ (map (fn x => x RS not_sym) 
-                      [NotAResultNotOK, NotAResultNotBA, NotAResultNotMF]);
-*)
-
 Addsimps ([BadArgNoMemVal,MemFailNoMemVal,InitValMemVal,NotAResultNotVal,
                   NotAResultNotOK, NotAResultNotBA, NotAResultNotMF]
                @ (map (fn x => x RS not_sym) 
                       [NotAResultNotOK, NotAResultNotBA, NotAResultNotMF]));
 
-(* Auxiliary rules *)
-
-qed_goal "MemValNotAResultE" MemoryParameters.thy
-   "[| x : MemVal; (x ~= NotAResult ==> P) |] ==> P"
-   (fn prems => [resolve_tac prems 1,
-                 cut_facts_tac (NotAResultNotVal::prems) 1,
-                 Force_tac 1
-                ]);
-
+val prems = goal thy "[| x : MemVal; (x ~= NotAResult ==> P) |] ==> P";
+by (resolve_tac prems 1);
+by (cut_facts_tac (NotAResultNotVal::prems) 1);
+by (Force_tac 1);
+qed "MemValNotAResultE";
--- a/src/HOL/TLA/Memory/MemoryParameters.thy	Thu Aug 03 19:28:37 2000 +0200
+++ b/src/HOL/TLA/Memory/MemoryParameters.thy	Thu Aug 03 19:29:03 2000 +0200
@@ -12,19 +12,8 @@
 MemoryParameters = Datatype + RPCMemoryParams +
 
 (* the memory operations *)
-(***
-datatype  Rd = read
-datatype  Wr = write
-***)
-
 datatype memOp = read Locs | write Locs Vals
 
-(***
-types
-  (* legal arguments for the memory *)
-  memArgType = "(Rd * Locs) + (Wr * Locs * Vals)"
-***)
-
 consts
   (* memory locations and contents *)
   MemLoc         :: Locs set
--- a/src/HOL/TLA/Memory/ProcedureInterface.ML	Thu Aug 03 19:28:37 2000 +0200
+++ b/src/HOL/TLA/Memory/ProcedureInterface.ML	Thu Aug 03 19:29:03 2000 +0200
@@ -16,45 +16,12 @@
 		      PLegalCaller_def, LegalCaller_def,
 		      PLegalReturner_def, LegalReturner_def];
 
-(* sample theorems (not used in the proof):
-   1. calls and returns are mutually exclusive
-
-qed_goal "CallNotReturn" ProcedureInterface.thy
-     "|- Call ch p v --> ~ Return ch p w"
-  (fn prems => [ auto_tac (temp_css addsimps2 [Call_def,Return_def]) ]);
-
-
-  2. enabledness of calls and returns
-
-qed_goal "Call_enabled" ProcedureInterface.thy
-   "!!p. basevars ((caller ch)!p) ==> |- ~ Calling ch p --> Enabled (Call ch p v)"
-   (fn _ => [action_simp_tac (simpset() addsimps [caller_def, Call_def]) 
-                             [] [base_enabled,Pair_inject] 1
-            ]);
+(* Calls and returns change their subchannel *)
+Goal "|- Call ch p v --> <Call ch p v>_((caller ch)!p)";
+by (auto_tac (mem_css addsimps2 [angle_def,Call_def,caller_def,Calling_def]));
+qed "Call_changed";
 
-qed_goal "Call_enabled_rew" ProcedureInterface.thy
-   "basevars ((caller ch)!p) ==> |- Enabled (Call ch p v) = (~Calling ch p)"
-   (fn [prem] => [auto_tac (mem_css addsimps2 [Call_def]),
-                  force_tac (mem_css addsimps2 [enabled_def]) 1,
-                  enabled_tac prem 1,
-                  action_simp_tac (simpset() addsimps [caller_def]) [] [Pair_inject] 1
-            ]);
+Goal "|- Return ch p v --> <Return ch p v>_((rtrner ch)!p)";
+by (auto_tac (mem_css addsimps2 [angle_def,Return_def,rtrner_def,Calling_def]));
+qed "Return_changed";
 
-qed_goal "Return_enabled" ProcedureInterface.thy
-   "!!p. basevars ((rtrner ch)!p) ==> |- Calling ch p --> Enabled (Return ch p v)"
-   (fn _ => [action_simp_tac (simpset() addsimps [rtrner_def, Return_def]) 
-                             [] [base_enabled,Pair_inject] 1
-            ]);
-
-*)
-
-(* Calls and returns change their subchannel *)
-qed_goal "Call_changed" ProcedureInterface.thy
-   "|- Call ch p v --> <Call ch p v>_((caller ch)!p)"
-   (fn _ => [ auto_tac (mem_css addsimps2 [angle_def,Call_def,caller_def,Calling_def]) ]);
-
-qed_goal "Return_changed" ProcedureInterface.thy
-   "|- Return ch p v --> <Return ch p v>_((rtrner ch)!p)"
-   (fn _ => [ auto_tac (mem_css addsimps2 [angle_def,Return_def,rtrner_def,Calling_def]) ]);
-
-
--- a/src/HOL/TLA/Memory/ProcedureInterface.thy	Thu Aug 03 19:28:37 2000 +0200
+++ b/src/HOL/TLA/Memory/ProcedureInterface.thy	Thu Aug 03 19:29:03 2000 +0200
@@ -69,11 +69,11 @@
 
   Calling_def	"Calling ch p  == PRED cbit< ch!p > ~= rbit< ch!p >"
   Call_def      "(ACT Call ch p v)   == ACT  ~ $Calling ch p
-                                     & (cbit<ch!p>` ~= $rbit<ch!p>)
-                                     & (arg<ch!p>` = $v)"
+                                     & (cbit<ch!p>$ ~= $rbit<ch!p>)
+                                     & (arg<ch!p>$ = $v)"
   Return_def    "(ACT Return ch p v) == ACT  $Calling ch p
-                                     & (rbit<ch!p>` = $cbit<ch!p>)
-                                     & (res<ch!p>` = $v)"
+                                     & (rbit<ch!p>$ = $cbit<ch!p>)
+                                     & (res<ch!p>$ = $v)"
   PLegalCaller_def      "PLegalCaller ch p == TEMP
                              Init(~ Calling ch p)
                              & [][ ? a. Call ch p a ]_((caller ch)!p)"
--- a/src/HOL/TLA/Memory/RPC.ML	Thu Aug 03 19:28:37 2000 +0200
+++ b/src/HOL/TLA/Memory/RPC.ML	Thu Aug 03 19:29:03 2000 +0200
@@ -17,40 +17,36 @@
    unanswered call for that process.
 *)
 
-qed_goal "RPCidle" RPC.thy
-   "|- ~$(Calling send p) --> ~RPCNext send rcv rst p"
-   (fn _ => [ auto_tac (mem_css addsimps2 (Return_def::RPC_action_defs)) ]);
+Goal "|- ~$(Calling send p) --> ~RPCNext send rcv rst p";
+by (auto_tac (mem_css addsimps2 (Return_def::RPC_action_defs)));
+qed "RPCidle";
 
-qed_goal "RPCbusy" RPC.thy
-   "|- $(Calling rcv p) & $(rst!p) = #rpcB --> ~RPCNext send rcv rst p"
-   (fn _ => [ auto_tac (mem_css addsimps2 RPC_action_defs) ]);
+Goal "|- $(Calling rcv p) & $(rst!p) = #rpcB --> ~RPCNext send rcv rst p";
+by (auto_tac (mem_css addsimps2 RPC_action_defs));
+qed "RPCbusy";
 
 (* RPC failure actions are visible. *)
-qed_goal "RPCFail_vis" RPC.thy
-   "|- RPCFail send rcv rst p --> \
-\      <RPCNext send rcv rst p>_(rst!p, rtrner send!p, caller rcv!p)"
-   (fn _ => [auto_tac (claset() addSDs [Return_changed],
-		       simpset() addsimps [angle_def,RPCNext_def,RPCFail_def])
-	    ]);
+Goal "|- RPCFail send rcv rst p --> \
+\        <RPCNext send rcv rst p>_(rst!p, rtrner send!p, caller rcv!p)";
+by (auto_tac (claset() addSDs [Return_changed],
+	     simpset() addsimps [angle_def,RPCNext_def,RPCFail_def]));
+qed "RPCFail_vis";
 
-qed_goal "RPCFail_Next_enabled" RPC.thy
-   "|- Enabled (RPCFail send rcv rst p) --> \
-\      Enabled (<RPCNext send rcv rst p>_(rst!p, rtrner send!p, caller rcv!p))"
-   (fn _ => [force_tac (mem_css addSEs2 [enabled_mono,RPCFail_vis]) 1]);
+Goal "|- Enabled (RPCFail send rcv rst p) --> \
+\        Enabled (<RPCNext send rcv rst p>_(rst!p, rtrner send!p, caller rcv!p))";
+by (force_tac (mem_css addSEs2 [enabled_mono,RPCFail_vis]) 1);
+qed "RPCFail_Next_enabled";
 
 (* Enabledness of some actions *)
-
-qed_goal "RPCFail_enabled" RPC.thy
-   "!!p. basevars (rtrner send!p, caller rcv!p, rst!p) ==> \
-\        |- ~Calling rcv p & Calling send p --> Enabled (RPCFail send rcv rst p)"
-   (fn _ => [action_simp_tac (simpset() addsimps [RPCFail_def,Return_def,caller_def,rtrner_def])
-                             [exI] [base_enabled,Pair_inject] 1
-	    ]);
+Goal "!!p. basevars (rtrner send!p, caller rcv!p, rst!p) ==> \
+\     |- ~Calling rcv p & Calling send p --> Enabled (RPCFail send rcv rst p)";
+by (action_simp_tac (simpset() addsimps [RPCFail_def,Return_def,caller_def,rtrner_def])
+                    [exI] [base_enabled,Pair_inject] 1);
+qed "RPCFail_enabled";
 
-qed_goal "RPCReply_enabled" RPC.thy
-   "!!p. basevars (rtrner send!p, caller rcv!p, rst!p) ==> \
-\        |- ~Calling rcv p & Calling send p & rst!p = #rpcB \
-\           --> Enabled (RPCReply send rcv rst p)"
-   (fn _ => [action_simp_tac (simpset() addsimps [RPCReply_def,Return_def,caller_def,rtrner_def])
-                             [exI] [base_enabled,Pair_inject] 1]);
-
+Goal "!!p. basevars (rtrner send!p, caller rcv!p, rst!p) ==> \
+\     |- ~Calling rcv p & Calling send p & rst!p = #rpcB \
+\        --> Enabled (RPCReply send rcv rst p)";
+by (action_simp_tac (simpset() addsimps [RPCReply_def,Return_def,caller_def,rtrner_def])
+                    [exI] [base_enabled,Pair_inject] 1);
+qed "RPCReply_enabled";
--- a/src/HOL/TLA/Memory/RPC.thy	Thu Aug 03 19:28:37 2000 +0200
+++ b/src/HOL/TLA/Memory/RPC.thy	Thu Aug 03 19:29:03 2000 +0200
@@ -72,7 +72,7 @@
                          & [][ RPCNext send rcv rst p ]_(rst!p, rtrner send!p, caller rcv!p)
                          & WF(RPCNext send rcv rst p)_(rst!p, rtrner send!p, caller rcv!p)"
 
-  RPCISpec_def      "RPCISpec send rcv rst == TEMP (! p. RPCIPSpec send rcv rst p)"
+  RPCISpec_def      "RPCISpec send rcv rst == TEMP (ALL p. RPCIPSpec send rcv rst p)"
 
 end
 
--- a/src/HOL/TLA/Memory/RPCParameters.ML	Thu Aug 03 19:28:37 2000 +0200
+++ b/src/HOL/TLA/Memory/RPCParameters.ML	Thu Aug 03 19:29:03 2000 +0200
@@ -7,11 +7,5 @@
 *)
 
 
-(*
-val RP_simps = MP_simps @ [RFNoMemVal, NotAResultNotRF, OKNotRF, BANotRF]
-                        @ (map (fn x => x RS not_sym) [NotAResultNotRF, OKNotRF, BANotRF])
-                        @ rpcState.simps @ rpcOp.simps;
-*)
-
 Addsimps ([RFNoMemVal, NotAResultNotRF, OKNotRF, BANotRF]
           @ (map (fn x => x RS not_sym) [NotAResultNotRF, OKNotRF, BANotRF]));
--- a/src/HOL/TLA/Memory/RPCParameters.thy	Thu Aug 03 19:28:37 2000 +0200
+++ b/src/HOL/TLA/Memory/RPCParameters.thy	Thu Aug 03 19:29:03 2000 +0200
@@ -16,17 +16,6 @@
 datatype  rpcOp = memcall memOp | othercall Vals
 datatype  rpcState = rpcA | rpcB
 
-(***
-types
-  (* type of RPC arguments other than memory calls *)
-  noMemArgType
-  (* legal arguments for (our instance of) the RPC component *)
-  rpcArgType = "(rpcOps * memArgType) + (rpcOps * noMemArgType)"
-
-arities
-  noMemArgType :: term
-***)
-
 consts
   (* some particular return values *)
   RPCFailure     :: Vals
@@ -36,10 +25,6 @@
      is legal for the receiver (i.e., the memory). This can now be a little
      simpler than for the generic RPC component. RelayArg returns an arbitrary
      memory call for illegal arguments. *)
-(***
-  IsLegalRcvArg  :: rpcArgType => bool
-  RPCRelayArg    :: rpcArgType => memArgType
-***)
   IsLegalRcvArg  :: rpcOp => bool
   RPCRelayArg    :: rpcOp => memOp
 
@@ -50,12 +35,6 @@
   OKNotRF           "OK ~= RPCFailure"
   BANotRF           "BadArg ~= RPCFailure"
 
-(***
-  IsLegalRcvArg_def "IsLegalRcvArg ra == EX marg. ra = Inl (remoteCall,marg)"
-  RPCRelayArg_def   "RPCRelayArg ra == 
-                         case ra of Inl (rm) => (snd rm)
-                                  | Inr (rn) => (read, @ l. True)"
-***)
 defs
   IsLegalRcvArg_def "IsLegalRcvArg ra ==
 		         case ra of (memcall m) => True
--- a/src/HOL/TLA/TLA.ML	Thu Aug 03 19:28:37 2000 +0200
+++ b/src/HOL/TLA/TLA.ML	Thu Aug 03 19:29:03 2000 +0200
@@ -8,11 +8,13 @@
 
 (* Specialize intensional introduction/elimination rules for temporal formulas *)
 
-qed_goal "tempI" TLA.thy "(!!sigma. sigma |= (F::temporal)) ==> |- F"
-  (fn [prem] => [ REPEAT (resolve_tac [prem,intI] 1) ]);
+val [prem] = goal thy "(!!sigma. sigma |= (F::temporal)) ==> |- F";
+by (REPEAT (resolve_tac [prem,intI] 1));
+qed "tempI";
 
-qed_goal "tempD" TLA.thy "|- (F::temporal) ==> sigma |= F"
-  (fn [prem] => [ rtac (prem RS intD) 1 ]);
+val [prem] = goal thy "|- (F::temporal) ==> sigma |= F";
+by (rtac (prem RS intD) 1);
+qed "tempD";
 
 
 (* ======== Functions to "unlift" temporal theorems ====== *)
@@ -75,13 +77,16 @@
 section "Simple temporal logic";
 
 (* []~F == []~Init F *)
-bind_thm("boxNotInit", rewrite_rule Init_simps (read_instantiate [("F", "LIFT ~F")] boxInit));
+bind_thm("boxNotInit", 
+         rewrite_rule Init_simps (read_instantiate [("F", "LIFT ~F")] boxInit));
 
-qed_goalw "dmdInit" TLA.thy [dmd_def] "TEMP <>F == TEMP <> Init F"
-  (fn _ => [rewtac (read_instantiate [("F", "LIFT ~F")] boxInit),
-            simp_tac (simpset() addsimps Init_simps) 1]);
+Goalw [dmd_def] "TEMP <>F == TEMP <> Init F";
+by (rewtac (read_instantiate [("F", "LIFT ~F")] boxInit));
+by (simp_tac (simpset() addsimps Init_simps) 1);
+qed "dmdInit";
 
-bind_thm("dmdNotInit", rewrite_rule Init_simps (read_instantiate [("F", "LIFT ~F")] dmdInit));
+bind_thm("dmdNotInit", 
+         rewrite_rule Init_simps (read_instantiate [("F", "LIFT ~F")] dmdInit));
 
 (* boxInit and dmdInit cannot be used as rewrites, because they loop.
    Non-looping instances for state predicates and actions are occasionally useful.
@@ -103,26 +108,30 @@
 bind_thm("STL2", reflT);
 
 (* The "polymorphic" (generic) variant *)
-qed_goal "STL2_gen" TLA.thy "|- []F --> Init F"
-  (fn _ => [rewtac (read_instantiate [("F", "F")] boxInit),
-            rtac STL2 1]);
+Goal "|- []F --> Init F";
+by (rewtac (read_instantiate [("F", "F")] boxInit));
+by (rtac STL2 1);
+qed "STL2_gen";
 
 (* see also STL2_pr below: "|- []P --> Init P & Init (P`)" *)
 
 
 (* Dual versions for <> *)
-qed_goalw "InitDmd" TLA.thy [dmd_def] "|- F --> <> F"
-   (fn _ => [ auto_tac (temp_css addSDs2 [STL2]) ]);
+Goalw [dmd_def] "|- F --> <> F";
+by (auto_tac (temp_css addSDs2 [STL2]));
+qed "InitDmd";
 
-qed_goal "InitDmd_gen" TLA.thy "|- Init F --> <>F"
-   (fn _ => [Clarsimp_tac 1,
-             dtac (temp_use InitDmd) 1,
-             asm_full_simp_tac (simpset() addsimps [dmdInitD]) 1]);
+Goal "|- Init F --> <>F";
+by (Clarsimp_tac 1);
+by (dtac (temp_use InitDmd) 1);
+by (asm_full_simp_tac (simpset() addsimps [dmdInitD]) 1);
+qed "InitDmd_gen";
 
 
 (* ------------------------ STL3 ------------------------------------------- *)
-qed_goal "STL3" TLA.thy "|- ([][]F) = ([]F)"
-   (K [force_tac (temp_css addEs2 [transT,STL2]) 1]);
+Goal "|- ([][]F) = ([]F)";
+by (force_tac (temp_css addEs2 [transT,STL2]) 1);
+qed "STL3";
 
 (* corresponding elimination rule introduces double boxes: 
    [| (sigma |= []F); (sigma |= [][]F) ==> PROP W |] ==> PROP W
@@ -131,31 +140,34 @@
 bind_thm("dup_boxD", (temp_unlift STL3) RS iffD1);
 
 (* dual versions for <> *)
-qed_goalw "DmdDmd" TLA.thy [dmd_def] "|- (<><>F) = (<>F)"
-   (fn _ => [ auto_tac (temp_css addsimps2 [STL3]) ]);
+Goal "|- (<><>F) = (<>F)";
+by (auto_tac (temp_css addsimps2 [dmd_def,STL3]));
+qed "DmdDmd";
 bind_thm("dup_dmdE", make_elim((temp_unlift DmdDmd) RS iffD2));
 bind_thm("dup_dmdD", (temp_unlift DmdDmd) RS iffD1);
 
 
 (* ------------------------ STL4 ------------------------------------------- *)
-qed_goal "STL4" TLA.thy "|- F --> G  ==> |- []F --> []G"
-   (fn [prem] => [Clarsimp_tac 1,
-		  rtac (temp_use normalT) 1,
-                  rtac (temp_use (prem RS necT)) 1,
-		  atac 1
-		 ]);
+val [prem] = goal thy "|- F --> G  ==> |- []F --> []G";
+by (Clarsimp_tac 1);
+by (rtac (temp_use normalT) 1);
+by (rtac (temp_use (prem RS necT)) 1);
+by (atac 1);
+qed "STL4";
 
 (* Unlifted version as an elimination rule *)
-qed_goal "STL4E" TLA.thy 
-         "[| sigma |= []F; |- F --> G |] ==> sigma |= []G"
-   (fn prems => [ REPEAT (resolve_tac (prems @ [temp_use STL4]) 1) ]);
+val prems = goal thy "[| sigma |= []F; |- F --> G |] ==> sigma |= []G";
+by (REPEAT (resolve_tac (prems @ [temp_use STL4]) 1));
+qed "STL4E";
 
-qed_goal "STL4_gen" TLA.thy "|- Init F --> Init G ==> |- []F --> []G"
-   (fn [prem] => [rtac (rewrite_rule [boxInitD] (prem RS STL4)) 1]);
+val [prem] = goal thy "|- Init F --> Init G ==> |- []F --> []G";
+by (rtac (rewrite_rule [boxInitD] (prem RS STL4)) 1);
+qed "STL4_gen";
 
-qed_goal "STL4E_gen" TLA.thy
-         "[| sigma |= []F; |- Init F --> Init G |] ==> sigma |= []G"
-   (fn prems => [ REPEAT (resolve_tac (prems @ [temp_use STL4_gen]) 1) ]);
+val prems = goal thy
+   "[| sigma |= []F; |- Init F --> Init G |] ==> sigma |= []G";
+by (REPEAT (resolve_tac (prems @ [temp_use STL4_gen]) 1));
+qed "STL4E_gen";
 
 (* see also STL4Edup below, which allows an auxiliary boxed formula:
        []A /\ F => G
@@ -164,22 +176,24 @@
 *)
 
 (* The dual versions for <> *)
-qed_goalw "DmdImpl" TLA.thy [dmd_def]
-   "|- F --> G ==> |- <>F --> <>G"
-   (fn [prem] => [fast_tac (temp_cs addSIs [prem] addSEs [STL4E]) 1]);
+val [prem] = goalw thy [dmd_def]
+   "|- F --> G ==> |- <>F --> <>G";
+by (fast_tac (temp_cs addSIs [prem] addSEs [STL4E]) 1);
+qed "DmdImpl";
 
-qed_goal "DmdImplE" TLA.thy
-   "[| sigma |= <>F; |- F --> G |] ==> sigma |= <>G"
-   (fn prems => [ REPEAT (resolve_tac (prems @ [temp_use DmdImpl]) 1) ]);
+val prems = goal thy "[| sigma |= <>F; |- F --> G |] ==> sigma |= <>G";
+by (REPEAT (resolve_tac (prems @ [temp_use DmdImpl]) 1));
+qed "DmdImplE";
 
 
 (* ------------------------ STL5 ------------------------------------------- *)
-qed_goal "STL5" TLA.thy "|- ([]F & []G) = ([](F & G))"
-   (fn _ => [Auto_tac,
-	     subgoal_tac "sigma |= [](G --> (F & G))" 1,
-	     etac (temp_use normalT) 1, atac 1,
-	     ALLGOALS (fast_tac (temp_cs addSEs [STL4E]))
-	    ]);
+Goal "|- ([]F & []G) = ([](F & G))";
+by Auto_tac;
+by (subgoal_tac "sigma |= [](G --> (F & G))" 1);
+by (etac (temp_use normalT) 1);
+by (ALLGOALS (fast_tac (temp_cs addSEs [STL4E])));
+qed "STL5";
+
 (* rewrite rule to split conjunctions under boxes *)
 bind_thm("split_box_conj", (temp_unlift STL5) RS sym);
 
@@ -187,10 +201,10 @@
    (NB: F and G must have the same type, i.e., both actions or temporals.)
    Use "addSE2" etc. if you want to add this to a claset, otherwise it will loop!
 *)
-qed_goal "box_conjE" TLA.thy
-   "[| sigma |= []F; sigma |= []G; sigma |= [](F&G) ==> PROP R |] ==> PROP R"
-   (fn prems => [ REPEAT (resolve_tac
-			   (prems @ [(temp_unlift STL5) RS iffD1, conjI]) 1) ]);
+val prems = goal thy
+   "[| sigma |= []F; sigma |= []G; sigma |= [](F&G) ==> PROP R |] ==> PROP R";
+by (REPEAT (resolve_tac (prems @ [(temp_unlift STL5) RS iffD1, conjI]) 1));
+qed "box_conjE";
 
 (* Instances of box_conjE for state predicates, actions, and temporals
    in case the general rule is "too polymorphic".
@@ -229,99 +243,104 @@
 bind_thm("all_box", standard((temp_unlift allT) RS sym));
 
 
-qed_goal "DmdOr" TLA.thy "|- (<>(F | G)) = (<>F | <>G)"
-   (fn _ => [auto_tac (temp_css addsimps2 [dmd_def,split_box_conj]),
-             TRYALL (EVERY' [etac swap, 
-                             merge_box_tac, 
-                             fast_tac (temp_cs addSEs [STL4E])])
-            ]);
+Goal "|- (<>(F | G)) = (<>F | <>G)";
+by (auto_tac (temp_css addsimps2 [dmd_def,split_box_conj]));
+by (ALLGOALS (EVERY' [etac swap, 
+                      merge_box_tac, 
+                      fast_tac (temp_cs addSEs [STL4E])]));
+qed "DmdOr";
 
-qed_goal "exT" TLA.thy "|- (? x. <>(F x)) = (<>(? x. F x))"
-   (fn _ => [ auto_tac (temp_css addsimps2 [dmd_def,Not_Rex,all_box]) ]);
+Goal "|- (EX x. <>(F x)) = (<>(EX x. F x))";
+by (auto_tac (temp_css addsimps2 [dmd_def,Not_Rex,all_box]));
+qed "exT";
 
 bind_thm("ex_dmd", standard((temp_unlift exT) RS sym));
 	     
 
-qed_goal "STL4Edup" TLA.thy
-   "!!sigma. [| sigma |= []A; sigma |= []F; |- F & []A --> G |] ==> sigma |= []G"
-   (fn _ => [etac dup_boxE 1,
-	     merge_box_tac 1,
-	     etac STL4E 1,
-	     atac 1
-	    ]);
+Goal "!!sigma. [| sigma |= []A; sigma |= []F; |- F & []A --> G |] ==> sigma |= []G";
+by (etac dup_boxE 1);
+by (merge_box_tac 1);
+by (etac STL4E 1);
+by (atac 1);
+qed "STL4Edup";
 
-qed_goalw "DmdImpl2" TLA.thy [dmd_def]
-   "!!sigma. [| sigma |= <>F; sigma |= [](F --> G) |] ==> sigma |= <>G"
-   (fn _ => [Auto_tac,
-	     etac notE 1,
-	     merge_box_tac 1,
-	     fast_tac (temp_cs addSEs [STL4E]) 1
-	    ]);
+Goalw [dmd_def]
+   "!!sigma. [| sigma |= <>F; sigma |= [](F --> G) |] ==> sigma |= <>G";
+by Auto_tac;
+by (etac notE 1);
+by (merge_box_tac 1);
+by (fast_tac (temp_cs addSEs [STL4E]) 1);
+qed "DmdImpl2";
 
-qed_goal "InfImpl" TLA.thy
-   "[| sigma |= []<>F; sigma |= []G; |- F & G --> H |] ==> sigma |= []<>H"
-   (fn [prem1,prem2,prem3] 
-       => [cut_facts_tac [prem1,prem2] 1,
-	   eres_inst_tac [("F","G")] dup_boxE 1,
-	   merge_box_tac 1,
-	   fast_tac (temp_cs addSEs [STL4E,DmdImpl2] addSIs [prem3]) 1
-	  ]);
+val [prem1,prem2,prem3] = goal thy
+  "[| sigma |= []<>F; sigma |= []G; |- F & G --> H |] ==> sigma |= []<>H";
+by (cut_facts_tac [prem1,prem2] 1);
+by (eres_inst_tac [("F","G")] dup_boxE 1);
+by (merge_box_tac 1);
+by (fast_tac (temp_cs addSEs [STL4E,DmdImpl2] addSIs [prem3]) 1);
+qed "InfImpl";
 
 (* ------------------------ STL6 ------------------------------------------- *)
 (* Used in the proof of STL6, but useful in itself. *)
-qed_goalw "BoxDmd" TLA.thy [dmd_def] "|- []F & <>G --> <>([]F & G)"
-  (fn _ => [ Clarsimp_tac 1,
-             etac dup_boxE 1,
-	     merge_box_tac 1,
-             etac swap 1,
-             fast_tac (temp_cs addSEs [STL4E]) 1 ]);
+Goalw [dmd_def] "|- []F & <>G --> <>([]F & G)";
+by (Clarsimp_tac 1);
+by (etac dup_boxE 1);
+by (merge_box_tac 1);
+by (etac swap 1);
+by (fast_tac (temp_cs addSEs [STL4E]) 1);
+qed "BoxDmd";
 
 (* weaker than BoxDmd, but more polymorphic (and often just right) *)
-qed_goalw "BoxDmd_simple" TLA.thy [dmd_def] "|- []F & <>G --> <>(F & G)"
-  (fn _ => [ Clarsimp_tac 1,
-	     merge_box_tac 1,
-             fast_tac (temp_cs addSEs [notE,STL4E]) 1
-	   ]);
+Goalw [dmd_def] "|- []F & <>G --> <>(F & G)";
+by (Clarsimp_tac 1);
+by (merge_box_tac 1);
+by (fast_tac (temp_cs addSEs [notE,STL4E]) 1);
+qed "BoxDmd_simple";
 
-qed_goalw "BoxDmd2_simple" TLA.thy [dmd_def] "|- []F & <>G --> <>(G & F)"
-  (fn _ => [ Clarsimp_tac 1,
-	     merge_box_tac 1,
-             fast_tac (temp_cs addSEs [notE,STL4E]) 1
-	   ]);
+Goalw [dmd_def] "|- []F & <>G --> <>(G & F)";
+by (Clarsimp_tac 1);
+by (merge_box_tac 1);
+by (fast_tac (temp_cs addSEs [notE,STL4E]) 1);
+qed "BoxDmd2_simple";
 
-qed_goal "DmdImpldup" TLA.thy 
-   "[| sigma |= []A; sigma |= <>F; |- []A & F --> G |] ==> sigma |= <>G"
-   (fn [p1,p2,p3] => [rtac ((p2 RS (p1 RS (temp_use BoxDmd))) RS DmdImplE) 1,
-                      rtac p3 1]);
+val [p1,p2,p3] = goal thy
+   "[| sigma |= []A; sigma |= <>F; |- []A & F --> G |] ==> sigma |= <>G";
+by (rtac ((p2 RS (p1 RS (temp_use BoxDmd))) RS DmdImplE) 1);
+by (rtac p3 1);
+qed "DmdImpldup";
 
-qed_goal "STL6" TLA.thy "|- <>[]F & <>[]G --> <>[](F & G)"
-  (fn _ => [auto_tac (temp_css addsimps2 [symmetric (temp_rewrite STL5)]),
-	    dtac (temp_use linT) 1, atac 1, etac thin_rl 1,
-	    rtac ((temp_unlift DmdDmd) RS iffD1) 1,
-	    etac disjE 1,
-	    etac DmdImplE 1, rtac BoxDmd 1,
-	    (* the second subgoal needs commutativity of &, which complicates the proof *)
-	    etac DmdImplE 1,
-	    Auto_tac,
-	    dtac (temp_use BoxDmd) 1, atac 1, etac thin_rl 1,
-	    fast_tac (temp_cs addSEs [DmdImplE]) 1
-	   ]);
+Goal "|- <>[]F & <>[]G --> <>[](F & G)";
+by (auto_tac (temp_css addsimps2 [symmetric (temp_rewrite STL5)]));
+by (dtac (temp_use linT) 1);
+by (atac 1); 
+by (etac thin_rl 1);
+by (rtac ((temp_unlift DmdDmd) RS iffD1) 1);
+by (etac disjE 1);
+by (etac DmdImplE 1);
+by (rtac BoxDmd 1);
+by (etac DmdImplE 1);
+by Auto_tac;
+by (dtac (temp_use BoxDmd) 1); 
+by (atac 1); 
+by (etac thin_rl 1);
+by (fast_tac (temp_cs addSEs [DmdImplE]) 1);
+qed "STL6";
 
 
 (* ------------------------ True / False ----------------------------------------- *)
 section "Simplification of constants";
 
-qed_goal "BoxConst" TLA.thy "|- ([]#P) = #P"
-  (fn _ => [rtac tempI 1,
-            case_tac "P" 1,
-            auto_tac (temp_css addSIs2 [necT] addDs2 [STL2_gen] 
-                               addsimps2 Init_simps)
-           ]);
+Goal "|- ([]#P) = #P";
+by (rtac tempI 1);
+by (case_tac "P" 1);
+by (auto_tac (temp_css addSIs2 [necT] addDs2 [STL2_gen] 
+                       addsimps2 Init_simps));
+qed "BoxConst";
 
-qed_goalw "DmdConst" TLA.thy [dmd_def] "|- (<>#P) = #P"
-  (fn _ => [case_tac "P" 1,
-            ALLGOALS (asm_full_simp_tac (simpset() addsimps [BoxConst]))
-           ]);
+Goalw [dmd_def] "|- (<>#P) = #P";
+by (case_tac "P" 1);
+by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [BoxConst])));
+qed "DmdConst";
 
 val temp_simps = map temp_rewrite [BoxConst, DmdConst];
 
@@ -334,11 +353,13 @@
 (* ------------------------ Further rewrites ----------------------------------------- *)
 section "Further rewrites";
 
-qed_goalw "NotBox" TLA.thy [dmd_def] "|- (~[]F) = (<>~F)"
-   (fn _ => [ Simp_tac 1 ]);
+Goalw [dmd_def] "|- (~[]F) = (<>~F)";
+by (Simp_tac 1);
+qed "NotBox";
 
-qed_goalw "NotDmd" TLA.thy [dmd_def] "|- (~<>F) = ([]~F)"
-   (fn _ => [ Simp_tac 1 ]);
+Goalw [dmd_def] "|- (~<>F) = ([]~F)";
+by (Simp_tac 1);
+qed "NotDmd";
 
 (* These are not by default included in temp_css, because they could be harmful,
    e.g. []F & ~[]F becomes []F & <>~F !! *)
@@ -346,48 +367,51 @@
                        @ (map (fn th => (temp_unlift th) RS eq_reflection)
 		         [NotBox, NotDmd]);
 
-qed_goal "BoxDmdBox" TLA.thy "|- ([]<>[]F) = (<>[]F)"
-   (fn _ => [ auto_tac (temp_css addSDs2 [STL2]),
-              rtac ccontr 1,
-              subgoal_tac "sigma |= <>[][]F & <>[]~[]F" 1,
-              etac thin_rl 1,
-              Auto_tac,
-	      dtac (temp_use STL6) 1, atac 1,
-	      Asm_full_simp_tac 1,
-	      ALLGOALS (asm_full_simp_tac (simpset() addsimps more_temp_simps))
-	    ]);
+Goal "|- ([]<>[]F) = (<>[]F)";
+by (auto_tac (temp_css addSDs2 [STL2]));
+by (rtac ccontr 1);
+by (subgoal_tac "sigma |= <>[][]F & <>[]~[]F" 1);
+by (etac thin_rl 1);
+by Auto_tac;
+by (dtac (temp_use STL6) 1); 
+by (atac 1);
+by (Asm_full_simp_tac 1);
+by (ALLGOALS (asm_full_simp_tac (simpset() addsimps more_temp_simps)));
+qed "BoxDmdBox";
 
-qed_goalw "DmdBoxDmd" TLA.thy [dmd_def] "|- (<>[]<>F) = ([]<>F)"
-  (fn _ => [ auto_tac (temp_css addsimps2 [rewrite_rule [dmd_def] BoxDmdBox]) ]);
+Goalw [dmd_def] "|- (<>[]<>F) = ([]<>F)";
+by (auto_tac (temp_css addsimps2 [rewrite_rule [dmd_def] BoxDmdBox]));
+qed "DmdBoxDmd";
 
 val more_temp_simps = more_temp_simps @ (map temp_rewrite [BoxDmdBox, DmdBoxDmd]);
 
 
 (* ------------------------ Miscellaneous ----------------------------------- *)
 
-qed_goal "BoxOr" TLA.thy 
-   "!!sigma. [| sigma |= []F | []G |] ==> sigma |= [](F | G)"
-   (fn _ => [ fast_tac (temp_cs addSEs [STL4E]) 1 ]);
+Goal "!!sigma. [| sigma |= []F | []G |] ==> sigma |= [](F | G)";
+by (fast_tac (temp_cs addSEs [STL4E]) 1);
+qed "BoxOr";
 
 (* "persistently implies infinitely often" *)
-qed_goal "DBImplBD" TLA.thy "|- <>[]F --> []<>F"
-  (fn _ => [Clarsimp_tac 1,
-	    rtac ccontr 1,
-            asm_full_simp_tac (simpset() addsimps more_temp_simps) 1,
-            dtac (temp_use STL6) 1, atac 1,
-            Asm_full_simp_tac 1
-	   ]);
+Goal "|- <>[]F --> []<>F";
+by (Clarsimp_tac 1);
+by (rtac ccontr 1);
+by (asm_full_simp_tac (simpset() addsimps more_temp_simps) 1);
+by (dtac (temp_use STL6) 1); 
+by (atac 1);
+by (Asm_full_simp_tac 1);
+qed "DBImplBD";
 
-qed_goal "BoxDmdDmdBox" TLA.thy
-   "|- []<>F & <>[]G --> []<>(F & G)"
-   (fn _ => [Clarsimp_tac 1,
-             rtac ccontr 1,
-	     rewrite_goals_tac more_temp_simps,
-	     dtac (temp_use STL6) 1, atac 1,
-	     subgoal_tac "sigma |= <>[]~F" 1,
-	     force_tac (temp_css addsimps2 [dmd_def]) 1,
-	     fast_tac (temp_cs addEs [DmdImplE,STL4E]) 1
-	    ]);
+Goal "|- []<>F & <>[]G --> []<>(F & G)";
+by (Clarsimp_tac 1);
+by (rtac ccontr 1);
+by (rewrite_goals_tac more_temp_simps);
+by (dtac (temp_use STL6) 1);
+by (atac 1);
+by (subgoal_tac "sigma |= <>[]~F" 1);
+ by (force_tac (temp_css addsimps2 [dmd_def]) 1);
+by (fast_tac (temp_cs addEs [DmdImplE,STL4E]) 1);
+qed "BoxDmdDmdBox";
 
 
 (* ------------------------------------------------------------------------- *)
@@ -396,79 +420,85 @@
 section "priming";
 
 (* ------------------------ TLA2 ------------------------------------------- *)
-qed_goal "STL2_pr" TLA.thy
-  "|- []P --> Init P & Init P`"
-  (fn _ => [fast_tac (temp_cs addSIs [primeI, STL2_gen]) 1]);
+Goal "|- []P --> Init P & Init P`";
+by (fast_tac (temp_cs addSIs [primeI, STL2_gen]) 1);
+qed "STL2_pr";
 
 (* Auxiliary lemma allows priming of boxed actions *)
-qed_goal "BoxPrime" TLA.thy "|- []P --> []($P & P$)"
-  (fn _ => [Clarsimp_tac 1,
-	    etac dup_boxE 1,
-            rewtac boxInit_act,
-            etac STL4E 1,
-	    auto_tac (temp_css addsimps2 Init_simps addSDs2 [STL2_pr])
-	   ]);
+Goal "|- []P --> []($P & P$)";
+by (Clarsimp_tac 1);
+by (etac dup_boxE 1);
+by (rewtac boxInit_act);
+by (etac STL4E 1);
+by (auto_tac (temp_css addsimps2 Init_simps addSDs2 [STL2_pr]));
+qed "BoxPrime";
 
-qed_goal "TLA2" TLA.thy "|- $P & P$ --> A  ==>  |- []P --> []A"
-  (fn prems => [Clarsimp_tac 1,
-                dtac (temp_use BoxPrime) 1,
-                auto_tac (temp_css addsimps2 [Init_stp_act_rev] addSIs2 prems addSEs2 [STL4E])
-               ]);
+val prems = goal thy "|- $P & P$ --> A  ==>  |- []P --> []A";
+by (Clarsimp_tac 1);
+by (dtac (temp_use BoxPrime) 1);
+by (auto_tac (temp_css addsimps2 [Init_stp_act_rev] 
+                       addSIs2 prems addSEs2 [STL4E]));
+qed "TLA2";
 
-qed_goal "TLA2E" TLA.thy 
-   "[| sigma |= []P; |- $P & P$ --> A |] ==> sigma |= []A"
-   (fn prems => [REPEAT (resolve_tac (prems @ (prems RL [temp_use TLA2])) 1)]);
+val prems = goal thy 
+  "[| sigma |= []P; |- $P & P$ --> A |] ==> sigma |= []A";
+by (REPEAT (resolve_tac (prems @ (prems RL [temp_use TLA2])) 1));
+qed "TLA2E";
 
-qed_goalw "DmdPrime" TLA.thy [dmd_def] "|- (<>P`) --> (<>P)"
-   (fn _ => [ fast_tac (temp_cs addSEs [TLA2E]) 1 ]);
+Goalw [dmd_def] "|- (<>P`) --> (<>P)";
+by (fast_tac (temp_cs addSEs [TLA2E]) 1);
+qed "DmdPrime";
 
 bind_thm("PrimeDmd", (temp_use InitDmd_gen) RS (temp_use DmdPrime));
 
 (* ------------------------ INV1, stable --------------------------------------- *)
 section "stable, invariant";
 
-qed_goal "ind_rule" TLA.thy
+val prems = goal thy
    "[| sigma |= []H; sigma |= Init P; |- H --> (Init P & ~[]F --> Init(P`) & F) |] \
-\   ==> sigma |= []F"
-   (fn prems => [rtac (temp_use indT) 1,
-		 REPEAT (resolve_tac (prems @ (prems RL [STL4E])) 1)]);
+\   ==> sigma |= []F";
+by (rtac (temp_use indT) 1);
+by (REPEAT (resolve_tac (prems @ (prems RL [STL4E])) 1));
+qed "ind_rule";
 
-qed_goalw "box_stp_act" TLA.thy [boxInit_act] "|- ([]$P) = ([]P)"
-  (K [simp_tac (simpset() addsimps Init_simps) 1]);
+Goalw [boxInit_act] "|- ([]$P) = ([]P)";
+by (simp_tac (simpset() addsimps Init_simps) 1);
+qed "box_stp_act";
 bind_thm("box_stp_actI", zero_var_indexes ((temp_use box_stp_act) RS iffD2));
 bind_thm("box_stp_actD", zero_var_indexes ((temp_use box_stp_act) RS iffD1));
 
 val more_temp_simps = (temp_rewrite box_stp_act)::more_temp_simps;
 
-qed_goalw "INV1" TLA.thy [stable_def,boxInit_stp,boxInit_act] 
-  "|- (Init P) --> (stable P) --> []P"
-  (K [Clarsimp_tac 1,
-      etac ind_rule 1,
-      auto_tac (temp_css addsimps2 Init_simps addEs2 [ind_rule])
-     ]);
+Goalw [stable_def,boxInit_stp,boxInit_act] 
+  "|- (Init P) --> (stable P) --> []P";
+by (Clarsimp_tac 1);
+by (etac ind_rule 1);
+by (auto_tac (temp_css addsimps2 Init_simps addEs2 [ind_rule]));
+qed "INV1";
 
-qed_goalw "StableT" TLA.thy [stable_def]
-   "|- $P & A --> P` ==> |- []A --> stable P"
-   (fn [prem] => [fast_tac (temp_cs addSEs [STL4E] addIs [prem]) 1]);
+Goalw [stable_def]
+   "!!P. |- $P & A --> P` ==> |- []A --> stable P";
+by (fast_tac (temp_cs addSEs [STL4E]) 1);
+qed "StableT";
 
-qed_goal "Stable" TLA.thy
-   "[| sigma |= []A; |- $P & A --> P` |] ==> sigma |= stable P"
-   (fn prems => [ REPEAT (resolve_tac (prems @ [temp_use StableT]) 1) ]);
+val prems = goal thy
+   "[| sigma |= []A; |- $P & A --> P` |] ==> sigma |= stable P";
+by (REPEAT (resolve_tac (prems @ [temp_use StableT]) 1));
+qed "Stable";
 
 (* Generalization of INV1 *)
-qed_goalw "StableBox" TLA.thy [stable_def]
-   "|- (stable P) --> [](Init P --> []P)"
-   (K [Clarsimp_tac 1,
-       etac dup_boxE 1,
-       force_tac (temp_css addsimps2 [stable_def] addEs2 [STL4E, INV1]) 1]);
+Goalw [stable_def] "|- (stable P) --> [](Init P --> []P)";
+by (Clarsimp_tac 1);
+by (etac dup_boxE 1);
+by (force_tac (temp_css addsimps2 [stable_def] addEs2 [STL4E, INV1]) 1);
+qed "StableBox";
 
-qed_goal "DmdStable" TLA.thy 
-   "|- (stable P) & <>P --> <>[]P"
-   (fn _ => [Clarsimp_tac 1,
-             rtac DmdImpl2 1,
-	     etac (temp_use StableBox) 2,
-	     asm_simp_tac (simpset() addsimps [dmdInitD]) 1
-	    ]);
+Goal "|- (stable P) & <>P --> <>[]P";
+by (Clarsimp_tac 1);
+by (rtac DmdImpl2 1);
+by (etac (temp_use StableBox) 2);
+by (asm_simp_tac (simpset() addsimps [dmdInitD]) 1);
+qed "DmdStable";
 
 (* ---------------- (Semi-)automatic invariant tactics ---------------------- *)
 
@@ -491,84 +521,89 @@
      (TRYALL (action_simp_tac (ss addsimps [Init_stp,Init_act]) [] [squareE])));
 
 
-qed_goalw "unless" TLA.thy [dmd_def]
-   "|- []($P --> P` | Q`) --> (stable P) | <>Q"
-   (fn _ => [clarsimp_tac (temp_css addSDs2 [BoxPrime]) 1,
-	     merge_box_tac 1,
-             etac swap 1,
-	     fast_tac (temp_cs addSEs [Stable]) 1
-	    ]);
+Goalw [dmd_def] "|- []($P --> P` | Q`) --> (stable P) | <>Q";
+by (clarsimp_tac (temp_css addSDs2 [BoxPrime]) 1);
+by (merge_box_tac 1);
+by (etac swap 1);
+by (fast_tac (temp_cs addSEs [Stable]) 1);
+qed "unless";
 
 
 (* --------------------- Recursive expansions --------------------------------------- *)
 section "recursive expansions";
 
 (* Recursive expansions of [] and <> for state predicates *)
-qed_goal "BoxRec" TLA.thy "|- ([]P) = (Init P & []P`)"
-   (fn _ => [auto_tac (temp_css addSIs2 [STL2_gen]),
-	     fast_tac (temp_cs addSEs [TLA2E]) 1,
-	     auto_tac (temp_css addsimps2 [stable_def] addSEs2 [INV1,STL4E])
-	    ]);
+Goal "|- ([]P) = (Init P & []P`)";
+by (auto_tac (temp_css addSIs2 [STL2_gen]));
+by (fast_tac (temp_cs addSEs [TLA2E]) 1);
+by (auto_tac (temp_css addsimps2 [stable_def] addSEs2 [INV1,STL4E]));
+qed "BoxRec";
 
-qed_goalw "DmdRec" TLA.thy [dmd_def, temp_rewrite BoxRec] "|- (<>P) = (Init P | <>P`)" 
-  (K [ auto_tac (temp_css addsimps2 Init_simps) ]);
+Goalw [dmd_def, temp_rewrite BoxRec] "|- (<>P) = (Init P | <>P`)";
+by (auto_tac (temp_css addsimps2 Init_simps));
+qed "DmdRec";
 
-qed_goal "DmdRec2" TLA.thy
- "!!sigma. [| sigma |= <>P; sigma |= []~P` |] ==> sigma |= Init P"
-   (K [ force_tac (temp_css addsimps2 [DmdRec,dmd_def]) 1]);
+Goal "!!sigma. [| sigma |= <>P; sigma |= []~P` |] ==> sigma |= Init P";
+by (force_tac (temp_css addsimps2 [DmdRec,dmd_def]) 1);
+qed "DmdRec2";
 
-(* The "-->" part of the following is a little intricate. *)
-qed_goal "InfinitePrime" TLA.thy "|- ([]<>P) = ([]<>P`)"
-   (fn _ => [Auto_tac,
-	     rtac classical 1,
-	     rtac (temp_use DBImplBD) 1,
-	     subgoal_tac "sigma |= <>[]P" 1,
-	     fast_tac (temp_cs addSEs [DmdImplE,TLA2E]) 1,
-	     subgoal_tac "sigma |= <>[](<>P & []~P`)" 1,
-	     force_tac (temp_css addsimps2 [boxInit_stp]
-			             addSEs2 [DmdImplE,STL4E,DmdRec2]) 1,
-	     force_tac (temp_css addSIs2 [STL6] addsimps2 more_temp_simps) 1,
-	     fast_tac (temp_cs addIs [DmdPrime] addSEs [STL4E]) 1
-	    ]);
+Goal "|- ([]<>P) = ([]<>P`)";
+by Auto_tac;
+by (rtac classical 1);
+by (rtac (temp_use DBImplBD) 1);
+by (subgoal_tac "sigma |= <>[]P" 1);
+ by (fast_tac (temp_cs addSEs [DmdImplE,TLA2E]) 1);
+ by (subgoal_tac "sigma |= <>[](<>P & []~P`)" 1);
+  by (force_tac (temp_css addsimps2 [boxInit_stp]
+                          addSEs2 [DmdImplE,STL4E,DmdRec2]) 1);
+ by (force_tac (temp_css addSIs2 [STL6] addsimps2 more_temp_simps) 1);
+by (fast_tac (temp_cs addIs [DmdPrime] addSEs [STL4E]) 1);
+qed "InfinitePrime";
 
-qed_goal "InfiniteEnsures" TLA.thy
-   "[| sigma |= []N; sigma |= []<>A; |- A & N --> P` |] ==> sigma |= []<>P"
-   (fn prems => [rewtac (temp_rewrite InfinitePrime),
-                 rtac InfImpl 1,
-                 REPEAT (resolve_tac prems 1)
-                ]);
+val prems = goalw thy [temp_rewrite InfinitePrime]
+  "[| sigma |= []N; sigma |= []<>A; |- A & N --> P` |] ==> sigma |= []<>P";
+by (rtac InfImpl 1);
+by (REPEAT (resolve_tac prems 1));
+qed "InfiniteEnsures";
 
 (* ------------------------ fairness ------------------------------------------- *)
 section "fairness";
 
 (* alternative definitions of fairness *)
-qed_goalw "WF_alt" TLA.thy [WF_def,dmd_def] 
-   "|- WF(A)_v = ([]<>~Enabled(<A>_v) | []<><A>_v)"
-   (fn _ => [ fast_tac temp_cs 1 ]);
+Goalw [WF_def,dmd_def] 
+  "|- WF(A)_v = ([]<>~Enabled(<A>_v) | []<><A>_v)";
+by (fast_tac temp_cs 1);
+qed "WF_alt";
 
-qed_goalw "SF_alt" TLA.thy [SF_def,dmd_def]
-   "|- SF(A)_v = (<>[]~Enabled(<A>_v) | []<><A>_v)"
-   (fn _ => [ fast_tac temp_cs 1 ]);
+Goalw [SF_def,dmd_def]
+  "|- SF(A)_v = (<>[]~Enabled(<A>_v) | []<><A>_v)";
+by (fast_tac temp_cs 1);
+qed "SF_alt";
 
 (* theorems to "box" fairness conditions *)
-qed_goal "BoxWFI" TLA.thy "|- WF(A)_v --> []WF(A)_v"
-   (fn _ => [ auto_tac (temp_css addsimps2 (WF_alt::more_temp_simps) 
-                                 addSIs2 [BoxOr]) ]);
+Goal "|- WF(A)_v --> []WF(A)_v";
+by (auto_tac (temp_css addsimps2 (WF_alt::more_temp_simps) 
+                       addSIs2 [BoxOr]));
+qed "BoxWFI";
 
-qed_goal "WF_Box" TLA.thy "|- ([]WF(A)_v) = WF(A)_v"
-  (fn prems => [ fast_tac (temp_cs addSIs [BoxWFI] addSDs [STL2]) 1 ]);
+Goal "|- ([]WF(A)_v) = WF(A)_v";
+by (fast_tac (temp_cs addSIs [BoxWFI] addSDs [STL2]) 1);
+qed "WF_Box";
 
-qed_goal "BoxSFI" TLA.thy "|- SF(A)_v --> []SF(A)_v"
-   (fn _ => [ auto_tac (temp_css addsimps2 (SF_alt::more_temp_simps) 
-                                 addSIs2 [BoxOr]) ]);
+Goal "|- SF(A)_v --> []SF(A)_v";
+by (auto_tac (temp_css addsimps2 (SF_alt::more_temp_simps) 
+                       addSIs2 [BoxOr]));
+qed "BoxSFI";
 
-qed_goal "SF_Box" TLA.thy "|- ([]SF(A)_v) = SF(A)_v"
-  (fn prems => [ fast_tac (temp_cs addSIs [BoxSFI] addSDs [STL2]) 1 ]);
+Goal "|- ([]SF(A)_v) = SF(A)_v";
+by (fast_tac (temp_cs addSIs [BoxSFI] addSDs [STL2]) 1);
+qed "SF_Box";
 
 val more_temp_simps = more_temp_simps @ (map temp_rewrite [WF_Box, SF_Box]);
 
-qed_goalw "SFImplWF" TLA.thy [SF_def,WF_def] "|- SF(A)_v --> WF(A)_v"
-  (fn _ => [ fast_tac (temp_cs addSDs [DBImplBD]) 1 ]);
+Goalw [SF_def,WF_def] "|- SF(A)_v --> WF(A)_v";
+by (fast_tac (temp_cs addSDs [DBImplBD]) 1);
+qed "SFImplWF";
 
 (* A tactic that "boxes" all fairness conditions. Apply more_temp_simps to "unbox". *)
 val box_fair_tac = SELECT_GOAL (REPEAT (dresolve_tac [BoxWFI,BoxSFI] 1));
@@ -578,354 +613,359 @@
 
 section "~>";
 
-qed_goalw "leadsto_init" TLA.thy [leadsto_def]
-   "|- (Init F) & (F ~> G) --> <>G"
-   (fn _ => [ auto_tac (temp_css addSDs2 [STL2]) ]);
+Goalw  [leadsto_def] "|- (Init F) & (F ~> G) --> <>G";
+by (auto_tac (temp_css addSDs2 [STL2]));
+qed "leadsto_init";
 
 (* |- F & (F ~> G) --> <>G *)
 bind_thm("leadsto_init_temp", 
          rewrite_rule Init_simps (read_instantiate [("'a","behavior")] leadsto_init));
 
-qed_goalw "streett_leadsto" TLA.thy [leadsto_def]
-   "|- ([]<>Init F --> []<>G) = (<>(F ~> G))" (K [
-             Auto_tac,
-             asm_full_simp_tac (simpset() addsimps more_temp_simps) 1,
-             fast_tac (temp_cs addSEs [DmdImplE,STL4E]) 1,
-             fast_tac (temp_cs addSIs [InitDmd] addSEs [STL4E]) 1,
-             subgoal_tac "sigma |= []<><>G" 1,
-             asm_full_simp_tac (simpset() addsimps more_temp_simps) 1,
-             dtac (temp_use BoxDmdDmdBox) 1, atac 1,
-             fast_tac (temp_cs addSEs [DmdImplE,STL4E]) 1
-            ]);
+Goalw [leadsto_def] "|- ([]<>Init F --> []<>G) = (<>(F ~> G))";
+by Auto_tac;
+by (asm_full_simp_tac (simpset() addsimps more_temp_simps) 1);
+by (fast_tac (temp_cs addSEs [DmdImplE,STL4E]) 1);
+by (fast_tac (temp_cs addSIs [InitDmd] addSEs [STL4E]) 1);
+by (subgoal_tac "sigma |= []<><>G" 1);
+by (asm_full_simp_tac (simpset() addsimps more_temp_simps) 1);
+by (dtac (temp_use BoxDmdDmdBox) 1); 
+by (atac 1);
+by (fast_tac (temp_cs addSEs [DmdImplE,STL4E]) 1);
+qed "streett_leadsto";
 
-qed_goal "leadsto_infinite" TLA.thy
-   "|- []<>F & (F ~> G) --> []<>G"
-   (fn _ => [Clarsimp_tac 1,
-             etac ((temp_use InitDmd) RS 
-                   ((temp_unlift streett_leadsto) RS iffD2 RS mp)) 1,
-             asm_simp_tac (simpset() addsimps [dmdInitD]) 1
-            ]);
+Goal "|- []<>F & (F ~> G) --> []<>G";
+by (Clarsimp_tac 1);
+by (etac ((temp_use InitDmd) RS 
+          ((temp_unlift streett_leadsto) RS iffD2 RS mp)) 1);
+by (asm_simp_tac (simpset() addsimps [dmdInitD]) 1);
+qed "leadsto_infinite";
 
 (* In particular, strong fairness is a Streett condition. The following
    rules are sometimes easier to use than WF2 or SF2 below.
 *)
-qed_goalw "leadsto_SF" TLA.thy [SF_def]
-  "|- (Enabled(<A>_v) ~> <A>_v) --> SF(A)_v"
-  (K [clarsimp_tac (temp_css addSEs2 [leadsto_infinite]) 1]);
+Goalw [SF_def] "|- (Enabled(<A>_v) ~> <A>_v) --> SF(A)_v";
+by (clarsimp_tac (temp_css addSEs2 [leadsto_infinite]) 1);
+qed "leadsto_SF";
 
-qed_goal "leadsto_WF" TLA.thy 
-  "|- (Enabled(<A>_v) ~> <A>_v) --> WF(A)_v"
-  (K [ clarsimp_tac (temp_css addSIs2 [SFImplWF, leadsto_SF]) 1 ]);
+Goal "|- (Enabled(<A>_v) ~> <A>_v) --> WF(A)_v";
+by (clarsimp_tac (temp_css addSIs2 [SFImplWF, leadsto_SF]) 1);
+qed "leadsto_WF";
 
 (* introduce an invariant into the proof of a leadsto assertion.
    []I --> ((P ~> Q)  =  (P /\ I ~> Q))
 *)
-qed_goalw "INV_leadsto" TLA.thy [leadsto_def]
-   "|- []I & (P & I ~> Q) --> (P ~> Q)"
-   (fn _ => [Clarsimp_tac 1,
-             etac STL4Edup 1, atac 1,
-	     auto_tac (temp_css addsimps2 Init_simps addSDs2 [STL2_gen])
-	    ]);
+Goalw [leadsto_def] "|- []I & (P & I ~> Q) --> (P ~> Q)";
+by (Clarsimp_tac 1);
+by (etac STL4Edup 1); 
+by (atac 1);
+by (auto_tac (temp_css addsimps2 Init_simps addSDs2 [STL2_gen]));
+qed "INV_leadsto";
 
-qed_goalw "leadsto_classical" TLA.thy [leadsto_def,dmd_def]
-   "|- (Init F & []~G ~> G) --> (F ~> G)"
-   (fn _ => [force_tac (temp_css addsimps2 Init_simps addSEs2 [STL4E]) 1]);
+Goalw [leadsto_def,dmd_def]
+  "|- (Init F & []~G ~> G) --> (F ~> G)";
+by (force_tac (temp_css addsimps2 Init_simps addSEs2 [STL4E]) 1);
+qed "leadsto_classical";
 
-qed_goalw "leadsto_false" TLA.thy [leadsto_def]
-  "|- (F ~> #False) = ([]~F)"
-  (fn _ => [ simp_tac (simpset() addsimps [boxNotInitD]) 1 ]);
+Goalw [leadsto_def] "|- (F ~> #False) = ([]~F)";
+by (simp_tac (simpset() addsimps [boxNotInitD]) 1);
+qed "leadsto_false";
 
-qed_goalw "leadsto_exists" TLA.thy [leadsto_def]
-  "|- ((? x. F x) ~> G) = (!x. (F x ~> G))"
-  (K [auto_tac (temp_css addsimps2 allT::Init_simps addSEs2 [STL4E])]);
-
+Goalw [leadsto_def] "|- ((EX x. F x) ~> G) = (ALL x. (F x ~> G))";
+by (auto_tac (temp_css addsimps2 allT::Init_simps addSEs2 [STL4E]));
+qed "leadsto_exists";
 
 (* basic leadsto properties, cf. Unity *)
 
-qed_goalw "ImplLeadsto_gen" TLA.thy [leadsto_def]
-   "|- [](Init F --> Init G) --> (F ~> G)"
-   (fn _ => [auto_tac (temp_css addSIs2 [InitDmd_gen] 
-                                addSEs2 [STL4E_gen] addsimps2 Init_simps)
-	    ]);
+Goalw [leadsto_def] "|- [](Init F --> Init G) --> (F ~> G)";
+by (auto_tac (temp_css addSIs2 [InitDmd_gen] addSEs2 [STL4E_gen]
+                       addsimps2 Init_simps));
+qed "ImplLeadsto_gen";
 
 bind_thm("ImplLeadsto",
          rewrite_rule Init_simps 
              (read_instantiate [("'a","behavior"), ("'b","behavior")] ImplLeadsto_gen));
 
-qed_goal "ImplLeadsto_simple" TLA.thy
-  "|- F --> G ==> |- F ~> G"
-  (fn [prem] => [auto_tac (temp_css addsimps2 [Init_def] 
-                                    addSIs2 [ImplLeadsto_gen,necT,prem])]);
+Goal "!!F G. |- F --> G ==> |- F ~> G";
+by (auto_tac (temp_css addsimps2 [Init_def] 
+                       addSIs2 [ImplLeadsto_gen,necT]));
+qed "ImplLeadsto_simple";
 
-qed_goalw "EnsuresLeadsto" TLA.thy [leadsto_def]
-   "|- A & $P --> Q` ==> |- []A --> (P ~> Q)" (fn [prem] => [
-		  clarsimp_tac (temp_css addSEs2 [INV_leadsto]) 1, 
-                  etac STL4E_gen 1,
-                  auto_tac (temp_css addsimps2 Init_defs
-                                     addSIs2 [PrimeDmd, prem])
-                 ]);
+val [prem] = goalw thy [leadsto_def]
+  "|- A & $P --> Q` ==> |- []A --> (P ~> Q)";
+by (clarsimp_tac (temp_css addSEs2 [INV_leadsto]) 1); 
+by (etac STL4E_gen 1);
+by (auto_tac (temp_css addsimps2 Init_defs addSIs2 [PrimeDmd,prem]));
+qed "EnsuresLeadsto";
 
-qed_goalw "EnsuresLeadsto2" TLA.thy [leadsto_def]
-   "|- []($P --> Q`) --> (P ~> Q)"
-   (fn _ => [Clarsimp_tac 1,
-             etac STL4E_gen 1,
-             auto_tac (temp_css addsimps2 Init_simps addSIs2 [PrimeDmd])
-            ]);
+Goalw  [leadsto_def] "|- []($P --> Q`) --> (P ~> Q)";
+by (Clarsimp_tac 1);
+by (etac STL4E_gen 1);
+by (auto_tac (temp_css addsimps2 Init_simps addSIs2 [PrimeDmd]));
+qed "EnsuresLeadsto2";
 
-qed_goalw "ensures" TLA.thy [leadsto_def]
+val [p1,p2] = goalw thy [leadsto_def]
   "[| |- $P & N --> P` | Q`; \
 \     |- ($P & N) & A --> Q` \
-\  |] ==> |- []N & []([]P --> <>A) --> (P ~> Q)"
-  (fn [p1,p2] => [Clarsimp_tac 1,
-                  etac STL4Edup 1, atac 1,
-                  Clarsimp_tac 1,
-                  subgoal_tac "sigmaa |= []($P --> P` | Q`)" 1,
-                   dtac (temp_use unless) 1,
-                   clarsimp_tac (temp_css addSDs2 [INV1]) 1,
-                   rtac ((temp_use (p2 RS DmdImpl)) RS (temp_use DmdPrime)) 1,
-                   force_tac (temp_css addSIs2 [BoxDmd_simple]
-                                       addsimps2 [split_box_conj,box_stp_act]) 1,
-                  force_tac (temp_css addEs2 [STL4E] addDs2 [p1]) 1
-                 ]);
+\  |] ==> |- []N & []([]P --> <>A) --> (P ~> Q)";
+by (Clarsimp_tac 1);
+by (etac STL4Edup 1); 
+by (atac 1);
+by (Clarsimp_tac 1);
+by (subgoal_tac "sigmaa |= []($P --> P` | Q`)" 1);
+ by (dtac (temp_use unless) 1);
+ by (clarsimp_tac (temp_css addSDs2 [INV1]) 1);
+ by (rtac ((temp_use (p2 RS DmdImpl)) RS (temp_use DmdPrime)) 1);
+ by (force_tac (temp_css addSIs2 [BoxDmd_simple]
+                         addsimps2 [split_box_conj,box_stp_act]) 1);
+by (force_tac (temp_css addEs2 [STL4E] addDs2 [p1]) 1);
+qed "ensures";
 
-qed_goal "ensures_simple" TLA.thy
+val prems = goal thy
   "[| |- $P & N --> P` | Q`; \
 \     |- ($P & N) & A --> Q` \
-\  |] ==> |- []N & []<>A --> (P ~> Q)"
-  (fn prems => [Clarsimp_tac 1,
-                rtac (temp_use ensures) 1,
-                TRYALL (ares_tac prems),
-                force_tac (temp_css addSEs2 [STL4E]) 1
-               ]);
+\  |] ==> |- []N & []<>A --> (P ~> Q)";
+by (Clarsimp_tac 1);
+by (rtac (temp_use ensures) 1);
+by (TRYALL (ares_tac prems));
+by (force_tac (temp_css addSEs2 [STL4E]) 1);
+qed "ensures_simple";
 
-qed_goal "EnsuresInfinite" TLA.thy
-   "[| sigma |= []<>P; sigma |= []A; |- A & $P --> Q` |] ==> sigma |= []<>Q"
-   (fn prems => [REPEAT (resolve_tac (prems @ [temp_use leadsto_infinite,
-					       temp_use EnsuresLeadsto]) 1)]);
+val prems = goal thy
+  "[| sigma |= []<>P; sigma |= []A; |- A & $P --> Q` |] ==> sigma |= []<>Q";
+by (REPEAT (resolve_tac (prems @ 
+                         (map temp_use [leadsto_infinite, EnsuresLeadsto])) 1));
+qed "EnsuresInfinite";
 
 
 (*** Gronning's lattice rules (taken from TLP) ***)
 section "Lattice rules";
 
-qed_goalw "LatticeReflexivity" TLA.thy [leadsto_def] "|- F ~> F"
-   (fn _ => [REPEAT (resolve_tac [necT,InitDmd_gen] 1)]);
-
-qed_goalw "LatticeTransitivity" TLA.thy [leadsto_def]
-   "|- (G ~> H) & (F ~> G) --> (F ~> H)"
-   (fn _ => [Clarsimp_tac 1,
-             etac dup_boxE 1,  (* [][](Init G --> H) *)
-	     merge_box_tac 1,
-	     clarsimp_tac (temp_css addSEs2 [STL4E]) 1,
-             rtac dup_dmdD 1,
-             subgoal_tac "sigmaa |= <>Init G" 1,
-             etac DmdImpl2 1, atac 1,
-             asm_simp_tac (simpset() addsimps [dmdInitD]) 1
-	    ]);
+Goalw [leadsto_def] "|- F ~> F";
+by (REPEAT (resolve_tac [necT,InitDmd_gen] 1));
+qed "LatticeReflexivity";
 
-qed_goalw "LatticeDisjunctionElim1" TLA.thy [leadsto_def]
-   "|- (F | G ~> H) --> (F ~> H)"
-   (fn _ => [ auto_tac (temp_css addsimps2 Init_simps addSEs2 [STL4E]) ]);
+Goalw [leadsto_def] "|- (G ~> H) & (F ~> G) --> (F ~> H)";
+by (Clarsimp_tac 1);
+by (etac dup_boxE 1);  (* [][](Init G --> H) *)
+by (merge_box_tac 1);
+by (clarsimp_tac (temp_css addSEs2 [STL4E]) 1);
+by (rtac dup_dmdD 1);
+by (subgoal_tac "sigmaa |= <>Init G" 1);
+ by (etac DmdImpl2 1); 
+ by (atac 1);
+by (asm_simp_tac (simpset() addsimps [dmdInitD]) 1);
+qed "LatticeTransitivity";
 
-qed_goalw "LatticeDisjunctionElim2" TLA.thy [leadsto_def]
-   "|- (F | G ~> H) --> (G ~> H)"
-   (fn _ => [ auto_tac (temp_css addsimps2 Init_simps addSEs2 [STL4E]) ]);
+Goalw [leadsto_def] "|- (F | G ~> H) --> (F ~> H)";
+by (auto_tac (temp_css addsimps2 Init_simps addSEs2 [STL4E]));
+qed "LatticeDisjunctionElim1";
 
-qed_goalw "LatticeDisjunctionIntro" TLA.thy [leadsto_def]
-   "|- (F ~> H) & (G ~> H) --> (F | G ~> H)"
-   (fn _ => [Clarsimp_tac 1,
-             merge_box_tac 1,
-	     auto_tac (temp_css addsimps2 Init_simps addSEs2 [STL4E])
-	    ]);
+Goalw [leadsto_def] "|- (F | G ~> H) --> (G ~> H)";
+by (auto_tac (temp_css addsimps2 Init_simps addSEs2 [STL4E]));
+qed "LatticeDisjunctionElim2";
 
-qed_goal "LatticeDisjunction" TLA.thy
-   "|- (F | G ~> H) = ((F ~> H) & (G ~> H))"
-   (fn _ => [auto_tac (temp_css addIs2 [LatticeDisjunctionIntro,
-                                LatticeDisjunctionElim1, LatticeDisjunctionElim2])]);
+Goalw [leadsto_def] "|- (F ~> H) & (G ~> H) --> (F | G ~> H)";
+by (Clarsimp_tac 1);
+by (merge_box_tac 1);
+by (auto_tac (temp_css addsimps2 Init_simps addSEs2 [STL4E]));
+qed "LatticeDisjunctionIntro";
 
-qed_goal "LatticeDiamond" TLA.thy
-   "|- (A ~> B | C) & (B ~> D) & (C ~> D) --> (A ~> D)"
-   (fn _ => [Clarsimp_tac 1,
-             subgoal_tac "sigma |= (B | C) ~> D" 1,
-	     eres_inst_tac [("G", "LIFT (B | C)")] (temp_use LatticeTransitivity) 1,
-	     ALLGOALS (fast_tac (temp_cs addSIs [LatticeDisjunctionIntro]))
-	    ]);
+Goal "|- (F | G ~> H) = ((F ~> H) & (G ~> H))";
+by (auto_tac (temp_css addIs2 [LatticeDisjunctionIntro,
+                               LatticeDisjunctionElim1, LatticeDisjunctionElim2]));
+qed "LatticeDisjunction";
 
-qed_goal "LatticeTriangle" TLA.thy
-   "|- (A ~> D | B) & (B ~> D) --> (A ~> D)"
-   (fn _ => [Clarsimp_tac 1,
-             subgoal_tac "sigma |= (D | B) ~> D" 1,
-	     eres_inst_tac [("G", "LIFT (D | B)")] (temp_use LatticeTransitivity) 1, atac 1,
-	     auto_tac (temp_css addSIs2 [LatticeDisjunctionIntro] 
-                                addIs2 [LatticeReflexivity])
-	    ]);
+Goal "|- (A ~> B | C) & (B ~> D) & (C ~> D) --> (A ~> D)";
+by (Clarsimp_tac 1);
+by (subgoal_tac "sigma |= (B | C) ~> D" 1);
+by (eres_inst_tac [("G", "LIFT (B | C)")] (temp_use LatticeTransitivity) 1);
+by (ALLGOALS (fast_tac (temp_cs addSIs [LatticeDisjunctionIntro])));
+qed "LatticeDiamond";
 
-qed_goal "LatticeTriangle2" TLA.thy
-   "|- (A ~> B | D) & (B ~> D) --> (A ~> D)"
-   (fn _ => [Clarsimp_tac 1,
-             subgoal_tac "sigma |= B | D ~> D" 1,
-	     eres_inst_tac [("G", "LIFT (B | D)")] (temp_use LatticeTransitivity) 1, atac 1,
-	     auto_tac (temp_css addSIs2 [LatticeDisjunctionIntro] 
-                                addIs2 [LatticeReflexivity])
-	    ]);
+Goal "|- (A ~> D | B) & (B ~> D) --> (A ~> D)";
+by (Clarsimp_tac 1);
+by (subgoal_tac "sigma |= (D | B) ~> D" 1);
+by (eres_inst_tac [("G", "LIFT (D | B)")] (temp_use LatticeTransitivity) 1); 
+by (atac 1);
+by (auto_tac (temp_css addIs2 [LatticeDisjunctionIntro,LatticeReflexivity]));
+qed "LatticeTriangle";
+
+Goal "|- (A ~> B | D) & (B ~> D) --> (A ~> D)";
+by (Clarsimp_tac 1);
+by (subgoal_tac "sigma |= B | D ~> D" 1);
+by (eres_inst_tac [("G", "LIFT (B | D)")] (temp_use LatticeTransitivity) 1); 
+by (atac 1);
+by (auto_tac (temp_css addIs2 [LatticeDisjunctionIntro,LatticeReflexivity]));
+qed "LatticeTriangle2";
 
 (*** Lamport's fairness rules ***)
 section "Fairness rules";
 
-qed_goal "WF1" TLA.thy
-   "[| |- $P & N  --> P` | Q`;   \
-\      |- ($P & N) & <A>_v --> Q`;   \
-\      |- $P & N --> $(Enabled(<A>_v)) |]   \
-\  ==> |- []N & WF(A)_v --> (P ~> Q)"  (fn prems => [
-             clarsimp_tac (temp_css addSDs2 [BoxWFI]) 1,
-             rtac (temp_use ensures) 1,
-             TRYALL (ares_tac prems),
-             etac STL4Edup 1, atac 1,
-             clarsimp_tac (temp_css addsimps2 [WF_def]) 1,
-             rtac (temp_use STL2) 1,
-             clarsimp_tac (temp_css addSEs2 [mp] addSIs2 [InitDmd]) 1,
-             resolve_tac ((map temp_use (prems RL [STL4])) RL [box_stp_actD]) 1,
-             asm_simp_tac (simpset() addsimps [split_box_conj,box_stp_actI]) 1
-            ]);
+val prems = goal thy
+  "[| |- $P & N  --> P` | Q`;   \
+\     |- ($P & N) & <A>_v --> Q`;   \
+\     |- $P & N --> $(Enabled(<A>_v)) |]   \
+\ ==> |- []N & WF(A)_v --> (P ~> Q)";
+by (clarsimp_tac (temp_css addSDs2 [BoxWFI]) 1);
+by (rtac (temp_use ensures) 1);
+by (TRYALL (ares_tac prems));
+by (etac STL4Edup 1); 
+by (atac 1);
+by (clarsimp_tac (temp_css addsimps2 [WF_def]) 1);
+by (rtac (temp_use STL2) 1);
+by (clarsimp_tac (temp_css addSEs2 [mp] addSIs2 [InitDmd]) 1);
+by (resolve_tac ((map temp_use (prems RL [STL4])) RL [box_stp_actD]) 1);
+by (asm_simp_tac (simpset() addsimps [split_box_conj,box_stp_actI]) 1);
+qed "WF1";
 
 (* Sometimes easier to use; designed for action B rather than state predicate Q *)
-qed_goalw "WF_leadsto" TLA.thy [leadsto_def]
-   "[| |- N & $P --> $Enabled (<A>_v);            \
-\      |- N & <A>_v --> B;                  \ 
-\      |- [](N & [~A]_v) --> stable P  |]  \
-\   ==> |- []N & WF(A)_v --> (P ~> B)"
-   (fn [prem1,prem2,prem3]
-       => [clarsimp_tac (temp_css addSDs2 [BoxWFI]) 1,
-           etac STL4Edup 1, atac 1,
-           Clarsimp_tac 1,
-           rtac (temp_use (prem2 RS DmdImpl)) 1,
-           rtac (temp_use BoxDmd_simple) 1, atac 1,
-           rtac classical 1,
-           rtac (temp_use STL2) 1,
-           clarsimp_tac (temp_css addsimps2 [WF_def] addSEs2 [mp] addSIs2 [InitDmd]) 1,
-           rtac ((temp_use (prem1 RS STL4)) RS box_stp_actD) 1,
-           asm_simp_tac (simpset() addsimps [split_box_conj,box_stp_act]) 1,
-           etac (temp_use INV1) 1,
-           rtac (temp_use prem3) 1,
-           asm_full_simp_tac (simpset() addsimps [split_box_conj,temp_use NotDmd,not_angle]) 1
-          ]);
+val [prem1,prem2,prem3] = goalw thy [leadsto_def]
+  "[| |- N & $P --> $Enabled (<A>_v);            \
+\     |- N & <A>_v --> B;                  \ 
+\     |- [](N & [~A]_v) --> stable P  |]  \
+\  ==> |- []N & WF(A)_v --> (P ~> B)";
+by (clarsimp_tac (temp_css addSDs2 [BoxWFI]) 1);
+by (etac STL4Edup 1); 
+by (atac 1);
+by (Clarsimp_tac 1);
+by (rtac (temp_use (prem2 RS DmdImpl)) 1);
+by (rtac (temp_use BoxDmd_simple) 1); 
+by (atac 1);
+by (rtac classical 1);
+by (rtac (temp_use STL2) 1);
+by (clarsimp_tac (temp_css addsimps2 [WF_def] addSEs2 [mp] addSIs2 [InitDmd]) 1);
+by (rtac ((temp_use (prem1 RS STL4)) RS box_stp_actD) 1);
+by (asm_simp_tac (simpset() addsimps [split_box_conj,box_stp_act]) 1);
+by (etac (temp_use INV1) 1);
+by (rtac (temp_use prem3) 1);
+by (asm_full_simp_tac (simpset() addsimps [split_box_conj,temp_use NotDmd,not_angle]) 1);
+qed "WF_leadsto";
 
-qed_goal "SF1" TLA.thy
-   "[| |- $P & N  --> P` | Q`;   \
-\      |- ($P & N) & <A>_v --> Q`;   \
-\      |- []P & []N & []F --> <>Enabled(<A>_v) |]   \
-\  ==> |- []N & SF(A)_v & []F --> (P ~> Q)"
-   (fn prems => [
-             clarsimp_tac (temp_css addSDs2 [BoxSFI]) 1,
-             rtac (temp_use ensures) 1,
-             TRYALL (ares_tac prems),
-             eres_inst_tac [("F","F")] dup_boxE 1,
-             merge_temp_box_tac 1,
-             etac STL4Edup 1, atac 1,
-             clarsimp_tac (temp_css addsimps2 [SF_def]) 1,
-             rtac (temp_use STL2) 1, etac mp 1,
-             resolve_tac (map temp_use (prems RL [STL4])) 1,
-             asm_simp_tac (simpset() addsimps [split_box_conj, STL3]) 1
-            ]);
+val prems = goal thy
+  "[| |- $P & N  --> P` | Q`;   \
+\     |- ($P & N) & <A>_v --> Q`;   \
+\     |- []P & []N & []F --> <>Enabled(<A>_v) |]   \
+\ ==> |- []N & SF(A)_v & []F --> (P ~> Q)";
+by (clarsimp_tac (temp_css addSDs2 [BoxSFI]) 1);
+by (rtac (temp_use ensures) 1);
+by (TRYALL (ares_tac prems));
+by (eres_inst_tac [("F","F")] dup_boxE 1);
+by (merge_temp_box_tac 1);
+by (etac STL4Edup 1); 
+by (atac 1);
+by (clarsimp_tac (temp_css addsimps2 [SF_def]) 1);
+by (rtac (temp_use STL2) 1); 
+by (etac mp 1);
+by (resolve_tac (map temp_use (prems RL [STL4])) 1);
+by (asm_simp_tac (simpset() addsimps [split_box_conj, STL3]) 1);
+qed "SF1";
 
-qed_goal "WF2" TLA.thy
-   "[| |- N & <B>_f --> <M>_g;   \
-\      |- $P & P` & <N & A>_f --> B;   \
-\      |- P & Enabled(<M>_g) --> Enabled(<A>_f);   \
-\      |- [](N & [~B]_f) & WF(A)_f & []F & <>[]Enabled(<M>_g) --> <>[]P |]   \
-\  ==> |- []N & WF(A)_f & []F --> WF(M)_g"
-(fn [prem1,prem2,prem3,prem4] => [
-	   clarsimp_tac (temp_css addSDs2 [BoxWFI, (temp_use BoxDmdBox) RS iffD2] 
-                            addsimps2 [read_instantiate [("A","M")] WF_def]) 1,
-           eres_inst_tac [("F","F")] dup_boxE 1,
-           merge_temp_box_tac 1,
-           etac STL4Edup 1, atac 1,
-           clarsimp_tac (temp_css addSIs2 [(temp_use BoxDmd_simple) RS (temp_use (prem1 RS DmdImpl))]) 1,
-           rtac classical 1,
-           subgoal_tac "sigmaa |= <>(($P & P` & N) & <A>_f)" 1,
-           force_tac (temp_css addsimps2 [angle_def] addSIs2 [prem2] addSEs2 [DmdImplE]) 1,
-           rtac (temp_use (rewrite_rule [temp_rewrite DmdDmd] (BoxDmd_simple RS DmdImpl))) 1,
-           asm_full_simp_tac (simpset() addsimps [temp_use NotDmd, not_angle]) 1,
-           merge_act_box_tac 1,
-           forward_tac [temp_use prem4] 1, TRYALL atac,
-           dtac (temp_use STL6) 1, atac 1, 
-           eres_inst_tac [("V","sigmaa |= <>[]P")] thin_rl 1,
-           eres_inst_tac [("V","sigmaa |= []F")] thin_rl 1,
-           dtac (temp_use BoxWFI) 1,
-           eres_inst_tac [("F", "ACT N & [~B]_f")] dup_boxE 1,
-           merge_temp_box_tac 1,
-           etac DmdImpldup 1, atac 1,
-           auto_tac (temp_css addsimps2 [split_box_conj,STL3,WF_Box,box_stp_act]),
-           force_tac (temp_css addSEs2 [read_instantiate [("P","P")] TLA2E]) 1,
-           rtac (temp_use STL2) 1,
-           force_tac (temp_css addsimps2 [WF_def,split_box_conj] addSEs2 [mp] 
-                               addSIs2 [InitDmd, prem3 RS STL4]) 1
-	  ]);
+val [prem1,prem2,prem3,prem4] = goal thy
+  "[| |- N & <B>_f --> <M>_g;   \
+\     |- $P & P` & <N & A>_f --> B;   \
+\     |- P & Enabled(<M>_g) --> Enabled(<A>_f);   \
+\     |- [](N & [~B]_f) & WF(A)_f & []F & <>[]Enabled(<M>_g) --> <>[]P |]   \
+\ ==> |- []N & WF(A)_f & []F --> WF(M)_g";
+by (clarsimp_tac (temp_css addSDs2 [BoxWFI, (temp_use BoxDmdBox) RS iffD2] 
+                           addsimps2 [read_instantiate [("A","M")] WF_def]) 1);
+by (eres_inst_tac [("F","F")] dup_boxE 1);
+by (merge_temp_box_tac 1);
+by (etac STL4Edup 1); 
+by (atac 1);
+by (clarsimp_tac (temp_css addSIs2
+         [(temp_use BoxDmd_simple) RS (temp_use (prem1 RS DmdImpl))]) 1);
+by (rtac classical 1);
+by (subgoal_tac "sigmaa |= <>(($P & P` & N) & <A>_f)" 1);
+ by (force_tac (temp_css addsimps2 [angle_def] addSIs2 [prem2] addSEs2 [DmdImplE]) 1);
+by (rtac (temp_use (rewrite_rule [temp_rewrite DmdDmd] (BoxDmd_simple RS DmdImpl))) 1);
+by (asm_full_simp_tac (simpset() addsimps [temp_use NotDmd, not_angle]) 1);
+by (merge_act_box_tac 1);
+by (forward_tac [temp_use prem4] 1); 
+by (TRYALL atac);
+by (dtac (temp_use STL6) 1); 
+by (atac 1);
+by (eres_inst_tac [("V","sigmaa |= <>[]P")] thin_rl 1);
+by (eres_inst_tac [("V","sigmaa |= []F")] thin_rl 1);
+by (dtac (temp_use BoxWFI) 1);
+by (eres_inst_tac [("F", "ACT N & [~B]_f")] dup_boxE 1);
+by (merge_temp_box_tac 1);
+by (etac DmdImpldup 1); 
+by (atac 1);
+by (auto_tac (temp_css addsimps2 [split_box_conj,STL3,WF_Box,box_stp_act]));
+ by (force_tac (temp_css addSEs2 [read_instantiate [("P","P")] TLA2E]) 1);
+by (rtac (temp_use STL2) 1);
+by (force_tac (temp_css addsimps2 [WF_def,split_box_conj] addSEs2 [mp] 
+                        addSIs2 [InitDmd, prem3 RS STL4]) 1);
+qed "WF2";
 
-qed_goal "SF2" TLA.thy
-   "[| |- N & <B>_f --> <M>_g;   \
-\      |- $P & P` & <N & A>_f --> B;   \
-\      |- P & Enabled(<M>_g) --> Enabled(<A>_f);   \
-\      |- [](N & [~B]_f) & SF(A)_f & []F & []<>Enabled(<M>_g) --> <>[]P |]   \
-\  ==> |- []N & SF(A)_f & []F --> SF(M)_g"
-(fn [prem1,prem2,prem3,prem4] => [
-	   clarsimp_tac (temp_css addSDs2 [BoxSFI] 
-                            addsimps2 [read_instantiate [("A","M")] SF_def]) 1,
-           eres_inst_tac [("F","F")] dup_boxE 1,
-           eres_inst_tac [("F","TEMP <>Enabled(<M>_g)")] dup_boxE 1,
-           merge_temp_box_tac 1,
-           etac STL4Edup 1, atac 1,
-           clarsimp_tac (temp_css addSIs2 [(temp_use BoxDmd_simple) RS (temp_use (prem1 RS DmdImpl))]) 1,
-           rtac classical 1,
-           subgoal_tac "sigmaa |= <>(($P & P` & N) & <A>_f)" 1,
-           force_tac (temp_css addsimps2 [angle_def] addSIs2 [prem2] addSEs2 [DmdImplE]) 1,
-           rtac (temp_use (rewrite_rule [temp_rewrite DmdDmd] (BoxDmd_simple RS DmdImpl))) 1,
-           asm_full_simp_tac (simpset() addsimps [temp_use NotDmd, not_angle]) 1,
-           merge_act_box_tac 1,
-           forward_tac [temp_use prem4] 1, TRYALL atac,
-           eres_inst_tac [("V","sigmaa |= []F")] thin_rl 1,
-           dtac (temp_use BoxSFI) 1,
-           eres_inst_tac [("F","TEMP <>Enabled(<M>_g)")] dup_boxE 1,
-           eres_inst_tac [("F", "ACT N & [~B]_f")] dup_boxE 1,
-           merge_temp_box_tac 1,
-           etac DmdImpldup 1, atac 1,
-           auto_tac (temp_css addsimps2 [split_box_conj,STL3,SF_Box,box_stp_act]),
-           force_tac (temp_css addSEs2 [read_instantiate [("P","P")] TLA2E]) 1,
-           rtac (temp_use STL2) 1,
-           force_tac (temp_css addsimps2 [SF_def,split_box_conj] addSEs2 [mp,InfImpl] 
-                               addSIs2 [prem3]) 1
-	  ]);
+val [prem1,prem2,prem3,prem4] = goal thy
+  "[| |- N & <B>_f --> <M>_g;   \
+\     |- $P & P` & <N & A>_f --> B;   \
+\     |- P & Enabled(<M>_g) --> Enabled(<A>_f);   \
+\     |- [](N & [~B]_f) & SF(A)_f & []F & []<>Enabled(<M>_g) --> <>[]P |]   \
+\ ==> |- []N & SF(A)_f & []F --> SF(M)_g";
+by (clarsimp_tac (temp_css addSDs2 [BoxSFI] 
+                           addsimps2 [read_instantiate [("A","M")] SF_def]) 1);
+by (eres_inst_tac [("F","F")] dup_boxE 1);
+by (eres_inst_tac [("F","TEMP <>Enabled(<M>_g)")] dup_boxE 1);
+by (merge_temp_box_tac 1);
+by (etac STL4Edup 1); 
+by (atac 1);
+by (clarsimp_tac (temp_css addSIs2
+        [(temp_use BoxDmd_simple) RS (temp_use (prem1 RS DmdImpl))]) 1);
+by (rtac classical 1);
+by (subgoal_tac "sigmaa |= <>(($P & P` & N) & <A>_f)" 1);
+ by (force_tac (temp_css addsimps2 [angle_def] addSIs2 [prem2] addSEs2 [DmdImplE]) 1);
+by (rtac (temp_use (rewrite_rule [temp_rewrite DmdDmd] (BoxDmd_simple RS DmdImpl))) 1);
+by (asm_full_simp_tac (simpset() addsimps [temp_use NotDmd, not_angle]) 1);
+by (merge_act_box_tac 1);
+by (forward_tac [temp_use prem4] 1); 
+by (TRYALL atac);
+by (eres_inst_tac [("V","sigmaa |= []F")] thin_rl 1);
+by (dtac (temp_use BoxSFI) 1);
+by (eres_inst_tac [("F","TEMP <>Enabled(<M>_g)")] dup_boxE 1);
+by (eres_inst_tac [("F", "ACT N & [~B]_f")] dup_boxE 1);
+by (merge_temp_box_tac 1);
+by (etac DmdImpldup 1); 
+by (atac 1);
+by (auto_tac (temp_css addsimps2 [split_box_conj,STL3,SF_Box,box_stp_act]));
+ by (force_tac (temp_css addSEs2 [read_instantiate [("P","P")] TLA2E]) 1);
+by (rtac (temp_use STL2) 1);
+by (force_tac (temp_css addsimps2 [SF_def,split_box_conj] addSEs2 [mp,InfImpl] 
+                        addSIs2 [prem3]) 1);
+qed "SF2";
 
 (* ------------------------------------------------------------------------- *)
 (***           Liveness proofs by well-founded orderings                   ***)
 (* ------------------------------------------------------------------------- *)
 section "Well-founded orderings";
 
-qed_goal "wf_leadsto" TLA.thy
+val p1::prems = goal thy
   "[| wf r;  \
-\     !!x. sigma |= F x ~> (G | (? y. #((y,x):r) & F y))   \
-\  |] ==> sigma |= F x ~> G"
-  (fn p1::prems =>
-     [rtac (p1 RS wf_induct) 1,
-      rtac (temp_use LatticeTriangle) 1,
-      resolve_tac prems 1,
-      auto_tac (temp_css addsimps2 [leadsto_exists]),
-      case_tac "(y,x):r" 1,
-       Force_tac 1,
-      force_tac (temp_css addsimps2 leadsto_def::Init_simps addSIs2 [necT]) 1]);
+\     !!x. sigma |= F x ~> (G | (EX y. #((y,x):r) & F y))   \
+\  |] ==> sigma |= F x ~> G";
+by (rtac (p1 RS wf_induct) 1);
+by (rtac (temp_use LatticeTriangle) 1);
+by (resolve_tac prems 1);
+by (auto_tac (temp_css addsimps2 [leadsto_exists]));
+by (case_tac "(y,x):r" 1);
+ by (Force_tac 1);
+by (force_tac (temp_css addsimps2 leadsto_def::Init_simps addSIs2 [necT]) 1);
+qed "wf_leadsto";
 
 (* If r is well-founded, state function v cannot decrease forever *)
-qed_goal "wf_not_box_decrease" TLA.thy
-  "!!r. wf r ==> |- [][ (v`, $v) : #r ]_v --> <>[][#False]_v"
-  (fn _ => [Clarsimp_tac 1,
-            rtac ccontr 1,
-            subgoal_tac "sigma |= (? x. v=#x) ~> #False" 1,
-             dtac ((temp_use leadsto_false) RS iffD1 RS (temp_use STL2_gen)) 1,
-             force_tac (temp_css addsimps2 Init_defs) 1,
-            clarsimp_tac (temp_css addsimps2 [leadsto_exists,not_square]@more_temp_simps) 1,
-            etac wf_leadsto 1,
-            rtac (temp_use ensures_simple) 1, TRYALL atac,
-            auto_tac (temp_css addsimps2 [square_def,angle_def])
-           ]);
+Goal "!!r. wf r ==> |- [][ (v`, $v) : #r ]_v --> <>[][#False]_v";
+by (Clarsimp_tac 1);
+by (rtac ccontr 1);
+by (subgoal_tac "sigma |= (EX x. v=#x) ~> #False" 1);
+ by (dtac ((temp_use leadsto_false) RS iffD1 RS (temp_use STL2_gen)) 1);
+ by (force_tac (temp_css addsimps2 Init_defs) 1);
+by (clarsimp_tac (temp_css addsimps2 [leadsto_exists,not_square]@more_temp_simps) 1);
+by (etac wf_leadsto 1);
+by (rtac (temp_use ensures_simple) 1); 
+by (TRYALL atac);
+by (auto_tac (temp_css addsimps2 [square_def,angle_def]));
+qed "wf_not_box_decrease";
 
 (* "wf r  ==>  |- <>[][ (v`, $v) : #r ]_v --> <>[][#False]_v" *)
 bind_thm("wf_not_dmd_box_decrease",
@@ -934,35 +974,36 @@
 (* If there are infinitely many steps where v decreases, then there
    have to be infinitely many non-stuttering steps where v doesn't decrease.
 *)
-qed_goal "wf_box_dmd_decrease" TLA.thy
-  "wf r ==> |- []<>((v`, $v) : #r) --> []<><(v`, $v) ~: #r>_v"
-  (fn [prem] => [
-            Clarsimp_tac 1,
-            rtac ccontr 1,
-            asm_full_simp_tac (simpset() addsimps not_angle::more_temp_simps) 1,
-            dtac (prem RS (temp_use wf_not_dmd_box_decrease)) 1,
-            dtac (temp_use BoxDmdDmdBox) 1, atac 1,
-            subgoal_tac "sigma |= []<>((#False)::action)" 1,
-            Force_tac 1,
-            etac STL4E 1,
-            rtac DmdImpl 1,
-            force_tac (temp_css addIs2 [prem RS wf_irrefl]) 1
-           ]);
+val [prem] = goal thy
+  "wf r ==> |- []<>((v`, $v) : #r) --> []<><(v`, $v) ~: #r>_v";
+by (Clarsimp_tac 1);
+by (rtac ccontr 1);
+by (asm_full_simp_tac (simpset() addsimps not_angle::more_temp_simps) 1);
+by (dtac (prem RS (temp_use wf_not_dmd_box_decrease)) 1);
+by (dtac (temp_use BoxDmdDmdBox) 1); 
+by (atac 1);
+by (subgoal_tac "sigma |= []<>((#False)::action)" 1);
+ by (Force_tac 1);
+by (etac STL4E 1);
+by (rtac DmdImpl 1);
+by (force_tac (temp_css addIs2 [prem RS wf_irrefl]) 1);
+qed "wf_box_dmd_decrease";
 
 (* In particular, for natural numbers, if n decreases infinitely often
    then it has to increase infinitely often.
 *)
-qed_goal "nat_box_dmd_decrease" TLA.thy
-  "!!n::nat stfun. |- []<>(n` < $n) --> []<>($n < n`)"
-  (K [Clarsimp_tac 1,
-      subgoal_tac "sigma |= []<><~( (n`,$n) : #less_than )>_n" 1,
-      etac thin_rl 1, etac STL4E 1, rtac DmdImpl 1,
-      clarsimp_tac (temp_css addsimps2 [angle_def]) 1,
-      rtac nat_less_cases 1,
-      Auto_tac,
-      rtac (temp_use wf_box_dmd_decrease) 1,
-      auto_tac (temp_css addSEs2 [STL4E,DmdImplE])
-     ]);
+Goal "!!n::nat stfun. |- []<>(n` < $n) --> []<>($n < n`)";
+by (Clarsimp_tac 1);
+by (subgoal_tac "sigma |= []<><~( (n`,$n) : #less_than )>_n" 1);
+ by (etac thin_rl 1); 
+ by (etac STL4E 1); 
+ by (rtac DmdImpl 1);
+ by (clarsimp_tac (temp_css addsimps2 [angle_def]) 1);
+ by (rtac nat_less_cases 1);
+ by Auto_tac;
+by (rtac (temp_use wf_box_dmd_decrease) 1);
+by (auto_tac (temp_css addSEs2 [STL4E,DmdImplE]));
+qed "nat_box_dmd_decrease";
 
 
 (* ------------------------------------------------------------------------- *)
@@ -970,47 +1011,51 @@
 (* ------------------------------------------------------------------------- *)
 section "Flexible quantification";
 
-qed_goal "aallI" TLA.thy 
-  "[| basevars vs; (!!x. basevars (x,vs) ==> sigma |= F x) |] ==> sigma |= (AALL x. F x)"
-  (fn [prem1,prem2] => [auto_tac (temp_css addsimps2 [aall_def] addSEs2 [eexE] 
-                                   addSIs2 [prem1] addSDs2 [prem2])]);
+val [prem1,prem2] = goal thy
+  "[| basevars vs; (!!x. basevars (x,vs) ==> sigma |= F x) |]\
+\  ==> sigma |= (AALL x. F x)";
+by (auto_tac (temp_css addsimps2 [aall_def] addSEs2 [eexE] 
+                       addSIs2 [prem1] addSDs2 [prem2]));
+qed "aallI";
 
-qed_goalw "aallE" TLA.thy [aall_def] "|- (AALL x. F x) --> F x"
-   (K [Clarsimp_tac 1, etac swap 1,
-       force_tac (temp_css addSIs2 [eexI]) 1]);
+Goalw [aall_def] "|- (AALL x. F x) --> F x";
+by (Clarsimp_tac 1);
+by (etac swap 1);
+by (force_tac (temp_css addSIs2 [eexI]) 1);
+qed "aallE";
 
 (* monotonicity of quantification *)
-qed_goal "eex_mono" TLA.thy
-  "[| sigma |= EEX x. F x; !!x. sigma |= F x --> G x |] ==> sigma |= EEX x. G x"
-  (fn [min,maj] => [rtac (unit_base RS (min RS eexE)) 1,
-                    rtac (temp_use eexI) 1,
-                    etac ((rewrite_rule intensional_rews maj) RS mp) 1
-                   ]);
+val [min,maj] = goal thy
+  "[| sigma |= EEX x. F x; !!x. sigma |= F x --> G x |] ==> sigma |= EEX x. G x";
+by (rtac (unit_base RS (min RS eexE)) 1);
+by (rtac (temp_use eexI) 1);
+by (etac ((rewrite_rule intensional_rews maj) RS mp) 1);
+qed "eex_mono";
 
-qed_goal "aall_mono" TLA.thy
-  "[| sigma |= AALL x. F(x); !!x. sigma |= F(x) --> G(x) |] ==> sigma |= AALL x. G(x)"
-  (fn [min,maj] => [rtac (unit_base RS aallI) 1,
-                    rtac ((rewrite_rule intensional_rews maj) RS mp) 1,
-                    rtac (min RS (temp_use aallE)) 1
-                   ]);
+val [min,maj] = goal thy
+  "[| sigma |= AALL x. F(x); !!x. sigma |= F(x) --> G(x) |] ==> sigma |= AALL x. G(x)";
+by (rtac (unit_base RS aallI) 1);
+by (rtac ((rewrite_rule intensional_rews maj) RS mp) 1);
+by (rtac (min RS (temp_use aallE)) 1);
+qed "aall_mono";
 
 (* Derived history introduction rule *)
-qed_goal "historyI" TLA.thy
+val [p1,p2,p3,p4,p5] = goal thy
   "[| sigma |= Init I; sigma |= []N; basevars vs; \
 \     (!!h. basevars(h,vs) ==> |- I & h = ha --> HI h); \
 \     (!!h s t. [| basevars(h,vs); N (s,t); h t = hb (h s) (s,t) |] ==> HN h (s,t)) \
-\  |] ==> sigma |= EEX h. Init (HI h) & [](HN h)" 
-  (fn [p1,p2,p3,p4,p5] 
-   => [rtac ((temp_use history) RS eexE) 1,
-       rtac p3 1,
-       rtac (temp_use eexI) 1,
-       Clarsimp_tac 1, rtac conjI 1,
-       cut_facts_tac [p2] 2,
-       merge_box_tac 2,
-       force_tac (temp_css addSEs2 [STL4E,p5]) 2,
-       cut_facts_tac [p1] 1,
-       force_tac (temp_css addsimps2 Init_defs addSEs2 [p4]) 1
-      ]);
+\  |] ==> sigma |= EEX h. Init (HI h) & [](HN h)";
+by (rtac ((temp_use history) RS eexE) 1);
+ by (rtac p3 1);
+by (rtac (temp_use eexI) 1);
+by (Clarsimp_tac 1); 
+by (rtac conjI 1);
+by (cut_facts_tac [p2] 2);
+by (merge_box_tac 2);
+by (force_tac (temp_css addSEs2 [STL4E,p5]) 2);
+by (cut_facts_tac [p1] 1);
+by (force_tac (temp_css addsimps2 Init_defs addSEs2 [p4]) 1);
+qed "historyI";
 
 (* ----------------------------------------------------------------------
    example of a history variable: existence of a clock
@@ -1022,4 +1067,3 @@
 (** solved **)
 
 ---------------------------------------------------------------------- *)
-
--- a/src/HOL/TLA/TLA.thy	Thu Aug 03 19:28:37 2000 +0200
+++ b/src/HOL/TLA/TLA.thy	Thu Aug 03 19:29:03 2000 +0200
@@ -81,7 +81,7 @@
   primeI     "|- []P --> Init P`"
   primeE     "|- [](Init P --> []F) --> Init P` --> (F --> []F)"
   indT       "|- [](Init P & ~[]F --> Init P` & F) --> Init P --> []F"
-  allT       "|- (! x. [](F x)) = ([](! x. F x))"
+  allT       "|- (ALL x. [](F x)) = ([](ALL x. F x))"
 
   necT       "|- F ==> |- []F"      (* polymorphic *)