src/HOL/Auth/Guard/Extensions.thy
changeset 58305 57752a91eec4
parent 56681 e8d5d60d655e
child 58860 fee7cfa69c50
--- a/src/HOL/Auth/Guard/Extensions.thy	Thu Sep 11 18:54:36 2014 +0200
+++ b/src/HOL/Auth/Guard/Extensions.thy	Thu Sep 11 18:54:36 2014 +0200
@@ -579,6 +579,7 @@
 apply (drule_tac P="%G. X:parts G" in knows_Cons_substD, safe)
 apply (erule initState_used)
 apply (case_tac a, auto)
+apply (rename_tac msg)
 apply (drule_tac B=A and X=msg and evs=evs in Gets_correct_Says)
 by (auto dest: Says_imp_used intro: used_ConsI)
 
@@ -592,6 +593,7 @@
 apply (frule_tac ev=a and evs=evs in one_step_Cons, simp, clarify)
 apply (drule_tac P="%G. X:parts G" in knows_max_Cons_substD, safe)
 apply (case_tac a, auto)
+apply (rename_tac msg)
 apply (drule_tac B=A and X=msg and evs=evs in Gets_correct_Says)
 by (auto simp: knows_max'_Cons dest: Says_imp_used intro: used_ConsI)