src/HOL/UNITY/SubstAx.ML
changeset 5253 82a5ca6290aa
parent 5240 bbcd79ef7cf2
child 5257 c03e3ba9cbcc
--- a/src/HOL/UNITY/SubstAx.ML	Wed Aug 05 10:56:58 1998 +0200
+++ b/src/HOL/UNITY/SubstAx.ML	Wed Aug 05 10:57:25 1998 +0200
@@ -16,31 +16,27 @@
 
 (*** Introduction rules: Basis, Trans, Union ***)
 
-Goal "leadsTo Acts A B ==> LeadsTo(Init,Acts) A B";
+Goal "leadsTo (Acts prg) A B ==> LeadsTo prg A B";
 by (simp_tac (simpset() addsimps [LeadsTo_def]) 1);
 by (blast_tac (claset() addIs [PSP_stable2, stable_reachable]) 1);
 qed "leadsTo_imp_LeadsTo";
 
-Goal "[| constrains Acts (reachable(Init,Acts) Int (A - A'))   \
-\                         (A Un A'); \
-\         transient  Acts (reachable(Init,Acts) Int (A-A')) |]   \
-\     ==> LeadsTo(Init,Acts) A A'";
-by (simp_tac (simpset() addsimps [LeadsTo_def]) 1);
+Goal "ensures (Acts prg) (reachable prg Int A) (reachable prg Int A') \
+\     ==> LeadsTo prg A A'";
+by (full_simp_tac (simpset() addsimps [ensures_def, LeadsTo_def]) 1);
 by (rtac (stable_reachable RS stable_ensures_Int RS leadsTo_Basis) 1);
-by (assume_tac 2);
-by (asm_simp_tac 
-    (simpset() addsimps [Int_Un_distrib RS sym, Diff_Int_distrib RS sym,
-			 stable_constrains_Int]) 1);
+by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [Diff_Int_distrib])));
+by (blast_tac (claset() addIs [constrains_weaken]) 1);
 qed "LeadsTo_Basis";
 
-Goal "[| LeadsTo(Init,Acts) A B;  LeadsTo(Init,Acts) B C |] \
-\      ==> LeadsTo(Init,Acts) A C";
+Goal "[| LeadsTo prg A B;  LeadsTo prg B C |] \
+\      ==> LeadsTo prg A C";
 by (full_simp_tac (simpset() addsimps [LeadsTo_def]) 1);
 by (blast_tac (claset() addIs [leadsTo_Trans]) 1);
 qed "LeadsTo_Trans";
 
 val [prem] = goalw thy [LeadsTo_def]
