author | wenzelm |
Wed, 06 Dec 2006 01:12:43 +0100 | |
changeset 21672 | 29c346b165d4 |
parent 16417 | 9bc16273c2d4 |
child 23746 | a455e69c31cc |
permissions | -rw-r--r-- |
6400 | 1 |
(* Title: HOL/Auth/Yahalom |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1996 University of Cambridge |
|
5 |
*) |
|
6 |
||
13956 | 7 |
header{*The Yahalom Protocol: A Flawed Version*} |
8 |
||
16417 | 9 |
theory Yahalom_Bad imports Public begin |
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
10 |
|
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
11 |
text{* |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
12 |
Demonstrates of why Oops is necessary. This protocol can be attacked because |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
13 |
it doesn't keep NB secret, but without Oops it can be "verified" anyway. |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
14 |
The issues are discussed in lcp's LICS 2000 invited lecture. |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
15 |
*} |
6400 | 16 |
|
11251 | 17 |
consts yahalom :: "event list set" |
6400 | 18 |
inductive "yahalom" |
11251 | 19 |
intros |
6400 | 20 |
(*Initial trace is empty*) |
11251 | 21 |
Nil: "[] \<in> yahalom" |
6400 | 22 |
|
23 |
(*The spy MAY say anything he CAN say. We do not expect him to |
|
24 |
invent new nonces here, but he can also use NS1. Common to |
|
25 |
all similar protocols.*) |
|
11251 | 26 |
Fake: "[| evsf \<in> yahalom; X \<in> synth (analz (knows Spy evsf)) |] |
27 |
==> Says Spy B X # evsf \<in> yahalom" |
|
6400 | 28 |
|
29 |
(*A message that has been sent can be received by the |
|
30 |
intended recipient.*) |
|
11251 | 31 |
Reception: "[| evsr \<in> yahalom; Says A B X \<in> set evsr |] |
32 |
==> Gets B X # evsr \<in> yahalom" |
|
6400 | 33 |
|
34 |
(*Alice initiates a protocol run*) |
|
11251 | 35 |
YM1: "[| evs1 \<in> yahalom; Nonce NA \<notin> used evs1 |] |
36 |
==> Says A B {|Agent A, Nonce NA|} # evs1 \<in> yahalom" |
|
6400 | 37 |
|
38 |
(*Bob's response to Alice's message.*) |
|
11251 | 39 |
YM2: "[| evs2 \<in> yahalom; Nonce NB \<notin> used evs2; |
40 |
Gets B {|Agent A, Nonce NA|} \<in> set evs2 |] |
|
41 |
==> Says B Server |
|
6400 | 42 |
{|Agent B, Nonce NB, Crypt (shrK B) {|Agent A, Nonce NA|}|} |
11251 | 43 |
# evs2 \<in> yahalom" |
6400 | 44 |
|
45 |
(*The Server receives Bob's message. He responds by sending a |
|
46 |
new session key to Alice, with a packet for forwarding to Bob.*) |
|
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
47 |
YM3: "[| evs3 \<in> yahalom; Key KAB \<notin> used evs3; KAB \<in> symKeys; |
11251 | 48 |
Gets Server |
6400 | 49 |
{|Agent B, Nonce NB, Crypt (shrK B) {|Agent A, Nonce NA|}|} |
11251 | 50 |
\<in> set evs3 |] |
6400 | 51 |
==> Says Server A |
52 |
{|Crypt (shrK A) {|Agent B, Key KAB, Nonce NA, Nonce NB|}, |
|
53 |
Crypt (shrK B) {|Agent A, Key KAB|}|} |
|
11251 | 54 |
# evs3 \<in> yahalom" |
6400 | 55 |
|
56 |
(*Alice receives the Server's (?) message, checks her Nonce, and |
|
57 |
uses the new session key to send Bob his Nonce. The premise |
|
11251 | 58 |
A \<noteq> Server is needed to prove Says_Server_not_range.*) |
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
59 |
YM4: "[| evs4 \<in> yahalom; A \<noteq> Server; K \<in> symKeys; |
6400 | 60 |
Gets A {|Crypt(shrK A) {|Agent B, Key K, Nonce NA, Nonce NB|}, X|} |
11251 | 61 |
\<in> set evs4; |
62 |
Says A B {|Agent A, Nonce NA|} \<in> set evs4 |] |
|
63 |
==> Says A B {|X, Crypt K (Nonce NB)|} # evs4 \<in> yahalom" |
|
64 |
||
65 |
||
66 |
declare Says_imp_knows_Spy [THEN analz.Inj, dest] |
|
67 |
declare parts.Body [dest] |
|
68 |
declare Fake_parts_insert_in_Un [dest] |
|
69 |
declare analz_into_parts [dest] |
|
70 |
||
71 |
||
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
72 |
text{*A "possibility property": there are traces that reach the end*} |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
73 |
lemma "[| A \<noteq> Server; Key K \<notin> used []; K \<in> symKeys |] |
14200
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents:
13956
diff
changeset
|
74 |
==> \<exists>X NB. \<exists>evs \<in> yahalom. |
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents:
13956
diff
changeset
|
75 |
Says A B {|X, Crypt K (Nonce NB)|} \<in> set evs" |
11251 | 76 |
apply (intro exI bexI) |
77 |
apply (rule_tac [2] yahalom.Nil |
|
78 |
[THEN yahalom.YM1, THEN yahalom.Reception, |
|
79 |
THEN yahalom.YM2, THEN yahalom.Reception, |
|
80 |
THEN yahalom.YM3, THEN yahalom.Reception, |
|
14200
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents:
13956
diff
changeset
|
81 |
THEN yahalom.YM4]) |
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents:
13956
diff
changeset
|
82 |
apply (possibility, simp add: used_Cons) |
11251 | 83 |
done |
84 |
||
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
85 |
subsection{*Regularity Lemmas for Yahalom*} |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
86 |
|
11251 | 87 |
lemma Gets_imp_Says: |
88 |
"[| Gets B X \<in> set evs; evs \<in> yahalom |] ==> \<exists>A. Says A B X \<in> set evs" |
|
89 |
by (erule rev_mp, erule yahalom.induct, auto) |
|
90 |
||
91 |
(*Must be proved separately for each protocol*) |
|
92 |
lemma Gets_imp_knows_Spy: |
|
93 |
"[| Gets B X \<in> set evs; evs \<in> yahalom |] ==> X \<in> knows Spy evs" |
|
94 |
by (blast dest!: Gets_imp_Says Says_imp_knows_Spy) |
|
95 |
||
96 |
declare Gets_imp_knows_Spy [THEN analz.Inj, dest] |
|
97 |
||
98 |
||
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
99 |
subsection{* For reasoning about the encrypted portion of messages *} |
11251 | 100 |
|
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
101 |
text{*Lets us treat YM4 using a similar argument as for the Fake case.*} |
11251 | 102 |
lemma YM4_analz_knows_Spy: |
103 |
"[| Gets A {|Crypt (shrK A) Y, X|} \<in> set evs; evs \<in> yahalom |] |
|
104 |
==> X \<in> analz (knows Spy evs)" |
|
105 |
by blast |
|
106 |
||
107 |
lemmas YM4_parts_knows_Spy = |
|
108 |
YM4_analz_knows_Spy [THEN analz_into_parts, standard] |
|
109 |
||
110 |
||
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
111 |
text{*Theorems of the form @{term "X \<notin> parts (knows Spy evs)"} imply |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
112 |
that NOBODY sends messages containing X!*} |
11251 | 113 |
|
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
114 |
text{*Spy never sees a good agent's shared key!*} |
11251 | 115 |
lemma Spy_see_shrK [simp]: |
116 |
"evs \<in> yahalom ==> (Key (shrK A) \<in> parts (knows Spy evs)) = (A \<in> bad)" |
|
117 |
apply (erule yahalom.induct, force, |
|
13507 | 118 |
drule_tac [6] YM4_parts_knows_Spy, simp_all, blast+) |
11251 | 119 |
done |
120 |
||
121 |
lemma Spy_analz_shrK [simp]: |
|
122 |
"evs \<in> yahalom ==> (Key (shrK A) \<in> analz (knows Spy evs)) = (A \<in> bad)" |
|
123 |
by auto |
|
124 |
||
125 |
lemma Spy_see_shrK_D [dest!]: |
|
126 |
"[|Key (shrK A) \<in> parts (knows Spy evs); evs \<in> yahalom|] ==> A \<in> bad" |
|
127 |
by (blast dest: Spy_see_shrK) |
|
128 |
||
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
129 |
text{*Nobody can have used non-existent keys! |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
130 |
Needed to apply @{text analz_insert_Key}*} |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
131 |
lemma new_keys_not_used [simp]: |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
132 |
"[|Key K \<notin> used evs; K \<in> symKeys; evs \<in> yahalom|] |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
133 |
==> K \<notin> keysFor (parts (spies evs))" |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
134 |
apply (erule rev_mp) |
11251 | 135 |
apply (erule yahalom.induct, force, |
136 |
frule_tac [6] YM4_parts_knows_Spy, simp_all) |
|
13926 | 137 |
txt{*Fake*} |
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
138 |
apply (force dest!: keysFor_parts_insert, auto) |
11251 | 139 |
done |
140 |
||
141 |
||
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
142 |
subsection{*Secrecy Theorems*} |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
143 |
|
11251 | 144 |
(**** |
145 |
The following is to prove theorems of the form |
|
146 |
||
147 |
Key K \<in> analz (insert (Key KAB) (knows Spy evs)) ==> |
|
148 |
Key K \<in> analz (knows Spy evs) |
|
149 |
||
150 |
A more general formula must be proved inductively. |
|
151 |
****) |
|
152 |
||
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
153 |
subsection{* Session keys are not used to encrypt other session keys *} |
11251 | 154 |
|
155 |
lemma analz_image_freshK [rule_format]: |
|
156 |
"evs \<in> yahalom ==> |
|
157 |
\<forall>K KK. KK <= - (range shrK) --> |
|
158 |
(Key K \<in> analz (Key`KK Un (knows Spy evs))) = |
|
159 |
(K \<in> KK | Key K \<in> analz (knows Spy evs))" |
|
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
160 |
by (erule yahalom.induct, |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
161 |
drule_tac [7] YM4_analz_knows_Spy, analz_freshK, spy_analz, blast) |
11251 | 162 |
|
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
163 |
lemma analz_insert_freshK: |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
164 |
"[| evs \<in> yahalom; KAB \<notin> range shrK |] ==> |
11655 | 165 |
(Key K \<in> analz (insert (Key KAB) (knows Spy evs))) = |
11251 | 166 |
(K = KAB | Key K \<in> analz (knows Spy evs))" |
167 |
by (simp only: analz_image_freshK analz_image_freshK_simps) |
|
168 |
||
169 |
||
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
170 |
text{*The Key K uniquely identifies the Server's message.*} |
11251 | 171 |
lemma unique_session_keys: |
172 |
"[| Says Server A |
|
173 |
{|Crypt (shrK A) {|Agent B, Key K, na, nb|}, X|} \<in> set evs; |
|
174 |
Says Server A' |
|
175 |
{|Crypt (shrK A') {|Agent B', Key K, na', nb'|}, X'|} \<in> set evs; |
|
176 |
evs \<in> yahalom |] |
|
177 |
==> A=A' & B=B' & na=na' & nb=nb'" |
|
178 |
apply (erule rev_mp, erule rev_mp) |
|
179 |
apply (erule yahalom.induct, simp_all) |
|
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
180 |
txt{*YM3, by freshness, and YM4*} |
11251 | 181 |
apply blast+ |
182 |
done |
|
183 |
||
184 |
||
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
185 |
text{* Crucial secrecy property: Spy does not see the keys sent in msg YM3 *} |
11251 | 186 |
lemma secrecy_lemma: |
187 |
"[| A \<notin> bad; B \<notin> bad; evs \<in> yahalom |] |
|
188 |
==> Says Server A |
|
189 |
{|Crypt (shrK A) {|Agent B, Key K, na, nb|}, |
|
190 |
Crypt (shrK B) {|Agent A, Key K|}|} |
|
191 |
\<in> set evs --> |
|
192 |
Key K \<notin> analz (knows Spy evs)" |
|
193 |
apply (erule yahalom.induct, force, drule_tac [6] YM4_analz_knows_Spy) |
|
13507 | 194 |
apply (simp_all add: pushes analz_insert_eq analz_insert_freshK, spy_analz) (*Fake*) |
11251 | 195 |
apply (blast dest: unique_session_keys) (*YM3*) |
196 |
done |
|
197 |
||
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
198 |
text{*Final version*} |
11251 | 199 |
lemma Spy_not_see_encrypted_key: |
200 |
"[| Says Server A |
|
201 |
{|Crypt (shrK A) {|Agent B, Key K, na, nb|}, |
|
202 |
Crypt (shrK B) {|Agent A, Key K|}|} |
|
203 |
\<in> set evs; |
|
204 |
A \<notin> bad; B \<notin> bad; evs \<in> yahalom |] |
|
205 |
==> Key K \<notin> analz (knows Spy evs)" |
|
206 |
by (blast dest: secrecy_lemma) |
|
207 |
||
6400 | 208 |
|
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
209 |
subsection{* Security Guarantee for A upon receiving YM3 *} |
11251 | 210 |
|
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
211 |
text{*If the encrypted message appears then it originated with the Server*} |
11251 | 212 |
lemma A_trusts_YM3: |
213 |
"[| Crypt (shrK A) {|Agent B, Key K, na, nb|} \<in> parts (knows Spy evs); |
|
214 |
A \<notin> bad; evs \<in> yahalom |] |
|
215 |
==> Says Server A |
|
216 |
{|Crypt (shrK A) {|Agent B, Key K, na, nb|}, |
|
217 |
Crypt (shrK B) {|Agent A, Key K|}|} |
|
218 |
\<in> set evs" |
|
219 |
apply (erule rev_mp) |
|
220 |
apply (erule yahalom.induct, force, |
|
221 |
frule_tac [6] YM4_parts_knows_Spy, simp_all) |
|
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
222 |
txt{*Fake, YM3*} |
11251 | 223 |
apply blast+ |
224 |
done |
|
225 |
||
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
226 |
text{*The obvious combination of @{text A_trusts_YM3} with |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
227 |
@{text Spy_not_see_encrypted_key}*} |
11251 | 228 |
lemma A_gets_good_key: |
229 |
"[| Crypt (shrK A) {|Agent B, Key K, na, nb|} \<in> parts (knows Spy evs); |
|
230 |
A \<notin> bad; B \<notin> bad; evs \<in> yahalom |] |
|
231 |
==> Key K \<notin> analz (knows Spy evs)" |
|
232 |
by (blast dest!: A_trusts_YM3 Spy_not_see_encrypted_key) |
|
233 |
||
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
234 |
subsection{* Security Guarantees for B upon receiving YM4 *} |
11251 | 235 |
|
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
236 |
text{*B knows, by the first part of A's message, that the Server distributed |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
237 |
the key for A and B. But this part says nothing about nonces.*} |
11251 | 238 |
lemma B_trusts_YM4_shrK: |
239 |
"[| Crypt (shrK B) {|Agent A, Key K|} \<in> parts (knows Spy evs); |
|
240 |
B \<notin> bad; evs \<in> yahalom |] |
|
241 |
==> \<exists>NA NB. Says Server A |
|
242 |
{|Crypt (shrK A) {|Agent B, Key K, Nonce NA, Nonce NB|}, |
|
243 |
Crypt (shrK B) {|Agent A, Key K|}|} |
|
244 |
\<in> set evs" |
|
245 |
apply (erule rev_mp) |
|
246 |
apply (erule yahalom.induct, force, |
|
247 |
frule_tac [6] YM4_parts_knows_Spy, simp_all) |
|
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
248 |
txt{*Fake, YM3*} |
11251 | 249 |
apply blast+ |
250 |
done |
|
251 |
||
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
252 |
subsection{*The Flaw in the Model*} |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
253 |
|
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
254 |
text{* Up to now, the reasoning is similar to standard Yahalom. Now the |
11251 | 255 |
doubtful reasoning occurs. We should not be assuming that an unknown |
256 |
key is secure, but the model allows us to: there is no Oops rule to |
|
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
257 |
let session keys become compromised.*} |
11251 | 258 |
|
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
259 |
text{*B knows, by the second part of A's message, that the Server distributed |
11251 | 260 |
the key quoting nonce NB. This part says nothing about agent names. |
261 |
Secrecy of K is assumed; the valid Yahalom proof uses (and later proves) |
|
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
262 |
the secrecy of NB.*} |
11251 | 263 |
lemma B_trusts_YM4_newK [rule_format]: |
264 |
"[|Key K \<notin> analz (knows Spy evs); evs \<in> yahalom|] |
|
265 |
==> Crypt K (Nonce NB) \<in> parts (knows Spy evs) --> |
|
266 |
(\<exists>A B NA. Says Server A |
|
267 |
{|Crypt (shrK A) {|Agent B, Key K, |
|
268 |
Nonce NA, Nonce NB|}, |
|
269 |
Crypt (shrK B) {|Agent A, Key K|}|} |
|
270 |
\<in> set evs)" |
|
271 |
apply (erule rev_mp) |
|
272 |
apply (erule yahalom.induct, force, |
|
273 |
frule_tac [6] YM4_parts_knows_Spy) |
|
274 |
apply (analz_mono_contra, simp_all) |
|
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
275 |
txt{*Fake*} |
11251 | 276 |
apply blast |
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
277 |
txt{*YM3*} |
11251 | 278 |
apply blast |
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
279 |
txt{*A is uncompromised because NB is secure |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
280 |
A's certificate guarantees the existence of the Server message*} |
11251 | 281 |
apply (blast dest!: Gets_imp_Says Crypt_Spy_analz_bad |
282 |
dest: Says_imp_spies |
|
283 |
parts.Inj [THEN parts.Fst, THEN A_trusts_YM3]) |
|
284 |
done |
|
285 |
||
286 |
||
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
287 |
text{*B's session key guarantee from YM4. The two certificates contribute to a |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
288 |
single conclusion about the Server's message. *} |
11251 | 289 |
lemma B_trusts_YM4: |
290 |
"[| Gets B {|Crypt (shrK B) {|Agent A, Key K|}, |
|
291 |
Crypt K (Nonce NB)|} \<in> set evs; |
|
292 |
Says B Server |
|
293 |
{|Agent B, Nonce NB, Crypt (shrK B) {|Agent A, Nonce NA|}|} |
|
294 |
\<in> set evs; |
|
295 |
A \<notin> bad; B \<notin> bad; evs \<in> yahalom |] |
|
296 |
==> \<exists>na nb. Says Server A |
|
297 |
{|Crypt (shrK A) {|Agent B, Key K, na, nb|}, |
|
298 |
Crypt (shrK B) {|Agent A, Key K|}|} |
|
299 |
\<in> set evs" |
|
300 |
by (blast dest: B_trusts_YM4_newK B_trusts_YM4_shrK Spy_not_see_encrypted_key |
|
301 |
unique_session_keys) |
|
302 |
||
303 |
||
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
304 |
text{*The obvious combination of @{text B_trusts_YM4} with |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
305 |
@{text Spy_not_see_encrypted_key}*} |
11251 | 306 |
lemma B_gets_good_key: |
307 |
"[| Gets B {|Crypt (shrK B) {|Agent A, Key K|}, |
|
308 |
Crypt K (Nonce NB)|} \<in> set evs; |
|
309 |
Says B Server |
|
310 |
{|Agent B, Nonce NB, Crypt (shrK B) {|Agent A, Nonce NA|}|} |
|
311 |
\<in> set evs; |
|
312 |
A \<notin> bad; B \<notin> bad; evs \<in> yahalom |] |
|
313 |
==> Key K \<notin> analz (knows Spy evs)" |
|
314 |
by (blast dest!: B_trusts_YM4 Spy_not_see_encrypted_key) |
|
315 |
||
316 |
||
317 |
(*** Authenticating B to A: these proofs are not considered. |
|
318 |
They are irrelevant to showing the need for Oops. ***) |
|
319 |
||
320 |
||
321 |
(*** Authenticating A to B using the certificate Crypt K (Nonce NB) ***) |
|
322 |
||
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
323 |
text{*Assuming the session key is secure, if both certificates are present then |
11251 | 324 |
A has said NB. We can't be sure about the rest of A's message, but only |
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
325 |
NB matters for freshness.*} |
11251 | 326 |
lemma A_Said_YM3_lemma [rule_format]: |
327 |
"evs \<in> yahalom |
|
328 |
==> Key K \<notin> analz (knows Spy evs) --> |
|
329 |
Crypt K (Nonce NB) \<in> parts (knows Spy evs) --> |
|
330 |
Crypt (shrK B) {|Agent A, Key K|} \<in> parts (knows Spy evs) --> |
|
331 |
B \<notin> bad --> |
|
332 |
(\<exists>X. Says A B {|X, Crypt K (Nonce NB)|} \<in> set evs)" |
|
333 |
apply (erule yahalom.induct, force, |
|
334 |
frule_tac [6] YM4_parts_knows_Spy) |
|
335 |
apply (analz_mono_contra, simp_all) |
|
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
336 |
txt{*Fake*} |
11251 | 337 |
apply blast |
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
338 |
txt{*YM3: by @{text new_keys_not_used}, the message |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
339 |
@{term "Crypt K (Nonce NB)"} could not exist*} |
11251 | 340 |
apply (force dest!: Crypt_imp_keysFor) |
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
341 |
txt{*YM4: was @{term "Crypt K (Nonce NB)"} the very last message? |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
342 |
If not, use the induction hypothesis*} |
11251 | 343 |
apply (simp add: ex_disj_distrib) |
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
344 |
txt{*yes: apply unicity of session keys*} |
11251 | 345 |
apply (blast dest!: Gets_imp_Says A_trusts_YM3 B_trusts_YM4_shrK |
346 |
Crypt_Spy_analz_bad |
|
347 |
dest: Says_imp_knows_Spy [THEN parts.Inj] unique_session_keys) |
|
348 |
done |
|
349 |
||
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
350 |
text{*If B receives YM4 then A has used nonce NB (and therefore is alive). |
11251 | 351 |
Moreover, A associates K with NB (thus is talking about the same run). |
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
352 |
Other premises guarantee secrecy of K.*} |
11251 | 353 |
lemma YM4_imp_A_Said_YM3 [rule_format]: |
354 |
"[| Gets B {|Crypt (shrK B) {|Agent A, Key K|}, |
|
355 |
Crypt K (Nonce NB)|} \<in> set evs; |
|
356 |
Says B Server |
|
357 |
{|Agent B, Nonce NB, Crypt (shrK B) {|Agent A, Nonce NA|}|} |
|
358 |
\<in> set evs; |
|
359 |
A \<notin> bad; B \<notin> bad; evs \<in> yahalom |] |
|
360 |
==> \<exists>X. Says A B {|X, Crypt K (Nonce NB)|} \<in> set evs" |
|
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
361 |
by (blast intro!: A_Said_YM3_lemma |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
362 |
dest: Spy_not_see_encrypted_key B_trusts_YM4 Gets_imp_Says) |
6400 | 363 |
|
364 |
end |