| author | berghofe | 
| Fri, 28 Apr 2006 17:58:33 +0200 | |
| changeset 19500 | 188d4e44c1a6 | 
| parent 16417 | 9bc16273c2d4 | 
| child 23746 | a455e69c31cc | 
| 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.  | 
|
| 2538 | 7  | 
Version incorporating Lowe's fix (inclusion of B's identity in round 2).  | 
| 2318 | 8  | 
*)  | 
9  | 
||
| 13956 | 10  | 
header{*Verifying the Needham-Schroeder-Lowe Public-Key Protocol*}
 | 
11  | 
||
| 16417 | 12  | 
theory NS_Public imports Public begin  | 
| 2318 | 13  | 
|
| 11104 | 14  | 
consts ns_public :: "event list set"  | 
| 2549 | 15  | 
|
| 2318 | 16  | 
inductive ns_public  | 
| 11104 | 17  | 
intros  | 
| 2318 | 18  | 
(*Initial trace is empty*)  | 
| 11104 | 19  | 
Nil: "[] \<in> ns_public"  | 
| 2318 | 20  | 
|
21  | 
(*The spy MAY say anything he CAN say. We do not expect him to  | 
|
22  | 
invent new nonces here, but he can also use NS1. Common to  | 
|
23  | 
all similar protocols.*)  | 
|
| 11366 | 24  | 
Fake: "\<lbrakk>evsf \<in> ns_public; X \<in> synth (analz (spies evsf))\<rbrakk>  | 
25  | 
\<Longrightarrow> Says Spy B X # evsf \<in> ns_public"  | 
|
| 2318 | 26  | 
|
27  | 
(*Alice initiates a protocol run, sending a nonce to Bob*)  | 
|
| 11104 | 28  | 
NS1: "\<lbrakk>evs1 \<in> ns_public; Nonce NA \<notin> used evs1\<rbrakk>  | 
| 13922 | 29  | 
\<Longrightarrow> Says A B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>)  | 
| 11104 | 30  | 
# evs1 \<in> ns_public"  | 
| 2318 | 31  | 
|
32  | 
(*Bob responds to Alice's message with a further nonce*)  | 
|
| 11104 | 33  | 
NS2: "\<lbrakk>evs2 \<in> ns_public; Nonce NB \<notin> used evs2;  | 
| 13922 | 34  | 
Says A' B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs2\<rbrakk>  | 
35  | 
\<Longrightarrow> Says B A (Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>)  | 
|
| 11104 | 36  | 
# evs2 \<in> ns_public"  | 
| 2318 | 37  | 
|
38  | 
(*Alice proves her existence by sending NB back to Bob.*)  | 
|
| 11104 | 39  | 
NS3: "\<lbrakk>evs3 \<in> ns_public;  | 
| 13922 | 40  | 
Says A B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs3;  | 
41  | 
Says B' A (Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>)  | 
|
| 11104 | 42  | 
\<in> set evs3\<rbrakk>  | 
| 13922 | 43  | 
\<Longrightarrow> Says A B (Crypt (pubEK B) (Nonce NB)) # evs3 \<in> ns_public"  | 
| 11104 | 44  | 
|
45  | 
declare knows_Spy_partsEs [elim]  | 
|
| 
14200
 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 
paulson 
parents: 
13956 
diff
changeset
 | 
46  | 
declare knows_Spy_partsEs [elim]  | 
| 
 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 
paulson 
parents: 
13956 
diff
changeset
 | 
47  | 
declare analz_into_parts [dest]  | 
| 
 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 
paulson 
parents: 
13956 
diff
changeset
 | 
48  | 
declare Fake_parts_insert_in_Un [dest]  | 
| 11104 | 49  | 
declare image_eq_UN [simp] (*accelerates proofs involving nested images*)  | 
50  | 
||
51  | 
(*A "possibility property": there are traces that reach the end*)  | 
|
| 13922 | 52  | 
lemma "\<exists>NB. \<exists>evs \<in> ns_public. Says A B (Crypt (pubEK B) (Nonce NB)) \<in> set evs"  | 
| 11104 | 53  | 
apply (intro exI bexI)  | 
54  | 
apply (rule_tac [2] ns_public.Nil [THEN ns_public.NS1, THEN ns_public.NS2,  | 
|
| 
14207
 
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
 
paulson 
parents: 
14200 
diff
changeset
 | 
55  | 
THEN ns_public.NS3], possibility)  | 
| 13926 | 56  | 
done  | 
| 11104 | 57  | 
|
58  | 
(** Theorems of the form X \<notin> parts (spies evs) imply that NOBODY  | 
|
59  | 
sends messages containing X! **)  | 
|
60  | 
||
61  | 
(*Spy never sees another agent's private key! (unless it's bad at start)*)  | 
|
| 13922 | 62  | 
lemma Spy_see_priEK [simp]:  | 
63  | 
"evs \<in> ns_public \<Longrightarrow> (Key (priEK A) \<in> parts (spies evs)) = (A \<in> bad)"  | 
|
| 11104 | 64  | 
by (erule ns_public.induct, auto)  | 
65  | 
||
| 13922 | 66  | 
lemma Spy_analz_priEK [simp]:  | 
67  | 
"evs \<in> ns_public \<Longrightarrow> (Key (priEK A) \<in> analz (spies evs)) = (A \<in> bad)"  | 
|
| 11104 | 68  | 
by auto  | 
69  | 
||
| 
14200
 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 
