| author | paulson |
| Tue, 01 Jul 1997 17:34:13 +0200 | |
| changeset 3476 | 1be4fee7606b |
| parent 3466 | 30791e5a69c4 |
| child 3516 | 470626799511 |
| permissions | -rw-r--r-- |
| 2318 | 1 |
(* Title: HOL/Auth/NS_Public |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1996 University of Cambridge |
|
5 |
||
6 |
Inductive relation "ns_public" for the Needham-Schroeder Public-Key protocol. |
|
7 |
Version incorporating Lowe's fix (inclusion of B's identify in round 2). |
|
8 |
*) |
|
9 |
||
10 |
open NS_Public; |
|
11 |
||
12 |
proof_timing:=true; |
|
13 |
HOL_quantifiers := false; |
|
14 |
||
15 |
AddIffs [Spy_in_lost]; |
|
16 |
||
17 |
(*Replacing the variable by a constant improves search speed by 50%!*) |
|
18 |
val Says_imp_sees_Spy' = |
|
19 |
read_instantiate_sg (sign_of thy) [("lost","lost")] Says_imp_sees_Spy;
|
|
20 |
||
21 |
(*A "possibility property": there are traces that reach the end*) |
|
22 |
goal thy |
|
23 |
"!!A B. A ~= B ==> EX NB. EX evs: ns_public. \ |
|
| 3465 | 24 |
\ Says A B (Crypt (pubK B) (Nonce NB)) : set evs"; |
| 2318 | 25 |
by (REPEAT (resolve_tac [exI,bexI] 1)); |
26 |
by (rtac (ns_public.Nil RS ns_public.NS1 RS ns_public.NS2 RS ns_public.NS3) 2); |
|
|
2516
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2497
diff
changeset
|
27 |
by possibility_tac; |
| 2318 | 28 |
result(); |
29 |
||
30 |
||
31 |
(**** Inductive proofs about ns_public ****) |
|
32 |
||
33 |
(*Nobody sends themselves messages*) |
|
| 3465 | 34 |
goal thy "!!evs. evs : ns_public ==> ALL A X. Says A A X ~: set evs"; |
| 2318 | 35 |
by (etac ns_public.induct 1); |
36 |
by (Auto_tac()); |
|
37 |
qed_spec_mp "not_Says_to_self"; |
|
38 |
Addsimps [not_Says_to_self]; |
|
39 |
AddSEs [not_Says_to_self RSN (2, rev_notE)]; |
|
40 |
||
41 |
||
42 |
(** Theorems of the form X ~: parts (sees lost Spy evs) imply that NOBODY |
|
43 |
sends messages containing X! **) |
|
44 |
||
45 |
(*Spy never sees another agent's private key! (unless it's lost at start)*) |
|
46 |
goal thy |
|
47 |
"!!evs. evs : ns_public \ |
|
48 |
\ ==> (Key (priK A) : parts (sees lost Spy evs)) = (A : lost)"; |
|
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
49 |
by (etac ns_public.induct 1); |
|
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
50 |
by (prove_simple_subgoals_tac 1); |
|
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
51 |
by (Fake_parts_insert_tac 1); |
| 2318 | 52 |
qed "Spy_see_priK"; |
53 |
Addsimps [Spy_see_priK]; |
|
54 |
||
55 |
goal thy |
|
56 |
"!!evs. evs : ns_public \ |
|
57 |
\ ==> (Key (priK A) : analz (sees lost Spy evs)) = (A : lost)"; |
|
58 |
by (auto_tac(!claset addDs [impOfSubs analz_subset_parts], !simpset)); |
|
59 |
qed "Spy_analz_priK"; |
|
60 |
Addsimps [Spy_analz_priK]; |
|
61 |
||
62 |
goal thy "!!A. [| Key (priK A) : parts (sees lost Spy evs); \ |
|
63 |
\ evs : ns_public |] ==> A:lost"; |
|
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
64 |
by (blast_tac (!claset addDs [Spy_see_priK]) 1); |
| 2318 | 65 |
qed "Spy_see_priK_D"; |
66 |
||
67 |
bind_thm ("Spy_analz_priK_D", analz_subset_parts RS subsetD RS Spy_see_priK_D);
|
|
68 |
AddSDs [Spy_see_priK_D, Spy_analz_priK_D]; |
|
69 |
||
70 |
||
|
2418
6b6a92d05fb2
New tactics: prove_unique_tac and analz_induct_tac
paulson
parents:
2374
diff
changeset
|
71 |
fun analz_induct_tac i = |
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
72 |
etac ns_public.induct i THEN |
|
2418
6b6a92d05fb2
New tactics: prove_unique_tac and analz_induct_tac
paulson
parents:
2374
diff
changeset
|
73 |
ALLGOALS (asm_simp_tac |
|
2516
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2497
diff
changeset
|
74 |
(!simpset addsimps [not_parts_not_analz] |
|
2451
ce85a2aafc7a
Extensive tidying and simplification, largely stemming from
paulson
parents:
2418
diff
changeset
|
75 |
setloop split_tac [expand_if])); |
| 2318 | 76 |
|
77 |
(**** Authenticity properties obtained from NS2 ****) |
|
78 |
||
79 |
(*It is impossible to re-use a nonce in both NS1 and NS2, provided the nonce |
|
80 |
is secret. (Honest users generate fresh nonces.)*) |
|
81 |
goal thy |
|
|
2536
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
82 |
"!!evs. [| Nonce NA ~: analz (sees lost Spy evs); \ |
|
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
83 |
\ Crypt (pubK B) {|Nonce NA, Agent A|} : parts (sees lost Spy evs); \
|
|
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
84 |
\ evs : ns_public |] \ |
|
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
85 |
\ ==> Crypt (pubK C) {|NA', Nonce NA, Agent D|} ~: parts (sees lost Spy evs)";
|
|
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
86 |
by (etac rev_mp 1); |
|
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
87 |
by (etac rev_mp 1); |
|
2418
6b6a92d05fb2
New tactics: prove_unique_tac and analz_induct_tac
paulson
parents:
2374
diff
changeset
|
88 |
by (analz_induct_tac 1); |
| 2318 | 89 |
(*NS3*) |
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
90 |
by (blast_tac (!claset addSEs partsEs) 4); |
| 2318 | 91 |
(*NS2*) |
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
92 |
by (blast_tac (!claset addSEs partsEs) 3); |
| 2318 | 93 |
(*Fake*) |
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
94 |
by (blast_tac (!claset addSIs [analz_insertI] |
|
2536
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
95 |
addDs [impOfSubs analz_subset_parts, |
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
96 |
impOfSubs Fake_parts_insert]) 2); |
| 2318 | 97 |
(*Base*) |
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
98 |
by (Blast_tac 1); |
|
2536
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
99 |
qed "no_nonce_NS1_NS2"; |
| 2318 | 100 |
|
101 |
||
| 2480 | 102 |
(*Unicity for NS1: nonce NA identifies agents A and B*) |
| 2318 | 103 |
goal thy |
|
2536
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
104 |
"!!evs. [| Nonce NA ~: analz (sees lost Spy evs); evs : ns_public |] \ |
|
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
105 |
\ ==> EX A' B'. ALL A B. \ |
| 2318 | 106 |
\ Crypt (pubK B) {|Nonce NA, Agent A|} : parts (sees lost Spy evs) --> \
|
|
2536
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
107 |
\ A=A' & B=B'"; |
|
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
108 |
by (etac rev_mp 1); |
|
2418
6b6a92d05fb2
New tactics: prove_unique_tac and analz_induct_tac
paulson
parents:
2374
diff
changeset
|
109 |
by (analz_induct_tac 1); |
| 2318 | 110 |
(*NS1*) |
| 2497 | 111 |
by (simp_tac (!simpset addsimps [all_conj_distrib]) 3); |
| 2318 | 112 |
by (expand_case_tac "NA = ?y" 3 THEN |
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
113 |
REPEAT (blast_tac (!claset addSEs partsEs) 3)); |
| 2318 | 114 |
(*Base*) |
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
115 |
by (Blast_tac 1); |
| 2318 | 116 |
(*Fake*) |
| 2497 | 117 |
by (simp_tac (!simpset addsimps [all_conj_distrib, parts_insert_sees]) 1); |
| 2374 | 118 |
by (step_tac (!claset addSIs [analz_insertI]) 1); |
| 2318 | 119 |
by (ex_strip_tac 1); |
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
120 |
by (blast_tac (!claset delrules [conjI] |
|
2516
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2497
diff
changeset
|
121 |
addSDs [impOfSubs Fake_parts_insert] |
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
122 |
addDs [impOfSubs analz_subset_parts]) 1); |
| 2318 | 123 |
val lemma = result(); |
124 |
||
125 |
goal thy |
|
126 |
"!!evs. [| Crypt(pubK B) {|Nonce NA, Agent A|} : parts(sees lost Spy evs); \
|
|
127 |
\ Crypt(pubK B') {|Nonce NA, Agent A'|} : parts(sees lost Spy evs); \
|
|
128 |
\ Nonce NA ~: analz (sees lost Spy evs); \ |
|
129 |
\ evs : ns_public |] \ |
|
130 |
\ ==> A=A' & B=B'"; |
|
| 2480 | 131 |
by (prove_unique_tac lemma 1); |
| 2318 | 132 |
qed "unique_NA"; |
133 |
||
134 |
||
135 |
(*Secrecy: Spy does not see the nonce sent in msg NS1 if A and B are secure*) |
|
136 |
goal thy |
|
|
3466
30791e5a69c4
Corrected indentations and margins after the renaming of "set_of_list"
paulson
parents:
3465
diff
changeset
|
137 |
"!!evs. [| Says A B (Crypt(pubK B) {|Nonce NA, Agent A|}) : set evs; \
|
|
2536
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
138 |
\ A ~: lost; B ~: lost; evs : ns_public |] \ |
|
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
139 |
\ ==> Nonce NA ~: analz (sees lost Spy evs)"; |
|
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
140 |
by (etac rev_mp 1); |
|
2418
6b6a92d05fb2
New tactics: prove_unique_tac and analz_induct_tac
paulson
parents:
2374
diff
changeset
|
141 |
by (analz_induct_tac 1); |
| 2318 | 142 |
(*NS3*) |
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
143 |
by (blast_tac (!claset addDs [Says_imp_sees_Spy' RS parts.Inj] |
|
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
144 |
addEs [no_nonce_NS1_NS2 RSN (2, rev_notE)]) 4); |
|
2536
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
145 |
(*NS2*) |
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
146 |
by (blast_tac (!claset addSEs [MPair_parts] |
|
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
147 |
addDs [Says_imp_sees_Spy' RS parts.Inj, |
|
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
148 |
parts.Body, unique_NA]) 3); |
| 2318 | 149 |
(*NS1*) |
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
150 |
by (blast_tac (!claset addSEs sees_Spy_partsEs |
| 2497 | 151 |
addIs [impOfSubs analz_subset_parts]) 2); |
| 2318 | 152 |
(*Fake*) |
| 2497 | 153 |
by (spy_analz_tac 1); |
|
2536
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
154 |
qed "Spy_not_see_NA"; |
| 2318 | 155 |
|
156 |
||
157 |
(*Authentication for A: if she receives message 2 and has used NA |
|
158 |
to start a run, then B has sent message 2.*) |
|
159 |
goal thy |
|
|
3466
30791e5a69c4
Corrected indentations and margins after the renaming of "set_of_list"
paulson
parents:
3465
diff
changeset
|
160 |
"!!evs. [| Says A B (Crypt (pubK B) {|Nonce NA, Agent A|}) : set evs; \
|
|
30791e5a69c4
Corrected indentations and margins after the renaming of "set_of_list"
paulson
parents:
3465
diff
changeset
|
161 |
\ Says B' A (Crypt(pubK A) {|Nonce NA, Nonce NB, Agent B|}) \
|
|
30791e5a69c4
Corrected indentations and margins after the renaming of "set_of_list"
paulson
parents:
3465
diff
changeset
|
162 |
\ : set evs; \ |
|
30791e5a69c4
Corrected indentations and margins after the renaming of "set_of_list"
paulson
parents:
3465
diff
changeset
|
163 |
\ A ~: lost; B ~: lost; evs : ns_public |] \ |
|
30791e5a69c4
Corrected indentations and margins after the renaming of "set_of_list"
paulson
parents:
3465
diff
changeset
|
164 |
\ ==> Says B A (Crypt(pubK A) {|Nonce NA, Nonce NB, Agent B|}) \
|
| 3465 | 165 |
\ : set evs"; |
|
2536
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
166 |
by (etac rev_mp 1); |
|
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
167 |
(*prepare induction over Crypt (pubK A) {|NA,NB,B|} : parts H*)
|
|
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
168 |
by (etac (Says_imp_sees_Spy' RS parts.Inj RS rev_mp) 1); |
| 2318 | 169 |
by (etac ns_public.induct 1); |
170 |
by (ALLGOALS Asm_simp_tac); |
|
171 |
(*NS1*) |
|
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
172 |
by (blast_tac (!claset addSEs sees_Spy_partsEs) 2); |
| 2318 | 173 |
(*Fake*) |
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
174 |
by (blast_tac (!claset addSDs [impOfSubs Fake_parts_insert] |
|
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
175 |
addDs [Spy_not_see_NA, |
|
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
176 |
impOfSubs analz_subset_parts]) 1); |
| 2318 | 177 |
qed "A_trusts_NS2"; |
178 |
||
179 |
(*If the encrypted message appears then it originated with Alice in NS1*) |
|
180 |
goal thy |
|
|
2536
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
181 |
"!!evs. [| Crypt (pubK B) {|Nonce NA, Agent A|} : parts (sees lost Spy evs); \
|
|
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
182 |
\ Nonce NA ~: analz (sees lost Spy evs); \ |
|
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
183 |
\ evs : ns_public |] \ |
| 3465 | 184 |
\ ==> Says A B (Crypt (pubK B) {|Nonce NA, Agent A|}) : set evs";
|
|
2536
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
185 |
by (etac rev_mp 1); |
|
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
186 |
by (etac rev_mp 1); |
| 2480 | 187 |
by (analz_induct_tac 1); |
| 2318 | 188 |
(*Fake*) |
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
189 |
by (blast_tac (!claset addSDs [impOfSubs Fake_parts_insert] |
|
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
190 |
addIs [analz_insertI] |
|
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
191 |
addDs [impOfSubs analz_subset_parts]) 2); |
| 2318 | 192 |
(*Base*) |
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
193 |
by (Blast_tac 1); |
|
2536
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
194 |
qed "B_trusts_NS1"; |
| 2318 | 195 |
|
196 |
||
197 |
||
198 |
(**** Authenticity properties obtained from NS2 ****) |
|
199 |
||
| 2480 | 200 |
(*Unicity for NS2: nonce NB identifies nonce NA and agents A, B |
| 2318 | 201 |
[unicity of B makes Lowe's fix work] |
202 |
[proof closely follows that for unique_NA] *) |
|
203 |
goal thy |
|
|
2536
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
204 |
"!!evs. [| Nonce NB ~: analz (sees lost Spy evs); evs : ns_public |] \ |
|
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
205 |
\ ==> EX A' NA' B'. ALL A NA B. \ |
|
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
206 |
\ Crypt (pubK A) {|Nonce NA, Nonce NB, Agent B|} \
|
|
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
207 |
\ : parts (sees lost Spy evs) --> A=A' & NA=NA' & B=B'"; |
|
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
208 |
by (etac rev_mp 1); |
|
2418
6b6a92d05fb2
New tactics: prove_unique_tac and analz_induct_tac
paulson
parents:
2374
diff
changeset
|
209 |
by (analz_induct_tac 1); |
| 2318 | 210 |
(*NS2*) |
| 2497 | 211 |
by (simp_tac (!simpset addsimps [all_conj_distrib]) 3); |
| 2318 | 212 |
by (expand_case_tac "NB = ?y" 3 THEN |
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
213 |
REPEAT (blast_tac (!claset addSEs partsEs) 3)); |
| 2318 | 214 |
(*Base*) |
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
215 |
by (Blast_tac 1); |
| 2318 | 216 |
(*Fake*) |
| 2497 | 217 |
by (simp_tac (!simpset addsimps [all_conj_distrib, parts_insert_sees]) 1); |
| 2374 | 218 |
by (step_tac (!claset addSIs [analz_insertI]) 1); |
| 2318 | 219 |
by (ex_strip_tac 1); |
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
220 |
by (blast_tac (!claset delrules [conjI] |
|
2516
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2497
diff
changeset
|
221 |
addSDs [impOfSubs Fake_parts_insert] |
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
222 |
addDs [impOfSubs analz_subset_parts]) 1); |
| 2318 | 223 |
val lemma = result(); |
224 |
||
225 |
goal thy |
|
226 |
"!!evs. [| Crypt(pubK A) {|Nonce NA, Nonce NB, Agent B|} \
|
|
227 |
\ : parts(sees lost Spy evs); \ |
|
228 |
\ Crypt(pubK A') {|Nonce NA', Nonce NB, Agent B'|} \
|
|
229 |
\ : parts(sees lost Spy evs); \ |
|
230 |
\ Nonce NB ~: analz (sees lost Spy evs); \ |
|
231 |
\ evs : ns_public |] \ |
|
232 |
\ ==> A=A' & NA=NA' & B=B'"; |
|
|
2418
6b6a92d05fb2
New tactics: prove_unique_tac and analz_induct_tac
paulson
parents:
2374
diff
changeset
|
233 |
by (prove_unique_tac lemma 1); |
| 2318 | 234 |
qed "unique_NB"; |
235 |
||
236 |
||
237 |
(*Secrecy: Spy does not see the nonce sent in msg NS2 if A and B are secure*) |
|
238 |
goal thy |
|
|
2536
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
239 |
"!!evs. [| Says B A (Crypt (pubK A) {|Nonce NA, Nonce NB, Agent B|}) \
|
|
3466
30791e5a69c4
Corrected indentations and margins after the renaming of "set_of_list"
paulson
parents:
3465
diff
changeset
|
240 |
\ : set evs; \ |
|
2536
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
241 |
\ A ~: lost; B ~: lost; evs : ns_public |] \ |
|
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
242 |
\ ==> Nonce NB ~: analz (sees lost Spy evs)"; |
|
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
243 |
by (etac rev_mp 1); |
|
2418
6b6a92d05fb2
New tactics: prove_unique_tac and analz_induct_tac
paulson
parents:
2374
diff
changeset
|
244 |
by (analz_induct_tac 1); |
| 2318 | 245 |
(*NS3*) |
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
246 |
by (blast_tac (!claset addDs [Says_imp_sees_Spy' RS parts.Inj, |
|
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
247 |
unique_NB]) 4); |
| 2318 | 248 |
(*NS1*) |
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
249 |
by (blast_tac (!claset addSEs sees_Spy_partsEs) 2); |
| 2318 | 250 |
(*Fake*) |
| 2497 | 251 |
by (spy_analz_tac 1); |
| 2318 | 252 |
(*NS2*) |
253 |
by (Step_tac 1); |
|
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
254 |
by (blast_tac (!claset addSEs sees_Spy_partsEs) 3); |
|
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
255 |
by (blast_tac (!claset addSDs [Says_imp_sees_Spy' RS parts.Inj] |
|
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
256 |
addEs [no_nonce_NS1_NS2 RSN (2, rev_notE)]) 2); |
|
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
257 |
by (blast_tac (!claset addSIs [impOfSubs analz_subset_parts]) 1); |
|
2536
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
258 |
qed "Spy_not_see_NB"; |
| 2318 | 259 |
|
260 |
||
261 |
(*Authentication for B: if he receives message 3 and has used NB |
|
262 |
in message 2, then A has sent message 3.*) |
|
263 |
goal thy |
|
|
2536
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
264 |
"!!evs. [| Says B A (Crypt (pubK A) {|Nonce NA, Nonce NB, Agent B|}) \
|
|
3466
30791e5a69c4
Corrected indentations and margins after the renaming of "set_of_list"
paulson
parents:
3465
diff
changeset
|
265 |
\ : set evs; \ |
|
30791e5a69c4
Corrected indentations and margins after the renaming of "set_of_list"
paulson
parents:
3465
diff
changeset
|
266 |
\ Says A' B (Crypt (pubK B) (Nonce NB)): set evs; \ |
|
2536
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
267 |
\ A ~: lost; B ~: lost; evs : ns_public |] \ |
| 3465 | 268 |
\ ==> Says A B (Crypt (pubK B) (Nonce NB)) : set evs"; |
|
2536
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
269 |
by (etac rev_mp 1); |
|
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
270 |
(*prepare induction over Crypt (pubK B) NB : parts H*) |
|
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
271 |
by (etac (Says_imp_sees_Spy' RS parts.Inj RS rev_mp) 1); |
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
272 |
by (analz_induct_tac 1); |
| 2318 | 273 |
(*NS1*) |
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
274 |
by (blast_tac (!claset addSEs sees_Spy_partsEs) 2); |
| 2318 | 275 |
(*Fake*) |
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
276 |
by (blast_tac (!claset addSDs [impOfSubs Fake_parts_insert] |
|
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
277 |
addDs [Spy_not_see_NB, |
|
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
278 |
impOfSubs analz_subset_parts]) 1); |
|
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
279 |
(*NS3; not clear why blast_tac needs to be preceeded by Step_tac*) |
| 2318 | 280 |
by (Step_tac 1); |
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
281 |
by (blast_tac (!claset addDs [Says_imp_sees_Spy' RS parts.Inj, |
|
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
282 |
Spy_not_see_NB, unique_NB]) 1); |
| 2318 | 283 |
qed "B_trusts_NS3"; |
284 |
||
285 |
||
286 |
(**** Overall guarantee for B*) |
|
287 |
||
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
288 |
(*Matches only NS2, not NS1 (or NS3)*) |
|
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
289 |
val Says_imp_sees_Spy'' = |
|
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
290 |
read_instantiate [("X","Crypt ?K {|?XX,?YY,?ZZ|}")] Says_imp_sees_Spy';
|
|
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
291 |
|
|
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
292 |
|
| 2318 | 293 |
(*If B receives NS3 and the nonce NB agrees with the nonce he joined with |
|
2536
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
294 |
NA, then A initiated the run using NA. SAME proof as B_trusts_NS3!*) |
| 2318 | 295 |
goal thy |
|
2536
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
296 |
"!!evs. [| Says B A (Crypt (pubK A) {|Nonce NA, Nonce NB, Agent B|}) \
|
|
3466
30791e5a69c4
Corrected indentations and margins after the renaming of "set_of_list"
paulson
parents:
3465
diff
changeset
|
297 |
\ : set evs; \ |
|
30791e5a69c4
Corrected indentations and margins after the renaming of "set_of_list"
paulson
parents:
3465
diff
changeset
|
298 |
\ Says A' B (Crypt (pubK B) (Nonce NB)): set evs; \ |
|
2536
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
299 |
\ A ~: lost; B ~: lost; evs : ns_public |] \ |
| 3465 | 300 |
\ ==> Says A B (Crypt (pubK B) {|Nonce NA, Agent A|}) : set evs";
|
|
2536
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
301 |
by (etac rev_mp 1); |
|
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
302 |
(*prepare induction over Crypt (pubK B) {|NB|} : parts H*)
|
|
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
303 |
by (etac (Says_imp_sees_Spy' RS parts.Inj RS rev_mp) 1); |
| 2318 | 304 |
by (etac ns_public.induct 1); |
305 |
by (ALLGOALS Asm_simp_tac); |
|
306 |
(*Fake, NS2, NS3*) |
|
307 |
(*NS1*) |
|
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
308 |
by (blast_tac (!claset addSEs sees_Spy_partsEs) 2); |
| 2318 | 309 |
(*Fake*) |
310 |
by (REPEAT_FIRST (resolve_tac [impI, conjI])); |
|
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
311 |
by (Blast_tac 1); |
| 2318 | 312 |
by (forward_tac [Spy_not_see_NB] 1 THEN REPEAT (assume_tac 1)); |
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
313 |
by (blast_tac (!claset addSIs [disjI2] |
|
2516
4d68fbe6378b
Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents:
2497
diff
changeset
|
314 |
addDs [impOfSubs analz_subset_parts, |
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
315 |
impOfSubs Fake_parts_insert]) 1); |
| 2318 | 316 |
(*NS3*) |
317 |
by (Step_tac 1); |
|
318 |
by (forward_tac [Spy_not_see_NB] 1 THEN REPEAT (assume_tac 1)); |
|
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
319 |
by (blast_tac (!claset addSDs [Says_imp_sees_Spy'' RS parts.Inj] |
|
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
320 |
addDs [unique_NB]) 1); |
|
2536
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
321 |
qed "B_trusts_protocol"; |
| 2318 | 322 |