- "(!!A. A : S ==> LeadsTo(Init,Acts) A B) ==> LeadsTo(Init,Acts) (Union S) B";
+ "(!!A. A : S ==> LeadsTo prg A B) ==> LeadsTo prg (Union S) B";
 by (Simp_tac 1);
 by (stac Int_Union 1);
 by (blast_tac (claset() addIs [leadsTo_UN,
@@ -50,42 +46,42 @@
 
 (*** Derived rules ***)
 
-Goal "id: Acts ==> LeadsTo(Init,Acts) A UNIV";
+Goal "id: (Acts prg) ==> LeadsTo prg A UNIV";
 by (asm_simp_tac (simpset() addsimps [LeadsTo_def, 
 				      Int_lower1 RS subset_imp_leadsTo]) 1);
 qed "LeadsTo_UNIV";
 Addsimps [LeadsTo_UNIV];
 
 (*Useful with cancellation, disjunction*)
-Goal "LeadsTo(Init,Acts) A (A' Un A') ==> LeadsTo(Init,Acts) A A'";
+Goal "LeadsTo prg A (A' Un A') ==> LeadsTo prg A A'";
 by (asm_full_simp_tac (simpset() addsimps Un_ac) 1);
 qed "LeadsTo_Un_duplicate";
 
-Goal "LeadsTo(Init,Acts) A (A' Un C Un C) ==> LeadsTo(Init,Acts) A (A' Un C)";
+Goal "LeadsTo prg A (A' Un C Un C) ==> LeadsTo prg A (A' Un C)";
 by (asm_full_simp_tac (simpset() addsimps Un_ac) 1);
 qed "LeadsTo_Un_duplicate2";
 
 val prems = goal thy
-   "(!!i. i : I ==> LeadsTo(Init,Acts) (A i) B) \
-\   ==> LeadsTo(Init,Acts) (UN i:I. A i) B";
+   "(!!i. i : I ==> LeadsTo prg (A i) B) \
+\   ==> LeadsTo prg (UN i:I. A i) B";
 by (simp_tac (simpset() addsimps [Union_image_eq RS sym]) 1);
 by (blast_tac (claset() addIs (LeadsTo_Union::prems)) 1);
 qed "LeadsTo_UN";
 
 (*Binary union introduction rule*)
-Goal "[| LeadsTo(Init,Acts) A C; LeadsTo(Init,Acts) B C |] ==> LeadsTo(Init,Acts) (A Un B) C";
+Goal "[| LeadsTo prg A C; LeadsTo prg B C |] ==> LeadsTo prg (A Un B) C";
 by (stac Un_eq_Union 1);
 by (blast_tac (claset() addIs [LeadsTo_Union]) 1);
 qed "LeadsTo_Un";
 
 
-Goal "[| reachable(Init,Acts) Int A <= B;  id: Acts |] \
-\     ==> LeadsTo(Init,Acts) A B";
+Goal "[| reachable prg Int A <= B;  id: (Acts prg) |] \
+\     ==> LeadsTo prg A B";
 by (simp_tac (simpset() addsimps [LeadsTo_def]) 1);
 by (blast_tac (claset() addIs [subset_imp_leadsTo]) 1);
 qed "Int_subset_imp_LeadsTo";
 
-Goal "[| A <= B;  id: Acts |] ==> LeadsTo(Init,Acts) A B";
+Goal "[| A <= B;  id: (Acts prg) |] ==> LeadsTo prg A B";
 by (simp_tac (simpset() addsimps [LeadsTo_def]) 1);
 by (blast_tac (claset() addIs [subset_imp_leadsTo]) 1);
 qed "subset_imp_LeadsTo";
@@ -93,61 +89,62 @@
 bind_thm ("empty_LeadsTo", empty_subsetI RS subset_imp_LeadsTo);
 Addsimps [empty_LeadsTo];
 
-Goal "[| reachable(Init,Acts) Int A = {};  id: Acts |] \
-\     ==> LeadsTo(Init,Acts) A B";
+Goal "[| reachable prg Int A = {};  id: (Acts prg) |] \
+\     ==> LeadsTo prg A B";
 by (asm_simp_tac (simpset() addsimps [Int_subset_imp_LeadsTo]) 1);
 qed "Int_empty_LeadsTo";
 
 
-Goal "[| LeadsTo(Init,Acts) A A';   \
-\        reachable(Init,Acts) Int A' <= B' |] \
-\     ==> LeadsTo(Init,Acts) A B'";
+Goal "[| LeadsTo prg A A';   \
+\        reachable prg Int A' <= B' |] \
+\     ==> LeadsTo prg A B'";
 by (full_simp_tac (simpset() addsimps [LeadsTo_def]) 1);
 by (blast_tac (claset() addIs [leadsTo_weaken_R]) 1);
 qed_spec_mp "LeadsTo_weaken_R";
 
 
-Goal "[| LeadsTo(Init,Acts) A A'; \
-\        reachable(Init,Acts) Int B <= A; id: Acts |]  \
-\     ==> LeadsTo(Init,Acts) B A'";
+Goal "[| LeadsTo prg A A'; \
+\        reachable prg Int B <= A; id: (Acts prg) |]  \
+\     ==> LeadsTo prg B A'";
 by (full_simp_tac (simpset() addsimps [LeadsTo_def]) 1);
 by (blast_tac (claset() addIs [leadsTo_weaken_L]) 1);
 qed_spec_mp "LeadsTo_weaken_L";
 
 
 (*Distributes over binary unions*)
-Goal "id: Acts ==> \
-\       LeadsTo(Init,Acts) (A Un B) C  =  \
-\       (LeadsTo(Init,Acts) A C & LeadsTo(Init,Acts) B C)";
+Goal "id: (Acts prg) ==> \
+\       LeadsTo prg (A Un B) C  =  \
+\       (LeadsTo prg A C & LeadsTo prg B C)";
 by (blast_tac (claset() addIs [LeadsTo_Un, LeadsTo_weaken_L]) 1);
 qed "LeadsTo_Un_distrib";
 
-Goal "id: Acts ==> \
-\       LeadsTo(Init,Acts) (UN i:I. A i) B  =  \
-\       (ALL i : I. LeadsTo(Init,Acts) (A i) B)";
+Goal "id: (Acts prg) ==> \
+\       LeadsTo prg (UN i:I. A i) B  =  \
+\       (ALL i : I. LeadsTo prg (A i) B)";
 by (blast_tac (claset() addIs [LeadsTo_UN, LeadsTo_weaken_L]) 1);
 qed "LeadsTo_UN_distrib";
 
-Goal "id: Acts ==> \
-\       LeadsTo(Init,Acts) (Union S) B  =  \
-\       (ALL A : S. LeadsTo(Init,Acts) A B)";
+Goal "id: (Acts prg) ==> \
+\       LeadsTo prg (Union S) B  =  \
+\       (ALL A : S. LeadsTo prg A B)";
 by (blast_tac (claset() addIs [LeadsTo_Union, LeadsTo_weaken_L]) 1);
 qed "LeadsTo_Union_distrib";
 
 
-Goal "[| LeadsTo(Init,Acts) A A'; id: Acts;   \
-\        reachable(Init,Acts) Int B  <= A;     \
-\        reachable(Init,Acts) Int A' <= B' |] \
-\     ==> LeadsTo(Init,Acts) B B'";
+Goal "[| LeadsTo prg A A'; id: (Acts prg);   \
+\        reachable prg Int B  <= A;     \
+\        reachable prg Int A' <= B' |] \
+\     ==> LeadsTo prg B B'";
 (*PROOF FAILED: why?*)
 by (blast_tac (claset() addIs [LeadsTo_Trans, LeadsTo_weaken_R,
 			       LeadsTo_weaken_L]) 1);
 qed "LeadsTo_weaken";
 
 
-(*Set difference: maybe combine with leadsTo_weaken_L??*)
-Goal "[| LeadsTo(Init,Acts) (A-B) C; LeadsTo(Init,Acts) B C; id: Acts |] \
-\       ==> LeadsTo(Init,Acts) A C";
+(*Set difference: maybe combine with leadsTo_weaken_L??
+  This is the most useful form of the "disjunction" rule*)
+Goal "[| LeadsTo prg (A-B) C; LeadsTo prg B C; id: (Acts prg) |] \
+\       ==> LeadsTo prg A C";
 by (blast_tac (claset() addIs [LeadsTo_Un, LeadsTo_weaken]) 1);
 qed "LeadsTo_Diff";
 
@@ -156,8 +153,8 @@
     see ball_constrains_UN in UNITY.ML***)
 
 val prems = goal thy
-   "(!! i. i:I ==> LeadsTo(Init,Acts) (A i) (A' i)) \
-\   ==> LeadsTo(Init,Acts) (UN i:I. A i) (UN i:I. A' i)";
+   "(!! i. i:I ==> LeadsTo prg (A i) (A' i)) \
+\   ==> LeadsTo prg (UN i:I. A i) (UN i:I. A' i)";
 by (simp_tac (simpset() addsimps [Union_image_eq RS sym]) 1);
 by (blast_tac (claset() addIs [LeadsTo_Union, LeadsTo_weaken_R] 
                         addIs prems) 1);
@@ -166,22 +163,22 @@
 
 (*Version with no index set*)
 val prems = goal thy
-   "(!! i. LeadsTo(Init,Acts) (A i) (A' i)) \
-\   ==> LeadsTo(Init,Acts) (UN i. A i) (UN i. A' i)";
+   "(!! i. LeadsTo prg (A i) (A' i)) \
+\   ==> LeadsTo prg (UN i. A i) (UN i. A' i)";
 by (blast_tac (claset() addIs [LeadsTo_UN_UN] 
                         addIs prems) 1);
 qed "LeadsTo_UN_UN_noindex";
 
 (*Version with no index set*)
-Goal "ALL i. LeadsTo(Init,Acts) (A i) (A' i) \
-\           ==> LeadsTo(Init,Acts) (UN i. A i) (UN i. A' i)";
+Goal "ALL i. LeadsTo prg (A i) (A' i) \
+\           ==> LeadsTo prg (UN i. A i) (UN i. A' i)";
 by (blast_tac (claset() addIs [LeadsTo_UN_UN]) 1);
 qed "all_LeadsTo_UN_UN";
 
 
 (*Binary union version*)
-Goal "[| LeadsTo(Init,Acts) A A'; LeadsTo(Init,Acts) B B' |] \
-\                 ==> LeadsTo(Init,Acts) (A Un B) (A' Un B')";
+Goal "[| LeadsTo prg A A'; LeadsTo prg B B' |] \
+\                 ==> LeadsTo prg (A Un B) (A' Un B')";
 by (blast_tac (claset() addIs [LeadsTo_Un, 
 			       LeadsTo_weaken_R]) 1);
 qed "LeadsTo_Un_Un";
@@ -189,28 +186,28 @@
 
 (** The cancellation law **)
 
-Goal "[| LeadsTo(Init,Acts) A (A' Un B); LeadsTo(Init,Acts) B B'; \
-\              id: Acts |]    \
-\           ==> LeadsTo(Init,Acts) A (A' Un B')";
+Goal "[| LeadsTo prg A (A' Un B); LeadsTo prg B B'; \
+\              id: (Acts prg) |]    \
+\           ==> LeadsTo prg A (A' Un B')";
 by (blast_tac (claset() addIs [LeadsTo_Un_Un, 
 			       subset_imp_LeadsTo, LeadsTo_Trans]) 1);
 qed "LeadsTo_cancel2";
 
-Goal "[| LeadsTo(Init,Acts) A (A' Un B); LeadsTo(Init,Acts) (B-A') B'; id: Acts |] \
-\           ==> LeadsTo(Init,Acts) A (A' Un B')";
+Goal "[| LeadsTo prg A (A' Un B); LeadsTo prg (B-A') B'; id: (Acts prg) |] \
+\           ==> LeadsTo prg A (A' Un B')";
 by (rtac LeadsTo_cancel2 1);
 by (assume_tac 2);
 by (ALLGOALS Asm_simp_tac);
 qed "LeadsTo_cancel_Diff2";
 
-Goal "[| LeadsTo(Init,Acts) A (B Un A'); LeadsTo(Init,Acts) B B'; id: Acts |] \
-\           ==> LeadsTo(Init,Acts) A (B' Un A')";
+Goal "[| LeadsTo prg A (B Un A'); LeadsTo prg B B'; id: (Acts prg) |] \
+\           ==> LeadsTo prg A (B' Un A')";
 by (asm_full_simp_tac (simpset() addsimps [Un_commute]) 1);
 by (blast_tac (claset() addSIs [LeadsTo_cancel2]) 1);
 qed "LeadsTo_cancel1";
 
-Goal "[| LeadsTo(Init,Acts) A (B Un A'); LeadsTo(Init,Acts) (B-A') B'; id: Acts |] \
-\           ==> LeadsTo(Init,Acts) A (B' Un A')";
+Goal "[| LeadsTo prg A (B Un A'); LeadsTo prg (B-A') B'; id: (Acts prg) |] \
+\           ==> LeadsTo prg A (B' Un A')";
 by (rtac LeadsTo_cancel1 1);
 by (assume_tac 2);
 by (ALLGOALS Asm_simp_tac);
@@ -220,7 +217,7 @@
 
 (** The impossibility law **)
 
-Goal "LeadsTo(Init,Acts) A {} ==> reachable(Init,Acts) Int A  = {}";
+Goal "LeadsTo prg A {} ==> reachable prg Int A  = {}";
 by (full_simp_tac (simpset() addsimps [LeadsTo_def]) 1);
 by (etac leadsTo_empty 1);
 qed "LeadsTo_empty";
@@ -229,20 +226,20 @@
 (** PSP: Progress-Safety-Progress **)
 
 (*Special case of PSP: Misra's "stable conjunction".  Doesn't need id:Acts. *)
-Goal "[| LeadsTo(Init,Acts) A A'; stable Acts B |] \
-\           ==> LeadsTo(Init,Acts) (A Int B) (A' Int B)";
+Goal "[| LeadsTo prg A A'; stable (Acts prg) B |] \
+\           ==> LeadsTo prg (A Int B) (A' Int B)";
 by (asm_full_simp_tac (simpset() addsimps [LeadsTo_def, Int_assoc RS sym, 
 					   PSP_stable]) 1);
 qed "R_PSP_stable";
 
-Goal "[| LeadsTo(Init,Acts) A A'; stable Acts B |] \
-\             ==> LeadsTo(Init,Acts) (B Int A) (B Int A')";
+Goal "[| LeadsTo prg A A'; stable (Acts prg) B |] \
+\             ==> LeadsTo prg (B Int A) (B Int A')";
 by (asm_simp_tac (simpset() addsimps (R_PSP_stable::Int_ac)) 1);
 qed "R_PSP_stable2";
 
 
-Goal "[| LeadsTo(Init,Acts) A A'; constrains Acts B B'; id: Acts |] \
-\           ==> LeadsTo(Init,Acts) (A Int B) ((A' Int B) Un (B' - B))";
+Goal "[| LeadsTo prg A A'; constrains (Acts prg) B B'; id: (Acts prg) |] \
+\           ==> LeadsTo prg (A Int B) ((A' Int B) Un (B' - B))";
 by (full_simp_tac (simpset() addsimps [LeadsTo_def]) 1);
 by (dtac PSP 1);
 by (etac constrains_reachable 1);
@@ -250,14 +247,14 @@
 by (ALLGOALS Blast_tac);
 qed "R_PSP";
 
-Goal "[| LeadsTo(Init,Acts) A A'; constrains Acts B B'; id: Acts |] \
-\           ==> LeadsTo(Init,Acts) (B Int A) ((B Int A') Un (B' - B))";
+Goal "[| LeadsTo prg A A'; constrains (Acts prg) B B'; id: (Acts prg) |] \
+\           ==> LeadsTo prg (B Int A) ((B Int A') Un (B' - B))";
 by (asm_simp_tac (simpset() addsimps (R_PSP::Int_ac)) 1);
 qed "R_PSP2";
 
 Goalw [unless_def]
-   "[| LeadsTo(Init,Acts) A A'; unless Acts B B'; id: Acts |] \
-\           ==> LeadsTo(Init,Acts) (A Int B) ((A' Int B) Un B')";
+   "[| LeadsTo prg A A'; unless (Acts prg) B B'; id: (Acts prg) |] \
+\           ==> LeadsTo prg (A Int B) ((A' Int B) Un B')";
 by (dtac R_PSP 1);
 by (assume_tac 1);
 by (asm_full_simp_tac (simpset() addsimps [Un_Diff_Diff, Int_Diff_Un]) 2);
@@ -272,10 +269,10 @@
 
 (** Meta or object quantifier ????? **)
 Goal "[| wf r;     \
-\              ALL m. LeadsTo(Init,Acts) (A Int f-``{m})                     \
+\              ALL m. LeadsTo prg (A Int f-``{m})                     \
 \                                       ((A Int f-``(r^-1 ^^ {m})) Un B);   \
-\              id: Acts |] \
-\           ==> LeadsTo(Init,Acts) A B";
+\              id: (Acts prg) |] \
+\           ==> LeadsTo prg A B";
 by (full_simp_tac (simpset() addsimps [LeadsTo_def]) 1);
 by (etac leadsTo_wf_induct 1);
 by (assume_tac 2);
@@ -284,10 +281,10 @@
 
 
 Goal "[| wf r;     \
-\        ALL m:I. LeadsTo(Init,Acts) (A Int f-``{m})                   \
+\        ALL m:I. LeadsTo prg (A Int f-``{m})                   \
 \                                    ((A Int f-``(r^-1 ^^ {m})) Un B);   \
-\        id: Acts |] \
-\     ==> LeadsTo(Init,Acts) A ((A - (f-``I)) Un B)";
+\        id: (Acts prg) |] \
+\     ==> LeadsTo prg A ((A - (f-``I)) Un B)";
 by (etac LeadsTo_wf_induct 1);
 by Safe_tac;
 by (case_tac "m:I" 1);
