| author | wenzelm | 
| Fri, 06 Nov 1998 14:04:54 +0100 | |
| changeset 5807 | bd2d9dd34dfd | 
| parent 5653 | 204083e3f368 | 
| child 6284 | 147db42c1009 | 
| 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: 
3474diff
changeset | 6 | Protocol goals: | 
| 
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
 paulson parents: 
3474diff
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: 
3474diff
changeset | 8 | parties (though A is not necessarily authenticated). | 
| 
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
 paulson parents: 
3474diff
changeset | 9 | |
| 3685 
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
 paulson parents: 
3683diff
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: 
3474diff
changeset | 11 | message is optional!) | 
| 3474 | 12 | |
| 3685 
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
 paulson parents: 
3683diff
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: 
3474diff
changeset | 14 | |
| 
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
 paulson parents: 
3474diff
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: 
3711diff
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: 
3474diff
changeset | 17 | rollback attacks). | 
| 3474 | 18 | *) | 
| 19 | ||
| 5433 | 20 | AddEs spies_partsEs; | 
| 21 | AddDs [impOfSubs analz_subset_parts]; | |
| 22 | AddDs [impOfSubs Fake_parts_insert]; | |
| 23 | ||
| 3772 | 24 | (*Automatically unfold the definition of "certificate"*) | 
| 25 | Addsimps [certificate_def]; | |
| 3480 
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
 paulson parents: 
3474diff
changeset | 26 | |
| 3474 | 27 | (*Injectiveness of key-generating functions*) | 
| 3677 
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
 paulson parents: 
3676diff
changeset | 28 | AddIffs [inj_PRF RS inj_eq, inj_sessionK RS inj_eq]; | 
| 3474 | 29 | |
| 3677 
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
 paulson parents: 
3676diff
changeset | 30 | (* invKey(sessionK x) = sessionK x*) | 
| 
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
 paulson parents: 
3676diff
changeset | 31 | Addsimps [isSym_sessionK, rewrite_rule [isSymKey_def] isSym_sessionK]; | 
| 3480 
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
 paulson parents: 
3474diff
changeset | 32 | |
| 3474 | 33 | |
| 34 | (*** clientK and serverK make symmetric keys; no clashes with pubK or priK ***) | |
| 35 | ||
| 5076 | 36 | Goal "pubK A ~= sessionK arg"; | 
| 4423 | 37 | by (rtac notI 1); | 
| 3474 | 38 | by (dres_inst_tac [("f","isSymKey")] arg_cong 1);
 | 
| 39 | by (Full_simp_tac 1); | |
| 3677 
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
 paulson parents: 
3676diff
changeset | 40 | qed "pubK_neq_sessionK"; | 
| 3474 | 41 | |
| 5076 | 42 | Goal "priK A ~= sessionK arg"; | 
| 4423 | 43 | by (rtac notI 1); | 
| 3474 | 44 | by (dres_inst_tac [("f","isSymKey")] arg_cong 1);
 | 
| 45 | by (Full_simp_tac 1); | |
| 3677 
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
 paulson parents: 
3676diff
changeset | 46 | qed "priK_neq_sessionK"; | 
| 3474 | 47 | |
| 3677 
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
 paulson parents: 
3676diff
changeset | 48 | 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: 
3506diff
changeset | 49 | AddIffs (keys_distinct @ (keys_distinct RL [not_sym])); | 
| 3474 | 50 | |
| 51 | ||
| 52 | (**** Protocol Proofs ****) | |
| 53 | ||
| 3772 | 54 | (*Possibility properties state that some traces run the protocol to the end. | 
| 55 | Four paths and 12 rules are considered.*) | |
| 3474 | 56 | |
| 3672 
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
 paulson parents: 
3519diff
changeset | 57 | |
| 3772 | 58 | (** These proofs assume that the Nonce_supply nonces | 
| 3672 
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
 paulson parents: 
3519diff
changeset | 59 | (which have the form @ N. Nonce N ~: used evs) | 
| 3772 | 60 | lie outside the range of PRF. It seems reasonable, but as it is needed | 
| 61 | only for the possibility theorems, it is not taken as an axiom. | |
| 3672 
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
 paulson parents: 
3519diff
changeset | 62 | **) | 
| 
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
 paulson parents: 
3519diff
changeset | 63 | |
| 
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
 paulson parents: 
3519diff
changeset | 64 | |
| 3685 
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
 paulson parents: 
3683diff
changeset | 65 | (*Possibility property ending with ClientAccepts.*) | 
| 5114 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 66 | Goal "[| ALL evs. (@ N. Nonce N ~: used evs) ~: range PRF; \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 67 | \ A ~= B |] \ | 
| 5433 | 68 | \ ==> EX SID M. EX evs: tls. \ | 
| 69 | \          Notes A {|Number SID, Agent A, Agent B, Nonce M|} : set evs";
 | |
| 3474 | 70 | by (REPEAT (resolve_tac [exI,bexI] 1)); | 
| 3745 
4c5d3b1ddc75
Client, Server certificates now sent using the separate Certificate rule,
 paulson parents: 
3729diff
changeset | 71 | by (rtac (tls.Nil RS tls.ClientHello RS tls.ServerHello RS tls.Certificate RS | 
| 
4c5d3b1ddc75
Client, Server certificates now sent using the separate Certificate rule,
 paulson parents: 
3729diff
changeset | 72 | tls.ClientKeyExch RS tls.ClientFinished RS tls.ServerFinished RS | 
| 
4c5d3b1ddc75
Client, Server certificates now sent using the separate Certificate rule,
 paulson parents: 
3729diff
changeset | 73 | tls.ClientAccepts) 2); | 
| 3474 | 74 | by possibility_tac; | 
| 3672 
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
 paulson parents: 
3519diff
changeset | 75 | by (REPEAT (Blast_tac 1)); | 
| 3474 | 76 | result(); | 
| 77 | ||
| 3685 
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
 paulson parents: 
3683diff
changeset | 78 | (*And one for ServerAccepts. Either FINISHED message may come first.*) | 
| 5114 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 79 | Goal "[| ALL evs. (@ N. Nonce N ~: used evs) ~: range PRF; \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 80 | \ A ~= B |] \ | 
| 5433 | 81 | \ ==> EX SID NA PA NB PB M. EX evs: tls. \ | 
| 82 | \          Notes B {|Number SID, Agent A, Agent B, Nonce M|} : set evs";
 | |
| 3474 | 83 | by (REPEAT (resolve_tac [exI,bexI] 1)); | 
| 3745 
4c5d3b1ddc75
Client, Server certificates now sent using the separate Certificate rule,
 paulson parents: 
3729diff
changeset | 84 | by (rtac (tls.Nil RS tls.ClientHello RS tls.ServerHello RS tls.Certificate RS | 
| 
4c5d3b1ddc75
Client, Server certificates now sent using the separate Certificate rule,
 paulson parents: 
3729diff
changeset | 85 | tls.ClientKeyExch RS tls.ServerFinished RS tls.ClientFinished RS | 
| 
4c5d3b1ddc75
Client, Server certificates now sent using the separate Certificate rule,
 paulson parents: 
3729diff
changeset | 86 | tls.ServerAccepts) 2); | 
| 3474 | 87 | by possibility_tac; | 
| 3672 
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
 paulson parents: 
3519diff
changeset | 88 | by (REPEAT (Blast_tac 1)); | 
| 3474 | 89 | result(); | 
| 90 | ||
| 91 | (*Another one, for CertVerify (which is optional)*) | |
| 5114 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 92 | Goal "[| ALL evs. (@ N. Nonce N ~: used evs) ~: range PRF; \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 93 | \ A ~= B |] \ | 
| 3760 
77f71f650433
Strengthened the possibility property for resumption so that it could have
 paulson parents: 
