| author | wenzelm | 
| Fri, 20 Sep 2024 19:51:08 +0200 | |
| changeset 80914 | d97fdabd9e2b | 
| parent 76368 | 943f99825f39 | 
| permissions | -rw-r--r-- | 
| 76299 | 1  | 
(* Title: HOL/Auth/NS_Public.thy  | 
| 2318 | 2  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
3  | 
Copyright 1996 University of Cambridge  | 
|
4  | 
*)  | 
|
5  | 
||
| 76299 | 6  | 
section\<open>The Needham-Schroeder Public-Key Protocol\<close>  | 
| 76297 | 7  | 
|
8  | 
text \<open>Flawed version, vulnerable to Lowe's attack.  | 
|
9  | 
From Burrows, Abadi and Needham. A Logic of Authentication.  | 
|
10  | 
Proc. Royal Soc. 426 (1989), p. 260\<close>  | 
|
| 13956 | 11  | 
|
| 16417 | 12  | 
theory NS_Public imports Public begin  | 
| 2318 | 13  | 
|
| 23746 | 14  | 
inductive_set ns_public :: "event list set"  | 
| 76297 | 15  | 
where  | 
| 11104 | 16  | 
Nil: "[] \<in> ns_public"  | 
| 76297 | 17  | 
\<comment> \<open>Initial trace is empty\<close>  | 
| 23746 | 18  | 
| Fake: "\<lbrakk>evsf \<in> ns_public; X \<in> synth (analz (spies evsf))\<rbrakk>  | 
| 11366 | 19  | 
\<Longrightarrow> Says Spy B X # evsf \<in> ns_public"  | 
| 76297 | 20  | 
\<comment> \<open>The spy can say almost anything.\<close>  | 
| 23746 | 21  | 
| NS1: "\<lbrakk>evs1 \<in> ns_public; Nonce NA \<notin> used evs1\<rbrakk>  | 
| 13922 | 22  | 
\<Longrightarrow> Says A B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>)  | 
| 76297 | 23  | 
# evs1 \<in> ns_public"  | 
24  | 
\<comment> \<open>Alice initiates a protocol run, sending a nonce to Bob\<close>  | 
|
| 23746 | 25  | 
| NS2: "\<lbrakk>evs2 \<in> ns_public; Nonce NB \<notin> used evs2;  | 
| 13922 | 26  | 
Says A' B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs2\<rbrakk>  | 
27  | 
\<Longrightarrow> Says B A (Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>)  | 
|
| 11104 | 28  | 
# evs2 \<in> ns_public"  | 
| 76297 | 29  | 
\<comment> \<open>Bob responds to Alice's message with a further nonce\<close>  | 
| 23746 | 30  | 
| NS3: "\<lbrakk>evs3 \<in> ns_public;  | 
| 13922 | 31  | 
Says A B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs3;  | 
| 76297 | 32  | 
Says B' A (Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs3\<rbrakk>  | 
| 13922 | 33  | 
\<Longrightarrow> Says A B (Crypt (pubEK B) (Nonce NB)) # evs3 \<in> ns_public"  | 
| 76297 | 34  | 
   \<comment> \<open>Alice proves her existence by sending @{term NB} back to Bob.\<close>
 | 
| 11104 | 35  | 
|
36  | 
declare knows_Spy_partsEs [elim]  | 
|
| 
14200
 
d8598e24f8fa
Removal of the Key_supply axiom (affects many possbility proofs) and minor
 
paulson 
parents: 
13956 
diff
changeset
 | 
37  | 
declare analz_into_parts [dest]  | 
| 56681 | 38  | 
declare Fake_parts_insert_in_Un [dest]  | 
| 11104 | 39  | 
|
| 76297 | 40  | 
text \<open>A "possibility property": there are traces that reach the end\<close>  | 
| 13922 | 41  | 
lemma "\<exists>NB. \<exists>evs \<in> ns_public. Says A B (Crypt (pubEK B) (Nonce NB)) \<in> set evs"  | 
| 76297 | 42  | 
apply (intro exI bexI)  | 
43  | 
apply (rule_tac [2] ns_public.Nil [THEN ns_public.NS1, THEN ns_public.NS2, THEN ns_public.NS3])  | 
|
44  | 
by possibility  | 
|
45  | 
||
46  | 
||
47  | 
subsection \<open>Inductive proofs about @{term ns_public}\<close>
 | 
