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