src/HOL/Auth/OtwayRees_AN.thy
author paulson
Tue, 28 Mar 2006 16:48:18 +0200
changeset 19334 96ca738055a6
parent 16417 9bc16273c2d4
child 23746 a455e69c31cc
permissions -rw-r--r--
Simplified version of Jia's filter. Now all constants are pooled, rather than relevance being compared against separate clauses. Rejects are no longer noted, and units cannot be added at the end.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2090
307ebbbec862 Abadi and Needham's variant of Otway-Rees
paulson
parents:
diff changeset
     1
(*  Title:      HOL/Auth/OtwayRees
307ebbbec862 Abadi and Needham's variant of Otway-Rees
paulson
parents:
diff changeset
     2
    ID:         $Id$
307ebbbec862 Abadi and Needham's variant of Otway-Rees
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
307ebbbec862 Abadi and Needham's variant of Otway-Rees
paulson
parents:
diff changeset
     4
    Copyright   1996  University of Cambridge
14207
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
     5
*)
2090
307ebbbec862 Abadi and Needham's variant of Otway-Rees
paulson
parents:
diff changeset
     6
14207
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
     7
header{*The Otway-Rees Protocol as Modified by Abadi and Needham*}
2090
307ebbbec862 Abadi and Needham's variant of Otway-Rees
paulson
parents:
diff changeset
     8
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 14238
diff changeset
     9
theory OtwayRees_AN imports Public begin
14207
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
    10
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
    11
text{*
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
    12
This simplified version has minimal encryption and explicit messages.
2090
307ebbbec862 Abadi and Needham's variant of Otway-Rees
paulson
parents:
diff changeset
    13
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
    14
Note that the formalization does not even assume that nonces are fresh.
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
    15
This is because the protocol does not rely on uniqueness of nonces for
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
    16
security, only for freshness, and the proof script does not prove freshness
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
    17
properties.
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
    18
2090
307ebbbec862 Abadi and Needham's variant of Otway-Rees
paulson
parents:
diff changeset
    19
From page 11 of
14207
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
    20
  Abadi and Needham (1996).  
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
    21
  Prudent Engineering Practice for Cryptographic Protocols.
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
    22
  IEEE Trans. SE 22 (1)
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
    23
*}
2090
307ebbbec862 Abadi and Needham's variant of Otway-Rees
paulson
parents:
diff changeset
    24
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
    25
consts  otway   :: "event list set"
3519
ab0a9fbed4c0 Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents: 3466
diff changeset
    26
inductive "otway"
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
    27
  intros
14238
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
    28
   Nil: --{*The empty trace*}
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
    29
        "[] \<in> otway"
2090
307ebbbec862 Abadi and Needham's variant of Otway-Rees
paulson
parents:
diff changeset
    30
14238
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
    31
   Fake: --{*The Spy may say anything he can say.  The sender field is correct,
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
    32
            but agents don't use that information.*}
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
    33
         "[| evsf \<in> otway;  X \<in> synth (analz (knows Spy evsf)) |]
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
    34
          ==> Says Spy B X  # evsf \<in> otway"
2090
307ebbbec862 Abadi and Needham's variant of Otway-Rees
paulson
parents:
diff changeset
    35
14238
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
    36
        
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
    37
   Reception: --{*A message that has been sent can be received by the
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
    38
                  intended recipient.*}
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
    39
	      "[| evsr \<in> otway;  Says A B X \<in>set evsr |]
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
    40
               ==> Gets B X # evsr \<in> otway"
6308
76f3865a2b1d Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents: 5434
diff changeset
    41
14238
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
    42
   OR1:  --{*Alice initiates a protocol run*}
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
    43
         "evs1 \<in> otway
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
    44
          ==> Says A B {|Agent A, Agent B, Nonce NA|} # evs1 \<in> otway"
2090
307ebbbec862 Abadi and Needham's variant of Otway-Rees
paulson
parents:
diff changeset
    45
14238
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
    46
   OR2:  --{*Bob's response to Alice's message.*}
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
    47
	 "[| evs2 \<in> otway;
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
    48
             Gets B {|Agent A, Agent B, Nonce NA|} \<in>set evs2 |]
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
    49
          ==> Says B Server {|Agent A, Agent B, Nonce NA, Nonce NB|}
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
    50
                 # evs2 \<in> otway"