|
| 11104 | 48  | 
|
49  | 
(** Theorems of the form X \<notin> parts (spies evs) imply that NOBODY  | 
|
50  | 
sends messages containing X! **)  | 
|
51  | 
||
| 76297 | 52  | 
text \<open>Spy never sees another agent's private key! (unless it's bad at start)\<close>  | 
53  | 
lemma Spy_see_priEK [simp]:  | 
|
54  | 
"evs \<in> ns_public \<Longrightarrow> (Key (priEK A) \<in> parts (spies evs)) = (A \<in> bad)"  | 
|
55  | 
by (erule ns_public.induct, auto)  | 
|
56  | 
||
57  | 
lemma Spy_analz_priEK [simp]:  | 
|
58  | 
"evs \<in> ns_public \<Longrightarrow> (Key (priEK A) \<in> analz (spies evs)) = (A \<in> bad)"  | 
|
59  | 
by auto  | 
|
60  | 
||
| 11104 | 61  | 
|
| 76297 | 62  | 
subsection \<open>Authenticity properties obtained from {term NS1}\<close>
 | 
| 11104 | 63  | 
|
| 76297 | 64  | 
text \<open>It is impossible to re-use a nonce in both {term NS1} and {term NS2}, provided the nonce
 | 
65  | 
is secret. (Honest users generate fresh nonces.)\<close>  | 
|
| 
76368
 
943f99825f39
Replaced some ugly legacy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
76299 
diff
changeset
 | 
66  | 
lemma no_nonce_NS1_NS2:  | 
| 
 
943f99825f39
Replaced some ugly legacy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
76299 
diff
changeset
 | 
67  | 
"\<lbrakk>evs \<in> ns_public;  | 
| 
 
943f99825f39
Replaced some ugly legacy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
76299 
diff
changeset
 | 
68  | 
Crypt (pubEK C) \<lbrace>NA', Nonce NA, Agent D\<rbrace> \<in> parts (spies evs);  | 
| 
 
943f99825f39
Replaced some ugly legacy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
76299 
diff
changeset
 | 
69  | 
Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace> \<in> parts (spies evs)\<rbrakk>  | 
| 
 
943f99825f39
Replaced some ugly legacy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
76299 
diff
changeset
 | 
70  | 
\<Longrightarrow> Nonce NA \<in> analz (spies evs)"  | 
| 76297 | 71  | 
by (induct rule: ns_public.induct) (auto intro: analz_insertI)  | 
| 11104 | 72  | 
|
73  | 
||
| 76297 | 74  | 
text \<open>Unicity for {term NS1}: nonce {term NA} identifies agents {term A} and {term B}\<close>
 | 
75  | 
lemma unique_NA:  | 
|
76  | 
assumes NA: "Crypt(pubEK B) \<lbrace>Nonce NA, Agent A \<rbrace> \<in> parts(spies evs)"  | 
|
77  | 
"Crypt(pubEK B') \<lbrace>Nonce NA, Agent A'\<rbrace> \<in> parts(spies evs)"  | 
|
78  | 
"Nonce NA \<notin> analz (spies evs)"  | 
|
79  | 
and evs: "evs \<in> ns_public"  | 
|
80  | 
shows "A=A' \<and> B=B'"  | 
|
81  | 
using evs NA  | 
|
82  | 
by (induction rule: ns_public.induct) (auto intro!: analz_insertI split: if_split_asm)  | 
|
| 11104 | 83  | 
|
84  | 
||
| 76297 | 85  | 
text \<open>Secrecy: Spy does not see the nonce sent in msg {term NS1} if {term A} and {term B} are secure
 | 