3758diff
changeset | 94 | \ ==> EX NB PMS. EX evs: tls. \ | 
| 3729 
6be7cf5086ab
Renamed XA, XB to PA, PB and removed the certificate from Client Verify
 paulson parents: 
3711diff
changeset | 95 | \  Says A B (Crypt (priK A) (Hash{|Nonce NB, Agent B, Nonce PMS|})) : set evs";
 | 
| 3474 | 96 | by (REPEAT (resolve_tac [exI,bexI] 1)); | 
| 3745 
4c5d3b1ddc75
Client, Server certificates now sent using the separate Certificate rule,
 paulson parents: 
3729diff
changeset | 97 | by (rtac (tls.Nil RS tls.ClientHello RS tls.ServerHello RS tls.Certificate RS | 
| 
4c5d3b1ddc75
Client, Server certificates now sent using the separate Certificate rule,
 paulson parents: 
3729diff
changeset | 98 | tls.ClientKeyExch RS tls.CertVerify) 2); | 
| 3474 | 99 | by possibility_tac; | 
| 3672 
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
 paulson parents: 
3519diff
changeset | 100 | by (REPEAT (Blast_tac 1)); | 
| 3474 | 101 | result(); | 
| 102 | ||
| 3685 
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
 paulson parents: 
3683diff
changeset | 103 | (*Another one, for session resumption (both ServerResume and ClientResume) *) | 
| 5114 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 104 | Goal "[| evs0 : tls; \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 105 | \        Notes A {|Number SID, Agent A, Agent B, Nonce M|} : set evs0; \
 | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 106 | \        Notes B {|Number SID, Agent A, Agent B, Nonce M|} : set evs0; \
 | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 107 | \ ALL evs. (@ N. Nonce N ~: used evs) ~: range PRF; \ | 
| 5433 | 108 | \ A ~= B |] \ | 
| 109 | \ ==> EX NA PA NB PB X. EX evs: tls. \ | |
| 110 | \           X = Hash{|Number SID, Nonce M,             \
 | |
| 111 | \ Nonce NA, Number PA, Agent A, \ | |
| 112 | \ Nonce NB, Number PB, Agent B|} & \ | |
| 113 | \ Says A B (Crypt (clientK(NA,NB,M)) X) : set evs & \ | |
| 114 | \ Says B A (Crypt (serverK(NA,NB,M)) X) : set evs"; | |
| 3685 
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
 paulson parents: 
3683diff
changeset | 115 | by (REPEAT (resolve_tac [exI,bexI] 1)); | 
| 3745 
4c5d3b1ddc75
Client, Server certificates now sent using the separate Certificate rule,
 paulson parents: 
3729diff
changeset | 116 | by (etac (tls.ClientHello RS tls.ServerHello RS tls.ServerResume RS | 
| 
4c5d3b1ddc75
Client, Server certificates now sent using the separate Certificate rule,
 paulson parents: 
3729diff
changeset | 117 | tls.ClientResume) 2); | 
| 3685 
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
 paulson parents: 
3683diff
changeset | 118 | by possibility_tac; | 
| 
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
 paulson parents: 
3683diff
changeset | 119 | by (REPEAT (Blast_tac 1)); | 
| 
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
 paulson parents: 
3683diff
changeset | 120 | result(); | 
| 
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
 paulson parents: 
3683diff
changeset | 121 | |
| 
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
 paulson parents: 
3683diff
changeset | 122 | |
| 3474 | 123 | |
| 124 | (**** Inductive proofs about tls ****) | |
| 125 | ||
| 126 | ||
| 3519 
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
 paulson parents: 
3515diff
changeset | 127 | (*Induction for regularity theorems. If induction formula has the form | 
| 3683 | 128 | X ~: analz (spies evs) --> ... then it shortens the proof by discarding | 
| 129 | needless information about analz (insert X (spies evs)) *) | |
| 3519 
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
 paulson parents: 
3515diff
changeset | 130 | fun parts_induct_tac i = | 
| 
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
 paulson parents: 
3515diff
changeset | 131 | etac tls.induct i | 
| 
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
 paulson parents: 
3515diff
changeset | 132 | THEN | 
| 
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
 paulson parents: 
3515diff
changeset | 133 | REPEAT (FIRSTGOAL analz_mono_contra_tac) | 
| 
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
 paulson parents: 
3515diff
changeset | 134 | THEN | 
| 4091 | 135 | fast_tac (claset() addss (simpset())) i THEN | 
| 4686 | 136 | ALLGOALS Asm_simp_tac; | 
| 3519 
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
 paulson parents: 
3515diff
changeset | 137 | |
| 
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
 paulson parents: 
3515diff
changeset | 138 | |
| 3683 | 139 | (** Theorems of the form X ~: parts (spies evs) imply that NOBODY | 
| 3474 | 140 | sends messages containing X! **) | 
| 141 | ||
| 3683 | 142 | (*Spy never sees another agent's private key! (unless it's bad at start)*) | 
| 5114 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 143 | Goal "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: 
3515diff
changeset | 144 | by (parts_induct_tac 1); | 
| 5433 | 145 | by (Blast_tac 1); | 
| 3474 | 146 | qed "Spy_see_priK"; | 
| 147 | Addsimps [Spy_see_priK]; | |
| 148 | ||
| 5114 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 149 | Goal "evs : tls ==> (Key (priK A) : analz (spies evs)) = (A : bad)"; | 
| 5433 | 150 | by Auto_tac; | 
| 3474 | 151 | qed "Spy_analz_priK"; | 
| 152 | Addsimps [Spy_analz_priK]; | |
| 153 | ||
| 4472 | 154 | AddSDs [Spy_see_priK RSN (2, rev_iffD1), | 
| 155 | Spy_analz_priK RSN (2, rev_iffD1)]; | |
| 3474 | 156 | |
| 157 | ||
| 3515 
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
 paulson parents: 
3506diff
changeset | 158 | (*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: 
3515diff
changeset | 159 | 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: 
3506diff
changeset | 160 | 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: 
3506diff
changeset | 161 | breach of security.*) | 
| 5076 | 162 | Goalw [certificate_def] | 
| 5359 | 163 | "[| certificate B KB : parts (spies evs); evs : tls |] ==> pubK B = KB"; | 
| 3772 | 164 | by (etac rev_mp 1); | 
| 3519 
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
 paulson parents: 
3515diff
changeset | 165 | by (parts_induct_tac 1); | 
| 5433 | 166 | by (Blast_tac 1); | 
| 3772 | 167 | qed "certificate_valid"; | 
| 3515 
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
 paulson parents: 
3506diff
changeset | 168 | |
| 
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
 paulson parents: 
3506diff
changeset | 169 | |
| 3745 
4c5d3b1ddc75
Client, Server certificates now sent using the separate Certificate rule,
 paulson parents: 
3729diff
changeset | 170 | (*Replace key KB in ClientKeyExch by (pubK B) *) | 
| 
4c5d3b1ddc75
Client, Server certificates now sent using the separate Certificate rule,
 paulson parents: 
3729diff
changeset | 171 | val ClientKeyExch_tac = | 
| 3772 | 172 | forward_tac [Says_imp_spies RS parts.Inj RS certificate_valid] | 
| 3515 
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
 paulson parents: 
3506diff
changeset | 173 | THEN' assume_tac | 
| 
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
 paulson parents: 
3506diff
changeset | 174 | THEN' hyp_subst_tac; | 
| 
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
 paulson parents: 
3506diff
changeset | 175 | |
| 
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
 paulson parents: 
3506diff
changeset | 176 | fun analz_induct_tac i = | 
| 
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
 paulson parents: 
3506diff
changeset | 177 | etac tls.induct i THEN | 
| 3745 
4c5d3b1ddc75
Client, Server certificates now sent using the separate Certificate rule,
 paulson parents: 
