author | paulson |
Tue, 01 Jul 1997 17:34:42 +0200 | |
changeset 3477 | 3aced7fa7d8b |
parent 3473 | c2334f9532ab |
child 3512 | 9dcb4daa15e8 |
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 |
||
6 |
Theory of Public Keys (common to all symmetric-key protocols) |
|
7 |
||
8 |
Server keys; initial states of agents; new nonces and keys; function "sees" |
|
9 |
*) |
|
10 |
||
11 |
||
12 |
open Public; |
|
13 |
||
14 |
(*Holds because Friend is injective: thus cannot prove for all f*) |
|
15 |
goal thy "(Friend x : Friend``A) = (x:A)"; |
|
16 |
by (Auto_tac()); |
|
17 |
qed "Friend_image_eq"; |
|
18 |
Addsimps [Friend_image_eq]; |
|
19 |
||
20 |
Addsimps [Un_insert_left, Un_insert_right]; |
|
21 |
||
22 |
(*By default only o_apply is built-in. But in the presence of eta-expansion |
|
23 |
this means that some terms displayed as (f o g) will be rewritten, and others |
|
24 |
will not!*) |
|
25 |
Addsimps [o_def]; |
|
26 |
||
27 |
goalw thy [keysFor_def] "keysFor (parts (initState lost C)) = {}"; |
|
28 |
by (agent.induct_tac "C" 1); |
|
29 |
by (auto_tac (!claset addIs [range_eqI], !simpset)); |
|
30 |
qed "keysFor_parts_initState"; |
|
31 |
Addsimps [keysFor_parts_initState]; |
|
32 |
||
33 |
goalw thy [keysFor_def] "keysFor (Key``E) = {}"; |
|
34 |
by (Auto_tac ()); |
|
35 |
qed "keysFor_image_Key"; |
|
36 |
Addsimps [keysFor_image_Key]; |
|
37 |
||
38 |
||
39 |
(*** Function "sees" ***) |
|
40 |
||
41 |
goal thy |
|
42 |
"!!evs. lost' <= lost ==> sees lost' A evs <= sees lost A evs"; |
|
43 |
by (list.induct_tac "evs" 1); |
|
44 |
by (agent.induct_tac "A" 1); |
|
45 |
by (event.induct_tac "a" 2); |
|
46 |
by (Auto_tac ()); |
|
47 |
qed "sees_mono"; |
|
48 |
||
2497 | 49 |
(*** Basic properties of pubK & priK ***) |
50 |
||
51 |
AddIffs [inj_pubK RS inj_eq]; |
|
3477 | 52 |
|
53 |
goal thy "!!A B. (priK A = priK B) = (A=B)"; |
|
54 |
by (Step_tac 1); |
|
55 |
by (dres_inst_tac [("f","invKey")] arg_cong 1); |
|
56 |
by (Full_simp_tac 1); |
|
57 |
qed "priK_inj_eq"; |
|
58 |
||
59 |
AddIffs [priK_inj_eq]; |
|
3473
c2334f9532ab
New and stronger lemmas; more default simp/cla rules
paulson
parents:
3465
diff
changeset
|
60 |
AddIffs [priK_neq_pubK, priK_neq_pubK RS not_sym]; |
c2334f9532ab
New and stronger lemmas; more default simp/cla rules
paulson
parents:
3465
diff
changeset
|
61 |
|
c2334f9532ab
New and stronger lemmas; more default simp/cla rules
paulson
parents:
3465
diff
changeset
|
62 |
goalw thy [isSymKey_def] "~ isSymKey (pubK A)"; |
c2334f9532ab
New and stronger lemmas; more default simp/cla rules
paulson
parents:
3465
diff
changeset
|
63 |
by (Simp_tac 1); |
c2334f9532ab
New and stronger lemmas; more default simp/cla rules
paulson
parents:
3465
diff
changeset
|
64 |
qed "not_isSymKey_pubK"; |
c2334f9532ab
New and stronger lemmas; more default simp/cla rules
paulson
parents:
3465
diff
changeset
|
65 |
|
c2334f9532ab
New and stronger lemmas; more default simp/cla rules
paulson
parents:
3465
diff
changeset
|
66 |
goalw thy [isSymKey_def] "~ isSymKey (priK A)"; |
c2334f9532ab
New and stronger lemmas; more default simp/cla rules
paulson
parents:
3465
diff
changeset
|
67 |
by (Simp_tac 1); |
c2334f9532ab
New and stronger lemmas; more default simp/cla rules
paulson
parents:
3465
diff
changeset
|
68 |
qed "not_isSymKey_priK"; |
c2334f9532ab
New and stronger lemmas; more default simp/cla rules
paulson
parents:
3465
diff
changeset
|
69 |
|
c2334f9532ab
New and stronger lemmas; more default simp/cla rules
paulson
parents:
3465
diff
changeset
|
70 |
AddIffs [not_isSymKey_pubK, not_isSymKey_priK]; |
c2334f9532ab
New and stronger lemmas; more default simp/cla rules
paulson
parents:
3465
diff
changeset
|
71 |
|
2497 | 72 |
|
2318 | 73 |
(*Agents see their own private keys!*) |
74 |
goal thy "A ~= Spy --> Key (priK A) : sees lost A evs"; |
|
75 |
by (list.induct_tac "evs" 1); |
|
76 |
by (agent.induct_tac "A" 1); |
|
77 |
by (Auto_tac ()); |
|
78 |
qed_spec_mp "sees_own_priK"; |
|
79 |
||
3473
c2334f9532ab
New and stronger lemmas; more default simp/cla rules
paulson
parents:
3465
diff
changeset
|
80 |
(*All public keys are visible to all*) |
c2334f9532ab
New and stronger lemmas; more default simp/cla rules
paulson
parents:
3465
diff
changeset
|
81 |
goal thy "Key (pubK A) : sees lost B evs"; |
2318 | 82 |
by (list.induct_tac "evs" 1); |
3473
c2334f9532ab
New and stronger lemmas; more default simp/cla rules
paulson
parents:
3465
diff
changeset
|
83 |
by (agent.induct_tac "B" 1); |
2318 | 84 |
by (Auto_tac ()); |
85 |
qed_spec_mp "sees_pubK"; |
|
86 |
||
87 |
(*Spy sees private keys of lost agents!*) |
|
88 |
goal thy "!!A. A: lost ==> Key (priK A) : sees lost Spy evs"; |
|
89 |
by (list.induct_tac "evs" 1); |
|
90 |
by (Auto_tac()); |
|
91 |
qed "Spy_sees_lost"; |
|
92 |
||
3473
c2334f9532ab
New and stronger lemmas; more default simp/cla rules
paulson
parents:
3465
diff
changeset
|
93 |
AddIffs [sees_pubK, sees_pubK RS analz.Inj]; |
2318 | 94 |
AddSIs [sees_own_priK, Spy_sees_lost]; |
95 |
||
96 |
||
97 |
(** Specialized rewrite rules for (sees lost A (Says...#evs)) **) |
|
98 |
||
99 |
goal thy "sees lost B (Says A B X # evs) = insert X (sees lost B evs)"; |
|
100 |
by (Simp_tac 1); |
|
101 |
qed "sees_own"; |
|
102 |
||
103 |
goal thy "!!A. Server ~= B ==> \ |
|
104 |
\ sees lost Server (Says A B X # evs) = sees lost Server evs"; |
|
105 |
by (Asm_simp_tac 1); |
|
106 |
qed "sees_Server"; |
|
107 |
||
108 |
goal thy "!!A. Friend i ~= B ==> \ |
|
109 |
\ sees lost (Friend i) (Says A B X # evs) = sees lost (Friend i) evs"; |
|
110 |
by (Asm_simp_tac 1); |
|
111 |
qed "sees_Friend"; |
|
112 |
||
113 |
goal thy "sees lost Spy (Says A B X # evs) = insert X (sees lost Spy evs)"; |
|
114 |
by (Simp_tac 1); |
|
115 |
qed "sees_Spy"; |
|
116 |
||
117 |
goal thy "sees lost A (Says A' B X # evs) <= insert X (sees lost A evs)"; |
|
118 |
by (simp_tac (!simpset setloop split_tac [expand_if]) 1); |
|
119 |
by (Fast_tac 1); |
|
120 |
qed "sees_Says_subset_insert"; |
|
121 |
||
122 |
goal thy "sees lost A evs <= sees lost A (Says A' B X # evs)"; |
|
123 |
by (simp_tac (!simpset setloop split_tac [expand_if]) 1); |
|
124 |
by (Fast_tac 1); |
|
125 |
qed "sees_subset_sees_Says"; |
|
126 |
||
127 |
(*Pushing Unions into parts. One of the agents A is B, and thus sees Y. |
|
128 |
Once used to prove new_keys_not_seen; now obsolete.*) |
|
129 |
goal thy "(UN A. parts (sees lost A (Says B C Y # evs))) = \ |
|
130 |
\ parts {Y} Un (UN A. parts (sees lost A evs))"; |
|
131 |
by (Step_tac 1); |
|
132 |
by (etac rev_mp 1); (*split_tac does not work on assumptions*) |
|
133 |
by (ALLGOALS |
|
3207 | 134 |
(fast_tac (!claset addss (!simpset addsimps [parts_Un, sees_Cons] |
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
135 |
setloop split_tac [expand_if])))); |
2318 | 136 |
qed "UN_parts_sees_Says"; |
137 |
||
3465 | 138 |
goal thy "Says A B X : set evs --> X : sees lost Spy evs"; |
2318 | 139 |
by (list.induct_tac "evs" 1); |
140 |
by (Auto_tac ()); |
|
141 |
qed_spec_mp "Says_imp_sees_Spy"; |
|
142 |
||
2537 | 143 |
(*Use with addSEs to derive contradictions from old Says events containing |
144 |
items known to be fresh*) |
|
145 |
val sees_Spy_partsEs = make_elim (Says_imp_sees_Spy RS parts.Inj):: partsEs; |
|
146 |
||
3442 | 147 |
(*For not_lost_tac*) |
148 |
goal thy "!!A. [| Crypt (pubK A) X : analz (sees lost Spy evs); A: lost |] \ |
|
149 |
\ ==> X : analz (sees lost Spy evs)"; |
|
150 |
by (blast_tac (!claset addSDs [analz.Decrypt]) 1); |
|
151 |
qed "Crypt_Spy_analz_lost"; |
|
2318 | 152 |
|
3442 | 153 |
(*Prove that the agent is uncompromised by the confidentiality of |
154 |
a component of a message she's said.*) |
|
155 |
fun not_lost_tac s = |
|
156 |
case_tac ("(" ^ s ^ ") : lost") THEN' |
|
157 |
SELECT_GOAL |
|
158 |
(REPEAT_DETERM (dtac (Says_imp_sees_Spy RS analz.Inj) 1) THEN |
|
159 |
REPEAT_DETERM (etac MPair_analz 1) THEN |
|
160 |
THEN_BEST_FIRST |
|
161 |
(dres_inst_tac [("A", s)] Crypt_Spy_analz_lost 1 THEN assume_tac 1) |
|
162 |
(has_fewer_prems 1, size_of_thm) |
|
163 |
(Step_tac 1)); |
|
2318 | 164 |
|
165 |
Addsimps [sees_own, sees_Server, sees_Friend, sees_Spy]; |
|
166 |
Delsimps [sees_Cons]; (**** NOTE REMOVAL -- laws above are cleaner ****) |
|
167 |
||
168 |
||
2497 | 169 |
(*** Fresh nonces ***) |
170 |
||
171 |
goal thy "Nonce N ~: parts (initState lost B)"; |
|
172 |
by (agent.induct_tac "B" 1); |
|
173 |
by (Auto_tac ()); |
|
174 |
qed "Nonce_notin_initState"; |
|
175 |
||
176 |
AddIffs [Nonce_notin_initState]; |
|
177 |
||
178 |
goalw thy [used_def] "!!X. X: parts (sees lost B evs) ==> X: used evs"; |
|
2516
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2497
diff
changeset
|
179 |
by (etac (impOfSubs parts_mono) 1); |
2497 | 180 |
by (Fast_tac 1); |
181 |
qed "usedI"; |
|
182 |
||
183 |
AddIs [usedI]; |
|
184 |
||
2516
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2497
diff
changeset
|
185 |
(** A supply of fresh nonces for possibility theorems. **) |
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2497
diff
changeset
|
186 |
|
2497 | 187 |
goalw thy [used_def] "EX N. ALL n. N<=n --> Nonce n ~: used evs"; |
188 |
by (list.induct_tac "evs" 1); |
|
189 |
by (res_inst_tac [("x","0")] exI 1); |
|
190 |
by (Step_tac 1); |
|
191 |
by (Full_simp_tac 1); |
|
2516
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2497
diff
changeset
|
192 |
(*Inductive step*) |
2497 | 193 |
by (event.induct_tac "a" 1); |
194 |
by (full_simp_tac (!simpset addsimps [UN_parts_sees_Says]) 1); |
|
195 |
by (msg.induct_tac "msg" 1); |
|
196 |
by (ALLGOALS (asm_simp_tac (!simpset addsimps [exI, parts_insert2]))); |
|
197 |
by (Step_tac 1); |
|
198 |
(*MPair case*) |
|
199 |
by (res_inst_tac [("x","Na+Nb")] exI 2); |
|
200 |
by (fast_tac (!claset addSEs [add_leE]) 2); |
|
201 |
(*Nonce case*) |
|
202 |
by (res_inst_tac [("x","N + Suc nat")] exI 1); |
|
2637
e9b203f854ae
reflecting my recent changes of the simplifier and classical reasoner
oheimb
parents:
2537
diff
changeset
|
203 |
by (fast_tac (!claset addSEs [add_leE] addaltern trans_tac) 1); |
2497 | 204 |
val lemma = result(); |
205 |
||
206 |
goal thy "Nonce (@ N. Nonce N ~: used evs) ~: used evs"; |
|
2516
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2497
diff
changeset
|
207 |
by (rtac (lemma RS exE) 1); |
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2497
diff
changeset
|
208 |
by (rtac selectI 1); |
2497 | 209 |
by (Fast_tac 1); |
210 |
qed "Nonce_supply"; |
|
211 |
||
2516
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2497
diff
changeset
|
212 |
(*Tactic for possibility theorems*) |
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2497
diff
changeset
|
213 |
val possibility_tac = |
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2497
diff
changeset
|
214 |
REPEAT |
2637
e9b203f854ae
reflecting my recent changes of the simplifier and classical reasoner
oheimb
parents:
2537
diff
changeset
|
215 |
(ALLGOALS (simp_tac (!simpset setSolver safe_solver)) |
2516
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2497
diff
changeset
|
216 |
THEN |
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2497
diff
changeset
|
217 |
REPEAT_FIRST (eq_assume_tac ORELSE' |
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2497
diff
changeset
|
218 |
resolve_tac [refl, conjI, Nonce_supply])); |
2497 | 219 |
|
2318 | 220 |
(** Power of the Spy **) |
221 |
||
222 |
(*The Spy can see more than anybody else, except for their initial state*) |
|
223 |
goal thy "sees lost A evs <= initState lost A Un sees lost Spy evs"; |
|
224 |
by (list.induct_tac "evs" 1); |
|
225 |
by (event.induct_tac "a" 2); |
|
226 |
by (ALLGOALS (fast_tac (!claset addDs [sees_Says_subset_insert RS subsetD] |
|
227 |
addss (!simpset)))); |
|
228 |
qed "sees_agent_subset_sees_Spy"; |
|
229 |
||
230 |
(*The Spy can see more than anybody else who's lost their key!*) |
|
231 |
goal thy "A: lost --> A ~= Server --> sees lost A evs <= sees lost Spy evs"; |
|
232 |
by (list.induct_tac "evs" 1); |
|
233 |
by (event.induct_tac "a" 2); |
|
234 |
by (agent.induct_tac "A" 1); |
|
235 |
by (auto_tac (!claset addDs [sees_Says_subset_insert RS subsetD], (!simpset))); |
|
236 |
qed_spec_mp "sees_lost_agent_subset_sees_Spy"; |
|
237 |
||
238 |
||
239 |
(** Simplifying parts (insert X (sees lost A evs)) |
|
240 |
= parts {X} Un parts (sees lost A evs) -- since general case loops*) |
|
241 |
||
242 |
val parts_insert_sees = |
|
243 |
parts_insert |> read_instantiate_sg (sign_of thy) |
|
244 |
[("H", "sees lost A evs")] |
|
245 |
|> standard; |