author | wenzelm |
Mon, 08 Jun 2020 21:55:14 +0200 | |
changeset 71926 | bee83c9d3306 |
parent 69597 | ff784d5a5bfb |
permissions | -rw-r--r-- |
33028 | 1 |
(* Title: HOL/SET_Protocol/Cardholder_Registration.thy |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
2 |
Author: Giampaolo Bella |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
3 |
Author: Fabio Massacci |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
4 |
Author: Lawrence C Paulson |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
5 |
Author: Piero Tramontano |
14199 | 6 |
*) |
7 |
||
63167 | 8 |
section\<open>The SET Cardholder Registration Protocol\<close> |
14199 | 9 |
|
33028 | 10 |
theory Cardholder_Registration |
11 |
imports Public_SET |
|
12 |
begin |
|
14199 | 13 |
|
63167 | 14 |
text\<open>Note: nonces seem to consist of 20 bytes. That includes both freshness |
14199 | 15 |
challenges (Chall-EE, etc.) and important secrets (CardSecret, PANsecret) |
63167 | 16 |
\<close> |
14199 | 17 |
|
63167 | 18 |
text\<open>Simplifications involving \<open>analz_image_keys_simps\<close> appear to |
14199 | 19 |
have become much slower. The cause is unclear. However, there is a big blow-up |
63167 | 20 |
and the rewriting is very sensitive to the set of rewrite rules given.\<close> |
14199 | 21 |
|
63167 | 22 |
subsection\<open>Predicate Formalizing the Encryption Association between Keys\<close> |
14199 | 23 |
|
67613 | 24 |
primrec KeyCryptKey :: "[key, key, event list] \<Rightarrow> bool" |
39758 | 25 |
where |
26 |
KeyCryptKey_Nil: |
|
27 |
"KeyCryptKey DK K [] = False" |
|
28 |
| KeyCryptKey_Cons: |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
29 |
\<comment> \<open>Says is the only important case. |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
30 |
1st case: CR5, where KC3 encrypts KC2. |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
31 |
2nd case: any use of priEK C. |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
32 |
Revision 1.12 has a more complicated version with separate treatment of |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
33 |
the dependency of KC1, KC2 and KC3 on priEK (CA i.) Not needed since |
63167 | 34 |
priEK C is never sent (and so can't be lost except at the start).\<close> |
39758 | 35 |
"KeyCryptKey DK K (ev # evs) = |
36 |
(KeyCryptKey DK K evs | |
|
37 |
(case ev of |
|
67613 | 38 |
Says A B Z \<Rightarrow> |
39 |
((\<exists>N X Y. A \<noteq> Spy \<and> |
|
40 |
DK \<in> symKeys \<and> |
|
61984 | 41 |
Z = \<lbrace>Crypt DK \<lbrace>Agent A, Nonce N, Key K, X\<rbrace>, Y\<rbrace>) | |
39758 | 42 |
(\<exists>C. DK = priEK C)) |
67613 | 43 |
| Gets A' X \<Rightarrow> False |
44 |
| Notes A' X \<Rightarrow> False))" |
|
14199 | 45 |
|
46 |
||
63167 | 47 |
subsection\<open>Predicate formalizing the association between keys and nonces\<close> |
14199 | 48 |
|
67613 | 49 |
primrec KeyCryptNonce :: "[key, key, event list] \<Rightarrow> bool" |
39758 | 50 |
where |
51 |
KeyCryptNonce_Nil: |
|
52 |
"KeyCryptNonce EK K [] = False" |
|
53 |
| KeyCryptNonce_Cons: |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
54 |
\<comment> \<open>Says is the only important case. |
14199 | 55 |
1st case: CR3, where KC1 encrypts NC2 (distinct from CR5 due to EXH); |
56 |
2nd case: CR5, where KC3 encrypts NC3; |
|
57 |
3rd case: CR6, where KC2 encrypts NC3; |
|
58 |
4th case: CR6, where KC2 encrypts NonceCCA; |
|
69597 | 59 |
5th case: any use of \<^term>\<open>priEK C\<close> (including CardSecret). |
14199 | 60 |
NB the only Nonces we need to keep secret are CardSecret and NonceCCA. |
63167 | 61 |
But we can't prove \<open>Nonce_compromise\<close> unless the relation covers ALL |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
62 |
nonces that the protocol keeps secret.\<close> |
14199 | 63 |
"KeyCryptNonce DK N (ev # evs) = |
64 |
(KeyCryptNonce DK N evs | |
|
65 |
(case ev of |
|
67613 | 66 |
Says A B Z \<Rightarrow> |
67 |
A \<noteq> Spy \<and> |
|
68 |
((\<exists>X Y. DK \<in> symKeys \<and> |
|
61984 | 69 |
Z = (EXHcrypt DK X \<lbrace>Agent A, Nonce N\<rbrace> Y)) | |
67613 | 70 |
(\<exists>X Y. DK \<in> symKeys \<and> |
61984 | 71 |
Z = \<lbrace>Crypt DK \<lbrace>Agent A, Nonce N, X\<rbrace>, Y\<rbrace>) | |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
72 |
(\<exists>K i X Y. |
67613 | 73 |
K \<in> symKeys \<and> |
74 |
Z = Crypt K \<lbrace>sign (priSK (CA i)) \<lbrace>Agent B, Nonce N, X\<rbrace>, Y\<rbrace> \<and> |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
75 |
(DK=K | KeyCryptKey DK K evs)) | |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
76 |
(\<exists>K C NC3 Y. |
67613 | 77 |
K \<in> symKeys \<and> |
14199 | 78 |
Z = Crypt K |
61984 | 79 |
\<lbrace>sign (priSK C) \<lbrace>Agent B, Nonce NC3, Agent C, Nonce N\<rbrace>, |
67613 | 80 |
Y\<rbrace> \<and> |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
81 |
(DK=K | KeyCryptKey DK K evs)) | |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
82 |
(\<exists>C. DK = priEK C)) |
67613 | 83 |
| Gets A' X \<Rightarrow> False |
84 |
| Notes A' X \<Rightarrow> False))" |
|
14199 | 85 |
|
86 |
||
63167 | 87 |
subsection\<open>Formal protocol definition\<close> |
14199 | 88 |
|
23755 | 89 |
inductive_set |
90 |
set_cr :: "event list set" |
|
91 |
where |
|
14199 | 92 |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
93 |
Nil: \<comment> \<open>Initial trace is empty\<close> |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
94 |
"[] \<in> set_cr" |
14199 | 95 |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
96 |
| Fake: \<comment> \<open>The spy MAY say anything he CAN say.\<close> |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
97 |
"[| evsf \<in> set_cr; X \<in> synth (analz (knows Spy evsf)) |] |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
98 |
==> Says Spy B X # evsf \<in> set_cr" |
14199 | 99 |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
100 |
| Reception: \<comment> \<open>If A sends a message X to B, then B might receive it\<close> |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
101 |
"[| evsr \<in> set_cr; Says A B X \<in> set evsr |] |
14199 | 102 |
==> Gets B X # evsr \<in> set_cr" |
103 |
||
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
104 |
| SET_CR1: \<comment> \<open>CardCInitReq: C initiates a run, sending a nonce to CCA\<close> |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
105 |
"[| evs1 \<in> set_cr; C = Cardholder k; Nonce NC1 \<notin> used evs1 |] |
61984 | 106 |
==> Says C (CA i) \<lbrace>Agent C, Nonce NC1\<rbrace> # evs1 \<in> set_cr" |
14199 | 107 |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
108 |
| SET_CR2: \<comment> \<open>CardCInitRes: CA responds sending NC1 and its certificates\<close> |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
109 |
"[| evs2 \<in> set_cr; |
61984 | 110 |
Gets (CA i) \<lbrace>Agent C, Nonce NC1\<rbrace> \<in> set evs2 |] |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
111 |
==> Says (CA i) C |
61984 | 112 |
\<lbrace>sign (priSK (CA i)) \<lbrace>Agent C, Nonce NC1\<rbrace>, |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
113 |
cert (CA i) (pubEK (CA i)) onlyEnc (priSK RCA), |
61984 | 114 |
cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)\<rbrace> |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
115 |
# evs2 \<in> set_cr" |
14199 | 116 |
|
23755 | 117 |
| SET_CR3: |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
118 |
\<comment> \<open>RegFormReq: C sends his PAN and a new nonce to CA. |
14199 | 119 |
C verifies that |
120 |
- nonce received is the same as that sent; |
|
121 |
- certificates are signed by RCA; |
|
122 |
- certificates are an encryption certificate (flag is onlyEnc) and a |
|
123 |
signature certificate (flag is onlySig); |
|
124 |
- certificates pertain to the CA that C contacted (this is done by |
|
125 |
checking the signature). |
|
126 |
C generates a fresh symmetric key KC1. |
|
69597 | 127 |
The point of encrypting \<^term>\<open>\<lbrace>Agent C, Nonce NC2, Hash (Pan(pan C))\<rbrace>\<close> |
63167 | 128 |
is not clear.\<close> |
14199 | 129 |
"[| evs3 \<in> set_cr; C = Cardholder k; |
130 |
Nonce NC2 \<notin> used evs3; |
|
131 |
Key KC1 \<notin> used evs3; KC1 \<in> symKeys; |
|
61984 | 132 |
Gets C \<lbrace>sign (invKey SKi) \<lbrace>Agent X, Nonce NC1\<rbrace>, |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
133 |
cert (CA i) EKi onlyEnc (priSK RCA), |
61984 | 134 |
cert (CA i) SKi onlySig (priSK RCA)\<rbrace> |
14199 | 135 |
\<in> set evs3; |
61984 | 136 |
Says C (CA i) \<lbrace>Agent C, Nonce NC1\<rbrace> \<in> set evs3|] |
137 |
==> Says C (CA i) (EXHcrypt KC1 EKi \<lbrace>Agent C, Nonce NC2\<rbrace> (Pan(pan C))) |
|
138 |
# Notes C \<lbrace>Key KC1, Agent (CA i)\<rbrace> |
|
14199 | 139 |
# evs3 \<in> set_cr" |
140 |
||
23755 | 141 |
| SET_CR4: |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
142 |
\<comment> \<open>RegFormRes: |
14199 | 143 |
CA responds sending NC2 back with a new nonce NCA, after checking that |
69597 | 144 |
- the digital envelope is correctly encrypted by \<^term>\<open>pubEK (CA i)\<close> |
14199 | 145 |
- the entire message is encrypted with the same key found inside the |
63167 | 146 |
envelope (here, KC1)\<close> |
14199 | 147 |
"[| evs4 \<in> set_cr; |
148 |
Nonce NCA \<notin> used evs4; KC1 \<in> symKeys; |
|
61984 | 149 |
Gets (CA i) (EXHcrypt KC1 EKi \<lbrace>Agent C, Nonce NC2\<rbrace> (Pan(pan X))) |
14199 | 150 |
\<in> set evs4 |] |
151 |
==> Says (CA i) C |
|
61984 | 152 |
\<lbrace>sign (priSK (CA i)) \<lbrace>Agent C, Nonce NC2, Nonce NCA\<rbrace>, |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
153 |
cert (CA i) (pubEK (CA i)) onlyEnc (priSK RCA), |
61984 | 154 |
cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)\<rbrace> |
14199 | 155 |
# evs4 \<in> set_cr" |
156 |
||
23755 | 157 |
| SET_CR5: |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
158 |
\<comment> \<open>CertReq: C sends his PAN, a new nonce, its proposed public signature key |
14199 | 159 |
and its half of the secret value to CA. |
160 |
We now assume that C has a fixed key pair, and he submits (pubSK C). |
|
161 |
The protocol does not require this key to be fresh. |
|
63167 | 162 |
The encryption below is actually EncX.\<close> |
14199 | 163 |
"[| evs5 \<in> set_cr; C = Cardholder k; |
164 |
Nonce NC3 \<notin> used evs5; Nonce CardSecret \<notin> used evs5; NC3\<noteq>CardSecret; |
|
165 |
Key KC2 \<notin> used evs5; KC2 \<in> symKeys; |
|
166 |
Key KC3 \<notin> used evs5; KC3 \<in> symKeys; KC2\<noteq>KC3; |
|
61984 | 167 |
Gets C \<lbrace>sign (invKey SKi) \<lbrace>Agent C, Nonce NC2, Nonce NCA\<rbrace>, |
14199 | 168 |
cert (CA i) EKi onlyEnc (priSK RCA), |
61984 | 169 |
cert (CA i) SKi onlySig (priSK RCA) \<rbrace> |
14199 | 170 |
\<in> set evs5; |
61984 | 171 |
Says C (CA i) (EXHcrypt KC1 EKi \<lbrace>Agent C, Nonce NC2\<rbrace> (Pan(pan C))) |
14199 | 172 |
\<in> set evs5 |] |
173 |
==> Says C (CA i) |
|
61984 | 174 |
\<lbrace>Crypt KC3 |
175 |
\<lbrace>Agent C, Nonce NC3, Key KC2, Key (pubSK C), |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
176 |
Crypt (priSK C) |
61984 | 177 |
(Hash \<lbrace>Agent C, Nonce NC3, Key KC2, |
178 |
Key (pubSK C), Pan (pan C), Nonce CardSecret\<rbrace>)\<rbrace>, |
|
179 |
Crypt EKi \<lbrace>Key KC3, Pan (pan C), Nonce CardSecret\<rbrace> \<rbrace> |
|
180 |
# Notes C \<lbrace>Key KC2, Agent (CA i)\<rbrace> |
|
181 |
# Notes C \<lbrace>Key KC3, Agent (CA i)\<rbrace> |
|
14199 | 182 |
# evs5 \<in> set_cr" |
183 |
||
184 |
||
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
185 |
\<comment> \<open>CertRes: CA responds sending NC3 back with its half of the secret value, |
14199 | 186 |
its signature certificate and the new cardholder signature |
187 |
certificate. CA checks to have never certified the key proposed by C. |
|
188 |
NOTE: In Merchant Registration, the corresponding rule (4) |
|
69597 | 189 |
uses the "sign" primitive. The encryption below is actually \<^term>\<open>EncK\<close>, |
190 |
which is just \<^term>\<open>Crypt K (sign SK X)\<close>.\<close> |
|
14199 | 191 |
|
23755 | 192 |
| SET_CR6: |
14199 | 193 |
"[| evs6 \<in> set_cr; |
194 |
Nonce NonceCCA \<notin> used evs6; |
|
195 |
KC2 \<in> symKeys; KC3 \<in> symKeys; cardSK \<notin> symKeys; |
|
196 |
Notes (CA i) (Key cardSK) \<notin> set evs6; |
|
197 |
Gets (CA i) |
|
61984 | 198 |
\<lbrace>Crypt KC3 \<lbrace>Agent C, Nonce NC3, Key KC2, Key cardSK, |
14199 | 199 |
Crypt (invKey cardSK) |
61984 | 200 |
(Hash \<lbrace>Agent C, Nonce NC3, Key KC2, |
201 |
Key cardSK, Pan (pan C), Nonce CardSecret\<rbrace>)\<rbrace>, |
|
202 |
Crypt (pubEK (CA i)) \<lbrace>Key KC3, Pan (pan C), Nonce CardSecret\<rbrace> \<rbrace> |
|
14199 | 203 |
\<in> set evs6 |] |
204 |
==> Says (CA i) C |
|
205 |
(Crypt KC2 |
|
61984 | 206 |
\<lbrace>sign (priSK (CA i)) |
207 |
\<lbrace>Agent C, Nonce NC3, Agent(CA i), Nonce NonceCCA\<rbrace>, |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
208 |
certC (pan C) cardSK (XOR(CardSecret,NonceCCA)) onlySig (priSK (CA i)), |
61984 | 209 |
cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)\<rbrace>) |
14199 | 210 |
# Notes (CA i) (Key cardSK) |
211 |
# evs6 \<in> set_cr" |
|
212 |
||
213 |
||
214 |
declare Says_imp_knows_Spy [THEN parts.Inj, dest] |
|
215 |
declare parts.Body [dest] |
|
216 |
declare analz_into_parts [dest] |
|
217 |
declare Fake_parts_insert_in_Un [dest] |
|
218 |
||
63167 | 219 |
text\<open>A "possibility property": there are traces that reach the end. |
220 |
An unconstrained proof with many subgoals.\<close> |
|
14199 | 221 |
|
222 |
lemma Says_to_Gets: |
|
223 |
"Says A B X # evs \<in> set_cr ==> Gets B X # Says A B X # evs \<in> set_cr" |
|
224 |
by (rule set_cr.Reception, auto) |
|
225 |
||
63167 | 226 |
text\<open>The many nonces and keys generated, some simultaneously, force us to |
227 |
introduce them explicitly as shown below.\<close> |
|
14199 | 228 |
lemma possibility_CR6: |
229 |
"[|NC1 < (NC2::nat); NC2 < NC3; NC3 < NCA ; |
|
230 |
NCA < NonceCCA; NonceCCA < CardSecret; |
|
231 |
KC1 < (KC2::key); KC2 < KC3; |
|
232 |
KC1 \<in> symKeys; Key KC1 \<notin> used []; |
|
233 |
KC2 \<in> symKeys; Key KC2 \<notin> used []; |
|
234 |
KC3 \<in> symKeys; Key KC3 \<notin> used []; |
|
235 |
C = Cardholder k|] |
|
236 |
==> \<exists>evs \<in> set_cr. |
|
237 |
Says (CA i) C |
|
238 |
(Crypt KC2 |
|
61984 | 239 |
\<lbrace>sign (priSK (CA i)) |
240 |
\<lbrace>Agent C, Nonce NC3, Agent(CA i), Nonce NonceCCA\<rbrace>, |
|
14199 | 241 |
certC (pan C) (pubSK (Cardholder k)) (XOR(CardSecret,NonceCCA)) |
242 |
onlySig (priSK (CA i)), |
|
61984 | 243 |
cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)\<rbrace>) |
14199 | 244 |
\<in> set evs" |
245 |
apply (intro exI bexI) |
|
14206
77bf175f5145
Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents:
14199
diff
changeset
|
246 |
apply (rule_tac [2] |
77bf175f5145
Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents:
14199
diff
changeset
|
247 |
set_cr.Nil |
77bf175f5145
Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents:
14199
diff
changeset
|
248 |
[THEN set_cr.SET_CR1 [of concl: C i NC1], |
77bf175f5145
Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents:
14199
diff
changeset
|
249 |
THEN Says_to_Gets, |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
250 |
THEN set_cr.SET_CR2 [of concl: i C NC1], |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
251 |
THEN Says_to_Gets, |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
252 |
THEN set_cr.SET_CR3 [of concl: C i KC1 _ NC2], |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
253 |
THEN Says_to_Gets, |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
254 |
THEN set_cr.SET_CR4 [of concl: i C NC2 NCA], |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
255 |
THEN Says_to_Gets, |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
256 |
THEN set_cr.SET_CR5 [of concl: C i KC3 NC3 KC2 CardSecret], |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
257 |
THEN Says_to_Gets, |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32404
diff
changeset
|
258 |
THEN set_cr.SET_CR6 [of concl: i C KC2]]) |
30607
c3d1590debd8
eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents:
30549
diff
changeset
|
259 |
apply basic_possibility |
14199 | 260 |
apply (simp_all (no_asm_simp) add: symKeys_neq_imp_neq) |
261 |
done |
|
262 |
||
63167 | 263 |
text\<open>General facts about message reception\<close> |
14199 | 264 |
lemma Gets_imp_Says: |
265 |
"[| Gets B X \<in> set evs; evs \<in> set_cr |] ==> \<exists>A. Says A B X \<in> set evs" |
|
266 |
apply (erule rev_mp) |
|
267 |
apply (erule set_cr.induct, auto) |
|
268 |
done |
|
269 |
||
270 |
lemma Gets_imp_knows_Spy: |
|
271 |
"[| Gets B X \<in> set evs; evs \<in> set_cr |] ==> X \<in> knows Spy evs" |
|
272 |
by (blast dest!: Gets_imp_Says Says_imp_knows_Spy) |
|
273 |
declare Gets_imp_knows_Spy [THEN parts.Inj, dest] |
|
274 |
||
275 |
||
63167 | 276 |
subsection\<open>Proofs on keys\<close> |
14199 | 277 |
|
63167 | 278 |
text\<open>Spy never sees an agent's private keys! (unless it's bad at start)\<close> |
14199 | 279 |
|
280 |
lemma Spy_see_private_Key [simp]: |
|
281 |
"evs \<in> set_cr |
|
282 |
==> (Key(invKey (publicKey b A)) \<in> parts(knows Spy evs)) = (A \<in> bad)" |
|
283 |
by (erule set_cr.induct, auto) |
|
284 |
||
285 |
lemma Spy_analz_private_Key [simp]: |
|
286 |
"evs \<in> set_cr ==> |
|
287 |
(Key(invKey (publicKey b A)) \<in> analz(knows Spy evs)) = (A \<in> bad)" |
|
288 |
by auto |
|
289 |
||
290 |
declare Spy_see_private_Key [THEN [2] rev_iffD1, dest!] |
|
291 |
declare Spy_analz_private_Key [THEN [2] rev_iffD1, dest!] |
|
292 |
||
293 |
||
63167 | 294 |
subsection\<open>Begin Piero's Theorems on Certificates\<close> |
295 |
text\<open>Trivial in the current model, where certificates by RCA are secure\<close> |
|
14199 | 296 |
|
297 |
lemma Crypt_valid_pubEK: |
|
61984 | 298 |
"[| Crypt (priSK RCA) \<lbrace>Agent C, Key EKi, onlyEnc\<rbrace> |
14199 | 299 |
\<in> parts (knows Spy evs); |
300 |
evs \<in> set_cr |] ==> EKi = pubEK C" |
|
301 |
apply (erule rev_mp) |
|
302 |
apply (erule set_cr.induct, auto) |
|
303 |
done |
|
304 |
||
305 |
lemma certificate_valid_pubEK: |
|
306 |
"[| cert C EKi onlyEnc (priSK RCA) \<in> parts (knows Spy evs); |
|
307 |
evs \<in> set_cr |] |
|
308 |
==> EKi = pubEK C" |
|
309 |
apply (unfold cert_def signCert_def) |
|
310 |
apply (blast dest!: Crypt_valid_pubEK) |
|
311 |
done |
|
312 |
||
313 |
lemma Crypt_valid_pubSK: |
|
61984 | 314 |
"[| Crypt (priSK RCA) \<lbrace>Agent C, Key SKi, onlySig\<rbrace> |
14199 | 315 |
\<in> parts (knows Spy evs); |
316 |
evs \<in> set_cr |] ==> SKi = pubSK C" |
|
317 |
apply (erule rev_mp) |
|
318 |
apply (erule set_cr.induct, auto) |
|
319 |
done |
|
320 |
||
321 |
lemma certificate_valid_pubSK: |
|
322 |
"[| cert C SKi onlySig (priSK RCA) \<in> parts (knows Spy evs); |
|
323 |
evs \<in> set_cr |] ==> SKi = pubSK C" |
|
324 |
apply (unfold cert_def signCert_def) |
|
325 |
apply (blast dest!: Crypt_valid_pubSK) |
|
326 |
done |
|
327 |
||
328 |
lemma Gets_certificate_valid: |
|
61984 | 329 |
"[| Gets A \<lbrace> X, cert C EKi onlyEnc (priSK RCA), |
330 |
cert C SKi onlySig (priSK RCA)\<rbrace> \<in> set evs; |
|
14199 | 331 |
evs \<in> set_cr |] |
67613 | 332 |
==> EKi = pubEK C \<and> SKi = pubSK C" |
14199 | 333 |
by (blast dest: certificate_valid_pubEK certificate_valid_pubSK) |
334 |
||
63167 | 335 |
text\<open>Nobody can have used non-existent keys!\<close> |
14199 | 336 |
lemma new_keys_not_used: |
337 |
"[|K \<in> symKeys; Key K \<notin> used evs; evs \<in> set_cr|] |
|
338 |
==> K \<notin> keysFor (parts (knows Spy evs))" |
|
339 |
apply (erule rev_mp) |
|
340 |
apply (erule rev_mp) |
|
341 |
apply (erule set_cr.induct) |
|
342 |
apply (frule_tac [8] Gets_certificate_valid) |
|
343 |
apply (frule_tac [6] Gets_certificate_valid, simp_all) |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
344 |
apply (force dest!: usedI keysFor_parts_insert) \<comment> \<open>Fake\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
345 |
apply (blast,auto) \<comment> \<open>Others\<close> |
14199 | 346 |
done |
347 |
||
348 |
||
63167 | 349 |
subsection\<open>New versions: as above, but generalized to have the KK argument\<close> |
14199 | 350 |
|
351 |
lemma gen_new_keys_not_used: |
|
352 |
"[|Key K \<notin> used evs; K \<in> symKeys; evs \<in> set_cr |] |
|
67613 | 353 |
==> Key K \<notin> used evs \<longrightarrow> K \<in> symKeys \<longrightarrow> |
354 |
K \<notin> keysFor (parts (Key`KK \<union> knows Spy evs))" |
|
14199 | 355 |
by (auto simp add: new_keys_not_used) |
356 |
||
357 |
lemma gen_new_keys_not_analzd: |
|
358 |
"[|Key K \<notin> used evs; K \<in> symKeys; evs \<in> set_cr |] |
|
67613 | 359 |
==> K \<notin> keysFor (analz (Key`KK \<union> knows Spy evs))" |
14199 | 360 |
by (blast intro: keysFor_mono [THEN [2] rev_subsetD] |
361 |
dest: gen_new_keys_not_used) |
|
362 |
||
363 |
lemma analz_Key_image_insert_eq: |
|
364 |
"[|K \<in> symKeys; Key K \<notin> used evs; evs \<in> set_cr |] |
|
365 |
==> analz (Key ` (insert K KK) \<union> knows Spy evs) = |
|
366 |
insert (Key K) (analz (Key ` KK \<union> knows Spy evs))" |
|
367 |
by (simp add: gen_new_keys_not_analzd) |
|
368 |
||
369 |
lemma Crypt_parts_imp_used: |
|
370 |
"[|Crypt K X \<in> parts (knows Spy evs); |
|
371 |
K \<in> symKeys; evs \<in> set_cr |] ==> Key K \<in> used evs" |
|
372 |
apply (rule ccontr) |
|
373 |
apply (force dest: new_keys_not_used Crypt_imp_invKey_keysFor) |
|
374 |
done |
|
375 |
||
376 |
lemma Crypt_analz_imp_used: |
|
377 |
"[|Crypt K X \<in> analz (knows Spy evs); |
|
378 |
K \<in> symKeys; evs \<in> set_cr |] ==> Key K \<in> used evs" |
|
379 |
by (blast intro: Crypt_parts_imp_used) |
|
380 |
||
381 |
||
14218 | 382 |
(*<*) |
63167 | 383 |
subsection\<open>Messages signed by CA\<close> |
14199 | 384 |
|
63167 | 385 |
text\<open>Message \<open>SET_CR2\<close>: C can check CA's signature if he has received |
386 |
CA's certificate.\<close> |
|
14199 | 387 |
lemma CA_Says_2_lemma: |
61984 | 388 |
"[| Crypt (priSK (CA i)) (Hash\<lbrace>Agent C, Nonce NC1\<rbrace>) |
14199 | 389 |
\<in> parts (knows Spy evs); |
390 |
evs \<in> set_cr; (CA i) \<notin> bad |] |
|
61984 | 391 |
==> \<exists>Y. Says (CA i) C \<lbrace>sign (priSK (CA i)) \<lbrace>Agent C, Nonce NC1\<rbrace>, Y\<rbrace> |
14199 | 392 |
\<in> set evs" |
393 |
apply (erule rev_mp) |
|
394 |
apply (erule set_cr.induct, auto) |
|
395 |
done |
|
396 |
||
63167 | 397 |
text\<open>Ever used?\<close> |
14199 | 398 |
lemma CA_Says_2: |
61984 | 399 |
"[| Crypt (invKey SK) (Hash\<lbrace>Agent C, Nonce NC1\<rbrace>) |
14199 | 400 |
\<in> parts (knows Spy evs); |
401 |
cert (CA i) SK onlySig (priSK RCA) \<in> parts (knows Spy evs); |
|
402 |
evs \<in> set_cr; (CA i) \<notin> bad |] |
|
61984 | 403 |
==> \<exists>Y. Says (CA i) C \<lbrace>sign (priSK (CA i)) \<lbrace>Agent C, Nonce NC1\<rbrace>, Y\<rbrace> |
14199 | 404 |
\<in> set evs" |
405 |
by (blast dest!: certificate_valid_pubSK intro!: CA_Says_2_lemma) |
|
406 |
||
407 |
||
63167 | 408 |
text\<open>Message \<open>SET_CR4\<close>: C can check CA's signature if he has received |
409 |
CA's certificate.\<close> |
|
14199 | 410 |
lemma CA_Says_4_lemma: |
61984 | 411 |
"[| Crypt (priSK (CA i)) (Hash\<lbrace>Agent C, Nonce NC2, Nonce NCA\<rbrace>) |
14199 | 412 |
\<in> parts (knows Spy evs); |
413 |
evs \<in> set_cr; (CA i) \<notin> bad |] |
|
61984 | 414 |
==> \<exists>Y. Says (CA i) C \<lbrace>sign (priSK (CA i)) |
415 |
\<lbrace>Agent C, Nonce NC2, Nonce NCA\<rbrace>, Y\<rbrace> \<in> set evs" |
|
14199 | 416 |
apply (erule rev_mp) |
417 |
apply (erule set_cr.induct, auto) |
|
418 |
done |
|
419 |
||
63167 | 420 |
text\<open>NEVER USED\<close> |
14199 | 421 |
lemma CA_Says_4: |
61984 | 422 |
"[| Crypt (invKey SK) (Hash\<lbrace>Agent C, Nonce NC2, Nonce NCA\<rbrace>) |
14199 | 423 |
\<in> parts (knows Spy evs); |
424 |
cert (CA i) SK onlySig (priSK RCA) \<in> parts (knows Spy evs); |
|
425 |
evs \<in> set_cr; (CA i) \<notin> bad |] |
|
61984 | 426 |
==> \<exists>Y. Says (CA i) C \<lbrace>sign (priSK (CA i)) |
427 |
\<lbrace>Agent C, Nonce NC2, Nonce NCA\<rbrace>, Y\<rbrace> \<in> set evs" |
|
14199 | 428 |
by (blast dest!: certificate_valid_pubSK intro!: CA_Says_4_lemma) |
429 |
||
430 |
||
63167 | 431 |
text\<open>Message \<open>SET_CR6\<close>: C can check CA's signature if he has |
432 |
received CA's certificate.\<close> |
|
14199 | 433 |
lemma CA_Says_6_lemma: |
434 |
"[| Crypt (priSK (CA i)) |
|
61984 | 435 |
(Hash\<lbrace>Agent C, Nonce NC3, Agent (CA i), Nonce NonceCCA\<rbrace>) |
14199 | 436 |
\<in> parts (knows Spy evs); |
437 |
evs \<in> set_cr; (CA i) \<notin> bad |] |
|
61984 | 438 |
==> \<exists>Y K. Says (CA i) C (Crypt K \<lbrace>sign (priSK (CA i)) |
439 |
\<lbrace>Agent C, Nonce NC3, Agent (CA i), Nonce NonceCCA\<rbrace>, Y\<rbrace>) \<in> set evs" |
|
14199 | 440 |
apply (erule rev_mp) |
441 |
apply (erule set_cr.induct, auto) |
|
442 |
done |
|
443 |
||
63167 | 444 |
text\<open>NEVER USED\<close> |
14199 | 445 |
lemma CA_Says_6: |
61984 | 446 |
"[| Crypt (invKey SK) (Hash\<lbrace>Agent C, Nonce NC3, Agent (CA i), Nonce NonceCCA\<rbrace>) |
14199 | 447 |
\<in> parts (knows Spy evs); |
448 |
cert (CA i) SK onlySig (priSK RCA) \<in> parts (knows Spy evs); |
|
449 |
evs \<in> set_cr; (CA i) \<notin> bad |] |
|
61984 | 450 |
==> \<exists>Y K. Says (CA i) C (Crypt K \<lbrace>sign (priSK (CA i)) |
451 |
\<lbrace>Agent C, Nonce NC3, Agent (CA i), Nonce NonceCCA\<rbrace>, Y\<rbrace>) \<in> set evs" |
|
14199 | 452 |
by (blast dest!: certificate_valid_pubSK intro!: CA_Says_6_lemma) |
14218 | 453 |
(*>*) |
14199 | 454 |
|
455 |
||
63167 | 456 |
subsection\<open>Useful lemmas\<close> |
14199 | 457 |
|
63167 | 458 |
text\<open>Rewriting rule for private encryption keys. Analogous rewriting rules |
459 |
for other keys aren't needed.\<close> |
|
14199 | 460 |
|
461 |
lemma parts_image_priEK: |
|
67613 | 462 |
"[|Key (priEK C) \<in> parts (Key`KK \<union> (knows Spy evs)); |
14199 | 463 |
evs \<in> set_cr|] ==> priEK C \<in> KK | C \<in> bad" |
464 |
by auto |
|
465 |
||
63167 | 466 |
text\<open>trivial proof because (priEK C) never appears even in (parts evs)\<close> |
14199 | 467 |
lemma analz_image_priEK: |
468 |
"evs \<in> set_cr ==> |
|
67613 | 469 |
(Key (priEK C) \<in> analz (Key`KK \<union> (knows Spy evs))) = |
14199 | 470 |
(priEK C \<in> KK | C \<in> bad)" |
471 |
by (blast dest!: parts_image_priEK intro: analz_mono [THEN [2] rev_subsetD]) |
|
472 |
||
473 |
||
63167 | 474 |
subsection\<open>Secrecy of Session Keys\<close> |
14199 | 475 |
|
63167 | 476 |
subsubsection\<open>Lemmas about the predicate KeyCryptKey\<close> |
14199 | 477 |
|
63167 | 478 |
text\<open>A fresh DK cannot be associated with any other |
479 |
(with respect to a given trace).\<close> |
|
14199 | 480 |
lemma DK_fresh_not_KeyCryptKey: |
67613 | 481 |
"[| Key DK \<notin> used evs; evs \<in> set_cr |] ==> \<not> KeyCryptKey DK K evs" |
14199 | 482 |
apply (erule rev_mp) |
483 |
apply (erule set_cr.induct) |
|
484 |
apply (simp_all (no_asm_simp)) |
|
485 |
apply (blast dest: Crypt_analz_imp_used)+ |
|
486 |
done |
|
487 |
||
63167 | 488 |
text\<open>A fresh K cannot be associated with any other. The assumption that |
14199 | 489 |
DK isn't a private encryption key may be an artifact of the particular |
63167 | 490 |
definition of KeyCryptKey.\<close> |
14199 | 491 |
lemma K_fresh_not_KeyCryptKey: |
67613 | 492 |
"[|\<forall>C. DK \<noteq> priEK C; Key K \<notin> used evs|] ==> \<not> KeyCryptKey DK K evs" |
14199 | 493 |
apply (induct evs) |
63648 | 494 |
apply (auto simp add: parts_insert2 split: event.split) |
14199 | 495 |
done |
496 |
||
497 |
||
63167 | 498 |
text\<open>This holds because if (priEK (CA i)) appears in any traffic then it must |
69597 | 499 |
be known to the Spy, by \<^term>\<open>Spy_see_private_Key\<close>\<close> |
14199 | 500 |
lemma cardSK_neq_priEK: |
501 |
"[|Key cardSK \<notin> analz (knows Spy evs); |
|
67613 | 502 |
Key cardSK \<in> parts (knows Spy evs); |
14199 | 503 |
evs \<in> set_cr|] ==> cardSK \<noteq> priEK C" |
504 |
by blast |
|
505 |
||
506 |
lemma not_KeyCryptKey_cardSK [rule_format (no_asm)]: |
|
507 |
"[|cardSK \<notin> symKeys; \<forall>C. cardSK \<noteq> priEK C; evs \<in> set_cr|] ==> |
|
67613 | 508 |
Key cardSK \<notin> analz (knows Spy evs) \<longrightarrow> \<not> KeyCryptKey cardSK K evs" |
14199 | 509 |
by (erule set_cr.induct, analz_mono_contra, auto) |
510 |
||
63167 | 511 |
text\<open>Lemma for message 5: pubSK C is never used to encrypt Keys.\<close> |
67613 | 512 |
lemma pubSK_not_KeyCryptKey [simp]: "\<not> KeyCryptKey (pubSK C) K evs" |
14199 | 513 |
apply (induct_tac "evs") |
63648 | 514 |
apply (auto simp add: parts_insert2 split: event.split) |
14199 | 515 |
done |
516 |
||
63167 | 517 |
text\<open>Lemma for message 6: either cardSK is compromised (when we don't care) |
14199 | 518 |
or else cardSK hasn't been used to encrypt K. Previously we treated |
519 |
message 5 in the same way, but the current model assumes that rule |
|
63167 | 520 |
\<open>SET_CR5\<close> is executed only by honest agents.\<close> |
14199 | 521 |
lemma msg6_KeyCryptKey_disj: |
61984 | 522 |
"[|Gets B \<lbrace>Crypt KC3 \<lbrace>Agent C, Nonce N, Key KC2, Key cardSK, X\<rbrace>, Y\<rbrace> |
14199 | 523 |
\<in> set evs; |
524 |
cardSK \<notin> symKeys; evs \<in> set_cr|] |
|
525 |
==> Key cardSK \<in> analz (knows Spy evs) | |
|
67613 | 526 |
(\<forall>K. \<not> KeyCryptKey cardSK K evs)" |
14199 | 527 |
by (blast dest: not_KeyCryptKey_cardSK intro: cardSK_neq_priEK) |
528 |
||
63167 | 529 |
text\<open>As usual: we express the property as a logical equivalence\<close> |
14199 | 530 |
lemma Key_analz_image_Key_lemma: |
67613 | 531 |
"P \<longrightarrow> (Key K \<in> analz (Key`KK \<union> H)) \<longrightarrow> (K \<in> KK | Key K \<in> analz H) |
14199 | 532 |
==> |
67613 | 533 |
P \<longrightarrow> (Key K \<in> analz (Key`KK \<union> H)) = (K \<in> KK | Key K \<in> analz H)" |
14199 | 534 |
by (blast intro: analz_mono [THEN [2] rev_subsetD]) |
535 |
||
63167 | 536 |
method_setup valid_certificate_tac = \<open> |
51798 | 537 |
Args.goal_spec >> (fn quant => fn ctxt => SIMPLE_METHOD'' quant |
30549 | 538 |
(fn i => |
59499 | 539 |
EVERY [forward_tac ctxt @{thms Gets_certificate_valid} i, |
58963
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents:
58889
diff
changeset
|
540 |
assume_tac ctxt i, |
60754 | 541 |
eresolve_tac ctxt [conjE] i, REPEAT (hyp_subst_tac ctxt i)])) |
63167 | 542 |
\<close> |
14199 | 543 |
|
63167 | 544 |
text\<open>The \<open>(no_asm)\<close> attribute is essential, since it retains |
545 |
the quantifier and allows the simprule's condition to itself be simplified.\<close> |
|
14199 | 546 |
lemma symKey_compromise [rule_format (no_asm)]: |
547 |
"evs \<in> set_cr ==> |
|
67613 | 548 |
(\<forall>SK KK. SK \<in> symKeys \<longrightarrow> (\<forall>K \<in> KK. \<not> KeyCryptKey K SK evs) \<longrightarrow> |
549 |
(Key SK \<in> analz (Key`KK \<union> (knows Spy evs))) = |
|
14199 | 550 |
(SK \<in> KK | Key SK \<in> analz (knows Spy evs)))" |
551 |
apply (erule set_cr.induct) |
|
552 |
apply (rule_tac [!] allI) + |
|
553 |
apply (rule_tac [!] impI [THEN Key_analz_image_Key_lemma, THEN impI])+ |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
554 |
apply (valid_certificate_tac [8]) \<comment> \<open>for message 5\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
555 |
apply (valid_certificate_tac [6]) \<comment> \<open>for message 5\<close> |
14199 | 556 |
apply (erule_tac [9] msg6_KeyCryptKey_disj [THEN disjE]) |
557 |
apply (simp_all |
|
558 |
del: image_insert image_Un imp_disjL |
|
559 |
add: analz_image_keys_simps analz_knows_absorb |
|
560 |
analz_Key_image_insert_eq notin_image_iff |
|
561 |
K_fresh_not_KeyCryptKey |
|
562 |
DK_fresh_not_KeyCryptKey ball_conj_distrib |
|
563 |
analz_image_priEK disj_simps) |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
564 |
\<comment> \<open>9 seconds on a 1.6GHz machine\<close> |
14199 | 565 |
apply spy_analz |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
566 |
apply blast \<comment> \<open>3\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
567 |
apply blast \<comment> \<open>5\<close> |
14199 | 568 |
done |
569 |
||
63167 | 570 |
text\<open>The remaining quantifiers seem to be essential. |
14199 | 571 |
NO NEED to assume the cardholder's OK: bad cardholders don't do anything |
63167 | 572 |
wrong!!\<close> |
14199 | 573 |
lemma symKey_secrecy [rule_format]: |
574 |
"[|CA i \<notin> bad; K \<in> symKeys; evs \<in> set_cr|] |
|
67613 | 575 |
==> \<forall>X c. Says (Cardholder c) (CA i) X \<in> set evs \<longrightarrow> |
576 |
Key K \<in> parts{X} \<longrightarrow> |
|
577 |
Cardholder c \<notin> bad \<longrightarrow> |
|
14199 | 578 |
Key K \<notin> analz (knows Spy evs)" |
579 |
apply (erule set_cr.induct) |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
580 |
apply (frule_tac [8] Gets_certificate_valid) \<comment> \<open>for message 5\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
581 |
apply (frule_tac [6] Gets_certificate_valid) \<comment> \<open>for message 3\<close> |
14199 | 582 |
apply (erule_tac [11] msg6_KeyCryptKey_disj [THEN disjE]) |
583 |
apply (simp_all del: image_insert image_Un imp_disjL |
|
584 |
add: symKey_compromise fresh_notin_analz_knows_Spy |
|
585 |
analz_image_keys_simps analz_knows_absorb |
|
586 |
analz_Key_image_insert_eq notin_image_iff |
|
587 |
K_fresh_not_KeyCryptKey |
|
588 |
DK_fresh_not_KeyCryptKey |
|
589 |
analz_image_priEK) |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
590 |
\<comment> \<open>2.5 seconds on a 1.6GHz machine\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
591 |
apply spy_analz \<comment> \<open>Fake\<close> |
14199 | 592 |
apply (auto intro: analz_into_parts [THEN usedI] in_parts_Says_imp_used) |
593 |
done |
|
594 |
||
595 |
||
63167 | 596 |
subsection\<open>Primary Goals of Cardholder Registration\<close> |
14199 | 597 |
|
63167 | 598 |
text\<open>The cardholder's certificate really was created by the CA, provided the |
599 |
CA is uncompromised\<close> |
|
14199 | 600 |
|
63167 | 601 |
text\<open>Lemma concerning the actual signed message digest\<close> |
14199 | 602 |
lemma cert_valid_lemma: |
61984 | 603 |
"[|Crypt (priSK (CA i)) \<lbrace>Hash \<lbrace>Nonce N, Pan(pan C)\<rbrace>, Key cardSK, N1\<rbrace> |
14199 | 604 |
\<in> parts (knows Spy evs); |
605 |
CA i \<notin> bad; evs \<in> set_cr|] |
|
606 |
==> \<exists>KC2 X Y. Says (CA i) C |
|
607 |
(Crypt KC2 |
|
61984 | 608 |
\<lbrace>X, certC (pan C) cardSK N onlySig (priSK (CA i)), Y\<rbrace>) |
14199 | 609 |
\<in> set evs" |
610 |
apply (erule rev_mp) |
|
611 |
apply (erule set_cr.induct) |
|
612 |
apply (simp_all (no_asm_simp)) |
|
613 |
apply auto |
|
614 |
done |
|
615 |
||
63167 | 616 |
text\<open>Pre-packaged version for cardholder. We don't try to confirm the values |
617 |
of KC2, X and Y, since they are not important.\<close> |
|
14199 | 618 |
lemma certificate_valid_cardSK: |
61984 | 619 |
"[|Gets C (Crypt KC2 \<lbrace>X, certC (pan C) cardSK N onlySig (invKey SKi), |
620 |
cert (CA i) SKi onlySig (priSK RCA)\<rbrace>) \<in> set evs; |
|
14199 | 621 |
CA i \<notin> bad; evs \<in> set_cr|] |
622 |
==> \<exists>KC2 X Y. Says (CA i) C |
|
623 |
(Crypt KC2 |
|
61984 | 624 |
\<lbrace>X, certC (pan C) cardSK N onlySig (priSK (CA i)), Y\<rbrace>) |
14199 | 625 |
\<in> set evs" |
626 |
by (force dest!: Gets_imp_knows_Spy [THEN parts.Inj, THEN parts.Body] |
|
627 |
certificate_valid_pubSK cert_valid_lemma) |
|
628 |
||
629 |
||
630 |
lemma Hash_imp_parts [rule_format]: |
|
631 |
"evs \<in> set_cr |
|
67613 | 632 |
==> Hash\<lbrace>X, Nonce N\<rbrace> \<in> parts (knows Spy evs) \<longrightarrow> |
14199 | 633 |
Nonce N \<in> parts (knows Spy evs)" |
634 |
apply (erule set_cr.induct, force) |
|
635 |
apply (simp_all (no_asm_simp)) |
|
636 |
apply (blast intro: parts_mono [THEN [2] rev_subsetD]) |
|
637 |
done |
|
638 |
||
639 |
lemma Hash_imp_parts2 [rule_format]: |
|
640 |
"evs \<in> set_cr |
|
67613 | 641 |
==> Hash\<lbrace>X, Nonce M, Y, Nonce N\<rbrace> \<in> parts (knows Spy evs) \<longrightarrow> |
642 |
Nonce M \<in> parts (knows Spy evs) \<and> Nonce N \<in> parts (knows Spy evs)" |
|
14199 | 643 |
apply (erule set_cr.induct, force) |
644 |
apply (simp_all (no_asm_simp)) |
|
645 |
apply (blast intro: parts_mono [THEN [2] rev_subsetD]) |
|
646 |
done |
|
647 |
||
648 |
||
63167 | 649 |
subsection\<open>Secrecy of Nonces\<close> |
14199 | 650 |
|
63167 | 651 |
subsubsection\<open>Lemmas about the predicate KeyCryptNonce\<close> |
14199 | 652 |
|
63167 | 653 |
text\<open>A fresh DK cannot be associated with any other |
654 |
(with respect to a given trace).\<close> |
|
14199 | 655 |
lemma DK_fresh_not_KeyCryptNonce: |
656 |
"[| DK \<in> symKeys; Key DK \<notin> used evs; evs \<in> set_cr |] |
|
67613 | 657 |
==> \<not> KeyCryptNonce DK K evs" |
14199 | 658 |
apply (erule rev_mp) |
659 |
apply (erule rev_mp) |
|
660 |
apply (erule set_cr.induct) |
|
661 |
apply (simp_all (no_asm_simp)) |
|
662 |
apply blast |
|
663 |
apply blast |
|
664 |
apply (auto simp add: DK_fresh_not_KeyCryptKey) |
|
665 |
done |
|
666 |
||
63167 | 667 |
text\<open>A fresh N cannot be associated with any other |
668 |
(with respect to a given trace).\<close> |
|
14199 | 669 |
lemma N_fresh_not_KeyCryptNonce: |
67613 | 670 |
"\<forall>C. DK \<noteq> priEK C ==> Nonce N \<notin> used evs \<longrightarrow> \<not> KeyCryptNonce DK N evs" |
14199 | 671 |
apply (induct_tac "evs") |
55417
01fbfb60c33e
adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
blanchet
parents:
51798
diff
changeset
|
672 |
apply (rename_tac [2] a evs') |
14199 | 673 |
apply (case_tac [2] "a") |
674 |
apply (auto simp add: parts_insert2) |
|
675 |
done |
|
676 |
||
677 |
lemma not_KeyCryptNonce_cardSK [rule_format (no_asm)]: |
|
678 |
"[|cardSK \<notin> symKeys; \<forall>C. cardSK \<noteq> priEK C; evs \<in> set_cr|] ==> |
|
67613 | 679 |
Key cardSK \<notin> analz (knows Spy evs) \<longrightarrow> \<not> KeyCryptNonce cardSK N evs" |
14199 | 680 |
apply (erule set_cr.induct, analz_mono_contra, simp_all) |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
681 |
apply (blast dest: not_KeyCryptKey_cardSK) \<comment> \<open>6\<close> |
14199 | 682 |
done |
683 |
||
63167 | 684 |
subsubsection\<open>Lemmas for message 5 and 6: |
14199 | 685 |
either cardSK is compromised (when we don't care) |
63167 | 686 |
or else cardSK hasn't been used to encrypt K.\<close> |
14199 | 687 |
|
63167 | 688 |
text\<open>Lemma for message 5: pubSK C is never used to encrypt Nonces.\<close> |
67613 | 689 |
lemma pubSK_not_KeyCryptNonce [simp]: "\<not> KeyCryptNonce (pubSK C) N evs" |
14199 | 690 |
apply (induct_tac "evs") |
63648 | 691 |
apply (auto simp add: parts_insert2 split: event.split) |
14199 | 692 |
done |
693 |
||
63167 | 694 |
text\<open>Lemma for message 6: either cardSK is compromised (when we don't care) |
695 |
or else cardSK hasn't been used to encrypt K.\<close> |
|
14199 | 696 |
lemma msg6_KeyCryptNonce_disj: |
61984 | 697 |
"[|Gets B \<lbrace>Crypt KC3 \<lbrace>Agent C, Nonce N, Key KC2, Key cardSK, X\<rbrace>, Y\<rbrace> |
14199 | 698 |
\<in> set evs; |
699 |
cardSK \<notin> symKeys; evs \<in> set_cr|] |
|
700 |
==> Key cardSK \<in> analz (knows Spy evs) | |
|
67613 | 701 |
((\<forall>K. \<not> KeyCryptKey cardSK K evs) \<and> |
702 |
(\<forall>N. \<not> KeyCryptNonce cardSK N evs))" |
|
14199 | 703 |
by (blast dest: not_KeyCryptKey_cardSK not_KeyCryptNonce_cardSK |
704 |
intro: cardSK_neq_priEK) |
|
705 |
||
706 |
||
63167 | 707 |
text\<open>As usual: we express the property as a logical equivalence\<close> |
14199 | 708 |
lemma Nonce_analz_image_Key_lemma: |
67613 | 709 |
"P \<longrightarrow> (Nonce N \<in> analz (Key`KK \<union> H)) \<longrightarrow> (Nonce N \<in> analz H) |
710 |
==> P \<longrightarrow> (Nonce N \<in> analz (Key`KK \<union> H)) = (Nonce N \<in> analz H)" |
|
14199 | 711 |
by (blast intro: analz_mono [THEN [2] rev_subsetD]) |
712 |
||
32404 | 713 |
|
63167 | 714 |
text\<open>The \<open>(no_asm)\<close> attribute is essential, since it retains |
715 |
the quantifier and allows the simprule's condition to itself be simplified.\<close> |
|
14199 | 716 |
lemma Nonce_compromise [rule_format (no_asm)]: |
717 |
"evs \<in> set_cr ==> |
|
67613 | 718 |
(\<forall>N KK. (\<forall>K \<in> KK. \<not> KeyCryptNonce K N evs) \<longrightarrow> |
719 |
(Nonce N \<in> analz (Key`KK \<union> (knows Spy evs))) = |
|
14199 | 720 |
(Nonce N \<in> analz (knows Spy evs)))" |
721 |
apply (erule set_cr.induct) |
|
722 |
apply (rule_tac [!] allI)+ |
|
723 |
apply (rule_tac [!] impI [THEN Nonce_analz_image_Key_lemma])+ |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
724 |
apply (frule_tac [8] Gets_certificate_valid) \<comment> \<open>for message 5\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
725 |
apply (frule_tac [6] Gets_certificate_valid) \<comment> \<open>for message 3\<close> |
14199 | 726 |
apply (frule_tac [11] msg6_KeyCryptNonce_disj) |
727 |
apply (erule_tac [13] disjE) |
|
728 |
apply (simp_all del: image_insert image_Un |
|
729 |
add: symKey_compromise |
|
730 |
analz_image_keys_simps analz_knows_absorb |
|
731 |
analz_Key_image_insert_eq notin_image_iff |
|
732 |
N_fresh_not_KeyCryptNonce |
|
733 |
DK_fresh_not_KeyCryptNonce K_fresh_not_KeyCryptKey |
|
734 |
ball_conj_distrib analz_image_priEK) |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
735 |
\<comment> \<open>14 seconds on a 1.6GHz machine\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
736 |
apply spy_analz \<comment> \<open>Fake\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
737 |
apply blast \<comment> \<open>3\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
738 |
apply blast \<comment> \<open>5\<close> |
63167 | 739 |
txt\<open>Message 6\<close> |
32404 | 740 |
apply (metis symKey_compromise) |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
741 |
\<comment> \<open>cardSK compromised\<close> |
63167 | 742 |
txt\<open>Simplify again--necessary because the previous simplification introduces |
743 |
some logical connectives\<close> |
|
32404 | 744 |
apply (force simp del: image_insert image_Un imp_disjL |
14199 | 745 |
simp add: analz_image_keys_simps symKey_compromise) |
14218 | 746 |
done |
14199 | 747 |
|
748 |
||
63167 | 749 |
subsection\<open>Secrecy of CardSecret: the Cardholder's secret\<close> |
14199 | 750 |
|
751 |
lemma NC2_not_CardSecret: |
|
61984 | 752 |
"[|Crypt EKj \<lbrace>Key K, Pan p, Hash \<lbrace>Agent D, Nonce N\<rbrace>\<rbrace> |
14199 | 753 |
\<in> parts (knows Spy evs); |
754 |
Key K \<notin> analz (knows Spy evs); |
|
755 |
Nonce N \<notin> analz (knows Spy evs); |
|
756 |
evs \<in> set_cr|] |
|
61984 | 757 |
==> Crypt EKi \<lbrace>Key K', Pan p', Nonce N\<rbrace> \<notin> parts (knows Spy evs)" |
14199 | 758 |
apply (erule rev_mp) |
759 |
apply (erule rev_mp) |
|
760 |
apply (erule rev_mp) |
|
761 |
apply (erule set_cr.induct, analz_mono_contra, simp_all) |
|
762 |
apply (blast dest: Hash_imp_parts)+ |
|
763 |
done |
|
764 |
||
765 |
lemma KC2_secure_lemma [rule_format]: |
|
61984 | 766 |
"[|U = Crypt KC3 \<lbrace>Agent C, Nonce N, Key KC2, X\<rbrace>; |
14199 | 767 |
U \<in> parts (knows Spy evs); |
768 |
evs \<in> set_cr|] |
|
67613 | 769 |
==> Nonce N \<notin> analz (knows Spy evs) \<longrightarrow> |
770 |
(\<exists>k i W. Says (Cardholder k) (CA i) \<lbrace>U,W\<rbrace> \<in> set evs \<and> |
|
771 |
Cardholder k \<notin> bad \<and> CA i \<notin> bad)" |
|
59807 | 772 |
apply (erule_tac P = "U \<in> H" for H in rev_mp) |
14199 | 773 |
apply (erule set_cr.induct) |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
774 |
apply (valid_certificate_tac [8]) \<comment> \<open>for message 5\<close> |
14199 | 775 |
apply (simp_all del: image_insert image_Un imp_disjL |
776 |
add: analz_image_keys_simps analz_knows_absorb |
|
777 |
analz_knows_absorb2 notin_image_iff) |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
778 |
\<comment> \<open>4 seconds on a 1.6GHz machine\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
779 |
apply (simp_all (no_asm_simp)) \<comment> \<open>leaves 4 subgoals\<close> |
14199 | 780 |
apply (blast intro!: analz_insertI)+ |
781 |
done |
|
782 |
||
783 |
lemma KC2_secrecy: |
|
61984 | 784 |
"[|Gets B \<lbrace>Crypt K \<lbrace>Agent C, Nonce N, Key KC2, X\<rbrace>, Y\<rbrace> \<in> set evs; |
14199 | 785 |
Nonce N \<notin> analz (knows Spy evs); KC2 \<in> symKeys; |
786 |
evs \<in> set_cr|] |
|
787 |
==> Key KC2 \<notin> analz (knows Spy evs)" |
|
788 |
by (force dest!: refl [THEN KC2_secure_lemma] symKey_secrecy) |
|
789 |
||
790 |
||
63167 | 791 |
text\<open>Inductive version\<close> |
14199 | 792 |
lemma CardSecret_secrecy_lemma [rule_format]: |
793 |
"[|CA i \<notin> bad; evs \<in> set_cr|] |
|
67613 | 794 |
==> Key K \<notin> analz (knows Spy evs) \<longrightarrow> |
61984 | 795 |
Crypt (pubEK (CA i)) \<lbrace>Key K, Pan p, Nonce CardSecret\<rbrace> |
67613 | 796 |
\<in> parts (knows Spy evs) \<longrightarrow> |
14199 | 797 |
Nonce CardSecret \<notin> analz (knows Spy evs)" |
798 |
apply (erule set_cr.induct, analz_mono_contra) |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
799 |
apply (valid_certificate_tac [8]) \<comment> \<open>for message 5\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
800 |
apply (valid_certificate_tac [6]) \<comment> \<open>for message 5\<close> |
14199 | 801 |
apply (frule_tac [9] msg6_KeyCryptNonce_disj [THEN disjE]) |
802 |
apply (simp_all |
|
803 |
del: image_insert image_Un imp_disjL |
|
804 |
add: analz_image_keys_simps analz_knows_absorb |
|
805 |
analz_Key_image_insert_eq notin_image_iff |
|
806 |
EXHcrypt_def Crypt_notin_image_Key |
|
807 |
N_fresh_not_KeyCryptNonce DK_fresh_not_KeyCryptNonce |
|
808 |
ball_conj_distrib Nonce_compromise symKey_compromise |
|
809 |
analz_image_priEK) |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
810 |
\<comment> \<open>2.5 seconds on a 1.6GHz machine\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
811 |
apply spy_analz \<comment> \<open>Fake\<close> |
14199 | 812 |
apply (simp_all (no_asm_simp)) |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
813 |
apply blast \<comment> \<open>1\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
814 |
apply (blast dest!: Gets_imp_knows_Spy [THEN analz.Inj]) \<comment> \<open>2\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
815 |
apply blast \<comment> \<open>3\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
816 |
apply (blast dest: NC2_not_CardSecret Gets_imp_knows_Spy [THEN analz.Inj] analz_symKeys_Decrypt) \<comment> \<open>4\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
817 |
apply blast \<comment> \<open>5\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
818 |
apply (blast dest: KC2_secrecy)+ \<comment> \<open>Message 6: two cases\<close> |
14199 | 819 |
done |
820 |
||
821 |
||
63167 | 822 |
text\<open>Packaged version for cardholder\<close> |
14199 | 823 |
lemma CardSecret_secrecy: |
824 |
"[|Cardholder k \<notin> bad; CA i \<notin> bad; |
|
825 |
Says (Cardholder k) (CA i) |
|
61984 | 826 |
\<lbrace>X, Crypt EKi \<lbrace>Key KC3, Pan p, Nonce CardSecret\<rbrace>\<rbrace> \<in> set evs; |
827 |
Gets A \<lbrace>Z, cert (CA i) EKi onlyEnc (priSK RCA), |
|
828 |
cert (CA i) SKi onlySig (priSK RCA)\<rbrace> \<in> set evs; |
|
14199 | 829 |
KC3 \<in> symKeys; evs \<in> set_cr|] |
830 |
==> Nonce CardSecret \<notin> analz (knows Spy evs)" |
|
831 |
apply (frule Gets_certificate_valid, assumption) |
|
832 |
apply (subgoal_tac "Key KC3 \<notin> analz (knows Spy evs) ") |
|
833 |
apply (blast dest: CardSecret_secrecy_lemma) |
|
834 |
apply (rule symKey_secrecy) |
|
835 |
apply (auto simp add: parts_insert2) |
|
836 |
done |
|
837 |
||
838 |
||
63167 | 839 |
subsection\<open>Secrecy of NonceCCA [the CA's secret]\<close> |
14199 | 840 |
|
841 |
lemma NC2_not_NonceCCA: |
|
61984 | 842 |
"[|Hash \<lbrace>Agent C', Nonce N', Agent C, Nonce N\<rbrace> |
14199 | 843 |
\<in> parts (knows Spy evs); |
844 |
Nonce N \<notin> analz (knows Spy evs); |
|
845 |
evs \<in> set_cr|] |
|
61984 | 846 |
==> Crypt KC1 \<lbrace>\<lbrace>Agent B, Nonce N\<rbrace>, Hash p\<rbrace> \<notin> parts (knows Spy evs)" |
14199 | 847 |
apply (erule rev_mp) |
848 |
apply (erule rev_mp) |
|
849 |
apply (erule set_cr.induct, analz_mono_contra, simp_all) |
|
850 |
apply (blast dest: Hash_imp_parts2)+ |
|
851 |
done |
|
852 |
||
853 |
||
63167 | 854 |
text\<open>Inductive version\<close> |
14199 | 855 |
lemma NonceCCA_secrecy_lemma [rule_format]: |
856 |
"[|CA i \<notin> bad; evs \<in> set_cr|] |
|
67613 | 857 |
==> Key K \<notin> analz (knows Spy evs) \<longrightarrow> |
14199 | 858 |
Crypt K |
61984 | 859 |
\<lbrace>sign (priSK (CA i)) |
860 |
\<lbrace>Agent C, Nonce N, Agent(CA i), Nonce NonceCCA\<rbrace>, |
|
861 |
X, Y\<rbrace> |
|
67613 | 862 |
\<in> parts (knows Spy evs) \<longrightarrow> |
14199 | 863 |
Nonce NonceCCA \<notin> analz (knows Spy evs)" |
864 |
apply (erule set_cr.induct, analz_mono_contra) |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
865 |
apply (valid_certificate_tac [8]) \<comment> \<open>for message 5\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
866 |
apply (valid_certificate_tac [6]) \<comment> \<open>for message 5\<close> |
14199 | 867 |
apply (frule_tac [9] msg6_KeyCryptNonce_disj [THEN disjE]) |
868 |
apply (simp_all |
|
869 |
del: image_insert image_Un imp_disjL |
|
870 |
add: analz_image_keys_simps analz_knows_absorb sign_def |
|
871 |
analz_Key_image_insert_eq notin_image_iff |
|
872 |
EXHcrypt_def Crypt_notin_image_Key |
|
873 |
N_fresh_not_KeyCryptNonce DK_fresh_not_KeyCryptNonce |
|
874 |
ball_conj_distrib Nonce_compromise symKey_compromise |
|
875 |
analz_image_priEK) |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
876 |
\<comment> \<open>3 seconds on a 1.6GHz machine\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
877 |
apply spy_analz \<comment> \<open>Fake\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
878 |
apply blast \<comment> \<open>1\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
879 |
apply (blast dest!: Gets_imp_knows_Spy [THEN analz.Inj]) \<comment> \<open>2\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
880 |
apply blast \<comment> \<open>3\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
881 |
apply (blast dest: NC2_not_NonceCCA) \<comment> \<open>4\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
882 |
apply blast \<comment> \<open>5\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
883 |
apply (blast dest: KC2_secrecy)+ \<comment> \<open>Message 6: two cases\<close> |
14199 | 884 |
done |
885 |
||
886 |
||
63167 | 887 |
text\<open>Packaged version for cardholder\<close> |
14199 | 888 |
lemma NonceCCA_secrecy: |
889 |
"[|Cardholder k \<notin> bad; CA i \<notin> bad; |
|
890 |
Gets (Cardholder k) |
|
891 |
(Crypt KC2 |
|
61984 | 892 |
\<lbrace>sign (priSK (CA i)) \<lbrace>Agent C, Nonce N, Agent(CA i), Nonce NonceCCA\<rbrace>, |
893 |
X, Y\<rbrace>) \<in> set evs; |
|
14199 | 894 |
Says (Cardholder k) (CA i) |
61984 | 895 |
\<lbrace>Crypt KC3 \<lbrace>Agent C, Nonce NC3, Key KC2, X'\<rbrace>, Y'\<rbrace> \<in> set evs; |
896 |
Gets A \<lbrace>Z, cert (CA i) EKi onlyEnc (priSK RCA), |
|
897 |
cert (CA i) SKi onlySig (priSK RCA)\<rbrace> \<in> set evs; |
|
14199 | 898 |
KC2 \<in> symKeys; evs \<in> set_cr|] |
899 |
==> Nonce NonceCCA \<notin> analz (knows Spy evs)" |
|
900 |
apply (frule Gets_certificate_valid, assumption) |
|
901 |
apply (subgoal_tac "Key KC2 \<notin> analz (knows Spy evs) ") |
|
902 |
apply (blast dest: NonceCCA_secrecy_lemma) |
|
903 |
apply (rule symKey_secrecy) |
|
904 |
apply (auto simp add: parts_insert2) |
|
905 |
done |
|
906 |
||
63167 | 907 |
text\<open>We don't bother to prove guarantees for the CA. He doesn't care about |
908 |
the PANSecret: it isn't his credit card!\<close> |
|
14199 | 909 |
|
910 |
||
63167 | 911 |
subsection\<open>Rewriting Rule for PANs\<close> |
14199 | 912 |
|
63167 | 913 |
text\<open>Lemma for message 6: either cardSK isn't a CA's private encryption key, |
14199 | 914 |
or if it is then (because it appears in traffic) that CA is bad, |
915 |
and so the Spy knows that key already. Either way, we can simplify |
|
69597 | 916 |
the expression \<^term>\<open>analz (insert (Key cardSK) X)\<close>.\<close> |
14199 | 917 |
lemma msg6_cardSK_disj: |
61984 | 918 |
"[|Gets A \<lbrace>Crypt K \<lbrace>c, n, k', Key cardSK, X\<rbrace>, Y\<rbrace> |
14199 | 919 |
\<in> set evs; evs \<in> set_cr |] |
920 |
==> cardSK \<notin> range(invKey o pubEK o CA) | Key cardSK \<in> knows Spy evs" |
|
921 |
by auto |
|
922 |
||
923 |
lemma analz_image_pan_lemma: |
|
67613 | 924 |
"(Pan P \<in> analz (Key`nE \<union> H)) \<longrightarrow> (Pan P \<in> analz H) ==> |
925 |
(Pan P \<in> analz (Key`nE \<union> H)) = (Pan P \<in> analz H)" |
|
14199 | 926 |
by (blast intro: analz_mono [THEN [2] rev_subsetD]) |
927 |
||
928 |
lemma analz_image_pan [rule_format]: |
|
929 |
"evs \<in> set_cr ==> |
|
67613 | 930 |
\<forall>KK. KK \<subseteq> - invKey ` pubEK ` range CA \<longrightarrow> |
931 |
(Pan P \<in> analz (Key`KK \<union> (knows Spy evs))) = |
|
14199 | 932 |
(Pan P \<in> analz (knows Spy evs))" |
933 |
apply (erule set_cr.induct) |
|
934 |
apply (rule_tac [!] allI impI)+ |
|
935 |
apply (rule_tac [!] analz_image_pan_lemma) |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
936 |
apply (valid_certificate_tac [8]) \<comment> \<open>for message 5\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
937 |
apply (valid_certificate_tac [6]) \<comment> \<open>for message 5\<close> |
14199 | 938 |
apply (erule_tac [9] msg6_cardSK_disj [THEN disjE]) |
939 |
apply (simp_all |
|
940 |
del: image_insert image_Un |
|
941 |
add: analz_image_keys_simps disjoint_image_iff |
|
942 |
notin_image_iff analz_image_priEK) |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
943 |
\<comment> \<open>6 seconds on a 1.6GHz machine\<close> |
14199 | 944 |
apply spy_analz |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
945 |
apply (simp add: insert_absorb) \<comment> \<open>6\<close> |
14199 | 946 |
done |
947 |
||
948 |
lemma analz_insert_pan: |
|
949 |
"[| evs \<in> set_cr; K \<notin> invKey ` pubEK ` range CA |] ==> |
|
950 |
(Pan P \<in> analz (insert (Key K) (knows Spy evs))) = |
|
951 |
(Pan P \<in> analz (knows Spy evs))" |
|
952 |
by (simp del: image_insert image_Un |
|
953 |
add: analz_image_keys_simps analz_image_pan) |
|
954 |
||
955 |
||
63167 | 956 |
text\<open>Confidentiality of the PAN\@. Maybe we could combine the statements of |
69597 | 957 |
this theorem with \<^term>\<open>analz_image_pan\<close>, requiring a single induction but |
63167 | 958 |
a much more difficult proof.\<close> |
14199 | 959 |
lemma pan_confidentiality: |
67613 | 960 |
"[| Pan (pan C) \<in> analz(knows Spy evs); C \<noteq>Spy; evs \<in> set_cr|] |
14199 | 961 |
==> \<exists>i X K HN. |
61984 | 962 |
Says C (CA i) \<lbrace>X, Crypt (pubEK (CA i)) \<lbrace>Key K, Pan (pan C), HN\<rbrace> \<rbrace> |
14199 | 963 |
\<in> set evs |
67613 | 964 |
\<and> (CA i) \<in> bad" |
14199 | 965 |
apply (erule rev_mp) |
966 |
apply (erule set_cr.induct) |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
967 |
apply (valid_certificate_tac [8]) \<comment> \<open>for message 5\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
968 |
apply (valid_certificate_tac [6]) \<comment> \<open>for message 5\<close> |
14199 | 969 |
apply (erule_tac [9] msg6_cardSK_disj [THEN disjE]) |
970 |
apply (simp_all |
|
971 |
del: image_insert image_Un |
|
972 |
add: analz_image_keys_simps analz_insert_pan analz_image_pan |
|
973 |
notin_image_iff analz_image_priEK) |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
974 |
\<comment> \<open>3.5 seconds on a 1.6GHz machine\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
975 |
apply spy_analz \<comment> \<open>fake\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
976 |
apply blast \<comment> \<open>3\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
977 |
apply blast \<comment> \<open>5\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63648
diff
changeset
|
978 |
apply (simp (no_asm_simp) add: insert_absorb) \<comment> \<open>6\<close> |
14199 | 979 |
done |
980 |
||
981 |
||
63167 | 982 |
subsection\<open>Unicity\<close> |
14199 | 983 |
|
984 |
lemma CR6_Says_imp_Notes: |
|
985 |
"[|Says (CA i) C (Crypt KC2 |
|
61984 | 986 |
\<lbrace>sign (priSK (CA i)) \<lbrace>Agent C, Nonce NC3, Agent (CA i), Nonce Y\<rbrace>, |
14199 | 987 |
certC (pan C) cardSK X onlySig (priSK (CA i)), |
61984 | 988 |
cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)\<rbrace>) \<in> set evs; |
14199 | 989 |
evs \<in> set_cr |] |
990 |
==> Notes (CA i) (Key cardSK) \<in> set evs" |
|
991 |
apply (erule rev_mp) |
|
992 |
apply (erule set_cr.induct) |
|
993 |
apply (simp_all (no_asm_simp)) |
|
994 |
done |
|
995 |
||
63167 | 996 |
text\<open>Unicity of cardSK: it uniquely identifies the other components. |
997 |
This holds because a CA accepts a cardSK at most once.\<close> |
|
14199 | 998 |
lemma cardholder_key_unicity: |
999 |
"[|Says (CA i) C (Crypt KC2 |
|
61984 | 1000 |
\<lbrace>sign (priSK (CA i)) \<lbrace>Agent C, Nonce NC3, Agent (CA i), Nonce Y\<rbrace>, |
14199 | 1001 |
certC (pan C) cardSK X onlySig (priSK (CA i)), |
61984 | 1002 |
cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)\<rbrace>) |
14199 | 1003 |
\<in> set evs; |
1004 |
Says (CA i) C' (Crypt KC2' |
|
61984 | 1005 |
\<lbrace>sign (priSK (CA i)) \<lbrace>Agent C', Nonce NC3', Agent (CA i), Nonce Y'\<rbrace>, |
14199 | 1006 |
certC (pan C') cardSK X' onlySig (priSK (CA i)), |
61984 | 1007 |
cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)\<rbrace>) |
14199 | 1008 |
\<in> set evs; |
67613 | 1009 |
evs \<in> set_cr |] ==> C=C' \<and> NC3=NC3' \<and> X=X' \<and> KC2=KC2' \<and> Y=Y'" |
14199 | 1010 |
apply (erule rev_mp) |
1011 |
apply (erule rev_mp) |
|
1012 |
apply (erule set_cr.induct) |
|
1013 |
apply (simp_all (no_asm_simp)) |
|
1014 |
apply (blast dest!: CR6_Says_imp_Notes) |
|
1015 |
done |
|
1016 |
||
1017 |
||
14218 | 1018 |
(*<*) |
63167 | 1019 |
text\<open>UNUSED unicity result\<close> |
14199 | 1020 |
lemma unique_KC1: |
61984 | 1021 |
"[|Says C B \<lbrace>Crypt KC1 X, Crypt EK \<lbrace>Key KC1, Y\<rbrace>\<rbrace> |
14199 | 1022 |
\<in> set evs; |
61984 | 1023 |
Says C B' \<lbrace>Crypt KC1 X', Crypt EK' \<lbrace>Key KC1, Y'\<rbrace>\<rbrace> |
14199 | 1024 |
\<in> set evs; |
67613 | 1025 |
C \<notin> bad; evs \<in> set_cr|] ==> B'=B \<and> Y'=Y" |
14199 | 1026 |
apply (erule rev_mp) |
1027 |
apply (erule rev_mp) |
|
1028 |
apply (erule set_cr.induct, auto) |
|
1029 |
done |
|
1030 |
||
63167 | 1031 |
text\<open>UNUSED unicity result\<close> |
14199 | 1032 |
lemma unique_KC2: |
61984 | 1033 |
"[|Says C B \<lbrace>Crypt K \<lbrace>Agent C, nn, Key KC2, X\<rbrace>, Y\<rbrace> \<in> set evs; |
1034 |
Says C B' \<lbrace>Crypt K' \<lbrace>Agent C, nn', Key KC2, X'\<rbrace>, Y'\<rbrace> \<in> set evs; |
|
67613 | 1035 |
C \<notin> bad; evs \<in> set_cr|] ==> B'=B \<and> X'=X" |
14199 | 1036 |
apply (erule rev_mp) |
1037 |
apply (erule rev_mp) |
|
1038 |
apply (erule set_cr.induct, auto) |
|
1039 |
done |
|
14218 | 1040 |
(*>*) |
1041 |
||
14199 | 1042 |
|
63167 | 1043 |
text\<open>Cannot show cardSK to be secret because it isn't assumed to be fresh |
1044 |
it could be a previously compromised cardSK [e.g. involving a bad CA]\<close> |
|
14199 | 1045 |
|
1046 |
||
1047 |
end |