3729diff
changeset | 178 | ClientKeyExch_tac (i+6) THEN (*ClientKeyExch*) | 
| 3515 
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
 paulson parents: 
3506diff
changeset | 179 | ALLGOALS (asm_simp_tac | 
| 4091 | 180 | (simpset() addcongs [if_weak_cong] | 
| 5535 | 181 | addsimps split_ifs @ pushes)) THEN | 
| 3515 
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
 paulson parents: 
3506diff
changeset | 182 | (*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: 
3506diff
changeset | 183 | 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: 
3506diff
changeset | 184 | ALLGOALS (asm_simp_tac | 
| 4091 | 185 | (simpset() addcongs [if_weak_cong] | 
| 4686 | 186 | addsimps [insert_absorb])); | 
| 3515 
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
 paulson parents: 
3506diff
changeset | 187 | |
| 
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
 paulson parents: 
3506diff
changeset | 188 | |
| 3758 
188a4fbfaf55
Exchanged the M and SID fields of the FINISHED messages to simplify proofs;
 paulson parents: 
3745diff
changeset | 189 | (*** Properties of items found in Notes ***) | 
| 3515 
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
 paulson parents: 
3506diff
changeset | 190 | |
| 5114 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 191 | Goal "[| Notes A {|Agent B, X|} : set evs;  evs : tls |]  \
 | 
| 5359 | 192 | \ ==> Crypt (pubK B) X : parts (spies evs)"; | 
| 3515 
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
 paulson parents: 
3506diff
changeset | 193 | by (etac rev_mp 1); | 
| 
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
 paulson parents: 
3506diff
changeset | 194 | by (analz_induct_tac 1); | 
| 4091 | 195 | by (blast_tac (claset() addIs [parts_insertI]) 1); | 
| 3683 | 196 | qed "Notes_Crypt_parts_spies"; | 
| 3515 
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
 paulson parents: 
3506diff
changeset | 197 | |
| 3758 
188a4fbfaf55
Exchanged the M and SID fields of the FINISHED messages to simplify proofs;
 paulson parents: 
3745diff
changeset | 198 | (*C may be either A or B*) | 
| 5114 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 199 | Goal "[| Notes C {|s, Agent A, Agent B, Nonce(PRF(PMS,NA,NB))|} : set evs; \
 | 
| 5359 | 200 | \ evs : tls |] \ | 
| 201 | \ ==> Crypt (pubK B) (Nonce PMS) : parts (spies evs)"; | |
| 3685 
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
 paulson parents: 
3683diff
changeset | 202 | by (etac rev_mp 1); | 
| 
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
 paulson parents: 
3683diff
changeset | 203 | by (parts_induct_tac 1); | 
| 3711 | 204 | by (ALLGOALS Clarify_tac); | 
| 3685 
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
 paulson parents: 
3683diff
changeset | 205 | (*Fake*) | 
| 4091 | 206 | by (blast_tac (claset() addIs [parts_insertI]) 1); | 
| 3685 
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
 paulson parents: 
3683diff
changeset | 207 | (*Client, Server Accept*) | 
| 5433 | 208 | by (REPEAT (blast_tac (claset() addSDs [Notes_Crypt_parts_spies]) 1)); | 
| 3685 
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
 paulson parents: 
3683diff
changeset | 209 | qed "Notes_master_imp_Crypt_PMS"; | 
| 
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
 paulson parents: 
3683diff
changeset | 210 | |
| 
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
 paulson parents: 
3683diff
changeset | 211 | (*Compared with the theorem above, both premise and conclusion are stronger*) | 
| 5114 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 212 | Goal "[| Notes A {|s, Agent A, Agent B, Nonce(PRF(PMS,NA,NB))|} : set evs;\
 | 
| 5359 | 213 | \ evs : tls |] \ | 
| 214 | \     ==> Notes A {|Agent B, Nonce PMS|} : set evs";
 | |
| 3685 
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
 paulson parents: 
3683diff
changeset | 215 | by (etac rev_mp 1); | 
| 
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
 paulson parents: 
3683diff
changeset | 216 | by (parts_induct_tac 1); | 
| 
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
 paulson parents: 
3683diff
changeset | 217 | (*ServerAccepts*) | 
| 
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
 paulson parents: 
