src/HOL/IOA/IOA.ML
changeset 4799 82b0ed20c2cb
parent 4772 8c7e7eaffbdf
child 4828 ee3317896a47
--- a/src/HOL/IOA/IOA.ML	Sat Apr 04 14:30:19 1998 +0200
+++ b/src/HOL/IOA/IOA.ML	Tue Apr 07 13:43:07 1998 +0200
@@ -55,7 +55,7 @@
 goalw IOA.thy (reachable_def::exec_rws)
 "!!A. [| reachable A s; (s,a,t) : trans_of(A) |] ==> reachable A t";
   by (asm_full_simp_tac (simpset() delsimps bex_simps) 1);
-  by Safe_tac;
+  by (safe_tac (claset() addSaltern ("split_all_tac", split_all_tac)));
   by (rename_tac "n ex1 ex2 a1 a4 a5 a2 a3" 1);
   by (res_inst_tac [("x","(%i. if i<n then ex1 i                    \
 \                            else (if i=n then Some a else None),    \
@@ -78,7 +78,7 @@
 \     !!s t a. [|reachable A s; P(s)|] ==> (s,a,t): trans_of(A) --> P(t) |] \
 \  ==> invariant A P";
   by (rewrite_goals_tac(reachable_def::Let_def::exec_rws));
-  by Safe_tac;
+  by (safe_tac (claset() addSaltern ("split_all_tac", split_all_tac)));
   by (rename_tac "n ex1 ex2" 1);
   by (res_inst_tac [("Q","reachable A (ex2 n)")] conjunct1 1);
   by (Asm_full_simp_tac 1);