@@ -296,29 +293,29 @@
 qed "R_bounded_induct";
 
 
-Goal "[| ALL m. LeadsTo(Init,Acts) (A Int f-``{m})                     \
-\                                  ((A Int f-``(lessThan m)) Un B);   \
-\        id: Acts |] \
-\     ==> LeadsTo(Init,Acts) A B";
+Goal "[| ALL m. LeadsTo prg (A Int f-``{m})                     \
+\                           ((A Int f-``(lessThan m)) Un B);   \
+\        id: (Acts prg) |] \
+\     ==> LeadsTo prg A B";
 by (rtac (wf_less_than RS LeadsTo_wf_induct) 1);
 by (assume_tac 2);
 by (Asm_simp_tac 1);
 qed "R_lessThan_induct";
 
-Goal "[| ALL m:(greaterThan l). LeadsTo(Init,Acts) (A Int f-``{m})   \
+Goal "[| ALL m:(greaterThan l). LeadsTo prg (A Int f-``{m})   \
 \                                        ((A Int f-``(lessThan m)) Un B);   \
-\              id: Acts |] \
-\           ==> LeadsTo(Init,Acts) A ((A Int (f-``(atMost l))) Un B)";
+\              id: (Acts prg) |] \
+\           ==> LeadsTo prg A ((A Int (f-``(atMost l))) Un B)";
 by (simp_tac (HOL_ss addsimps [Diff_eq RS sym, vimage_Compl, Compl_greaterThan RS sym]) 1);
 by (rtac (wf_less_than RS R_bounded_induct) 1);
 by (assume_tac 2);
 by (Asm_simp_tac 1);
 qed "R_lessThan_bounded_induct";
 
