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