3683diff
changeset | 218 | by (Fast_tac 1); (*Blast_tac's very slow here*) | 
| 
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
 paulson parents: 
3683diff
changeset | 219 | qed "Notes_master_imp_Notes_PMS"; | 
| 
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
 paulson parents: 
3683diff
changeset | 220 | |
| 
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
 paulson parents: 
3683diff
changeset | 221 | |
| 
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
 paulson parents: 
3683diff
changeset | 222 | (*** Protocol goal: if B receives CertVerify, then A sent it ***) | 
| 3474 | 223 | |
| 3745 
4c5d3b1ddc75
Client, Server certificates now sent using the separate Certificate rule,
 paulson parents: 
3729diff
changeset | 224 | (*B can check A's signature if he has received A's certificate.*) | 
| 5114 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 225 | Goal "[| X : parts (spies evs); \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 226 | \        X = Crypt (priK A) (Hash{|nb, Agent B, pms|});  \
 | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 227 | \ evs : tls; A ~: bad |] \ | 
| 5359 | 228 | \ ==> Says A B X : set evs"; | 
| 3745 
4c5d3b1ddc75
Client, Server certificates now sent using the separate Certificate rule,
 paulson parents: 
3729diff
changeset | 229 | by (etac rev_mp 1); | 
| 3480 
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
 paulson parents: 
3474diff
changeset | 230 | by (hyp_subst_tac 1); | 
| 3519 
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
 paulson parents: 
3515diff
changeset | 231 | by (parts_induct_tac 1); | 
| 5433 | 232 | by (Blast_tac 1); | 
| 3745 
4c5d3b1ddc75
Client, Server certificates now sent using the separate Certificate rule,
 paulson parents: 
3729diff
changeset | 233 | val lemma = result(); | 
| 
4c5d3b1ddc75
Client, Server certificates now sent using the separate Certificate rule,
 paulson parents: 
3729diff
changeset | 234 | |
| 
4c5d3b1ddc75
Client, Server certificates now sent using the separate Certificate rule,
 paulson parents: 
3729diff
changeset | 235 | (*Final version: B checks X using the distributed KA instead of priK A*) | 
| 5114 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 236 | Goal "[| X : parts (spies evs); \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 237 | \        X = Crypt (invKey KA) (Hash{|nb, Agent B, pms|}); \
 | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 238 | \ certificate A KA : parts (spies evs); \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 239 | \ evs : tls; A ~: bad |] \ | 
| 5359 | 240 | \ ==> Says A B X : set evs"; | 
| 4091 | 241 | by (blast_tac (claset() addSDs [certificate_valid] addSIs [lemma]) 1); | 
| 3745 
4c5d3b1ddc75
Client, Server certificates now sent using the separate Certificate rule,
 paulson parents: 
3729diff
changeset | 242 | qed "TrustCertVerify"; | 
| 3474 | 243 | |
| 244 | ||
| 3685 
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
 paulson parents: 
3683diff
changeset | 245 | (*If CertVerify is present then A has chosen PMS.*) | 
| 5114 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 246 | Goal "[| Crypt (priK A) (Hash{|nb, Agent B, Nonce PMS|}) \
 | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 247 | \ : parts (spies evs); \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 248 | \ evs : tls; A ~: bad |] \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 249 | \     ==> Notes A {|Agent B, Nonce PMS|} : set evs";
 | 
| 4423 | 250 | by (etac rev_mp 1); | 
| 3519 
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
 paulson parents: 
3515diff
changeset | 251 | by (parts_induct_tac 1); | 
| 5433 | 252 | by (Blast_tac 1); | 
| 3745 
4c5d3b1ddc75
Client, Server certificates now sent using the separate Certificate rule,
 paulson parents: 
3729diff
changeset | 253 | val lemma = result(); | 
| 
4c5d3b1ddc75
Client, Server certificates now sent using the separate Certificate rule,
 paulson parents: 
3729diff
changeset | 254 | |
| 
4c5d3b1ddc75
Client, Server certificates now sent using the separate Certificate rule,
 paulson parents: 
3729diff
changeset | 255 | (*Final version using the distributed KA instead of priK A*) | 
| 5114 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 256 | Goal "[| Crypt (invKey KA) (Hash{|nb, Agent B, Nonce PMS|}) \
 | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 257 | \ : parts (spies evs); \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 258 | \ certificate A KA : parts (spies evs); \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 259 | \ evs : tls; A ~: bad |] \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 260 | \     ==> Notes A {|Agent B, Nonce PMS|} : set evs";
 | 
| 4091 | 261 | by (blast_tac (claset() addSDs [certificate_valid] addSIs [lemma]) 1); | 
| 3515 
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
 paulson parents: 
3506diff
changeset | 262 | qed "UseCertVerify"; | 
| 3474 | 263 | |
| 3480 
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
 paulson parents: 
3474diff
changeset | 264 | |
| 5114 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 265 | Goal "evs : tls ==> Notes A {|Agent B, Nonce (PRF x)|} ~: set evs";
 | 
| 3672 
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
 paulson parents: 
3519diff
changeset | 266 | by (parts_induct_tac 1); | 
| 3745 
4c5d3b1ddc75
Client, Server certificates now sent using the separate Certificate rule,
 paulson parents: 
3729diff
changeset | 267 | (*ClientKeyExch: PMS is assumed to differ from any PRF.*) | 
| 3672 
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
 paulson parents: 
3519diff
changeset | 268 | by (Blast_tac 1); | 
| 
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
 paulson parents: 
3519diff
changeset | 269 | qed "no_Notes_A_PRF"; | 
| 
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
 paulson parents: 
3519diff
changeset | 270 | Addsimps [no_Notes_A_PRF]; | 
| 
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
 paulson parents: 
3519diff
changeset | 271 | |
| 
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
 paulson parents: 
3519diff
changeset | 272 | |
| 5359 | 273 | Goal "[| Nonce (PRF (PMS,NA,NB)) : parts (spies evs); evs : tls |] \ | 
| 274 | \ ==> Nonce PMS : parts (spies evs)"; | |
| 3672 
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
 paulson parents: 
3519diff
changeset | 275 | by (etac rev_mp 1); | 
| 
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
 paulson parents: 
3519diff
changeset | 276 | by (parts_induct_tac 1); | 
| 5433 | 277 | (*Easy, e.g. by freshness*) | 
| 278 | by (REPEAT (blast_tac (claset() addDs [Notes_Crypt_parts_spies]) 2)); | |
| 279 | (*Fake*) | |
| 280 | by (blast_tac (claset() addIs [parts_insertI]) 1); | |
| 3672 
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
 paulson parents: 
3519diff
changeset | 281 | qed "MS_imp_PMS"; | 
| 
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
 paulson parents: 
3519diff
changeset | 282 | AddSDs [MS_imp_PMS]; | 
| 
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
 paulson parents: 
3519diff
changeset | 283 | |
| 
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
 paulson parents: 
3519diff
changeset | 284 | |
| 3474 | 285 | |
| 3672 
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
 paulson parents: 
3519diff
changeset | 286 | (*** 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: 
3506diff
changeset | 287 | |
| 3672 
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
 paulson parents: 
3519diff
changeset | 288 | (*PMS determines B. Proof borrowed from NS_Public/unique_NA and from Yahalom*) | 
| 5114 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 289 | Goal "[| Nonce PMS ~: analz (spies evs); evs : tls |] \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 290 | \ ==> EX B'. ALL B. \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 291 | \ 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: 
3506diff
changeset | 292 | by (etac rev_mp 1); | 
| 3519 
ab0a9fbed4c0
Changing "lost" from a parameter of protocol definitions to a constant.
 paulson parents: 
3515diff
changeset | 293 | by (parts_induct_tac 1); | 
| 5433 | 294 | by (Blast_tac 1); | 
| 3745 
4c5d3b1ddc75
Client, Server certificates now sent using the separate Certificate rule,
 paulson parents: 
3729diff
changeset | 295 | (*ClientKeyExch*) | 
| 
4c5d3b1ddc75
Client, Server certificates now sent using the separate Certificate rule,
 paulson parents: 
3729diff
changeset | 296 | by (ClientKeyExch_tac 1); | 
| 4091 | 297 | 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: 
3519diff
changeset | 298 | by (expand_case_tac "PMS = ?y" 1 THEN | 
| 4091 | 299 | blast_tac (claset() addSEs partsEs) 1); | 
| 3515 
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
 paulson parents: 
3506diff
changeset | 300 | val lemma = result(); | 
| 
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
 paulson parents: 
3506diff
changeset | 301 | |
| 5114 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 302 | Goal "[| Crypt(pubK B) (Nonce PMS) : parts (spies evs); \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 303 | \ Crypt(pubK B') (Nonce PMS) : parts (spies evs); \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 304 | \ Nonce PMS ~: analz (spies evs); \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 305 | \ evs : tls |] \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 306 | \ ==> B=B'"; | 
| 3515 
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
 paulson parents: 
3506diff
changeset | 307 | by (prove_unique_tac lemma 1); | 
| 3704 | 308 | qed "Crypt_unique_PMS"; | 
| 309 | ||
| 3515 
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
 paulson parents: 
3506diff
changeset | 310 | |
| 3685 
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
 paulson parents: 
3683diff
changeset | 311 | (** It is frustrating that we need two versions of the unicity results. | 
| 3704 | 312 |     But Notes A {|Agent B, Nonce PMS|} determines both A and B.  Sometimes
 | 
| 313 | we have only the weaker assertion Crypt(pubK B) (Nonce PMS), which | |
| 314 | determines B alone, and only if PMS is secret. | |
| 3685 
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
 paulson parents: 
3683diff
changeset | 315 | **) | 
| 3515 
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
 paulson parents: 
3506diff
changeset | 316 | |
| 3677 
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
 paulson parents: 
3676diff
changeset | 317 | (*In A's internal Note, PMS determines A and B.*) | 
| 5359 | 318 | Goal "evs : tls ==> EX A' B'. ALL A B. \ | 
| 319 | \                   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: 
3515diff
changeset | 320 | by (parts_induct_tac 1); | 
| 4091 | 321 | by (asm_simp_tac (simpset() addsimps [all_conj_distrib]) 1); | 
| 3745 
4c5d3b1ddc75
Client, Server certificates now sent using the separate Certificate rule,
 paulson parents: 
3729diff
changeset | 322 | (*ClientKeyExch: if PMS is fresh, then it can't appear in Notes A X.*) | 
| 3672 
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
 paulson parents: 
3519diff
changeset | 323 | by (expand_case_tac "PMS = ?y" 1 THEN | 
| 4091 | 324 | 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: 
3506diff
changeset | 325 | val lemma = result(); | 
| 
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
 paulson parents: 
3506diff
changeset | 326 | |
| 5114 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 327 | Goal "[| Notes A  {|Agent B,  Nonce PMS|} : set evs;  \
 | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 328 | \        Notes A' {|Agent B', Nonce PMS|} : set evs;  \
 | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 329 | \ evs : tls |] \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 330 | \ ==> A=A' & B=B'"; | 
| 3515 
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
 paulson parents: 
3506diff
changeset | 331 | by (prove_unique_tac lemma 1); | 
| 3672 
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
 paulson parents: 
3519diff
changeset | 332 | qed "Notes_unique_PMS"; | 
| 3515 
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
 paulson parents: 
3506diff
changeset | 333 | |
| 
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
 paulson parents: 
3506diff
changeset | 334 | |
| 3474 | 335 | |
| 3772 | 336 | (**** Secrecy Theorems ****) | 
| 337 | ||
| 338 | (*Key compromise lemma needed to prove analz_image_keys. | |
| 339 | No collection of keys can help the spy get new private keys.*) | |
| 5359 | 340 | Goal "evs : tls \ | 
| 341 | \ ==> ALL KK. (Key(priK B) : analz (Key``KK Un (spies evs))) = \ | |
| 342 | \ (priK B : KK | B : bad)"; | |
| 3772 | 343 | by (etac tls.induct 1); | 
| 344 | by (ALLGOALS | |
| 345 | (asm_simp_tac (analz_image_keys_ss | |
| 5535 | 346 | addsimps certificate_def::keys_distinct))); | 
| 3772 | 347 | (*Fake*) | 
| 4422 
21238c9d363e
Simplified proofs using rewrites for f``A where f is injective
 paulson parents: 
4201diff
changeset | 348 | by (spy_analz_tac 1); | 
| 3772 | 349 | qed_spec_mp "analz_image_priK"; | 
| 350 | ||
| 351 | ||
| 352 | (*slightly speeds up the big simplification below*) | |
| 5114 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 353 | Goal "KK <= range sessionK ==> priK B ~: KK"; | 
| 3772 | 354 | by (Blast_tac 1); | 
| 355 | val range_sessionkeys_not_priK = result(); | |
| 356 | ||
| 357 | (*Lemma for the trivial direction of the if-and-only-if*) | |
| 5114 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 358 | Goal "(X : analz (G Un H)) --> (X : analz H) ==> \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 359 | \ (X : analz (G Un H)) = (X : analz H)"; | 
| 4091 | 360 | by (blast_tac (claset() addIs [impOfSubs analz_mono]) 1); | 
| 3961 | 361 | val analz_image_keys_lemma = result(); | 
| 3772 | 362 | |
| 363 | (** Strangely, the following version doesn't work: | |
| 5114 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 364 | \ ALL Z. (Nonce N : analz (Key``(sessionK``Z) Un (spies evs))) = \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 365 | \ (Nonce N : analz (spies evs))"; | 
| 3772 | 366 | **) | 
| 367 | ||
| 5114 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 368 | Goal "evs : tls ==> \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 369 | \ ALL KK. KK <= range sessionK --> \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 370 | \ (Nonce N : analz (Key``KK Un (spies evs))) = \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 371 | \ (Nonce N : analz (spies evs))"; | 
| 3772 | 372 | by (etac tls.induct 1); | 
| 373 | by (ClientKeyExch_tac 7); | |
| 374 | by (REPEAT_FIRST (resolve_tac [allI, impI])); | |
| 3961 | 375 | by (REPEAT_FIRST (rtac analz_image_keys_lemma)); | 
| 5076 | 376 | by (ALLGOALS (*4.5 seconds*) | 
| 3772 | 377 | (asm_simp_tac (analz_image_keys_ss | 
| 5535 | 378 | addsimps split_ifs @ pushes @ | 
| 379 | [range_sessionkeys_not_priK, | |
| 380 | analz_image_priK, certificate_def]))); | |
| 4091 | 381 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [insert_absorb]))); | 
| 3772 | 382 | (*Fake*) | 
| 4422 
21238c9d363e
Simplified proofs using rewrites for f``A where f is injective
 paulson parents: 
4201diff
changeset | 383 | by (spy_analz_tac 1); | 
| 3772 | 384 | qed_spec_mp "analz_image_keys"; | 
| 385 | ||
| 386 | (*Knowing some session keys is no help in getting new nonces*) | |
| 5114 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 387 | Goal "evs : tls ==> \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 388 | \ Nonce N : analz (insert (Key (sessionK z)) (spies evs)) = \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 389 | \ (Nonce N : analz (spies evs))"; | 
| 3772 | 390 | by (asm_simp_tac (analz_image_keys_ss addsimps [analz_image_keys]) 1); | 
| 391 | qed "analz_insert_key"; | |
| 392 | Addsimps [analz_insert_key]; | |
| 393 | ||
| 394 | ||
| 395 | (*** Protocol goal: serverK(Na,Nb,M) and clientK(Na,Nb,M) remain secure ***) | |
| 396 | ||
| 397 | (** Some lemmas about session keys, comprising clientK and serverK **) | |
| 398 | ||
| 399 | ||
| 400 | (*Lemma: session keys are never used if PMS is fresh. | |
| 401 | Nonces don't have to agree, allowing session resumption. | |
| 402 | Converse doesn't hold; revealing PMS doesn't force the keys to be sent. | |
| 403 | THEY ARE NOT SUITABLE AS SAFE ELIM RULES.*) | |
| 5114 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 404 | Goal "[| Nonce PMS ~: parts (spies evs); \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 405 | \ K = sessionK((Na, Nb, PRF(PMS,NA,NB)), b); \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 406 | \ evs : tls |] \ | 
| 3772 | 407 | \ ==> Key K ~: parts (spies evs) & (ALL Y. Crypt K Y ~: parts (spies evs))"; | 
| 408 | by (etac rev_mp 1); | |
| 409 | by (hyp_subst_tac 1); | |
| 410 | by (analz_induct_tac 1); | |
| 411 | (*SpyKeys*) | |
| 5433 | 412 | by (Blast_tac 3); | 
| 3772 | 413 | (*Fake*) | 
| 5433 | 414 | by (blast_tac (claset() addIs [parts_insertI]) 2); | 
| 3772 | 415 | (** LEVEL 6 **) | 
| 416 | (*Oops*) | |
| 5433 | 417 | by (Force_tac 6); | 
| 3772 | 418 | by (REPEAT | 
| 4091 | 419 | (blast_tac (claset() addSDs [Notes_Crypt_parts_spies, | 
| 5433 | 420 | Notes_master_imp_Crypt_PMS]) 1)); | 
| 3772 | 421 | val lemma = result(); | 
| 422 | ||
| 5114 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 423 | Goal "[| Key (sessionK((Na, Nb, PRF(PMS,NA,NB)), b)) : parts (spies evs); \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 424 | \ evs : tls |] \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 425 | \ ==> Nonce PMS : parts (spies evs)"; | 
| 4091 | 426 | by (blast_tac (claset() addDs [lemma]) 1); | 
| 3772 | 427 | qed "PMS_sessionK_not_spied"; | 
| 428 | ||
| 5114 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 429 | Goal "[| Crypt (sessionK((Na, Nb, PRF(PMS,NA,NB)), b)) Y \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 430 | \ : parts (spies evs); evs : tls |] \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 431 | \ ==> Nonce PMS : parts (spies evs)"; | 
| 4091 | 432 | by (blast_tac (claset() addDs [lemma]) 1); | 
| 3772 | 433 | qed "PMS_Crypt_sessionK_not_spied"; | 
| 434 | ||
| 5433 | 435 | (*Write keys are never sent if M (MASTER SECRET) is secure. | 
| 436 | Converse fails; betraying M doesn't force the keys to be sent! | |
| 3772 | 437 | The strong Oops condition can be weakened later by unicity reasoning, | 
| 5433 | 438 | with some effort. | 
| 439 | NO LONGER USED: see clientK_not_spied and serverK_not_spied*) | |
| 5114 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 440 | Goal "[| ALL A. Says A Spy (Key (sessionK((NA,NB,M),b))) ~: set evs; \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 441 | \ Nonce M ~: analz (spies evs); evs : tls |] \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 442 | \ ==> Key (sessionK((NA,NB,M),b)) ~: parts (spies evs)"; | 
| 3772 | 443 | by (etac rev_mp 1); | 
| 444 | by (etac rev_mp 1); | |
| 5433 | 445 | by (analz_induct_tac 1); (*7 seconds*) | 
| 3772 | 446 | (*SpyKeys*) | 
| 4091 | 447 | by (blast_tac (claset() addDs [Says_imp_spies RS analz.Inj]) 3); | 
| 3772 | 448 | (*Fake*) | 
| 449 | by (spy_analz_tac 2); | |
| 450 | (*Base*) | |
| 451 | by (Blast_tac 1); | |
| 452 | qed "sessionK_not_spied"; | |
| 453 | ||
| 454 | ||
| 3745 
4c5d3b1ddc75
Client, Server certificates now sent using the separate Certificate rule,
 paulson parents: 
3729diff
changeset | 455 | (*If A sends ClientKeyExch to an honest B, then the PMS will stay secret.*) | 
| 5114 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 456 | Goal "[| evs : tls; A ~: bad; B ~: bad |] \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 457 | \     ==> Notes A {|Agent B, Nonce PMS|} : set evs  -->   \
 | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 458 | \ Nonce PMS ~: analz (spies evs)"; | 
| 3745 
4c5d3b1ddc75
Client, Server certificates now sent using the separate Certificate rule,
 paulson parents: 
3729diff
changeset | 459 | by (analz_induct_tac 1); (*11 seconds*) | 
| 3677 
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
 paulson parents: 
3676diff
changeset | 460 | (*ClientAccepts and ServerAccepts: because PMS ~: range PRF*) | 
| 4091 | 461 | by (REPEAT (fast_tac (claset() addss (simpset())) 6)); | 
| 3745 
4c5d3b1ddc75
Client, Server certificates now sent using the separate Certificate rule,
 paulson parents: 
3729diff
changeset | 462 | (*ClientHello, ServerHello, ClientKeyExch, ServerResume: | 
| 3687 
fb7d096d7884
Simplified SpyKeys to use sessionK instead of clientK and serverK
 paulson parents: 
3686diff
changeset | 463 | mostly freshness reasoning*) | 
| 4091 | 464 | by (REPEAT (blast_tac (claset() addSEs partsEs | 
| 4201 | 465 | addDs [Notes_Crypt_parts_spies, | 
| 466 | Says_imp_spies RS analz.Inj]) 3)); | |
| 3677 
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
 paulson parents: 
3676diff
changeset | 467 | (*SpyKeys*) | 
| 4091 | 468 | by (fast_tac (claset() addss (simpset())) 2); | 
| 3677 
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
 paulson parents: 
3676diff
changeset | 469 | (*Fake*) | 
| 
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
 paulson parents: 
3676diff
changeset | 470 | by (spy_analz_tac 1); | 
| 
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
 paulson parents: 
3676diff
changeset | 471 | 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: 
3676diff
changeset | 472 | |
| 
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
 paulson parents: 
3676diff
changeset | 473 | |
| 3745 
4c5d3b1ddc75
Client, Server certificates now sent using the separate Certificate rule,
 paulson parents: 
3729diff
changeset | 474 | (*If A sends ClientKeyExch to an honest B, then the MASTER SECRET | 
| 3677 
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
 paulson parents: 
3676diff
changeset | 475 | will stay secret.*) | 
| 5114 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 476 | Goal "[| evs : tls; A ~: bad; B ~: bad |] \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 477 | \     ==> Notes A {|Agent B, Nonce PMS|} : set evs  -->   \
 | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 478 | \ Nonce (PRF(PMS,NA,NB)) ~: analz (spies evs)"; | 
| 3745 
4c5d3b1ddc75
Client, Server certificates now sent using the separate Certificate rule,
 paulson parents: 
3729diff
changeset | 479 | by (analz_induct_tac 1); (*13 seconds*) | 
| 3677 
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
 paulson parents: 
3676diff
changeset | 480 | (*ClientAccepts and ServerAccepts: because PMS was already visible*) | 
| 4091 | 481 | by (REPEAT (blast_tac (claset() addDs [Spy_not_see_PMS, | 
| 4201 | 482 | Says_imp_spies RS analz.Inj, | 
| 483 | Notes_imp_spies RS analz.Inj]) 6)); | |
| 3677 
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
 paulson parents: 
3676diff
changeset | 484 | (*ClientHello*) | 
| 
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
 paulson parents: 
3676diff
changeset | 485 | by (Blast_tac 3); | 
| 
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
 paulson parents: 
3676diff
changeset | 486 | (*SpyKeys: by secrecy of the PMS, Spy cannot make the MS*) | 
| 4091 | 487 | by (blast_tac (claset() addSDs [Spy_not_see_PMS, | 
| 4422 
21238c9d363e
Simplified proofs using rewrites for f``A where f is injective
 paulson parents: 
4201diff
changeset | 488 | Says_imp_spies RS analz.Inj]) 2); | 
| 3677 
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
 paulson parents: 
3676diff
changeset | 489 | (*Fake*) | 
| 
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
 paulson parents: 
3676diff
changeset | 490 | by (spy_analz_tac 1); | 
| 3745 
4c5d3b1ddc75
Client, Server certificates now sent using the separate Certificate rule,
 paulson parents: 
3729diff
changeset | 491 | (*ServerHello and ClientKeyExch: mostly freshness reasoning*) | 
| 4091 | 492 | by (REPEAT (blast_tac (claset() addSEs partsEs | 
| 4201 | 493 | addDs [Notes_Crypt_parts_spies, | 
| 494 | Says_imp_spies RS analz.Inj]) 1)); | |
| 3677 
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
 paulson parents: 
3676diff
changeset | 495 | 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: 
3676diff
changeset | 496 | |
| 
f2569416d18b
Now with the sessionK constant and new events ClientAccepts and ServerAccepts
 paulson parents: 
3676diff
changeset | 497 | |
| 3704 | 498 | (*** Weakening the Oops conditions for leakage of clientK ***) | 
| 499 | ||
| 5433 | 500 | (*If A created PMS then nobody else (except the Spy in replays) | 
| 501 | would send a message using a clientK generated from that PMS.*) | |
| 502 | Goal "[| Says A' B' (Crypt (clientK(Na,Nb,PRF(PMS,NA,NB))) Y) : set evs; \ | |
| 503 | \        Notes A {|Agent B, Nonce PMS|} : set evs;   \
 | |
| 504 | \ evs : tls; A' ~= Spy |] \ | |
| 505 | \ ==> A = A'"; | |
| 506 | by (etac rev_mp 1); | |
| 507 | by (etac rev_mp 1); | |
| 3745 
4c5d3b1ddc75
Client, Server certificates now sent using the separate Certificate rule,
 paulson parents: 
3729diff
changeset | 508 | by (analz_induct_tac 1); (*8 seconds*) | 
| 3711 | 509 | by (ALLGOALS Clarify_tac); | 
| 3704 | 510 | (*ClientFinished, ClientResume: by unicity of PMS*) | 
| 511 | by (REPEAT | |
| 4091 | 512 | (blast_tac (claset() addSDs [Notes_master_imp_Notes_PMS] | 
| 4201 | 513 | addIs [Notes_unique_PMS RS conjunct1]) 2)); | 
| 3745 
4c5d3b1ddc75
Client, Server certificates now sent using the separate Certificate rule,
 paulson parents: 
3729diff
changeset | 514 | (*ClientKeyExch*) | 
| 4472 | 515 | by (blast_tac (claset() addSDs [PMS_Crypt_sessionK_not_spied, | 
| 516 | Says_imp_spies RS parts.Inj]) 1); | |
| 5433 | 517 | qed "Says_clientK_unique"; | 
| 3704 | 518 | |
| 519 | ||
| 520 | (*If A created PMS and has not leaked her clientK to the Spy, | |
| 5433 | 521 | then it is completely secure: not even in parts!*) | 
| 522 | Goal "[| Notes A {|Agent B, Nonce PMS|} : set evs;  \
 | |
| 523 | \ Says A Spy (Key (clientK(Na,Nb,PRF(PMS,NA,NB)))) ~: set evs; \ | |
| 524 | \ A ~: bad; B ~: bad; \ | |
| 525 | \ evs : tls |] \ | |
| 526 | \ ==> Key (clientK(Na,Nb,PRF(PMS,NA,NB))) ~: parts (spies evs)"; | |
| 527 | by (etac rev_mp 1); | |
| 528 | by (etac rev_mp 1); | |
| 529 | by (analz_induct_tac 1); (*17 seconds*) | |
| 3704 | 530 | (*Oops*) | 
| 5433 | 531 | by (blast_tac (claset() addIs [Says_clientK_unique]) 4); | 
| 3745 
4c5d3b1ddc75
Client, Server certificates now sent using the separate Certificate rule,
 paulson parents: 
3729diff
changeset | 532 | (*ClientKeyExch*) | 
| 5433 | 533 | by (blast_tac (claset() addSDs [PMS_sessionK_not_spied]) 3); | 
| 534 | (*SpyKeys*) | |
| 535 | by (blast_tac (claset() addSEs [Spy_not_see_MS RSN (2,rev_notE)]) 2); | |
| 536 | (*Fake*) | |
| 537 | by (spy_analz_tac 1); | |
| 538 | qed "clientK_not_spied"; | |
| 3704 | 539 | |
| 540 | ||
| 541 | (*** Weakening the Oops conditions for leakage of serverK ***) | |
| 542 | ||
| 543 | (*If A created PMS for B, then nobody other than B or the Spy would | |
| 544 | send a message using a serverK generated from that PMS.*) | |
| 5433 | 545 | Goal "[| Says B' A' (Crypt (serverK(Na,Nb,PRF(PMS,NA,NB))) Y) : set evs; \ | 
| 546 | \        Notes A {|Agent B, Nonce PMS|} : set evs;  \
 | |
| 547 | \ evs : tls; A ~: bad; B ~: bad; B' ~= Spy |] \ | |
| 548 | \ ==> B = B'"; | |
| 549 | by (etac rev_mp 1); | |
| 550 | by (etac rev_mp 1); | |
| 3745 
4c5d3b1ddc75
Client, Server certificates now sent using the separate Certificate rule,
 paulson parents: 
3729diff
changeset | 551 | by (analz_induct_tac 1); (*9 seconds*) | 
| 3711 | 552 | by (ALLGOALS Clarify_tac); | 
| 3704 | 553 | (*ServerResume, ServerFinished: by unicity of PMS*) | 
| 554 | by (REPEAT | |
| 5433 | 555 | (blast_tac (claset() addSDs [Notes_master_imp_Crypt_PMS] | 
| 4201 | 556 | addDs [Spy_not_see_PMS, | 
| 557 | Notes_Crypt_parts_spies, | |
| 558 | Crypt_unique_PMS]) 2)); | |
| 3745 
4c5d3b1ddc75
Client, Server certificates now sent using the separate Certificate rule,
 paulson parents: 
3729diff
changeset | 559 | (*ClientKeyExch*) | 
| 4472 | 560 | by (blast_tac (claset() addSDs [PMS_Crypt_sessionK_not_spied, | 
| 561 | Says_imp_spies RS parts.Inj]) 1); | |
| 5433 | 562 | qed "Says_serverK_unique"; | 
| 3704 | 563 | |
| 564 | (*If A created PMS for B, and B has not leaked his serverK to the Spy, | |
| 5433 | 565 | then it is completely secure: not even in parts!*) | 
| 566 | Goal "[| Notes A {|Agent B, Nonce PMS|} : set evs;  \
 | |
| 567 | \ Says B Spy (Key(serverK(Na,Nb,PRF(PMS,NA,NB)))) ~: set evs; \ | |
| 568 | \ evs : tls; A ~: bad; B ~: bad |] \ | |
| 569 | \ ==> Key (serverK(Na,Nb,PRF(PMS,NA,NB))) ~: parts (spies evs)"; | |
| 570 | by (etac rev_mp 1); | |
| 571 | by (etac rev_mp 1); | |
| 572 | by (analz_induct_tac 1); (*6 seconds*) | |
| 3704 | 573 | (*Oops*) | 
| 5433 | 574 | by (blast_tac (claset() addIs [Says_serverK_unique]) 4); | 
| 3745 
4c5d3b1ddc75
Client, Server certificates now sent using the separate Certificate rule,
 paulson parents: 