86  | 
The major premise "Says A B ..." makes it a dest-rule, hence the given assumption order. \<close>  | 
|
87  | 
theorem Spy_not_see_NA:  | 
|
88  | 
assumes NA: "Says A B (Crypt(pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs"  | 
|
89  | 
"A \<notin> bad" "B \<notin> bad"  | 
|
90  | 
and evs: "evs \<in> ns_public"  | 
|
91  | 
shows "Nonce NA \<notin> analz (spies evs)"  | 
|
92  | 
using evs NA  | 
|
93  | 
proof (induction rule: ns_public.induct)  | 
|
94  | 
case (Fake evsf X B)  | 
|
95  | 
then show ?case  | 
|
96  | 
by spy_analz  | 
|
97  | 
next  | 
|
98  | 
case (NS2 evs2 NB A' B NA A)  | 
|
99  | 
then show ?case  | 
|
100  | 
by simp (metis Says_imp_analz_Spy analz_into_parts parts.simps unique_NA usedI)  | 
|
101  | 
next  | 
|
102  | 
case (NS3 evs3 A B NA B' NB)  | 
|
103  | 
then show ?case  | 
|
104  | 
by simp (meson Says_imp_analz_Spy analz_into_parts no_nonce_NS1_NS2)  | 
|
105  | 
qed auto  | 
|
| 11104 | 106  | 
|
| 2318 | 107  | 
|
| 76297 | 108  | 
text \<open>Authentication for {term A}: if she receives message 2 and has used {term NA}
 | 
109  | 
  to start a run, then {term B} has sent message 2.\<close>
 | 
|
| 
76368
 
943f99825f39
Replaced some ugly legacy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
76299 
diff
changeset
 | 
110  | 
lemma A_trusts_NS2_lemma:  | 
| 
 
943f99825f39
Replaced some ugly legacy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
76299 
diff
changeset
 | 
111  | 
"\<lbrakk>evs \<in> ns_public;  | 
| 
 
943f99825f39
Replaced some ugly legacy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
76299 
diff
changeset
 | 
112  | 
Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace> \<in> parts (spies evs);  | 
| 
 
943f99825f39
Replaced some ugly legacy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
76299 
diff
changeset
 | 
113  | 
Says A B (Crypt(pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs;  | 
| 
 
943f99825f39
Replaced some ugly legacy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
76299 
diff
changeset
 | 
114  | 
A \<notin> bad; B \<notin> bad\<rbrakk>  | 
| 
 
943f99825f39
Replaced some ugly legacy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
76299 
diff
changeset
 | 
115  | 
\<Longrightarrow> Says B A (Crypt(pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs"  | 
| 
 
943f99825f39
Replaced some ugly legacy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
76299 
diff
changeset
 | 
116  | 
by (induct rule: ns_public.induct) (auto dest: Spy_not_see_NA unique_NA)  | 
| 11104 | 117  | 
|
| 76297 | 118  | 
theorem A_trusts_NS2:  | 
119  | 
"\<lbrakk>Says A B (Crypt(pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs;  | 
|
| 13922 | 120  | 
Says B' A (Crypt(pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs;  | 
| 76297 | 121  | 
A \<notin> bad; B \<notin> bad; evs \<in> ns_public\<rbrakk>  | 
| 13922 | 122  | 
\<Longrightarrow> Says B A (Crypt(pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs"  | 
| 76297 | 123  | 
by (blast intro: A_trusts_NS2_lemma)  | 
| 11104 | 124  | 
|
125  | 
||
| 76297 | 126  | 
text \<open>If the encrypted message appears then it originated with Alice in {term NS1}\<close>
 | 
| 
76368
 
943f99825f39
Replaced some ugly legacy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
76299 
diff
changeset
 | 
127  | 
lemma B_trusts_NS1:  | 
| 
 
943f99825f39
Replaced some ugly legacy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
76299 
diff
changeset
 | 
128  | 
"\<lbrakk>evs \<in> ns_public;  | 
| 
 
943f99825f39
Replaced some ugly legacy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
76299 
diff
changeset
 | 
129  | 
Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace> \<in> parts (spies evs);  | 
| 
 
943f99825f39
Replaced some ugly legacy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
76299 
diff
changeset
 | 
130  | 
Nonce NA \<notin> analz (spies evs)\<rbrakk>  | 
| 
 
943f99825f39
Replaced some ugly legacy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
76299 
diff
changeset
 | 
131  | 
\<Longrightarrow> Says A B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs"  | 
| 
 
943f99825f39
Replaced some ugly legacy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
76299 
diff
changeset
 | 
132  | 
by (induct evs rule: ns_public.induct) (use analz_insertI in \<open>auto split: if_split_asm\<close>)  | 
| 76297 | 133  | 
|
134  | 
||
135  | 
subsection \<open>Authenticity properties obtained from {term NS2}\<close>
 | 
|
136  | 
||
137  | 
text \<open>Unicity for {term NS2}: nonce {term NB} identifies nonce {term NA} and agent {term A} 
 | 
|
138  | 
  [proof closely follows that for @{thm [source] unique_NA}]\<close>
 | 
|
139  | 
||
140  | 
lemma unique_NB [dest]:  | 
|
141  | 
assumes NB: "Crypt(pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace> \<in> parts(spies evs)"  | 
|
142  | 
"Crypt(pubEK A') \<lbrace>Nonce NA', Nonce NB, Agent B'\<rbrace> \<in> parts(spies evs)"  | 
|
143  | 
"Nonce NB \<notin> analz (spies evs)"  | 
|
144  | 
and evs: "evs \<in> ns_public"  | 
|
145  | 
shows "A=A' \<and> NA=NA' \<and> B=B'"  | 
|
146  | 
using evs NB  | 
|
147  | 
by (induction rule: ns_public.induct) (auto intro!: analz_insertI split: if_split_asm)  | 
|
| 11104 | 148  | 
|
149  | 
||
| 76297 | 150  | 
text \<open>{term NB} remains secret\<close>
 | 
151  | 
theorem Spy_not_see_NB [dest]:  | 
|
152  | 
assumes NB: "Says B A (Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs"  | 
|
153  | 
"A \<notin> bad" "B \<notin> bad"  | 
|
154  | 
and evs: "evs \<in> ns_public"  | 
|
155  | 
shows "Nonce NB \<notin> analz (spies evs)"  | 
|
156  | 
using evs NB evs  | 
|
157  | 
proof (induction rule: ns_public.induct)  | 
|
158  | 
case Fake  | 
|
159  | 
then show ?case by spy_analz  | 
|
160  | 
next  | 
|
161  | 
case NS2  | 
|
162  | 
then show ?case  | 
|
163  | 
by (auto intro!: no_nonce_NS1_NS2)  | 
|
164  | 
qed auto  | 
|
| 11104 | 165  | 
|
166  | 
||
| 76297 | 167  | 
text \<open>Authentication for {term B}: if he receives message 3 and has used {term NB}
 | 
168  | 
  in message 2, then {term A} has sent message 3.\<close>
 | 
|
| 
76368
 
943f99825f39
Replaced some ugly legacy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
76299 
diff
changeset
 | 
169  | 
lemma B_trusts_NS3_lemma:  | 
| 76297 | 170  | 
"\<lbrakk>evs \<in> ns_public;  | 
171  | 
Crypt (pubEK B) (Nonce NB) \<in> parts (spies evs);  | 
|
172  | 
Says B A (Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs;  | 
|
173  | 
A \<notin> bad; B \<notin> bad\<rbrakk>  | 
|
174  | 
\<Longrightarrow> Says A B (Crypt (pubEK B) (Nonce NB)) \<in> set evs"  | 
|
175  | 
proof (induction rule: ns_public.induct)  | 
|
176  | 
case (NS3 evs3 A B NA B' NB)  | 
|
177  | 
then show ?case  | 
|
178  | 
by simp (blast intro: no_nonce_NS1_NS2)  | 
|
179  | 
qed auto  | 
|
| 11104 | 180  | 
|
181  | 
theorem B_trusts_NS3:  | 
|
| 13922 | 182  | 
"\<lbrakk>Says B A (Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs;  | 
| 76297 | 183  | 
Says A' B (Crypt (pubEK B) (Nonce NB)) \<in> set evs;  | 
184  | 
A \<notin> bad; B \<notin> bad; evs \<in> ns_public\<rbrakk>  | 
|
| 13922 | 185  | 
\<Longrightarrow> Says A B (Crypt (pubEK B) (Nonce NB)) \<in> set evs"  | 
| 76297 | 186  | 
by (blast intro: B_trusts_NS3_lemma)  | 
| 11104 | 187  | 
|
| 76297 | 188  | 
|
189  | 
subsection\<open>Overall guarantee for {term B}\<close>
 | 
|
| 11104 | 190  | 
|
| 76297 | 191  | 
text \<open>If NS3 has been sent and the nonce NB agrees with the nonce B joined with  | 
192  | 
NA, then A initiated the run using NA.\<close>  | 
|
| 11104 | 193  | 
theorem B_trusts_protocol:  | 
194  | 
"\<lbrakk>A \<notin> bad; B \<notin> bad; evs \<in> ns_public\<rbrakk> \<Longrightarrow>  | 
|
| 13922 | 195  | 
Crypt (pubEK B) (Nonce NB) \<in> parts (spies evs) \<longrightarrow>  | 
196  | 
Says B A (Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs \<longrightarrow>  | 
|
197  | 
Says A B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs"  | 
|
| 11104 | 198  | 
by (erule ns_public.induct, auto)  | 
| 2318 | 199  | 
|
200  | 
end  |