author | kleing |
Fri, 09 Feb 2001 16:22:30 +0100 | |
changeset 11086 | e714862ecc0a |
parent 8054 | 2ce57ef2a4aa |
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 |
||
4470 | 10 |
AddEs spies_partsEs; |
11 |
AddDs [impOfSubs analz_subset_parts]; |
|
12 |
AddDs [impOfSubs Fake_parts_insert]; |
|
13 |
||
3683 | 14 |
AddIffs [Spy_in_bad]; |
2318 | 15 |
|
8054
2ce57ef2a4aa
used image_eq_UN to speed up slow proofs of base cases
paulson
parents:
5434
diff
changeset
|
16 |
Addsimps [image_eq_UN]; (*accelerates proofs involving nested images*) |
2ce57ef2a4aa
used image_eq_UN to speed up slow proofs of base cases
paulson
parents:
5434
diff
changeset
|
17 |
|
2318 | 18 |
(*A "possibility property": there are traces that reach the end*) |
5434
9b4bed3f394c
Got rid of not_Says_to_self and most uses of ~= in definitions and theorems
paulson
parents:
5421
diff
changeset
|
19 |
Goal |
9b4bed3f394c
Got rid of not_Says_to_self and most uses of ~= in definitions and theorems
paulson
parents:
5421
diff
changeset
|
20 |
"EX NB. EX evs: ns_public. Says A B (Crypt (pubK B) (Nonce NB)) : set evs"; |
2318 | 21 |
by (REPEAT (resolve_tac [exI,bexI] 1)); |
22 |
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
|
23 |
by possibility_tac; |
2318 | 24 |
result(); |
25 |
||
26 |
||
27 |
(**** Inductive proofs about ns_public ****) |
|
28 |
||
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3516
diff
changeset
|
29 |
(*Induction for regularity theorems. If induction formula has the form |
3683 | 30 |
X ~: analz (spies evs) --> ... then it shortens the proof by discarding |
31 |
needless information about analz (insert X (spies evs)) *) |
|
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3516
diff
changeset
|
32 |
fun parts_induct_tac i = |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3516
diff
changeset
|
33 |
etac ns_public.induct i |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3516
diff
changeset
|
34 |
THEN |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3516
diff
changeset
|
35 |
REPEAT (FIRSTGOAL analz_mono_contra_tac) |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3516
diff
changeset
|
36 |
THEN |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3516
diff
changeset
|
37 |
prove_simple_subgoals_tac i; |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3516
diff
changeset
|
38 |
|
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3516
diff
changeset
|
39 |
|
3683 | 40 |
(** Theorems of the form X ~: parts (spies evs) imply that NOBODY |
2318 | 41 |
sends messages containing X! **) |
42 |
||
3683 | 43 |
(*Spy never sees another agent's private key! (unless it's bad at start)*) |
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
44 |
Goal "evs: ns_public ==> (Key (priK A) : parts (spies evs)) = (A : bad)"; |
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3516
diff
changeset
|
45 |
by (parts_induct_tac 1); |
4470 | 46 |
by (Blast_tac 1); |
2318 | 47 |
qed "Spy_see_priK"; |
48 |
Addsimps [Spy_see_priK]; |
|
49 |
||
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
50 |
Goal "evs: ns_public ==> (Key (priK A) : analz (spies evs)) = (A : bad)"; |
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4470
diff
changeset
|
51 |
by Auto_tac; |
2318 | 52 |
qed "Spy_analz_priK"; |
53 |
Addsimps [Spy_analz_priK]; |
|
54 |
||
4470 | 55 |
AddSDs [Spy_see_priK RSN (2, rev_iffD1), |
56 |
Spy_analz_priK RSN (2, rev_iffD1)]; |
|
2318 | 57 |
|
58 |
||
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3516
diff
changeset
|
59 |
(**** Authenticity properties obtained from NS2 ****) |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3516
diff
changeset
|
60 |
|
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3516
diff
changeset
|
61 |
(*It is impossible to re-use a nonce in both NS1 and NS2, provided the nonce |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3516
diff
changeset
|
62 |
is secret. (Honest users generate fresh nonces.)*) |
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
63 |
Goal "[| Crypt (pubK B) {|Nonce NA, Agent A|} : parts (spies evs); \ |
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
64 |
\ Nonce NA ~: analz (spies evs); evs : ns_public |] \ |
3683 | 65 |
\ ==> Crypt (pubK C) {|NA', Nonce NA, Agent D|} ~: parts (spies evs)"; |
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3516
diff
changeset
|
66 |
by (etac rev_mp 1); |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3516
diff
changeset
|
67 |
by (etac rev_mp 1); |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3516
diff
changeset
|
68 |
by (parts_induct_tac 1); |
4470 | 69 |
by (ALLGOALS Blast_tac); |
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3516
diff
changeset
|
70 |
qed "no_nonce_NS1_NS2"; |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3516
diff
changeset
|
71 |
|
4470 | 72 |
(*Adding it to the claset slows down proofs...*) |
73 |
val nonce_NS1_NS2_E = no_nonce_NS1_NS2 RSN (2, rev_notE); |
|
74 |
||
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3516
diff
changeset
|
75 |
|
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3516
diff
changeset
|
76 |
(*Unicity for NS1: nonce NA identifies agents A and B*) |
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
77 |
Goal "[| Nonce NA ~: analz (spies evs); evs : ns_public |] \ |
3709 | 78 |
\ ==> EX A' B'. ALL A B. \ |
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
79 |
\ Crypt (pubK B) {|Nonce NA, Agent A|} : parts (spies evs) --> \ |
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
80 |
\ A=A' & B=B'"; |
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3516
diff
changeset
|
81 |
by (etac rev_mp 1); |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3516
diff
changeset
|
82 |
by (parts_induct_tac 1); |
4556
e7a4683c0026
Tidying, mostly to do with handling a more specific version of Fake_parts_insert
paulson
parents:
4477
diff
changeset
|
83 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps [all_conj_distrib]))); |
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3516
diff
changeset
|
84 |
(*NS1*) |
4470 | 85 |
by (expand_case_tac "NA = ?y" 2 THEN Blast_tac 2); |
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3516
diff
changeset
|
86 |
(*Fake*) |
3709 | 87 |
by (Clarify_tac 1); |
4470 | 88 |
by (Blast_tac 1); |
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3516
diff
changeset
|
89 |
val lemma = result(); |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3516
diff
changeset
|
90 |
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
91 |
Goal "[| Crypt(pubK B) {|Nonce NA, Agent A|} : parts(spies evs); \ |
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
92 |
\ Crypt(pubK B') {|Nonce NA, Agent A'|} : parts(spies evs); \ |
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
93 |
\ Nonce NA ~: analz (spies evs); \ |
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
94 |
\ evs : ns_public |] \ |
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
95 |
\ ==> A=A' & B=B'"; |
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3516
diff
changeset
|
96 |
by (prove_unique_tac lemma 1); |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3516
diff
changeset
|
97 |
qed "unique_NA"; |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3516
diff
changeset
|
98 |
|
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3516
diff
changeset
|
99 |
|
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3516
diff
changeset
|
100 |
(*Tactic for proving secrecy theorems*) |
4686 | 101 |
fun analz_induct_tac i = |
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
102 |
etac ns_public.induct i THEN |
4686 | 103 |
ALLGOALS Asm_simp_tac; |
2318 | 104 |
|
105 |
||
106 |
(*Secrecy: Spy does not see the nonce sent in msg NS1 if A and B are secure*) |
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
107 |
Goal "[| Says A B (Crypt(pubK B) {|Nonce NA, Agent A|}) : set evs; \ |
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
108 |
\ A ~: bad; B ~: bad; evs : ns_public |] \ |
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
109 |
\ ==> Nonce NA ~: analz (spies evs)"; |
2536
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
110 |
by (etac rev_mp 1); |
2418
6b6a92d05fb2
New tactics: prove_unique_tac and analz_induct_tac
paulson
parents:
2374
diff
changeset
|
111 |
by (analz_induct_tac 1); |
2318 | 112 |
(*NS3*) |
4470 | 113 |
by (blast_tac (claset() addEs [nonce_NS1_NS2_E]) 4); |
2536
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
114 |
(*NS2*) |
4470 | 115 |
by (blast_tac (claset() addDs [unique_NA]) 3); |
2318 | 116 |
(*NS1*) |
4470 | 117 |
by (Blast_tac 2); |
2318 | 118 |
(*Fake*) |
2497 | 119 |
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
|
120 |
qed "Spy_not_see_NA"; |
2318 | 121 |
|
122 |
||
123 |
(*Authentication for A: if she receives message 2 and has used NA |
|
124 |
to start a run, then B has sent message 2.*) |
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
125 |
Goal "[| Says A B (Crypt(pubK B) {|Nonce NA, Agent A|}) : set evs; \ |
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
126 |
\ Says B' A (Crypt(pubK A) {|Nonce NA, Nonce NB, Agent B|}) \ |
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
127 |
\ : set evs; \ |
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
128 |
\ A ~: bad; B ~: bad; evs : ns_public |] \ |
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
129 |
\ ==> Says B A (Crypt(pubK A) {|Nonce NA, Nonce NB, Agent B|}) \ |
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
130 |
\ : set evs"; |
2536
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
131 |
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
|
132 |
(*prepare induction over Crypt (pubK A) {|NA,NB,B|} : parts H*) |
3683 | 133 |
by (etac (Says_imp_spies RS parts.Inj RS rev_mp) 1); |
2318 | 134 |
by (etac ns_public.induct 1); |
135 |
by (ALLGOALS Asm_simp_tac); |
|
136 |
(*NS1*) |
|
4470 | 137 |
by (Blast_tac 2); |
2318 | 138 |
(*Fake*) |
4422
21238c9d363e
Simplified proofs using rewrites for f``A where f is injective
paulson
parents:
4197
diff
changeset
|
139 |
by (blast_tac (claset() addDs [Spy_not_see_NA]) 1); |
2318 | 140 |
qed "A_trusts_NS2"; |
141 |
||
4556
e7a4683c0026
Tidying, mostly to do with handling a more specific version of Fake_parts_insert
paulson
parents:
4477
diff
changeset
|
142 |
|
2318 | 143 |
(*If the encrypted message appears then it originated with Alice in NS1*) |
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
144 |
Goal "[| Crypt (pubK B) {|Nonce NA, Agent A|} : parts (spies evs); \ |
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
145 |
\ Nonce NA ~: analz (spies evs); \ |
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
146 |
\ evs : ns_public |] \ |
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
147 |
\==> 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
|
148 |
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
|
149 |
by (etac rev_mp 1); |
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3516
diff
changeset
|
150 |
by (parts_induct_tac 1); |
4470 | 151 |
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
|
152 |
qed "B_trusts_NS1"; |
2318 | 153 |
|
154 |
||
155 |
||
156 |
(**** Authenticity properties obtained from NS2 ****) |
|
157 |
||
2480 | 158 |
(*Unicity for NS2: nonce NB identifies nonce NA and agents A, B |
2318 | 159 |
[unicity of B makes Lowe's fix work] |
160 |
[proof closely follows that for unique_NA] *) |
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
161 |
Goal "[| Nonce NB ~: analz (spies evs); evs : ns_public |] \ |
3709 | 162 |
\ ==> EX A' NA' B'. ALL A NA B. \ |
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
163 |
\ Crypt (pubK A) {|Nonce NA, Nonce NB, Agent B|} : parts (spies evs) \ |
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
164 |
\ --> A=A' & NA=NA' & B=B'"; |
2536
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
165 |
by (etac rev_mp 1); |
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3516
diff
changeset
|
166 |
by (parts_induct_tac 1); |
4556
e7a4683c0026
Tidying, mostly to do with handling a more specific version of Fake_parts_insert
paulson
parents:
4477
diff
changeset
|
167 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps [all_conj_distrib]))); |
2318 | 168 |
(*NS2*) |
4470 | 169 |
by (expand_case_tac "NB = ?y" 2 THEN Blast_tac 2); |
2318 | 170 |
(*Fake*) |
3709 | 171 |
by (Clarify_tac 1); |
4470 | 172 |
by (Blast_tac 1); |
2318 | 173 |
val lemma = result(); |
174 |
||
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
175 |
Goal "[| Crypt(pubK A) {|Nonce NA, Nonce NB, Agent B|} \ |
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
176 |
\ : parts(spies evs); \ |
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
177 |
\ Crypt(pubK A') {|Nonce NA', Nonce NB, Agent B'|} \ |
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
178 |
\ : parts(spies evs); \ |
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
179 |
\ Nonce NB ~: analz (spies evs); \ |
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
180 |
\ evs : ns_public |] \ |
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
181 |
\ ==> A=A' & NA=NA' & B=B'"; |
2418
6b6a92d05fb2
New tactics: prove_unique_tac and analz_induct_tac
paulson
parents:
2374
diff
changeset
|
182 |
by (prove_unique_tac lemma 1); |
2318 | 183 |
qed "unique_NB"; |
184 |
||
4470 | 185 |
AddDs [unique_NB]; |
186 |
||
2318 | 187 |
|
188 |
(*Secrecy: Spy does not see the nonce sent in msg NS2 if A and B are secure*) |
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
189 |
Goal "[| Says B A (Crypt (pubK A) {|Nonce NA, Nonce NB, Agent B|}) \ |
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
190 |
\ : set evs; \ |
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
191 |
\ A ~: bad; B ~: bad; evs : ns_public |] \ |
3683 | 192 |
\ ==> Nonce NB ~: analz (spies evs)"; |
2536
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
193 |
by (etac rev_mp 1); |
2418
6b6a92d05fb2
New tactics: prove_unique_tac and analz_induct_tac
paulson
parents:
2374
diff
changeset
|
194 |
by (analz_induct_tac 1); |
2318 | 195 |
(*NS3*) |
4470 | 196 |
by (Blast_tac 4); |
3709 | 197 |
(*NS2: by freshness and unicity of NB*) |
4470 | 198 |
by (blast_tac (claset() addEs [nonce_NS1_NS2_E]) 3); |
2318 | 199 |
(*NS1*) |
4470 | 200 |
by (Blast_tac 2); |
2318 | 201 |
(*Fake*) |
2497 | 202 |
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
|
203 |
qed "Spy_not_see_NB"; |
2318 | 204 |
|
4422
21238c9d363e
Simplified proofs using rewrites for f``A where f is injective
paulson
parents:
4197
diff
changeset
|
205 |
AddDs [Spy_not_see_NB]; |
21238c9d363e
Simplified proofs using rewrites for f``A where f is injective
paulson
parents:
4197
diff
changeset
|
206 |
|
2318 | 207 |
|
208 |
(*Authentication for B: if he receives message 3 and has used NB |
|
209 |
in message 2, then A has sent message 3.*) |
|
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
210 |
Goal "[| Says B A (Crypt (pubK A) {|Nonce NA, Nonce NB, Agent B|}) \ |
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
211 |
\ : set evs; \ |
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
212 |
\ Says A' B (Crypt (pubK B) (Nonce NB)): set evs; \ |
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
213 |
\ A ~: bad; B ~: bad; evs : ns_public |] \ |
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
214 |
\ ==> 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
|
215 |
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
|
216 |
(*prepare induction over Crypt (pubK B) NB : parts H*) |
3683 | 217 |
by (etac (Says_imp_spies RS parts.Inj RS rev_mp) 1); |
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3516
diff
changeset
|
218 |
by (parts_induct_tac 1); |
3709 | 219 |
by (ALLGOALS Clarify_tac); |
4470 | 220 |
by (ALLGOALS Blast_tac); |
2318 | 221 |
qed "B_trusts_NS3"; |
222 |
||
223 |
||
224 |
(**** Overall guarantee for B*) |
|
225 |
||
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
226 |
(*Matches only NS2, not NS1 (or NS3)*) |
3683 | 227 |
val Says_imp_spies' = |
228 |
read_instantiate [("X","Crypt ?K {|?XX,?YY,?ZZ|}")] Says_imp_spies; |
|
3121
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
229 |
|
cbb6c0c1c58a
Conversion to use blast_tac (with other improvements)
paulson
parents:
2637
diff
changeset
|
230 |
|
2318 | 231 |
(*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
|
232 |
NA, then A initiated the run using NA. SAME proof as B_trusts_NS3!*) |
5114
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
233 |
Goal "[| Says B A (Crypt (pubK A) {|Nonce NA, Nonce NB, Agent B|}) \ |
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
234 |
\ : set evs; \ |
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
235 |
\ Says A' B (Crypt (pubK B) (Nonce NB)): set evs; \ |
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
236 |
\ A ~: bad; B ~: bad; evs : ns_public |] \ |
c729d4c299c1
Deleted leading parameters thanks to new Goal command
paulson
parents:
5076
diff
changeset
|
237 |
\ ==> 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
|
238 |
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
|
239 |
(*prepare induction over Crypt (pubK B) {|NB|} : parts H*) |
3683 | 240 |
by (etac (Says_imp_spies RS parts.Inj RS rev_mp) 1); |
2318 | 241 |
by (etac ns_public.induct 1); |
242 |
by (ALLGOALS Asm_simp_tac); |
|
3709 | 243 |
by (ALLGOALS Clarify_tac); |
4470 | 244 |
(*NS3 holds because NB determines A and NA*) |
245 |
by (ALLGOALS Blast_tac); |
|
2536
1e04eb7f7eb1
Tidied proofs by using "etac rev_mp" instead of applying rev_mp to result()
paulson
parents:
2516
diff
changeset
|
246 |
qed "B_trusts_protocol"; |
2318 | 247 |