3729diff
changeset | 575 | (*ClientKeyExch*) | 
| 5433 | 576 | by (blast_tac (claset() addSDs [PMS_sessionK_not_spied]) 3); | 
| 577 | (*SpyKeys*) | |
| 578 | by (blast_tac (claset() addSEs [Spy_not_see_MS RSN (2,rev_notE)]) 2); | |
| 579 | (*Fake*) | |
| 580 | by (spy_analz_tac 1); | |
| 581 | qed "serverK_not_spied"; | |
| 3704 | 582 | |
| 583 | ||
| 3685 
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
 paulson parents: 
3683diff
changeset | 584 | (*** 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: 
3711diff
changeset | 585 | 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: 
3711diff
changeset | 586 | to compare PA with what she originally sent. | 
| 3474 | 587 | ***) | 
| 588 | ||
| 3685 
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
 paulson parents: 
3683diff
changeset | 589 | (*The mention of her name (A) in X assures A that B knows who she is.*) | 
| 5114 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 590 | Goal "[| X = Crypt (serverK(Na,Nb,M)) \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 591 | \              (Hash{|Number SID, Nonce M,             \
 | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 592 | \ Nonce Na, Number PA, Agent A, \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 593 | \ Nonce Nb, Number PB, Agent B|}); \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 594 | \ M = PRF(PMS,NA,NB); \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 595 | \ evs : tls; A ~: bad; B ~: bad |] \ | 