-Goal "[| ALL m:(lessThan l). LeadsTo(Init,Acts) (A Int f-``{m})   \
+Goal "[| ALL m:(lessThan l). LeadsTo prg (A Int f-``{m})   \
 \                              ((A Int f-``(greaterThan m)) Un B);   \
-\        id: Acts |] \
-\     ==> LeadsTo(Init,Acts) A ((A Int (f-``(atLeast l))) Un B)";
+\        id: (Acts prg) |] \
+\     ==> LeadsTo prg A ((A Int (f-``(atLeast l))) Un B)";
 by (res_inst_tac [("f","f"),("f1", "%k. l - k")]
     (wf_less_than RS wf_inv_image RS LeadsTo_wf_induct) 1);
 by (assume_tac 2);
@@ -333,19 +330,19 @@
 
 (*** Completion: Binary and General Finite versions ***)
 
-Goal "[| LeadsTo(Init,Acts) A A';  stable Acts A';   \
-\        LeadsTo(Init,Acts) B B';  stable Acts B';  id: Acts |] \
-\     ==> LeadsTo(Init,Acts) (A Int B) (A' Int B')";
+Goal "[| LeadsTo prg A A';  stable (Acts prg) A';   \
+\        LeadsTo prg B B';  stable (Acts prg) B';  id: (Acts prg) |] \
+\     ==> LeadsTo prg (A Int B) (A' Int B')";
 by (full_simp_tac (simpset() addsimps [LeadsTo_def]) 1);
 by (blast_tac (claset() addIs [stable_completion RS leadsTo_weaken] 
                         addSIs [stable_Int, stable_reachable]) 1);
 qed "R_stable_completion";
 
 
