author | paulson |
Tue, 29 Apr 2003 12:36:49 +0200 | |
changeset 13935 | 4822d9597d1e |
parent 13926 | 6e62e5357a10 |
child 13956 | 8fe7e12290e1 |
permissions | -rw-r--r-- |
2318 | 1 |
(* Title: HOL/Auth/Public |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1996 University of Cambridge |
|
5 |
||
3512
9dcb4daa15e8
Moving common declarations and proofs from theories "Shared"
paulson
parents:
3478
diff
changeset
|
6 |
Theory of Public Keys (common to all public-key protocols) |
2318 | 7 |
|
3512
9dcb4daa15e8
Moving common declarations and proofs from theories "Shared"
paulson
parents:
3478
diff
changeset
|
8 |
Private and public keys; initial states of agents |
2318 | 9 |
*) |
10 |
||
13922 | 11 |
theory Public = Event: |
12 |
||
13 |
subsection{*Asymmetric Keys*} |
|
2318 | 14 |
|
15 |
consts |
|
13922 | 16 |
(*the bool is TRUE if a signing key*) |
17 |
publicKey :: "[bool,agent] => key" |
|
2318 | 18 |
|
19 |
syntax |
|
13922 | 20 |
pubEK :: "agent => key" |
21 |
pubSK :: "agent => key" |
|
22 |
||
23 |
privateKey :: "[bool,agent] => key" |
|
24 |
priEK :: "agent => key" |
|
25 |
priSK :: "agent => key" |
|
26 |
||
27 |
translations |
|
28 |
"pubEK" == "publicKey False" |
|
29 |
"pubSK" == "publicKey True" |
|
30 |
||
31 |
(*BEWARE!! priEK, priSK DON'T WORK with inj, range, image, etc.*) |
|
32 |
"privateKey b A" == "invKey (publicKey b A)" |
|
33 |
"priEK A" == "privateKey False A" |
|
34 |
"priSK A" == "privateKey True A" |
|
35 |
||
36 |
||
37 |
text{*These translations give backward compatibility. They represent the |
|
38 |
simple situation where the signature and encryption keys are the same.*} |
|
39 |
syntax |
|
40 |
pubK :: "agent => key" |
|
41 |
priK :: "agent => key" |
|
42 |
||
43 |
translations |
|
44 |
"pubK A" == "pubEK A" |
|
45 |
"priK A" == "invKey (pubEK A)" |
|
46 |
||
47 |
||
48 |
axioms |
|
49 |
(*By freeness of agents, no two agents have the same key. Since true\<noteq>false, |
|
50 |
no agent has identical signing and encryption keys*) |
|
51 |
injective_publicKey: |
|
52 |
"publicKey b A = publicKey c A' ==> b=c & A=A'" |
|
53 |
||
54 |
(*No private key equals any public key (essential to ensure that private |
|
55 |
keys are private!) *) |
|
56 |
privateKey_neq_publicKey [iff]: "privateKey b A \<noteq> publicKey c A'" |
|
57 |
||
13926 | 58 |
declare privateKey_neq_publicKey [THEN not_sym, iff] |
13922 | 59 |
|
60 |
||
13926 | 61 |
subsection{*Basic properties of @{term pubK} and @{term priK}*} |
13922 | 62 |
|
63 |
lemma [iff]: "(publicKey b A = publicKey c A') = (b=c & A=A')" |
|
64 |
by (blast dest!: injective_publicKey) |
|
65 |
||
66 |
lemma not_symKeys_pubK [iff]: "publicKey b A \<notin> symKeys" |
|
13926 | 67 |
by (simp add: symKeys_def) |
13922 | 68 |
|
69 |
lemma not_symKeys_priK [iff]: "privateKey b A \<notin> symKeys" |
|
13926 | 70 |
by (simp add: symKeys_def) |
13922 | 71 |
|
72 |
lemma symKey_neq_priEK: "K \<in> symKeys ==> K \<noteq> priEK A" |
|
73 |
by auto |
|
74 |
||
75 |
lemma symKeys_neq_imp_neq: "(K \<in> symKeys) \<noteq> (K' \<in> symKeys) ==> K \<noteq> K'" |
|
13926 | 76 |
by blast |
13922 | 77 |
|
78 |
lemma symKeys_invKey_iff [iff]: "(invKey K \<in> symKeys) = (K \<in> symKeys)" |
|
13926 | 79 |
by (unfold symKeys_def, auto) |
2318 | 80 |
|
13922 | 81 |
lemma analz_symKeys_Decrypt: |
82 |
"[| Crypt K X \<in> analz H; K \<in> symKeys; Key K \<in> analz H |] |
|
83 |
==> X \<in> analz H" |
|
84 |
by (auto simp add: symKeys_def) |
|
85 |
||
86 |
||
87 |
||
88 |
subsection{*"Image" equations that hold for injective functions*} |
|
89 |
||
90 |
lemma invKey_image_eq [simp]: "(invKey x \<in> invKey`A) = (x \<in> A)" |
|
13926 | 91 |
by auto |
13922 | 92 |
|
93 |
(*holds because invKey is injective*) |
|
94 |
lemma publicKey_image_eq [simp]: |
|
95 |
"(publicKey b x \<in> publicKey c ` AA) = (b=c & x \<in> AA)" |
|
96 |
by auto |
|
97 |
||
98 |
lemma privateKey_notin_image_publicKey [simp]: "privateKey b x \<notin> publicKey c ` AA" |
|
13926 | 99 |
by auto |
13922 | 100 |
|
101 |
lemma privateKey_image_eq [simp]: |
|
102 |
"(privateKey b A \<in> invKey ` publicKey c ` AS) = (b=c & A\<in>AS)" |
|
103 |
by auto |
|
104 |
||
105 |
lemma publicKey_notin_image_privateKey [simp]: "publicKey b A \<notin> invKey ` publicKey c ` AS" |
|
13926 | 106 |
by auto |
13922 | 107 |
|
108 |
||
109 |
subsection{*Symmetric Keys*} |
|
110 |
||
111 |
text{*For some protocols, it is convenient to equip agents with symmetric as |
|
112 |
well as asymmetric keys. The theory @{text Shared} assumes that all keys |
|
113 |
are symmetric.*} |
|
114 |
||
115 |
consts |
|
116 |
shrK :: "agent => key" --{*long-term shared keys*} |
|
117 |
||
118 |
axioms |
|
119 |
inj_shrK: "inj shrK" --{*No two agents have the same key*} |
|
120 |
sym_shrK [iff]: "shrK X \<in> symKeys" --{*All shared keys are symmetric*} |
|
121 |
||
122 |
(*Injectiveness: Agents' long-term keys are distinct.*) |
|
123 |
declare inj_shrK [THEN inj_eq, iff] |
|
124 |
||
125 |
lemma priK_neq_shrK [iff]: "shrK A \<noteq> privateKey b C" |
|
13926 | 126 |
by (simp add: symKeys_neq_imp_neq) |
13922 | 127 |
|
128 |
declare priK_neq_shrK [THEN not_sym, simp] |
|
129 |
||
130 |
lemma pubK_neq_shrK [iff]: "shrK A \<noteq> publicKey b C" |
|
13926 | 131 |
by (simp add: symKeys_neq_imp_neq) |
13922 | 132 |
|
133 |
declare pubK_neq_shrK [THEN not_sym, simp] |
|
134 |
||
135 |
lemma priEK_noteq_shrK [simp]: "priEK A \<noteq> shrK B" |
|
136 |
by auto |
|
137 |
||
138 |
lemma publicKey_notin_image_shrK [simp]: "publicKey b x \<notin> shrK ` AA" |
|
139 |
by auto |
|
140 |
||
141 |
lemma privateKey_notin_image_shrK [simp]: "privateKey b x \<notin> shrK ` AA" |
|
142 |
by auto |
|
143 |
||
144 |
lemma shrK_notin_image_publicKey [simp]: "shrK x \<notin> publicKey b ` AA" |
|
145 |
by auto |
|
146 |
||
147 |
lemma shrK_notin_image_privateKey [simp]: "shrK x \<notin> invKey ` publicKey b ` AA" |
|
148 |
by auto |
|
149 |
||
150 |
lemma shrK_image_eq [simp]: "(shrK x \<in> shrK ` AA) = (x \<in> AA)" |
|
151 |
by auto |
|
152 |
||
153 |
||
154 |
subsection{*Initial States of Agents*} |
|
155 |
||
156 |
text{*Note: for all practical purposes, all that matters is the initial |
|
157 |
knowledge of the Spy. All other agents are automata, merely following the |
|
158 |
protocol.*} |
|
2318 | 159 |
|
5183 | 160 |
primrec |
2318 | 161 |
(*Agents know their private key and all public keys*) |
13922 | 162 |
initState_Server: |
163 |
"initState Server = |
|
164 |
{Key (priEK Server), Key (priSK Server)} \<union> |
|
165 |
(Key ` range pubEK) \<union> (Key ` range pubSK) \<union> (Key ` range shrK)" |
|
166 |
||
167 |
initState_Friend: |
|
168 |
"initState (Friend i) = |
|
169 |
{Key (priEK(Friend i)), Key (priSK(Friend i)), Key (shrK(Friend i))} \<union> |
|
170 |
(Key ` range pubEK) \<union> (Key ` range pubSK)" |
|
171 |
||
172 |
initState_Spy: |
|
173 |
"initState Spy = |
|
174 |
(Key ` invKey ` pubEK ` bad) \<union> (Key ` invKey ` pubSK ` bad) \<union> |
|
175 |
(Key ` shrK ` bad) \<union> |
|
176 |
(Key ` range pubEK) \<union> (Key ` range pubSK)" |
|
177 |
||
178 |
||
179 |
text{*These lemmas allow reasoning about @{term "used evs"} rather than |
|
13935 | 180 |
@{term "knows Spy evs"}, which is useful when there are private Notes. |
181 |
Because they depend upon the definition of @{term initState}, they cannot |
|
182 |
be moved up.*} |
|
13922 | 183 |
|
184 |
lemma used_parts_subset_parts [rule_format]: |
|
185 |
"\<forall>X \<in> used evs. parts {X} \<subseteq> used evs" |
|
186 |
apply (induct evs) |
|
187 |
prefer 2 |
|
188 |
apply (simp add: used_Cons) |
|
189 |
apply (rule ballI) |
|
13935 | 190 |
apply (case_tac a, auto) |
191 |
apply (auto dest!: parts_cut) |
|
13922 | 192 |
txt{*Base case*} |
13935 | 193 |
apply (simp add: used_Nil) |
13922 | 194 |
done |
195 |
||
196 |
lemma MPair_used_D: "{|X,Y|} \<in> used H ==> X \<in> used H & Y \<in> used H" |
|
197 |
by (drule used_parts_subset_parts, simp, blast) |
|
198 |
||
199 |
lemma MPair_used [elim!]: |
|
200 |
"[| {|X,Y|} \<in> used H; |
|
201 |
[| X \<in> used H; Y \<in> used H |] ==> P |] |
|
202 |
==> P" |
|
203 |
by (blast dest: MPair_used_D) |
|
204 |
||
205 |
||
206 |
text{*Rewrites should not refer to @{term "initState(Friend i)"} because |
|
207 |
that expression is not in normal form.*} |
|
208 |
||
209 |
lemma keysFor_parts_initState [simp]: "keysFor (parts (initState C)) = {}" |
|
210 |
apply (unfold keysFor_def) |
|
211 |
apply (induct_tac "C") |
|
212 |
apply (auto intro: range_eqI) |
|
213 |
done |
|
214 |
||
215 |
lemma Crypt_notin_initState: "Crypt K X \<notin> parts (initState B)" |
|
216 |
by (induct B, auto) |
|
217 |
||
13935 | 218 |
lemma Crypt_notin_used_empty [simp]: "Crypt K X \<notin> used []" |
219 |
by (simp add: Crypt_notin_initState used_Nil) |
|
220 |
||
13922 | 221 |
|
222 |
(*** Basic properties of shrK ***) |
|
223 |
||
224 |
(*Agents see their own shared keys!*) |
|
225 |
lemma shrK_in_initState [iff]: "Key (shrK A) \<in> initState A" |
|
13926 | 226 |
by (induct_tac "A", auto) |
13922 | 227 |
|
228 |
lemma shrK_in_knows [iff]: "Key (shrK A) \<in> knows A evs" |
|
229 |
by (simp add: initState_subset_knows [THEN subsetD]) |
|
230 |
||
231 |
lemma shrK_in_used [iff]: "Key (shrK A) \<in> used evs" |
|
13926 | 232 |
by (rule initState_into_used, blast) |
13922 | 233 |
|
234 |
(** Fresh keys never clash with long-term shared keys **) |
|
235 |
||
236 |
(*Used in parts_induct_tac and analz_Fake_tac to distinguish session keys |
|
237 |
from long-term shared keys*) |
|
238 |
lemma Key_not_used: "Key K \<notin> used evs ==> K \<notin> range shrK" |
|
13926 | 239 |
by blast |
13922 | 240 |
|
241 |
lemma shrK_neq: "Key K \<notin> used evs ==> shrK B \<noteq> K" |
|
13926 | 242 |
by blast |
13922 | 243 |
|
2318 | 244 |
|
245 |
||
13922 | 246 |
subsection{*Function @{term spies} *} |
247 |
||
248 |
text{*Agents see their own private keys!*} |
|
249 |
lemma priK_in_initState [iff]: "Key (privateKey b A) \<in> initState A" |
|
13926 | 250 |
by (induct_tac "A", auto) |
13922 | 251 |
|
252 |
text{*Agents see all public keys!*} |
|
253 |
lemma publicKey_in_initState [iff]: "Key (publicKey b A) \<in> initState B" |
|
13926 | 254 |
by (case_tac "B", auto) |
13922 | 255 |
|
256 |
text{*All public keys are visible*} |
|
257 |
lemma spies_pubK [iff]: "Key (publicKey b A) \<in> spies evs" |
|
258 |
apply (induct_tac "evs") |
|
259 |
apply (simp_all add: imageI knows_Cons split add: event.split) |
|
260 |
done |
|
261 |
||
262 |
declare spies_pubK [THEN analz.Inj, iff] |
|
263 |
||
264 |
text{*Spy sees private keys of bad agents!*} |
|
265 |
lemma Spy_spies_bad_privateKey [intro!]: |
|
266 |
"A \<in> bad ==> Key (privateKey b A) \<in> spies evs" |
|
267 |
apply (induct_tac "evs") |
|
268 |
apply (simp_all add: imageI knows_Cons split add: event.split) |
|
269 |
done |
|
270 |
||
271 |
text{*Spy sees long-term shared keys of bad agents!*} |
|
272 |
lemma Spy_spies_bad_shrK [intro!]: |
|
273 |
"A \<in> bad ==> Key (shrK A) \<in> spies evs" |
|
274 |
apply (induct_tac "evs") |
|
275 |
apply (simp_all add: imageI knows_Cons split add: event.split) |
|
276 |
done |
|
277 |
||
278 |
lemma publicKey_into_used [iff] :"Key (publicKey b A) \<in> used evs" |
|
279 |
apply (rule initState_into_used) |
|
280 |
apply (rule publicKey_in_initState [THEN parts.Inj]) |
|
281 |
done |
|
282 |
||
283 |
lemma privateKey_into_used [iff]: "Key (privateKey b A) \<in> used evs" |
|
284 |
apply(rule initState_into_used) |
|
285 |
apply(rule priK_in_initState [THEN parts.Inj]) |
|
286 |
done |
|
287 |
||
288 |
||
289 |
subsection{*Fresh Nonces*} |
|
290 |
||
291 |
lemma Nonce_notin_initState [iff]: "Nonce N \<notin> parts (initState B)" |
|
13926 | 292 |
by (induct_tac "B", auto) |
13922 | 293 |
|
294 |
lemma Nonce_notin_used_empty [simp]: "Nonce N \<notin> used []" |
|
13926 | 295 |
by (simp add: used_Nil) |
13922 | 296 |
|
11104 | 297 |
|
13922 | 298 |
subsection{*Supply fresh nonces for possibility theorems*} |
299 |
||
300 |
text{*In any trace, there is an upper bound N on the greatest nonce in use*} |
|
301 |
lemma Nonce_supply_lemma: "EX N. ALL n. N<=n --> Nonce n \<notin> used evs" |
|
302 |
apply (induct_tac "evs") |
|
13926 | 303 |
apply (rule_tac x = 0 in exI) |
13922 | 304 |
apply (simp_all (no_asm_simp) add: used_Cons split add: event.split) |
305 |
apply safe |
|
306 |
apply (rule msg_Nonce_supply [THEN exE], blast elim!: add_leE)+ |
|
307 |
done |
|
308 |
||
309 |
lemma Nonce_supply1: "EX N. Nonce N \<notin> used evs" |
|
13926 | 310 |
by (rule Nonce_supply_lemma [THEN exE], blast) |
13922 | 311 |
|
312 |
lemma Nonce_supply: "Nonce (@ N. Nonce N \<notin> used evs) \<notin> used evs" |
|
313 |
apply (rule Nonce_supply_lemma [THEN exE]) |
|
13926 | 314 |
apply (rule someI, fast) |
13922 | 315 |
done |
316 |
||
317 |
subsection{*Specialized rewriting for the analz_image_... theorems*} |
|
318 |
||
319 |
lemma insert_Key_singleton: "insert (Key K) H = Key ` {K} Un H" |
|
13926 | 320 |
by blast |
11104 | 321 |
|
13922 | 322 |
lemma insert_Key_image: "insert (Key K) (Key`KK \<union> C) = Key ` (insert K KK) \<union> C" |
13926 | 323 |
by blast |
13922 | 324 |
|
325 |
ML |
|
326 |
{* |
|
327 |
val Key_not_used = thm "Key_not_used"; |
|
328 |
val insert_Key_singleton = thm "insert_Key_singleton"; |
|
329 |
val insert_Key_image = thm "insert_Key_image"; |
|
330 |
*} |
|
331 |
||
332 |
(* |
|
333 |
val not_symKeys_pubK = thm "not_symKeys_pubK"; |
|
334 |
val not_symKeys_priK = thm "not_symKeys_priK"; |
|
335 |
val symKeys_neq_imp_neq = thm "symKeys_neq_imp_neq"; |
|
336 |
val analz_symKeys_Decrypt = thm "analz_symKeys_Decrypt"; |
|
337 |
val invKey_image_eq = thm "invKey_image_eq"; |
|
338 |
val pubK_image_eq = thm "pubK_image_eq"; |
|
339 |
val priK_pubK_image_eq = thm "priK_pubK_image_eq"; |
|
340 |
val keysFor_parts_initState = thm "keysFor_parts_initState"; |
|
341 |
val priK_in_initState = thm "priK_in_initState"; |
|
342 |
val spies_pubK = thm "spies_pubK"; |
|
343 |
val Spy_spies_bad = thm "Spy_spies_bad"; |
|
344 |
val Nonce_notin_initState = thm "Nonce_notin_initState"; |
|
345 |
val Nonce_notin_used_empty = thm "Nonce_notin_used_empty"; |
|
346 |
||
347 |
*) |
|
348 |
||
349 |
||
350 |
lemma invKey_K [simp]: "K \<in> symKeys ==> invKey K = K" |
|
11287 | 351 |
by (simp add: symKeys_def) |
352 |
||
13922 | 353 |
lemma Crypt_imp_keysFor :"[|K \<in> symKeys; Crypt K X \<in> H|] ==> K \<in> keysFor H" |
13926 | 354 |
by (drule Crypt_imp_invKey_keysFor, simp) |
13922 | 355 |
|
356 |
||
357 |
subsection{*Specialized Methods for Possibility Theorems*} |
|
358 |
||
359 |
ML |
|
360 |
{* |
|
361 |
val Nonce_supply1 = thm "Nonce_supply1"; |
|
362 |
val Nonce_supply = thm "Nonce_supply"; |
|
363 |
||
364 |
(*Tactic for possibility theorems (Isar interface)*) |
|
365 |
fun gen_possibility_tac ss state = state |> |
|
366 |
REPEAT (*omit used_Says so that Nonces start from different traces!*) |
|
367 |
(ALLGOALS (simp_tac (ss delsimps [used_Says])) |
|
368 |
THEN |
|
369 |
REPEAT_FIRST (eq_assume_tac ORELSE' |
|
370 |
resolve_tac [refl, conjI, Nonce_supply])) |
|
371 |
||
372 |
(*Tactic for possibility theorems (ML script version)*) |
|
373 |
fun possibility_tac state = gen_possibility_tac (simpset()) state |
|
374 |
*} |
|
11104 | 375 |
|
376 |
method_setup possibility = {* |
|
11270
a315a3862bb4
better treatment of methods: uses Method.ctxt_args to refer to current
paulson
parents:
11104
diff
changeset
|
377 |
Method.ctxt_args (fn ctxt => |
a315a3862bb4
better treatment of methods: uses Method.ctxt_args to refer to current
paulson
parents:
11104
diff
changeset
|
378 |
Method.METHOD (fn facts => |
a315a3862bb4
better treatment of methods: uses Method.ctxt_args to refer to current
paulson
parents:
11104
diff
changeset
|
379 |
gen_possibility_tac (Simplifier.get_local_simpset ctxt))) *} |
11104 | 380 |
"for proving possibility theorems" |
2318 | 381 |
|
13922 | 382 |
|
383 |
||
13926 | 384 |
lemmas analz_image_freshK_simps = |
385 |
simp_thms mem_simps --{*these two allow its use with @{text "only:"}*} |
|
386 |
disj_comms |
|
387 |
image_insert [THEN sym] image_Un [THEN sym] empty_subsetI insert_subset |
|
388 |
analz_insert_eq Un_upper2 [THEN analz_mono, THEN [2] rev_subsetD] |
|
389 |
insert_Key_singleton |
|
390 |
Key_not_used insert_Key_image Un_assoc [THEN sym] |
|
391 |
||
13922 | 392 |
ML |
393 |
{* |
|
394 |
val analz_image_freshK_ss = |
|
395 |
simpset() delsimps [image_insert, image_Un] |
|
396 |
delsimps [imp_disjL] (*reduces blow-up*) |
|
13926 | 397 |
addsimps thms"analz_image_freshK_simps" |
13922 | 398 |
*} |
399 |
||
13926 | 400 |
axioms |
401 |
Key_supply_ax: "finite KK ==> \<exists>K\<in>symKeys. K \<notin> KK & Key K \<notin> used evs" |
|
402 |
--{*Unlike the corresponding property of nonces, this cannot be proved. |
|
403 |
We have infinitely many agents and there is nothing to stop their |
|
404 |
long-term keys from exhausting all the natural numbers. The axiom |
|
405 |
assumes that their keys are dispersed so as to leave room for infinitely |
|
406 |
many fresh session keys. We could, alternatively, restrict agents to |
|
407 |
an unspecified finite number. We could however replace @{term"used evs"} |
|
408 |
by @{term "used []"}.*} |
|
409 |
||
410 |
||
411 |
lemma Key_supply1: "\<exists>K\<in>symKeys. Key K \<notin> used evs" |
|
412 |
by (rule Finites.emptyI [THEN Key_supply_ax, THEN bexE], blast) |
|
413 |
||
414 |
||
2318 | 415 |
end |