| 5433 | 596 | \ ==> Says B Spy (Key(serverK(Na,Nb,M))) ~: set evs --> \ | 
| 5114 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 597 | \         Notes A {|Agent B, Nonce PMS|} : set evs --> \
 | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 598 | \ X : parts (spies evs) --> Says B A X : set evs"; | 
| 3480 
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
 paulson parents: 
3474diff
changeset | 599 | by (hyp_subst_tac 1); | 
| 5114 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 600 | by (analz_induct_tac 1); (*10 seconds*) | 
| 4091 | 601 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [all_conj_distrib]))); | 
| 3711 | 602 | by (ALLGOALS Clarify_tac); | 
| 4472 | 603 | (*ClientKeyExch*) | 
| 604 | by (blast_tac (claset() addSDs [PMS_Crypt_sessionK_not_spied]) 2); | |
| 3480 
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
 paulson parents: 
3474diff
changeset | 605 | (*Fake: the Spy doesn't have the critical session key!*) | 
| 5433 | 606 | by (blast_tac (claset() addEs [serverK_not_spied RSN (2,rev_notE)]) 1); | 
| 3474 | 607 | qed_spec_mp "TrustServerFinished"; | 
| 608 | ||
| 609 | ||
| 3685 
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
 paulson parents: 
3683diff
changeset | 610 | (*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: 
3683diff
changeset | 611 | 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: 
3506diff
changeset | 612 | 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: 
3515diff
changeset | 613 | that B sends his message to A. If CLIENT KEY EXCHANGE were augmented | 
| 3704 | 614 | to bind A's identity with PMS, then we could replace A' by A below.*) | 
| 5114 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 615 | Goal "[| M = PRF(PMS,NA,NB); evs : tls; A ~: bad; B ~: bad |] \ | 
| 5433 | 616 | \ ==> Says B Spy (Key(serverK(Na,Nb,M))) ~: set evs --> \ | 
| 5114 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 617 | \         Notes A {|Agent B, Nonce PMS|} : set evs -->              \
 | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 618 | \ Crypt (serverK(Na,Nb,M)) Y : parts (spies evs) --> \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 619 | \ (EX A'. Says B A' (Crypt (serverK(Na,Nb,M)) Y) : set evs)"; | 
| 3672 
56e4365a0c99
TLS now with a distinction between premaster secret and master secret
 paulson parents: 
3519diff
changeset | 620 | by (hyp_subst_tac 1); | 
| 3686 
4b484805b4c4
First working version with Oops event for session keys
 paulson parents: 
3685diff
changeset | 621 | by (analz_induct_tac 1); (*20 seconds*) | 
| 4091 | 622 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [ex_disj_distrib]))); | 
| 3711 | 623 | by (ALLGOALS Clarify_tac); | 
| 3704 | 624 | (*ServerResume, ServerFinished: by unicity of PMS*) | 
| 625 | by (REPEAT | |
| 5433 | 626 | (blast_tac (claset() addSDs [Notes_master_imp_Crypt_PMS] | 
| 4201 | 627 | addDs [Spy_not_see_PMS, | 
| 628 | Notes_Crypt_parts_spies, | |
| 629 | Crypt_unique_PMS]) 3)); | |
| 3745 
4c5d3b1ddc75
Client, Server certificates now sent using the separate Certificate rule,
 paulson parents: 