-Goal "[| finite I;  id: Acts |]                     \
-\     ==> (ALL i:I. LeadsTo(Init,Acts) (A i) (A' i)) -->  \
-\         (ALL i:I. stable Acts (A' i)) -->         \
-\         LeadsTo(Init,Acts) (INT i:I. A i) (INT i:I. A' i)";
+Goal "[| finite I;  id: (Acts prg) |]                     \
+\     ==> (ALL i:I. LeadsTo prg (A i) (A' i)) -->  \
+\         (ALL i:I. stable (Acts prg) (A' i)) -->         \
+\         LeadsTo prg (INT i:I. A i) (INT i:I. A' i)";
 by (etac finite_induct 1);
 by (Asm_simp_tac 1);
 by (asm_simp_tac 
@@ -354,10 +351,10 @@
 qed_spec_mp "R_finite_stable_completion";
 
 
-Goal "[| LeadsTo(Init,Acts) A (A' Un C);  constrains Acts A' (A' Un C); \
-\        LeadsTo(Init,Acts) B (B' Un C);  constrains Acts B' (B' Un C); \
-\        id: Acts |] \
-\     ==> LeadsTo(Init,Acts) (A Int B) ((A' Int B') Un C)";
+Goal "[| LeadsTo prg A (A' Un C);  constrains (Acts prg) A' (A' Un C); \
+\        LeadsTo prg B (B' Un C);  constrains (Acts prg) B' (B' Un C); \
+\        id: (Acts prg) |] \
+\     ==> LeadsTo prg (A Int B) ((A' Int B') Un C)";
 by (full_simp_tac (simpset() addsimps [LeadsTo_def, Int_Un_distrib]) 1);
 by (dtac completion 1);
 by (assume_tac 2);
@@ -368,10 +365,10 @@
 qed "R_completion";
 
 
-Goal "[| finite I;  id: Acts |] \
-\     ==> (ALL i:I. LeadsTo(Init,Acts) (A i) (A' i Un C)) -->  \
-\         (ALL i:I. constrains Acts (A' i) (A' i Un C)) --> \
-\         LeadsTo(Init,Acts) (INT i:I. A i) ((INT i:I. A' i) Un C)";
+Goal "[| finite I;  id: (Acts prg) |] \
+\     ==> (ALL i:I. LeadsTo prg (A i) (A' i Un C)) -->  \
+\         (ALL i:I. constrains (Acts prg) (A' i) (A' i Un C)) --> \
+\         LeadsTo prg (INT i:I. A i) ((INT i:I. A' i) Un C)";
 by (etac finite_induct 1);
 by (ALLGOALS Asm_simp_tac);
 by (Clarify_tac 1);
@@ -384,32 +381,37 @@
 (*** Specialized laws for handling invariants ***)
 
 Goalw [transient_def]
-    "[| reachable(Init,Acts) <= INV;  transient Acts (INV Int A) |]  \
-\    ==> transient Acts (reachable(Init,Acts) Int A)";
+    "[| reachable prg <= INV;  transient (Acts prg) (INV Int A) |]  \
+\    ==> transient (Acts prg) (reachable prg Int A)";
 by (Clarify_tac 1);
 by (rtac bexI 1); 
 by (assume_tac 2);
 by (Blast_tac 1);
 qed "reachable_transient";
 
-(*Uses the premise "invariant (Init,Acts)".  Raw theorem-proving on this
+(*Uses the premise "invariant prg".  Raw theorem-proving on this
   inclusion is slow: the term contains a copy of the program.*)
-Goal "[| invariant (Init,Acts) INV;      \
-\        constrains Acts (INV Int (A - A')) (A Un A'); \
-\        transient  Acts (INV Int (A-A')) |]   \
-\     ==> LeadsTo(Init,Acts) A A'";
+Goal "[| invariant prg INV;      \
+\        constrains (Acts prg) (INV Int (A-A')) (A Un A'); \
+\        transient  (Acts prg) (INV Int (A-A')) |]   \
+\ ==> ensures (Acts prg) (reachable prg Int A) (reachable prg Int A')";
 bd invariant_includes_reachable 1;
-by (rtac LeadsTo_Basis 1);
+by (rtac ensuresI 1);
+by (ALLGOALS 
+    (full_simp_tac (simpset() addsimps [Int_Un_distrib RS sym, 
+					Diff_Int_distrib RS sym])));
 by (blast_tac (claset() addSIs [reachable_transient]) 2);
-by (rewtac constrains_def);
-by (Blast_tac 1);
-qed "invariant_LeadsTo_Basis";
+br (stable_reachable RS stable_constrains_Int) 1;
+by (blast_tac (claset() addIs [constrains_weaken]) 1);
+qed "invariant_ensuresI";
+
+bind_thm ("invariant_LeadsTo_Basis", invariant_ensuresI RS LeadsTo_Basis);
 
 
-Goal "[| invariant (Init,Acts) INV;      \
-\        LeadsTo(Init,Acts) A A'; id: Acts;   \
+Goal "[| invariant prg INV;      \
+\        LeadsTo prg A A'; id: (Acts prg);   \
 \        INV Int B  <= A;  INV Int A' <= B' |] \
-\     ==> LeadsTo(Init,Acts) B B'";
+\     ==> LeadsTo prg B B'";
 by (blast_tac (claset() addDs [invariant_includes_reachable]
 			addIs [LeadsTo_weaken]) 1);
 qed "invariant_LeadsTo_weaken";
@@ -425,8 +427,8 @@
    SELECT_GOAL
       (EVERY [TRY (rtac stableI 1),
 	      rtac constrainsI 1,
-	      rewtac main_def,
-	      REPEAT_FIRST (eresolve_tac [insertE, emptyE]),
+	      full_simp_tac (simpset() addsimps [main_def]) 1,
+	      REPEAT_FIRST (eresolve_tac [disjE]),
 	      rewrite_goals_tac cmd_defs,
 	      ALLGOALS (SELECT_GOAL Auto_tac)]);