author | wenzelm |
Fri, 22 Apr 2022 10:11:06 +0200 | |
changeset 75444 | 331f96a67924 |
parent 69597 | ff784d5a5bfb |
child 76287 | cdc14f94c754 |
permissions | -rw-r--r-- |
37936 | 1 |
(* Title: HOL/Auth/Public.thy |
2318 | 2 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
3 |
Copyright 1996 University of Cambridge |
|
4 |
||
3512
9dcb4daa15e8
Moving common declarations and proofs from theories "Shared"
paulson
parents:
3478
diff
changeset
|
5 |
Theory of Public Keys (common to all public-key protocols) |
2318 | 6 |
|
3512
9dcb4daa15e8
Moving common declarations and proofs from theories "Shared"
paulson
parents:
3478
diff
changeset
|
7 |
Private and public keys; initial states of agents |
2318 | 8 |
*) |
9 |
||
32630 | 10 |
theory Public |
11 |
imports Event |
|
12 |
begin |
|
13922 | 13 |
|
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
14 |
lemma invKey_K: "K \<in> symKeys ==> invKey K = K" |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
15 |
by (simp add: symKeys_def) |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
16 |
|
61830 | 17 |
subsection\<open>Asymmetric Keys\<close> |
2318 | 18 |
|
58310 | 19 |
datatype keymode = Signature | Encryption |
18749
31c2af8b0c60
replacement of bool by a datatype (making problems first-order). More lemma names
paulson
parents:
18570
diff
changeset
|
20 |
|
2318 | 21 |
consts |
67613 | 22 |
publicKey :: "[keymode,agent] \<Rightarrow> key" |
2318 | 23 |
|
20768 | 24 |
abbreviation |
67613 | 25 |
pubEK :: "agent \<Rightarrow> key" where |
20768 | 26 |
"pubEK == publicKey Encryption" |
13922 | 27 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20768
diff
changeset
|
28 |
abbreviation |
67613 | 29 |
pubSK :: "agent \<Rightarrow> key" where |
20768 | 30 |
"pubSK == publicKey Signature" |
13922 | 31 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20768
diff
changeset
|
32 |
abbreviation |
67613 | 33 |
privateKey :: "[keymode, agent] \<Rightarrow> key" where |
20768 | 34 |
"privateKey b A == invKey (publicKey b A)" |
13922 | 35 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20768
diff
changeset
|
36 |
abbreviation |
13922 | 37 |
(*BEWARE!! priEK, priSK DON'T WORK with inj, range, image, etc.*) |
67613 | 38 |
priEK :: "agent \<Rightarrow> key" where |
20768 | 39 |
"priEK A == privateKey Encryption A" |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20768
diff
changeset
|
40 |
|
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20768
diff
changeset
|
41 |
abbreviation |
67613 | 42 |
priSK :: "agent \<Rightarrow> key" where |
20768 | 43 |
"priSK A == privateKey Signature A" |
13922 | 44 |
|
45 |
||
61830 | 46 |
text\<open>These abbreviations give backward compatibility. They represent the |
47 |
simple situation where the signature and encryption keys are the same.\<close> |
|
20768 | 48 |
|
49 |
abbreviation |
|
67613 | 50 |
pubK :: "agent \<Rightarrow> key" where |
20768 | 51 |
"pubK A == pubEK A" |
52 |
||
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20768
diff
changeset
|
53 |
abbreviation |
67613 | 54 |
priK :: "agent \<Rightarrow> key" where |
20768 | 55 |
"priK A == invKey (pubEK A)" |
13922 | 56 |
|
57 |
||
61830 | 58 |
text\<open>By freeness of agents, no two agents have the same key. Since |
69597 | 59 |
\<^term>\<open>True\<noteq>False\<close>, no agent has identical signing and encryption keys\<close> |
14126
28824746d046
Tidying and replacement of some axioms by specifications
paulson
parents:
13956
diff
changeset
|
60 |
specification (publicKey) |
13922 | 61 |
injective_publicKey: |
67613 | 62 |
"publicKey b A = publicKey c A' ==> b=c \<and> A=A'" |
18749
31c2af8b0c60
replacement of bool by a datatype (making problems first-order). More lemma names
paulson
parents:
18570
diff
changeset
|
63 |
apply (rule exI [of _ |
67613 | 64 |
"\<lambda>b A. 2 * case_agent 0 (\<lambda>n. n + 2) 1 A + case_keymode 0 1 b"]) |
18749
31c2af8b0c60
replacement of bool by a datatype (making problems first-order). More lemma names
paulson
parents:
18570
diff
changeset
|
65 |
apply (auto simp add: inj_on_def split: agent.split keymode.split) |
23315 | 66 |
apply presburger |
67 |
apply presburger |
|
14126
28824746d046
Tidying and replacement of some axioms by specifications
paulson
parents:
13956
diff
changeset
|
68 |
done |
13922 | 69 |
|
14126
28824746d046
Tidying and replacement of some axioms by specifications
paulson
parents:
13956
diff
changeset
|
70 |
|
41774 | 71 |
axiomatization where |
13922 | 72 |
(*No private key equals any public key (essential to ensure that private |
73 |
keys are private!) *) |
|
74 |
privateKey_neq_publicKey [iff]: "privateKey b A \<noteq> publicKey c A'" |
|
75 |
||
18749
31c2af8b0c60
replacement of bool by a datatype (making problems first-order). More lemma names
paulson
parents:
18570
diff
changeset
|
76 |
lemmas publicKey_neq_privateKey = privateKey_neq_publicKey [THEN not_sym] |
31c2af8b0c60
replacement of bool by a datatype (making problems first-order). More lemma names
paulson
parents:
18570
diff
changeset
|
77 |
declare publicKey_neq_privateKey [iff] |
13922 | 78 |
|
79 |
||
69597 | 80 |
subsection\<open>Basic properties of \<^term>\<open>pubK\<close> and \<^term>\<open>priK\<close>\<close> |
13922 | 81 |
|
67613 | 82 |
lemma publicKey_inject [iff]: "(publicKey b A = publicKey c A') = (b=c \<and> A=A')" |
13922 | 83 |
by (blast dest!: injective_publicKey) |
84 |
||
85 |
lemma not_symKeys_pubK [iff]: "publicKey b A \<notin> symKeys" |
|
13926 | 86 |
by (simp add: symKeys_def) |
13922 | 87 |
|
88 |
lemma not_symKeys_priK [iff]: "privateKey b A \<notin> symKeys" |
|
13926 | 89 |
by (simp add: symKeys_def) |
13922 | 90 |
|
91 |
lemma symKey_neq_priEK: "K \<in> symKeys ==> K \<noteq> priEK A" |
|
92 |
by auto |
|
93 |
||
94 |
lemma symKeys_neq_imp_neq: "(K \<in> symKeys) \<noteq> (K' \<in> symKeys) ==> K \<noteq> K'" |
|
13926 | 95 |
by blast |
13922 | 96 |
|
97 |
lemma symKeys_invKey_iff [iff]: "(invKey K \<in> symKeys) = (K \<in> symKeys)" |
|
13926 | 98 |
by (unfold symKeys_def, auto) |
2318 | 99 |
|
13922 | 100 |
lemma analz_symKeys_Decrypt: |
101 |
"[| Crypt K X \<in> analz H; K \<in> symKeys; Key K \<in> analz H |] |
|
102 |
==> X \<in> analz H" |
|
103 |
by (auto simp add: symKeys_def) |
|
104 |
||
105 |
||
106 |
||
61830 | 107 |
subsection\<open>"Image" equations that hold for injective functions\<close> |
13922 | 108 |
|
109 |
lemma invKey_image_eq [simp]: "(invKey x \<in> invKey`A) = (x \<in> A)" |
|
13926 | 110 |
by auto |
13922 | 111 |
|
112 |
(*holds because invKey is injective*) |
|
113 |
lemma publicKey_image_eq [simp]: |
|
67613 | 114 |
"(publicKey b x \<in> publicKey c ` AA) = (b=c \<and> x \<in> AA)" |
13922 | 115 |
by auto |
116 |
||
117 |
lemma privateKey_notin_image_publicKey [simp]: "privateKey b x \<notin> publicKey c ` AA" |
|
13926 | 118 |
by auto |
13922 | 119 |
|
120 |
lemma privateKey_image_eq [simp]: |
|
67613 | 121 |
"(privateKey b A \<in> invKey ` publicKey c ` AS) = (b=c \<and> A\<in>AS)" |
13922 | 122 |
by auto |
123 |
||
124 |
lemma publicKey_notin_image_privateKey [simp]: "publicKey b A \<notin> invKey ` publicKey c ` AS" |
|
13926 | 125 |
by auto |
13922 | 126 |
|
127 |
||
61830 | 128 |
subsection\<open>Symmetric Keys\<close> |
13922 | 129 |
|
61830 | 130 |
text\<open>For some protocols, it is convenient to equip agents with symmetric as |
131 |
well as asymmetric keys. The theory \<open>Shared\<close> assumes that all keys |
|
132 |
are symmetric.\<close> |
|
13922 | 133 |
|
134 |
consts |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
135 |
shrK :: "agent => key" \<comment> \<open>long-term shared keys\<close> |
13922 | 136 |
|
14126
28824746d046
Tidying and replacement of some axioms by specifications
paulson
parents:
13956
diff
changeset
|
137 |
specification (shrK) |
28824746d046
Tidying and replacement of some axioms by specifications
paulson
parents:
13956
diff
changeset
|
138 |
inj_shrK: "inj shrK" |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
139 |
\<comment> \<open>No two agents have the same long-term key\<close> |
55416 | 140 |
apply (rule exI [of _ "case_agent 0 (\<lambda>n. n + 2) 1"]) |
14126
28824746d046
Tidying and replacement of some axioms by specifications
paulson
parents:
13956
diff
changeset
|
141 |
apply (simp add: inj_on_def split: agent.split) |
28824746d046
Tidying and replacement of some axioms by specifications
paulson
parents:
13956
diff
changeset
|
142 |
done |
28824746d046
Tidying and replacement of some axioms by specifications
paulson
parents:
13956
diff
changeset
|
143 |
|
41774 | 144 |
axiomatization where |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
145 |
sym_shrK [iff]: "shrK X \<in> symKeys" \<comment> \<open>All shared keys are symmetric\<close> |
13922 | 146 |
|
61830 | 147 |
text\<open>Injectiveness: Agents' long-term keys are distinct.\<close> |
18570 | 148 |
lemmas shrK_injective = inj_shrK [THEN inj_eq] |
149 |
declare shrK_injective [iff] |
|
13922 | 150 |
|
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
151 |
lemma invKey_shrK [simp]: "invKey (shrK A) = shrK A" |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
152 |
by (simp add: invKey_K) |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
153 |
|
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
154 |
lemma analz_shrK_Decrypt: |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
155 |
"[| Crypt (shrK A) X \<in> analz H; Key(shrK A) \<in> analz H |] ==> X \<in> analz H" |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
156 |
by auto |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
157 |
|
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
158 |
lemma analz_Decrypt': |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
159 |
"[| Crypt K X \<in> analz H; K \<in> symKeys; Key K \<in> analz H |] ==> X \<in> analz H" |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
160 |
by (auto simp add: invKey_K) |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
161 |
|
13922 | 162 |
lemma priK_neq_shrK [iff]: "shrK A \<noteq> privateKey b C" |
13926 | 163 |
by (simp add: symKeys_neq_imp_neq) |
13922 | 164 |
|
18749
31c2af8b0c60
replacement of bool by a datatype (making problems first-order). More lemma names
paulson
parents:
18570
diff
changeset
|
165 |
lemmas shrK_neq_priK = priK_neq_shrK [THEN not_sym] |
31c2af8b0c60
replacement of bool by a datatype (making problems first-order). More lemma names
paulson
parents:
18570
diff
changeset
|
166 |
declare shrK_neq_priK [simp] |
13922 | 167 |
|
168 |
lemma pubK_neq_shrK [iff]: "shrK A \<noteq> publicKey b C" |
|
13926 | 169 |
by (simp add: symKeys_neq_imp_neq) |
13922 | 170 |
|
18749
31c2af8b0c60
replacement of bool by a datatype (making problems first-order). More lemma names
paulson
parents:
18570
diff
changeset
|
171 |
lemmas shrK_neq_pubK = pubK_neq_shrK [THEN not_sym] |
31c2af8b0c60
replacement of bool by a datatype (making problems first-order). More lemma names
paulson
parents:
18570
diff
changeset
|
172 |
declare shrK_neq_pubK [simp] |
13922 | 173 |
|
174 |
lemma priEK_noteq_shrK [simp]: "priEK A \<noteq> shrK B" |
|
175 |
by auto |
|
176 |
||
177 |
lemma publicKey_notin_image_shrK [simp]: "publicKey b x \<notin> shrK ` AA" |
|
178 |
by auto |
|
179 |
||
180 |
lemma privateKey_notin_image_shrK [simp]: "privateKey b x \<notin> shrK ` AA" |
|
181 |
by auto |
|
182 |
||
183 |
lemma shrK_notin_image_publicKey [simp]: "shrK x \<notin> publicKey b ` AA" |
|
184 |
by auto |
|
185 |
||
186 |
lemma shrK_notin_image_privateKey [simp]: "shrK x \<notin> invKey ` publicKey b ` AA" |
|
187 |
by auto |
|
188 |
||
189 |
lemma shrK_image_eq [simp]: "(shrK x \<in> shrK ` AA) = (x \<in> AA)" |
|
190 |
by auto |
|
191 |
||
61830 | 192 |
text\<open>For some reason, moving this up can make some proofs loop!\<close> |
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
193 |
declare invKey_K [simp] |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
194 |
|
13922 | 195 |
|
61830 | 196 |
subsection\<open>Initial States of Agents\<close> |
13922 | 197 |
|
61830 | 198 |
text\<open>Note: for all practical purposes, all that matters is the initial |
13922 | 199 |
knowledge of the Spy. All other agents are automata, merely following the |
61830 | 200 |
protocol.\<close> |
2318 | 201 |
|
39246 | 202 |
overloading |
203 |
initState \<equiv> initState |
|
204 |
begin |
|
205 |
||
206 |
primrec initState where |
|
2318 | 207 |
(*Agents know their private key and all public keys*) |
13922 | 208 |
initState_Server: |
209 |
"initState Server = |
|
210 |
{Key (priEK Server), Key (priSK Server)} \<union> |
|
211 |
(Key ` range pubEK) \<union> (Key ` range pubSK) \<union> (Key ` range shrK)" |
|
212 |
||
39246 | 213 |
| initState_Friend: |
13922 | 214 |
"initState (Friend i) = |
215 |
{Key (priEK(Friend i)), Key (priSK(Friend i)), Key (shrK(Friend i))} \<union> |
|
216 |
(Key ` range pubEK) \<union> (Key ` range pubSK)" |
|
217 |
||
39246 | 218 |
| initState_Spy: |
13922 | 219 |
"initState Spy = |
220 |
(Key ` invKey ` pubEK ` bad) \<union> (Key ` invKey ` pubSK ` bad) \<union> |
|
221 |
(Key ` shrK ` bad) \<union> |
|
222 |
(Key ` range pubEK) \<union> (Key ` range pubSK)" |
|
223 |
||
39246 | 224 |
end |
225 |
||
13922 | 226 |
|
69597 | 227 |
text\<open>These lemmas allow reasoning about \<^term>\<open>used evs\<close> rather than |
228 |
\<^term>\<open>knows Spy evs\<close>, which is useful when there are private Notes. |
|
229 |
Because they depend upon the definition of \<^term>\<open>initState\<close>, they cannot |
|
61830 | 230 |
be moved up.\<close> |
13922 | 231 |
|
232 |
lemma used_parts_subset_parts [rule_format]: |
|
233 |
"\<forall>X \<in> used evs. parts {X} \<subseteq> used evs" |
|
234 |
apply (induct evs) |
|
235 |
prefer 2 |
|
39251
8756b44582e2
Tidied up proofs using sledgehammer, also deleting unnecessary semicolons
paulson
parents:
39246
diff
changeset
|
236 |
apply (simp add: used_Cons split: event.split) |
8756b44582e2
Tidied up proofs using sledgehammer, also deleting unnecessary semicolons
paulson
parents:
39246
diff
changeset
|
237 |
apply (metis Un_iff empty_subsetI insert_subset le_supI1 le_supI2 parts_subset_iff) |
61830 | 238 |
txt\<open>Base case\<close> |
39251
8756b44582e2
Tidied up proofs using sledgehammer, also deleting unnecessary semicolons
paulson
parents:
39246
diff
changeset
|
239 |
apply (auto dest!: parts_cut simp add: used_Nil) |
13922 | 240 |
done |
241 |
||
67613 | 242 |
lemma MPair_used_D: "\<lbrace>X,Y\<rbrace> \<in> used H ==> X \<in> used H \<and> Y \<in> used H" |
13922 | 243 |
by (drule used_parts_subset_parts, simp, blast) |
244 |
||
61830 | 245 |
text\<open>There was a similar theorem in Event.thy, so perhaps this one can |
246 |
be moved up if proved directly by induction.\<close> |
|
13922 | 247 |
lemma MPair_used [elim!]: |
61956 | 248 |
"[| \<lbrace>X,Y\<rbrace> \<in> used H; |
13922 | 249 |
[| X \<in> used H; Y \<in> used H |] ==> P |] |
250 |
==> P" |
|
251 |
by (blast dest: MPair_used_D) |
|
252 |
||
253 |
||
69597 | 254 |
text\<open>Rewrites should not refer to \<^term>\<open>initState(Friend i)\<close> because |
61830 | 255 |
that expression is not in normal form.\<close> |
13922 | 256 |
|
257 |
lemma keysFor_parts_initState [simp]: "keysFor (parts (initState C)) = {}" |
|
258 |
apply (unfold keysFor_def) |
|
259 |
apply (induct_tac "C") |
|
260 |
apply (auto intro: range_eqI) |
|
261 |
done |
|
262 |
||
263 |
lemma Crypt_notin_initState: "Crypt K X \<notin> parts (initState B)" |
|
264 |
by (induct B, auto) |
|
265 |
||
13935 | 266 |
lemma Crypt_notin_used_empty [simp]: "Crypt K X \<notin> used []" |
267 |
by (simp add: Crypt_notin_initState used_Nil) |
|
268 |
||
13922 | 269 |
(*** Basic properties of shrK ***) |
270 |
||
271 |
(*Agents see their own shared keys!*) |
|
272 |
lemma shrK_in_initState [iff]: "Key (shrK A) \<in> initState A" |
|
13926 | 273 |
by (induct_tac "A", auto) |
13922 | 274 |
|
275 |
lemma shrK_in_knows [iff]: "Key (shrK A) \<in> knows A evs" |
|
276 |
by (simp add: initState_subset_knows [THEN subsetD]) |
|
277 |
||
278 |
lemma shrK_in_used [iff]: "Key (shrK A) \<in> used evs" |
|
13926 | 279 |
by (rule initState_into_used, blast) |
13922 | 280 |
|
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
281 |
|
13922 | 282 |
(** Fresh keys never clash with long-term shared keys **) |
283 |
||
284 |
(*Used in parts_induct_tac and analz_Fake_tac to distinguish session keys |
|
285 |
from long-term shared keys*) |
|
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
286 |
lemma Key_not_used [simp]: "Key K \<notin> used evs ==> K \<notin> range shrK" |
13926 | 287 |
by blast |
13922 | 288 |
|
289 |
lemma shrK_neq: "Key K \<notin> used evs ==> shrK B \<noteq> K" |
|
13926 | 290 |
by blast |
13922 | 291 |
|
18749
31c2af8b0c60
replacement of bool by a datatype (making problems first-order). More lemma names
paulson
parents:
18570
diff
changeset
|
292 |
lemmas neq_shrK = shrK_neq [THEN not_sym] |
31c2af8b0c60
replacement of bool by a datatype (making problems first-order). More lemma names
paulson
parents:
18570
diff
changeset
|
293 |
declare neq_shrK [simp] |
2318 | 294 |
|
295 |
||
69597 | 296 |
subsection\<open>Function \<^term>\<open>spies\<close>\<close> |
13922 | 297 |
|
18749
31c2af8b0c60
replacement of bool by a datatype (making problems first-order). More lemma names
paulson
parents:
18570
diff
changeset
|
298 |
lemma not_SignatureE [elim!]: "b \<noteq> Signature \<Longrightarrow> b = Encryption" |
31c2af8b0c60
replacement of bool by a datatype (making problems first-order). More lemma names
paulson
parents:
18570
diff
changeset
|
299 |
by (cases b, auto) |
31c2af8b0c60
replacement of bool by a datatype (making problems first-order). More lemma names
paulson
parents:
18570
diff
changeset
|
300 |
|
61830 | 301 |
text\<open>Agents see their own private keys!\<close> |
13922 | 302 |
lemma priK_in_initState [iff]: "Key (privateKey b A) \<in> initState A" |
18749
31c2af8b0c60
replacement of bool by a datatype (making problems first-order). More lemma names
paulson
parents:
18570
diff
changeset
|
303 |
by (cases A, auto) |
13922 | 304 |
|
61830 | 305 |
text\<open>Agents see all public keys!\<close> |
13922 | 306 |
lemma publicKey_in_initState [iff]: "Key (publicKey b A) \<in> initState B" |
18749
31c2af8b0c60
replacement of bool by a datatype (making problems first-order). More lemma names
paulson
parents:
18570
diff
changeset
|
307 |
by (cases B, auto) |
13922 | 308 |
|
61830 | 309 |
text\<open>All public keys are visible\<close> |
13922 | 310 |
lemma spies_pubK [iff]: "Key (publicKey b A) \<in> spies evs" |
311 |
apply (induct_tac "evs") |
|
63648 | 312 |
apply (auto simp add: imageI knows_Cons split: event.split) |
13922 | 313 |
done |
314 |
||
18749
31c2af8b0c60
replacement of bool by a datatype (making problems first-order). More lemma names
paulson
parents:
18570
diff
changeset
|
315 |
lemmas analz_spies_pubK = spies_pubK [THEN analz.Inj] |
31c2af8b0c60
replacement of bool by a datatype (making problems first-order). More lemma names
paulson
parents:
18570
diff
changeset
|
316 |
declare analz_spies_pubK [iff] |
13922 | 317 |
|
61830 | 318 |
text\<open>Spy sees private keys of bad agents!\<close> |
13922 | 319 |
lemma Spy_spies_bad_privateKey [intro!]: |
320 |
"A \<in> bad ==> Key (privateKey b A) \<in> spies evs" |
|
321 |
apply (induct_tac "evs") |
|
63648 | 322 |
apply (auto simp add: imageI knows_Cons split: event.split) |
13922 | 323 |
done |
324 |
||
61830 | 325 |
text\<open>Spy sees long-term shared keys of bad agents!\<close> |
13922 | 326 |
lemma Spy_spies_bad_shrK [intro!]: |
327 |
"A \<in> bad ==> Key (shrK A) \<in> spies evs" |
|
328 |
apply (induct_tac "evs") |
|
63648 | 329 |
apply (simp_all add: imageI knows_Cons split: event.split) |
13922 | 330 |
done |
331 |
||
332 |
lemma publicKey_into_used [iff] :"Key (publicKey b A) \<in> used evs" |
|
333 |
apply (rule initState_into_used) |
|
334 |
apply (rule publicKey_in_initState [THEN parts.Inj]) |
|
335 |
done |
|
336 |
||
337 |
lemma privateKey_into_used [iff]: "Key (privateKey b A) \<in> used evs" |
|
338 |
apply(rule initState_into_used) |
|
339 |
apply(rule priK_in_initState [THEN parts.Inj]) |
|
340 |
done |
|
341 |
||
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
342 |
(*For case analysis on whether or not an agent is compromised*) |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
343 |
lemma Crypt_Spy_analz_bad: |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
344 |
"[| Crypt (shrK A) X \<in> analz (knows Spy evs); A \<in> bad |] |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
345 |
==> X \<in> analz (knows Spy evs)" |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
346 |
by force |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
347 |
|
13922 | 348 |
|
61830 | 349 |
subsection\<open>Fresh Nonces\<close> |
13922 | 350 |
|
351 |
lemma Nonce_notin_initState [iff]: "Nonce N \<notin> parts (initState B)" |
|
13926 | 352 |
by (induct_tac "B", auto) |
13922 | 353 |
|
354 |
lemma Nonce_notin_used_empty [simp]: "Nonce N \<notin> used []" |
|
13926 | 355 |
by (simp add: used_Nil) |
13922 | 356 |
|
11104 | 357 |
|
61830 | 358 |
subsection\<open>Supply fresh nonces for possibility theorems\<close> |
13922 | 359 |
|
61830 | 360 |
text\<open>In any trace, there is an upper bound N on the greatest nonce in use\<close> |
67613 | 361 |
lemma Nonce_supply_lemma: "\<exists>N. \<forall>n. N\<le>n \<longrightarrow> Nonce n \<notin> used evs" |
13922 | 362 |
apply (induct_tac "evs") |
13926 | 363 |
apply (rule_tac x = 0 in exI) |
63648 | 364 |
apply (simp_all (no_asm_simp) add: used_Cons split: event.split) |
13922 | 365 |
apply safe |
366 |
apply (rule msg_Nonce_supply [THEN exE], blast elim!: add_leE)+ |
|
367 |
done |
|
368 |
||
67613 | 369 |
lemma Nonce_supply1: "\<exists>N. Nonce N \<notin> used evs" |
13926 | 370 |
by (rule Nonce_supply_lemma [THEN exE], blast) |
13922 | 371 |
|
67613 | 372 |
lemma Nonce_supply: "Nonce (SOME N. Nonce N \<notin> used evs) \<notin> used evs" |
13922 | 373 |
apply (rule Nonce_supply_lemma [THEN exE]) |
13926 | 374 |
apply (rule someI, fast) |
13922 | 375 |
done |
376 |
||
69597 | 377 |
subsection\<open>Specialized Rewriting for Theorems About \<^term>\<open>analz\<close> and Image\<close> |
13922 | 378 |
|
67613 | 379 |
lemma insert_Key_singleton: "insert (Key K) H = Key ` {K} \<union> H" |
13926 | 380 |
by blast |
11104 | 381 |
|
13922 | 382 |
lemma insert_Key_image: "insert (Key K) (Key`KK \<union> C) = Key ` (insert K KK) \<union> C" |
13926 | 383 |
by blast |
13922 | 384 |
|
385 |
||
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
386 |
lemma Crypt_imp_keysFor :"[|Crypt K X \<in> H; K \<in> symKeys|] ==> K \<in> keysFor H" |
13926 | 387 |
by (drule Crypt_imp_invKey_keysFor, simp) |
13922 | 388 |
|
61830 | 389 |
text\<open>Lemma for the trivial direction of the if-and-only-if of the |
390 |
Session Key Compromise Theorem\<close> |
|
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
391 |
lemma analz_image_freshK_lemma: |
67613 | 392 |
"(Key K \<in> analz (Key`nE \<union> H)) \<longrightarrow> (K \<in> nE | Key K \<in> analz H) ==> |
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
393 |
(Key K \<in> analz (Key`nE \<union> H)) = (K \<in> nE | Key K \<in> analz H)" |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
394 |
by (blast intro: analz_mono [THEN [2] rev_subsetD]) |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
395 |
|
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
396 |
lemmas analz_image_freshK_simps = |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
397 |
simp_thms mem_simps \<comment> \<open>these two allow its use with \<open>only:\<close>\<close> |
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
398 |
disj_comms |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
399 |
image_insert [THEN sym] image_Un [THEN sym] empty_subsetI insert_subset |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
400 |
analz_insert_eq Un_upper2 [THEN analz_mono, THEN subsetD] |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
401 |
insert_Key_singleton |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
402 |
Key_not_used insert_Key_image Un_assoc [THEN sym] |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
403 |
|
61830 | 404 |
ML \<open> |
24122 | 405 |
structure Public = |
406 |
struct |
|
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
407 |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
41774
diff
changeset
|
408 |
val analz_image_freshK_ss = |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
41774
diff
changeset
|
409 |
simpset_of |
69597 | 410 |
(\<^context> delsimps [image_insert, image_Un] |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
41774
diff
changeset
|
411 |
delsimps [@{thm imp_disjL}] (*reduces blow-up*) |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
41774
diff
changeset
|
412 |
addsimps @{thms analz_image_freshK_simps}) |
13922 | 413 |
|
23894
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23315
diff
changeset
|
414 |
(*Tactic for possibility theorems*) |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23315
diff
changeset
|
415 |
fun possibility_tac ctxt = |
13922 | 416 |
REPEAT (*omit used_Says so that Nonces start from different traces!*) |
56073
29e308b56d23
enhanced simplifier solver for preconditions of rewrite rule, can now deal with conjunctions
nipkow
parents:
55416
diff
changeset
|
417 |
(ALLGOALS (simp_tac (ctxt setSolver safe_solver delsimps [@{thm used_Says}])) |
13922 | 418 |
THEN |
419 |
REPEAT_FIRST (eq_assume_tac ORELSE' |
|
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58310
diff
changeset
|
420 |
resolve_tac ctxt [refl, conjI, @{thm Nonce_supply}])) |
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
421 |
|
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
422 |
(*For harder protocols (such as Recur) where we have to set up some |
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
423 |
nonces and keys initially*) |
23894
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23315
diff
changeset
|
424 |
fun basic_possibility_tac ctxt = |
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
425 |
REPEAT |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
41774
diff
changeset
|
426 |
(ALLGOALS (asm_simp_tac (ctxt setSolver safe_solver)) |
14207
f20fbb141673
Conversion of all main protocols from "Shared" to "Public".
paulson
parents:
14200
diff
changeset
|
427 |
THEN |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58310
diff
changeset
|
428 |
REPEAT_FIRST (resolve_tac ctxt [refl, conjI])) |
24122 | 429 |
|
430 |
end |
|
61830 | 431 |
\<close> |
11104 | 432 |
|
61830 | 433 |
method_setup analz_freshK = \<open> |
30549 | 434 |
Scan.succeed (fn ctxt => |
30510
4120fc59dd85
unified type Proof.method and pervasive METHOD combinators;
wenzelm
parents:
24122
diff
changeset
|
435 |
(SIMPLE_METHOD |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58310
diff
changeset
|
436 |
(EVERY [REPEAT_FIRST (resolve_tac ctxt [allI, ballI, impI]), |
60754 | 437 |
REPEAT_FIRST (resolve_tac ctxt @{thms analz_image_freshK_lemma}), |
61830 | 438 |
ALLGOALS (asm_simp_tac (put_simpset Public.analz_image_freshK_ss ctxt))])))\<close> |
24122 | 439 |
"for proving the Session Key Compromise theorem" |
440 |
||
441 |
||
61830 | 442 |
subsection\<open>Specialized Methods for Possibility Theorems\<close> |
24122 | 443 |
|
61830 | 444 |
method_setup possibility = \<open> |
445 |
Scan.succeed (SIMPLE_METHOD o Public.possibility_tac)\<close> |
|
23894
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23315
diff
changeset
|
446 |
"for proving possibility theorems" |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23315
diff
changeset
|
447 |
|
61830 | 448 |
method_setup basic_possibility = \<open> |
449 |
Scan.succeed (SIMPLE_METHOD o Public.basic_possibility_tac)\<close> |
|
11104 | 450 |
"for proving possibility theorems" |
2318 | 451 |
|
452 |
end |