3729diff
changeset | 630 | (*ClientKeyExch*) | 
| 4472 | 631 | by (blast_tac (claset() addSDs [PMS_Crypt_sessionK_not_spied]) 2); | 
| 3515 
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
 paulson parents: 
3506diff
changeset | 632 | (*Fake: the Spy doesn't have the critical session key!*) | 
| 5433 | 633 | by (blast_tac (claset() addEs [serverK_not_spied RSN (2,rev_notE)]) 1); | 
| 3515 
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
 paulson parents: 
3506diff
changeset | 634 | qed_spec_mp "TrustServerMsg"; | 
| 
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
 paulson parents: 
3506diff
changeset | 635 | |
| 
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
 paulson parents: 
3506diff
changeset | 636 | |
| 
d8a71f6eaf40
Now uses the Notes constructor to distinguish the Client (who has chosen M)
 paulson parents: 
3506diff
changeset | 637 | (*** 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: 
3519diff
changeset | 638 | 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: 
3506diff
changeset | 639 | 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: 
3711diff
changeset | 640 | ClientFinished, then B can then check the quoted values PA, PB, etc. | 
| 3506 | 641 | ***) | 
| 3704 | 642 | |
| 5114 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 643 | Goal "[| M = PRF(PMS,NA,NB); evs : tls; A ~: bad; B ~: bad |] \ | 
| 5433 | 644 | \ ==> Says A Spy (Key(clientK(Na,Nb,M))) ~: set evs --> \ | 
| 5359 | 645 | \         Notes A {|Agent B, Nonce PMS|} : set evs -->               \
 | 
| 646 | \ Crypt (clientK(Na,Nb,M)) Y : parts (spies evs) --> \ | |
| 647 | \ Says A B (Crypt (clientK(Na,Nb,M)) Y) : set evs"; | |
| 3772 | 648 | by (hyp_subst_tac 1); | 
| 3745 
4c5d3b1ddc75
Client, Server certificates now sent using the separate Certificate rule,
 paulson parents: 
3729diff
changeset | 649 | by (analz_induct_tac 1); (*15 seconds*) | 
| 3711 | 650 | by (ALLGOALS Clarify_tac); | 
| 3704 | 651 | (*ClientFinished, ClientResume: by unicity of PMS*) | 
| 4091 | 652 | by (REPEAT (blast_tac (claset() delrules [conjI] | 
| 4201 | 653 | addSDs [Notes_master_imp_Notes_PMS] | 
| 654 | addDs [Notes_unique_PMS]) 3)); | |
| 4472 | 655 | (*ClientKeyExch*) | 
| 656 | by (blast_tac (claset() addSDs [PMS_Crypt_sessionK_not_spied]) 2); | |
| 3480 
d59bbf053258
More realistic model: the Spy can compute clientK and serverK
 paulson parents: 
3474diff
changeset | 657 | (*Fake: the Spy doesn't have the critical session key!*) | 
| 5433 | 658 | by (blast_tac (claset() addEs [clientK_not_spied RSN (2,rev_notE)]) 1); | 
| 659 | qed_spec_mp "TrustClientMsg"; | |
| 3506 | 660 | |
| 661 | ||
| 3685 
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
 paulson parents: 
3683diff
changeset | 662 | |
| 
5b8c0c8f576e
Full version of TLS including session resumption, but no Oops
 paulson parents: 
3683diff
changeset | 663 | (*** 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: 
3683diff
changeset | 664 | 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: 
3711diff
changeset | 665 | values PA, PB, etc. Even this one requires A to be uncompromised. | 
| 3506 | 666 | ***) | 
| 5114 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 667 | Goal "[| M = PRF(PMS,NA,NB); \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 668 | \ Says A Spy (Key(clientK(Na,Nb,M))) ~: set evs;\ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 669 | \ Says A' B (Crypt (clientK(Na,Nb,M)) Y) : set evs; \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 670 | \ certificate A KA : parts (spies evs); \ | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 671 | \        Says A'' B (Crypt (invKey KA) (Hash{|nb, Agent B, Nonce PMS|}))\
 | 
| 
c729d4c299c1
Deleted leading parameters thanks to new Goal command
 paulson parents: 
5076diff
changeset | 672 | \ : set evs; \ | 
| 5359 | 673 | \ evs : tls; A ~: bad; B ~: bad |] \ | 
| 674 | \ ==> Says A B (Crypt (clientK(Na,Nb,M)) Y) : set evs"; | |
| 4091 | 675 | by (blast_tac (claset() addSIs [TrustClientMsg, UseCertVerify] | 
| 4201 | 676 | 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: 
3506diff
changeset | 677 | qed "AuthClientFinished"; | 
| 3687 
fb7d096d7884
Simplified SpyKeys to use sessionK instead of clientK and serverK
 paulson parents: 
3686diff
changeset | 678 | |
| 
fb7d096d7884
Simplified SpyKeys to use sessionK instead of clientK and serverK
 paulson parents: 
3686diff
changeset | 679 | (*22/9/97: loads in 622s, which is 10 minutes 22 seconds*) | 
| 3711 | 680 | (*24/9/97: loads in 672s, which is 11 minutes 12 seconds [stronger theorems]*) | 
| 3745 
4c5d3b1ddc75
Client, Server certificates now sent using the separate Certificate rule,
 paulson parents: 
3729diff
changeset | 681 | (*29/9/97: loads in 481s, after removing Certificate from ClientKeyExch*) | 
| 3758 
188a4fbfaf55
Exchanged the M and SID fields of the FINISHED messages to simplify proofs;
 paulson parents: 
3745diff
changeset | 682 | (*30/9/97: loads in 476s, after removing unused theorems*) | 
| 3760 
77f71f650433
Strengthened the possibility property for resumption so that it could have
 paulson parents: 
3758diff
changeset | 683 | (*30/9/97: loads in 448s, after fixing ServerResume*) | 
| 5433 | 684 | |
| 685 | (*08/9/97: loads in 189s (pike), after much reorganization, | |
| 686 | back to 621s on albatross?*) |