src/HOL/Auth/Guard/P1.thy
changeset 71989 bad75618fb82
parent 67613 ce654b0e6d69
child 76287 cdc14f94c754
--- a/src/HOL/Auth/Guard/P1.thy	Thu Jul 02 08:49:04 2020 +0000
+++ b/src/HOL/Auth/Guard/P1.thy	Thu Jul 02 12:10:58 2020 +0000
@@ -282,6 +282,7 @@
 
 lemma insertion_resilience: "\<forall>L. L \<in> valid A n B \<longrightarrow> Suc i < len L
 \<longrightarrow> ins (L,Suc i,M) \<notin> valid A n B"
+supply [[simproc del: defined_all]]
 apply (induct i)
 (* i = 0 *)
 apply clarify
@@ -481,6 +482,7 @@
 
 lemma pro_imp_Guard [rule_format]: "[| evs \<in> p1; B \<notin> bad; A \<notin> bad |] ==>
 pro B ofr A r I (cons M L) J C \<in> set evs \<longrightarrow> Guard ofr {priK A} (spies evs)"
+supply [[simproc del: defined_all]]
 apply (erule p1.induct) (* +3 subgoals *)
 (* Nil *)
 apply simp