author | paulson |
Mon, 29 Sep 1997 11:46:33 +0200 | |
changeset 3729 | 6be7cf5086ab |
parent 3711 | 2f86b403d975 |
child 3745 | 4c5d3b1ddc75 |
permissions | -rw-r--r-- |
3474 | 1 |
(* Title: HOL/Auth/TLS |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1997 University of Cambridge |
|
5 |
||
3480
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
6 |
Protocol goals: |
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
7 |
* M, serverK(NA,NB,M) and clientK(NA,NB,M) will be known only to the two |
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
8 |
parties (though A is not necessarily authenticated). |
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
9 |
|
3685
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
10 |
* B upon receiving CertVerify knows that A is present (But this |
3480
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
11 |
message is optional!) |
3474 | 12 |
|
3685
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
13 |
* A upon receiving ServerFinished knows that B is present |
3480
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
14 |
|
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
15 |
* Each party who has received a FINISHED message can trust that the other |
3729
6be7cf5086ab
Renamed XA, XB to PA, PB and removed the certificate from Client Verify
paulson
parents:
3711
diff
changeset
|
16 |
party agrees on all message components, including PA and PB (thus foiling |
3480
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
17 |
rollback attacks). |
3474 | 18 |
*) |
19 |
||
3704 | 20 |
|
3474 | 21 |
open TLS; |
22 |
||
23 |
proof_timing:=true; |
|
24 |
HOL_quantifiers := false; |
|
25 |
||
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
26 |
(** We mostly DO NOT unfold the definition of "certificate". The attached |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
27 |
lemmas unfold it lazily, when "certificate B KB" occurs in appropriate |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
28 |
contexts. |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
29 |
**) |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
30 |
|
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
31 |
goalw thy [certificate_def] |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
32 |
"parts (insert (certificate B KB) H) = \ |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
33 |
\ parts (insert (Crypt (priK Server) {|Agent B, Key KB|}) H)"; |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
34 |
by (rtac refl 1); |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
35 |
qed "parts_insert_certificate"; |
3474 | 36 |
|
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
37 |
goalw thy [certificate_def] |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
38 |
"analz (insert (certificate B KB) H) = \ |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
39 |
\ analz (insert (Crypt (priK Server) {|Agent B, Key KB|}) H)"; |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
40 |
by (rtac refl 1); |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
41 |
qed "analz_insert_certificate"; |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
42 |
Addsimps [parts_insert_certificate, analz_insert_certificate]; |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
43 |
|
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
44 |
goalw thy [certificate_def] |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
45 |
"(X = certificate B KB) = (Crypt (priK Server) {|Agent B, Key KB|} = X)"; |
3480
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
46 |
by (Blast_tac 1); |
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
47 |
qed "eq_certificate_iff"; |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
48 |
AddIffs [eq_certificate_iff]; |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
49 |
|
3480
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
50 |
|
3474 | 51 |
(*Injectiveness of key-generating functions*) |
3677
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
52 |
AddIffs [inj_PRF RS inj_eq, inj_sessionK RS inj_eq]; |
3474 | 53 |
|
3677
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
54 |
(* invKey(sessionK x) = sessionK x*) |
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
55 |
Addsimps [isSym_sessionK, rewrite_rule [isSymKey_def] isSym_sessionK]; |
3480
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
56 |
|
3474 | 57 |
|
58 |
(*** clientK and serverK make symmetric keys; no clashes with pubK or priK ***) |
|
59 |
||
3677
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
60 |
goal thy "pubK A ~= sessionK arg"; |
3474 | 61 |
br notI 1; |
62 |
by (dres_inst_tac [("f","isSymKey")] arg_cong 1); |
|
63 |
by (Full_simp_tac 1); |
|
3677
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
64 |
qed "pubK_neq_sessionK"; |
3474 | 65 |
|
3677
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
66 |
goal thy "priK A ~= sessionK arg"; |
3474 | 67 |
br notI 1; |
68 |
by (dres_inst_tac [("f","isSymKey")] arg_cong 1); |
|
69 |
by (Full_simp_tac 1); |
|
3677
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
70 |
qed "priK_neq_sessionK"; |
3474 | 71 |
|
3677
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
72 |
val keys_distinct = [pubK_neq_sessionK, priK_neq_sessionK]; |
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
73 |
AddIffs (keys_distinct @ (keys_distinct RL [not_sym])); |
3474 | 74 |
|
75 |
||
76 |
(**** Protocol Proofs ****) |
|
77 |
||
78 |
(*A "possibility property": there are traces that reach the end. |
|
79 |
This protocol has three end points and six messages to consider.*) |
|
80 |
||
3672
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
81 |
|
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
82 |
(** These proofs make the further assumption that the Nonce_supply nonces |
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
83 |
(which have the form @ N. Nonce N ~: used evs) |
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
84 |
lie outside the range of PRF. This assumption seems reasonable, but |
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
85 |
as it is needed only for the possibility theorems, it is not taken |
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
86 |
as an axiom. |
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
87 |
**) |
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
88 |
|
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
89 |
|
3685
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
90 |
(*Possibility property ending with ClientAccepts.*) |
3474 | 91 |
goal thy |
3672
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
92 |
"!!A B. [| ALL evs. (@ N. Nonce N ~: used evs) ~: range PRF; \ |
3685
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
93 |
\ A ~= B |] ==> EX SID M. EX evs: tls. \ |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
94 |
\ Notes A {|Number SID, Agent A, Agent B, Nonce M|} : set evs"; |
3474 | 95 |
by (REPEAT (resolve_tac [exI,bexI] 1)); |
96 |
by (rtac (tls.Nil RS tls.ClientHello RS tls.ServerHello RS tls.ClientCertKeyEx |
|
3685
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
97 |
RS tls.ClientFinished RS tls.ServerFinished RS tls.ClientAccepts) 2); |
3474 | 98 |
by possibility_tac; |
3672
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
99 |
by (REPEAT (Blast_tac 1)); |
3474 | 100 |
result(); |
101 |
||
3685
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
102 |
(*And one for ServerAccepts. Either FINISHED message may come first.*) |
3474 | 103 |
goal thy |
3672
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
104 |
"!!A B. [| ALL evs. (@ N. Nonce N ~: used evs) ~: range PRF; \ |
3729
6be7cf5086ab
Renamed XA, XB to PA, PB and removed the certificate from Client Verify
paulson
parents:
3711
diff
changeset
|
105 |
\ A ~= B |] ==> EX SID NA PA NB PB M. EX evs: tls. \ |
3685
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
106 |
\ Notes B {|Number SID, Agent A, Agent B, Nonce M|} : set evs"; |
3474 | 107 |
by (REPEAT (resolve_tac [exI,bexI] 1)); |
108 |
by (rtac (tls.Nil RS tls.ClientHello RS tls.ServerHello RS tls.ClientCertKeyEx |
|
3685
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
109 |
RS tls.ServerFinished RS tls.ClientFinished RS tls.ServerAccepts) 2); |
3474 | 110 |
by possibility_tac; |
3672
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
111 |
by (REPEAT (Blast_tac 1)); |
3474 | 112 |
result(); |
113 |
||
114 |
(*Another one, for CertVerify (which is optional)*) |
|
115 |
goal thy |
|
3672
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
116 |
"!!A B. [| ALL evs. (@ N. Nonce N ~: used evs) ~: range PRF; \ |
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
117 |
\ A ~= B |] ==> EX NB PMS. EX evs: tls. \ |
3729
6be7cf5086ab
Renamed XA, XB to PA, PB and removed the certificate from Client Verify
paulson
parents:
3711
diff
changeset
|
118 |
\ Says A B (Crypt (priK A) (Hash{|Nonce NB, Agent B, Nonce PMS|})) : set evs"; |
3474 | 119 |
by (REPEAT (resolve_tac [exI,bexI] 1)); |
3506 | 120 |
by (rtac (tls.Nil RS tls.ClientHello RS tls.ServerHello RS tls.ClientCertKeyEx |
121 |
RS tls.CertVerify) 2); |
|
3474 | 122 |
by possibility_tac; |
3672
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
123 |
by (REPEAT (Blast_tac 1)); |
3474 | 124 |
result(); |
125 |
||
3685
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
126 |
(*Another one, for session resumption (both ServerResume and ClientResume) *) |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
127 |
goal thy |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
128 |
"!!A B. [| evs0 : tls; \ |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
129 |
\ Notes A {|Number SID, Agent A, Agent B, Nonce M|} : set evs0; \ |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
130 |
\ Notes B {|Number SID, Agent A, Agent B, Nonce M|} : set evs0; \ |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
131 |
\ ALL evs. (@ N. Nonce N ~: used evs) ~: range PRF; \ |
3729
6be7cf5086ab
Renamed XA, XB to PA, PB and removed the certificate from Client Verify
paulson
parents:
3711
diff
changeset
|
132 |
\ A ~= B |] ==> EX NA PA NB PB. EX evs: tls. \ |
3685
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
133 |
\ Says A B (Crypt (clientK(NA,NB,M)) \ |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
134 |
\ (Hash{|Nonce M, Number SID, \ |
3729
6be7cf5086ab
Renamed XA, XB to PA, PB and removed the certificate from Client Verify
paulson
parents:
3711
diff
changeset
|
135 |
\ Nonce NA, Number PA, Agent A, \ |
6be7cf5086ab
Renamed XA, XB to PA, PB and removed the certificate from Client Verify
paulson
parents:
3711
diff
changeset
|
136 |
\ Nonce NB, Number PB, Agent B|})) : set evs"; |
3685
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
137 |
by (REPEAT (resolve_tac [exI,bexI] 1)); |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
138 |
by (etac (tls.ClientHello RS tls.ServerResume RS tls.ClientResume) 2); |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
139 |
by possibility_tac; |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
140 |
by (REPEAT (Blast_tac 1)); |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
141 |
result(); |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
142 |
|
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
143 |
|
3474 | 144 |
|
145 |
(**** Inductive proofs about tls ****) |
|
146 |
||
147 |
(*Nobody sends themselves messages*) |
|
148 |
goal thy "!!evs. evs : tls ==> ALL A X. Says A A X ~: set evs"; |
|
149 |
by (etac tls.induct 1); |
|
150 |
by (Auto_tac()); |
|
151 |
qed_spec_mp "not_Says_to_self"; |
|
152 |
Addsimps [not_Says_to_self]; |
|
153 |
AddSEs [not_Says_to_self RSN (2, rev_notE)]; |
|
154 |
||
155 |
||
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
156 |
(*Induction for regularity theorems. If induction formula has the form |
3683 | 157 |
X ~: analz (spies evs) --> ... then it shortens the proof by discarding |
158 |
needless information about analz (insert X (spies evs)) *) |
|
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
159 |
fun parts_induct_tac i = |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
160 |
etac tls.induct i |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
161 |
THEN |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
162 |
REPEAT (FIRSTGOAL analz_mono_contra_tac) |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
163 |
THEN |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
164 |
fast_tac (!claset addss (!simpset)) i THEN |
3704 | 165 |
ALLGOALS (asm_simp_tac (!simpset setloop split_tac [expand_if])); |
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
166 |
|
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
167 |
|
3683 | 168 |
(** Theorems of the form X ~: parts (spies evs) imply that NOBODY |
3474 | 169 |
sends messages containing X! **) |
170 |
||
3683 | 171 |
(*Spy never sees another agent's private key! (unless it's bad at start)*) |
3474 | 172 |
goal thy |
3683 | 173 |
"!!evs. evs : tls ==> (Key (priK A) : parts (spies evs)) = (A : bad)"; |
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
174 |
by (parts_induct_tac 1); |
3474 | 175 |
by (Fake_parts_insert_tac 1); |
176 |
qed "Spy_see_priK"; |
|
177 |
Addsimps [Spy_see_priK]; |
|
178 |
||
179 |
goal thy |
|
3683 | 180 |
"!!evs. evs : tls ==> (Key (priK A) : analz (spies evs)) = (A : bad)"; |
3474 | 181 |
by (auto_tac(!claset addDs [impOfSubs analz_subset_parts], !simpset)); |
182 |
qed "Spy_analz_priK"; |
|
183 |
Addsimps [Spy_analz_priK]; |
|
184 |
||
3683 | 185 |
goal thy "!!A. [| Key (priK A) : parts (spies evs); \ |
186 |
\ evs : tls |] ==> A:bad"; |
|
3474 | 187 |
by (blast_tac (!claset addDs [Spy_see_priK]) 1); |
188 |
qed "Spy_see_priK_D"; |
|
189 |
||
190 |
bind_thm ("Spy_analz_priK_D", analz_subset_parts RS subsetD RS Spy_see_priK_D); |
|
191 |
AddSDs [Spy_see_priK_D, Spy_analz_priK_D]; |
|
192 |
||
193 |
||
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
194 |
(*This lemma says that no false certificates exist. One might extend the |
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
195 |
model to include bogus certificates for the agents, but there seems |
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
196 |
little point in doing so: the loss of their private keys is a worse |
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
197 |
breach of security.*) |
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
198 |
goalw thy [certificate_def] |
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
199 |
"!!evs. evs : tls \ |
3683 | 200 |
\ ==> certificate B KB : parts (spies evs) --> KB = pubK B"; |
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
201 |
by (parts_induct_tac 1); |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
202 |
by (Fake_parts_insert_tac 1); |
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
203 |
bind_thm ("Server_cert_pubB", result() RSN (2, rev_mp)); |
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
204 |
|
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
205 |
|
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
206 |
(*Replace key KB in ClientCertKeyEx by (pubK B) *) |
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
207 |
val ClientCertKeyEx_tac = |
3683 | 208 |
forward_tac [Says_imp_spies RS parts.Inj RS |
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
209 |
parts.Snd RS parts.Snd RS parts.Snd RS Server_cert_pubB] |
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
210 |
THEN' assume_tac |
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
211 |
THEN' hyp_subst_tac; |
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
212 |
|
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
213 |
fun analz_induct_tac i = |
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
214 |
etac tls.induct i THEN |
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
215 |
ClientCertKeyEx_tac (i+7) THEN (*ClientFinished*) |
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
216 |
ClientCertKeyEx_tac (i+6) THEN (*CertVerify*) |
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
217 |
ClientCertKeyEx_tac (i+5) THEN (*ClientCertKeyEx*) |
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
218 |
ALLGOALS (asm_simp_tac |
3672
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
219 |
(!simpset addcongs [if_weak_cong] |
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
220 |
setloop split_tac [expand_if])) THEN |
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
221 |
(*Remove instances of pubK B: the Spy already knows all public keys. |
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
222 |
Combining the two simplifier calls makes them run extremely slowly.*) |
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
223 |
ALLGOALS (asm_simp_tac |
3672
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
224 |
(!simpset addcongs [if_weak_cong] |
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
225 |
addsimps [insert_absorb] |
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
226 |
setloop split_tac [expand_if])); |
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
227 |
|
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
228 |
|
3685
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
229 |
(*** Notes are made under controlled circumstances ***) |
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
230 |
|
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
231 |
goal thy "!!evs. [| Notes A {|Agent B, X|} : set evs; evs : tls |] \ |
3683 | 232 |
\ ==> Crypt (pubK B) X : parts (spies evs)"; |
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
233 |
by (etac rev_mp 1); |
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
234 |
by (analz_induct_tac 1); |
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
235 |
by (blast_tac (!claset addIs [parts_insertI]) 1); |
3683 | 236 |
qed "Notes_Crypt_parts_spies"; |
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
237 |
|
3685
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
238 |
(*C might be either A or B*) |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
239 |
goal thy |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
240 |
"!!evs. [| Notes C {|Number SID, Agent A, Agent B, Nonce M|} : set evs; \ |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
241 |
\ evs : tls \ |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
242 |
\ |] ==> M : range PRF"; |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
243 |
by (etac rev_mp 1); |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
244 |
by (parts_induct_tac 1); |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
245 |
by (Auto_tac()); |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
246 |
qed "Notes_master_range_PRF"; |
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
247 |
|
3685
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
248 |
(*C might be either A or B*) |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
249 |
goal thy |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
250 |
"!!evs. [| Notes C {|Number SID, Agent A, Agent B, Nonce(PRF(PMS,NA,NB))|} \ |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
251 |
\ : set evs; evs : tls \ |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
252 |
\ |] ==> Crypt (pubK B) (Nonce PMS) : parts (spies evs)"; |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
253 |
by (etac rev_mp 1); |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
254 |
by (parts_induct_tac 1); |
3711 | 255 |
by (ALLGOALS Clarify_tac); |
3685
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
256 |
(*Fake*) |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
257 |
by (blast_tac (!claset addIs [parts_insertI]) 1); |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
258 |
(*Client, Server Accept*) |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
259 |
by (REPEAT (blast_tac (!claset addSEs spies_partsEs |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
260 |
addSDs [Notes_Crypt_parts_spies]) 1)); |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
261 |
qed "Notes_master_imp_Crypt_PMS"; |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
262 |
|
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
263 |
(*Compared with the theorem above, both premise and conclusion are stronger*) |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
264 |
goal thy |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
265 |
"!!evs. [| Notes A {|Number SID, Agent A, Agent B, Nonce(PRF(PMS,NA,NB))|} \ |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
266 |
\ : set evs; evs : tls \ |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
267 |
\ |] ==> Notes A {|Agent B, Nonce PMS|} : set evs"; |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
268 |
by (etac rev_mp 1); |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
269 |
by (parts_induct_tac 1); |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
270 |
(*ServerAccepts*) |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
271 |
by (Fast_tac 1); (*Blast_tac's very slow here*) |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
272 |
qed "Notes_master_imp_Notes_PMS"; |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
273 |
|
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
274 |
|
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
275 |
(*Every Nonce that's hashed is already in past traffic; this event |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
276 |
occurs in CertVerify. The condition NB ~: range PRF excludes the |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
277 |
MASTER SECRET from consideration; it is created using PRF.*) |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
278 |
goal thy "!!evs. [| Hash {|Nonce NB, X|} : parts (spies evs); \ |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
279 |
\ NB ~: range PRF; evs : tls |] \ |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
280 |
\ ==> Nonce NB : parts (spies evs)"; |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
281 |
by (etac rev_mp 1); |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
282 |
by (etac rev_mp 1); |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
283 |
by (parts_induct_tac 1); |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
284 |
by (ALLGOALS (asm_simp_tac (!simpset addsimps [parts_insert_spies]))); |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
285 |
(*Server/Client Resume: wrong sort of nonce!*) |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
286 |
by (REPEAT (blast_tac (!claset addSDs [Notes_master_range_PRF]) 5)); |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
287 |
(*FINISHED messages are trivial because M : range PRF*) |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
288 |
by (REPEAT (Blast_tac 3)); |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
289 |
(*CertVerify is the only interesting case*) |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
290 |
by (blast_tac (!claset addSEs spies_partsEs) 2); |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
291 |
by (Fake_parts_insert_tac 1); |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
292 |
qed "Hash_Nonce_CV"; |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
293 |
|
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
294 |
|
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
295 |
|
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
296 |
(*** Protocol goal: if B receives CertVerify, then A sent it ***) |
3474 | 297 |
|
3506 | 298 |
(*B can check A's signature if he has received A's certificate. |
299 |
Perhaps B~=Spy is unnecessary, but there's no obvious proof if the first |
|
3480
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
300 |
message is Fake. We don't need guarantees for the Spy anyway. We must |
3683 | 301 |
assume A~:bad; otherwise, the Spy can forge A's signature.*) |
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
302 |
goal thy |
3729
6be7cf5086ab
Renamed XA, XB to PA, PB and removed the certificate from Client Verify
paulson
parents:
3711
diff
changeset
|
303 |
"!!evs. [| X = Crypt (priK A) (Hash{|Nonce NB, Agent B, Nonce PMS|}); \ |
3683 | 304 |
\ evs : tls; A ~: bad; B ~= Spy |] \ |
3729
6be7cf5086ab
Renamed XA, XB to PA, PB and removed the certificate from Client Verify
paulson
parents:
3711
diff
changeset
|
305 |
\ ==> Says B A {|Nonce NB, Number SID, Number PB, certificate B KB|} \ |
3506 | 306 |
\ : set evs --> \ |
3683 | 307 |
\ X : parts (spies evs) --> Says A B X : set evs"; |
3480
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
308 |
by (hyp_subst_tac 1); |
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
309 |
by (parts_induct_tac 1); |
3474 | 310 |
by (Fake_parts_insert_tac 1); |
3480
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
311 |
(*ServerHello: nonce NB cannot be in X because it's fresh!*) |
3672
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
312 |
by (blast_tac (!claset addSDs [Hash_Nonce_CV] |
3683 | 313 |
addSEs spies_partsEs) 1); |
3474 | 314 |
qed_spec_mp "TrustCertVerify"; |
315 |
||
316 |
||
3685
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
317 |
(*If CertVerify is present then A has chosen PMS.*) |
3506 | 318 |
goal thy |
3729
6be7cf5086ab
Renamed XA, XB to PA, PB and removed the certificate from Client Verify
paulson
parents:
3711
diff
changeset
|
319 |
"!!evs. [| Crypt (priK A) (Hash{|Nonce NB, Agent B, Nonce PMS|}) \ |
3683 | 320 |
\ : parts (spies evs); \ |
321 |
\ evs : tls; A ~: bad |] \ |
|
3672
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
322 |
\ ==> Notes A {|Agent B, Nonce PMS|} : set evs"; |
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
323 |
be rev_mp 1; |
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
324 |
by (parts_induct_tac 1); |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
325 |
by (Fake_parts_insert_tac 1); |
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
326 |
qed "UseCertVerify"; |
3474 | 327 |
|
3480
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
328 |
|
3687
fb7d096d7884
Simplified SpyKeys to use sessionK instead of clientK and serverK
paulson
parents:
3686
diff
changeset
|
329 |
(*Key compromise lemma needed to prove analz_image_keys. |
fb7d096d7884
Simplified SpyKeys to use sessionK instead of clientK and serverK
paulson
parents:
3686
diff
changeset
|
330 |
No collection of keys can help the spy get new private keys.*) |
3480
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
331 |
goal thy |
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
332 |
"!!evs. evs : tls ==> \ |
3683 | 333 |
\ ALL KK. (Key(priK B) : analz (Key``KK Un (spies evs))) = \ |
3687
fb7d096d7884
Simplified SpyKeys to use sessionK instead of clientK and serverK
paulson
parents:
3686
diff
changeset
|
334 |
\ (priK B : KK | B : bad)"; |
3480
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
335 |
by (etac tls.induct 1); |
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
336 |
by (ALLGOALS |
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
337 |
(asm_simp_tac (analz_image_keys_ss |
3677
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
338 |
addsimps (analz_insert_certificate::keys_distinct)))); |
3480
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
339 |
(*Fake*) |
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
340 |
by (spy_analz_tac 2); |
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
341 |
(*Base*) |
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
342 |
by (Blast_tac 1); |
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
343 |
qed_spec_mp "analz_image_priK"; |
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
344 |
|
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
345 |
|
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
346 |
(*Lemma for the trivial direction of the if-and-only-if*) |
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
347 |
goal thy |
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
348 |
"!!evs. (X : analz (G Un H)) --> (X : analz H) ==> \ |
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
349 |
\ (X : analz (G Un H)) = (X : analz H)"; |
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
350 |
by (blast_tac (!claset addIs [impOfSubs analz_mono]) 1); |
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
351 |
val lemma = result(); |
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
352 |
|
3677
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
353 |
(*slightly speeds up the big simplification below*) |
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
354 |
goal thy "!!evs. KK <= range sessionK ==> priK B ~: KK"; |
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
355 |
by (Blast_tac 1); |
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
356 |
val range_sessionkeys_not_priK = result(); |
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
357 |
|
3687
fb7d096d7884
Simplified SpyKeys to use sessionK instead of clientK and serverK
paulson
parents:
3686
diff
changeset
|
358 |
(** It is a mystery to me why the following formulation is actually slower |
fb7d096d7884
Simplified SpyKeys to use sessionK instead of clientK and serverK
paulson
parents:
3686
diff
changeset
|
359 |
in simplification: |
fb7d096d7884
Simplified SpyKeys to use sessionK instead of clientK and serverK
paulson
parents:
3686
diff
changeset
|
360 |
|
fb7d096d7884
Simplified SpyKeys to use sessionK instead of clientK and serverK
paulson
parents:
3686
diff
changeset
|
361 |
\ ALL Z. (Nonce N : analz (Key``(sessionK``Z) Un (spies evs))) = \ |
fb7d096d7884
Simplified SpyKeys to use sessionK instead of clientK and serverK
paulson
parents:
3686
diff
changeset
|
362 |
\ (Nonce N : analz (spies evs))"; |
fb7d096d7884
Simplified SpyKeys to use sessionK instead of clientK and serverK
paulson
parents:
3686
diff
changeset
|
363 |
|
fb7d096d7884
Simplified SpyKeys to use sessionK instead of clientK and serverK
paulson
parents:
3686
diff
changeset
|
364 |
More so as it can take advantage of unconditional rewrites such as |
fb7d096d7884
Simplified SpyKeys to use sessionK instead of clientK and serverK
paulson
parents:
3686
diff
changeset
|
365 |
priK B ~: sessionK``Z |
fb7d096d7884
Simplified SpyKeys to use sessionK instead of clientK and serverK
paulson
parents:
3686
diff
changeset
|
366 |
**) |
fb7d096d7884
Simplified SpyKeys to use sessionK instead of clientK and serverK
paulson
parents:
3686
diff
changeset
|
367 |
|
3480
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
368 |
goal thy |
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
369 |
"!!evs. evs : tls ==> \ |
3677
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
370 |
\ ALL KK. KK <= range sessionK --> \ |
3683 | 371 |
\ (Nonce N : analz (Key``KK Un (spies evs))) = \ |
372 |
\ (Nonce N : analz (spies evs))"; |
|
3480
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
373 |
by (etac tls.induct 1); |
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
374 |
by (ClientCertKeyEx_tac 6); |
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
375 |
by (REPEAT_FIRST (resolve_tac [allI, impI])); |
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
376 |
by (REPEAT_FIRST (rtac lemma)); |
3677
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
377 |
writeln"SLOW simplification: 55 secs??"; |
3683 | 378 |
(*ClientCertKeyEx is to blame, causing case splits for A, B: bad*) |
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
379 |
by (ALLGOALS |
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
380 |
(asm_simp_tac (analz_image_keys_ss |
3677
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
381 |
addsimps [range_sessionkeys_not_priK, |
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
382 |
analz_image_priK, analz_insert_certificate]))); |
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
383 |
by (ALLGOALS (asm_simp_tac (!simpset addsimps [insert_absorb]))); |
3480
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
384 |
(*Fake*) |
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
385 |
by (spy_analz_tac 2); |
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
386 |
(*Base*) |
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
387 |
by (Blast_tac 1); |
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
388 |
qed_spec_mp "analz_image_keys"; |
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
389 |
|
3687
fb7d096d7884
Simplified SpyKeys to use sessionK instead of clientK and serverK
paulson
parents:
3686
diff
changeset
|
390 |
(*Knowing some session keys is no help in getting new nonces*) |
fb7d096d7884
Simplified SpyKeys to use sessionK instead of clientK and serverK
paulson
parents:
3686
diff
changeset
|
391 |
goal thy |
fb7d096d7884
Simplified SpyKeys to use sessionK instead of clientK and serverK
paulson
parents:
3686
diff
changeset
|
392 |
"!!evs. evs : tls ==> \ |
fb7d096d7884
Simplified SpyKeys to use sessionK instead of clientK and serverK
paulson
parents:
3686
diff
changeset
|
393 |
\ Nonce N : analz (insert (Key (sessionK z)) (spies evs)) = \ |
fb7d096d7884
Simplified SpyKeys to use sessionK instead of clientK and serverK
paulson
parents:
3686
diff
changeset
|
394 |
\ (Nonce N : analz (spies evs))"; |
fb7d096d7884
Simplified SpyKeys to use sessionK instead of clientK and serverK
paulson
parents:
3686
diff
changeset
|
395 |
by (asm_simp_tac (analz_image_keys_ss addsimps [analz_image_keys]) 1); |
fb7d096d7884
Simplified SpyKeys to use sessionK instead of clientK and serverK
paulson
parents:
3686
diff
changeset
|
396 |
qed "analz_insert_key"; |
fb7d096d7884
Simplified SpyKeys to use sessionK instead of clientK and serverK
paulson
parents:
3686
diff
changeset
|
397 |
Addsimps [analz_insert_key]; |
3480
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
398 |
|
3672
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
399 |
goal thy "!!evs. evs : tls ==> Notes A {|Agent B, Nonce (PRF x)|} ~: set evs"; |
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
400 |
by (parts_induct_tac 1); |
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
401 |
(*ClientCertKeyEx: PMS is assumed to differ from any PRF.*) |
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
402 |
by (Blast_tac 1); |
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
403 |
qed "no_Notes_A_PRF"; |
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
404 |
Addsimps [no_Notes_A_PRF]; |
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
405 |
|
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
406 |
|
3683 | 407 |
goal thy "!!evs. [| Nonce (PRF (PMS,NA,NB)) : parts (spies evs); \ |
3672
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
408 |
\ evs : tls |] \ |
3683 | 409 |
\ ==> Nonce PMS : parts (spies evs)"; |
3672
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
410 |
by (etac rev_mp 1); |
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
411 |
by (parts_induct_tac 1); |
3683 | 412 |
by (ALLGOALS (asm_simp_tac (!simpset addsimps [parts_insert_spies]))); |
3672
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
413 |
by (Fake_parts_insert_tac 1); |
3677
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
414 |
(*Six others, all trivial or by freshness*) |
3683 | 415 |
by (REPEAT (blast_tac (!claset addSDs [Notes_Crypt_parts_spies] |
416 |
addSEs spies_partsEs) 1)); |
|
3672
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
417 |
qed "MS_imp_PMS"; |
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
418 |
AddSDs [MS_imp_PMS]; |
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
419 |
|
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
420 |
|
3474 | 421 |
|
422 |
(*** Protocol goal: serverK(NA,NB,M) and clientK(NA,NB,M) remain secure ***) |
|
423 |
||
3677
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
424 |
(** Some lemmas about session keys, comprising clientK and serverK **) |
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
425 |
|
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
426 |
|
3704 | 427 |
(*Lemma: session keys are never used if PMS is fresh. |
428 |
Nonces don't have to agree, allowing session resumption. |
|
429 |
Converse doesn't hold; revealing PMS doesn't force the keys to be sent. |
|
430 |
THEY ARE NOT SUITABLE AS SAFE ELIM RULES.*) |
|
431 |
goal thy |
|
432 |
"!!evs. [| Nonce PMS ~: parts (spies evs); \ |
|
433 |
\ K = sessionK((Na, Nb, PRF(PMS,NA,NB)), b); \ |
|
434 |
\ evs : tls |] \ |
|
435 |
\ ==> Key K ~: parts (spies evs) & (ALL Y. Crypt K Y ~: parts (spies evs))"; |
|
436 |
by (etac rev_mp 1); |
|
437 |
by (hyp_subst_tac 1); |
|
438 |
by (analz_induct_tac 1); |
|
439 |
(*SpyKeys*) |
|
440 |
by (blast_tac (!claset addSEs spies_partsEs) 3); |
|
441 |
(*Fake*) |
|
442 |
by (simp_tac (!simpset addsimps [parts_insert_spies]) 2); |
|
443 |
by (Fake_parts_insert_tac 2); |
|
444 |
(** LEVEL 6 **) |
|
445 |
(*Oops*) |
|
446 |
by (fast_tac (!claset addSEs [MPair_parts] |
|
447 |
addDs [Says_imp_spies RS parts.Inj] |
|
448 |
addss (!simpset)) 6); |
|
449 |
by (REPEAT |
|
450 |
(blast_tac (!claset addSDs [Notes_Crypt_parts_spies, |
|
451 |
Notes_master_imp_Crypt_PMS] |
|
452 |
addSEs spies_partsEs) 1)); |
|
453 |
val lemma = result(); |
|
454 |
||
455 |
goal thy |
|
456 |
"!!evs. [| Nonce PMS ~: parts (spies evs); evs : tls |] \ |
|
457 |
\ ==> Key (sessionK((Na, Nb, PRF(PMS,NA,NB)), b)) ~: parts (spies evs)"; |
|
458 |
by (blast_tac (!claset addDs [lemma]) 1); |
|
459 |
qed "PMS_sessionK_not_spied"; |
|
460 |
||
461 |
goal thy |
|
462 |
"!!evs. [| Nonce PMS ~: parts (spies evs); evs : tls |] \ |
|
463 |
\ ==> Crypt (sessionK((Na, Nb, PRF(PMS,NA,NB)), b)) Y ~: parts (spies evs)"; |
|
464 |
by (blast_tac (!claset addDs [lemma]) 1); |
|
465 |
qed "PMS_Crypt_sessionK_not_spied"; |
|
466 |
||
467 |
||
468 |
(*Lemma: write keys are never sent if M (MASTER SECRET) is secure. |
|
469 |
Converse doesn't hold; betraying M doesn't force the keys to be sent! |
|
470 |
The strong Oops condition can be weakened later by unicity reasoning, |
|
471 |
with some effort.*) |
|
3474 | 472 |
goal thy |
3686
4b484805b4c4
First working version with Oops event for session keys
paulson
parents:
3685
diff
changeset
|
473 |
"!!evs. [| ALL A. Says A Spy (Key (sessionK((NA,NB,M),b))) ~: set evs; \ |
4b484805b4c4
First working version with Oops event for session keys
paulson
parents:
3685
diff
changeset
|
474 |
\ Nonce M ~: analz (spies evs); evs : tls |] \ |
4b484805b4c4
First working version with Oops event for session keys
paulson
parents:
3685
diff
changeset
|
475 |
\ ==> Key (sessionK((NA,NB,M),b)) ~: parts (spies evs)"; |
4b484805b4c4
First working version with Oops event for session keys
paulson
parents:
3685
diff
changeset
|
476 |
by (etac rev_mp 1); |
3480
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
477 |
by (etac rev_mp 1); |
3704 | 478 |
by (analz_induct_tac 1); (*30 seconds??*) |
3686
4b484805b4c4
First working version with Oops event for session keys
paulson
parents:
3685
diff
changeset
|
479 |
(*Oops*) |
4b484805b4c4
First working version with Oops event for session keys
paulson
parents:
3685
diff
changeset
|
480 |
by (Blast_tac 4); |
3480
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
481 |
(*SpyKeys*) |
3683 | 482 |
by (blast_tac (!claset addDs [Says_imp_spies RS analz.Inj]) 3); |
3480
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
483 |
(*Fake*) |
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
484 |
by (spy_analz_tac 2); |
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
485 |
(*Base*) |
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
486 |
by (Blast_tac 1); |
3704 | 487 |
qed "sessionK_not_spied"; |
488 |
Addsimps [sessionK_not_spied]; |
|
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
489 |
|
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
490 |
|
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
491 |
(*NEEDED??*) |
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
492 |
goal thy |
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
493 |
"!!evs. [| Says A B {|certA, Crypt KB (Nonce M)|} : set evs; \ |
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
494 |
\ A ~= Spy; evs : tls |] ==> KB = pubK B"; |
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
495 |
be rev_mp 1; |
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
496 |
by (analz_induct_tac 1); |
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
497 |
qed "A_Crypt_pubB"; |
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
498 |
|
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
499 |
|
3672
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
500 |
(*** Unicity results for PMS, the pre-master-secret ***) |
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
501 |
|
3672
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
502 |
(*PMS determines B. Proof borrowed from NS_Public/unique_NA and from Yahalom*) |
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
503 |
goal thy |
3683 | 504 |
"!!evs. [| Nonce PMS ~: analz (spies evs); evs : tls |] \ |
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
505 |
\ ==> EX B'. ALL B. \ |
3683 | 506 |
\ Crypt (pubK B) (Nonce PMS) : parts (spies evs) --> B=B'"; |
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
507 |
by (etac rev_mp 1); |
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
508 |
by (parts_induct_tac 1); |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
509 |
by (Fake_parts_insert_tac 1); |
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
510 |
(*ClientCertKeyEx*) |
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
511 |
by (ClientCertKeyEx_tac 1); |
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
512 |
by (asm_simp_tac (!simpset addsimps [all_conj_distrib]) 1); |
3672
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
513 |
by (expand_case_tac "PMS = ?y" 1 THEN |
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
514 |
blast_tac (!claset addSEs partsEs) 1); |
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
515 |
val lemma = result(); |
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
516 |
|
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
517 |
goal thy |
3683 | 518 |
"!!evs. [| Crypt(pubK B) (Nonce PMS) : parts (spies evs); \ |
519 |
\ Crypt(pubK B') (Nonce PMS) : parts (spies evs); \ |
|
520 |
\ Nonce PMS ~: analz (spies evs); \ |
|
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
521 |
\ evs : tls |] \ |
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
522 |
\ ==> B=B'"; |
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
523 |
by (prove_unique_tac lemma 1); |
3704 | 524 |
qed "Crypt_unique_PMS"; |
525 |
||
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
526 |
|
3685
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
527 |
(** It is frustrating that we need two versions of the unicity results. |
3704 | 528 |
But Notes A {|Agent B, Nonce PMS|} determines both A and B. Sometimes |
529 |
we have only the weaker assertion Crypt(pubK B) (Nonce PMS), which |
|
530 |
determines B alone, and only if PMS is secret. |
|
3685
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
531 |
**) |
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
532 |
|
3677
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
533 |
(*In A's internal Note, PMS determines A and B.*) |
3704 | 534 |
goal thy "!!evs. evs : tls \ |
535 |
\ ==> EX A' B'. ALL A B. \ |
|
536 |
\ Notes A {|Agent B, Nonce PMS|} : set evs --> A=A' & B=B'"; |
|
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
537 |
by (parts_induct_tac 1); |
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
538 |
by (asm_simp_tac (!simpset addsimps [all_conj_distrib]) 1); |
3672
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
539 |
(*ClientCertKeyEx: if PMS is fresh, then it can't appear in Notes A X.*) |
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
540 |
by (expand_case_tac "PMS = ?y" 1 THEN |
3683 | 541 |
blast_tac (!claset addSDs [Notes_Crypt_parts_spies] addSEs partsEs) 1); |
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
542 |
val lemma = result(); |
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
543 |
|
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
544 |
goal thy |
3672
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
545 |
"!!evs. [| Notes A {|Agent B, Nonce PMS|} : set evs; \ |
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
546 |
\ Notes A' {|Agent B', Nonce PMS|} : set evs; \ |
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
547 |
\ evs : tls |] \ |
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
548 |
\ ==> A=A' & B=B'"; |
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
549 |
by (prove_unique_tac lemma 1); |
3672
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
550 |
qed "Notes_unique_PMS"; |
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
551 |
|
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
552 |
|
3474 | 553 |
|
3677
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
554 |
(*If A sends ClientCertKeyEx to an honest B, then the PMS will stay secret.*) |
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
555 |
goal thy |
3683 | 556 |
"!!evs. [| evs : tls; A ~: bad; B ~: bad |] \ |
3677
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
557 |
\ ==> Notes A {|Agent B, Nonce PMS|} : set evs --> \ |
3683 | 558 |
\ Nonce PMS ~: analz (spies evs)"; |
3685
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
559 |
by (analz_induct_tac 1); (*30 seconds??*) |
3677
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
560 |
(*ClientAccepts and ServerAccepts: because PMS ~: range PRF*) |
3686
4b484805b4c4
First working version with Oops event for session keys
paulson
parents:
3685
diff
changeset
|
561 |
by (EVERY (map (fast_tac (!claset addss (!simpset))) [7,6])); |
3687
fb7d096d7884
Simplified SpyKeys to use sessionK instead of clientK and serverK
paulson
parents:
3686
diff
changeset
|
562 |
(*ClientHello, ServerHello, ClientCertKeyEx, ServerResume: |
fb7d096d7884
Simplified SpyKeys to use sessionK instead of clientK and serverK
paulson
parents:
3686
diff
changeset
|
563 |
mostly freshness reasoning*) |
3677
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
564 |
by (REPEAT (blast_tac (!claset addSEs partsEs |
3683 | 565 |
addDs [Notes_Crypt_parts_spies, |
3677
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
566 |
impOfSubs analz_subset_parts, |
3687
fb7d096d7884
Simplified SpyKeys to use sessionK instead of clientK and serverK
paulson
parents:
3686
diff
changeset
|
567 |
Says_imp_spies RS analz.Inj]) 3)); |
3677
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
568 |
(*SpyKeys*) |
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
569 |
by (fast_tac (!claset addss (!simpset)) 2); |
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
570 |
(*Fake*) |
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
571 |
by (spy_analz_tac 1); |
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
572 |
bind_thm ("Spy_not_see_PMS", result() RSN (2, rev_mp)); |
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
573 |
|
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
574 |
|
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
575 |
(*If A sends ClientCertKeyEx to an honest B, then the MASTER SECRET |
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
576 |
will stay secret.*) |
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
577 |
goal thy |
3683 | 578 |
"!!evs. [| evs : tls; A ~: bad; B ~: bad |] \ |
3677
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
579 |
\ ==> Notes A {|Agent B, Nonce PMS|} : set evs --> \ |
3683 | 580 |
\ Nonce (PRF(PMS,NA,NB)) ~: analz (spies evs)"; |
3677
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
581 |
by (analz_induct_tac 1); (*35 seconds*) |
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
582 |
(*ClientAccepts and ServerAccepts: because PMS was already visible*) |
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
583 |
by (REPEAT (blast_tac (!claset addDs [Spy_not_see_PMS, |
3683 | 584 |
Says_imp_spies RS analz.Inj, |
585 |
Notes_imp_spies RS analz.Inj]) 6)); |
|
3677
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
586 |
(*ClientHello*) |
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
587 |
by (Blast_tac 3); |
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
588 |
(*SpyKeys: by secrecy of the PMS, Spy cannot make the MS*) |
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
589 |
by (blast_tac (!claset addSDs [Spy_not_see_PMS, |
3683 | 590 |
Says_imp_spies RS analz.Inj]) 2); |
3677
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
591 |
(*Fake*) |
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
592 |
by (spy_analz_tac 1); |
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
593 |
(*ServerHello and ClientCertKeyEx: mostly freshness reasoning*) |
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
594 |
by (REPEAT (blast_tac (!claset addSEs partsEs |
3683 | 595 |
addDs [Notes_Crypt_parts_spies, |
3677
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
596 |
impOfSubs analz_subset_parts, |
3683 | 597 |
Says_imp_spies RS analz.Inj]) 1)); |
3677
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
598 |
bind_thm ("Spy_not_see_MS", result() RSN (2, rev_mp)); |
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
599 |
|
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
600 |
|
3704 | 601 |
(*** Weakening the Oops conditions for leakage of clientK ***) |
602 |
||
603 |
(*If A created PMS then nobody other than the Spy would send a message |
|
604 |
using a clientK generated from that PMS.*) |
|
605 |
goal thy |
|
606 |
"!!evs. [| evs : tls; A' ~= Spy |] \ |
|
607 |
\ ==> Notes A {|Agent B, Nonce PMS|} : set evs --> \ |
|
608 |
\ Says A' B' (Crypt (clientK(Na,Nb,PRF(PMS,NA,NB))) Y) : set evs --> \ |
|
609 |
\ A = A'"; |
|
610 |
by (analz_induct_tac 1); (*17 seconds*) |
|
3711 | 611 |
by (ALLGOALS Clarify_tac); |
3704 | 612 |
(*ClientFinished, ClientResume: by unicity of PMS*) |
613 |
by (REPEAT |
|
614 |
(blast_tac (!claset addSDs [Notes_master_imp_Notes_PMS] |
|
615 |
addIs [Notes_unique_PMS RS conjunct1]) 2)); |
|
616 |
(*ClientCertKeyEx*) |
|
617 |
by (blast_tac (!claset addSEs [PMS_Crypt_sessionK_not_spied RSN (2,rev_notE)] |
|
618 |
addSDs [Says_imp_spies RS parts.Inj]) 1); |
|
619 |
bind_thm ("Says_clientK_unique", |
|
620 |
result() RSN(2,rev_mp) RSN(2,rev_mp)); |
|
621 |
||
622 |
||
623 |
(*If A created PMS and has not leaked her clientK to the Spy, |
|
624 |
then nobody has.*) |
|
625 |
goal thy |
|
626 |
"!!evs. evs : tls \ |
|
627 |
\ ==> Says A Spy (Key(clientK(Na,Nb,PRF(PMS,NA,NB)))) ~: set evs --> \ |
|
628 |
\ Notes A {|Agent B, Nonce PMS|} : set evs --> \ |
|
629 |
\ (ALL A. Says A Spy (Key(clientK(Na,Nb,PRF(PMS,NA,NB)))) ~: set evs) "; |
|
630 |
by (etac tls.induct 1); |
|
631 |
(*This roundabout proof sequence avoids looping in simplification*) |
|
632 |
by (ALLGOALS Simp_tac); |
|
3711 | 633 |
by (ALLGOALS Clarify_tac); |
3704 | 634 |
by (Fake_parts_insert_tac 1); |
635 |
by (ALLGOALS Asm_simp_tac); |
|
636 |
(*Oops*) |
|
637 |
by (blast_tac (!claset addIs [Says_clientK_unique]) 2); |
|
638 |
(*ClientCertKeyEx*) |
|
639 |
by (blast_tac (!claset addSEs ((PMS_sessionK_not_spied RSN (2,rev_notE)) :: |
|
640 |
spies_partsEs)) 1); |
|
641 |
qed_spec_mp "clientK_Oops_ALL"; |
|
642 |
||
643 |
||
644 |
||
645 |
(*** Weakening the Oops conditions for leakage of serverK ***) |
|
646 |
||
647 |
(*If A created PMS for B, then nobody other than B or the Spy would |
|
648 |
send a message using a serverK generated from that PMS.*) |
|
649 |
goal thy |
|
650 |
"!!evs. [| evs : tls; A ~: bad; B ~: bad; B' ~= Spy |] \ |
|
651 |
\ ==> Notes A {|Agent B, Nonce PMS|} : set evs --> \ |
|
652 |
\ Says B' A' (Crypt (serverK(Na,Nb,PRF(PMS,NA,NB))) Y) : set evs --> \ |
|
653 |
\ B = B'"; |
|
654 |
by (analz_induct_tac 1); (*17 seconds*) |
|
3711 | 655 |
by (ALLGOALS Clarify_tac); |
3704 | 656 |
(*ServerResume, ServerFinished: by unicity of PMS*) |
657 |
by (REPEAT |
|
658 |
(blast_tac (!claset addSEs [MPair_parts] |
|
659 |
addSDs [Notes_master_imp_Crypt_PMS, |
|
660 |
Says_imp_spies RS parts.Inj] |
|
661 |
addDs [Spy_not_see_PMS, |
|
662 |
Notes_Crypt_parts_spies, |
|
663 |
Crypt_unique_PMS]) 2)); |
|
664 |
(*ClientCertKeyEx*) |
|
665 |
by (blast_tac (!claset addSEs [PMS_Crypt_sessionK_not_spied RSN (2,rev_notE)] |
|
666 |
addSDs [Says_imp_spies RS parts.Inj]) 1); |
|
667 |
bind_thm ("Says_serverK_unique", |
|
668 |
result() RSN(2,rev_mp) RSN(2,rev_mp)); |
|
669 |
||
670 |
(*If A created PMS for B, and B has not leaked his serverK to the Spy, |
|
671 |
then nobody has.*) |
|
672 |
goal thy |
|
673 |
"!!evs. [| evs : tls; A ~: bad; B ~: bad |] \ |
|
674 |
\ ==> Says B Spy (Key(serverK(Na,Nb,PRF(PMS,NA,NB)))) ~: set evs --> \ |
|
675 |
\ Notes A {|Agent B, Nonce PMS|} : set evs --> \ |
|
676 |
\ (ALL A. Says A Spy (Key(serverK(Na,Nb,PRF(PMS,NA,NB)))) ~: set evs) "; |
|
677 |
by (etac tls.induct 1); |
|
678 |
(*This roundabout proof sequence avoids looping in simplification*) |
|
679 |
by (ALLGOALS Simp_tac); |
|
3711 | 680 |
by (ALLGOALS Clarify_tac); |
3704 | 681 |
by (Fake_parts_insert_tac 1); |
682 |
by (ALLGOALS Asm_simp_tac); |
|
683 |
(*Oops*) |
|
684 |
by (blast_tac (!claset addIs [Says_serverK_unique]) 2); |
|
685 |
(*ClientCertKeyEx*) |
|
686 |
by (blast_tac (!claset addSEs ((PMS_sessionK_not_spied RSN (2,rev_notE)) :: |
|
687 |
spies_partsEs)) 1); |
|
688 |
qed_spec_mp "serverK_Oops_ALL"; |
|
689 |
||
690 |
||
691 |
||
3685
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
692 |
(*** Protocol goals: if A receives ServerFinished, then B is present |
3729
6be7cf5086ab
Renamed XA, XB to PA, PB and removed the certificate from Client Verify
paulson
parents:
3711
diff
changeset
|
693 |
and has used the quoted values PA, PB, etc. Note that it is up to A |
6be7cf5086ab
Renamed XA, XB to PA, PB and removed the certificate from Client Verify
paulson
parents:
3711
diff
changeset
|
694 |
to compare PA with what she originally sent. |
3474 | 695 |
***) |
696 |
||
3685
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
697 |
(*The mention of her name (A) in X assures A that B knows who she is.*) |
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
698 |
goal thy |
3686
4b484805b4c4
First working version with Oops event for session keys
paulson
parents:
3685
diff
changeset
|
699 |
"!!evs. [| ALL A. Says A Spy (Key (serverK(Na,Nb,M))) ~: set evs; \ |
4b484805b4c4
First working version with Oops event for session keys
paulson
parents:
3685
diff
changeset
|
700 |
\ X = Crypt (serverK(Na,Nb,M)) \ |
3676
cbaec955056b
Addition of SessionIDs to the Hello and Finished messages
paulson
parents:
3672
diff
changeset
|
701 |
\ (Hash{|Nonce M, Number SID, \ |
3729
6be7cf5086ab
Renamed XA, XB to PA, PB and removed the certificate from Client Verify
paulson
parents:
3711
diff
changeset
|
702 |
\ Nonce NA, Number PA, Agent A, \ |
6be7cf5086ab
Renamed XA, XB to PA, PB and removed the certificate from Client Verify
paulson
parents:
3711
diff
changeset
|
703 |
\ Nonce NB, Number PB, Agent B|}); \ |
3676
cbaec955056b
Addition of SessionIDs to the Hello and Finished messages
paulson
parents:
3672
diff
changeset
|
704 |
\ M = PRF(PMS,NA,NB); \ |
3683 | 705 |
\ evs : tls; A ~: bad; B ~: bad |] \ |
3676
cbaec955056b
Addition of SessionIDs to the Hello and Finished messages
paulson
parents:
3672
diff
changeset
|
706 |
\ ==> Notes A {|Agent B, Nonce PMS|} : set evs --> \ |
3683 | 707 |
\ X : parts (spies evs) --> Says B A X : set evs"; |
3686
4b484805b4c4
First working version with Oops event for session keys
paulson
parents:
3685
diff
changeset
|
708 |
by (etac rev_mp 1); |
3480
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
709 |
by (hyp_subst_tac 1); |
3704 | 710 |
by (analz_induct_tac 1); (*27 seconds*) |
711 |
by (ALLGOALS (asm_simp_tac (!simpset addsimps [all_conj_distrib]))); |
|
3711 | 712 |
(*proves ServerResume*) |
713 |
by (ALLGOALS Clarify_tac); |
|
3677
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
714 |
(*ClientCertKeyEx*) |
3704 | 715 |
by (fast_tac (*blast_tac gives PROOF FAILED*) |
716 |
(!claset addSEs [PMS_Crypt_sessionK_not_spied RSN (2,rev_notE)]) 2); |
|
3480
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
717 |
(*Fake: the Spy doesn't have the critical session key!*) |
3683 | 718 |
by (subgoal_tac "Key (serverK(Na,Nb,PRF(PMS,NA,NB))) ~: analz(spies evsa)" 1); |
3672
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
719 |
by (asm_simp_tac (!simpset addsimps [Spy_not_see_MS, |
3480
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
720 |
not_parts_not_analz]) 2); |
3474 | 721 |
by (Fake_parts_insert_tac 1); |
3704 | 722 |
val lemma = normalize_thm [RSspec, RSmp] (result()); |
723 |
||
724 |
(*Final version*) |
|
725 |
goal thy |
|
726 |
"!!evs. [| X = Crypt (serverK(Na,Nb,M)) \ |
|
727 |
\ (Hash{|Nonce M, Number SID, \ |
|
3729
6be7cf5086ab
Renamed XA, XB to PA, PB and removed the certificate from Client Verify
paulson
parents:
3711
diff
changeset
|
728 |
\ Nonce NA, Number PA, Agent A, \ |
6be7cf5086ab
Renamed XA, XB to PA, PB and removed the certificate from Client Verify
paulson
parents:
3711
diff
changeset
|
729 |
\ Nonce NB, Number PB, Agent B|}); \ |
3704 | 730 |
\ M = PRF(PMS,NA,NB); \ |
731 |
\ X : parts (spies evs); \ |
|
732 |
\ Notes A {|Agent B, Nonce PMS|} : set evs; \ |
|
733 |
\ Says B Spy (Key (serverK(Na,Nb,M))) ~: set evs; \ |
|
734 |
\ evs : tls; A ~: bad; B ~: bad |] \ |
|
735 |
\ ==> Says B A X : set evs"; |
|
736 |
by (blast_tac (!claset addIs [lemma] |
|
737 |
addEs [serverK_Oops_ALL RSN(2, rev_notE)]) 1); |
|
3474 | 738 |
qed_spec_mp "TrustServerFinished"; |
739 |
||
740 |
||
3704 | 741 |
|
3685
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
742 |
(*This version refers not to ServerFinished but to any message from B. |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
743 |
We don't assume B has received CertVerify, and an intruder could |
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
744 |
have changed A's identity in all other messages, so we can't be sure |
3519
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents:
3515
diff
changeset
|
745 |
that B sends his message to A. If CLIENT KEY EXCHANGE were augmented |
3704 | 746 |
to bind A's identity with PMS, then we could replace A' by A below.*) |
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
747 |
goal thy |
3686
4b484805b4c4
First working version with Oops event for session keys
paulson
parents:
3685
diff
changeset
|
748 |
"!!evs. [| ALL A. Says A Spy (Key (serverK(Na,Nb,M))) ~: set evs; \ |
4b484805b4c4
First working version with Oops event for session keys
paulson
parents:
3685
diff
changeset
|
749 |
\ evs : tls; A ~: bad; B ~: bad; \ |
3672
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
750 |
\ M = PRF(PMS,NA,NB) |] \ |
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
751 |
\ ==> Notes A {|Agent B, Nonce PMS|} : set evs --> \ |
3683 | 752 |
\ Crypt (serverK(Na,Nb,M)) Y : parts (spies evs) --> \ |
3672
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
753 |
\ (EX A'. Says B A' (Crypt (serverK(Na,Nb,M)) Y) : set evs)"; |
3686
4b484805b4c4
First working version with Oops event for session keys
paulson
parents:
3685
diff
changeset
|
754 |
by (etac rev_mp 1); |
3672
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
755 |
by (hyp_subst_tac 1); |
3686
4b484805b4c4
First working version with Oops event for session keys
paulson
parents:
3685
diff
changeset
|
756 |
by (analz_induct_tac 1); (*20 seconds*) |
3704 | 757 |
by (ALLGOALS (asm_simp_tac (!simpset addsimps [ex_disj_distrib]))); |
3711 | 758 |
by (ALLGOALS Clarify_tac); |
3704 | 759 |
(*ServerResume, ServerFinished: by unicity of PMS*) |
760 |
by (REPEAT |
|
761 |
(blast_tac (!claset addSEs [MPair_parts] |
|
762 |
addSDs [Notes_master_imp_Crypt_PMS, |
|
763 |
Says_imp_spies RS parts.Inj] |
|
764 |
addDs [Spy_not_see_PMS, |
|
765 |
Notes_Crypt_parts_spies, |
|
766 |
Crypt_unique_PMS]) 3)); |
|
3677
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
767 |
(*ClientCertKeyEx*) |
3704 | 768 |
by (blast_tac |
769 |
(!claset addSEs [PMS_Crypt_sessionK_not_spied RSN (2,rev_notE)]) 2); |
|
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
770 |
(*Fake: the Spy doesn't have the critical session key!*) |
3683 | 771 |
by (subgoal_tac "Key (serverK(Na,Nb,PRF(PMS,NA,NB))) ~: analz(spies evsa)" 1); |
3672
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
772 |
by (asm_simp_tac (!simpset addsimps [Spy_not_see_MS, |
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
773 |
not_parts_not_analz]) 2); |
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
774 |
by (Fake_parts_insert_tac 1); |
3704 | 775 |
val lemma = normalize_thm [RSspec, RSmp] (result()); |
776 |
||
777 |
(*Final version*) |
|
778 |
goal thy |
|
779 |
"!!evs. [| M = PRF(PMS,NA,NB); \ |
|
780 |
\ Crypt (serverK(Na,Nb,M)) Y : parts (spies evs); \ |
|
781 |
\ Notes A {|Agent B, Nonce PMS|} : set evs; \ |
|
782 |
\ Says B Spy (Key (serverK(Na,Nb,M))) ~: set evs; \ |
|
783 |
\ evs : tls; A ~: bad; B ~: bad |] \ |
|
784 |
\ ==> EX A'. Says B A' (Crypt (serverK(Na,Nb,M)) Y) : set evs"; |
|
785 |
by (blast_tac (!claset addIs [lemma] |
|
786 |
addEs [serverK_Oops_ALL RSN(2, rev_notE)]) 1); |
|
787 |
||
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
788 |
qed_spec_mp "TrustServerMsg"; |
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
789 |
|
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
790 |
|
3704 | 791 |
|
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
792 |
(*** Protocol goal: if B receives any message encrypted with clientK |
3672
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
793 |
then A has sent it, ASSUMING that A chose PMS. Authentication is |
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
794 |
assumed here; B cannot verify it. But if the message is |
3729
6be7cf5086ab
Renamed XA, XB to PA, PB and removed the certificate from Client Verify
paulson
parents:
3711
diff
changeset
|
795 |
ClientFinished, then B can then check the quoted values PA, PB, etc. |
3506 | 796 |
***) |
3704 | 797 |
|
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
798 |
goal thy |
3683 | 799 |
"!!evs. [| evs : tls; A ~: bad; B ~: bad |] \ |
3686
4b484805b4c4
First working version with Oops event for session keys
paulson
parents:
3685
diff
changeset
|
800 |
\ ==> (ALL A. Says A Spy (Key(clientK(Na,Nb,PRF(PMS,NA,NB)))) ~: set evs) -->\ |
4b484805b4c4
First working version with Oops event for session keys
paulson
parents:
3685
diff
changeset
|
801 |
\ Notes A {|Agent B, Nonce PMS|} : set evs --> \ |
3683 | 802 |
\ Crypt (clientK(Na,Nb,PRF(PMS,NA,NB))) Y : parts (spies evs) --> \ |
3672
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
803 |
\ Says A B (Crypt (clientK(Na,Nb,PRF(PMS,NA,NB))) Y) : set evs"; |
3686
4b484805b4c4
First working version with Oops event for session keys
paulson
parents:
3685
diff
changeset
|
804 |
by (analz_induct_tac 1); (*23 seconds*) |
3711 | 805 |
by (ALLGOALS Clarify_tac); |
3704 | 806 |
(*ClientFinished, ClientResume: by unicity of PMS*) |
807 |
by (REPEAT (blast_tac (!claset delrules [conjI] |
|
808 |
addSDs [Notes_master_imp_Notes_PMS] |
|
809 |
addDs [Notes_unique_PMS]) 3)); |
|
3677
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
paulson
parents:
3676
diff
changeset
|
810 |
(*ClientCertKeyEx*) |
3704 | 811 |
by (fast_tac (*blast_tac gives PROOF FAILED*) |
812 |
(!claset addSEs [PMS_Crypt_sessionK_not_spied RSN (2,rev_notE)]) 2); |
|
3480
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
813 |
(*Fake: the Spy doesn't have the critical session key!*) |
3683 | 814 |
by (subgoal_tac "Key (clientK(Na,Nb,PRF(PMS,NA,NB))) ~: analz(spies evsa)" 1); |
3672
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
815 |
by (asm_simp_tac (!simpset addsimps [Spy_not_see_MS, |
3480
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
paulson
parents:
3474
diff
changeset
|
816 |
not_parts_not_analz]) 2); |
3474 | 817 |
by (Fake_parts_insert_tac 1); |
3704 | 818 |
val lemma = normalize_thm [RSspec, RSmp] (result()); |
819 |
||
820 |
(*Final version*) |
|
821 |
goal thy |
|
822 |
"!!evs. [| Crypt (clientK(Na,Nb,PRF(PMS,NA,NB))) Y : parts (spies evs); \ |
|
823 |
\ Notes A {|Agent B, Nonce PMS|} : set evs; \ |
|
824 |
\ Says A Spy (Key(clientK(Na,Nb,PRF(PMS,NA,NB)))) ~: set evs; \ |
|
825 |
\ evs : tls; A ~: bad; B ~: bad |] \ |
|
826 |
\ ==> Says A B (Crypt (clientK(Na,Nb,PRF(PMS,NA,NB))) Y) : set evs"; |
|
827 |
by (blast_tac (!claset addIs [lemma] |
|
828 |
addEs [clientK_Oops_ALL RSN(2, rev_notE)]) 1); |
|
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
829 |
qed_spec_mp "TrustClientMsg"; |
3506 | 830 |
|
831 |
||
3685
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
832 |
|
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
833 |
(*** Protocol goal: if B receives ClientFinished, and if B is able to |
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
paulson
parents:
3683
diff
changeset
|
834 |
check a CertVerify from A, then A has used the quoted |
3729
6be7cf5086ab
Renamed XA, XB to PA, PB and removed the certificate from Client Verify
paulson
parents:
3711
diff
changeset
|
835 |
values PA, PB, etc. Even this one requires A to be uncompromised. |
3506 | 836 |
***) |
837 |
goal thy |
|
3704 | 838 |
"!!evs. [| Says A Spy (Key(clientK(Na,Nb,PRF(PMS,NA,NB)))) ~: set evs;\ |
3686
4b484805b4c4
First working version with Oops event for session keys
paulson
parents:
3685
diff
changeset
|
839 |
\ Says A' B (Crypt (clientK(Na,Nb,PRF(PMS,NA,NB))) Y) : set evs; \ |
3729
6be7cf5086ab
Renamed XA, XB to PA, PB and removed the certificate from Client Verify
paulson
parents:
3711
diff
changeset
|
840 |
\ Says B A {|Nonce NB, Number SID, Number PB, certificate B KB|} \ |
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
841 |
\ : set evs; \ |
3729
6be7cf5086ab
Renamed XA, XB to PA, PB and removed the certificate from Client Verify
paulson
parents:
3711
diff
changeset
|
842 |
\ Says A'' B (Crypt (priK A) (Hash{|Nonce NB, Agent B, Nonce PMS|}))\ |
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
843 |
\ : set evs; \ |
3683 | 844 |
\ evs : tls; A ~: bad; B ~: bad |] \ |
3672
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
paulson
parents:
3519
diff
changeset
|
845 |
\ ==> Says A B (Crypt (clientK(Na,Nb,PRF(PMS,NA,NB))) Y) : set evs"; |
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
846 |
by (blast_tac (!claset addSIs [TrustClientMsg, UseCertVerify] |
3683 | 847 |
addDs [Says_imp_spies RS parts.Inj]) 1); |
3515
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
paulson
parents:
3506
diff
changeset
|
848 |
qed "AuthClientFinished"; |
3687
fb7d096d7884
Simplified SpyKeys to use sessionK instead of clientK and serverK
paulson
parents:
3686
diff
changeset
|
849 |
|
fb7d096d7884
Simplified SpyKeys to use sessionK instead of clientK and serverK
paulson
parents:
3686
diff
changeset
|
850 |
(*22/9/97: loads in 622s, which is 10 minutes 22 seconds*) |
3711 | 851 |
(*24/9/97: loads in 672s, which is 11 minutes 12 seconds [stronger theorems]*) |
3704 | 852 |
|
853 |