author | paulson <lp15@cam.ac.uk> |
Thu, 13 Oct 2022 15:49:09 +0100 | |
changeset 76288 | b82ac7ef65ec |
parent 76287 | cdc14f94c754 |
permissions | -rw-r--r-- |
41775 | 1 |
(* Title: HOL/Auth/Guard/P2.thy |
2 |
Author: Frederic Blanqui, University of Cambridge Computer Laboratory |
|
3 |
Copyright 2002 University of Cambridge |
|
13508 | 4 |
|
41775 | 5 |
From G. Karjoth, N. Asokan and C. Gulcu |
6 |
"Protecting the computation results of free-roaming agents" |
|
7 |
Mobiles Agents 1998, LNCS 1477. |
|
8 |
*) |
|
13508 | 9 |
|
61830 | 10 |
section\<open>Protocol P2\<close> |
13508 | 11 |
|
16417 | 12 |
theory P2 imports Guard_Public List_Msg begin |
13508 | 13 |
|
61830 | 14 |
subsection\<open>Protocol Definition\<close> |
13508 | 15 |
|
16 |
||
61830 | 17 |
text\<open>Like P1 except the definitions of \<open>chain\<close>, \<open>shop\<close>, |
18 |
\<open>next_shop\<close> and \<open>nonce\<close>\<close> |
|
13508 | 19 |
|
61830 | 20 |
subsubsection\<open>offer chaining: |
21 |
B chains his offer for A with the head offer of L for sending it to C\<close> |
|
13508 | 22 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
23746
diff
changeset
|
23 |
definition chain :: "agent => nat => agent => msg => agent => msg" where |
13508 | 24 |
"chain B ofr A L C == |
25 |
let m1= sign B (Nonce ofr) in |
|
61956 | 26 |
let m2= Hash \<lbrace>head L, Agent C\<rbrace> in |
27 |
\<lbrace>Crypt (pubK A) m1, m2\<rbrace>" |
|
13508 | 28 |
|
29 |
declare Let_def [simp] |
|
30 |
||
31 |
lemma chain_inj [iff]: "(chain B ofr A L C = chain B' ofr' A' L' C') |
|
32 |
= (B=B' & ofr=ofr' & A=A' & head L = head L' & C=C')" |
|
33 |
by (auto simp: chain_def Let_def) |
|
34 |
||
67613 | 35 |
lemma Nonce_in_chain [iff]: "Nonce ofr \<in> parts {chain B ofr A L C}" |
13508 | 36 |
by (auto simp: chain_def sign_def) |
37 |
||
61830 | 38 |
subsubsection\<open>agent whose key is used to sign an offer\<close> |
13508 | 39 |
|
35418 | 40 |
fun shop :: "msg => msg" where |
61956 | 41 |
"shop \<lbrace>Crypt K \<lbrace>B,ofr,Crypt K' H\<rbrace>,m2\<rbrace> = Agent (agt K')" |
13508 | 42 |
|
43 |
lemma shop_chain [simp]: "shop (chain B ofr A L C) = Agent B" |
|
44 |
by (simp add: chain_def sign_def) |
|
45 |
||
61830 | 46 |
subsubsection\<open>nonce used in an offer\<close> |
13508 | 47 |
|
35418 | 48 |
fun nonce :: "msg => msg" where |
61956 | 49 |
"nonce \<lbrace>Crypt K \<lbrace>B,ofr,CryptH\<rbrace>,m2\<rbrace> = ofr" |
13508 | 50 |
|
51 |
lemma nonce_chain [simp]: "nonce (chain B ofr A L C) = Nonce ofr" |
|
52 |
by (simp add: chain_def sign_def) |
|
53 |
||
61830 | 54 |
subsubsection\<open>next shop\<close> |
13508 | 55 |
|
35418 | 56 |
fun next_shop :: "msg => agent" where |
61956 | 57 |
"next_shop \<lbrace>m1,Hash \<lbrace>headL,Agent C\<rbrace>\<rbrace> = C" |
13508 | 58 |
|
59 |
lemma "next_shop (chain B ofr A L C) = C" |
|
60 |
by (simp add: chain_def sign_def) |
|
61 |
||
61830 | 62 |
subsubsection\<open>anchor of the offer list\<close> |
13508 | 63 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
23746
diff
changeset
|
64 |
definition anchor :: "agent => nat => agent => msg" where |
13508 | 65 |
"anchor A n B == chain A n A (cons nil nil) B" |
66 |
||
67 |
lemma anchor_inj [iff]: |
|
67613 | 68 |
"(anchor A n B = anchor A' n' B') = (A=A' \<and> n=n' \<and> B=B')" |
13508 | 69 |
by (auto simp: anchor_def) |
70 |
||
67613 | 71 |
lemma Nonce_in_anchor [iff]: "Nonce n \<in> parts {anchor A n B}" |
13508 | 72 |
by (auto simp: anchor_def) |
73 |
||
74 |
lemma shop_anchor [simp]: "shop (anchor A n B) = Agent A" |
|
75 |
by (simp add: anchor_def) |
|
76 |
||
61830 | 77 |
subsubsection\<open>request event\<close> |
13508 | 78 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
23746
diff
changeset
|
79 |
definition reqm :: "agent => nat => nat => msg => agent => msg" where |
61956 | 80 |
"reqm A r n I B == \<lbrace>Agent A, Number r, cons (Agent A) (cons (Agent B) I), |
81 |
cons (anchor A n B) nil\<rbrace>" |
|
13508 | 82 |
|
83 |
lemma reqm_inj [iff]: "(reqm A r n I B = reqm A' r' n' I' B') |
|
84 |
= (A=A' & r=r' & n=n' & I=I' & B=B')" |
|
85 |
by (auto simp: reqm_def) |
|
86 |
||
67613 | 87 |
lemma Nonce_in_reqm [iff]: "Nonce n \<in> parts {reqm A r n I B}" |
13508 | 88 |
by (auto simp: reqm_def) |
89 |
||
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
23746
diff
changeset
|
90 |
definition req :: "agent => nat => nat => msg => agent => event" where |
13508 | 91 |
"req A r n I B == Says A B (reqm A r n I B)" |
92 |
||
93 |
lemma req_inj [iff]: "(req A r n I B = req A' r' n' I' B') |
|
94 |
= (A=A' & r=r' & n=n' & I=I' & B=B')" |
|
95 |
by (auto simp: req_def) |
|
96 |
||
61830 | 97 |
subsubsection\<open>propose event\<close> |
13508 | 98 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
23746
diff
changeset
|
99 |
definition prom :: "agent => nat => agent => nat => msg => msg => |
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
23746
diff
changeset
|
100 |
msg => agent => msg" where |
61956 | 101 |
"prom B ofr A r I L J C == \<lbrace>Agent A, Number r, |
102 |
app (J, del (Agent B, I)), cons (chain B ofr A L C) L\<rbrace>" |
|
13508 | 103 |
|
104 |
lemma prom_inj [dest]: "prom B ofr A r I L J C = prom B' ofr' A' r' I' L' J' C' |
|
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
105 |
\<Longrightarrow> B=B' & ofr=ofr' & A=A' & r=r' & L=L' & C=C'" |
13508 | 106 |
by (auto simp: prom_def) |
107 |
||
67613 | 108 |
lemma Nonce_in_prom [iff]: "Nonce ofr \<in> parts {prom B ofr A r I L J C}" |
13508 | 109 |
by (auto simp: prom_def) |
110 |
||
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
23746
diff
changeset
|
111 |
definition pro :: "agent => nat => agent => nat => msg => msg => |
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
23746
diff
changeset
|
112 |
msg => agent => event" where |
13508 | 113 |
"pro B ofr A r I L J C == Says B C (prom B ofr A r I L J C)" |
114 |
||
115 |
lemma pro_inj [dest]: "pro B ofr A r I L J C = pro B' ofr' A' r' I' L' J' C' |
|
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
116 |
\<Longrightarrow> B=B' & ofr=ofr' & A=A' & r=r' & L=L' & C=C'" |
13508 | 117 |
by (auto simp: pro_def dest: prom_inj) |
118 |
||
61830 | 119 |
subsubsection\<open>protocol\<close> |
13508 | 120 |
|
23746 | 121 |
inductive_set p2 :: "event list set" |
122 |
where |
|
13508 | 123 |
|
67613 | 124 |
Nil: "[] \<in> p2" |
13508 | 125 |
|
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
126 |
| Fake: "\<lbrakk>evsf \<in> p2; X \<in> synth (analz (spies evsf))\<rbrakk> \<Longrightarrow> Says Spy B X # evsf \<in> p2" |
13508 | 127 |
|
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
128 |
| Request: "\<lbrakk>evsr \<in> p2; Nonce n \<notin> used evsr; I \<in> agl\<rbrakk> \<Longrightarrow> req A r n I B # evsr \<in> p2" |
13508 | 129 |
|
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
130 |
| Propose: "\<lbrakk>evsp \<in> p2; Says A' B \<lbrace>Agent A,Number r,I,cons M L\<rbrace> \<in> set evsp; |
67613 | 131 |
I \<in> agl; J \<in> agl; isin (Agent C, app (J, del (Agent B, I))); |
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
132 |
Nonce ofr \<notin> used evsp\<rbrakk> \<Longrightarrow> pro B ofr A r I (cons M L) J C # evsp \<in> p2" |
13508 | 133 |
|
61830 | 134 |
subsubsection\<open>valid offer lists\<close> |
13508 | 135 |
|
23746 | 136 |
inductive_set |
67613 | 137 |
valid :: "agent \<Rightarrow> nat \<Rightarrow> agent \<Rightarrow> msg set" |
23746 | 138 |
for A :: agent and n :: nat and B :: agent |
139 |
where |
|
67613 | 140 |
Request [intro]: "cons (anchor A n B) nil \<in> valid A n B" |
13508 | 141 |
|
67613 | 142 |
| Propose [intro]: "L \<in> valid A n B |
143 |
\<Longrightarrow> cons (chain (next_shop (head L)) ofr A L C) L \<in> valid A n B" |
|
13508 | 144 |
|
61830 | 145 |
subsubsection\<open>basic properties of valid\<close> |
13508 | 146 |
|
67613 | 147 |
lemma valid_not_empty: "L \<in> valid A n B \<Longrightarrow> \<exists>M L'. L = cons M L'" |
13508 | 148 |
by (erule valid.cases, auto) |
149 |
||
67613 | 150 |
lemma valid_pos_len: "L \<in> valid A n B \<Longrightarrow> 0 < len L" |
13508 | 151 |
by (erule valid.induct, auto) |
152 |
||
61830 | 153 |
subsubsection\<open>list of offers\<close> |
13508 | 154 |
|
67613 | 155 |
fun offers :: "msg \<Rightarrow> msg" |
35418 | 156 |
where |
61956 | 157 |
"offers (cons M L) = cons \<lbrace>shop M, nonce M\<rbrace> (offers L)" |
35418 | 158 |
| "offers other = nil" |
13508 | 159 |
|
160 |
||
61830 | 161 |
subsection\<open>Properties of Protocol P2\<close> |
13508 | 162 |
|
61830 | 163 |
text\<open>same as \<open>P1_Prop\<close> except that publicly verifiable forward |
164 |
integrity is replaced by forward privacy\<close> |
|
13508 | 165 |
|
61830 | 166 |
subsection\<open>strong forward integrity: |
167 |
except the last one, no offer can be modified\<close> |
|
13508 | 168 |
|
67613 | 169 |
lemma strong_forward_integrity: "\<forall>L. Suc i < len L |
170 |
\<longrightarrow> L \<in> valid A n B \<longrightarrow> repl (L,Suc i,M) \<in> valid A n B \<longrightarrow> M = ith (L,Suc i)" |
|
13508 | 171 |
apply (induct i) |
172 |
(* i = 0 *) |
|
173 |
apply clarify |
|
174 |
apply (frule len_not_empty, clarsimp) |
|
175 |
apply (frule len_not_empty, clarsimp) |
|
67613 | 176 |
apply (ind_cases "\<lbrace>x,xa,l'a\<rbrace> \<in> valid A n B" for x xa l'a) |
177 |
apply (ind_cases "\<lbrace>x,M,l'a\<rbrace> \<in> valid A n B" for x l'a) |
|
13508 | 178 |
apply (simp add: chain_def) |
179 |
(* i > 0 *) |
|
180 |
apply clarify |
|
181 |
apply (frule len_not_empty, clarsimp) |
|
67613 | 182 |
apply (ind_cases "\<lbrace>x,repl(l',Suc na,M)\<rbrace> \<in> valid A n B" for x l' na) |
13508 | 183 |
apply (frule len_not_empty, clarsimp) |
67613 | 184 |
apply (ind_cases "\<lbrace>x,l'\<rbrace> \<in> valid A n B" for x l') |
13508 | 185 |
by (drule_tac x=l' in spec, simp, blast) |
186 |
||
61830 | 187 |
subsection\<open>insertion resilience: |
188 |
except at the beginning, no offer can be inserted\<close> |
|
13508 | 189 |
|
67613 | 190 |
lemma chain_isnt_head [simp]: "L \<in> valid A n B \<Longrightarrow> |
191 |
head L \<noteq> chain (next_shop (head L)) ofr A L C" |
|
13508 | 192 |
by (erule valid.induct, auto simp: chain_def sign_def anchor_def) |
193 |
||
67613 | 194 |
lemma insertion_resilience: "\<forall>L. L \<in> valid A n B \<longrightarrow> Suc i < len L |
195 |
\<longrightarrow> ins (L,Suc i,M) \<notin> valid A n B" |
|
71989
bad75618fb82
extraction of equations x = t from premises beneath meta-all
haftmann
parents:
67613
diff
changeset
|
196 |
supply [[simproc del: defined_all]] |
13508 | 197 |
apply (induct i) |
198 |
(* i = 0 *) |
|
199 |
apply clarify |
|
200 |
apply (frule len_not_empty, clarsimp) |
|
67613 | 201 |
apply (ind_cases "\<lbrace>x,l'\<rbrace> \<in> valid A n B" for x l', simp) |
202 |
apply (ind_cases "\<lbrace>x,M,l'\<rbrace> \<in> valid A n B" for x l', clarsimp) |
|
203 |
apply (ind_cases "\<lbrace>head l',l'\<rbrace> \<in> valid A n B" for l', simp, simp) |
|
13508 | 204 |
(* i > 0 *) |
205 |
apply clarify |
|
206 |
apply (frule len_not_empty, clarsimp) |
|
67613 | 207 |
apply (ind_cases "\<lbrace>x,l'\<rbrace> \<in> valid A n B" for x l') |
13508 | 208 |
apply (frule len_not_empty, clarsimp) |
67613 | 209 |
apply (ind_cases "\<lbrace>x,ins(l',Suc na,M)\<rbrace> \<in> valid A n B" for x l' na) |
13508 | 210 |
apply (frule len_not_empty, clarsimp) |
211 |
by (drule_tac x=l' in spec, clarsimp) |
|
212 |
||
61830 | 213 |
subsection\<open>truncation resilience: |
214 |
only shop i can truncate at offer i\<close> |
|
13508 | 215 |
|
67613 | 216 |
lemma truncation_resilience: "\<forall>L. L \<in> valid A n B \<longrightarrow> Suc i < len L |
217 |
\<longrightarrow> cons M (trunc (L,Suc i)) \<in> valid A n B \<longrightarrow> shop M = shop (ith (L,i))" |
|
13508 | 218 |
apply (induct i) |
219 |
(* i = 0 *) |
|
220 |
apply clarify |
|
221 |
apply (frule len_not_empty, clarsimp) |
|
67613 | 222 |
apply (ind_cases "\<lbrace>x,l'\<rbrace> \<in> valid A n B" for x l') |
13508 | 223 |
apply (frule len_not_empty, clarsimp) |
67613 | 224 |
apply (ind_cases "\<lbrace>M,l'\<rbrace> \<in> valid A n B" for l') |
13508 | 225 |
apply (frule len_not_empty, clarsimp, simp) |
226 |
(* i > 0 *) |
|
227 |
apply clarify |
|
228 |
apply (frule len_not_empty, clarsimp) |
|
67613 | 229 |
apply (ind_cases "\<lbrace>x,l'\<rbrace> \<in> valid A n B" for x l') |
13508 | 230 |
apply (frule len_not_empty, clarsimp) |
231 |
by (drule_tac x=l' in spec, clarsimp) |
|
232 |
||
61830 | 233 |
subsection\<open>declarations for tactics\<close> |
13508 | 234 |
|
235 |
declare knows_Spy_partsEs [elim] |
|
236 |
declare Fake_parts_insert [THEN subsetD, dest] |
|
237 |
declare initState.simps [simp del] |
|
238 |
||
61830 | 239 |
subsection\<open>get components of a message\<close> |
13508 | 240 |
|
67613 | 241 |
lemma get_ML [dest]: "Says A' B \<lbrace>A,R,I,M,L\<rbrace> \<in> set evs \<Longrightarrow> |
242 |
M \<in> parts (spies evs) \<and> L \<in> parts (spies evs)" |
|
13508 | 243 |
by blast |
244 |
||
61830 | 245 |
subsection\<open>general properties of p2\<close> |
13508 | 246 |
|
247 |
lemma reqm_neq_prom [iff]: |
|
67613 | 248 |
"reqm A r n I B \<noteq> prom B' ofr A' r' I' (cons M L) J C" |
13508 | 249 |
by (auto simp: reqm_def prom_def) |
250 |
||
251 |
lemma prom_neq_reqm [iff]: |
|
67613 | 252 |
"prom B' ofr A' r' I' (cons M L) J C \<noteq> reqm A r n I B" |
13508 | 253 |
by (auto simp: reqm_def prom_def) |
254 |
||
67613 | 255 |
lemma req_neq_pro [iff]: "req A r n I B \<noteq> pro B' ofr A' r' I' (cons M L) J C" |
13508 | 256 |
by (auto simp: req_def pro_def) |
257 |
||
67613 | 258 |
lemma pro_neq_req [iff]: "pro B' ofr A' r' I' (cons M L) J C \<noteq> req A r n I B" |
13508 | 259 |
by (auto simp: req_def pro_def) |
260 |
||
67613 | 261 |
lemma p2_has_no_Gets: "evs \<in> p2 \<Longrightarrow> \<forall>A X. Gets A X \<notin> set evs" |
13508 | 262 |
by (erule p2.induct, auto simp: req_def pro_def) |
263 |
||
264 |
lemma p2_is_Gets_correct [iff]: "Gets_correct p2" |
|
265 |
by (auto simp: Gets_correct_def dest: p2_has_no_Gets) |
|
266 |
||
267 |
lemma p2_is_one_step [iff]: "one_step p2" |
|
76288
b82ac7ef65ec
Removal of the "unfold" method in favour of "unfolding"
paulson <lp15@cam.ac.uk>
parents:
76287
diff
changeset
|
268 |
unfolding one_step_def by (clarify, ind_cases "ev#evs \<in> p2" for ev evs, auto) |
13508 | 269 |
|
67613 | 270 |
lemma p2_has_only_Says' [rule_format]: "evs \<in> p2 \<Longrightarrow> |
271 |
ev \<in> set evs \<longrightarrow> (\<exists>A B X. ev=Says A B X)" |
|
13508 | 272 |
by (erule p2.induct, auto simp: req_def pro_def) |
273 |
||
274 |
lemma p2_has_only_Says [iff]: "has_only_Says p2" |
|
275 |
by (auto simp: has_only_Says_def dest: p2_has_only_Says') |
|
276 |
||
277 |
lemma p2_is_regular [iff]: "regular p2" |
|
278 |
apply (simp only: regular_def, clarify) |
|
279 |
apply (erule_tac p2.induct) |
|
280 |
apply (simp_all add: initState.simps knows.simps pro_def prom_def |
|
281 |
req_def reqm_def anchor_def chain_def sign_def) |
|
282 |
by (auto dest: no_Key_in_agl no_Key_in_appdel parts_trans) |
|
283 |
||
61830 | 284 |
subsection\<open>private keys are safe\<close> |
13508 | 285 |
|
286 |
lemma priK_parts_Friend_imp_bad [rule_format,dest]: |
|
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
287 |
"\<lbrakk>evs \<in> p2; Friend B \<noteq> A\<rbrakk> |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
288 |
\<Longrightarrow> (Key (priK A) \<in> parts (knows (Friend B) evs)) \<longrightarrow> (A \<in> bad)" |
13508 | 289 |
apply (erule p2.induct) |
290 |
apply (simp_all add: initState.simps knows.simps pro_def prom_def |
|
17778 | 291 |
req_def reqm_def anchor_def chain_def sign_def) |
13508 | 292 |
apply (blast dest: no_Key_in_agl) |
293 |
apply (auto del: parts_invKey disjE dest: parts_trans |
|
294 |
simp add: no_Key_in_appdel) |
|
295 |
done |
|
296 |
||
297 |
lemma priK_analz_Friend_imp_bad [rule_format,dest]: |
|
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
298 |
"\<lbrakk>evs \<in> p2; Friend B \<noteq> A\<rbrakk> |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
299 |
\<Longrightarrow> (Key (priK A) \<in> analz (knows (Friend B) evs)) \<longrightarrow> (A \<in> bad)" |
13508 | 300 |
by auto |
301 |
||
302 |
lemma priK_notin_knows_max_Friend: |
|
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
303 |
"\<lbrakk>evs \<in> p2; A \<notin> bad; A \<noteq> Friend C\<rbrakk> |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
304 |
\<Longrightarrow> Key (priK A) \<notin> analz (knows_max (Friend C) evs)" |
13508 | 305 |
apply (rule not_parts_not_analz, simp add: knows_max_def, safe) |
306 |
apply (drule_tac H="spies' evs" in parts_sub) |
|
307 |
apply (rule_tac p=p2 in knows_max'_sub_spies', simp+) |
|
308 |
apply (drule_tac H="spies evs" in parts_sub) |
|
309 |
by (auto dest: knows'_sub_knows [THEN subsetD] priK_notin_initState_Friend) |
|
310 |
||
61830 | 311 |
subsection\<open>general guardedness properties\<close> |
13508 | 312 |
|
67613 | 313 |
lemma agl_guard [intro]: "I \<in> agl \<Longrightarrow> I \<in> guard n Ks" |
13508 | 314 |
by (erule agl.induct, auto) |
315 |
||
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
316 |
lemma Says_to_knows_max'_guard: "\<lbrakk>Says A' C \<lbrace>A'',r,I,L\<rbrace> \<in> set evs; |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
317 |
Guard n Ks (knows_max' C evs)\<rbrakk> \<Longrightarrow> L \<in> guard n Ks" |
13508 | 318 |
by (auto dest: Says_to_knows_max') |
319 |
||
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
320 |
lemma Says_from_knows_max'_guard: "\<lbrakk>Says C A' \<lbrace>A'',r,I,L\<rbrace> \<in> set evs; |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
321 |
Guard n Ks (knows_max' C evs)\<rbrakk> \<Longrightarrow> L \<in> guard n Ks" |
13508 | 322 |
by (auto dest: Says_from_knows_max') |
323 |
||
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
324 |
lemma Says_Nonce_not_used_guard: "\<lbrakk>Says A' B \<lbrace>A'',r,I,L\<rbrace> \<in> set evs; |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
325 |
Nonce n \<notin> used evs\<rbrakk> \<Longrightarrow> L \<in> guard n Ks" |
13508 | 326 |
by (drule not_used_not_parts, auto) |
327 |
||
61830 | 328 |
subsection\<open>guardedness of messages\<close> |
13508 | 329 |
|
67613 | 330 |
lemma chain_guard [iff]: "chain B ofr A L C \<in> guard n {priK A}" |
13508 | 331 |
by (case_tac "ofr=n", auto simp: chain_def sign_def) |
332 |
||
67613 | 333 |
lemma chain_guard_Nonce_neq [intro]: "n \<noteq> ofr |
334 |
\<Longrightarrow> chain B ofr A' L C \<in> guard n {priK A}" |
|
13508 | 335 |
by (auto simp: chain_def sign_def) |
336 |
||
67613 | 337 |
lemma anchor_guard [iff]: "anchor A n' B \<in> guard n {priK A}" |
13508 | 338 |
by (case_tac "n'=n", auto simp: anchor_def) |
339 |
||
67613 | 340 |
lemma anchor_guard_Nonce_neq [intro]: "n \<noteq> n' |
341 |
\<Longrightarrow> anchor A' n' B \<in> guard n {priK A}" |
|
13508 | 342 |
by (auto simp: anchor_def) |
343 |
||
67613 | 344 |
lemma reqm_guard [intro]: "I \<in> agl \<Longrightarrow> reqm A r n' I B \<in> guard n {priK A}" |
13508 | 345 |
by (case_tac "n'=n", auto simp: reqm_def) |
346 |
||
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
347 |
lemma reqm_guard_Nonce_neq [intro]: "\<lbrakk>n \<noteq> n'; I \<in> agl\<rbrakk> |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
348 |
\<Longrightarrow> reqm A' r n' I B \<in> guard n {priK A}" |
13508 | 349 |
by (auto simp: reqm_def) |
350 |
||
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
351 |
lemma prom_guard [intro]: "\<lbrakk>I \<in> agl; J \<in> agl; L \<in> guard n {priK A}\<rbrakk> |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
352 |
\<Longrightarrow> prom B ofr A r I L J C \<in> guard n {priK A}" |
13508 | 353 |
by (auto simp: prom_def) |
354 |
||
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
355 |
lemma prom_guard_Nonce_neq [intro]: "\<lbrakk>n \<noteq> ofr; I \<in> agl; J \<in> agl; |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
356 |
L \<in> guard n {priK A}\<rbrakk> \<Longrightarrow> prom B ofr A' r I L J C \<in> guard n {priK A}" |
13508 | 357 |
by (auto simp: prom_def) |
358 |
||
61830 | 359 |
subsection\<open>Nonce uniqueness\<close> |
13508 | 360 |
|
67613 | 361 |
lemma uniq_Nonce_in_chain [dest]: "Nonce k \<in> parts {chain B ofr A L C} \<Longrightarrow> k=ofr" |
13508 | 362 |
by (auto simp: chain_def sign_def) |
363 |
||
67613 | 364 |
lemma uniq_Nonce_in_anchor [dest]: "Nonce k \<in> parts {anchor A n B} \<Longrightarrow> k=n" |
13508 | 365 |
by (auto simp: anchor_def chain_def sign_def) |
366 |
||
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
367 |
lemma uniq_Nonce_in_reqm [dest]: "\<lbrakk>Nonce k \<in> parts {reqm A r n I B}; |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
368 |
I \<in> agl\<rbrakk> \<Longrightarrow> k=n" |
13508 | 369 |
by (auto simp: reqm_def dest: no_Nonce_in_agl) |
370 |
||
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
371 |
lemma uniq_Nonce_in_prom [dest]: "\<lbrakk>Nonce k \<in> parts {prom B ofr A r I L J C}; |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
372 |
I \<in> agl; J \<in> agl; Nonce k \<notin> parts {L}\<rbrakk> \<Longrightarrow> k=ofr" |
13508 | 373 |
by (auto simp: prom_def dest: no_Nonce_in_agl no_Nonce_in_appdel) |
374 |
||
61830 | 375 |
subsection\<open>requests are guarded\<close> |
13508 | 376 |
|
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
377 |
lemma req_imp_Guard [rule_format]: "\<lbrakk>evs \<in> p2; A \<notin> bad\<rbrakk> \<Longrightarrow> |
67613 | 378 |
req A r n I B \<in> set evs \<longrightarrow> Guard n {priK A} (spies evs)" |
13508 | 379 |
apply (erule p2.induct, simp) |
380 |
apply (simp add: req_def knows.simps, safe) |
|
381 |
apply (erule in_synth_Guard, erule Guard_analz, simp) |
|
382 |
by (auto simp: req_def pro_def dest: Says_imp_knows_Spy) |
|
383 |
||
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
384 |
lemma req_imp_Guard_Friend: "\<lbrakk>evs \<in> p2; A \<notin> bad; req A r n I B \<in> set evs\<rbrakk> |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
385 |
\<Longrightarrow> Guard n {priK A} (knows_max (Friend C) evs)" |
13508 | 386 |
apply (rule Guard_knows_max') |
387 |
apply (rule_tac H="spies evs" in Guard_mono) |
|
388 |
apply (rule req_imp_Guard, simp+) |
|
389 |
apply (rule_tac B="spies' evs" in subset_trans) |
|
390 |
apply (rule_tac p=p2 in knows_max'_sub_spies', simp+) |
|
391 |
by (rule knows'_sub_knows) |
|
392 |
||
61830 | 393 |
subsection\<open>propositions are guarded\<close> |
13508 | 394 |
|
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
395 |
lemma pro_imp_Guard [rule_format]: "\<lbrakk>evs \<in> p2; B \<notin> bad; A \<notin> bad\<rbrakk> \<Longrightarrow> |
67613 | 396 |
pro B ofr A r I (cons M L) J C \<in> set evs \<longrightarrow> Guard ofr {priK A} (spies evs)" |
71989
bad75618fb82
extraction of equations x = t from premises beneath meta-all
haftmann
parents:
67613
diff
changeset
|
397 |
supply [[simproc del: defined_all]] |
13508 | 398 |
apply (erule p2.induct) (* +3 subgoals *) |
399 |
(* Nil *) |
|
400 |
apply simp |
|
401 |
(* Fake *) |
|
402 |
apply (simp add: pro_def, safe) (* +4 subgoals *) |
|
403 |
(* 1 *) |
|
404 |
apply (erule in_synth_Guard, drule Guard_analz, simp, simp) |
|
405 |
(* 2 *) |
|
406 |
apply simp |
|
407 |
(* 3 *) |
|
408 |
apply (simp, simp add: req_def pro_def, blast) |
|
409 |
(* 4 *) |
|
410 |
apply (simp add: pro_def) |
|
411 |
apply (blast dest: prom_inj Says_Nonce_not_used_guard Nonce_not_used_Guard) |
|
412 |
(* 5 *) |
|
413 |
apply simp |
|
414 |
apply safe (* +1 subgoal *) |
|
415 |
apply (simp add: pro_def) |
|
416 |
apply (blast dest: prom_inj Says_Nonce_not_used_guard) |
|
417 |
(* 6 *) |
|
418 |
apply (simp add: pro_def) |
|
419 |
apply (blast dest: Says_imp_knows_Spy) |
|
420 |
(* Request *) |
|
421 |
apply (simp add: pro_def) |
|
422 |
apply (blast dest: prom_inj Says_Nonce_not_used_guard Nonce_not_used_Guard) |
|
423 |
(* Propose *) |
|
424 |
apply simp |
|
425 |
apply safe (* +1 subgoal *) |
|
426 |
(* 1 *) |
|
427 |
apply (simp add: pro_def) |
|
428 |
apply (blast dest: prom_inj Says_Nonce_not_used_guard) |
|
429 |
(* 2 *) |
|
430 |
apply (simp add: pro_def) |
|
431 |
by (blast dest: Says_imp_knows_Spy) |
|
432 |
||
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
433 |
lemma pro_imp_Guard_Friend: "\<lbrakk>evs \<in> p2; B \<notin> bad; A \<notin> bad; |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
434 |
pro B ofr A r I (cons M L) J C \<in> set evs\<rbrakk> |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
435 |
\<Longrightarrow> Guard ofr {priK A} (knows_max (Friend D) evs)" |
13508 | 436 |
apply (rule Guard_knows_max') |
437 |
apply (rule_tac H="spies evs" in Guard_mono) |
|
438 |
apply (rule pro_imp_Guard, simp+) |
|
439 |
apply (rule_tac B="spies' evs" in subset_trans) |
|
440 |
apply (rule_tac p=p2 in knows_max'_sub_spies', simp+) |
|
441 |
by (rule knows'_sub_knows) |
|
442 |
||
61830 | 443 |
subsection\<open>data confidentiality: |
444 |
no one other than the originator can decrypt the offers\<close> |
|
13508 | 445 |
|
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
446 |
lemma Nonce_req_notin_spies: "\<lbrakk>evs \<in> p2; req A r n I B \<in> set evs; A \<notin> bad\<rbrakk> |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
447 |
\<Longrightarrow> Nonce n \<notin> analz (spies evs)" |
13508 | 448 |
by (frule req_imp_Guard, simp+, erule Guard_Nonce_analz, simp+) |
449 |
||
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
450 |
lemma Nonce_req_notin_knows_max_Friend: "\<lbrakk>evs \<in> p2; req A r n I B \<in> set evs; |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
451 |
A \<notin> bad; A \<noteq> Friend C\<rbrakk> \<Longrightarrow> Nonce n \<notin> analz (knows_max (Friend C) evs)" |
13508 | 452 |
apply (clarify, frule_tac C=C in req_imp_Guard_Friend, simp+) |
453 |
apply (simp add: knows_max_def, drule Guard_invKey_keyset, simp+) |
|
454 |
by (drule priK_notin_knows_max_Friend, auto simp: knows_max_def) |
|
455 |
||
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
456 |
lemma Nonce_pro_notin_spies: "\<lbrakk>evs \<in> p2; B \<notin> bad; A \<notin> bad; |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
457 |
pro B ofr A r I (cons M L) J C \<in> set evs\<rbrakk> \<Longrightarrow> Nonce ofr \<notin> analz (spies evs)" |
13508 | 458 |
by (frule pro_imp_Guard, simp+, erule Guard_Nonce_analz, simp+) |
459 |
||
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
460 |
lemma Nonce_pro_notin_knows_max_Friend: "\<lbrakk>evs \<in> p2; B \<notin> bad; A \<notin> bad; |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
461 |
A \<noteq> Friend D; pro B ofr A r I (cons M L) J C \<in> set evs\<rbrakk> |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
462 |
\<Longrightarrow> Nonce ofr \<notin> analz (knows_max (Friend D) evs)" |
13508 | 463 |
apply (clarify, frule_tac A=A in pro_imp_Guard_Friend, simp+) |
464 |
apply (simp add: knows_max_def, drule Guard_invKey_keyset, simp+) |
|
465 |
by (drule priK_notin_knows_max_Friend, auto simp: knows_max_def) |
|
466 |
||
61830 | 467 |
subsection\<open>forward privacy: |
468 |
only the originator can know the identity of the shops\<close> |
|
13508 | 469 |
|
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
470 |
lemma forward_privacy_Spy: "\<lbrakk>evs \<in> p2; B \<notin> bad; A \<notin> bad; |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
471 |
pro B ofr A r I (cons M L) J C \<in> set evs\<rbrakk> |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
472 |
\<Longrightarrow> sign B (Nonce ofr) \<notin> analz (spies evs)" |
13508 | 473 |
by (auto simp:sign_def dest: Nonce_pro_notin_spies) |
474 |
||
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
475 |
lemma forward_privacy_Friend: "\<lbrakk>evs \<in> p2; B \<notin> bad; A \<notin> bad; A \<noteq> Friend D; |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
476 |
pro B ofr A r I (cons M L) J C \<in> set evs\<rbrakk> |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
477 |
\<Longrightarrow> sign B (Nonce ofr) \<notin> analz (knows_max (Friend D) evs)" |
13508 | 478 |
by (auto simp:sign_def dest:Nonce_pro_notin_knows_max_Friend ) |
479 |
||
61830 | 480 |
subsection\<open>non repudiability: an offer signed by B has been sent by B\<close> |
13508 | 481 |
|
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
482 |
lemma Crypt_reqm: "\<lbrakk>Crypt (priK A) X \<in> parts {reqm A' r n I B}; I \<in> agl\<rbrakk> \<Longrightarrow> A=A'" |
13508 | 483 |
by (auto simp: reqm_def anchor_def chain_def sign_def dest: no_Crypt_in_agl) |
484 |
||
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
485 |
lemma Crypt_prom: "\<lbrakk>Crypt (priK A) X \<in> parts {prom B ofr A' r I L J C}; |
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
486 |
I \<in> agl; J \<in> agl\<rbrakk> \<Longrightarrow> A=B | Crypt (priK A) X \<in> parts {L}" |
13508 | 487 |
apply (simp add: prom_def anchor_def chain_def sign_def) |
488 |
by (blast dest: no_Crypt_in_agl no_Crypt_in_appdel) |
|
489 |
||
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
490 |
lemma Crypt_safeness: "\<lbrakk>evs \<in> p2; A \<notin> bad\<rbrakk> \<Longrightarrow> Crypt (priK A) X \<in> parts (spies evs) |
67613 | 491 |
\<longrightarrow> (\<exists>B Y. Says A B Y \<in> set evs & Crypt (priK A) X \<in> parts {Y})" |
13508 | 492 |
apply (erule p2.induct) |
493 |
(* Nil *) |
|
494 |
apply simp |
|
495 |
(* Fake *) |
|
496 |
apply clarsimp |
|
67613 | 497 |
apply (drule_tac P="\<lambda>G. Crypt (priK A) X \<in> G" in parts_insert_substD, simp) |
13508 | 498 |
apply (erule disjE) |
499 |
apply (drule_tac K="priK A" in Crypt_synth, simp+, blast, blast) |
|
500 |
(* Request *) |
|
501 |
apply (simp add: req_def, clarify) |
|
67613 | 502 |
apply (drule_tac P="\<lambda>G. Crypt (priK A) X \<in> G" in parts_insert_substD, simp) |
13508 | 503 |
apply (erule disjE) |
504 |
apply (frule Crypt_reqm, simp, clarify) |
|
505 |
apply (rule_tac x=B in exI, rule_tac x="reqm A r n I B" in exI, simp, blast) |
|
506 |
(* Propose *) |
|
507 |
apply (simp add: pro_def, clarify) |
|
67613 | 508 |
apply (drule_tac P="\<lambda>G. Crypt (priK A) X \<in> G" in parts_insert_substD, simp) |
13508 | 509 |
apply (rotate_tac -1, erule disjE) |
510 |
apply (frule Crypt_prom, simp, simp) |
|
511 |
apply (rotate_tac -1, erule disjE) |
|
512 |
apply (rule_tac x=C in exI) |
|
513 |
apply (rule_tac x="prom B ofr Aa r I (cons M L) J C" in exI, blast) |
|
67613 | 514 |
apply (subgoal_tac "cons M L \<in> parts (spies evsp)") |
13508 | 515 |
apply (drule_tac G="{cons M L}" and H="spies evsp" in parts_trans, blast, blast) |
516 |
apply (drule Says_imp_spies, rotate_tac -1, drule parts.Inj) |
|
517 |
apply (drule parts.Snd, drule parts.Snd, drule parts.Snd) |
|
518 |
by auto |
|
519 |
||
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
520 |
lemma Crypt_Hash_imp_sign: "\<lbrakk>evs \<in> p2; A \<notin> bad\<rbrakk> \<Longrightarrow> |
67613 | 521 |
Crypt (priK A) (Hash X) \<in> parts (spies evs) |
522 |
\<longrightarrow> (\<exists>B Y. Says A B Y \<in> set evs \<and> sign A X \<in> parts {Y})" |
|
13508 | 523 |
apply (erule p2.induct) |
524 |
(* Nil *) |
|
525 |
apply simp |
|
526 |
(* Fake *) |
|
527 |
apply clarsimp |
|
67613 | 528 |
apply (drule_tac P="\<lambda>G. Crypt (priK A) (Hash X) \<in> G" in parts_insert_substD) |
13508 | 529 |
apply simp |
530 |
apply (erule disjE) |
|
531 |
apply (drule_tac K="priK A" in Crypt_synth, simp+, blast, blast) |
|
532 |
(* Request *) |
|
533 |
apply (simp add: req_def, clarify) |
|
67613 | 534 |
apply (drule_tac P="\<lambda>G. Crypt (priK A) (Hash X) \<in> G" in parts_insert_substD) |
13508 | 535 |
apply simp |
536 |
apply (erule disjE) |
|
537 |
apply (frule Crypt_reqm, simp+) |
|
538 |
apply (rule_tac x=B in exI, rule_tac x="reqm Aa r n I B" in exI) |
|
539 |
apply (simp add: reqm_def sign_def anchor_def no_Crypt_in_agl) |
|
540 |
apply (simp add: chain_def sign_def, blast) |
|
541 |
(* Propose *) |
|
542 |
apply (simp add: pro_def, clarify) |
|
67613 | 543 |
apply (drule_tac P="\<lambda>G. Crypt (priK A) (Hash X) \<in> G" in parts_insert_substD) |
13508 | 544 |
apply simp |
545 |
apply (rotate_tac -1, erule disjE) |
|
546 |
apply (simp add: prom_def sign_def no_Crypt_in_agl no_Crypt_in_appdel) |
|
547 |
apply (simp add: chain_def sign_def) |
|
548 |
apply (rotate_tac -1, erule disjE) |
|
549 |
apply (rule_tac x=C in exI) |
|
550 |
apply (rule_tac x="prom B ofr Aa r I (cons M L) J C" in exI) |
|
551 |
apply (simp add: prom_def chain_def sign_def) |
|
552 |
apply (erule impE) |
|
553 |
apply (blast dest: get_ML parts_sub) |
|
554 |
apply (blast del: MPair_parts)+ |
|
555 |
done |
|
556 |
||
76287
cdc14f94c754
Elimination of the archaic ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71989
diff
changeset
|
557 |
lemma sign_safeness: "\<lbrakk>evs \<in> p2; A \<notin> bad\<rbrakk> \<Longrightarrow> sign A X \<in> parts (spies evs) |
67613 | 558 |
\<longrightarrow> (\<exists>B Y. Says A B Y \<in> set evs \<and> sign A X \<in> parts {Y})" |
13508 | 559 |
apply (clarify, simp add: sign_def, frule parts.Snd) |
560 |
apply (blast dest: Crypt_Hash_imp_sign [unfolded sign_def]) |
|
561 |
done |
|
562 |
||
62390 | 563 |
end |