paulson 
parents: 
13956 
diff
changeset
 | 
70  | 
subsection{*Authenticity properties obtained from NS2*}
 | 
| 11104 | 71  | 
|
72  | 
||
73  | 
(*It is impossible to re-use a nonce in both NS1 and NS2, provided the nonce  | 
|
74  | 
is secret. (Honest users generate fresh nonces.)*)  | 
|
75  | 
lemma no_nonce_NS1_NS2 [rule_format]:  | 
|
76  | 
"evs \<in> ns_public  | 
|
| 13922 | 77  | 
\<Longrightarrow> Crypt (pubEK C) \<lbrace>NA', Nonce NA, Agent D\<rbrace> \<in> parts (spies evs) \<longrightarrow>  | 
78  | 
Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace> \<in> parts (spies evs) \<longrightarrow>  | 
|
| 11104 | 79  | 
Nonce NA \<in> analz (spies evs)"  | 
80  | 
apply (erule ns_public.induct, simp_all)  | 
|
81  | 
apply (blast intro: analz_insertI)+  | 
|
82  | 
done  | 
|
83  | 
||
84  | 
(*Unicity for NS1: nonce NA identifies agents A and B*)  | 
|
85  | 
lemma unique_NA:  | 
|
| 13922 | 86  | 
"\<lbrakk>Crypt(pubEK B) \<lbrace>Nonce NA, Agent A \<rbrace> \<in> parts(spies evs);  | 
87  | 
Crypt(pubEK B') \<lbrace>Nonce NA, Agent A'\<rbrace> \<in> parts(spies evs);  | 
|
| 11104 | 88  | 
Nonce NA \<notin> analz (spies evs); evs \<in> ns_public\<rbrakk>  | 
89  | 
\<Longrightarrow> A=A' \<and> B=B'"  | 
|
90  | 
apply (erule rev_mp, erule rev_mp, erule rev_mp)  | 
|
91  | 
apply (erule ns_public.induct, simp_all)  | 
|
92  | 
(*Fake, NS1*)  | 
|
93  | 
apply (blast intro: analz_insertI)+  | 
|
94  | 
done  | 
|
95  | 
||
96  | 
||
97  | 
(*Secrecy: Spy does not see the nonce sent in msg NS1 if A and B are secure  | 
|
98  | 
The major premise "Says A B ..." makes it a dest-rule, so we use  | 
|
99  | 
(erule rev_mp) rather than rule_format. *)  | 
|
100  | 
theorem Spy_not_see_NA:  | 
|
| 13922 | 101  | 
"\<lbrakk>Says A B (Crypt(pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs;  | 
| 11104 | 102  | 
A \<notin> bad; B \<notin> bad; evs \<in> ns_public\<rbrakk>  | 
103  | 
\<Longrightarrow> Nonce NA \<notin> analz (spies evs)"  | 
|
104  | 
apply (erule rev_mp)  | 
|
| 13507 | 105  | 
apply (erule ns_public.induct, simp_all, spy_analz)  | 
| 11104 | 106  | 
apply (blast dest: unique_NA intro: no_nonce_NS1_NS2)+  | 
107  | 
done  | 
|
108  | 
||
| 2318 | 109  | 
|
| 11104 | 110  | 
(*Authentication for A: if she receives message 2 and has used NA  | 
111  | 
to start a run, then B has sent message 2.*)  | 
|
112  | 
lemma A_trusts_NS2_lemma [rule_format]:  | 
|
113  | 
"\<lbrakk>A \<notin> bad; B \<notin> bad; evs \<in> ns_public\<rbrakk>  | 
|
| 13922 | 114  | 
\<Longrightarrow> Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace> \<in> parts (spies evs) \<longrightarrow>  | 
115  | 
Says A B (Crypt(pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs \<longrightarrow>  | 
|
116  | 
Says B A (Crypt(pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs"  | 
|
| 11104 | 117  | 
apply (erule ns_public.induct, simp_all)  | 
118  | 
(*Fake, NS1*)  | 
|
119  | 
apply (blast dest: Spy_not_see_NA)+  | 
|
120  | 
done  | 
|
121  | 
||
122  | 
theorem A_trusts_NS2:  | 
|
| 13922 | 123  | 
"\<lbrakk>Says A B (Crypt(pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs;  | 
124  | 
Says B' A (Crypt(pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs;  | 
|
| 11104 | 125  | 
A \<notin> bad; B \<notin> bad; evs \<in> ns_public\<rbrakk>  | 
| 13922 | 126  | 
\<Longrightarrow> Says B A (Crypt(pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs"  | 
| 11104 | 127  | 
by (blast intro: A_trusts_NS2_lemma)  | 
128  | 
||
129  | 
||
130  | 
(*If the encrypted message appears then it originated with Alice in NS1*)  | 
|
131  | 
lemma B_trusts_NS1 [rule_format]:  | 
|
132  | 
"evs \<in> ns_public  | 
|
| 13922 | 133  | 
\<Longrightarrow> Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace> \<in> parts (spies evs) \<longrightarrow>  | 
| 11104 | 134  | 
Nonce NA \<notin> analz (spies evs) \<longrightarrow>  | 
| 13922 | 135  | 
Says A B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs"  | 
| 11104 | 136  | 
apply (erule ns_public.induct, simp_all)  | 
137  | 
(*Fake*)  | 
|
138  | 
apply (blast intro!: analz_insertI)  | 
|
139  | 
done  | 
|
140  | 
||
141  | 
||
| 
14200
 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 
paulson 
parents: 
13956 
diff
changeset
 | 
142  | 
subsection{*Authenticity properties obtained from NS2*}
 | 
| 11104 | 143  | 
|
144  | 
(*Unicity for NS2: nonce NB identifies nonce NA and agents A, B  | 
|
145  | 
[unicity of B makes Lowe's fix work]  | 
|
146  | 
[proof closely follows that for unique_NA] *)  | 
|
147  | 
||
148  | 
lemma unique_NB [dest]:  | 
|
| 13922 | 149  | 
"\<lbrakk>Crypt(pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace> \<in> parts(spies evs);  | 
150  | 
Crypt(pubEK A') \<lbrace>Nonce NA', Nonce NB, Agent B'\<rbrace> \<in> parts(spies evs);  | 
|
| 11104 | 151  | 
Nonce NB \<notin> analz (spies evs); evs \<in> ns_public\<rbrakk>  | 
152  | 
\<Longrightarrow> A=A' \<and> NA=NA' \<and> B=B'"  | 
|
153  | 
apply (erule rev_mp, erule rev_mp, erule rev_mp)  | 
|
154  | 
apply (erule ns_public.induct, simp_all)  | 
|
155  | 
(*Fake, NS2*)  | 
|
156  | 
apply (blast intro: analz_insertI)+  | 
|
157  | 
done  | 
|
158  | 
||
159  | 
||
160  | 
(*Secrecy: Spy does not see the nonce sent in msg NS2 if A and B are secure*)  | 
|
161  | 
theorem Spy_not_see_NB [dest]:  | 
|
| 13922 | 162  | 
"\<lbrakk>Says B A (Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs;  | 
| 11104 | 163  | 
A \<notin> bad; B \<notin> bad; evs \<in> ns_public\<rbrakk>  | 
164  | 
\<Longrightarrow> Nonce NB \<notin> analz (spies evs)"  | 
|
165  | 
apply (erule rev_mp)  | 
|
| 13507 | 166  | 
apply (erule ns_public.induct, simp_all, spy_analz)  | 
| 11104 | 167  | 
apply (blast intro: no_nonce_NS1_NS2)+  | 
168  | 
done  | 
|
169  | 
||
170  | 
||
171  | 
(*Authentication for B: if he receives message 3 and has used NB  | 
|
172  | 
in message 2, then A has sent message 3.*)  | 
|
173  | 
lemma B_trusts_NS3_lemma [rule_format]:  | 
|
174  | 
"\<lbrakk>A \<notin> bad; B \<notin> bad; evs \<in> ns_public\<rbrakk> \<Longrightarrow>  | 
|
| 13922 | 175  | 
Crypt (pubEK B) (Nonce NB) \<in> parts (spies evs) \<longrightarrow>  | 
176  | 
Says B A (Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs \<longrightarrow>  | 
|
177  | 
Says A B (Crypt (pubEK B) (Nonce NB)) \<in> set evs"  | 
|
| 11104 | 178  | 
by (erule ns_public.induct, auto)  | 
179  | 
||
180  | 
theorem B_trusts_NS3:  | 
|
| 13922 | 181  | 
"\<lbrakk>Says B A (Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs;  | 
182  | 
Says A' B (Crypt (pubEK B) (Nonce NB)) \<in> set evs;  | 
|
| 11104 | 183  | 
A \<notin> bad; B \<notin> bad; evs \<in> ns_public\<rbrakk>  | 
| 13922 | 184  | 
\<Longrightarrow> Says A B (Crypt (pubEK B) (Nonce NB)) \<in> set evs"  | 
| 11104 | 185  | 
by (blast intro: B_trusts_NS3_lemma)  | 
186  | 
||
| 
14200
 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 
paulson 
parents: 
13956 
diff
changeset
 | 
187  | 
subsection{*Overall guarantee for B*}
 | 
| 11104 | 188  | 
|
189  | 
(*If NS3 has been sent and the nonce NB agrees with the nonce B joined with  | 
|
190  | 
NA, then A initiated the run using NA.*)  | 
|
191  | 
theorem B_trusts_protocol:  | 
|
192  | 
"\<lbrakk>A \<notin> bad; B \<notin> bad; evs \<in> ns_public\<rbrakk> \<Longrightarrow>  | 
|
| 13922 | 193  | 
Crypt (pubEK B) (Nonce NB) \<in> parts (spies evs) \<longrightarrow>  | 
194  | 
Says B A (Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs \<longrightarrow>  | 
|
195  | 
Says A B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs"  | 
|
| 11104 | 196  | 
by (erule ns_public.induct, auto)  | 
| 2318 | 197  | 
|
198  | 
end  |