--- a/src/HOL/Auth/Public.thy Thu Jul 23 18:44:08 2009 +0200
+++ b/src/HOL/Auth/Public.thy Thu Jul 23 18:44:09 2009 +0200
@@ -407,7 +407,7 @@
(*Tactic for possibility theorems*)
fun possibility_tac ctxt =
REPEAT (*omit used_Says so that Nonces start from different traces!*)
- (ALLGOALS (simp_tac (local_simpset_of ctxt delsimps [@{thm used_Says}]))
+ (ALLGOALS (simp_tac (simpset_of ctxt delsimps [@{thm used_Says}]))
THEN
REPEAT_FIRST (eq_assume_tac ORELSE'
resolve_tac [refl, conjI, @{thm Nonce_supply}]))
@@ -416,7 +416,7 @@
nonces and keys initially*)
fun basic_possibility_tac ctxt =
REPEAT
- (ALLGOALS (asm_simp_tac (local_simpset_of ctxt setSolver safe_solver))
+ (ALLGOALS (asm_simp_tac (simpset_of ctxt setSolver safe_solver))
THEN
REPEAT_FIRST (resolve_tac [refl, conjI]))