2090
307ebbbec862 Abadi and Needham's variant of Otway-Rees
paulson
parents:
diff changeset
    51
14238
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
    52
   OR3:  --{*The Server receives Bob's message.  Then he sends a new
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
    53
           session key to Bob with a packet for forwarding to Alice.*}
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
    54
	 "[| evs3 \<in> otway;  Key KAB \<notin> used evs3;
6308
76f3865a2b1d Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents: 5434
diff changeset
    55
             Gets Server {|Agent A, Agent B, Nonce NA, Nonce NB|}
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
    56
               \<in>set evs3 |]
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
    57
          ==> Says Server B
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
    58
               {|Crypt (shrK A) {|Nonce NA, Agent A, Agent B, Key KAB|},
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2451
diff changeset
    59
                 Crypt (shrK B) {|Nonce NB, Agent A, Agent B, Key KAB|}|}
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
    60
              # evs3 \<in> otway"
2090
307ebbbec862 Abadi and Needham's variant of Otway-Rees
paulson
parents:
diff changeset
    61
14238
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
    62
   OR4:  --{*Bob receives the Server's (?) message and compares the Nonces with
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
    63
	     those in the message he previously sent the Server.
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
    64
             Need @{term "B \<noteq> Server"} because we allow messages to self.*}
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
    65
	 "[| evs4 \<in> otway;  B \<noteq> Server;
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
    66
             Says B Server {|Agent A, Agent B, Nonce NA, Nonce NB|} \<in>set evs4;
6308
76f3865a2b1d Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents: 5434
diff changeset
    67
             Gets B {|X, Crypt(shrK B){|Nonce NB,Agent A,Agent B,Key K|}|}
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
    68
               \<in>set evs4 |]
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
    69
          ==> Says B A X # evs4 \<in> otway"
2090
307ebbbec862 Abadi and Needham's variant of Otway-Rees
paulson
parents:
diff changeset
    70
14238
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
    71
   Oops: --{*This message models possible leaks of session keys.  The nonces
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
    72
             identify the protocol run.*}
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
    73
	 "[| evso \<in> otway;
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
    74
             Says Server B
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
    75
                      {|Crypt (shrK A) {|Nonce NA, Agent A, Agent B, Key K|},
2284
80ebd1a213fd Swapped arguments of Crypt (for clarity and because it is conventional)
paulson
parents: 2131
diff changeset
    76
                        Crypt (shrK B) {|Nonce NB, Agent A, Agent B, Key K|}|}
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
    77
               \<in>set evso |]
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
    78
          ==> Notes Spy {|Nonce NA, Nonce NB, Key K|} # evso \<in> otway"
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
    79
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
    80
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
    81
declare Says_imp_knows_Spy [THEN analz.Inj, dest]
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
    82
declare parts.Body  [dest]
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
    83
declare analz_into_parts [dest]
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
    84
declare Fake_parts_insert_in_Un  [dest]
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
    85
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
    86
14238
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
    87
text{*A "possibility property": there are traces that reach the end*}
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 13507
diff changeset
    88
lemma "[| B \<noteq> Server; Key K \<notin> used [] |]
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 13507
diff changeset
    89
      ==> \<exists>evs \<in> otway.
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
    90
           Says B A (Crypt (shrK A) {|Nonce NA, Agent A, Agent B, Key K|})
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
    91
             \<in> set evs"
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
    92
apply (intro exI bexI)
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
    93
apply (rule_tac [2] otway.Nil
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
    94
                    [THEN otway.OR1, THEN otway.Reception,
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
    95
                     THEN otway.OR2, THEN otway.Reception,
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 13507
diff changeset
    96
                     THEN otway.OR3, THEN otway.Reception, THEN otway.OR4])
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 13507
diff changeset
    97
apply (possibility, simp add: used_Cons) 
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
    98
done
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
    99
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   100
lemma Gets_imp_Says [dest!]:
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   101
     "[| Gets B X \<in> set evs; evs \<in> otway |] ==> \<exists>A. Says A B X \<in> set evs"
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   102
by (erule rev_mp, erule otway.induct, auto)
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   103
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   104
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   105
14238
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
   106
text{* For reasoning about the encrypted portion of messages *}
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   107
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   108
lemma OR4_analz_knows_Spy:
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   109
     "[| Gets B {|X, Crypt(shrK B) X'|} \<in> set evs;  evs \<in> otway |]
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   110
      ==> X \<in> analz (knows Spy evs)"
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   111
by blast
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   112
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   113
14238
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
   114
text{*Theorems of the form @{term "X \<notin> parts (spies evs)"} imply that
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
   115
NOBODY sends messages containing X! *}
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   116
14238
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
   117
text{*Spy never sees a good agent's shared key!*}
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   118
lemma Spy_see_shrK [simp]:
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   119
     "evs \<in> otway ==> (Key (shrK A) \<in> parts (knows Spy evs)) = (A \<in> bad)"
14238
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
   120
by (erule otway.induct, simp_all, blast+)
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   121
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   122
lemma Spy_analz_shrK [simp]:
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   123
     "evs \<in> otway ==> (Key (shrK A) \<in> analz (knows Spy evs)) = (A \<in> bad)"
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   124
by auto
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   125
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   126
lemma Spy_see_shrK_D [dest!]:
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   127
     "[|Key (shrK A) \<in> parts (knows Spy evs);  evs \<in> otway|] ==> A \<in> bad"
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   128
by (blast dest: Spy_see_shrK)
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   129
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   130
14238
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
   131
subsection{*Proofs involving analz *}
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   132
14238
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
   133
text{*Describes the form of K and NA when the Server sends this message.*}
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   134
lemma Says_Server_message_form:
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   135
     "[| Says Server B
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   136
            {|Crypt (shrK A) {|NA, Agent A, Agent B, Key K|},
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   137
              Crypt (shrK B) {|NB, Agent A, Agent B, Key K|}|}
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   138
           \<in> set evs;
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   139
         evs \<in> otway |]
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   140
      ==> K \<notin> range shrK & (\<exists>i. NA = Nonce i) & (\<exists>j. NB = Nonce j)"
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   141
apply (erule rev_mp)
14207
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
   142
apply (erule otway.induct, auto)
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   143
done
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   144
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   145
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   146
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   147
(****
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   148
 The following is to prove theorems of the form
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   149
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   150
  Key K \<in> analz (insert (Key KAB) (knows Spy evs)) ==>
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   151
  Key K \<in> analz (knows Spy evs)
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   152
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   153
 A more general formula must be proved inductively.
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   154
****)
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   155
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   156
14238
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
   157
text{* Session keys are not used to encrypt other session keys *}
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   158
14238
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
   159
text{*The equality makes the induction hypothesis easier to apply*}
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   160
lemma analz_image_freshK [rule_format]:
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   161
 "evs \<in> otway ==>
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   162
   \<forall>K KK. KK <= -(range shrK) -->
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   163
          (Key K \<in> analz (Key`KK Un (knows Spy evs))) =
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   164
          (K \<in> KK | Key K \<in> analz (knows Spy evs))"
14207
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
   165
apply (erule otway.induct) 
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
   166
apply (frule_tac [8] Says_Server_message_form)
f20fbb141673 Conversion of all main protocols from "Shared" to "Public".
paulson
parents: 14200
diff changeset
   167
apply (drule_tac [7] OR4_analz_knows_Spy, analz_freshK, spy_analz, auto) 
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   168
done
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   169
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   170
lemma analz_insert_freshK:
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   171
  "[| evs \<in> otway;  KAB \<notin> range shrK |] ==>
11655
923e4d0d36d5 tuned parentheses in relational expressions;
wenzelm
parents: 11251
diff changeset
   172
      (Key K \<in> analz (insert (Key KAB) (knows Spy evs))) =
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   173
      (K = KAB | Key K \<in> analz (knows Spy evs))"
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   174
by (simp only: analz_image_freshK analz_image_freshK_simps)
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   175
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   176
14238
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
   177
text{*The Key K uniquely identifies the Server's message.*}
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   178
lemma unique_session_keys:
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   179
     "[| Says Server B
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   180
          {|Crypt (shrK A) {|NA, Agent A, Agent B, K|},
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   181
            Crypt (shrK B) {|NB, Agent A, Agent B, K|}|}
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   182
         \<in> set evs;
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   183
        Says Server B'
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   184
          {|Crypt (shrK A') {|NA', Agent A', Agent B', K|},
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   185
            Crypt (shrK B') {|NB', Agent A', Agent B', K|}|}
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   186
         \<in> set evs;
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   187
        evs \<in> otway |]
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   188
     ==> A=A' & B=B' & NA=NA' & NB=NB'"
13507
febb8e5d2a9d tidying of Isar scripts
paulson
parents: 11655
diff changeset
   189
apply (erule rev_mp, erule rev_mp, erule otway.induct, simp_all)
14238
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
   190
apply blast+  --{*OR3 and OR4*}
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   191
done
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   192
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   193
14238
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
   194
subsection{*Authenticity properties relating to NA*}
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   195
14238
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
   196
text{*If the encrypted message appears then it originated with the Server!*}
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   197
lemma NA_Crypt_imp_Server_msg [rule_format]:
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   198
    "[| A \<notin> bad;  A \<noteq> B;  evs \<in> otway |]
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   199
     ==> Crypt (shrK A) {|NA, Agent A, Agent B, Key K|} \<in> parts (knows Spy evs)
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   200
       --> (\<exists>NB. Says Server B
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   201
                    {|Crypt (shrK A) {|NA, Agent A, Agent B, Key K|},
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   202
                      Crypt (shrK B) {|NB, Agent A, Agent B, Key K|}|}
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   203
                    \<in> set evs)"
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   204
apply (erule otway.induct, force)
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   205
apply (simp_all add: ex_disj_distrib)
14238
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
   206
apply blast+  --{*Fake, OR3*}
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   207
done
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   208
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   209
14238
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
   210
text{*Corollary: if A receives B's OR4 message then it originated with the
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
   211
      Server. Freshness may be inferred from nonce NA.*}
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   212
lemma A_trusts_OR4:
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   213
     "[| Says B' A (Crypt (shrK A) {|NA, Agent A, Agent B, Key K|}) \<in> set evs;
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   214
         A \<notin> bad;  A \<noteq> B;  evs \<in> otway |]
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   215
      ==> \<exists>NB. Says Server B
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   216
                  {|Crypt (shrK A) {|NA, Agent A, Agent B, Key K|},
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   217
                    Crypt (shrK B) {|NB, Agent A, Agent B, Key K|}|}
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   218
                 \<in> set evs"
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   219
by (blast intro!: NA_Crypt_imp_Server_msg)
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   220
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   221
14238
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
   222
text{*Crucial secrecy property: Spy does not see the keys sent in msg OR3
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   223
    Does not in itself guarantee security: an attack could violate
14238
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
   224
    the premises, e.g. by having @{term "A=Spy"}*}
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   225
lemma secrecy_lemma:
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   226
     "[| A \<notin> bad;  B \<notin> bad;  evs \<in> otway |]
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   227
      ==> Says Server B
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   228
           {|Crypt (shrK A) {|NA, Agent A, Agent B, Key K|},
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   229
             Crypt (shrK B) {|NB, Agent A, Agent B, Key K|}|}
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   230
          \<in> set evs -->
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   231
          Notes Spy {|NA, NB, Key K|} \<notin> set evs -->
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   232
          Key K \<notin> analz (knows Spy evs)"
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   233
apply (erule otway.induct, force)
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   234
apply (frule_tac [7] Says_Server_message_form)
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   235
apply (drule_tac [6] OR4_analz_knows_Spy)
14238
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
   236
apply (simp_all add: analz_insert_eq analz_insert_freshK pushes)
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
   237
apply spy_analz  --{*Fake*}
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
   238
apply (blast dest: unique_session_keys)+  --{*OR3, OR4, Oops*}
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   239
done
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   240
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   241
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   242
lemma Spy_not_see_encrypted_key:
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   243
     "[| Says Server B
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   244
            {|Crypt (shrK A) {|NA, Agent A, Agent B, Key K|},
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   245
              Crypt (shrK B) {|NB, Agent A, Agent B, Key K|}|}
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   246
           \<in> set evs;
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   247
         Notes Spy {|NA, NB, Key K|} \<notin> set evs;
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   248
         A \<notin> bad;  B \<notin> bad;  evs \<in> otway |]
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   249
      ==> Key K \<notin> analz (knows Spy evs)"
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   250
by (blast dest: Says_Server_message_form secrecy_lemma)
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   252
14238
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
   253
text{*A's guarantee.  The Oops premise quantifies over NB because A cannot know
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
   254
  what it is.*}
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   255
lemma A_gets_good_key:
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   256
     "[| Says B' A (Crypt (shrK A) {|NA, Agent A, Agent B, Key K|}) \<in> set evs;
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   257
         \<forall>NB. Notes Spy {|NA, NB, Key K|} \<notin> set evs;
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   258
         A \<notin> bad;  B \<notin> bad;  A \<noteq> B;  evs \<in> otway |]
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   259
      ==> Key K \<notin> analz (knows Spy evs)"
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   260
by (blast dest!: A_trusts_OR4 Spy_not_see_encrypted_key)
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   261
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   262
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   263
14238
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
   264
subsection{*Authenticity properties relating to NB*}
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   265
14238
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
   266
text{*If the encrypted message appears then it originated with the Server!*}
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   267
lemma NB_Crypt_imp_Server_msg [rule_format]:
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   268
 "[| B \<notin> bad;  A \<noteq> B;  evs \<in> otway |]
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   269
  ==> Crypt (shrK B) {|NB, Agent A, Agent B, Key K|} \<in> parts (knows Spy evs)
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   270
      --> (\<exists>NA. Says Server B
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   271
                   {|Crypt (shrK A) {|NA, Agent A, Agent B, Key K|},
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   272
                     Crypt (shrK B) {|NB, Agent A, Agent B, Key K|}|}
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   273
                   \<in> set evs)"
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   274
apply (erule otway.induct, force, simp_all add: ex_disj_distrib)
14238
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
   275
apply blast+  --{*Fake, OR3*}
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   276
done
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   277
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   278
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   279
14238
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
   280
text{*Guarantee for B: if it gets a well-formed certificate then the Server
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
   281
  has sent the correct message in round 3.*}
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   282
lemma B_trusts_OR3:
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   283
     "[| Says S B {|X, Crypt (shrK B) {|NB, Agent A, Agent B, Key K|}|}
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   284
           \<in> set evs;
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   285
         B \<notin> bad;  A \<noteq> B;  evs \<in> otway |]
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   286
      ==> \<exists>NA. Says Server B
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   287
                   {|Crypt (shrK A) {|NA, Agent A, Agent B, Key K|},
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   288
                     Crypt (shrK B) {|NB, Agent A, Agent B, Key K|}|}
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   289
                   \<in> set evs"
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   290
by (blast intro!: NB_Crypt_imp_Server_msg)
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   291
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   292
14238
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
   293
text{*The obvious combination of @{text B_trusts_OR3} with 
59b02c1efd01 improved presentation
paulson
parents: 14207
diff changeset
   294
      @{text Spy_not_see_encrypted_key}*}
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   295
lemma B_gets_good_key:
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   296
     "[| Gets B {|X, Crypt (shrK B) {|NB, Agent A, Agent B, Key K|}|}
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   297
          \<in> set evs;
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   298
         \<forall>NA. Notes Spy {|NA, NB, Key K|} \<notin> set evs;
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   299
         A \<notin> bad;  B \<notin> bad;  A \<noteq> B;  evs \<in> otway |]
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   300
      ==> Key K \<notin> analz (knows Spy evs)"
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   301
by (blast dest: B_trusts_OR3 Spy_not_see_encrypted_key)
2090
307ebbbec862 Abadi and Needham's variant of Otway-Rees
paulson
parents:
diff changeset
   302
307ebbbec862 Abadi and Needham's variant of Otway-Rees
paulson
parents:
diff changeset
   303
end